src/Pure/Isar/proof_display.ML
author wenzelm
Tue, 16 Oct 2012 20:23:00 +0200
changeset 49866 619acbd72664
parent 49863 b5fb6e7f8d81
child 49867 d3053a55bfcb
permissions -rw-r--r--
more proof method text position information;
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
30628
4078276bcace removed obsolete pprint operations;
wenzelm
parents: 30364
diff changeset
     9
  val pp_context: Proof.context -> Pretty.T
4078276bcace removed obsolete pprint operations;
wenzelm
parents: 30364
diff changeset
    10
  val pp_thm: thm -> Pretty.T
4078276bcace removed obsolete pprint operations;
wenzelm
parents: 30364
diff changeset
    11
  val pp_typ: theory -> typ -> Pretty.T
4078276bcace removed obsolete pprint operations;
wenzelm
parents: 30364
diff changeset
    12
  val pp_term: theory -> term -> Pretty.T
4078276bcace removed obsolete pprint operations;
wenzelm
parents: 30364
diff changeset
    13
  val pp_ctyp: ctyp -> Pretty.T
4078276bcace removed obsolete pprint operations;
wenzelm
parents: 30364
diff changeset
    14
  val pp_cterm: cterm -> Pretty.T
33515
d066e8369a33 print_theorems: suppress concealed (global) facts, unless "!" option is given;
wenzelm
parents: 33389
diff changeset
    15
  val print_theorems_diff: bool -> theory -> theory -> unit
d066e8369a33 print_theorems: suppress concealed (global) facts, unless "!" option is given;
wenzelm
parents: 33389
diff changeset
    16
  val print_theorems: bool -> theory -> unit
22335
6c4204df6863 pretty_full_theory: expose in signature.
aspinall
parents: 22095
diff changeset
    17
  val pretty_full_theory: bool -> theory -> Pretty.T
28092
5886e9359aa8 no pervasive bindings;
wenzelm
parents: 28087
diff changeset
    18
  val print_theory: theory -> unit
20311
3666316adad6 moved debug option to proof_display.ML (again);
wenzelm
parents: 20253
diff changeset
    19
  val string_of_rule: Proof.context -> string -> thm -> string
49860
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
    20
  val pretty_goal_header: thm -> Pretty.T
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
    21
  val string_of_goal: Proof.context -> thm -> string
49866
619acbd72664 more proof method text position information;
wenzelm
parents: 49863
diff changeset
    22
  val method_error: string -> Position.T ->
619acbd72664 more proof method text position information;
wenzelm
parents: 49863
diff changeset
    23
    {context: Proof.context, facts: thm list, goal: thm} -> 'a Seq.result
46728
85f8e3932712 display proof results as "state", to suppress odd squiggles in the Prover IDE (see also 9240be8c8c69);
wenzelm
parents: 45269
diff changeset
    24
  val print_results: Markup.T -> bool -> Proof.context ->
85f8e3932712 display proof results as "state", to suppress odd squiggles in the Prover IDE (see also 9240be8c8c69);
wenzelm
parents: 45269
diff changeset
    25
    ((string * string) * (string * thm list) list) -> unit
44192
a32ca9165928 less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents: 42717
diff changeset
    26
  val print_consts: bool -> Proof.context -> (string * typ -> bool) -> (string * typ) list -> unit
17369
dec2ddbb3edf Printing of Isar proof elements etc.
wenzelm
parents:
diff changeset
    27
end
dec2ddbb3edf Printing of Isar proof elements etc.
wenzelm
parents:
diff changeset
    28
33389
bb3a5fa94a91 modernized structure Proof_Display;
wenzelm
parents: 32145
diff changeset
    29
structure Proof_Display: PROOF_DISPLAY =
17369
dec2ddbb3edf Printing of Isar proof elements etc.
wenzelm
parents:
diff changeset
    30
struct
dec2ddbb3edf Printing of Isar proof elements etc.
wenzelm
parents:
diff changeset
    31
