| author | wenzelm | 
| Tue, 30 Mar 2010 00:47:52 +0200 | |
| changeset 36016 | 4f5c7a19ebe0 | 
| parent 32966 | 5b21661fe618 | 
| child 36687 | 58020b59baf7 | 
| permissions | -rw-r--r-- | 
| 6118 | 1  | 
(* Title: Pure/General/pretty.ML  | 
| 8806 | 2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 
10952
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
3  | 
Author: Markus Wenzel, TU Munich  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
4  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
5  | 
Generic pretty printing module.  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
6  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
7  | 
Loosely based on  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
D. C. Oppen, "Pretty Printing",  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
9  | 
ACM Transactions on Programming Languages and Systems (1980), 465-483.  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
10  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
11  | 
The object to be printed is given as a tree with indentation and line  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
12  | 
breaking information. A "break" inserts a newline if the text until  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
13  | 
the next break is too long to fit on the current line. After the newline,  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
14  | 
text is indented to the level of the enclosing block. Normally, if a block  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
15  | 
is broken then all enclosing blocks will also be broken. Only "inconsistent  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
16  | 
breaks" are provided.  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
17  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
18  | 
The stored length of a block is used in breakdist (to treat each inner block as  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
19  | 
a unit for breaking).  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
20  | 
*)  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
21  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
22  | 
signature PRETTY =  | 
| 
14832
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
23  | 
sig  | 
| 
23660
 
18765718cf62
type output = string indicates raw system output;
 
wenzelm 
parents: 
23659 
diff
changeset
 | 
24  | 
val default_indent: string -> int -> output  | 
| 23698 | 25  | 
val add_mode: string -> (string -> int -> output) -> unit  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
26  | 
type T  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
27  | 
val str: string -> T  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
28  | 
val brk: int -> T  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
29  | 
val fbrk: T  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
30  | 
val breaks: T list -> T list  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
31  | 
val fbreaks: T list -> T list  | 
| 23645 | 32  | 
val blk: int * T list -> T  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
33  | 
val block: T list -> T  | 
| 23645 | 34  | 
val strs: string list -> T  | 
| 23617 | 35  | 
val markup: Markup.T -> T list -> T  | 
| 26703 | 36  | 
val mark: Markup.T -> T -> T  | 
| 23617 | 37  | 
val keyword: string -> T  | 
38  | 
val command: string -> T  | 
|
| 23638 | 39  | 
val markup_chunks: Markup.T -> T list -> T  | 
| 18802 | 40  | 
val chunks: T list -> T  | 
| 19266 | 41  | 
val chunks2: T list -> T  | 
| 23617 | 42  | 
val block_enclose: T * T -> T list -> T  | 
| 18802 | 43  | 
val quote: T -> T  | 
44  | 
val backquote: T -> T  | 
|
45  | 
val separate: string -> T list -> T list  | 
|
46  | 
val commas: T list -> T list  | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
47  | 
val enclose: string -> string -> T list -> T  | 
| 18802 | 48  | 
val enum: string -> string -> string -> T list -> T  | 
| 30620 | 49  | 
val position: Position.T -> T  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
50  | 
val list: string -> string -> T list -> T  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
51  | 
val str_list: string -> string -> string list -> T  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
52  | 
val big_list: string -> T list -> T  | 
| 9730 | 53  | 
val indent: int -> T -> T  | 
| 23645 | 54  | 
val unbreakable: T -> T  | 
55  | 
val setmargin: int -> unit  | 
|
| 
32966
 
5b21661fe618
indicate CRITICAL nature of various setmp combinators;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
56  | 
  val setmp_margin_CRITICAL: int -> ('a -> 'b) -> 'a -> 'b
 | 
| 23645 | 57  | 
val setdepth: int -> unit  | 
| 30620 | 58  | 
val to_ML: T -> ML_Pretty.pretty  | 
59  | 
val from_ML: ML_Pretty.pretty -> T  | 
|
| 23645 | 60  | 
val symbolicN: string  | 
| 16714 | 61  | 
val output_buffer: T -> Buffer.T  | 
| 
23660
 
