| author | Cezary Kaliszyk <cezarykaliszyk@gmail.com> | 
| Fri, 03 Feb 2012 15:51:10 +0100 | |
| changeset 46404 | 7736068b9f56 | 
| parent 45269 | 6f8949e6c71a | 
| child 46728 | 85f8e3932712 | 
| permissions | -rw-r--r-- | 
| 17369 | 1  | 
(* Title: Pure/Isar/proof_display.ML  | 
2  | 
Author: Makarius  | 
|
3  | 
||
| 
19432
 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 
wenzelm 
parents: 
18799 
diff
changeset
 | 
4  | 
Printing of theorems, goals, results etc.  | 
| 17369 | 5  | 
*)  | 
6  | 
||
7  | 
signature PROOF_DISPLAY =  | 
|
8  | 
sig  | 
|
| 30628 | 9  | 
val pp_context: Proof.context -> Pretty.T  | 
10  | 
val pp_thm: thm -> Pretty.T  | 
|
11  | 
val pp_typ: theory -> typ -> Pretty.T  | 
|
12  | 
val pp_term: theory -> term -> Pretty.T  | 
|
13  | 
val pp_ctyp: ctyp -> Pretty.T  | 
|
14  | 
val pp_cterm: cterm -> Pretty.T  | 
|
| 
33515
 
d066e8369a33
print_theorems: suppress concealed (global) facts, unless "!" option is given;
 
wenzelm 
parents: 
33389 
diff
changeset
 | 
15  | 
val print_theorems_diff: bool -> theory -> theory -> unit  | 
| 
 
d066e8369a33
print_theorems: suppress concealed (global) facts, unless "!" option is given;
 
wenzelm 
parents: 
33389 
diff
changeset
 | 
16  | 
val print_theorems: bool -> theory -> unit  | 
| 22335 | 17  | 
val pretty_full_theory: bool -> theory -> Pretty.T  | 
| 28092 | 18  | 
val print_theory: theory -> unit  | 
| 20311 | 19  | 
val string_of_rule: Proof.context -> string -> thm -> string  | 
| 27857 | 20  | 
val print_results: bool -> Proof.context -> ((string * string) * (string * thm list) list) -> unit  | 
| 
44192
 
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
 
wenzelm 
parents: 
42717 
diff
changeset
 | 
21  | 
val print_consts: bool -> Proof.context -> (string * typ -> bool) -> (string * typ) list -> unit  | 
| 17369 | 22  | 
end  | 
23  | 
||
| 33389 | 24  | 
structure Proof_Display: PROOF_DISPLAY =  | 
| 17369 | 25  | 
struct  | 
26  | 
||
| 20235 | 27  | 
(* toplevel pretty printing *)  | 
| 20211 | 28  | 
|
| 30628 | 29  | 
fun pp_context ctxt =  | 
| 
42717
 
0bbb56867091
proper configuration options Proof_Context.debug and Proof_Context.verbose;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
30  | 
(if Config.get ctxt Proof_Context.debug then  | 
| 42360 | 31  | 
Pretty.quote (Pretty.big_list "proof context:" (Proof_Context.pretty_context ctxt))  | 
| 20211 | 32  | 
else Pretty.str "<context>");  | 
33  | 
||
| 
45269
 
6f8949e6c71a
more robust ML pretty printing (cf. b6c527c64789);
 
wenzelm 
parents: 
44240 
diff
changeset
 | 
34  | 
fun default_context thy0 =  | 
| 
 
6f8949e6c71a
more robust ML pretty printing (cf. b6c527c64789);
 
wenzelm 
parents: 
44240 
diff
changeset
 | 
