| author | wenzelm | 
| Fri, 29 Dec 2006 17:24:45 +0100 | |
| changeset 21934 | ba683b0b2456 | 
| parent 21721 | 908a93216f00 | 
| child 22846 | fb79144af9a3 | 
| permissions | -rw-r--r-- | 
| 1591 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 1 | (* Title: Pure/display.ML | 
| 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 4 | Copyright 1993 University of Cambridge | 
| 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 5 | |
| 19420 | 6 | Printing of theorems, goals, results etc. | 
| 1591 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 7 | *) | 
| 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 8 | |
| 11883 | 9 | signature BASIC_DISPLAY = | 
| 10 | sig | |
| 12081 | 11 | val goals_limit: int ref | 
| 11883 | 12 | val show_hyps: bool ref | 
| 13 | val show_tags: bool ref | |
| 14 | val string_of_thm: thm -> string | |
| 15 | val print_thm: thm -> unit | |
| 16 | val print_thms: thm list -> unit | |
| 17 | val prth: thm -> thm | |
| 18 | val prthq: thm Seq.seq -> thm Seq.seq | |
| 19 | val prths: thm list -> thm list | |
| 20 | val string_of_ctyp: ctyp -> string | |
| 21 | val print_ctyp: ctyp -> unit | |
| 22 | val string_of_cterm: cterm -> string | |
| 23 | val print_cterm: cterm -> unit | |
| 24 | val print_syntax: theory -> unit | |
| 25 | val show_consts: bool ref | |
| 26 | end; | |
| 27 | ||
| 1591 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 28 | signature DISPLAY = | 
| 4950 | 29 | sig | 
| 11883 | 30 | include BASIC_DISPLAY | 
| 19591 | 31 | val raw_string_of_sort: sort -> string | 
| 32 | val raw_string_of_typ: typ -> string | |
| 33 | val raw_string_of_term: term -> string | |
| 14876 | 34 | val pretty_flexpair: Pretty.pp -> term * term -> Pretty.T | 
| 17447 | 35 | val pretty_thm_aux: Pretty.pp -> bool -> bool -> term list -> thm -> Pretty.T | 
| 11883 | 36 | val pretty_thm: thm -> Pretty.T | 
| 37 | val pretty_thms: thm list -> Pretty.T | |
| 16437 
aa87badf7a3c
removed pretty_theory, pprint_theory (see context.ML or thy_info.ML);
 wenzelm parents: 
16364diff
changeset | 38 | val pretty_thm_sg: theory -> thm -> Pretty.T | 
| 
aa87badf7a3c
removed pretty_theory, pprint_theory (see context.ML or thy_info.ML);
 wenzelm parents: 
16364diff
changeset | 39 | val pretty_thms_sg: theory -> thm list -> Pretty.T | 
| 11883 | 40 | val pretty_ctyp: ctyp -> Pretty.T | 
| 41 | val pretty_cterm: cterm -> Pretty.T | |
| 20629 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 wenzelm parents: 
20226diff
changeset | 42 | val pretty_full_theory: bool -> theory -> Pretty.T list | 
| 14876 | 43 | val pretty_goals_aux: Pretty.pp -> string -> bool * bool -> int -> thm -> Pretty.T list | 
| 11883 | 44 | val pretty_goals: int -> thm -> Pretty.T list | 
| 45 | val print_goals: int -> thm -> unit | |
| 46 | val current_goals_markers: (string * string * string) ref | |
| 12418 | 47 | val pretty_current_goals: int -> int -> thm -> Pretty.T list | 
| 48 | val print_current_goals_default: int -> int -> thm -> unit | |
| 49 | val print_current_goals_fn: (int -> int -> thm -> unit) ref | |
| 4950 | 50 | end; | 
| 1591 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 51 | |
| 4950 | 52 | structure Display: DISPLAY = | 
| 1591 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 53 | struct | 
| 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 54 | |
| 19591 | 55 | (** raw output **) | 
| 56 | ||
| 57 | val raw_string_of_sort = Sign.string_of_sort ProtoPure.thy; | |
| 58 | val raw_string_of_typ = Sign.string_of_typ ProtoPure.thy; | |
| 59 | val raw_string_of_term = Sign.string_of_term ProtoPure.thy; | |
| 60 | ||
| 61 | ||
| 11883 | 62 | |
| 6087 | 63 | (** print thm **) | 
| 64 | ||
| 16534 | 65 | val goals_limit = ref 10; (*max number of goals to print*) | 
| 12831 
a2a3896f9c48
reset show_hyps by default (in accordance to existing Isar practice);
 wenzelm parents: 