18765718cf62
type output = string indicates raw system output;
 
wenzelm 
parents: 
23659 
diff
changeset
 | 
62  | 
val output: T -> output  | 
| 23645 | 63  | 
val string_of: T -> string  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
64  | 
val str_of: T -> string  | 
| 23645 | 65  | 
val writeln: T -> unit  | 
| 
14832
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
66  | 
type pp  | 
| 14972 | 67  | 
val pp: (term -> T) * (typ -> T) * (sort -> T) * (class list -> T) * (arity -> T) -> pp  | 
| 
14832
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
68  | 
val term: pp -> term -> T  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
69  | 
val typ: pp -> typ -> T  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
70  | 
val sort: pp -> sort -> T  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
71  | 
val classrel: pp -> class list -> T  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
72  | 
val arity: pp -> arity -> T  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
73  | 
val string_of_term: pp -> term -> string  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
74  | 
val string_of_typ: pp -> typ -> string  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
75  | 
val string_of_sort: pp -> sort -> string  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
76  | 
val string_of_classrel: pp -> class list -> string  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
77  | 
val string_of_arity: pp -> arity -> string  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
78  | 
end;  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
79  | 
|
| 23617 | 80  | 
structure Pretty: PRETTY =  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
81  | 
struct  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
82  | 
|
| 23617 | 83  | 
(** print mode operations **)  | 
84  | 
||
85  | 
fun default_indent (_: string) = Symbol.spaces;  | 
|
86  | 
||
87  | 
local  | 
|
| 23698 | 88  | 
  val default = {indent = default_indent};
 | 
| 32738 | 89  | 
  val modes = Unsynchronized.ref (Symtab.make [("", default)]);
 | 
| 23617 | 90  | 
in  | 
| 
23922
 
707639e9497d
marked some CRITICAL sections (for multithreading);
 
wenzelm 
parents: 
23787 
diff
changeset
 | 
91  | 
fun add_mode name indent = CRITICAL (fn () =>  | 
| 32738 | 92  | 
    Unsynchronized.change modes (Symtab.update_new (name, {indent = indent})));
 | 
| 23617 | 93  | 
fun get_mode () =  | 
| 24612 | 94  | 
the_default default (Library.get_first (Symtab.lookup (! modes)) (print_mode_value ()));  | 
| 23617 | 95  | 
end;  | 
96  | 
||
97  | 
fun mode_indent x y = #indent (get_mode ()) x y;  | 
|
| 23645 | 98  | 
|
99  | 
val output_spaces = Output.output o Symbol.spaces;  | 
|
100  | 
val add_indent = Buffer.add o output_spaces;  | 
|
| 23617 | 101  | 
|
102  | 
||
| 
10952
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
103  | 
|
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
104  | 
(** printing items: compound phrases, strings, and breaks **)  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
105  | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
106  | 
datatype T =  | 
| 
30667
 
53fbf7c679b0
Block markup: maintain output version within tree values (in accordance with String) -- changes operational behaviour wrt. print_mode;
 
wenzelm 
parents: 
30624 
diff
changeset
 | 
107  | 
Block of (output * output) * T list * int * int | (*markup output, body, indentation, length*)  | 
| 
 
53fbf7c679b0
Block markup: maintain output version within tree values (in accordance with String) -- changes operational behaviour wrt. print_mode;
 
wenzelm 
parents: 
30624 
diff
changeset
 | 
108  | 
String of output * int | (*text, length*)  | 
| 
 
53fbf7c679b0
Block markup: maintain output version within tree values (in accordance with String) -- changes operational behaviour wrt. print_mode;
 
wenzelm 
parents: 
30624 
diff
changeset
 | 
