| author | wenzelm | 
| Tue, 09 Dec 2014 18:29:45 +0100 | |
| changeset 59117 | caddfa6ca534 | 
| parent 57605 | 8e0a7eaffe47 | 
| child 59184 | 830bb7ddb3ab | 
| 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: 
18799diff
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 | |
| 57605 | 15 | val pretty_theorems_diff: bool -> theory list -> theory -> Pretty.T list | 
| 16 | val pretty_theorems: bool -> theory -> Pretty.T list | |
| 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 | 
| 49860 | 20 | val pretty_goal_header: thm -> Pretty.T | 
| 21 | val string_of_goal: Proof.context -> thm -> string | |
| 51584 | 22 | val pretty_goal_facts: Proof.context -> string -> thm list -> Pretty.T | 
| 49866 | 23 | val method_error: string -> Position.T -> | 
| 24 |     {context: Proof.context, facts: thm list, goal: thm} -> 'a Seq.result
 | |
| 56932 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 25 | val print_results: bool -> Position.T -> Proof.context -> | 
| 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 26 | ((string * string) * (string * thm list) list) -> unit | 
| 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 27 | val print_consts: bool -> Position.T -> Proof.context -> | 
| 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 28 | (string * typ -> bool) -> (string * typ) list -> unit | 
| 17369 | 29 | end | 
| 30 | ||
| 33389 | 31 | structure Proof_Display: PROOF_DISPLAY = | 
| 17369 | 32 | struct | 
| 33 | ||
| 20235 | 34 | (* toplevel pretty printing *) | 
| 20211 | 35 | |
| 30628 | 36 | fun pp_context ctxt = | 
| 42717 
0bbb56867091
proper configuration options Proof_Context.debug and Proof_Context.verbose;
 wenzelm parents: 
42360diff
changeset | 37 | (if Config.get ctxt Proof_Context.debug then | 
| 42360 | 38 | Pretty.quote (Pretty.big_list "proof context:" (Proof_Context.pretty_context ctxt)) | 
| 20211 | 39 | else Pretty.str "<context>"); | 
| 40 | ||
| 45269 
6f8949e6c71a
more robust ML pretty printing (cf. b6c527c64789);
 wenzelm parents: 
44240diff
changeset | 41 | fun default_context thy0 = | 
| 
6f8949e6c71a
more robust ML pretty printing (cf. b6c527c64789);
 wenzelm parents: 
44240diff
changeset | 42 | (case Context.thread_data () of | 
| 
6f8949e6c71a
more robust ML pretty printing (cf. b6c527c64789);
 wenzelm parents: 
44240diff
changeset | 43 | SOME (Context.Proof ctxt) => ctxt | 
| 
6f8949e6c71a
more robust ML pretty printing (cf. b6c527c64789);
 wenzelm parents: 
44240diff
changeset | 44 | | SOME (Context.Theory thy) => | 
| 
6f8949e6c71a
more robust ML pretty printing (cf. b6c527c64789);
 wenzelm parents: 
44240diff
changeset | 45 | (case try Syntax.init_pretty_global thy of | 
| 
6f8949e6c71a
more robust ML pretty printing (cf. b6c527c64789);
 wenzelm parents: 
44240diff
changeset | 46 | SOME ctxt => ctxt | 
| 
6f8949e6c71a
more robust ML pretty printing (cf. b6c527c64789);
 wenzelm parents: 
44240diff
changeset | 47 | | NONE => Syntax.init_pretty_global thy0) | 
| 
6f8949e6c71a
more robust ML pretty printing (cf. b6c527c64789);
 wenzelm parents: 
44240diff
changeset | 48 | | NONE => Syntax.init_pretty_global thy0); | 
| 44240 
4b1a63678839
improved default context for ML toplevel pretty-printing;
 wenzelm parents: 
44192diff
changeset | 49 | |
| 30628 | 50 | fun pp_thm th = | 
| 50126 
3dec88149176
theorem status about oracles/futures is no longer printed by default;
 wenzelm parents: 
49867diff
changeset | 51 | let val ctxt = default_context (Thm.theory_of_thm th); | 
| 
3dec88149176
theorem status about oracles/futures is no longer printed by default;
 wenzelm parents: 
49867diff
changeset | 52 |   in Display.pretty_thm_raw ctxt {quote = true, show_hyps = false} th end;
 | 
| 20211 | 53 | |
| 44240 
4b1a63678839
improved default context for ML toplevel pretty-printing;
 wenzelm parents: 
44192diff
changeset | 54 | 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: 
44192diff
changeset | 55 | fun pp_term thy t = Pretty.quote (Syntax.pretty_term (default_context thy) t); | 
| 30628 | 56 | |
| 57 | fun pp_ctyp cT = pp_typ (Thm.theory_of_ctyp cT) (Thm.typ_of cT); | |
| 58 | fun pp_cterm ct = pp_term (Thm.theory_of_cterm ct) (Thm.term_of ct); | |
| 20211 | 59 | |
| 19432 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 60 | |
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 61 | (* theorems and theory *) | 
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 62 | |
| 33515 
d066e8369a33
print_theorems: suppress concealed (global) facts, unless "!" option is given;
 wenzelm parents: 
