| author | wenzelm | 
| Fri, 18 Jul 2008 22:03:20 +0200 | |
| changeset 27654 | 0f8e2dcabbf9 | 
| parent 27302 | 8d12ac6a3e1c | 
| child 28288 | 09c812966e7f | 
| 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 show_consts: bool ref  | 
|
15  | 
end;  | 
|
16  | 
||
| 
1591
 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 
paulson 
parents:  
diff
changeset
 | 
17  | 
signature DISPLAY =  | 
| 4950 | 18  | 
sig  | 
| 11883 | 19  | 
include BASIC_DISPLAY  | 
| 14876 | 20  | 
val pretty_flexpair: Pretty.pp -> term * term -> Pretty.T  | 
| 17447 | 21  | 
val pretty_thm_aux: Pretty.pp -> bool -> bool -> term list -> thm -> Pretty.T  | 
| 11883 | 22  | 
val pretty_thm: thm -> Pretty.T  | 
| 26928 | 23  | 
val string_of_thm: thm -> string  | 
| 11883 | 24  | 
val pretty_thms: thm list -> Pretty.T  | 
| 
16437
 
aa87badf7a3c
removed pretty_theory, pprint_theory (see context.ML or thy_info.ML);
 
wenzelm 
parents: 
16364 
diff
changeset
 | 
25  | 
val pretty_thm_sg: theory -> thm -> Pretty.T  | 
| 
 
aa87badf7a3c
removed pretty_theory, pprint_theory (see context.ML or thy_info.ML);
 
wenzelm 
parents: 
16364 
diff
changeset
 | 
26  | 
val pretty_thms_sg: theory -> thm list -> Pretty.T  | 
| 26928 | 27  | 
val print_thm: thm -> unit  | 
28  | 
val print_thms: thm list -> unit  | 
|
29  | 
val prth: thm -> thm  | 
|
30  | 
val prthq: thm Seq.seq -> thm Seq.seq  | 
|
31  | 
val prths: thm list -> thm list  | 
|
| 11883 | 32  | 
val pretty_ctyp: ctyp -> Pretty.T  | 
| 26928 | 33  | 
val string_of_ctyp: ctyp -> string  | 
34  | 
val print_ctyp: ctyp -> unit  | 
|
| 11883 | 35  | 
val pretty_cterm: cterm -> Pretty.T  | 
| 26928 | 36  | 
val string_of_cterm: cterm -> string  | 
37  | 
val print_cterm: cterm -> unit  | 
|
38  | 
val print_syntax: theory -> unit  | 
|
| 
20629
 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 
wenzelm 
parents: 
20226 
diff
changeset
 | 
39  | 
val pretty_full_theory: bool -> theory -> Pretty.T list  | 
| 23634 | 40  | 
val pretty_goals_aux: Pretty.pp -> Markup.T -> bool * bool -> int -> thm -> Pretty.T list  | 
| 11883 | 41  | 
val pretty_goals: int -> thm -> Pretty.T list  | 
42  | 
val print_goals: int -> thm -> unit  | 
|
| 4950 | 43  | 
end;  | 
| 
1591
 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 
paulson 
parents:  
diff
changeset
 | 
44  | 
|
| 4950 | 45  | 
structure Display: DISPLAY =  | 
| 
1591
 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 
paulson 
parents:  
diff
changeset
 | 
46  | 
struct  | 
| 
 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 
paulson 
parents:  
diff
changeset
 | 
47  | 
|
| 11883 | 48  | 
|
| 6087 | 49  | 
(** print thm **)  | 
50  | 
||
| 16534 | 51  | 
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: 
12418 
diff
changeset
 | 
52  | 
val show_hyps = ref false; (*false: print meta-hypotheses as dots*)  | 
| 11883 | 53  | 
val show_tags = ref false; (*false: suppress tags*)  | 
| 6087 | 54  | 
|
| 23657 | 55  | 
fun pretty_tag (name, arg) = Pretty.strs [name, Library.quote arg];  | 
| 6087 | 56  | 
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
 | 