109  | 
Break of bool * int; (*mandatory flag, width if not taken*)  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
110  | 
|
| 23645 | 111  | 
fun length (Block (_, _, _, len)) = len  | 
112  | 
| length (String (_, len)) = len  | 
|
113  | 
| length (Break (_, wd)) = wd;  | 
|
114  | 
||
| 9730 | 115  | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
116  | 
|
| 23645 | 117  | 
(** derived operations to create formatting expressions **)  | 
118  | 
||
119  | 
val str = String o Output.output_width;  | 
|
120  | 
||
121  | 
fun brk wd = Break (false, wd);  | 
|
122  | 
val fbrk = Break (true, 2);  | 
|
123  | 
||
124  | 
fun breaks prts = Library.separate (brk 1) prts;  | 
|
125  | 
fun fbreaks prts = Library.separate fbrk prts;  | 
|
126  | 
||
| 
30667
 
53fbf7c679b0
Block markup: maintain output version within tree values (in accordance with String) -- changes operational behaviour wrt. print_mode;
 
wenzelm 
parents: 
30624 
diff
changeset
 | 
127  | 
fun block_markup m (indent, es) =  | 
| 23645 | 128  | 
let  | 
129  | 
fun sum [] k = k  | 
|
130  | 
| sum (e :: es) k = sum es (length e + k);  | 
|
131  | 
in Block (m, es, indent, sum es 0) end;  | 
|
132  | 
||
| 
30667
 
53fbf7c679b0
Block markup: maintain output version within tree values (in accordance with String) -- changes operational behaviour wrt. print_mode;
 
wenzelm 
parents: 
30624 
diff
changeset
 | 
133  | 
fun markup_block m arg = block_markup (Markup.output m) arg;  | 
| 
 
53fbf7c679b0
Block markup: maintain output version within tree values (in accordance with String) -- changes operational behaviour wrt. print_mode;
 
wenzelm 
parents: 
30624 
diff
changeset
 | 
134  | 
|
| 23645 | 135  | 
val blk = markup_block Markup.none;  | 
136  | 
fun block prts = blk (2, prts);  | 
|
137  | 
val strs = block o breaks o map str;  | 
|
138  | 
||
139  | 
fun markup m prts = markup_block m (0, prts);  | 
|
| 26703 | 140  | 
fun mark m prt = markup m [prt];  | 
141  | 
fun keyword name = mark (Markup.keyword name) (str name);  | 
|
142  | 
fun command name = mark (Markup.command name) (str name);  | 
|
| 23645 | 143  | 
|
144  | 
fun markup_chunks m prts = markup m (fbreaks prts);  | 
|
145  | 
val chunks = markup_chunks Markup.none;  | 
|
146  | 
fun chunks2 prts = blk (0, flat (Library.separate [fbrk, fbrk] (map single prts)));  | 
|
147  | 
||
148  | 
fun block_enclose (p1, p2) ps = chunks [(block o fbreaks) (p1 :: ps), p2];  | 
|
149  | 
||
150  | 
fun quote prt = blk (1, [str "\"", prt, str "\""]);  | 
|
151  | 
fun backquote prt = blk (1, [str "`", prt, str "`"]);  | 
|
152  | 
||
153  | 
fun separate sep prts =  | 
|
154  | 
flat (Library.separate [str sep, brk 1] (map single prts));  | 
|
155  | 
||
156  | 
val commas = separate ",";  | 
|
157  | 
||
158  | 
fun enclose lpar rpar prts =  | 
|
159  | 
block (str lpar :: (prts @ [str rpar]));  | 
|
160  | 
||
161  | 
fun enum sep lpar rpar prts = enclose lpar rpar (separate sep prts);  | 
|
| 
10952
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
162  | 
|
| 30620 | 163  | 
val position =  | 
164  | 
  enum "," "{" "}" o map (fn (x, y) => str (x ^ "=" ^ y)) o Position.properties_of;
 | 
