| author | wenzelm | 
| Fri, 15 Nov 2024 16:50:44 +0100 | |
| changeset 81453 | b99b531f13e6 | 
| parent 80873 | e71cb37c7395 | 
| child 81507 | 08574da77b4a | 
| 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 | |
| 61267 | 9 | val pretty_theory: bool -> Proof.context -> Pretty.T | 
| 61252 | 10 | val pretty_definitions: bool -> Proof.context -> Pretty.T | 
| 61266 | 11 | val pretty_theorems_diff: bool -> theory list -> Proof.context -> Pretty.T list | 
| 12 | val pretty_theorems: bool -> Proof.context -> Pretty.T list | |
| 20311 | 13 | val string_of_rule: Proof.context -> string -> thm -> string | 
| 49860 | 14 | val pretty_goal_header: thm -> Pretty.T | 
| 15 | val string_of_goal: Proof.context -> thm -> string | |
| 76061 | 16 | val pretty_goal_facts: Proof.context -> string -> thm list option -> Pretty.T list | 
| 76064 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 17 | val pretty_goal_inst: Proof.context -> term list list -> thm -> Pretty.T list | 
| 49866 | 18 | val method_error: string -> Position.T -> | 
| 19 |     {context: Proof.context, facts: thm list, goal: thm} -> 'a Seq.result
 | |
| 74156 | 20 | val show_results: bool Config.T | 
| 78469 
53b59fa42696
prefer Output.writeln for theory "results", as opposed to Output.state for genuine proof states (see f8c412a45af8, c668735fb8b5, ecf80e37ed1a);
 wenzelm parents: 
78465diff
changeset | 21 |   val print_results: {interactive: bool, pos: Position.T, proof_state: bool} -> Proof.context ->
 | 