57  | 
|
| 14876 | 58  | 
fun pretty_flexpair pp (t, u) = Pretty.block  | 
59  | 
[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
 | 
60  | 
|
| 17447 | 61  | 
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
 | 
62  | 
let  | 
| 17447 | 63  | 
val th = Thm.strip_shyps raw_th;  | 
| 16290 | 64  | 
    val {hyps, tpairs, prop, der = (ora, _), ...} = Thm.rep_thm th;
 | 
| 6087 | 65  | 
val xshyps = Thm.extra_shyps th;  | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
20629 
diff
changeset
 | 
66  | 
val tags = Thm.get_tags th;  | 
| 6087 | 67  | 
|
| 
13658
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
12831 
diff
changeset
 | 
68  | 
val q = if quote then Pretty.quote else I;  | 
| 22878 | 69  | 
val prt_term = q o Pretty.term pp;  | 
| 6279 | 70  | 
|
| 19300 | 71  | 
val hyps' = if ! show_hyps then hyps else subtract (op aconv) asms hyps;  | 
| 17475 | 72  | 
val ora' = ora andalso (! show_hyps orelse not (! quick_and_dirty));  | 
| 17447 | 73  | 
val hlen = length xshyps + length hyps' + length tpairs;  | 
| 
1591
 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 
paulson 
parents:  
diff
changeset
 | 
74  | 
val hsymbs =  | 
| 17475 | 75  | 
if hlen = 0 andalso not ora' then []  | 
| 17447 | 76  | 
else if ! show_hyps orelse show_hyps' then  | 
| 
1591
 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 
paulson 
parents:  
diff
changeset
 | 
77  | 
[Pretty.brk 2, Pretty.list "[" "]"  | 
| 17447 | 78  | 
(map (q o pretty_flexpair pp) tpairs @ map prt_term hyps' @  | 
| 14876 | 79  | 
map (Pretty.sort pp) xshyps @  | 
| 17475 | 80  | 
(if ora' then [Pretty.str "!"] else []))]  | 
| 6889 | 81  | 
      else [Pretty.brk 2, Pretty.str ("[" ^ implode (replicate hlen ".") ^
 | 
| 17475 | 82  | 
(if ora' then "!" else "") ^ "]")];  | 
| 6087 | 83  | 
val tsymbs =  | 
84  | 
if null tags orelse not (! show_tags) then []  | 
|
85  | 
else [Pretty.brk 1, pretty_tags tags];  | 
|
| 6279 | 86  | 
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
 | 
87  | 
|
| 20226 | 88  | 
fun pretty_thm th =  | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26928 
diff
changeset
 | 
89  | 
pretty_thm_aux (Syntax.pp_global (Thm.theory_of_thm th)) true false [] th;  | 
| 6889 | 90  | 
|
| 
1591
 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 
paulson 
parents:  
diff
changeset
 | 
91  | 
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
 | 
92  | 
|
| 6087 | 93  | 
fun pretty_thms [th] = pretty_thm th  | 
94  | 
| 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
 | 
95  | 
|
| 
16437
 
aa87badf7a3c
removed pretty_theory, pprint_theory (see context.ML or thy_info.ML);
 
wenzelm 
parents: 
16364 
diff
changeset
 | 
96  | 
val pretty_thm_sg = pretty_thm oo Thm.transfer;  | 
| 
 
aa87badf7a3c
removed pretty_theory, pprint_theory (see context.ML or thy_info.ML);
 
wenzelm 
parents: 
16364 
diff
changeset
 | 
97  | 
val pretty_thms_sg = pretty_thms oo (map o Thm.transfer);  | 
| 10010 | 98  | 
|
| 6087 | 99  | 
|
100  | 
(* top-level commands for printing theorems *)  | 
|
101  | 
||
102  | 
val print_thm = Pretty.writeln o pretty_thm;  | 
|
103  | 
val print_thms = Pretty.writeln o pretty_thms;  | 
|
| 
1591
 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 
paulson 
parents:  
diff
changeset
 | 
104  | 
|
| 
 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 
paulson 
parents:  
diff
changeset
 | 
105  | 
fun prth th = (print_thm th; th);  | 
| 16534 | 106  | 
fun prthq thq = (Seq.print (K print_thm) 100000 thq; thq);  | 
107  | 
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
 | 
108  | 
|
| 
 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 
paulson 
parents:  
diff
changeset
 | 
109  | 
|
| 
 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 
paulson 
parents:  
diff
changeset
 | 
110  | 
(* other printing commands *)  | 
| 
 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 
paulson 
parents:  
diff
changeset
 | 
111  | 
|
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26928 
diff
changeset
 | 
112  | 
fun pretty_ctyp cT = Syntax.pretty_typ_global (Thm.theory_of_ctyp cT) (Thm.typ_of cT);  | 
| 
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26928 
diff
changeset
 | 
113  | 
fun string_of_ctyp cT = Syntax.string_of_typ_global (Thm.theory_of_ctyp cT) (Thm.typ_of cT);  | 
| 
1591
 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 
paulson 
parents:  
diff
changeset
 | 
114  | 
val print_ctyp = writeln o string_of_ctyp;  | 
| 
 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 
paulson 
parents:  
diff
changeset
 | 
115  | 
|
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26928 
diff
changeset
 | 
116  | 
fun pretty_cterm ct = Syntax.pretty_term_global (Thm.theory_of_cterm ct) (Thm.term_of ct);  | 
| 
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26928 
diff
changeset
 | 
117  | 
fun string_of_cterm ct = Syntax.string_of_term_global (Thm.theory_of_cterm ct) (Thm.term_of ct);  | 
| 
1591
 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 
paulson 
parents:  
diff
changeset
 | 
118  | 
val print_cterm = writeln o string_of_cterm;  | 
| 
 
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  | 
|
| 4250 | 121  | 
|
122  | 
(** print theory **)  | 
|
| 
1591
 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 
paulson 
parents:  
diff
changeset
 | 
123  | 
|
| 
16437
 
aa87badf7a3c
removed pretty_theory, pprint_theory (see context.ML or thy_info.ML);
 
wenzelm 
parents: 
16364 
diff
changeset
 | 
124  | 
val print_syntax = Syntax.print_syntax o Sign.syn_of;  | 
| 4498 | 125  | 
|
126  | 
||
| 11883 | 127  | 
(* pretty_full_theory *)  | 
| 4250 | 128  | 
|
| 
20629
 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 
wenzelm 
parents: 
20226 
diff
changeset
 | 
129  | 
fun pretty_full_theory verbose thy =  | 
| 4250 | 130  | 
let  | 
| 24920 | 131  | 
val ctxt = ProofContext.init thy;  | 
132  | 
||
133  | 
fun prt_cls c = Syntax.pretty_sort ctxt [c];  | 
|
134  | 
fun prt_sort S = Syntax.pretty_sort ctxt S;  | 
|
135  | 
fun prt_arity t (c, (_, Ss)) = Syntax.pretty_arity ctxt (t, Ss, [c]);  | 
|
136  | 
fun prt_typ ty = Pretty.quote (Syntax.pretty_typ ctxt ty);  | 
|
| 19806 | 137  | 
val prt_typ_no_tvars = prt_typ o Logic.unvarifyT;  | 
| 24920 | 138  | 
fun prt_term t = Pretty.quote (Syntax.pretty_term ctxt t);  | 
| 18936 | 139  | 
val prt_term_no_vars = prt_term o Logic.unvarify;  | 
| 19698 | 140  | 
fun prt_const (c, ty) = [Pretty.str c, Pretty.str " ::", Pretty.brk 1, prt_typ_no_tvars ty];  | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26928 
diff
changeset
 | 
141  | 
val prt_const' = Defs.pretty_const (Syntax.pp ctxt);  | 
| 4250 | 142  | 
|
| 14794 | 143  | 
fun pretty_classrel (c, []) = prt_cls c  | 
144  | 
| pretty_classrel (c, cs) = Pretty.block  | 
|
145  | 
(prt_cls c :: Pretty.str " <" :: Pretty.brk 1 ::  | 
|
146  | 
Pretty.commas (map prt_cls cs));  | 
|
| 4250 | 147  | 
|
148  | 
fun pretty_default S = Pretty.block  | 
|
| 14794 | 149  | 
[Pretty.str "default sort:", Pretty.brk 1, prt_sort S];  | 
| 4250 | 150  | 
|
| 14996 | 151  | 
val tfrees = map (fn v => TFree (v, []));  | 
| 27302 | 152  | 
fun pretty_type syn (t, ((Type.LogicalType n, _), _)) =  | 
| 15531 | 153  | 
if syn then NONE  | 
| 24848 | 154  | 
else SOME (prt_typ (Type (t, tfrees (Name.invents Name.context Name.aT n))))  | 
| 27302 | 155  | 
| pretty_type syn (t, ((Type.Abbreviation (vs, U, syn'), _), _)) =  | 
| 15531 | 156  | 
if syn <> syn' then NONE  | 
157  | 
else SOME (Pretty.block  | 
|
| 14996 | 158  | 
[prt_typ (Type (t, tfrees vs)), Pretty.str " =", Pretty.brk 1, prt_typ U])  | 
| 27302 | 159  | 
| pretty_type syn (t, ((Type.Nonterminal, _), _)) =  | 
| 15531 | 160  | 
if not syn then NONE  | 
161  | 
else SOME (prt_typ (Type (t, [])));  | 
|
| 4250 | 162  | 
|
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19420 
diff
changeset
 | 
163  | 
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: 
14178 
diff
changeset
 | 
164  | 
|
| 19698 | 165  | 
fun pretty_abbrev (c, (ty, t)) = Pretty.block  | 
166  | 
(prt_const (c, ty) @ [Pretty.str " ==", Pretty.brk 1, prt_term_no_vars t]);  | 
|
167  | 
||
168  | 
fun pretty_axm (a, t) = Pretty.block [Pretty.str (a ^ ":"), Pretty.brk 1, prt_term_no_vars t];  | 
|
| 19365 | 169  | 
|
| 19702 | 170  | 
fun pretty_finals reds = Pretty.block  | 
| 19703 | 171  | 
(Pretty.str "final:" :: Pretty.brk 1 :: Pretty.commas (map (prt_const' o fst) reds));  | 
| 19702 | 172  | 
|
| 19698 | 173  | 
fun pretty_reduct (lhs, rhs) = Pretty.block  | 
| 19702 | 174  | 
([prt_const' lhs, Pretty.str " ->", Pretty.brk 2] @  | 
175  | 
Pretty.commas (map prt_const' (sort_wrt #1 rhs)));  | 
|
| 4250 | 176  | 
|
| 19698 | 177  | 
fun pretty_restrict (const, name) =  | 
178  | 
      Pretty.block ([prt_const' const, Pretty.brk 2, Pretty.str ("(from " ^ quote name ^ ")")]);
 | 
|
| 8720 | 179  | 
|
| 24665 | 180  | 
val axioms = (Theory.axiom_space thy, Theory.axiom_table thy);  | 
181  | 
val oracles = (Theory.oracle_space thy, Theory.oracle_table thy);  | 
|
182  | 
val defs = Theory.defs_of thy;  | 
|
| 19698 | 183  | 
    val {restricts, reducts} = Defs.dest defs;
 | 
| 18061 | 184  | 
    val {naming, syn = _, tsig, consts} = Sign.rep_sg thy;
 | 
| 18936 | 185  | 
    val {constants, constraints} = Consts.dest consts;
 | 
| 19698 | 186  | 
val extern_const = NameSpace.extern (#1 constants);  | 
| 26637 | 187  | 
    val {classes, default, types, ...} = Type.rep_tsig tsig;
 | 
| 19698 | 188  | 
val (class_space, class_algebra) = classes;  | 
| 19642 | 189  | 
    val {classes, arities} = Sorts.rep_algebra class_algebra;
 | 
| 14996 | 190  | 
|
| 
24774
 
bc31c318e673
print_theory: observe Markup.internal_property of consts, discontinued special treatment of internal names elsewhere;
 
wenzelm 
parents: 
24665 
diff
changeset
 | 
191  | 
val clsses = NameSpace.dest_table (class_space, Symtab.make (Graph.dest classes));  | 
| 
 
bc31c318e673
print_theory: observe Markup.internal_property of consts, discontinued special treatment of internal names elsewhere;
 
wenzelm 
parents: 
24665 
diff
changeset
 | 
192  | 
val tdecls = NameSpace.dest_table types;  | 
| 
 
bc31c318e673
print_theory: observe Markup.internal_property of consts, discontinued special treatment of internal names elsewhere;
 
wenzelm 
parents: 
24665 
diff
changeset
 | 
193  | 
val arties = NameSpace.dest_table (Sign.type_space thy, arities);  | 
| 
20629
 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 
wenzelm 
parents: 
20226 
diff
changeset
 | 
194  | 
|
| 
24774
 
bc31c318e673
print_theory: observe Markup.internal_property of consts, discontinued special treatment of internal names elsewhere;
 
wenzelm 
parents: 
24665 
diff
changeset
 | 
195  | 
fun prune_const c = not verbose andalso  | 
| 
 
bc31c318e673
print_theory: observe Markup.internal_property of consts, discontinued special treatment of internal names elsewhere;
 
wenzelm 
parents: 
24665 
diff
changeset
 | 
196  | 
member (op =) (Consts.the_tags consts c) Markup.property_internal;  | 
| 
 
bc31c318e673
print_theory: observe Markup.internal_property of consts, discontinued special treatment of internal names elsewhere;
 
wenzelm 
parents: 
24665 
diff
changeset
 | 
197  | 
val cnsts = NameSpace.extern_table (#1 constants,  | 
| 
 
bc31c318e673
print_theory: observe Markup.internal_property of consts, discontinued special treatment of internal names elsewhere;
 
wenzelm 
parents: 
24665 
diff
changeset
 | 
198  | 
Symtab.make (filter_out (prune_const o fst) (Symtab.dest (#2 constants))));  | 
| 
 
bc31c318e673
print_theory: observe Markup.internal_property of consts, discontinued special treatment of internal names elsewhere;
 
wenzelm 
parents: 
24665 
diff
changeset
 | 
199  | 
|
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19420 
diff
changeset
 | 
200  | 
val log_cnsts = map_filter (fn (c, (ty, NONE)) => SOME (c, ty) | _ => NONE) cnsts;  | 
| 25405 | 201  | 
val abbrevs = map_filter (fn (c, (ty, SOME t)) => SOME (c, (ty, t)) | _ => NONE) cnsts;  | 
| 
24774
 
bc31c318e673
print_theory: observe Markup.internal_property of consts, discontinued special treatment of internal names elsewhere;
 
wenzelm 
parents: 
24665 
diff
changeset
 | 
202  | 
val cnstrs = NameSpace.extern_table constraints;  | 
| 19698 | 203  | 
|
| 
24774
 
bc31c318e673
print_theory: observe Markup.internal_property of consts, discontinued special treatment of internal names elsewhere;
 
wenzelm 
parents: 
24665 
diff
changeset
 | 
204  | 
val axms = NameSpace.extern_table axioms;  | 
| 
 
bc31c318e673
print_theory: observe Markup.internal_property of consts, discontinued special treatment of internal names elsewhere;
 
wenzelm 
parents: 
24665 
diff
changeset
 | 
205  | 
val oras = NameSpace.extern_table oracles;  | 
| 
 
bc31c318e673
print_theory: observe Markup.internal_property of consts, discontinued special treatment of internal names elsewhere;
 
wenzelm 
parents: 
24665 
diff
changeset
 | 
206  | 
|
| 
 
bc31c318e673
print_theory: observe Markup.internal_property of consts, discontinued special treatment of internal names elsewhere;
 
wenzelm 
parents: 
24665 
diff
changeset
 | 
207  | 
val (reds0, (reds1, reds2)) = filter_out (prune_const o fst o fst) reducts  | 
| 
 
bc31c318e673
print_theory: observe Markup.internal_property of consts, discontinued special treatment of internal names elsewhere;
 
wenzelm 
parents: 
24665 
diff
changeset
 | 
208  | 
|> map (fn (lhs, rhs) =>  | 
| 
 
bc31c318e673
print_theory: observe Markup.internal_property of consts, discontinued special treatment of internal names elsewhere;
 
wenzelm 
parents: 
24665 
diff
changeset
 | 
209  | 
(apfst extern_const lhs, map (apfst extern_const) (filter_out (prune_const o fst) rhs)))  | 
| 19702 | 210  | 
|> sort_wrt (#1 o #1)  | 
211  | 
|> List.partition (null o #2)  | 
|
212  | 
||> List.partition (Defs.plain_args o #2 o #1);  | 
|
| 19698 | 213  | 
val rests = restricts |> map (apfst (apfst extern_const)) |> sort_wrt (#1 o #1);  | 
| 4250 | 214  | 
in  | 
| 
20629
 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 
wenzelm 
parents: 
20226 
diff
changeset
 | 
215  | 
    [Pretty.strs ("names:" :: Context.names_of thy)] @
 | 
| 
 
8f6cc81ba4a3
pretty_full_theory: suppress internal entities by default;
 
wenzelm 
parents: 
20226 
diff
changeset
 | 
216  | 
[Pretty.strs ["name prefix:", NameSpace.path_of naming],  | 
| 16364 | 217  | 
Pretty.big_list "classes:" (map pretty_classrel clsses),  | 
| 8720 | 218  | 
pretty_default default,  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19420 
diff
changeset
 | 
219  | 
Pretty.big_list "syntactic types:" (map_filter (pretty_type true) tdecls),  | 
| 
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19420 
diff
changeset
 | 
220  | 
Pretty.big_list "logical types:" (map_filter (pretty_type false) tdecls),  | 
| 16534 | 221  | 
Pretty.big_list "type arities:" (pretty_arities arties),  | 
| 19698 | 222  | 
Pretty.big_list "logical consts:" (map (Pretty.block o prt_const) log_cnsts),  | 
| 19365 | 223  | 
Pretty.big_list "abbreviations:" (map pretty_abbrev abbrevs),  | 
| 19698 | 224  | 
Pretty.big_list "constraints:" (map (Pretty.block o prt_const) cnstrs),  | 
225  | 
Pretty.big_list "axioms:" (map pretty_axm axms),  | 
|
| 19702 | 226  | 
      Pretty.strs ("oracles:" :: (map #1 oras)),
 | 
| 19698 | 227  | 
Pretty.big_list "definitions:"  | 
| 19702 | 228  | 
[pretty_finals reds0,  | 
229  | 
Pretty.big_list "non-overloaded:" (map pretty_reduct reds1),  | 
|
230  | 
Pretty.big_list "overloaded:" (map pretty_reduct reds2),  | 
|
231  | 
Pretty.big_list "pattern restrictions:" (map pretty_restrict rests)]]  | 
|
| 4250 | 232  | 
end;  | 
233  | 
||
234  | 
||
| 11883 | 235  | 
|
236  | 
(** print_goals **)  | 
|
237  | 
||
238  | 
(* print_goals etc. *)  | 
|
239  | 
||
| 16534 | 240  | 
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: 
3811 
diff
changeset
 | 
241  | 
|
| 
1591
 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 
paulson 
parents:  
diff
changeset
 | 
242  | 
|
| 11883 | 243  | 
(*print thm A1,...,An/B in "goal style" -- premises as numbered subgoals*)  | 
244  | 
||
245  | 
local  | 
|
246  | 
||
| 18980 | 247  | 
fun ins_entry (x, y) =  | 
248  | 
AList.default (op =) (x, []) #>  | 
|
249  | 
AList.map_entry (op =) x (insert (op =) y);  | 
|
| 12081 | 250  | 
|
| 18980 | 251  | 
val add_consts = Term.fold_aterms  | 
252  | 
(fn Const (c, T) => ins_entry (T, (c, T))  | 
|
253  | 
| _ => I);  | 
|
| 11883 | 254  | 
|
| 18980 | 255  | 
val add_vars = Term.fold_aterms  | 
256  | 
(fn Free (x, T) => ins_entry (T, (x, ~1))  | 
|
257  | 
| Var (xi, T) => ins_entry (T, xi)  | 
|
258  | 
| _ => I);  | 
|
| 11883 | 259  | 
|
| 18980 | 260  | 
val add_varsT = Term.fold_atyps  | 
261  | 
(fn TFree (x, S) => ins_entry (S, (x, ~1))  | 
|
262  | 
| TVar (xi, S) => ins_entry (S, xi)  | 
|
263  | 
| _ => I);  | 
|
| 11883 | 264  | 
|
| 16490 | 265  | 
fun sort_idxs vs = map (apsnd (sort (prod_ord string_ord int_ord))) vs;  | 
| 12081 | 266  | 
fun sort_cnsts cs = map (apsnd (sort_wrt fst)) cs;  | 
267  | 
||
| 18980 | 268  | 
fun consts_of t = sort_cnsts (add_consts t []);  | 
269  | 
fun vars_of t = sort_idxs (add_vars t []);  | 
|
270  | 
fun varsT_of t = rev (sort_idxs (Term.fold_types add_varsT t []));  | 
|
| 12081 | 271  | 
|
272  | 
in  | 
|
| 11883 | 273  | 
|
| 23634 | 274  | 
fun pretty_goals_aux pp markup (msg, main) maxgoals state =  | 
| 12081 | 275  | 
let  | 
276  | 
fun prt_atoms prt prtT (X, xs) = Pretty.block  | 
|
277  | 
[Pretty.block (Pretty.commas (map prt xs)), Pretty.str " ::",  | 
|
278  | 
Pretty.brk 1, prtT X];  | 
|
| 11883 | 279  | 
|
| 14876 | 280  | 
fun prt_var (x, ~1) = Pretty.term pp (Syntax.free x)  | 
281  | 
| prt_var xi = Pretty.term pp (Syntax.var xi);  | 
|
| 12081 | 282  | 
|
| 14876 | 283  | 
fun prt_varT (x, ~1) = Pretty.typ pp (TFree (x, []))  | 
284  | 
| prt_varT xi = Pretty.typ pp (TVar (xi, []));  | 
|
| 12081 | 285  | 
|
| 14876 | 286  | 
val prt_consts = prt_atoms (Pretty.term pp o Const) (Pretty.typ pp);  | 
287  | 
val prt_vars = prt_atoms prt_var (Pretty.typ pp);  | 
|
288  | 
val prt_varsT = prt_atoms prt_varT (Pretty.sort pp);  | 
|
| 11883 | 289  | 
|
290  | 
||
| 12081 | 291  | 
fun pretty_list _ _ [] = []  | 
292  | 
| pretty_list name prt lst = [Pretty.big_list name (map prt lst)];  | 
|
| 11883 | 293  | 
|
| 23634 | 294  | 
fun pretty_subgoal (n, A) = Pretty.markup markup  | 
295  | 
      [Pretty.str (" " ^ string_of_int n ^ ". "), Pretty.term pp A];
 | 
|
| 12081 | 296  | 
fun pretty_subgoals As = map pretty_subgoal (1 upto length As ~~ As);  | 
| 11883 | 297  | 
|
| 14876 | 298  | 
val pretty_ffpairs = pretty_list "flex-flex pairs:" (pretty_flexpair pp);  | 
| 11883 | 299  | 
|
| 12081 | 300  | 
val pretty_consts = pretty_list "constants:" prt_consts o consts_of;  | 
301  | 
val pretty_vars = pretty_list "variables:" prt_vars o vars_of;  | 
|
302  | 
val pretty_varsT = pretty_list "type variables:" prt_varsT o varsT_of;  | 
|
| 11883 | 303  | 
|
304  | 
||
| 16290 | 305  | 
    val {prop, tpairs, ...} = Thm.rep_thm state;
 | 
| 
13658
 
c9ad3e64ddcf
Changed handling of flex-flex constraints: now stored in separate
 
berghofe 
parents: 
12831 
diff
changeset
 | 
306  | 
val (As, B) = Logic.strip_horn prop;  | 
| 12081 | 307  | 
val ngoals = length As;  | 
| 11883 | 308  | 
|
| 12081 | 309  | 
fun pretty_gs (types, sorts) =  | 
| 14876 | 310  | 
(if main then [Pretty.term pp B] else []) @  | 
| 12081 | 311  | 
(if ngoals = 0 then [Pretty.str "No subgoals!"]  | 
312  | 
else if ngoals > maxgoals then  | 
|
| 15570 | 313  | 
pretty_subgoals (Library.take (maxgoals, As)) @  | 
| 12081 | 314  | 
          (if msg then [Pretty.str ("A total of " ^ string_of_int ngoals ^ " subgoals...")]
 | 
315  | 
else [])  | 
|
316  | 
else pretty_subgoals As) @  | 
|
317  | 
pretty_ffpairs tpairs @  | 
|
318  | 
(if ! show_consts then pretty_consts prop else []) @  | 
|
319  | 
(if types then pretty_vars prop else []) @  | 
|
320  | 
(if sorts then pretty_varsT prop else []);  | 
|
321  | 
in  | 
|
322  | 
setmp show_no_free_types true  | 
|
| 
14178
 
14a12da7288e
Makes interactive proof scripting recognize the show_all_types flag.
 
skalberg 
parents: 
13658 
diff
changeset
 | 
323  | 
(setmp show_types (! show_types orelse ! show_sorts orelse ! show_all_types)  | 
| 12081 | 324  | 
(setmp show_sorts false pretty_gs))  | 
| 
14178
 
14a12da7288e
Makes interactive proof scripting recognize the show_all_types flag.
 
skalberg 
parents: 
13658 
diff
changeset
 | 
325  | 
(! show_types orelse ! show_sorts orelse ! show_all_types, ! show_sorts)  | 
| 11883 | 326  | 
end;  | 
327  | 
||
| 23634 | 328  | 
fun pretty_goals n th =  | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26928 
diff
changeset
 | 
329  | 
pretty_goals_aux (Syntax.pp_global (Thm.theory_of_thm th)) Markup.none (true, true) n th;  | 
| 12081 | 330  | 
|
| 23634 | 331  | 
val print_goals = (Pretty.writeln o Pretty.chunks) oo pretty_goals;  | 
| 12081 | 332  | 
|
| 
1591
 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 
paulson 
parents:  
diff
changeset
 | 
333  | 
end;  | 
| 
 
7fa0265dcd13
New module for display/printing operations, taken from drule.ML
 
paulson 
parents:  
diff
changeset
 | 
334  | 
|
| 11883 | 335  | 
|
336  | 
end;  | 
|
337  | 
||
338  | 
structure BasicDisplay: BASIC_DISPLAY = Display;  | 
|
339  | 
open BasicDisplay;  |