20235
3cbf73ed92f8 ProofContext.legacy_pretty_thm;
wenzelm
parents: 20211
diff changeset
    32
(* toplevel pretty printing *)
20211
c7f907f41f7c moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 19482
diff changeset
    33
30628
4078276bcace removed obsolete pprint operations;
wenzelm
parents: 30364
diff changeset
    34
fun pp_context ctxt =
42717
0bbb56867091 proper configuration options Proof_Context.debug and Proof_Context.verbose;
wenzelm
parents: 42360
diff changeset
    35
 (if Config.get ctxt Proof_Context.debug then
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 41551
diff changeset
    36
    Pretty.quote (Pretty.big_list "proof context:" (Proof_Context.pretty_context ctxt))
20211
c7f907f41f7c moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 19482
diff changeset
    37
  else Pretty.str "<context>");
c7f907f41f7c moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 19482
diff changeset
    38
45269
6f8949e6c71a more robust ML pretty printing (cf. b6c527c64789);
wenzelm
parents: 44240
diff changeset
    39
fun default_context thy0 =
6f8949e6c71a more robust ML pretty printing (cf. b6c527c64789);
wenzelm
parents: 44240
diff changeset
    40
  (case Context.thread_data () of
6f8949e6c71a more robust ML pretty printing (cf. b6c527c64789);
wenzelm
parents: 44240
diff changeset
    41
    SOME (Context.Proof ctxt) => ctxt
6f8949e6c71a more robust ML pretty printing (cf. b6c527c64789);
wenzelm
parents: 44240
diff changeset
    42
  | SOME (Context.Theory thy) =>
6f8949e6c71a more robust ML pretty printing (cf. b6c527c64789);
wenzelm
parents: 44240
diff changeset
    43
      (case try Syntax.init_pretty_global thy of
6f8949e6c71a more robust ML pretty printing (cf. b6c527c64789);
wenzelm
parents: 44240
diff changeset
    44
        SOME ctxt => ctxt
6f8949e6c71a more robust ML pretty printing (cf. b6c527c64789);
wenzelm
parents: 44240
diff changeset
    45
      | NONE => Syntax.init_pretty_global thy0)
6f8949e6c71a more robust ML pretty printing (cf. b6c527c64789);
wenzelm
parents: 44240
diff changeset
    46
  | NONE => Syntax.init_pretty_global thy0);
44240
4b1a63678839 improved default context for ML toplevel pretty-printing;
wenzelm
parents: 44192
diff changeset
    47
30628
4078276bcace removed obsolete pprint operations;
wenzelm
parents: 30364
diff changeset
    48
fun pp_thm th =
30724
2e54e89bcce7 pretty_rule/print_results: no thm status here -- it is potentially slow and mostly uninformative/confusing as long as proofs are still unfinished;
wenzelm
parents: 30628
diff changeset
    49
  let
44240
4b1a63678839 improved default context for ML toplevel pretty-printing;
wenzelm
parents: 44192
diff changeset
    50
    val ctxt = default_context (Thm.theory_of_thm th);
30724
2e54e89bcce7 pretty_rule/print_results: no thm status here -- it is potentially slow and mostly uninformative/confusing as long as proofs are still unfinished;
wenzelm
parents: 30628
diff changeset
    51
    val flags = {quote = true, show_hyps = false, show_status = true};
44240
4b1a63678839 improved default context for ML toplevel pretty-printing;
wenzelm
parents: 44192
diff changeset
    52
  in Display.pretty_thm_raw ctxt flags th end;
20211
c7f907f41f7c moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 19482
diff changeset
    53
44240
4b1a63678839 improved default context for ML toplevel pretty-printing;
wenzelm
parents: 44192
diff changeset
    54
fun pp_typ thy T = Pretty.quote (Syntax.pretty_typ (default_context thy) T);
4b1a63678839 improved default context for ML toplevel pretty-printing;
wenzelm
parents: 44192
diff changeset
    55
fun pp_term thy t = Pretty.quote (Syntax.pretty_term (default_context thy) t);
30628
4078276bcace removed obsolete pprint operations;
wenzelm
parents: 30364
diff changeset
    56
