| author | haftmann | 
| Mon, 03 Dec 2007 16:04:14 +0100 | |
| changeset 25517 | 36d710d1dbce | 
| parent 24920 | 2a45e400fdad | 
| child 26336 | a0e2b706ce73 | 
| permissions | -rw-r--r-- | 
| 17369 | 1 | (* Title: Pure/Isar/proof_display.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Makarius | |
| 4 | ||
| 19432 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 5 | Printing of theorems, goals, results etc. | 
| 17369 | 6 | *) | 
| 7 | ||
| 19432 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 8 | signature BASIC_PROOF_DISPLAY = | 
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 9 | sig | 
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 10 | val print_theorems: theory -> unit | 
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 11 | val print_theory: theory -> unit | 
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 12 | end | 
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 13 | |
| 17369 | 14 | signature PROOF_DISPLAY = | 
| 15 | sig | |
| 19432 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 16 | include BASIC_PROOF_DISPLAY | 
| 20311 | 17 | val pprint_context: Proof.context -> pprint_args -> unit | 
| 20211 | 18 | val pprint_typ: theory -> typ -> pprint_args -> unit | 
| 19 | val pprint_term: theory -> term -> pprint_args -> unit | |
| 20 | val pprint_ctyp: ctyp -> pprint_args -> unit | |
| 21 | val pprint_cterm: cterm -> pprint_args -> unit | |
| 22 | val pprint_thm: thm -> pprint_args -> unit | |
| 19432 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 23 | val print_theorems_diff: theory -> theory -> unit | 
| 22335 | 24 | val pretty_full_theory: bool -> theory -> Pretty.T | 
| 20311 | 25 | val string_of_rule: Proof.context -> string -> thm -> string | 
| 26 | val print_results: bool -> Proof.context -> | |
| 17369 | 27 | ((string * string) * (string * thm list) list) -> unit | 
| 20311 | 28 | val present_results: Proof.context -> | 
| 17369 | 29 | ((string * string) * (string * thm list) list) -> unit | 
| 22872 | 30 | val pretty_consts: Proof.context -> | 
| 31 | (string * typ -> bool) -> (string * typ) list -> Pretty.T | |
| 17369 | 32 | end | 
| 33 | ||
| 34 | structure ProofDisplay: PROOF_DISPLAY = | |
| 35 | struct | |
| 36 | ||
| 20235 | 37 | (* toplevel pretty printing *) | 
| 20211 | 38 | |
| 39 | fun pprint_context ctxt = Pretty.pprint | |
| 20311 | 40 | (if ! ProofContext.debug then | 
| 20211 | 41 | Pretty.quote (Pretty.big_list "proof context:" (ProofContext.pretty_context ctxt)) | 
| 42 | else Pretty.str "<context>"); | |
| 43 | ||
| 44 | fun pprint pretty thy = Pretty.pprint o Pretty.quote o pretty (ProofContext.init thy); | |
| 45 | ||
| 24920 | 46 | val pprint_typ = pprint Syntax.pretty_typ; | 
| 47 | val pprint_term = pprint Syntax.pretty_term; | |
| 20211 | 48 | fun pprint_ctyp cT = pprint_typ (Thm.theory_of_ctyp cT) (Thm.typ_of cT); | 
| 49 | fun pprint_cterm ct = pprint_term (Thm.theory_of_cterm ct) (Thm.term_of ct); | |
| 22872 | 50 | val pprint_thm = Pretty.pprint o ProofContext.pretty_thm_legacy; | 
| 20211 | 51 | |
| 19432 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 52 | |
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 53 | (* theorems and theory *) | 
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 54 | |
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 55 | fun pretty_theorems_diff prev_thms thy = | 
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 56 | let | 
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 57 | val ctxt = ProofContext.init thy; | 
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 58 | val (space, thms) = PureThy.theorems_of thy; | 
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 59 | val diff_thmss = Symtab.fold (fn fact => | 
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 60 | if not (Symtab.member Thm.eq_thms prev_thms fact) then cons fact else I) thms []; | 
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 61 | val thmss = diff_thmss |> map (apfst (NameSpace.extern space)) |> Library.sort_wrt #1; | 
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 62 | in Pretty.big_list "theorems:" (map (ProofContext.pretty_fact ctxt) thmss) end; | 
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 63 | |
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 64 | fun print_theorems_diff prev_thy thy = | 
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 65 | Pretty.writeln (pretty_theorems_diff (#2 (PureThy.theorems_of prev_thy)) thy); | 
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 66 | |
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 67 | fun pretty_theorems thy = pretty_theorems_diff Symtab.empty thy; | 
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 68 | |
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 69 | val print_theorems = Pretty.writeln o pretty_theorems; | 
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 70 | |
| 22872 | 71 | fun pretty_full_theory verbose thy = | 
| 72 | Pretty.chunks (Display.pretty_full_theory verbose thy @ [pretty_theorems thy]); | |
| 20621 | 73 | |
| 22872 | 74 | val print_theory = Pretty.writeln o pretty_full_theory false; | 
| 22335 | 75 | |
| 19432 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 76 | |
| 17369 | 77 | (* refinement rule *) | 
| 78 | ||
| 79 | fun pretty_rule ctxt s thm = | |
| 80 | Pretty.block [Pretty.str (s ^ " attempt to solve goal by exported rule:"), | |
| 81 | Pretty.fbrk, ProofContext.pretty_thm ctxt thm]; | |
| 82 | ||
| 83 | val string_of_rule = Pretty.string_of ooo pretty_rule; | |
| 84 | ||
| 85 | ||
| 86 | (* results *) | |
| 87 | ||
| 88 | local | |
| 89 | ||
| 90 | fun pretty_facts ctxt = | |
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19432diff
changeset | 91 | flat o (separate [Pretty.fbrk, Pretty.str "and "]) o | 
| 17369 | 92 | map (single o ProofContext.pretty_fact ctxt); | 
| 93 | ||
| 94 | fun pretty_results ctxt ((kind, ""), facts) = | |
| 95 | Pretty.block (Pretty.str kind :: Pretty.brk 1 :: pretty_facts ctxt facts) | |
| 96 | | pretty_results ctxt ((kind, name), [fact]) = Pretty.blk (1, | |
| 97 | [Pretty.str (kind ^ " " ^ name ^ ":"), Pretty.fbrk, ProofContext.pretty_fact ctxt fact]) | |
| 98 | | pretty_results ctxt ((kind, name), facts) = Pretty.blk (1, | |
| 99 | [Pretty.str (kind ^ " " ^ name ^ ":"), Pretty.fbrk, | |
| 100 |         Pretty.blk (1, Pretty.str "(" :: pretty_facts ctxt facts @ [Pretty.str ")"])]);
 | |
| 101 | ||
| 102 | fun name_results "" res = res | |
| 103 | | name_results name res = | |
| 104 | let | |
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19432diff
changeset | 105 | val n = length (maps snd res); | 
| 17369 | 106 | fun name_res (a, ths) i = | 
| 107 | let | |
| 108 | val m = length ths; | |
| 109 | val j = i + m; | |
| 110 | in | |
| 111 | if a <> "" then ((a, ths), j) | |
| 112 | else if m = n then ((name, ths), j) | |
| 113 | else if m = 1 then | |
| 114 | ((PureThy.string_of_thmref (NameSelection (name, [Single i])), ths), j) | |
| 115 | else ((PureThy.string_of_thmref (NameSelection (name, [FromTo (i, j - 1)])), ths), j) | |
| 116 | end; | |
| 17756 | 117 | in fst (fold_map name_res res 1) end; | 
| 17369 | 118 | |
| 119 | in | |
| 120 | ||
| 121 | fun print_results true = Pretty.writeln oo pretty_results | |
| 122 | | print_results false = K (K ()); | |
| 123 | ||
| 124 | fun present_results ctxt ((kind, name), res) = | |
| 21437 | 125 | if kind = "" orelse kind = Thm.internalK then () | 
| 17369 | 126 | else (print_results true ctxt ((kind, name), res); | 
| 22095 
07875394618e
moved ML context stuff to from Context to ML_Context;
 wenzelm parents: 
22086diff
changeset | 127 | ML_Context.setmp (SOME (Context.Proof ctxt)) | 
| 17369 | 128 | (Present.results kind) (name_results name res)); | 
| 129 | ||
| 130 | end; | |
| 131 | ||
| 20889 | 132 | |
| 133 | (* consts *) | |
| 134 | ||
| 135 | local | |
| 136 | ||
| 137 | fun pretty_var ctxt (x, T) = | |
| 138 | Pretty.block [Pretty.str x, Pretty.str " ::", Pretty.brk 1, | |
| 24920 | 139 | Pretty.quote (Syntax.pretty_typ ctxt T)]; | 
| 20889 | 140 | |
| 141 | fun pretty_vars ctxt kind vs = Pretty.big_list kind (map (pretty_var ctxt) vs); | |
| 142 | ||
| 143 | in | |
| 144 | ||
| 145 | fun pretty_consts ctxt pred cs = | |
| 146 | (case filter pred (#1 (ProofContext.inferred_fixes ctxt)) of | |
| 147 | [] => pretty_vars ctxt "constants" cs | |
| 148 | | ps => Pretty.chunks [pretty_vars ctxt "parameters" ps, pretty_vars ctxt "constants" cs]); | |
| 149 | ||
| 150 | end; | |
| 151 | ||
| 17369 | 152 | end; | 
| 19432 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 153 | |
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 154 | structure BasicProofDisplay: BASIC_PROOF_DISPLAY = ProofDisplay; | 
| 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 155 | open BasicProofDisplay; |