33389diff
changeset | 63 | fun pretty_theorems_diff verbose prev_thys thy = | 
| 19432 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 64 | let | 
| 42360 | 65 | 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: 
39284diff
changeset | 66 | val facts = Global_Theory.facts_of thy; | 
| 56158 
c2c6d560e7b2
more explicit treatment of verbose mode, which includes concealed entries;
 wenzelm parents: 
55763diff
changeset | 67 | val thmss = Facts.dest_static verbose (map Global_Theory.facts_of prev_thys) facts; | 
| 57605 | 68 | val prts = map #1 (sort_wrt (#1 o #2) (map (`pretty_fact) thmss)); | 
| 69 | in if null prts then [] else [Pretty.big_list "theorems:" prts] end; | |
| 19432 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 70 | |
| 56868 
b5fb264d53ba
clarified print operations for "terms" and "theorems";
 wenzelm parents: 
56158diff
changeset | 71 | fun pretty_theorems verbose thy = | 
| 
b5fb264d53ba
clarified print operations for "terms" and "theorems";
 wenzelm parents: 
56158diff
changeset | 72 | pretty_theorems_diff verbose (Theory.parents_of thy) thy; | 
| 19432 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 73 | |
| 22872 | 74 | fun pretty_full_theory verbose thy = | 
| 57605 | 75 | Pretty.chunks (Display.pretty_full_theory verbose thy @ pretty_theorems verbose thy); | 
| 20621 | 76 | |
| 22872 | 77 | val print_theory = Pretty.writeln o pretty_full_theory false; | 
| 22335 | 78 | |
| 19432 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 79 | |
| 17369 | 80 | (* refinement rule *) | 
| 81 | ||
| 82 | fun pretty_rule ctxt s thm = | |
| 83 | Pretty.block [Pretty.str (s ^ " attempt to solve goal by exported rule:"), | |
| 50126 
3dec88149176
theorem status about oracles/futures is no longer printed by default;
 wenzelm parents: 
49867diff
changeset | 84 | Pretty.fbrk, Display.pretty_thm ctxt thm]; | 
| 17369 | 85 | |
| 86 | val string_of_rule = Pretty.string_of ooo pretty_rule; | |
| 87 | ||
| 88 | ||
| 49860 | 89 | (* goals *) | 
| 90 | ||
| 91 | local | |
| 92 | ||
| 93 | fun subgoals 0 = [] | |
| 94 | | subgoals 1 = [Pretty.brk 1, Pretty.str "(1 subgoal)"] | |
| 95 |   | subgoals n = [Pretty.brk 1, Pretty.str ("(" ^ string_of_int n ^ " subgoals)")];
 | |
| 96 | ||
| 97 | in | |
| 98 | ||
| 99 | fun pretty_goal_header goal = | |
| 55763 | 100 | Pretty.block ([Pretty.keyword1 "goal"] @ subgoals (Thm.nprems_of goal) @ [Pretty.str ":"]); | 
| 49860 | 101 | |
| 102 | end; | |
| 103 | ||
| 104 | fun string_of_goal ctxt goal = | |
| 51958 
bca32217b304
retain goal display options when printing error messages, to avoid breakdown for huge goals;
 wenzelm parents: 
51601diff
changeset | 105 | Pretty.string_of (Pretty.chunks [pretty_goal_header goal, Goal_Display.pretty_goal ctxt goal]); | 
| 49860 | 106 | |
| 107 | ||
| 51584 | 108 | (* goal facts *) | 
| 109 | ||
| 110 | fun pretty_goal_facts ctxt s ths = | |
| 111 | (Pretty.block o Pretty.fbreaks) | |
| 51601 | 112 | [if s = "" then Pretty.str "this:" | 
| 55763 | 113 | else Pretty.block [Pretty.keyword1 s, Pretty.brk 1, Pretty.str "this:"], | 
| 51601 | 114 |      Proof_Context.pretty_fact ctxt ("", ths)];
 | 