4078276bcace removed obsolete pprint operations;
wenzelm
parents: 30364
diff changeset
    57
fun pp_ctyp cT = pp_typ (Thm.theory_of_ctyp cT) (Thm.typ_of cT);
4078276bcace removed obsolete pprint operations;
wenzelm
parents: 30364
diff changeset
    58
fun pp_cterm ct = pp_term (Thm.theory_of_cterm ct) (Thm.term_of ct);
20211
c7f907f41f7c moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 19482
diff changeset
    59
19432
cae7cc072994 added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
wenzelm
parents: 18799
diff changeset
    60
cae7cc072994 added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
wenzelm
parents: 18799
diff changeset
    61
(* theorems and theory *)
cae7cc072994 added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
wenzelm
parents: 18799
diff changeset
    62
33515
d066e8369a33 print_theorems: suppress concealed (global) facts, unless "!" option is given;
wenzelm
parents: 33389
diff changeset
    63
fun pretty_theorems_diff verbose prev_thys thy =
19432
cae7cc072994 added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
wenzelm
parents: 18799
diff changeset
    64
  let
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 41551
diff changeset
    65
    val pretty_fact = Proof_Context.pretty_fact (Proof_Context.init_global thy);
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39284
diff changeset
    66
    val facts = Global_Theory.facts_of thy;
33515
d066e8369a33 print_theorems: suppress concealed (global) facts, unless "!" option is given;
wenzelm
parents: 33389
diff changeset
    67
    val thmss =
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39284
diff changeset
    68
      Facts.dest_static (map Global_Theory.facts_of prev_thys) facts