|
165  | 
||
| 23645 | 166  | 
val list = enum ",";  | 
167  | 
fun str_list lpar rpar strs = list lpar rpar (map str strs);  | 
|
168  | 
||
169  | 
fun big_list name prts = block (fbreaks (str name :: prts));  | 
|
170  | 
||
171  | 
fun indent 0 prt = prt  | 
|
172  | 
| indent n prt = blk (0, [str (Symbol.spaces n), prt]);  | 
|
173  | 
||
174  | 
fun unbreakable (Break (_, wd)) = String (output_spaces wd, wd)  | 
|
175  | 
| unbreakable (Block (m, es, indent, wd)) = Block (m, map unbreakable es, indent, wd)  | 
|
176  | 
| unbreakable (e as String _) = e;  | 
|
177  | 
||
178  | 
||
179  | 
||
180  | 
(** formatting **)  | 
|
181  | 
||
182  | 
(* margin *)  | 
|
183  | 
||
184  | 
fun make_margin_info m =  | 
|
185  | 
 {margin = m,                   (*right margin, or page width*)
 | 
|
186  | 
breakgain = m div 20, (*minimum added space required of a break*)  | 
|
187  | 
emergencypos = m div 2}; (*position too far to right*)  | 
|
188  | 
||
| 32738 | 189  | 
val margin_info = Unsynchronized.ref (make_margin_info 76);  | 
| 23645 | 190  | 
fun setmargin m = margin_info := make_margin_info m;  | 
| 
32966
 
5b21661fe618
indicate CRITICAL nature of various setmp combinators;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
191  | 
fun setmp_margin_CRITICAL m f = setmp_CRITICAL margin_info (make_margin_info m) f;  | 
| 23645 | 192  | 
|
193  | 
||
194  | 
(* depth limitation *)  | 
|
195  | 
||
| 32738 | 196  | 
val depth = Unsynchronized.ref 0; (*maximum depth; 0 means no limit*)  | 
| 23645 | 197  | 
fun setdepth dp = (depth := dp);  | 
198  | 
||
199  | 
local  | 
|
| 32784 | 200  | 
fun pruning dp (Block (m, bes, indent, _)) =  | 
| 23645 | 201  | 
if dp > 0  | 
| 
30667
 
53fbf7c679b0
Block markup: maintain output version within tree values (in accordance with String) -- changes operational behaviour wrt. print_mode;
 
wenzelm 
parents: 
30624 
diff
changeset
 | 
202  | 
then block_markup m (indent, map (pruning (dp - 1)) bes)  | 
| 23645 | 203  | 
else str "..."  | 
| 32784 | 204  | 
| pruning _ e = e;  | 
| 23645 | 205  | 
in  | 
206  | 
fun prune e = if ! depth > 0 then pruning (! depth) e else e;  | 
|
207  | 
end;  | 
|
208  | 
||
209  | 
||
210  | 
(* formatted output *)  | 
|
211  | 
||
212  | 
local  | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
213  | 
|
| 17756 | 214  | 
type text = {tx: Buffer.T, ind: Buffer.T, pos: int, nl: int};
 | 
215  | 
||
216  | 
val empty: text =  | 
|
| 
10952
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
217  | 
 {tx = Buffer.empty,
 | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
218  | 
ind = Buffer.empty,  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
219  | 
pos = 0,  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
220  | 
nl = 0};  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
221  | 
|
| 32784 | 222  | 
fun newline {tx, ind = _, pos = _, nl} : text =
 | 
| 
14832
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
223  | 
 {tx = Buffer.add (Output.output "\n") tx,
 | 
| 
10952
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
224  | 
ind = Buffer.empty,  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
225  | 
pos = 0,  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
226  | 
nl = nl + 1};  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
227  | 
|
| 
23628
 
41cdbfb9f77b
markup: emit as control information -- no indent text;
 
wenzelm 
parents: 
23617 
diff
changeset
 | 
228  | 
fun control s {tx, ind, pos: int, nl} : text =
 | 
| 
 
41cdbfb9f77b
markup: emit as control information -- no indent text;
 
wenzelm 
parents: 
23617 
diff
changeset
 | 