| 56932 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 22 | ((string * string) * (string * thm list) list) -> unit | 
| 78465 | 23 | val print_theorem: Position.T -> Proof.context -> string * thm list -> unit | 
| 56932 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 24 | 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 | 25 | (string * typ -> bool) -> (string * typ) list -> unit | 
| 78032 | 26 | val markup_extern_label: Position.T -> (Markup.T * xstring) option | 
| 27 | val print_label: Position.T -> string | |
| 28 | val print_context_tracing: Context.generic * Position.T -> string | |
| 17369 | 29 | end | 
| 30 | ||
| 33389 | 31 | structure Proof_Display: PROOF_DISPLAY = | 
| 17369 | 32 | struct | 
| 33 | ||
| 61252 | 34 | (** theory content **) | 
| 35 | ||
| 61266 | 36 | fun pretty_theory verbose ctxt = | 
| 37 | let | |
| 38 | val thy = Proof_Context.theory_of ctxt; | |
| 39 | ||
| 40 | fun prt_cls c = Syntax.pretty_sort ctxt [c]; | |
| 41 | fun prt_sort S = Syntax.pretty_sort ctxt S; | |
| 42 | fun prt_arity t (c, Ss) = Syntax.pretty_arity ctxt (t, Ss, [c]); | |
| 43 | fun prt_typ ty = Pretty.quote (Syntax.pretty_typ ctxt ty); | |
| 44 | val prt_typ_no_tvars = prt_typ o Logic.unvarifyT_global; | |
| 45 | fun prt_term t = Pretty.quote (Syntax.pretty_term ctxt t); | |
| 46 | val prt_term_no_vars = prt_term o Logic.unvarify_global; | |
| 47 | fun prt_const (c, ty) = [Pretty.mark_str c, Pretty.str " ::", Pretty.brk 1, prt_typ_no_tvars ty]; | |
| 48 | ||
| 49 | fun pretty_classrel (c, []) = prt_cls c | |
| 50 | | pretty_classrel (c, cs) = Pretty.block | |
| 51 | (prt_cls c :: Pretty.str " <" :: Pretty.brk 1 :: Pretty.commas (map prt_cls cs)); | |
| 52 | ||
| 53 | fun pretty_default S = Pretty.block | |
| 54 | [Pretty.str "default sort:", Pretty.brk 1, prt_sort S]; | |
| 55 | ||
| 56 | val tfrees = map (fn v => TFree (v, [])); | |
| 79469 | 57 | fun pretty_type syn (t, Type.Logical_Type n) = | 
| 61266 | 58 | if syn then NONE | 
| 59 | else SOME (prt_typ (Type (t, tfrees (Name.invent Name.context Name.aT n)))) | |
| 60 | | pretty_type syn (t, Type.Abbreviation (vs, U, syn')) = | |
| 61 | if syn <> syn' then NONE | |
| 62 | else SOME (Pretty.block | |
| 63 | [prt_typ (Type (t, tfrees vs)), Pretty.str " =", Pretty.brk 1, prt_typ U]) | |
| 64 | | pretty_type syn (t, Type.Nonterminal) = | |
| 65 | if not syn then NONE | |
| 66 | else SOME (prt_typ (Type (t, []))); | |
| 67 | ||
| 68 | val pretty_arities = maps (fn (t, ars) => map (prt_arity t) ars); | |
| 69 | ||
| 70 | fun pretty_abbrev (c, (ty, t)) = Pretty.block | |
| 64596 | 71 | (prt_const (c, ty) @ [Pretty.str " \<equiv>", Pretty.brk 1, prt_term_no_vars t]); | 
| 19432 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 72 | |
| 61266 | 73 | val tsig = Sign.tsig_of thy; | 
| 74 | val consts = Sign.consts_of thy; | |
| 75 |     val {const_space, constants, constraints} = Consts.dest consts;
 | |
| 76 |     val {classes, default, types, ...} = Type.rep_tsig tsig;
 | |
| 77 | val type_space = Type.type_space tsig; | |
| 78 | val (class_space, class_algebra) = classes; | |
| 79 | val classes = Sorts.classes_of class_algebra; | |
| 80 | val arities = Sorts.arities_of class_algebra; | |
| 81 | ||
| 82 | val clsses = | |
| 83 | Name_Space.extern_entries verbose ctxt class_space | |
| 84 | (map (fn ((c, _), cs) => (c, Sign.minimize_sort thy cs)) (Graph.dest classes)) | |
| 85 | |> map (apfst #1); | |
| 86 | val tdecls = Name_Space.extern_table verbose ctxt types |> map (apfst #1); | |
| 87 | val arties = | |
| 88 | Name_Space.extern_entries verbose ctxt type_space (Symtab.dest arities) | |
| 89 | |> map (apfst #1); | |
| 90 | ||
| 91 | val cnsts = Name_Space.markup_entries verbose ctxt const_space constants; | |
| 92 | val log_cnsts = map_filter (fn (c, (ty, NONE)) => SOME (c, ty) | _ => NONE) cnsts; | |
| 93 | val abbrevs = map_filter (fn (c, (ty, SOME t)) => SOME (c, (ty, t)) | _ => NONE) cnsts; | |
| 94 | val cnstrs = Name_Space.markup_entries verbose ctxt const_space constraints; | |
| 95 | in | |
| 61267 | 96 | Pretty.chunks | 
| 97 | [Pretty.big_list "classes:" (map pretty_classrel clsses), | |
| 61266 | 98 | pretty_default default, | 
| 99 | Pretty.big_list "syntactic types:" (map_filter (pretty_type true) tdecls), | |
| 100 | Pretty.big_list "logical types:" (map_filter (pretty_type false) tdecls), | |
| 101 | Pretty.big_list "type arities:" (pretty_arities arties), | |
| 102 | Pretty.big_list "logical consts:" (map (Pretty.block o prt_const) log_cnsts), | |
| 103 | Pretty.big_list "abbreviations:" (map pretty_abbrev abbrevs), | |
| 104 | Pretty.big_list "constraints:" (map (Pretty.block o prt_const) cnstrs), | |
| 105 | Pretty.block | |
| 106 | (Pretty.breaks (Pretty.str "oracles:" :: | |
| 61267 | 107 | map Pretty.mark_str (Thm.extern_oracles verbose ctxt)))] | 
| 108 | end; | |
| 109 | ||
| 110 | ||
| 111 | (* theorems *) | |
| 112 | ||
| 113 | fun pretty_theorems_diff verbose prev_thys ctxt = | |
| 114 | let | |
| 115 | val pretty_fact = Proof_Context.pretty_fact ctxt; | |
| 116 | val facts = Global_Theory.facts_of (Proof_Context.theory_of ctxt); | |
| 117 | val thmss = Facts.dest_static verbose (map Global_Theory.facts_of prev_thys) facts; | |
| 118 | val prts = map #1 (sort_by (#1 o #2) (map (`pretty_fact) thmss)); | |
| 119 | in if null prts then [] else [Pretty.big_list "theorems:" prts] end; | |
| 120 | ||
| 121 | fun pretty_theorems verbose ctxt = | |
| 122 | pretty_theorems_diff verbose (Theory.parents_of (Proof_Context.theory_of ctxt)) ctxt; | |
| 20621 | 123 | |
| 61252 | 124 | |
| 125 | (* definitions *) | |
| 126 | ||
| 127 | fun pretty_definitions verbose ctxt = | |
| 128 | let | |
| 129 | val thy = Proof_Context.theory_of ctxt; | |
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61258diff
changeset | 130 | val context = Proof_Context.defs_context ctxt; | 
| 61252 | 131 | |
| 61253 | 132 | val const_space = Proof_Context.const_space ctxt; | 
| 133 | val type_space = Proof_Context.type_space ctxt; | |
| 61256 | 134 | val item_space = fn Defs.Const => const_space | Defs.Type => type_space; | 
| 61253 | 135 | fun prune_item (k, c) = not verbose andalso Name_Space.is_concealed (item_space k) c; | 
| 136 | ||
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61258diff
changeset | 137 | fun extern_item (kind, name) = | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61258diff
changeset | 138 | let val xname = Name_Space.extern ctxt (item_space kind) name | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61258diff
changeset | 139 | in (xname, (kind, name)) end; | 
| 61253 | 140 | |
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61258diff
changeset | 141 | fun extern_item_ord ((xname1, (kind1, _)), (xname2, (kind2, _))) = | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61258diff
changeset | 142 | (case Defs.item_kind_ord (kind1, kind2) of | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61258diff
changeset | 143 | EQUAL => string_ord (xname1, xname2) | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61258diff
changeset | 144 | | ord => ord); | 
| 61253 | 145 | |
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61258diff
changeset | 146 | fun sort_items f = sort (extern_item_ord o apply2 f); | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61258diff
changeset | 147 | |
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61258diff
changeset | 148 | fun pretty_entry ((_: string, item), args) = Defs.pretty_entry context (item, args); | 
| 61252 | 149 | |
| 150 | fun pretty_reduct (lhs, rhs) = Pretty.block | |
| 61253 | 151 | ([pretty_entry lhs, Pretty.str " ->", Pretty.brk 2] @ | 
| 152 | Pretty.commas (map pretty_entry (sort_items fst rhs))); | |
| 61252 | 153 | |
| 61253 | 154 | fun pretty_restrict (entry, name) = | 
| 155 |       Pretty.block ([pretty_entry entry, Pretty.brk 2, Pretty.str ("(from " ^ quote name ^ ")")]);
 | |
| 61252 | 156 | |
| 157 | val defs = Theory.defs_of thy; | |
| 158 |     val {restricts, reducts} = Defs.dest defs;
 | |
| 22335 | 159 | |
| 61257 | 160 | val (reds1, reds2) = | 
| 61253 | 161 | filter_out (prune_item o #1 o #1) reducts | 
| 61252 | 162 | |> map (fn (lhs, rhs) => | 
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61258diff
changeset | 163 | (apfst extern_item lhs, map (apfst extern_item) (filter_out (prune_item o fst) rhs))) | 
| 61253 | 164 | |> sort_items (#1 o #1) | 
| 61257 | 165 | |> filter_out (null o #2) | 
| 166 | |> List.partition (Defs.plain_args o #2 o #1); | |
| 61253 | 167 | |
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61258diff
changeset | 168 | val rests = restricts |> map (apfst (apfst extern_item)) |> sort_items (#1 o #1); | 
| 61252 | 169 | in | 
| 170 | Pretty.big_list "definitions:" | |
| 61258 | 171 | (map (Pretty.text_fold o single) | 
| 172 | [Pretty.big_list "non-overloaded LHS:" (map pretty_reduct reds1), | |
| 173 | Pretty.big_list "overloaded LHS:" (map pretty_reduct reds2), | |
| 174 | Pretty.big_list "pattern restrictions due to incomplete LHS:" (map pretty_restrict rests)]) | |
| 61252 | 175 | end; | 
| 176 | ||
| 177 | ||
| 178 | ||
| 179 | (** proof items **) | |
| 19432 
cae7cc072994
added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
 wenzelm parents: 
18799diff
changeset | 180 | |
| 17369 | 181 | (* refinement rule *) | 
| 182 | ||
| 183 | fun pretty_rule ctxt s thm = | |
| 184 | Pretty.block [Pretty.str (s ^ " attempt to solve goal by exported rule:"), | |
| 61268 | 185 | Pretty.fbrk, Thm.pretty_thm ctxt thm]; | 
| 17369 | 186 | |
| 187 | val string_of_rule = Pretty.string_of ooo pretty_rule; | |
| 188 | ||
| 189 | ||
| 49860 | 190 | (* goals *) | 
| 191 | ||
| 192 | local | |
| 193 | ||
| 194 | fun subgoals 0 = [] | |
| 195 | | subgoals 1 = [Pretty.brk 1, Pretty.str "(1 subgoal)"] | |
| 196 |   | subgoals n = [Pretty.brk 1, Pretty.str ("(" ^ string_of_int n ^ " subgoals)")];
 | |
| 197 | ||
| 198 | in | |
| 199 | ||
| 200 | fun pretty_goal_header goal = | |
| 55763 | 201 | Pretty.block ([Pretty.keyword1 "goal"] @ subgoals (Thm.nprems_of goal) @ [Pretty.str ":"]); | 
| 49860 | 202 | |
| 203 | end; | |
| 204 | ||
| 205 | 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 | 206 | Pretty.string_of (Pretty.chunks [pretty_goal_header goal, Goal_Display.pretty_goal ctxt goal]); | 
| 49860 | 207 | |
| 208 | ||
| 51584 | 209 | (* goal facts *) | 
| 210 | ||
| 76061 | 211 | fun pretty_goal_facts _ _ NONE = [] | 
| 212 | | pretty_goal_facts ctxt s (SOME ths) = | |
| 213 | (single o Pretty.block o Pretty.fbreaks) | |
| 214 | [if s = "" then Pretty.str "this:" | |
| 215 | else Pretty.block [Pretty.keyword1 s, Pretty.brk 1, Pretty.str "this:"], | |
| 216 |          Proof_Context.pretty_fact ctxt ("", ths)];
 | |
| 51584 | 217 | |
| 218 | ||
| 76064 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 219 | (* goal instantiation *) | 
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 220 | |
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 221 | fun pretty_goal_inst ctxt propss goal = | 
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 222 | let | 
| 76079 | 223 | val title = "goal instantiation:"; | 
| 76064 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 224 | |
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 225 | fun prt_inst env = | 
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 226 | if Envir.is_empty env then [] | 
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 227 | else | 
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 228 | let | 
| 76067 
e39c1da9d904
proper Envir.subst operations: env is already normalized, using Envir.norm may lead to non-termination;
 wenzelm parents: 
76066diff
changeset | 229 |           val Envir.Envir {tyenv, tenv, ...} = env;
 | 
| 
e39c1da9d904
proper Envir.subst operations: env is already normalized, using Envir.norm may lead to non-termination;
 wenzelm parents: 
76066diff
changeset | 230 | |
| 76064 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 231 | val prt_type = Syntax.pretty_typ ctxt; | 
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 232 | val prt_term = Syntax.pretty_term ctxt; | 
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 233 | |
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 234 | fun instT v = | 
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 235 | let | 
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 236 | val T = TVar v; | 
| 76067 
e39c1da9d904
proper Envir.subst operations: env is already normalized, using Envir.norm may lead to non-termination;
 wenzelm parents: 
76066diff
changeset | 237 | val T' = Envir.subst_type tyenv T; | 
| 76064 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 238 | in if T = T' then NONE else SOME (prt_type T, prt_type T') end; | 
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 239 | |
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 240 | fun inst v = | 
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 241 | let | 
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 242 | val t = Var v; | 
| 79540 
afa75b58166a
more robust (amending 1600fb749c54), to support the following corner case:
 wenzelm parents: 
79469diff
changeset | 243 | val t' = | 
| 
afa75b58166a
more robust (amending 1600fb749c54), to support the following corner case:
 wenzelm parents: 
79469diff
changeset | 244 | Envir.subst_term (tyenv, tenv) t handle TYPE _ => | 
| 
afa75b58166a
more robust (amending 1600fb749c54), to support the following corner case:
 wenzelm parents: 
79469diff
changeset | 245 | Envir.subst_term (tyenv, tenv) (Envir.subst_term_types tyenv t); | 
| 76064 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 246 | in if t aconv t' then NONE else SOME (prt_term t, prt_term t') end; | 
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 247 | |
| 76079 | 248 | fun inst_pair (x, y) = Pretty.item [x, Pretty.str " \<leadsto>", Pretty.brk 1, y]; | 
| 76064 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 249 | |
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 250 | val prts = | 
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 251 | ((fold o fold) Term.add_tvars propss [] |> sort Term_Ord.tvar_ord |> map_filter instT) @ | 
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 252 | ((fold o fold) Term.add_vars propss [] |> sort Term_Ord.var_ord |> map_filter inst); | 
| 76079 | 253 | in if null prts then [] else [Pretty.big_list title (map inst_pair prts)] end; | 
| 76064 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 254 | |
| 76081 
730638d4e37a
clarified failure: warning for logical error, exception for program breakdown;
 wenzelm parents: 
76079diff
changeset | 255 | exception LOST_STRUCTURE; | 
| 76064 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 256 | fun goal_matcher () = | 
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 257 | let | 
| 76081 
730638d4e37a
clarified failure: warning for logical error, exception for program breakdown;
 wenzelm parents: 
76079diff
changeset | 258 | val concl = | 
| 
730638d4e37a
clarified failure: warning for logical error, exception for program breakdown;
 wenzelm parents: 
76079diff
changeset | 259 | Logic.unprotect (Thm.concl_of goal) | 
| 
730638d4e37a
clarified failure: warning for logical error, exception for program breakdown;
 wenzelm parents: 
76079diff
changeset | 260 | handle TERM _ => raise LOST_STRUCTURE; | 
| 76064 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 261 | val goal_propss = filter_out null propss; | 
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 262 | val results = | 
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 263 | Logic.dest_conjunction_balanced (length goal_propss) concl | 
| 76081 
730638d4e37a
clarified failure: warning for logical error, exception for program breakdown;
 wenzelm parents: 
76079diff
changeset | 264 | |> map2 Logic.dest_conjunction_balanced (map length goal_propss) | 
| 
730638d4e37a
clarified failure: warning for logical error, exception for program breakdown;
 wenzelm parents: 
76079diff
changeset | 265 | handle TERM _ => raise LOST_STRUCTURE; | 
| 76064 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 266 | in | 
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 267 | Unify.matcher (Context.Proof ctxt) | 
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 268 | (map (Soft_Type_System.purge ctxt) (flat goal_propss)) (flat results) | 
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 269 | end; | 
| 76081 
730638d4e37a
clarified failure: warning for logical error, exception for program breakdown;
 wenzelm parents: 
76079diff
changeset | 270 | |
| 
730638d4e37a
clarified failure: warning for logical error, exception for program breakdown;
 wenzelm parents: 
76079diff
changeset | 271 | fun failure msg = (warning (title ^ " " ^ msg); []); | 
| 76064 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 272 | in | 
| 76095 
7cac5565e79b
discontinued somewhat pointless option: Proof_Display.pretty_goal_inst should always work smoothly (and not crash unexpectedly);
 wenzelm parents: 
76081diff
changeset | 273 | (case goal_matcher () of | 
| 
7cac5565e79b
discontinued somewhat pointless option: Proof_Display.pretty_goal_inst should always work smoothly (and not crash unexpectedly);
 wenzelm parents: 
76081diff
changeset | 274 | SOME env => prt_inst env | 
| 
7cac5565e79b
discontinued somewhat pointless option: Proof_Display.pretty_goal_inst should always work smoothly (and not crash unexpectedly);
 wenzelm parents: 
76081diff
changeset | 275 | | NONE => failure "match failed") | 
| 
7cac5565e79b
discontinued somewhat pointless option: Proof_Display.pretty_goal_inst should always work smoothly (and not crash unexpectedly);
 wenzelm parents: 
76081diff
changeset | 276 | handle LOST_STRUCTURE => failure "lost goal structure" | 
| 76064 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 277 | end; | 
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 278 | |
| 
28ddebb43d93
show goal instantiation, notably for 'schematic_goal' command (inactive by default);
 wenzelm parents: 
76061diff
changeset | 279 | |
| 49860 | 280 | (* method_error *) | 
| 281 | ||
| 49866 | 282 | fun method_error kind pos {context = ctxt, facts, goal} = Seq.Error (fn () =>
 | 
| 49860 | 283 | let | 
| 49866 | 284 | val pr_header = | 
| 285 | "Failed to apply " ^ (if kind = "" then "" else kind ^ " ") ^ | |
| 286 | "proof method" ^ Position.here pos ^ ":\n"; | |
| 49860 | 287 | val pr_facts = | 
| 288 | if null facts then "" | |
| 76061 | 289 | else Pretty.string_of (Pretty.chunks (pretty_goal_facts ctxt "using" (SOME facts))) ^ "\n"; | 
| 49860 | 290 | val pr_goal = string_of_goal ctxt goal; | 
| 291 | in pr_header ^ pr_facts ^ pr_goal end); | |
| 292 | ||
| 293 | ||
| 17369 | 294 | (* results *) | 
| 295 | ||
| 74156 | 296 | val interactive = | 
| 297 |   Config.declare_bool ("interactive", \<^here>) (K false);
 | |
| 298 | ||
| 299 | val show_results = | |
| 76065 | 300 | Attrib.setup_config_bool \<^binding>\<open>show_results\<close> | 
| 76073 
951abf9db857
option "show_states" for more verbosity of batch-builds;
 wenzelm parents: 
76071diff
changeset | 301 | (fn context => | 
| 
951abf9db857
option "show_states" for more verbosity of batch-builds;
 wenzelm parents: 
76071diff
changeset | 302 | Config.get_generic context interactive orelse | 
| 
951abf9db857
option "show_states" for more verbosity of batch-builds;
 wenzelm parents: 
76071diff
changeset | 303 | Options.default_bool \<^system_option>\<open>show_states\<close>); | 
| 74156 | 304 | |
| 305 | fun no_print int ctxt = not (Config.get (Config.put interactive int ctxt) show_results); | |
| 306 | ||
| 28092 | 307 | local | 
| 308 | ||
| 80873 | 309 | fun pretty_fact_name pos (kind, "") = Pretty.mark_position pos (Pretty.keyword1 kind) | 
| 56932 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 310 | | pretty_fact_name pos (kind, name) = | 
| 80873 | 311 | Pretty.block [Pretty.mark_position pos (Pretty.keyword1 kind), Pretty.brk 1, | 
| 41551 | 312 | 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 | 313 | |
| 17369 | 314 | fun pretty_facts ctxt = | 
| 55763 | 315 | 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 | 316 | map (single o Proof_Context.pretty_fact ctxt); | 
| 17369 | 317 | |
| 27857 | 318 | in | 
| 319 | ||
| 78469 
53b59fa42696
prefer Output.writeln for theory "results", as opposed to Output.state for genuine proof states (see f8c412a45af8, c668735fb8b5, ecf80e37ed1a);
 wenzelm parents: 
78465diff
changeset | 320 | fun print_results {interactive, pos, proof_state} ctxt ((kind, name), facts) =
 | 
| 
53b59fa42696
prefer Output.writeln for theory "results", as opposed to Output.state for genuine proof states (see f8c412a45af8, c668735fb8b5, ecf80e37ed1a);
 wenzelm parents: 
78465diff
changeset | 321 | let val print = if proof_state then Output.state o Pretty.string_of else Pretty.writeln in | 
| 
53b59fa42696
prefer Output.writeln for theory "results", as opposed to Output.state for genuine proof states (see f8c412a45af8, c668735fb8b5, ecf80e37ed1a);
 wenzelm parents: 
78465diff
changeset | 322 | if kind = "" orelse no_print interactive ctxt then () | 
| 
53b59fa42696
prefer Output.writeln for theory "results", as opposed to Output.state for genuine proof states (see f8c412a45af8, c668735fb8b5, ecf80e37ed1a);
 wenzelm parents: 
78465diff
changeset | 323 | else if name = "" then | 
| 
53b59fa42696
prefer Output.writeln for theory "results", as opposed to Output.state for genuine proof states (see f8c412a45af8, c668735fb8b5, ecf80e37ed1a);
 wenzelm parents: 
78465diff
changeset | 324 | |
| 80873 | 325 | (Pretty.block (Pretty.mark_position pos (Pretty.keyword1 kind) :: Pretty.brk 1 :: | 
| 78469 
53b59fa42696
prefer Output.writeln for theory "results", as opposed to Output.state for genuine proof states (see f8c412a45af8, c668735fb8b5, ecf80e37ed1a);
 wenzelm parents: 
78465diff
changeset | 326 | pretty_facts ctxt facts)) | 
| 
53b59fa42696
prefer Output.writeln for theory "results", as opposed to Output.state for genuine proof states (see f8c412a45af8, c668735fb8b5, ecf80e37ed1a);
 wenzelm parents: 
78465diff
changeset | 327 | else | 
| 
53b59fa42696
prefer Output.writeln for theory "results", as opposed to Output.state for genuine proof states (see f8c412a45af8, c668735fb8b5, ecf80e37ed1a);
 wenzelm parents: 
78465diff
changeset | 328 | |
| 
53b59fa42696
prefer Output.writeln for theory "results", as opposed to Output.state for genuine proof states (see f8c412a45af8, c668735fb8b5, ecf80e37ed1a);
 wenzelm parents: 
78465diff
changeset | 329 | (case facts of | 
| 80328 | 330 | [fact] => Pretty.block1 [pretty_fact_name pos (kind, name), Pretty.fbrk, | 
| 331 | Proof_Context.pretty_fact ctxt fact] | |
| 332 | | _ => Pretty.block1 [pretty_fact_name pos (kind, name), Pretty.fbrk, | |
| 333 |             Pretty.block1 (Pretty.str "(" :: pretty_facts ctxt facts @ [Pretty.str ")"])])
 | |
| 78469 
53b59fa42696
prefer Output.writeln for theory "results", as opposed to Output.state for genuine proof states (see f8c412a45af8, c668735fb8b5, ecf80e37ed1a);
 wenzelm parents: 
78465diff
changeset | 334 | end; | 
| 17369 | 335 | |
| 78465 | 336 | fun print_theorem pos ctxt fact = | 
| 78469 
53b59fa42696
prefer Output.writeln for theory "results", as opposed to Output.state for genuine proof states (see f8c412a45af8, c668735fb8b5, ecf80e37ed1a);
 wenzelm parents: 
78465diff
changeset | 337 |   print_results {interactive = true, pos = pos, proof_state = false}
 | 
| 
53b59fa42696
prefer Output.writeln for theory "results", as opposed to Output.state for genuine proof states (see f8c412a45af8, c668735fb8b5, ecf80e37ed1a);
 wenzelm parents: 
78465diff
changeset | 338 | ctxt ((Thm.theoremK, ""), [fact]); | 
| 78465 | 339 | |
| 17369 | 340 | end; | 
| 341 | ||
| 20889 | 342 | |
| 343 | (* consts *) | |
| 344 | ||
| 345 | local | |
| 346 | ||
| 347 | fun pretty_var ctxt (x, T) = | |
| 348 | Pretty.block [Pretty.str x, Pretty.str " ::", Pretty.brk 1, | |
| 24920 | 349 | Pretty.quote (Syntax.pretty_typ ctxt T)]; | 
| 20889 | 350 | |
| 56932 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 351 | fun pretty_vars pos ctxt kind vs = | 
| 62188 | 352 | Pretty.block | 
| 80873 | 353 | (Pretty.fbreaks (Pretty.mark_position pos (Pretty.mark_str kind) :: map (pretty_var ctxt) vs)); | 
| 62188 | 354 | |
| 355 | val fixes = (Markup.keyword2, "fixes"); | |
| 356 | val consts = (Markup.keyword1, "consts"); | |
| 20889 | 357 | |
| 56932 
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
 wenzelm parents: 
56897diff
changeset | 358 | fun pretty_consts pos ctxt pred cs = | 
| 42360 | 359 | (case filter pred (#1 (Proof_Context.inferred_fixes ctxt)) of | 
| 62188 | 360 | [] => pretty_vars pos ctxt consts cs | 
| 361 | | ps => Pretty.chunks [pretty_vars pos ctxt fixes ps, pretty_vars pos ctxt consts cs]); | |
| 20889 | 362 | |
| 44192 
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
 wenzelm parents: 
42717diff
changeset | 363 | in | 
| 
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
 wenzelm parents: 
42717diff
changeset | 364 | |
| 74156 | 365 | fun print_consts int pos ctxt pred cs = | 
| 366 | if null cs orelse no_print int ctxt then () | |
| 78469 
53b59fa42696
prefer Output.writeln for theory "results", as opposed to Output.state for genuine proof states (see f8c412a45af8, c668735fb8b5, ecf80e37ed1a);
 wenzelm parents: 
78465diff
changeset | 367 | else Pretty.writeln (pretty_consts pos ctxt pred cs); | 
| 44192 
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
 wenzelm parents: 
42717diff
changeset | 368 | |
| 20889 | 369 | end; | 
| 370 | ||
| 78032 | 371 | |
| 372 | (* position label *) | |
| 373 | ||
| 374 | val command_prefix = Markup.commandN ^ "."; | |
| 375 | ||
| 376 | fun markup_extern_label pos = | |
| 377 | Position.label_of pos |> Option.map (fn label => | |
| 378 | (case try (unprefix command_prefix) label of | |
| 379 | SOME cmd => (Markup.keyword1, Long_Name.base_name cmd) | |
| 380 | | NONE => (Markup.empty, label))); | |
| 381 | ||
| 382 | fun print_label pos = | |
| 383 | (case markup_extern_label pos of | |
| 384 | NONE => "" | |
| 385 | | SOME (m, s) => Markup.markup m s); | |
| 386 | ||
| 387 | ||
| 388 | (* context tracing *) | |
| 389 | ||
| 390 | fun context_type (Context.Theory _) = "theory" | |
| 391 | | context_type (Context.Proof ctxt) = | |
| 392 | if can Local_Theory.assert ctxt then "local_theory" else "Proof.context"; | |
| 393 | ||
| 394 | fun print_context_tracing (context, pos) = | |
| 395 | context_type context ^ (case print_label pos of "" => "" | s => " " ^ s) ^ Position.here pos; | |
| 396 | ||
| 17369 | 397 | end; |