33515
d066e8369a33 print_theorems: suppress concealed (global) facts, unless "!" option is given;
wenzelm
parents: 33389
diff changeset
    69
      |> not verbose ? filter_out (Facts.is_concealed facts o #1);
28210
c164d1892553 more procise printing of fact names;
wenzelm
parents: 28092
diff changeset
    70
  in Pretty.big_list "theorems:" (map #1 (sort_wrt (#1 o #2) (map (`pretty_fact) thmss))) end;
19432
cae7cc072994 added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
wenzelm
parents: 18799
diff changeset
    71
33515
d066e8369a33 print_theorems: suppress concealed (global) facts, unless "!" option is given;
wenzelm
parents: 33389
diff changeset
    72
fun print_theorems_diff verbose prev_thy thy =
d066e8369a33 print_theorems: suppress concealed (global) facts, unless "!" option is given;
wenzelm
parents: 33389
diff changeset
    73
  Pretty.writeln (pretty_theorems_diff verbose [prev_thy] thy);
19432
cae7cc072994 added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
wenzelm
parents: 18799
diff changeset
    74
33515
d066e8369a33 print_theorems: suppress concealed (global) facts, unless "!" option is given;
wenzelm
parents: 33389
diff changeset
    75
fun pretty_theorems verbose thy = pretty_theorems_diff verbose (Theory.parents_of thy) thy;
d066e8369a33 print_theorems: suppress concealed (global) facts, unless "!" option is given;
wenzelm
parents: 33389
diff changeset
    76
val print_theorems = Pretty.writeln oo pretty_theorems;
19432
cae7cc072994 added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
wenzelm
parents: 18799
diff changeset
    77
22872
d7189dc8939c tuned ProofDisplay.pretty_full_theory;
wenzelm
parents: 22335
diff changeset
    78
fun pretty_full_theory verbose thy =
33515
d066e8369a33 print_theorems: suppress concealed (global) facts, unless "!" option is given;
wenzelm
parents: 33389
diff changeset
    79
  Pretty.chunks (Display.pretty_full_theory verbose thy @ [pretty_theorems verbose thy]);
20621
29d57880ba00 'print_theory': bang option for full verbosity;
wenzelm
parents: 20311
diff changeset
    80
22872
d7189dc8939c tuned ProofDisplay.pretty_full_theory;
wenzelm
parents: 22335
diff changeset
    81
val print_theory = Pretty.writeln o pretty_full_theory false;
22335
6c4204df6863 pretty_full_theory: expose in signature.
aspinall
parents: 22095
diff changeset
    82
19432
cae7cc072994 added print_theorems/theory, print_theorems_diff (from pure_thy.ML);
wenzelm
parents: 18799
diff changeset
    83
17369
dec2ddbb3edf Printing of Isar proof elements etc.
wenzelm
parents:
diff changeset
    84
(* refinement rule *)
dec2ddbb3edf Printing of Isar proof elements etc.
wenzelm
parents:
diff changeset
    85
dec2ddbb3edf Printing of Isar proof elements etc.
wenzelm
parents:
diff changeset
    86
fun pretty_rule ctxt s thm =
dec2ddbb3edf Printing of Isar proof elements etc.
wenzelm
parents:
diff changeset
    87
  Pretty.block [Pretty.str (s ^ " attempt to solve goal by exported rule:"),
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30724
diff changeset
    88
    Pretty.fbrk, Display.pretty_thm_aux ctxt false thm];
17369
dec2ddbb3edf Printing of Isar proof elements etc.
wenzelm
parents:
diff changeset
    89
dec2ddbb3edf Printing of Isar proof elements etc.
wenzelm
parents:
diff changeset
    90
val string_of_rule = Pretty.string_of ooo pretty_rule;
dec2ddbb3edf Printing of Isar proof elements etc.
wenzelm
parents:
diff changeset
    91
dec2ddbb3edf Printing of Isar proof elements etc.
wenzelm
parents:
diff changeset
    92
49860
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
    93
(* goals *)
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
    94
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
    95
local
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
    96
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
    97
fun subgoals 0 = []
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
    98
  | subgoals 1 = [Pretty.brk 1, Pretty.str "(1 subgoal)"]
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
    99
  | 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
   100
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   101
in
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   102
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   103
fun pretty_goal_header goal =
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   104
  Pretty.block ([Pretty.command "goal"] @ subgoals (Thm.nprems_of goal) @ [Pretty.str ":"]);
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   105
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   106
end;
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   107
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   108
fun string_of_goal ctxt goal =
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   109
  Pretty.string_of (Pretty.chunks
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   110
    [pretty_goal_header goal, Goal_Display.pretty_goal {main = true, limit = false} ctxt goal]);
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   111
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   112
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   113
(* method_error *)
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   114
49866
619acbd72664 more proof method text position information;
wenzelm
parents: 49863
diff changeset
   115
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
   116
  let
49866
619acbd72664 more proof method text position information;
wenzelm
parents: 49863
diff changeset
   117
    val pr_header =
619acbd72664 more proof method text position information;
wenzelm
parents: 49863
diff changeset
   118
      "Failed to apply " ^ (if kind = "" then "" else kind ^ " ") ^
619acbd72664 more proof method text position information;
wenzelm
parents: 49863
diff changeset
   119
      "proof method" ^ Position.here pos ^ ":\n";
49860
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   120
    val pr_facts =
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   121
      if null facts then ""
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   122
      else
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   123
        Pretty.string_of
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   124
          (Pretty.chunks
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   125
            (Pretty.block [Pretty.command "using", Pretty.brk 1, Pretty.str "this:"] ::
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   126
              map (Display.pretty_thm ctxt) facts)) ^ "\n";
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   127
    val pr_goal = string_of_goal ctxt goal;
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   128
  in pr_header ^ pr_facts ^ pr_goal end);
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   129
9a0fe50e4534 further attempts to unify/simplify goal output;
wenzelm
parents: 46728
diff changeset
   130
17369
dec2ddbb3edf Printing of Isar proof elements etc.
wenzelm
parents:
diff changeset
   131
(* results *)
dec2ddbb3edf Printing of Isar proof elements etc.
wenzelm
parents:
diff changeset
   132
28092
5886e9359aa8 no pervasive bindings;
wenzelm
parents: 28087
diff changeset
   133
local
5886e9359aa8 no pervasive bindings;
wenzelm
parents: 28087
diff changeset
   134
41551
791b139a6c1e tuned markup;
wenzelm
parents: 39557
diff changeset
   135
fun pretty_fact_name (kind, "") = Pretty.command kind
791b139a6c1e tuned markup;
wenzelm
parents: 39557
diff changeset
   136
  | pretty_fact_name (kind, name) =
791b139a6c1e tuned markup;
wenzelm
parents: 39557
diff changeset
   137
      Pretty.block [Pretty.command kind, Pretty.brk 1,
791b139a6c1e tuned markup;
wenzelm
parents: 39557
diff changeset
   138
        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
   139
17369
dec2ddbb3edf Printing of Isar proof elements etc.
wenzelm
parents:
diff changeset
   140
fun pretty_facts ctxt =
41551
791b139a6c1e tuned markup;
wenzelm
parents: 39557
diff changeset
   141
  flat o (separate [Pretty.fbrk, Pretty.keyword "and", Pretty.str " "]) o
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 41551
diff changeset
   142
    map (single o Proof_Context.pretty_fact_aux ctxt false);
17369
dec2ddbb3edf Printing of Isar proof elements etc.
wenzelm
parents:
diff changeset
   143
27857
fdf43ffceae0 removed obsolete present_results;
wenzelm
parents: 27176
diff changeset
   144
in
fdf43ffceae0 removed obsolete present_results;
wenzelm
parents: 27176
diff changeset
   145
46728
85f8e3932712 display proof results as "state", to suppress odd squiggles in the Prover IDE (see also 9240be8c8c69);
wenzelm
parents: 45269
diff changeset
   146
fun print_results markup do_print ctxt ((kind, name), facts) =
33643
b275f26a638b eliminated obsolete "internal" kind -- collapsed to unspecific "";
wenzelm
parents: 33515
diff changeset
   147
  if not do_print orelse kind = "" then ()
28092
5886e9359aa8 no pervasive bindings;
wenzelm
parents: 28087
diff changeset
   148
  else if name = "" then
46728
85f8e3932712 display proof results as "state", to suppress odd squiggles in the Prover IDE (see also 9240be8c8c69);
wenzelm
parents: 45269
diff changeset
   149
    (Pretty.writeln o Pretty.mark markup)
85f8e3932712 display proof results as "state", to suppress odd squiggles in the Prover IDE (see also 9240be8c8c69);
wenzelm
parents: 45269
diff changeset
   150
      (Pretty.block (Pretty.command kind :: Pretty.brk 1 :: pretty_facts ctxt facts))
85f8e3932712 display proof results as "state", to suppress odd squiggles in the Prover IDE (see also 9240be8c8c69);
wenzelm
parents: 45269
diff changeset
   151
  else
85f8e3932712 display proof results as "state", to suppress odd squiggles in the Prover IDE (see also 9240be8c8c69);
wenzelm
parents: 45269
diff changeset
   152
    (Pretty.writeln o Pretty.mark markup)
85f8e3932712 display proof results as "state", to suppress odd squiggles in the Prover IDE (see also 9240be8c8c69);
wenzelm
parents: 45269
diff changeset
   153
      (case facts of
85f8e3932712 display proof results as "state", to suppress odd squiggles in the Prover IDE (see also 9240be8c8c69);
wenzelm
parents: 45269
diff changeset
   154
        [fact] => Pretty.blk (1, [pretty_fact_name (kind, name), Pretty.fbrk,
85f8e3932712 display proof results as "state", to suppress odd squiggles in the Prover IDE (see also 9240be8c8c69);
wenzelm
parents: 45269
diff changeset
   155
          Proof_Context.pretty_fact_aux ctxt false fact])
85f8e3932712 display proof results as "state", to suppress odd squiggles in the Prover IDE (see also 9240be8c8c69);
wenzelm
parents: 45269
diff changeset
   156
      | _ => Pretty.blk (1, [pretty_fact_name (kind, name), Pretty.fbrk,
85f8e3932712 display proof results as "state", to suppress odd squiggles in the Prover IDE (see also 9240be8c8c69);
wenzelm
parents: 45269
diff changeset
   157
          Pretty.blk (1, Pretty.str "(" :: pretty_facts ctxt facts @ [Pretty.str ")"])]));
17369
dec2ddbb3edf Printing of Isar proof elements etc.
wenzelm
parents:
diff changeset
   158
dec2ddbb3edf Printing of Isar proof elements etc.
wenzelm
parents:
diff changeset
   159
end;
dec2ddbb3edf Printing of Isar proof elements etc.
wenzelm
parents:
diff changeset
   160
20889
f625a65bfed5 added pretty_consts (from specification.ML);
wenzelm
parents: 20621
diff changeset
   161
f625a65bfed5 added pretty_consts (from specification.ML);
wenzelm
parents: 20621
diff changeset
   162
(* consts *)
f625a65bfed5 added pretty_consts (from specification.ML);
wenzelm
parents: 20621
diff changeset
   163
f625a65bfed5 added pretty_consts (from specification.ML);
wenzelm
parents: 20621
diff changeset
   164
local
f625a65bfed5 added pretty_consts (from specification.ML);
wenzelm
parents: 20621
diff changeset
   165
f625a65bfed5 added pretty_consts (from specification.ML);
wenzelm
parents: 20621
diff changeset
   166
fun pretty_var ctxt (x, T) =
f625a65bfed5 added pretty_consts (from specification.ML);
wenzelm
parents: 20621
diff changeset
   167
  Pretty.block [Pretty.str x, Pretty.str " ::", Pretty.brk 1,
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 22872
diff changeset
   168
    Pretty.quote (Syntax.pretty_typ ctxt T)];
20889
f625a65bfed5 added pretty_consts (from specification.ML);
wenzelm
parents: 20621
diff changeset
   169
f625a65bfed5 added pretty_consts (from specification.ML);
wenzelm
parents: 20621
diff changeset
   170
fun pretty_vars ctxt kind vs = Pretty.big_list kind (map (pretty_var ctxt) vs);
f625a65bfed5 added pretty_consts (from specification.ML);
wenzelm
parents: 20621
diff changeset
   171
f625a65bfed5 added pretty_consts (from specification.ML);
wenzelm
parents: 20621
diff changeset
   172
fun pretty_consts ctxt pred cs =
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 41551
diff changeset
   173
  (case filter pred (#1 (Proof_Context.inferred_fixes ctxt)) of
20889
f625a65bfed5 added pretty_consts (from specification.ML);
wenzelm
parents: 20621
diff changeset
   174
    [] => pretty_vars ctxt "constants" cs
f625a65bfed5 added pretty_consts (from specification.ML);
wenzelm
parents: 20621
diff changeset
   175
  | ps => Pretty.chunks [pretty_vars ctxt "parameters" ps, pretty_vars ctxt "constants" cs]);
f625a65bfed5 added pretty_consts (from specification.ML);
wenzelm
parents: 20621
diff changeset
   176
44192
a32ca9165928 less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents: 42717
diff changeset
   177
in
a32ca9165928 less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents: 42717
diff changeset
   178
a32ca9165928 less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents: 42717
diff changeset
   179
fun print_consts do_print ctxt pred cs =
a32ca9165928 less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents: 42717
diff changeset
   180
  if not do_print orelse null cs then ()
a32ca9165928 less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents: 42717
diff changeset
   181
  else Pretty.writeln (pretty_consts ctxt pred cs);
a32ca9165928 less verbosity in batch mode -- spam reduction and notable performance improvement;
wenzelm
parents: 42717
diff changeset
   182
20889
f625a65bfed5 added pretty_consts (from specification.ML);
wenzelm
parents: 20621
diff changeset
   183
end;
f625a65bfed5 added pretty_consts (from specification.ML);
wenzelm
parents: 20621
diff changeset
   184
17369
dec2ddbb3edf Printing of Isar proof elements etc.
wenzelm
parents:
diff changeset
   185
end;