| author | wenzelm | 
| Tue, 27 Apr 2010 16:09:15 +0200 | |
| changeset 36421 | 066e35d1c0d7 | 
| parent 33769 | 6d8630fab26a | 
| child 36744 | 6e1f3d609a68 | 
| permissions | -rw-r--r-- | 
| 
7765
 
fa28bac7903c
New function thm_deps for visualizing dependencies of theorems.
 
berghofe 
parents:  
diff
changeset
 | 
1  | 
(* Title: Pure/Thy/thm_deps.ML  | 
| 
 
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  | 
| 26697 | 9  | 
val thm_deps: thm list -> unit  | 
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 | 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 | 16  | 
(* thm_deps *)  | 
17  | 
||
| 
7765
 
fa28bac7903c
New function thm_deps for visualizing dependencies of theorems.
 
berghofe 
parents:  
diff
changeset
 | 
18  | 
fun thm_deps thms =  | 
| 28810 | 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  | 
| 
30364
 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
23  | 
val prefix = #1 (Library.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 :: _ =>  | 
| 
 
3b460b6eadae
simplified thm_deps -- no need to build a graph datastructure;
 
wenzelm 
parents: 
28817 
diff
changeset
 | 
27  | 
(case ThyInfo.lookup_theory a of  | 
| 
 
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,
 | 
| 
28826
 
3b460b6eadae
simplified thm_deps -- no need to build a graph datastructure;
 
wenzelm 
parents: 
28817 
diff
changeset
 | 
37  | 
ID = name,  | 
| 
 
3b460b6eadae
simplified thm_deps -- no need to build a graph datastructure;
 
wenzelm 
parents: 
28817 
diff
changeset
 | 
38  | 
dir = space_implode "/" (session @ prefix),  | 
| 
 
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 = "",  | 
| 
 
3b460b6eadae
simplified thm_deps -- no need to build a graph datastructure;
 
wenzelm 
parents: 
28817 
diff
changeset
 | 
41  | 
parents = parents};  | 
| 
 
3b460b6eadae
simplified thm_deps -- no need to build a graph datastructure;
 
wenzelm 
parents: 
28817 
diff
changeset
 | 
42  | 
in cons entry end;  | 
| 
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
 | 
43  | 
val deps = Proofterm.fold_body_thms add_dep (map Thm.proof_body_of thms) [];  | 
| 
28826
 
3b460b6eadae
simplified thm_deps -- no need to build a graph datastructure;
 
wenzelm 
parents: 
28817 
diff
changeset
 | 
44  | 
in Present.display_graph (sort_wrt #ID deps) end;  | 
| 26697 | 45  | 
|
46  | 
||
47  | 
(* unused_thms *)  | 
|
48  | 
||
49  | 
fun unused_thms (base_thys, thys) =  | 
|
| 26185 | 50  | 
let  | 
| 33170 | 51  | 
fun add_fact space (name, ths) =  | 
| 26697 | 52  | 
if exists (fn thy => PureThy.defined_fact thy name) base_thys then I  | 
| 33170 | 53  | 
else  | 
54  | 
        let val {concealed, group, ...} = Name_Space.the_entry space name in
 | 
|
55  | 
fold_rev (fn th =>  | 
|
56  | 
(case Thm.get_name th of  | 
|
57  | 
"" => I  | 
|
58  | 
| a => cons (a, (th, concealed, group)))) ths  | 
|
59  | 
end;  | 
|
60  | 
fun add_facts facts = Facts.fold_static (add_fact (Facts.space_of facts)) facts;  | 
|
61  | 
||
| 
33769
 
6d8630fab26a
unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
 
wenzelm 
parents: 
33642 
diff
changeset
 | 
62  | 
val new_thms =  | 
| 33170 | 63  | 
fold (add_facts o PureThy.facts_of) thys []  | 
| 26699 | 64  | 
|> sort_distinct (string_ord o pairself #1);  | 
| 28810 | 65  | 
|
| 
33769
 
6d8630fab26a
unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
 
wenzelm 
parents: 
33642 
diff
changeset
 | 
66  | 
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
 | 
67  | 
Proofterm.fold_body_thms  | 
| 33170 | 68  | 
(fn (a, prop, _) => a <> "" ? Symtab.insert_list (op =) (a, prop))  | 
| 
33769
 
6d8630fab26a
unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
 
wenzelm 
parents: 
33642 
diff
changeset
 | 
69  | 
(map (Proofterm.strip_thm o Thm.proof_body_of o #1 o #2) new_thms) Symtab.empty;  | 
| 33170 | 70  | 
|
71  | 
fun is_unused (a, th) =  | 
|
| 
33769
 
6d8630fab26a
unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
 
wenzelm 
parents: 
33642 
diff
changeset
 | 
72  | 
not (member (op aconv) (Symtab.lookup_list used a) (Thm.prop_of th));  | 
| 28810 | 73  | 
|
| 26185 | 74  | 
(* groups containing at least one used theorem *)  | 
| 33170 | 75  | 
val used_groups = fold (fn (a, (th, _, group)) =>  | 
76  | 
if is_unused (a, th) then I  | 
|
77  | 
else  | 
|
78  | 
(case group of  | 
|
79  | 
NONE => I  | 
|
| 
33769
 
6d8630fab26a
unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
 
wenzelm 
parents: 
33642 
diff
changeset
 | 
80  | 
| 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
 | 
81  | 
|
| 
 
6d8630fab26a
unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
 
wenzelm 
parents: 
33642 
diff
changeset
 | 
82  | 
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
 | 
83  | 
if not concealed andalso  | 
| 
 
6d8630fab26a
unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
 
wenzelm 
parents: 
33642 
diff
changeset
 | 
84  | 
member (op =) [Thm.theoremK, Thm.lemmaK, Thm.corollaryK] (Thm.get_kind th) andalso  | 
| 
 
6d8630fab26a
unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
 
wenzelm 
parents: 
33642 
diff
changeset
 | 
85  | 
is_unused (a, th)  | 
| 
 
6d8630fab26a
unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
 
wenzelm 
parents: 
33642 
diff
changeset
 | 
86  | 
then  | 
| 33170 | 87  | 
(case group of  | 
| 
33769
 
6d8630fab26a
unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
 
wenzelm 
parents: 
33642 
diff
changeset
 | 
88  | 
NONE => ((a, th) :: thms, seen_groups)  | 
| 26185 | 89  | 
| SOME grp =>  | 
| 
33769
 
6d8630fab26a
unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
 
wenzelm 
parents: 
33642 
diff
changeset
 | 
90  | 
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
 | 
91  | 
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
 | 
92  | 
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
 | 
93  | 
else q) new_thms ([], Inttab.empty);  | 
| 26185 | 94  | 
in rev thms' end;  | 
95  | 
||
| 
7765
 
fa28bac7903c
New function thm_deps for visualizing dependencies of theorems.
 
berghofe 
parents:  
diff
changeset
 | 
96  | 
end;  | 
| 
 
fa28bac7903c
New function thm_deps for visualizing dependencies of theorems.
 
berghofe 
parents:  
diff
changeset
 | 
97  |