229  | 
 {tx = Buffer.add s tx,
 | 
| 
 
41cdbfb9f77b
markup: emit as control information -- no indent text;
 
wenzelm 
parents: 
23617 
diff
changeset
 | 
230  | 
ind = ind,  | 
| 
 
41cdbfb9f77b
markup: emit as control information -- no indent text;
 
wenzelm 
parents: 
23617 
diff
changeset
 | 
231  | 
pos = pos,  | 
| 
 
41cdbfb9f77b
markup: emit as control information -- no indent text;
 
wenzelm 
parents: 
23617 
diff
changeset
 | 
232  | 
nl = nl};  | 
| 
 
41cdbfb9f77b
markup: emit as control information -- no indent text;
 
wenzelm 
parents: 
23617 
diff
changeset
 | 
233  | 
|
| 17756 | 234  | 
fun string (s, len) {tx, ind, pos: int, nl} : text =
 | 
| 
10952
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
235  | 
 {tx = Buffer.add s tx,
 | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
236  | 
ind = Buffer.add s ind,  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
237  | 
pos = pos + len,  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
238  | 
nl = nl};  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
239  | 
|
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
240  | 
fun blanks wd = string (output_spaces wd, wd);  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
241  | 
|
| 17756 | 242  | 
fun indentation (buf, len) {tx, ind, pos, nl} : text =
 | 
| 
10952
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
243  | 
let val s = Buffer.content buf in  | 
| 23617 | 244  | 
   {tx = Buffer.add (mode_indent s len) tx,
 | 
| 
10952
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
245  | 
ind = Buffer.add s ind,  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
246  | 
pos = pos + len,  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
247  | 
nl = nl}  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
248  | 
end;  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
249  | 
|
| 
10952
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
250  | 
(*Add the lengths of the expressions until the next Break; if no Break then  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
251  | 
include "after", to account for text following this block.*)  | 
| 23617 | 252  | 
fun breakdist (Block (_, _, _, len) :: es, after) = len + breakdist (es, after)  | 
| 32784 | 253  | 
| breakdist (String (_, len) :: es, after) = len + breakdist (es, after)  | 
254  | 
| breakdist (Break _ :: _, _) = 0  | 
|
| 
10952
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
255  | 
| breakdist ([], after) = after;  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
256  | 
|
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
257  | 
(*Search for the next break (at this or higher levels) and force it to occur.*)  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
258  | 
fun forcenext [] = []  | 
| 32784 | 259  | 
| forcenext (Break _ :: es) = Break (true, 0) :: es  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
260  | 
| forcenext (e :: es) = e :: forcenext es;  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
261  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
262  | 
(*es is list of expressions to print;  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
263  | 
blockin is the indentation of the current block;  | 
| 
10952
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
264  | 
after is the width of the following context until next break.*)  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
265  | 
fun format ([], _, _) text = text  | 
| 32784 | 266  | 
  | format (e :: es, block as (_, blockin), after) (text as {ind, pos, nl, ...}) =
 | 
| 
10952
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
267  | 
(case e of  | 
| 32784 | 268  | 
Block ((bg, en), bes, indent, _) =>  | 
| 
10952
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
269  | 
let  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
270  | 
            val {emergencypos, ...} = ! margin_info;
 | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
271  | 
val pos' = pos + indent;  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
272  | 
val pos'' = pos' mod emergencypos;  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
273  | 
val block' =  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
274  | 
if pos' < emergencypos then (ind |> add_indent indent, pos')  | 
| 23645 | 275  | 
else (add_indent pos'' Buffer.empty, pos'');  | 
| 23617 | 276  | 
val btext: text = text  | 
| 
23628
 
41cdbfb9f77b
markup: emit as control information -- no indent text;
 
wenzelm 
parents: 
23617 
diff
changeset
 | 