| 51584 | 115 | |
| 116 | ||
| 49860 | 117 | (* method_error *) | 
| 118 | ||
| 49866 | 119 | fun method_error kind pos {context = ctxt, facts, goal} = Seq.Error (fn () =>
 | 
| 49860 | 120 | let | 
| 49866 | 121 | val pr_header = | 
| 122 | "Failed to apply " ^ (if kind = "" then "" else kind ^ " ") ^ | |
| 123 | "proof method" ^ Position.here pos ^ ":\n"; | |
| 49860 | 124 | val pr_facts = | 
| 125 | if null facts then "" | |
| 51584 | 126 | else Pretty.string_of (pretty_goal_facts ctxt "using" facts) ^ "\n"; | 
| 49860 | 127 | val pr_goal = string_of_goal ctxt goal; | 
| 128 | in pr_header ^ pr_facts ^ pr_goal end); | |
| 129 | ||
| 130 | ||
| 17369 | 131 | (* results *) | 
| 132 | ||
| 56932 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 133 | fun position_markup pos = Pretty.mark (Position.markup pos Markup.position); | 
| 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 134 | |
| 28092 | 135 | local | 
| 136 | ||
| 56932 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 137 | fun pretty_fact_name pos (kind, "") = position_markup pos (Pretty.keyword1 kind) | 
| 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 138 | | pretty_fact_name pos (kind, name) = | 
| 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 139 | Pretty.block [position_markup pos (Pretty.keyword1 kind), Pretty.brk 1, | 
| 41551 | 140 | 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: 
27857diff
changeset | 141 | |
| 17369 | 142 | fun pretty_facts ctxt = | 
| 55763 | 143 | flat o (separate [Pretty.fbrk, Pretty.keyword2 "and", Pretty.str " "]) o | 
| 50126 
3dec88149176
theorem status about oracles/futures is no longer printed by default;
 wenzelm parents: 
49867diff
changeset | 144 | map (single o Proof_Context.pretty_fact ctxt); | 
| 17369 | 145 | |
| 27857 | 146 | in | 
| 147 | ||
| 56932 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 148 | fun print_results do_print pos ctxt ((kind, name), facts) = | 
| 33643 
b275f26a638b
eliminated obsolete "internal" kind -- collapsed to unspecific "";
 wenzelm parents: 
33515diff
changeset | 149 | if not do_print orelse kind = "" then () | 
| 28092 | 150 | else if name = "" then | 
| 56897 
c668735fb8b5
print results as "state", to avoid intrusion into the source text;
 wenzelm parents: 
56894diff
changeset | 151 | (Pretty.writeln o Pretty.mark Markup.state) | 
| 56932 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 152 | (Pretty.block (position_markup pos (Pretty.keyword1 kind) :: Pretty.brk 1 :: | 
| 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 153 | pretty_facts ctxt facts)) | 
| 46728 
85f8e3932712
display proof results as "state", to suppress odd squiggles in the Prover IDE (see also 9240be8c8c69);
 wenzelm parents: 
45269diff
changeset | 154 | else | 
| 56897 
c668735fb8b5
print results as "state", to avoid intrusion into the source text;
 wenzelm parents: 
56894diff
changeset | 155 | (Pretty.writeln o Pretty.mark Markup.state) | 
| 46728 
85f8e3932712
display proof results as "state", to suppress odd squiggles in the Prover IDE (see also 9240be8c8c69);
 wenzelm parents: 
45269diff
changeset | 156 | (case facts of | 
| 56932 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 157 | [fact] => Pretty.blk (1, [pretty_fact_name pos (kind, name), Pretty.fbrk, | 
| 50126 
3dec88149176
theorem status about oracles/futures is no longer printed by default;
 wenzelm parents: 
49867diff
changeset | 158 | Proof_Context.pretty_fact ctxt fact]) | 
| 56932 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 159 | | _ => Pretty.blk (1, [pretty_fact_name pos (kind, name), Pretty.fbrk, | 
| 46728 
85f8e3932712
display proof results as "state", to suppress odd squiggles in the Prover IDE (see also 9240be8c8c69);
 wenzelm parents: 
45269diff
changeset | 160 |           Pretty.blk (1, Pretty.str "(" :: pretty_facts ctxt facts @ [Pretty.str ")"])]));
 | 
| 17369 | 161 | |
| 162 | end; | |
| 163 | ||
| 20889 | 164 | |
| 165 | (* consts *) | |
| 166 | ||
| 167 | local | |
| 168 | ||
| 169 | fun pretty_var ctxt (x, T) = | |
| 170 | Pretty.block [Pretty.str x, Pretty.str " ::", Pretty.brk 1, | |
| 24920 | 171 | Pretty.quote (Syntax.pretty_typ ctxt T)]; | 
| 20889 | 172 | |
| 56932 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 173 | fun pretty_vars pos ctxt kind vs = | 
| 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 174 | Pretty.block (Pretty.fbreaks (position_markup pos (Pretty.str kind) :: map (pretty_var ctxt) vs)); | 
| 20889 | 175 | |
| 56932 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 176 | fun pretty_consts pos ctxt pred cs = | 
| 42360 | 177 | (case filter pred (#1 (Proof_Context.inferred_fixes ctxt)) of | 
| 56932 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 178 | [] => pretty_vars pos ctxt "constants" cs | 
| 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 179 | | ps => Pretty.chunks [pretty_vars pos ctxt "parameters" ps, pretty_vars pos ctxt "constants" cs]); | 
| 20889 | 180 | |
| 44192 
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
 wenzelm parents: 
42717diff
changeset | 181 | in | 
| 
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
 wenzelm parents: 
42717diff
changeset | 182 | |
| 56932 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 183 | fun print_consts do_print pos ctxt pred cs = | 
| 44192 
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
 wenzelm parents: 
42717diff
changeset | 184 | if not do_print orelse null cs then () | 
| 56932 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 185 | else Pretty.writeln (Pretty.mark Markup.state (pretty_consts pos ctxt pred cs)); | 
| 44192 
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
 wenzelm parents: 
42717diff
changeset | 186 | |
| 20889 | 187 | end; | 
| 188 | ||
| 17369 | 189 | end; |