35  | 
(case Context.thread_data () of  | 
| 
 
6f8949e6c71a
more robust ML pretty printing (cf. b6c527c64789);
 
wenzelm 
parents: 
44240 
diff
changeset
 | 
36  | 
SOME (Context.Proof ctxt) => ctxt  | 
| 
 
6f8949e6c71a
more robust ML pretty printing (cf. b6c527c64789);
 
wenzelm 
parents: 
44240 
diff
changeset
 | 
37  | 
| SOME (Context.Theory thy) =>  | 
| 
 
6f8949e6c71a
more robust ML pretty printing (cf. b6c527c64789);
 
wenzelm 
parents: 
44240 
diff
changeset
 | 
38  | 
(case try Syntax.init_pretty_global thy of  | 
| 
 
6f8949e6c71a
more robust ML pretty printing (cf. b6c527c64789);
 
wenzelm 
parents: 
44240 
diff
changeset
 | 
39  | 
SOME ctxt => ctxt  | 
| 
 
6f8949e6c71a
more robust ML pretty printing (cf. b6c527c64789);
 
wenzelm 
parents: 
44240 
diff
changeset
 | 
40  | 
| NONE => Syntax.init_pretty_global thy0)  | 
| 
 
6f8949e6c71a
more robust ML pretty printing (cf. b6c527c64789);
 
wenzelm 
parents: 
44240 
diff
changeset
 | 
41  | 
| NONE => Syntax.init_pretty_global thy0);  | 
| 
44240
 
4b1a63678839
improved default context for ML toplevel pretty-printing;
 
wenzelm 
parents: 
44192 
diff
changeset
 | 
42  | 
|
| 30628 | 43  | 
fun pp_thm th =  | 
| 
30724
 
2e54e89bcce7
pretty_rule/print_results: no thm status here -- it is potentially slow and mostly uninformative/confusing as long as proofs are still unfinished;
 
wenzelm 
parents: 
30628 
diff
changeset
 | 
44  | 
let  | 
| 
44240
 
4b1a63678839
improved default context for ML toplevel pretty-printing;
 
wenzelm 
parents: 
44192 
diff
changeset
 | 
45  | 
val ctxt = default_context (Thm.theory_of_thm th);  | 
| 
30724
 
2e54e89bcce7
pretty_rule/print_results: no thm status here -- it is potentially slow and mostly uninformative/confusing as long as proofs are still unfinished;
 
wenzelm 
parents: 
30628 
diff
changeset
 | 
46  | 
    val flags = {quote = true, show_hyps = false, show_status = true};
 | 
| 
44240
 
4b1a63678839
improved default context for ML toplevel pretty-printing;
 
wenzelm 
parents: 
44192 
diff
changeset
 | 
47  | 
in Display.pretty_thm_raw ctxt flags th end;  | 
| 20211 | 48  | 
|
| 
44240
 
4b1a63678839
improved default context for ML toplevel pretty-printing;
 
wenzelm 
parents: 
44192 
diff
changeset
 | 
49  | 
fun pp_typ thy T = Pretty.quote (Syntax.pretty_typ (default_context thy) T);  | 
| 
 
4b1a63678839
improved default context for ML toplevel pretty-printing;
 
wenzelm 
parents: 
44192 
diff
changeset
 | 
50  | 
fun pp_term thy t = Pretty.quote (Syntax.pretty_term (default_context thy) t);  | 
| 30628 | 51  | 
|
52  | 
fun pp_ctyp cT = pp_typ (Thm.theory_of_ctyp cT) (Thm.typ_of cT);  | 
|
53  | 
fun pp_cterm ct = pp_term (Thm.theory_of_cterm ct) (Thm.term_of ct);  | 
|
| 20211 | 54  | 
|
| 
19432
 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 
wenzelm 
parents: 
18799 
diff
changeset
 | 
55  | 
|
| 
 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 
wenzelm 
parents: 
18799 
diff
changeset
 | 
56  | 
(* theorems and theory *)  | 
| 
 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 
wenzelm 
parents: 
18799 
diff
changeset
 | 
57  | 
|
| 
33515
 
d066e8369a33
print_theorems: suppress concealed (global) facts, unless "!" option is given;
 
wenzelm 
parents: 
33389 
diff
changeset
 | 
58  | 
fun pretty_theorems_diff verbose prev_thys thy =  | 
| 
19432
 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 
wenzelm 
parents: 
18799 
diff
changeset
 | 
59  | 
let  | 
| 42360 | 60  | 
val pretty_fact = Proof_Context.pretty_fact (Proof_Context.init_global thy);  | 
| 
39557
 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 