277  | 
|> control bg  | 
| 23617 | 278  | 
|> format (bes, block', breakdist (es, after))  | 
| 
23628
 
41cdbfb9f77b
markup: emit as control information -- no indent text;
 
wenzelm 
parents: 
23617 
diff
changeset
 | 
279  | 
|> control en;  | 
| 
10952
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
280  | 
(*if this block was broken then force the next break*)  | 
| 23617 | 281  | 
val es' = if nl < #nl btext then forcenext es else es;  | 
282  | 
in format (es', block, after) btext end  | 
|
| 
10952
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
283  | 
| String str => format (es, block, after) (string str text)  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
284  | 
| Break (force, wd) =>  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
285  | 
          let val {margin, breakgain, ...} = ! margin_info in
 | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
286  | 
(*no break if text to next break fits on this line  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
287  | 
or if breaking would add only breakgain to space*)  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
288  | 
format (es, block, after)  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
289  | 
(if not force andalso  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
290  | 
pos + wd <= Int.max (margin - breakdist (es, after), blockin + breakgain)  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
291  | 
then text |> blanks wd (*just insert wd blanks*)  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
292  | 
else text |> newline |> indentation block)  | 
| 
 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 
wenzelm 
parents: 
9730 
diff
changeset
 | 
293  | 
end);  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
294  | 
|
| 23645 | 295  | 
in  | 
| 19266 | 296  | 
|
| 23645 | 297  | 
fun formatted e = #tx (format ([prune e], (Buffer.empty, 0), 0) empty);  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
298  | 
|
| 23645 | 299  | 
end;  | 
| 
14832
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
300  | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
301  | 
|
| 23645 | 302  | 
(* special output *)  | 
| 18802 | 303  | 
|
| 23645 | 304  | 
(*symbolic markup -- no formatting*)  | 
305  | 
fun symbolic prt =  | 
|
306  | 
let  | 
|
| 
30667
 
53fbf7c679b0
Block markup: maintain output version within tree values (in accordance with String) -- changes operational behaviour wrt. print_mode;
 
wenzelm 
parents: 
30624 
diff
changeset
 | 
307  | 
fun out (Block ((bg, en), [], _, _)) = Buffer.add bg #> Buffer.add en  | 
| 
 
53fbf7c679b0
Block markup: maintain output version within tree values (in accordance with String) -- changes operational behaviour wrt. print_mode;
 
wenzelm 
parents: 
30624 
diff
changeset
 | 
308  | 
| out (Block ((bg, en), prts, indent, _)) =  | 
| 
 
53fbf7c679b0
Block markup: maintain output version within tree values (in accordance with String) -- changes operational behaviour wrt. print_mode;
 
wenzelm 
parents: 
30624 
diff
changeset
 | 
309  | 
Buffer.add bg #> Buffer.markup (Markup.block indent) (fold out prts) #> Buffer.add en  | 
| 23645 | 310  | 
| out (String (s, _)) = Buffer.add s  | 
| 23787 | 311  | 
| out (Break (false, wd)) = Buffer.markup (Markup.break wd) (Buffer.add (output_spaces wd))  | 
312  | 
| out (Break (true, _)) = Buffer.markup Markup.fbreak (Buffer.add (output_spaces 1));  | 
|
| 23645 | 313  | 
in out prt Buffer.empty end;  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
314  | 
|
| 23645 | 315  | 
(*unformatted output*)  | 
316  | 
fun unformatted prt =  | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
317  | 
let  | 
| 
30667
 
53fbf7c679b0
Block markup: maintain output version within tree values (in accordance with String) -- changes operational behaviour wrt. print_mode;
 
wenzelm 
parents: 
30624 
diff
changeset
 | 
318  | 
fun fmt (Block ((bg, en), prts, _, _)) = Buffer.add bg #> fold fmt prts #> Buffer.add en  | 
| 23617 | 319  | 
| fmt (String (s, _)) = Buffer.add s  | 
320  | 
| fmt (Break (false, wd)) = Buffer.add (output_spaces wd)  | 
|
321  | 
| fmt (Break (true, _)) = Buffer.add (output_spaces 1);  | 
|
| 23645 | 322  | 
in fmt (prune prt) Buffer.empty end;  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
323  | 
|
| 
30624
 
e755b8b76365
simplified datatype ML_Pretty.pretty: model Isabelle not Poly/ML;
 
wenzelm 
parents: 
30620 
diff
changeset
 | 
324  | 
|
| 
 
e755b8b76365
simplified datatype ML_Pretty.pretty: model Isabelle not Poly/ML;
 
wenzelm 
parents: 
30620 
diff
changeset
 | 
325  | 
(* ML toplevel pretty printing *)  | 
| 23645 | 326  | 
|
| 
30667
 
53fbf7c679b0
Block markup: maintain output version within tree values (in accordance with String) -- changes operational behaviour wrt. print_mode;
 
wenzelm 
parents: 
30624 
diff
changeset
 | 
327  | 
fun to_ML (Block (m, prts, ind, _)) = ML_Pretty.Block (m, map to_ML prts, ind)  | 
| 
30624
 
e755b8b76365
simplified datatype ML_Pretty.pretty: model Isabelle not Poly/ML;
 
wenzelm 
parents: 
30620 
diff
changeset
 | 
328  | 
| to_ML (String s) = ML_Pretty.String s  | 
| 
 
e755b8b76365
simplified datatype ML_Pretty.pretty: model Isabelle not Poly/ML;
 
wenzelm 
parents: 
30620 
diff
changeset
 | 
329  | 
| to_ML (Break b) = ML_Pretty.Break b;  | 
| 30620 | 330  | 
|
| 
30667
 
53fbf7c679b0
Block markup: maintain output version within tree values (in accordance with String) -- changes operational behaviour wrt. print_mode;
 
wenzelm 
parents: 
30624 
diff
changeset
 | 
331  | 
fun from_ML (ML_Pretty.Block (m, prts, ind)) = block_markup m (ind, map from_ML prts)  | 
| 
30624
 
e755b8b76365
simplified datatype ML_Pretty.pretty: model Isabelle not Poly/ML;
 
wenzelm 
parents: 
30620 
diff
changeset
 | 
332  | 
| from_ML (ML_Pretty.String s) = String s  | 
| 
 
e755b8b76365
simplified datatype ML_Pretty.pretty: model Isabelle not Poly/ML;
 
wenzelm 
parents: 
30620 
diff
changeset
 | 
333  | 
| from_ML (ML_Pretty.Break b) = Break b;  | 
| 30620 | 334  | 
|
| 23645 | 335  | 
|
336  | 
(* output interfaces *)  | 
|
337  | 
||
338  | 
val symbolicN = "pretty_symbolic";  | 
|
339  | 
||
340  | 
fun output_buffer prt =  | 
|
341  | 
if print_mode_active symbolicN then symbolic prt  | 
|
342  | 
else formatted prt;  | 
|
343  | 
||
344  | 
val output = Buffer.content o output_buffer;  | 
|
345  | 
val string_of = Output.escape o output;  | 
|
346  | 
val str_of = Output.escape o Buffer.content o unformatted;  | 
|
347  | 
val writeln = Output.writeln o string_of;  | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
348  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
349  | 
|
| 
14832
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
350  | 
|
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
351  | 
(** abstract pretty printing context **)  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
352  | 
|
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
353  | 
datatype pp =  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
354  | 
PP of (term -> T) * (typ -> T) * (sort -> T) * (class list -> T) * (arity -> T);  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
355  | 
|
| 14972 | 356  | 
val pp = PP;  | 
| 
14832
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
357  | 
|
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
358  | 
fun pp_fun f (PP x) = f x;  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
359  | 
|
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
360  | 
val term = pp_fun #1;  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
361  | 
val typ = pp_fun #2;  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
362  | 
val sort = pp_fun #3;  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
363  | 
val classrel = pp_fun #4;  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
364  | 
val arity = pp_fun #5;  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
365  | 
|
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
366  | 
val string_of_term = string_of oo term;  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
367  | 
val string_of_typ = string_of oo typ;  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
368  | 
val string_of_sort = string_of oo sort;  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
369  | 
val string_of_classrel = string_of oo classrel;  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
370  | 
val string_of_arity = string_of oo arity;  | 
| 
 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 
wenzelm 
parents: 
12421 
diff
changeset
 | 
371  | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
372  | 
end;  |