author | huffman |
Wed, 08 Jun 2005 01:40:39 +0200 | |
changeset 16319 | 1ff2965cc2e7 |
parent 16290 | e661e37a4d50 |
child 16334 | b773132e3715 |
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 |
|
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
6 |
Printing of theories, theorems, etc. |
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 |
14876 | 31 |
val pretty_flexpair: Pretty.pp -> term * term -> Pretty.T |
32 |
val pretty_thm_aux: Pretty.pp -> bool -> thm -> Pretty.T |
|
6976 | 33 |
val pretty_thm_no_quote: thm -> Pretty.T |
11883 | 34 |
val pretty_thm: thm -> Pretty.T |
35 |
val pretty_thms: thm list -> Pretty.T |
|
36 |
val pretty_thm_sg: Sign.sg -> thm -> Pretty.T |
|
37 |
val pretty_thms_sg: Sign.sg -> thm list -> Pretty.T |
|
38 |
val pprint_thm: thm -> pprint_args -> unit |
|
39 |
val pretty_ctyp: ctyp -> Pretty.T |
|
40 |
val pprint_ctyp: ctyp -> pprint_args -> unit |
|
41 |
val pretty_cterm: cterm -> Pretty.T |
|
42 |
val pprint_cterm: cterm -> pprint_args -> unit |
|
43 |
val pretty_theory: theory -> Pretty.T |
|
44 |
val pprint_theory: theory -> pprint_args -> unit |
|
8720 | 45 |
val pretty_full_theory: theory -> Pretty.T list |
11883 | 46 |
val pretty_name_space: string * NameSpace.T -> Pretty.T |
14876 | 47 |
val pretty_goals_aux: Pretty.pp -> string -> bool * bool -> int -> thm -> Pretty.T list |
11883 | 48 |
val pretty_goals: int -> thm -> Pretty.T list |
49 |
val print_goals: int -> thm -> unit |
|
50 |
val current_goals_markers: (string * string * string) ref |
|
12418 | 51 |
val pretty_current_goals: int -> int -> thm -> Pretty.T list |
52 |
val print_current_goals_default: int -> int -> thm -> unit |
|
53 |
val print_current_goals_fn: (int -> int -> thm -> unit) ref |
|
4950 | 54 |
end; |
1591
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
55 |
|
4950 | 56 |
structure Display: DISPLAY = |
1591
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
57 |
struct |
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
58 |
|
11883 | 59 |
|
6087 | 60 |
(** print thm **) |
61 |
||
12081 | 62 |
val goals_limit = ref 10; (*max number of goals to print -- set by user*) |
12831
a2a3896f9c48
reset show_hyps by default (in accordance to existing Isar practice);
wenzelm
parents:
12418
diff
changeset
|
63 |
val show_hyps = ref false; (*false: print meta-hypotheses as dots*) |
11883 | 64 |
val show_tags = ref false; (*false: suppress tags*) |
6087 | 65 |
|
14598
7009f59711e3
Replaced quote by Library.quote, since quote now refers to Symbol.quote
berghofe
parents:
14472
diff
changeset
|
66 |
fun pretty_tag (name, args) = Pretty.strs (name :: map Library.quote args); |
6087 | 67 |
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
|
68 |
|
14876 | 69 |
fun pretty_flexpair pp (t, u) = Pretty.block |
70 |
[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:
12831
diff
changeset
|
71 |
|
14876 | 72 |
fun pretty_thm_aux pp quote th = |
1591
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
73 |
let |
16290 | 74 |
val {hyps, tpairs, prop, der = (ora, _), ...} = Thm.rep_thm th; |
6087 | 75 |
val xshyps = Thm.extra_shyps th; |
76 |
val (_, tags) = Thm.get_name_tags th; |
|
77 |
||
13658
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
berghofe
parents:
12831
diff
changeset
|
78 |
val q = if quote then Pretty.quote else I; |
14876 | 79 |
val prt_term = q o (Pretty.term pp); |
6279 | 80 |
|
13658
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
berghofe
parents:
12831
diff
changeset
|
81 |
val hlen = length xshyps + length hyps + length tpairs; |
1591
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
82 |
val hsymbs = |
9500 | 83 |
if hlen = 0 andalso not ora then [] |
1591
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
84 |
else if ! show_hyps then |
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
85 |
[Pretty.brk 2, Pretty.list "[" "]" |
14876 | 86 |
(map (q o pretty_flexpair pp) tpairs @ map prt_term hyps @ |
87 |
map (Pretty.sort pp) xshyps @ |
|
9500 | 88 |
(if ora then [Pretty.str "!"] else []))] |
6889 | 89 |
else [Pretty.brk 2, Pretty.str ("[" ^ implode (replicate hlen ".") ^ |
9500 | 90 |
(if ora then "!" else "") ^ "]")]; |
6087 | 91 |
val tsymbs = |
92 |
if null tags orelse not (! show_tags) then [] |
|
93 |
else [Pretty.brk 1, pretty_tags tags]; |
|
6279 | 94 |
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
|
95 |
|
12067 | 96 |
fun gen_pretty_thm quote th = |
14876 | 97 |
pretty_thm_aux (Sign.pp (Thm.sign_of_thm th)) quote th; |
6889 | 98 |
|
12067 | 99 |
val pretty_thm = gen_pretty_thm true; |
100 |
val pretty_thm_no_quote = gen_pretty_thm false; |
|
6889 | 101 |
|
102 |
||
1591
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
103 |
val string_of_thm = Pretty.string_of o pretty_thm; |
6279 | 104 |
val pprint_thm = Pretty.pprint o pretty_thm; |
1591
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
105 |
|
6087 | 106 |
fun pretty_thms [th] = pretty_thm th |
107 |
| 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
|
108 |
|
10010 | 109 |
val pretty_thm_sg = pretty_thm oo Thm.transfer_sg; |
110 |
val pretty_thms_sg = pretty_thms oo (map o Thm.transfer_sg); |
|
111 |
||
6087 | 112 |
|
113 |
(* top-level commands for printing theorems *) |
|
114 |
||
115 |
val print_thm = Pretty.writeln o pretty_thm; |
|
116 |
val print_thms = Pretty.writeln o pretty_thms; |
|
1591
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
117 |
|
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
118 |
fun prth th = (print_thm th; th); |
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
119 |
|
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
120 |
(*Print and return a sequence of theorems, separated by blank lines. *) |
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
121 |
fun prthq thseq = |
4270 | 122 |
(Seq.print (fn _ => print_thm) 100000 thseq; thseq); |
1591
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 |
(*Print and return a list of theorems, separated by blank lines. *) |
15570 | 125 |
fun prths ths = (List.app (fn th => (print_thm th; writeln "")) ths; ths); |
1591
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
126 |
|
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
127 |
|
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
128 |
(* other printing commands *) |
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
129 |
|
4126 | 130 |
fun pretty_ctyp cT = |
131 |
let val {sign, T} = rep_ctyp cT in Sign.pretty_typ sign T end; |
|
132 |
||
1591
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
133 |
fun pprint_ctyp cT = |
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
134 |
let val {sign, T} = rep_ctyp cT in Sign.pprint_typ sign T end; |
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
135 |
|
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
136 |
fun string_of_ctyp cT = |
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
137 |
let val {sign, T} = rep_ctyp cT in Sign.string_of_typ sign T end; |
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
138 |
|
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
139 |
val print_ctyp = writeln o string_of_ctyp; |
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
140 |
|
3547 | 141 |
fun pretty_cterm ct = |
142 |
let val {sign, t, ...} = rep_cterm ct in Sign.pretty_term sign t end; |
|
143 |
||
1591
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
144 |
fun pprint_cterm ct = |
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
145 |
let val {sign, t, ...} = rep_cterm ct in Sign.pprint_term sign t end; |
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
146 |
|
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
147 |
fun string_of_cterm ct = |
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
148 |
let val {sign, t, ...} = rep_cterm ct in Sign.string_of_term sign t end; |
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
149 |
|
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
150 |
val print_cterm = writeln o string_of_cterm; |
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
151 |
|
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
152 |
|
4250 | 153 |
|
154 |
(** print theory **) |
|
1591
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
155 |
|
6390 | 156 |
val pretty_theory = Sign.pretty_sg o Theory.sign_of; |
157 |
val pprint_theory = Sign.pprint_sg o Theory.sign_of; |
|
1591
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
158 |
|
6390 | 159 |
val print_syntax = Syntax.print_syntax o Theory.syn_of; |
1591
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
160 |
|
4250 | 161 |
|
4498 | 162 |
(* pretty_name_space *) |
163 |
||
164 |
fun pretty_name_space (kind, space) = |
|
165 |
let |
|
166 |
fun prt_entry (name, accs) = Pretty.block |
|
14598
7009f59711e3
Replaced quote by Library.quote, since quote now refers to Symbol.quote
berghofe
parents:
14472
diff
changeset
|
167 |
(Pretty.str (Library.quote name ^ " =") :: Pretty.brk 1 :: |
7009f59711e3
Replaced quote by Library.quote, since quote now refers to Symbol.quote
berghofe
parents:
14472
diff
changeset
|
168 |
Pretty.commas (map (Pretty.quote o Pretty.str) accs)); |
4498 | 169 |
in |
170 |
Pretty.fbreaks (Pretty.str (kind ^ ":") :: map prt_entry (NameSpace.dest space)) |
|
171 |
|> Pretty.block |
|
172 |
end; |
|
173 |
||
174 |
||
11883 | 175 |
(* pretty_full_theory *) |
4250 | 176 |
|
8720 | 177 |
fun pretty_full_theory thy = |
4250 | 178 |
let |
8720 | 179 |
val sg = Theory.sign_of thy; |
180 |
||
4250 | 181 |
fun prt_cls c = Sign.pretty_sort sg [c]; |
182 |
fun prt_sort S = Sign.pretty_sort sg S; |
|
183 |
fun prt_arity t (c, Ss) = Sign.pretty_arity sg (t, Ss, [c]); |
|
14996 | 184 |
fun prt_typ ty = Pretty.quote (Sign.pretty_typ sg ty); |
16290 | 185 |
val prt_typ_no_tvars = prt_typ o Type.freeze_type; |
8720 | 186 |
fun prt_term t = Pretty.quote (Sign.pretty_term sg t); |
4250 | 187 |
|
14794 | 188 |
fun pretty_classrel (c, []) = prt_cls c |
189 |
| pretty_classrel (c, cs) = Pretty.block |
|
190 |
(prt_cls c :: Pretty.str " <" :: Pretty.brk 1 :: |
|
191 |
Pretty.commas (map prt_cls cs)); |
|
4250 | 192 |
|
193 |
fun pretty_default S = Pretty.block |
|
14794 | 194 |
[Pretty.str "default sort:", Pretty.brk 1, prt_sort S]; |
4250 | 195 |
|
15531 | 196 |
fun pretty_witness NONE = Pretty.str "universal non-emptiness witness: -" |
197 |
| pretty_witness (SOME (T, S)) = Pretty.block |
|
14794 | 198 |
[Pretty.str "universal non-emptiness witness:", Pretty.brk 1, |
199 |
prt_typ_no_tvars T, Pretty.str " ::", Pretty.brk 1, prt_sort S]; |
|
200 |
||
14996 | 201 |
val tfrees = map (fn v => TFree (v, [])); |
15000 | 202 |
fun pretty_type syn (t, Type.LogicalType n) = |
15531 | 203 |
if syn then NONE |
204 |
else SOME (prt_typ (Type (t, tfrees (Term.invent_names [] "'a" n)))) |
|
15000 | 205 |
| pretty_type syn (t, Type.Abbreviation (vs, U, syn')) = |
15531 | 206 |
if syn <> syn' then NONE |
207 |
else SOME (Pretty.block |
|
14996 | 208 |
[prt_typ (Type (t, tfrees vs)), Pretty.str " =", Pretty.brk 1, prt_typ U]) |
15000 | 209 |
| pretty_type syn (t, Type.Nonterminal) = |
15531 | 210 |
if not syn then NONE |
211 |
else SOME (prt_typ (Type (t, []))); |
|
4250 | 212 |
|
15570 | 213 |
val pretty_arities = List.concat o map (fn (t, ars) => map (prt_arity t) ars); |
14223
0ee05eef881b
Added support for making constants final, that is, ensuring that no
skalberg
parents:
14178
diff
changeset
|
214 |
|
14794 | 215 |
fun pretty_const (c, (ty, _)) = Pretty.block |
10737 | 216 |
[Pretty.str c, Pretty.str " ::", Pretty.brk 1, prt_typ_no_tvars ty]; |
4250 | 217 |
|
14996 | 218 |
fun pretty_final (c, []) = Pretty.str c |
219 |
| pretty_final (c, tys) = Pretty.block |
|
220 |
(Pretty.str c :: Pretty.str " ::" :: Pretty.brk 1 :: |
|
221 |
Pretty.commas (map prt_typ_no_tvars tys)); |
|
222 |
||
8720 | 223 |
fun prt_axm (a, t) = Pretty.block [Pretty.str (a ^ ":"), Pretty.brk 1, prt_term t]; |
224 |
||
225 |
||
16290 | 226 |
val {sign = _, defs, axioms, oracles, parents = _, ancestors = _} = Theory.rep_theory thy; |
16127 | 227 |
val {self = _, tsig, consts, naming, spaces, data} = Sign.rep_sg sg; |
14794 | 228 |
val {classes, default, types, arities, log_types = _, witness} = Type.rep_tsig tsig; |
14996 | 229 |
|
230 |
val spcs = Library.sort_wrt #1 spaces; |
|
15000 | 231 |
val tdecls = Symtab.dest types |> map (fn (t, (d, _)) => |
16127 | 232 |
(Sign.extern sg Sign.typeK t, (t, d))) |> Library.sort_wrt #1 |> map #2; |
233 |
val cnsts = Sign.extern_table sg Sign.constK consts; |
|
16290 | 234 |
val finals = Sign.extern_table sg Sign.constK (Defs.finals defs); |
16127 | 235 |
val axms = Sign.extern_table sg Theory.axiomK axioms; |
236 |
val oras = Sign.extern_table sg Theory.oracleK oracles; |
|
4250 | 237 |
in |
8720 | 238 |
[Pretty.strs ("stamps:" :: Sign.stamp_names_of sg), |
239 |
Pretty.strs ("data:" :: Sign.data_kinds data), |
|
16127 | 240 |
Pretty.strs ["name prefix:", NameSpace.path_of naming], |
14996 | 241 |
Pretty.big_list "name spaces:" (map pretty_name_space spcs), |
14794 | 242 |
Pretty.big_list "classes:" (map pretty_classrel (Graph.dest classes)), |
8720 | 243 |
pretty_default default, |
14794 | 244 |
pretty_witness witness, |
15570 | 245 |
Pretty.big_list "syntactic types:" (List.mapPartial (pretty_type true) tdecls), |
246 |
Pretty.big_list "logical types:" (List.mapPartial (pretty_type false) tdecls), |
|
14996 | 247 |
Pretty.big_list "type arities:" (pretty_arities (Symtab.dest arities)), |
14794 | 248 |
Pretty.big_list "consts:" (map pretty_const cnsts), |
14996 | 249 |
Pretty.big_list "finals consts:" (map pretty_final finals), |
8720 | 250 |
Pretty.big_list "axioms:" (map prt_axm axms), |
251 |
Pretty.strs ("oracles:" :: (map #1 oras))] |
|
4250 | 252 |
end; |
253 |
||
254 |
||
11883 | 255 |
|
256 |
(** print_goals **) |
|
257 |
||
258 |
(* print_goals etc. *) |
|
259 |
||
8908 | 260 |
(*show consts with types in proof state output?*) |
3851
fe9932a7cd46
print_goals: optional output of const types (set show_consts);
wenzelm
parents:
3811
diff
changeset
|
261 |
val show_consts = ref false; |
fe9932a7cd46
print_goals: optional output of const types (set show_consts);
wenzelm
parents:
3811
diff
changeset
|
262 |
|
1591
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
263 |
|
11883 | 264 |
(*print thm A1,...,An/B in "goal style" -- premises as numbered subgoals*) |
265 |
||
266 |
local |
|
267 |
||
12081 | 268 |
fun ins_entry (x, y) [] = [(x, [y])] |
269 |
| ins_entry (x, y) ((pair as (x', ys')) :: pairs) = |
|
270 |
if x = x' then (x', y ins ys') :: pairs |
|
271 |
else pair :: ins_entry (x, y) pairs; |
|
272 |
||
273 |
fun add_consts (Const (c, T), env) = ins_entry (T, (c, T)) env |
|
274 |
| add_consts (t $ u, env) = add_consts (u, add_consts (t, env)) |
|
275 |
| add_consts (Abs (_, _, t), env) = add_consts (t, env) |
|
276 |
| add_consts (_, env) = env; |
|
11883 | 277 |
|
12081 | 278 |
fun add_vars (Free (x, T), env) = ins_entry (T, (x, ~1)) env |
279 |
| add_vars (Var (xi, T), env) = ins_entry (T, xi) env |
|
280 |
| add_vars (Abs (_, _, t), env) = add_vars (t, env) |
|
281 |
| add_vars (t $ u, env) = add_vars (u, add_vars (t, env)) |
|
282 |
| add_vars (_, env) = env; |
|
11883 | 283 |
|
15574
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
284 |
fun add_varsT (Type (_, Ts), env) = foldr add_varsT env Ts |
12081 | 285 |
| add_varsT (TFree (x, S), env) = ins_entry (S, (x, ~1)) env |
286 |
| add_varsT (TVar (xi, S), env) = ins_entry (S, xi) env; |
|
11883 | 287 |
|
14472
cba7c0a3ffb3
Removing the datatype declaration of "order" allows the standard General.order
paulson
parents:
14223
diff
changeset
|
288 |
fun sort_idxs vs = map (apsnd (sort (prod_ord String.compare Int.compare))) vs; |
12081 | 289 |
fun sort_cnsts cs = map (apsnd (sort_wrt fst)) cs; |
290 |
||
291 |
fun consts_of t = sort_cnsts (add_consts (t, [])); |
|
292 |
fun vars_of t = sort_idxs (add_vars (t, [])); |
|
293 |
fun varsT_of t = rev (sort_idxs (it_term_types add_varsT (t, []))); |
|
294 |
||
295 |
in |
|
11883 | 296 |
|
14876 | 297 |
fun pretty_goals_aux pp begin_goal (msg, main) maxgoals state = |
12081 | 298 |
let |
299 |
fun prt_atoms prt prtT (X, xs) = Pretty.block |
|
300 |
[Pretty.block (Pretty.commas (map prt xs)), Pretty.str " ::", |
|
301 |
Pretty.brk 1, prtT X]; |
|
11883 | 302 |
|
14876 | 303 |
fun prt_var (x, ~1) = Pretty.term pp (Syntax.free x) |
304 |
| prt_var xi = Pretty.term pp (Syntax.var xi); |
|
12081 | 305 |
|
14876 | 306 |
fun prt_varT (x, ~1) = Pretty.typ pp (TFree (x, [])) |
307 |
| prt_varT xi = Pretty.typ pp (TVar (xi, [])); |
|
12081 | 308 |
|
14876 | 309 |
val prt_consts = prt_atoms (Pretty.term pp o Const) (Pretty.typ pp); |
310 |
val prt_vars = prt_atoms prt_var (Pretty.typ pp); |
|
311 |
val prt_varsT = prt_atoms prt_varT (Pretty.sort pp); |
|
11883 | 312 |
|
313 |
||
12081 | 314 |
fun pretty_list _ _ [] = [] |
315 |
| pretty_list name prt lst = [Pretty.big_list name (map prt lst)]; |
|
11883 | 316 |
|
14876 | 317 |
fun pretty_subgoal (n, A) = Pretty.blk (0, |
318 |
[Pretty.str (begin_goal ^ " " ^ string_of_int n ^ ". "), Pretty.term pp A]); |
|
12081 | 319 |
fun pretty_subgoals As = map pretty_subgoal (1 upto length As ~~ As); |
11883 | 320 |
|
14876 | 321 |
val pretty_ffpairs = pretty_list "flex-flex pairs:" (pretty_flexpair pp); |
11883 | 322 |
|
12081 | 323 |
val pretty_consts = pretty_list "constants:" prt_consts o consts_of; |
324 |
val pretty_vars = pretty_list "variables:" prt_vars o vars_of; |
|
325 |
val pretty_varsT = pretty_list "type variables:" prt_varsT o varsT_of; |
|
11883 | 326 |
|
327 |
||
16290 | 328 |
val {prop, tpairs, ...} = Thm.rep_thm state; |
13658
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
berghofe
parents:
12831
diff
changeset
|
329 |
val (As, B) = Logic.strip_horn prop; |
12081 | 330 |
val ngoals = length As; |
11883 | 331 |
|
12081 | 332 |
fun pretty_gs (types, sorts) = |
14876 | 333 |
(if main then [Pretty.term pp B] else []) @ |
12081 | 334 |
(if ngoals = 0 then [Pretty.str "No subgoals!"] |
335 |
else if ngoals > maxgoals then |
|
15570 | 336 |
pretty_subgoals (Library.take (maxgoals, As)) @ |
12081 | 337 |
(if msg then [Pretty.str ("A total of " ^ string_of_int ngoals ^ " subgoals...")] |
338 |
else []) |
|
339 |
else pretty_subgoals As) @ |
|
340 |
pretty_ffpairs tpairs @ |
|
341 |
(if ! show_consts then pretty_consts prop else []) @ |
|
342 |
(if types then pretty_vars prop else []) @ |
|
343 |
(if sorts then pretty_varsT prop else []); |
|
344 |
in |
|
345 |
setmp show_no_free_types true |
|
14178
14a12da7288e
Makes interactive proof scripting recognize the show_all_types flag.
skalberg
parents:
13658
diff
changeset
|
346 |
(setmp show_types (! show_types orelse ! show_sorts orelse ! show_all_types) |
12081 | 347 |
(setmp show_sorts false pretty_gs)) |
14178
14a12da7288e
Makes interactive proof scripting recognize the show_all_types flag.
skalberg
parents:
13658
diff
changeset
|
348 |
(! show_types orelse ! show_sorts orelse ! show_all_types, ! show_sorts) |
11883 | 349 |
end; |
350 |
||
12081 | 351 |
fun pretty_goals_marker bg n th = |
14876 | 352 |
pretty_goals_aux (Sign.pp (Thm.sign_of_thm th)) bg (true, true) n th; |
12081 | 353 |
|
354 |
val pretty_goals = pretty_goals_marker ""; |
|
12418 | 355 |
val print_goals = (Pretty.writeln o Pretty.chunks) oo pretty_goals_marker ""; |
12081 | 356 |
|
1591
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
357 |
end; |
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
paulson
parents:
diff
changeset
|
358 |
|
11883 | 359 |
|
360 |
(* print_current_goals *) |
|
361 |
||
362 |
val current_goals_markers = ref ("", "", ""); |
|
363 |
||
12418 | 364 |
fun pretty_current_goals n m th = |
11883 | 365 |
let |
366 |
val ref (begin_state, end_state, begin_goal) = current_goals_markers; |
|
367 |
val ngoals = nprems_of th; |
|
368 |
in |
|
12418 | 369 |
(if begin_state = "" then [] else [Pretty.str begin_state]) @ |
370 |
[Pretty.str ("Level " ^ string_of_int n ^ |
|
11883 | 371 |
(if ngoals > 0 then " (" ^ string_of_int ngoals ^ " subgoal" ^ |
372 |
(if ngoals <> 1 then "s" else "") ^ ")" |
|
12418 | 373 |
else ""))] @ |
374 |
pretty_goals_marker begin_goal m th @ |
|
375 |
(if end_state = "" then [] else [Pretty.str end_state]) |
|
11883 | 376 |
end; |
377 |
||
12418 | 378 |
fun print_current_goals_default n m th = |
379 |
Pretty.writeln (Pretty.chunks (pretty_current_goals n m th)); |
|
380 |
||
11883 | 381 |
val print_current_goals_fn = ref print_current_goals_default; |
382 |
||
383 |
end; |
|
384 |
||
385 |
structure BasicDisplay: BASIC_DISPLAY = Display; |
|
386 |
open BasicDisplay; |