wenzelm 
parents: 
39284 
diff
changeset
 | 
61  | 
val facts = Global_Theory.facts_of thy;  | 
| 
33515
 
d066e8369a33
print_theorems: suppress concealed (global) facts, unless "!" option is given;
 
wenzelm 
parents: 
33389 
diff
changeset
 | 
62  | 
val thmss =  | 
| 
39557
 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 
wenzelm 
parents: 
39284 
diff
changeset
 | 
63  | 
Facts.dest_static (map Global_Theory.facts_of prev_thys) facts  | 
| 
33515
 
d066e8369a33
print_theorems: suppress concealed (global) facts, unless "!" option is given;
 
wenzelm 
parents: 
33389 
diff
changeset
 | 
64  | 
|> not verbose ? filter_out (Facts.is_concealed facts o #1);  | 
| 28210 | 65  | 
in Pretty.big_list "theorems:" (map #1 (sort_wrt (#1 o #2) (map (`pretty_fact) thmss))) end;  | 
| 
19432
 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 
wenzelm 
parents: 
18799 
diff
changeset
 | 
66  | 
|
| 
33515
 
d066e8369a33
print_theorems: suppress concealed (global) facts, unless "!" option is given;
 
wenzelm 
parents: 
33389 
diff
changeset
 | 
67  | 
fun print_theorems_diff verbose prev_thy thy =  | 
| 
 
d066e8369a33
print_theorems: suppress concealed (global) facts, unless "!" option is given;
 
wenzelm 
parents: 
33389 
diff
changeset
 | 
68  | 
Pretty.writeln (pretty_theorems_diff verbose [prev_thy] thy);  | 
| 
19432
 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 
wenzelm 
parents: 
18799 
diff
changeset
 | 
69  | 
|
| 
33515
 
d066e8369a33
print_theorems: suppress concealed (global) facts, unless "!" option is given;
 
wenzelm 
parents: 
33389 
diff
changeset
 | 
70  | 
fun pretty_theorems verbose thy = pretty_theorems_diff verbose (Theory.parents_of thy) thy;  | 
| 
 
d066e8369a33
print_theorems: suppress concealed (global) facts, unless "!" option is given;
 
wenzelm 
parents: 
33389 
diff
changeset
 | 
71  | 
val print_theorems = Pretty.writeln oo pretty_theorems;  | 
| 
19432
 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 
wenzelm 
parents: 
18799 
diff
changeset
 | 
72  | 
|
| 22872 | 73  | 
fun pretty_full_theory verbose thy =  | 
| 
33515
 
d066e8369a33
print_theorems: suppress concealed (global) facts, unless "!" option is given;
 
wenzelm 
parents: 
33389 
diff
changeset
 | 
74  | 
Pretty.chunks (Display.pretty_full_theory verbose thy @ [pretty_theorems verbose thy]);  | 
| 20621 | 75  | 
|
| 22872 | 76  | 
val print_theory = Pretty.writeln o pretty_full_theory false;  | 
| 22335 | 77  | 
|
| 
19432
 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 
wenzelm 
parents: 
18799 
diff
changeset
 | 
78  | 
|
| 17369 | 79  | 
(* refinement rule *)  | 
80  | 
||
81  | 
fun pretty_rule ctxt s thm =  | 
|
82  | 
Pretty.block [Pretty.str (s ^ " attempt to solve goal by exported rule:"),  | 
|
| 
32091
 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 
wenzelm 
parents: 
30724 
diff
changeset
 | 
83  | 
Pretty.fbrk, Display.pretty_thm_aux ctxt false thm];  | 
| 17369 | 84  | 
|
85  | 
val string_of_rule = Pretty.string_of ooo pretty_rule;  | 
|
86  | 
||
87  | 
||
88  | 
(* results *)  | 
|
89  | 
||
| 28092 | 90  | 
local  | 
91  | 
||
| 41551 | 92  | 
fun pretty_fact_name (kind, "") = Pretty.command kind  | 
93  | 
| pretty_fact_name (kind, name) =  | 
|
94  | 
Pretty.block [Pretty.command kind, Pretty.brk 1,  | 
|
95  | 
Pretty.str (Long_Name.base_name name), Pretty.str ":"];  | 
|
| 
28087
 
a9cccdd9d521
pretty_fact/results: display base only, since results now come with full names (note that Facts.extern is not really well-defined unless we present the real target context);
 
wenzelm 
parents: 
27857 
diff
changeset
 | 
96  | 
|
| 17369 | 97  | 
fun pretty_facts ctxt =  | 
| 41551 | 98  | 
flat o (separate [Pretty.fbrk, Pretty.keyword "and", Pretty.str " "]) o  | 
| 42360 | 99  | 
map (single o Proof_Context.pretty_fact_aux ctxt false);  | 
| 17369 | 100  | 
|
| 27857 | 101  | 
in  | 
102  | 
||
| 28092 | 103  | 
fun print_results do_print ctxt ((kind, name), facts) =  | 
| 
33643
 
b275f26a638b
eliminated obsolete "internal" kind -- collapsed to unspecific "";
 
wenzelm 
parents: 
33515 
diff
changeset
 | 
104  | 
if not do_print orelse kind = "" then ()  | 
| 28092 | 105  | 
else if name = "" then  | 
| 39284 | 106  | 
Pretty.writeln (Pretty.block (Pretty.command kind :: Pretty.brk 1 :: pretty_facts ctxt facts))  | 
| 28092 | 107  | 
else Pretty.writeln  | 
108  | 
(case facts of  | 
|
109  | 
[fact] => Pretty.blk (1, [pretty_fact_name (kind, name), Pretty.fbrk,  | 
|
| 42360 | 110  | 
Proof_Context.pretty_fact_aux ctxt false fact])  | 
| 28092 | 111  | 
| _ => Pretty.blk (1, [pretty_fact_name (kind, name), Pretty.fbrk,  | 
112  | 
        Pretty.blk (1, Pretty.str "(" :: pretty_facts ctxt facts @ [Pretty.str ")"])]));
 | 
|
| 17369 | 113  | 
|
114  | 
end;  | 
|
115  | 
||
| 20889 | 116  | 
|
117  | 
(* consts *)  | 
|
118  | 
||
119  | 
local  | 
|
120  | 
||
121  | 
fun pretty_var ctxt (x, T) =  | 
|
122  | 
Pretty.block [Pretty.str x, Pretty.str " ::", Pretty.brk 1,  | 
|
| 24920 | 123  | 
Pretty.quote (Syntax.pretty_typ ctxt T)];  | 
| 20889 | 124  | 
|
125  | 
fun pretty_vars ctxt kind vs = Pretty.big_list kind (map (pretty_var ctxt) vs);  | 
|
126  | 
||
127  | 
fun pretty_consts ctxt pred cs =  | 
|
| 42360 | 128  | 
(case filter pred (#1 (Proof_Context.inferred_fixes ctxt)) of  | 
| 20889 | 129  | 
[] => pretty_vars ctxt "constants" cs  | 
130  | 
| ps => Pretty.chunks [pretty_vars ctxt "parameters" ps, pretty_vars ctxt "constants" cs]);  | 
|
131  | 
||
| 
44192
 
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
 
wenzelm 
parents: 
42717 
diff
changeset
 | 
132  | 
in  | 
| 
 
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
 
wenzelm 
parents: 
42717 
diff
changeset
 | 
133  | 
|
| 
 
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
 
wenzelm 
parents: 
42717 
diff
changeset
 | 
134  | 
fun print_consts do_print ctxt pred cs =  | 
| 
 
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
 
wenzelm 
parents: 
42717 
diff
changeset
 | 
135  | 
if not do_print orelse null cs then ()  | 
| 
 
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
 
wenzelm 
parents: 
42717 
diff
changeset
 | 
136  | 
else Pretty.writeln (pretty_consts ctxt pred cs);  | 
| 
 
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
 
wenzelm 
parents: 
42717 
diff
changeset
 | 
137  | 
|
| 20889 | 138  | 
end;  | 
139  | 
||
| 17369 | 140  | 
end;  |