12418diff
changeset | 66 | val show_hyps = ref false; (*false: print meta-hypotheses as dots*) | 
| 11883 | 67 | val show_tags = ref false; (*false: suppress tags*) | 
| 6087 | 68 | |
| 14598 
7009f59711e3
Replaced quote by Library.quote, since quote now refers to Symbol.quote
 berghofe parents: 
14472diff
changeset | 69 | fun pretty_tag (name, args) = Pretty.strs (name :: map Library.quote args); | 
| 6087 | 70 | val pretty_tags = Pretty.list "[" "]" o map pretty_tag; | 
| 1591 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 71 | |
| 14876 | 72 | fun pretty_flexpair pp (t, u) = Pretty.block | 
| 73 | [Pretty.term pp t, Pretty.str " =?=", Pretty.brk 1, Pretty.term pp u]; | |
| 13658 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 berghofe parents: 
12831diff
changeset | 74 | |
| 17447 | 75 | fun pretty_thm_aux pp quote show_hyps' asms raw_th = | 
| 1591 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 76 | let | 
| 17447 | 77 | val th = Thm.strip_shyps raw_th; | 
| 16290 | 78 |     val {hyps, tpairs, prop, der = (ora, _), ...} = Thm.rep_thm th;
 | 
| 6087 | 79 | val xshyps = Thm.extra_shyps th; | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
20629diff
changeset | 80 | val tags = Thm.get_tags th; | 
| 6087 | 81 | |
| 13658 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 berghofe parents: 
12831diff
changeset | 82 | val q = if quote then Pretty.quote else I; | 
| 14876 | 83 | val prt_term = q o (Pretty.term pp); | 
| 6279 | 84 | |
| 19300 | 85 | val hyps' = if ! show_hyps then hyps else subtract (op aconv) asms hyps; | 
| 17475 | 86 | val ora' = ora andalso (! show_hyps orelse not (! quick_and_dirty)); | 
| 17447 | 87 | val hlen = length xshyps + length hyps' + length tpairs; | 
| 1591 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 88 | val hsymbs = | 
| 17475 | 89 | if hlen = 0 andalso not ora' then [] | 
| 17447 | 90 | else if ! show_hyps orelse show_hyps' then | 
| 1591 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 91 | [Pretty.brk 2, Pretty.list "[" "]" | 
| 17447 | 92 | (map (q o pretty_flexpair pp) tpairs @ map prt_term hyps' @ | 
| 14876 | 93 | map (Pretty.sort pp) xshyps @ | 
| 17475 | 94 | (if ora' then [Pretty.str "!"] else []))] | 
| 6889 | 95 |       else [Pretty.brk 2, Pretty.str ("[" ^ implode (replicate hlen ".") ^
 | 
| 17475 | 96 | (if ora' then "!" else "") ^ "]")]; | 
| 6087 | 97 | val tsymbs = | 
| 98 | if null tags orelse not (! show_tags) then [] | |
| 99 | else [Pretty.brk 1, pretty_tags tags]; | |
| 6279 | 100 | in Pretty.block (prt_term prop :: (hsymbs @ tsymbs)) end; | 
| 1591 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 101 | |
| 20226 | 102 | fun pretty_thm th = | 
| 103 | pretty_thm_aux (Sign.pp (Thm.theory_of_thm th)) true false [] th; | |
| 6889 | 104 | |
| 1591 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 105 | val string_of_thm = Pretty.string_of o pretty_thm; | 
| 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 106 | |
| 6087 | 107 | fun pretty_thms [th] = pretty_thm th | 
| 108 | | pretty_thms ths = Pretty.block (Pretty.fbreaks (map pretty_thm ths)); | |
| 1591 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 109 | |
| 16437 
aa87badf7a3c
removed pretty_theory, pprint_theory (see context.ML or thy_info.ML);
 wenzelm parents: 
16364diff
changeset | 110 | val pretty_thm_sg = pretty_thm oo Thm.transfer; | 
| 
aa87badf7a3c
removed pretty_theory, pprint_theory (see context.ML or thy_info.ML);
 wenzelm parents: 
16364diff
changeset | 111 | val pretty_thms_sg = pretty_thms oo (map o Thm.transfer); | 
| 10010 | 112 | |
| 6087 | 113 | |
| 114 | (* top-level commands for printing theorems *) | |
| 115 | ||
| 116 | val print_thm = Pretty.writeln o pretty_thm; | |
| 117 | val print_thms = Pretty.writeln o pretty_thms; | |
| 1591 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 118 | |
| 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 119 | fun prth th = (print_thm th; th); | 
| 16534 | 120 | fun prthq thq = (Seq.print (K print_thm) 100000 thq; thq); | 
| 121 | fun prths ths = (prthq (Seq.of_list ths); ths); | |
| 1591 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 122 | |
| 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 123 | |
| 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 124 | (* other printing commands *) | 
| 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 125 | |
| 4126 | 126 | fun pretty_ctyp cT = | 
| 16437 
aa87badf7a3c
removed pretty_theory, pprint_theory (see context.ML or thy_info.ML);
 wenzelm parents: 
16364diff
changeset | 127 |   let val {thy, T, ...} = rep_ctyp cT in Sign.pretty_typ thy T end;
 | 
| 4126 | 128 | |
| 1591 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 129 | fun string_of_ctyp cT = | 
| 16437 
aa87badf7a3c
removed pretty_theory, pprint_theory (see context.ML or thy_info.ML);
 wenzelm parents: 
16364diff
changeset | 130 |   let val {thy, T, ...} = rep_ctyp cT in Sign.string_of_typ thy T end;
 | 
| 1591 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 131 | |
| 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 132 | val print_ctyp = writeln o string_of_ctyp; | 
| 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 133 | |
| 3547 | 134 | fun pretty_cterm ct = | 
| 16437 
aa87badf7a3c
removed pretty_theory, pprint_theory (see context.ML or thy_info.ML);
 wenzelm parents: 
16364diff
changeset | 135 |   let val {thy, t, ...} = rep_cterm ct in Sign.pretty_term thy t end;
 | 
| 3547 | 136 | |
| 1591 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 137 | fun string_of_cterm ct = | 
| 16437 
aa87badf7a3c
removed pretty_theory, pprint_theory (see context.ML or thy_info.ML);
 wenzelm parents: 
16364diff
changeset | 138 |   let val {thy, t, ...} = rep_cterm ct in Sign.string_of_term thy t end;
 | 
| 1591 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 139 | |
| 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 140 | val print_cterm = writeln o string_of_cterm; | 
| 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 141 | |
| 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 142 | |
| 4250 | 143 | |
| 144 | (** print theory **) | |
| 1591 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 145 | |
| 16437 
aa87badf7a3c
removed pretty_theory, pprint_theory (see context.ML or thy_info.ML);
 wenzelm parents: 
16364diff
changeset | 146 | val print_syntax = Syntax.print_syntax o Sign.syn_of; | 
| 4498 | 147 | |
| 148 | ||
| 11883 | 149 | (* pretty_full_theory *) | 
| 4250 | 150 | |
| 20629 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 wenzelm parents: 
20226diff
changeset | 151 | fun pretty_full_theory verbose thy = | 
| 4250 | 152 | let | 
| 16437 
aa87badf7a3c
removed pretty_theory, pprint_theory (see context.ML or thy_info.ML);
 wenzelm parents: 
16364diff
changeset | 153 | fun prt_cls c = Sign.pretty_sort thy [c]; | 
| 
aa87badf7a3c
removed pretty_theory, pprint_theory (see context.ML or thy_info.ML);
 wenzelm parents: 
16364diff
changeset | 154 | fun prt_sort S = Sign.pretty_sort thy S; | 
| 19521 | 155 | fun prt_arity t (c, (_, Ss)) = Sign.pretty_arity thy (t, Ss, [c]); | 
| 16437 
aa87badf7a3c
removed pretty_theory, pprint_theory (see context.ML or thy_info.ML);
 wenzelm parents: 
16364diff
changeset | 156 | fun prt_typ ty = Pretty.quote (Sign.pretty_typ thy ty); | 
| 19806 | 157 | val prt_typ_no_tvars = prt_typ o Logic.unvarifyT; | 
| 16437 
aa87badf7a3c
removed pretty_theory, pprint_theory (see context.ML or thy_info.ML);
 wenzelm parents: 
16364diff
changeset | 158 | fun prt_term t = Pretty.quote (Sign.pretty_term thy t); | 
| 18936 | 159 | val prt_term_no_vars = prt_term o Logic.unvarify; | 
| 19698 | 160 | fun prt_const (c, ty) = [Pretty.str c, Pretty.str " ::", Pretty.brk 1, prt_typ_no_tvars ty]; | 
| 161 | val prt_const' = Defs.pretty_const (Sign.pp thy); | |
| 4250 | 162 | |
| 14794 | 163 | fun pretty_classrel (c, []) = prt_cls c | 
| 164 | | pretty_classrel (c, cs) = Pretty.block | |
| 165 | (prt_cls c :: Pretty.str " <" :: Pretty.brk 1 :: | |
| 166 | Pretty.commas (map prt_cls cs)); | |
| 4250 | 167 | |
| 168 | fun pretty_default S = Pretty.block | |
| 14794 | 169 | [Pretty.str "default sort:", Pretty.brk 1, prt_sort S]; | 
| 4250 | 170 | |
| 15531 | 171 | fun pretty_witness NONE = Pretty.str "universal non-emptiness witness: -" | 
| 172 | | pretty_witness (SOME (T, S)) = Pretty.block | |
| 14794 | 173 | [Pretty.str "universal non-emptiness witness:", Pretty.brk 1, | 
| 174 | prt_typ_no_tvars T, Pretty.str " ::", Pretty.brk 1, prt_sort S]; | |
| 175 | ||
| 14996 | 176 | val tfrees = map (fn v => TFree (v, [])); | 
| 16364 | 177 | fun pretty_type syn (t, (Type.LogicalType n, _)) = | 
| 15531 | 178 | if syn then NONE | 
| 20076 | 179 | else SOME (prt_typ (Type (t, tfrees (Name.invent_list [] "'a" n)))) | 
| 16364 | 180 | | pretty_type syn (t, (Type.Abbreviation (vs, U, syn'), _)) = | 
| 15531 | 181 | if syn <> syn' then NONE | 
| 182 | else SOME (Pretty.block | |
| 14996 | 183 | [prt_typ (Type (t, tfrees vs)), Pretty.str " =", Pretty.brk 1, prt_typ U]) | 
| 16364 | 184 | | pretty_type syn (t, (Type.Nonterminal, _)) = | 
| 15531 | 185 | if not syn then NONE | 
| 186 | else SOME (prt_typ (Type (t, []))); | |
| 4250 | 187 | |
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19420diff
changeset | 188 | val pretty_arities = maps (fn (t, ars) => map (prt_arity t) ars); | 
| 14223 
0ee05eef881b
Added support for making constants final, that is, ensuring that no
 skalberg parents: 
14178diff
changeset | 189 | |
| 19698 | 190 | fun pretty_abbrev (c, (ty, t)) = Pretty.block | 
| 191 | (prt_const (c, ty) @ [Pretty.str " ==", Pretty.brk 1, prt_term_no_vars t]); | |
| 192 | ||
| 193 | fun pretty_axm (a, t) = Pretty.block [Pretty.str (a ^ ":"), Pretty.brk 1, prt_term_no_vars t]; | |
| 19365 | 194 | |
| 19702 | 195 | fun pretty_finals reds = Pretty.block | 
| 19703 | 196 | (Pretty.str "final:" :: Pretty.brk 1 :: Pretty.commas (map (prt_const' o fst) reds)); | 
| 19702 | 197 | |
| 19698 | 198 | fun pretty_reduct (lhs, rhs) = Pretty.block | 
| 19702 | 199 | ([prt_const' lhs, Pretty.str " ->", Pretty.brk 2] @ | 
| 200 | Pretty.commas (map prt_const' (sort_wrt #1 rhs))); | |
| 4250 | 201 | |
| 19698 | 202 | fun pretty_restrict (const, name) = | 
| 203 |       Pretty.block ([prt_const' const, Pretty.brk 2, Pretty.str ("(from " ^ quote name ^ ")")]);
 | |
| 8720 | 204 | |
| 19698 | 205 |     val {axioms, defs, oracles} = Theory.rep_theory thy;
 | 
| 206 |     val {restricts, reducts} = Defs.dest defs;
 | |
| 18061 | 207 |     val {naming, syn = _, tsig, consts} = Sign.rep_sg thy;
 | 
| 18936 | 208 |     val {constants, constraints} = Consts.dest consts;
 | 
| 19698 | 209 | val extern_const = NameSpace.extern (#1 constants); | 
| 210 |     val {classes, default, types, witness, ...} = Type.rep_tsig tsig;
 | |
| 211 | val (class_space, class_algebra) = classes; | |
| 19642 | 212 |     val {classes, arities} = Sorts.rep_algebra class_algebra;
 | 
| 14996 | 213 | |
| 20629 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 wenzelm parents: 
20226diff
changeset | 214 | fun prune xs = if verbose then xs else filter_out (can Name.dest_internal o fst) xs; | 
| 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 wenzelm parents: 
20226diff
changeset | 215 | fun prune' xs = if verbose then xs else filter_out (can Name.dest_internal o fst o fst) xs; | 
| 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 wenzelm parents: 
20226diff
changeset | 216 | fun dest_table tab = prune (NameSpace.dest_table tab); | 
| 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 wenzelm parents: 
20226diff
changeset | 217 | fun extern_table tab = prune (NameSpace.extern_table tab); | 
| 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 wenzelm parents: 
20226diff
changeset | 218 | |
| 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 wenzelm parents: 
20226diff
changeset | 219 | val clsses = dest_table (class_space, Symtab.make (Graph.dest classes)); | 
| 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 wenzelm parents: 
20226diff
changeset | 220 | val tdecls = dest_table types; | 
| 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 wenzelm parents: 
20226diff
changeset | 221 | val arties = dest_table (Sign.type_space thy, arities); | 
| 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 wenzelm parents: 
20226diff
changeset | 222 | val cnsts = extern_table constants; | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19420diff
changeset | 223 | val log_cnsts = map_filter (fn (c, (ty, NONE)) => SOME (c, ty) | _ => NONE) cnsts; | 
| 21721 | 224 | val abbrevs = map_filter (fn (c, (ty, SOME (t, _))) => SOME (c, (ty, t)) | _ => NONE) cnsts; | 
| 20629 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 wenzelm parents: 
20226diff
changeset | 225 | val cnstrs = extern_table constraints; | 
| 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 wenzelm parents: 
20226diff
changeset | 226 | val axms = extern_table axioms; | 
| 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 wenzelm parents: 
20226diff
changeset | 227 | val oras = extern_table oracles; | 
| 19698 | 228 | |
| 20629 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 wenzelm parents: 
20226diff
changeset | 229 | val (reds0, (reds1, reds2)) = prune' reducts |> map (fn (lhs, rhs) => | 
| 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 wenzelm parents: 
20226diff
changeset | 230 | (apfst extern_const lhs, map (apfst extern_const) (prune rhs))) | 
| 19702 | 231 | |> sort_wrt (#1 o #1) | 
| 232 | |> List.partition (null o #2) | |
| 233 | ||> List.partition (Defs.plain_args o #2 o #1); | |
| 19698 | 234 | val rests = restricts |> map (apfst (apfst extern_const)) |> sort_wrt (#1 o #1); | 
| 4250 | 235 | in | 
| 20629 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 wenzelm parents: 
20226diff
changeset | 236 |     [Pretty.strs ("names:" :: Context.names_of thy)] @
 | 
| 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 wenzelm parents: 
20226diff
changeset | 237 | (if verbose then | 
| 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 wenzelm parents: 
20226diff
changeset | 238 |       [Pretty.strs ("theory data:" :: Context.theory_data_of thy),
 | 
| 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 wenzelm parents: 
20226diff
changeset | 239 |        Pretty.strs ("proof data:" :: Context.proof_data_of thy)] else []) @
 | 
| 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 wenzelm parents: 
20226diff
changeset | 240 | [Pretty.strs ["name prefix:", NameSpace.path_of naming], | 
| 16364 | 241 | Pretty.big_list "classes:" (map pretty_classrel clsses), | 
| 8720 | 242 | pretty_default default, | 
| 14794 | 243 | pretty_witness witness, | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19420diff
changeset | 244 | Pretty.big_list "syntactic types:" (map_filter (pretty_type true) tdecls), | 
| 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19420diff
changeset | 245 | Pretty.big_list "logical types:" (map_filter (pretty_type false) tdecls), | 
| 16534 | 246 | Pretty.big_list "type arities:" (pretty_arities arties), | 
| 19698 | 247 | Pretty.big_list "logical consts:" (map (Pretty.block o prt_const) log_cnsts), | 
| 19365 | 248 | Pretty.big_list "abbreviations:" (map pretty_abbrev abbrevs), | 
| 19698 | 249 | Pretty.big_list "constraints:" (map (Pretty.block o prt_const) cnstrs), | 
| 250 | Pretty.big_list "axioms:" (map pretty_axm axms), | |
| 19702 | 251 |       Pretty.strs ("oracles:" :: (map #1 oras)),
 | 
| 19698 | 252 | Pretty.big_list "definitions:" | 
| 19702 | 253 | [pretty_finals reds0, | 
| 254 | Pretty.big_list "non-overloaded:" (map pretty_reduct reds1), | |
| 255 | Pretty.big_list "overloaded:" (map pretty_reduct reds2), | |
| 256 | Pretty.big_list "pattern restrictions:" (map pretty_restrict rests)]] | |
| 4250 | 257 | end; | 
| 258 | ||
| 259 | ||
| 11883 | 260 | |
| 261 | (** print_goals **) | |
| 262 | ||
| 263 | (* print_goals etc. *) | |
| 264 | ||
| 16534 | 265 | val show_consts = ref false; (*true: show consts with types in proof state output*) | 
| 3851 
fe9932a7cd46
print_goals: optional output of const types (set show_consts);
 wenzelm parents: 
3811diff
changeset | 266 | |
| 1591 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 267 | |
| 11883 | 268 | (*print thm A1,...,An/B in "goal style" -- premises as numbered subgoals*) | 
| 269 | ||
| 270 | local | |
| 271 | ||
| 18980 | 272 | fun ins_entry (x, y) = | 
| 273 | AList.default (op =) (x, []) #> | |
| 274 | AList.map_entry (op =) x (insert (op =) y); | |
| 12081 | 275 | |
| 18980 | 276 | val add_consts = Term.fold_aterms | 
| 277 | (fn Const (c, T) => ins_entry (T, (c, T)) | |
| 278 | | _ => I); | |
| 11883 | 279 | |
| 18980 | 280 | val add_vars = Term.fold_aterms | 
| 281 | (fn Free (x, T) => ins_entry (T, (x, ~1)) | |
| 282 | | Var (xi, T) => ins_entry (T, xi) | |
| 283 | | _ => I); | |
| 11883 | 284 | |
| 18980 | 285 | val add_varsT = Term.fold_atyps | 
| 286 | (fn TFree (x, S) => ins_entry (S, (x, ~1)) | |
| 287 | | TVar (xi, S) => ins_entry (S, xi) | |
| 288 | | _ => I); | |
| 11883 | 289 | |
| 16490 | 290 | fun sort_idxs vs = map (apsnd (sort (prod_ord string_ord int_ord))) vs; | 
| 12081 | 291 | fun sort_cnsts cs = map (apsnd (sort_wrt fst)) cs; | 
| 292 | ||
| 18980 | 293 | fun consts_of t = sort_cnsts (add_consts t []); | 
| 294 | fun vars_of t = sort_idxs (add_vars t []); | |
| 295 | fun varsT_of t = rev (sort_idxs (Term.fold_types add_varsT t [])); | |
| 12081 | 296 | |
| 297 | in | |
| 11883 | 298 | |
| 14876 | 299 | fun pretty_goals_aux pp begin_goal (msg, main) maxgoals state = | 
| 12081 | 300 | let | 
| 301 | fun prt_atoms prt prtT (X, xs) = Pretty.block | |
| 302 | [Pretty.block (Pretty.commas (map prt xs)), Pretty.str " ::", | |
| 303 | Pretty.brk 1, prtT X]; | |
| 11883 | 304 | |
| 14876 | 305 | fun prt_var (x, ~1) = Pretty.term pp (Syntax.free x) | 
| 306 | | prt_var xi = Pretty.term pp (Syntax.var xi); | |
| 12081 | 307 | |
| 14876 | 308 | fun prt_varT (x, ~1) = Pretty.typ pp (TFree (x, [])) | 
| 309 | | prt_varT xi = Pretty.typ pp (TVar (xi, [])); | |
| 12081 | 310 | |
| 14876 | 311 | val prt_consts = prt_atoms (Pretty.term pp o Const) (Pretty.typ pp); | 
| 312 | val prt_vars = prt_atoms prt_var (Pretty.typ pp); | |
| 313 | val prt_varsT = prt_atoms prt_varT (Pretty.sort pp); | |
| 11883 | 314 | |
| 315 | ||
| 12081 | 316 | fun pretty_list _ _ [] = [] | 
| 317 | | pretty_list name prt lst = [Pretty.big_list name (map prt lst)]; | |
| 11883 | 318 | |
| 14876 | 319 | fun pretty_subgoal (n, A) = Pretty.blk (0, | 
| 320 | [Pretty.str (begin_goal ^ " " ^ string_of_int n ^ ". "), Pretty.term pp A]); | |
| 12081 | 321 | fun pretty_subgoals As = map pretty_subgoal (1 upto length As ~~ As); | 
| 11883 | 322 | |
| 14876 | 323 | val pretty_ffpairs = pretty_list "flex-flex pairs:" (pretty_flexpair pp); | 
| 11883 | 324 | |
| 12081 | 325 | val pretty_consts = pretty_list "constants:" prt_consts o consts_of; | 
| 326 | val pretty_vars = pretty_list "variables:" prt_vars o vars_of; | |
| 327 | val pretty_varsT = pretty_list "type variables:" prt_varsT o varsT_of; | |
| 11883 | 328 | |
| 329 | ||
| 16290 | 330 |     val {prop, tpairs, ...} = Thm.rep_thm state;
 | 
| 13658 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 berghofe parents: 
12831diff
changeset | 331 | val (As, B) = Logic.strip_horn prop; | 
| 12081 | 332 | val ngoals = length As; | 
| 11883 | 333 | |
| 12081 | 334 | fun pretty_gs (types, sorts) = | 
| 14876 | 335 | (if main then [Pretty.term pp B] else []) @ | 
| 12081 | 336 | (if ngoals = 0 then [Pretty.str "No subgoals!"] | 
| 337 | else if ngoals > maxgoals then | |
| 15570 | 338 | pretty_subgoals (Library.take (maxgoals, As)) @ | 
| 12081 | 339 |           (if msg then [Pretty.str ("A total of " ^ string_of_int ngoals ^ " subgoals...")]
 | 
| 340 | else []) | |
| 341 | else pretty_subgoals As) @ | |
| 342 | pretty_ffpairs tpairs @ | |
| 343 | (if ! show_consts then pretty_consts prop else []) @ | |
| 344 | (if types then pretty_vars prop else []) @ | |
| 345 | (if sorts then pretty_varsT prop else []); | |
| 346 | in | |
| 347 | setmp show_no_free_types true | |
| 14178 
14a12da7288e
Makes interactive proof scripting recognize the show_all_types flag.
 skalberg parents: 
13658diff
changeset | 348 | (setmp show_types (! show_types orelse ! show_sorts orelse ! show_all_types) | 
| 12081 | 349 | (setmp show_sorts false pretty_gs)) | 
| 14178 
14a12da7288e
Makes interactive proof scripting recognize the show_all_types flag.
 skalberg parents: 
13658diff
changeset | 350 | (! show_types orelse ! show_sorts orelse ! show_all_types, ! show_sorts) | 
| 11883 | 351 | end; | 
| 352 | ||
| 12081 | 353 | fun pretty_goals_marker bg n th = | 
| 16437 
aa87badf7a3c
removed pretty_theory, pprint_theory (see context.ML or thy_info.ML);
 wenzelm parents: 
16364diff
changeset | 354 | pretty_goals_aux (Sign.pp (Thm.theory_of_thm th)) bg (true, true) n th; | 
| 12081 | 355 | |
| 356 | val pretty_goals = pretty_goals_marker ""; | |
| 12418 | 357 | val print_goals = (Pretty.writeln o Pretty.chunks) oo pretty_goals_marker ""; | 
| 12081 | 358 | |
| 1591 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 359 | end; | 
| 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 paulson parents: diff
changeset | 360 | |
| 11883 | 361 | |
| 362 | (* print_current_goals *) | |
| 363 | ||
| 364 | val current_goals_markers = ref ("", "", "");
 | |
| 365 | ||
| 12418 | 366 | fun pretty_current_goals n m th = | 
| 11883 | 367 | let | 
| 368 | val ref (begin_state, end_state, begin_goal) = current_goals_markers; | |
| 369 | val ngoals = nprems_of th; | |
| 370 | in | |
| 12418 | 371 | (if begin_state = "" then [] else [Pretty.str begin_state]) @ | 
| 372 |     [Pretty.str ("Level " ^ string_of_int n ^
 | |
| 11883 | 373 |       (if ngoals > 0 then " (" ^ string_of_int ngoals ^ " subgoal" ^
 | 
| 374 | (if ngoals <> 1 then "s" else "") ^ ")" | |
| 12418 | 375 | else ""))] @ | 
| 376 | pretty_goals_marker begin_goal m th @ | |
| 377 | (if end_state = "" then [] else [Pretty.str end_state]) | |
| 11883 | 378 | end; | 
| 379 | ||
| 12418 | 380 | fun print_current_goals_default n m th = | 
| 381 | Pretty.writeln (Pretty.chunks (pretty_current_goals n m th)); | |
| 382 | ||
| 11883 | 383 | val print_current_goals_fn = ref print_current_goals_default; | 
| 384 | ||
| 385 | end; | |
| 386 | ||
| 387 | structure BasicDisplay: BASIC_DISPLAY = Display; | |
| 388 | open BasicDisplay; |