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