| author | wenzelm |
| Sun, 20 Dec 2015 12:48:56 +0100 | |
| changeset 61874 | a942e237c9e8 |
| parent 61870 | 26f976d5dc4a |
| child 61877 | 276ad4354069 |
| 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 |
| 61864 | 15 |
is broken then all enclosing blocks will also be broken. |
|
6116
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
wenzelm
parents:
diff
changeset
|
16 |
|
| 61864 | 17 |
The stored length of a block is used in break_dist (to treat each inner block as |
|
6116
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
wenzelm
parents:
diff
changeset
|
18 |
a unit for breaking). |
|
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
wenzelm
parents:
diff
changeset
|
19 |
*) |
|
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 |
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
|
22 |
sig |
|
40131
7cbebd636e79
explicitly qualify type Output.output, which is a slightly odd internal feature;
wenzelm
parents:
38474
diff
changeset
|
23 |
val default_indent: string -> int -> Output.output |
|
7cbebd636e79
explicitly qualify type Output.output, which is a slightly odd internal feature;
wenzelm
parents:
38474
diff
changeset
|
24 |
val add_mode: string -> (string -> int -> Output.output) -> unit |
|
6116
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
wenzelm
parents:
diff
changeset
|
25 |
type T |
| 61864 | 26 |
val make_block: Output.output * Output.output -> bool -> int -> T list -> T |
|
6116
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 |
|
61862
e2a9e46ac0fb
support pretty break indent, like underlying ML systems;
wenzelm
parents:
56334
diff
changeset
|
29 |
val brk_indent: int -> int -> T |
|
6116
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
wenzelm
parents:
diff
changeset
|
30 |
val fbrk: T |
|
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
wenzelm
parents:
diff
changeset
|
31 |
val breaks: T list -> T list |
|
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
wenzelm
parents:
diff
changeset
|
32 |
val fbreaks: T list -> T list |
| 23645 | 33 |
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
|
34 |
val block: T list -> T |
| 23645 | 35 |
val strs: string list -> T |
| 23617 | 36 |
val markup: Markup.T -> T list -> T |
| 26703 | 37 |
val mark: Markup.T -> T -> T |
| 42266 | 38 |
val mark_str: Markup.T * string -> T |
39 |
val marks_str: Markup.T list * string -> T |
|
|
51570
3633828d80fc
basic support for Pretty.item, which is considered as logical markup and interpreted in Isabelle/Scala, but ignored elsewhere (TTY, latex etc.);
wenzelm
parents:
51511
diff
changeset
|
40 |
val item: T list -> T |
| 52693 | 41 |
val text_fold: T list -> T |
| 55763 | 42 |
val keyword1: string -> T |
43 |
val keyword2: string -> T |
|
| 50162 | 44 |
val text: string -> T list |
45 |
val paragraph: T list -> T |
|
46 |
val para: string -> T |
|
| 18802 | 47 |
val quote: T -> T |
48 |
val backquote: T -> T |
|
| 55033 | 49 |
val cartouche: T -> T |
| 18802 | 50 |
val separate: string -> T list -> T list |
51 |
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
|
52 |
val enclose: string -> string -> T list -> T |
| 18802 | 53 |
val enum: string -> string -> string -> T list -> T |
| 30620 | 54 |
val position: Position.T -> T |
|
6116
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
wenzelm
parents:
diff
changeset
|
55 |
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
|
56 |
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
|
57 |
val big_list: string -> T list -> T |
| 9730 | 58 |
val indent: int -> T -> T |
| 23645 | 59 |
val unbreakable: T -> T |
| 36745 | 60 |
val margin_default: int Unsynchronized.ref |
| 23645 | 61 |
val symbolicN: string |
| 36745 | 62 |
val output_buffer: int option -> T -> Buffer.T |
|
40131
7cbebd636e79
explicitly qualify type Output.output, which is a slightly odd internal feature;
wenzelm
parents:
38474
diff
changeset
|
63 |
val output: int option -> T -> Output.output |
| 36745 | 64 |
val string_of_margin: int -> T -> string |
| 23645 | 65 |
val string_of: T -> string |
|
49656
7ff712de5747
treat wrapped markup elements as raw markup delimiters;
wenzelm
parents:
49565
diff
changeset
|
66 |
val writeln: T -> unit |
|
7ff712de5747
treat wrapped markup elements as raw markup delimiters;
wenzelm
parents:
49565
diff
changeset
|
67 |
val symbolic_output: T -> Output.output |
|
49565
ea4308b7ef0f
ML support for generic graph display, with browser and graphview backends (via print modes);
wenzelm
parents:
49554
diff
changeset
|
68 |
val symbolic_string_of: T -> string |
|
6116
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
wenzelm
parents:
diff
changeset
|
69 |
val str_of: T -> string |
|
56334
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
70 |
val markup_chunks: Markup.T -> T list -> T |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
71 |
val chunks: T list -> T |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
72 |
val chunks2: T list -> T |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
73 |
val block_enclose: T * T -> T list -> T |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
74 |
val writeln_chunks: T list -> unit |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
75 |
val writeln_chunks2: T list -> unit |
| 36748 | 76 |
val to_ML: T -> ML_Pretty.pretty |
77 |
val from_ML: ML_Pretty.pretty -> T |
|
|
14832
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 |
||
| 61865 | 85 |
fun default_indent (_: string) = Symbol.spaces; |
| 23617 | 86 |
|
87 |
local |
|
| 23698 | 88 |
val default = {indent = default_indent};
|
| 43684 | 89 |
val modes = Synchronized.var "Pretty.modes" (Symtab.make [("", default)]);
|
| 23617 | 90 |
in |
| 43684 | 91 |
fun add_mode name indent = |
|
46894
e2ad717ec889
allow redefining pretty/markup modes (not output due to bootstrap issues) -- to support reloading of theory src/HOL/src/Tools/Code_Generator;
wenzelm
parents:
45666
diff
changeset
|
92 |
Synchronized.change modes (fn tab => |
|
e2ad717ec889
allow redefining pretty/markup modes (not output due to bootstrap issues) -- to support reloading of theory src/HOL/src/Tools/Code_Generator;
wenzelm
parents:
45666
diff
changeset
|
93 |
(if not (Symtab.defined tab name) then () |
|
e2ad717ec889
allow redefining pretty/markup modes (not output due to bootstrap issues) -- to support reloading of theory src/HOL/src/Tools/Code_Generator;
wenzelm
parents:
45666
diff
changeset
|
94 |
else warning ("Redefining pretty mode " ^ quote name);
|
|
e2ad717ec889
allow redefining pretty/markup modes (not output due to bootstrap issues) -- to support reloading of theory src/HOL/src/Tools/Code_Generator;
wenzelm
parents:
45666
diff
changeset
|
95 |
Symtab.update (name, {indent = indent}) tab));
|
| 23617 | 96 |
fun get_mode () = |
| 43684 | 97 |
the_default default |
98 |
(Library.get_first (Symtab.lookup (Synchronized.value modes)) (print_mode_value ())); |
|
| 23617 | 99 |
end; |
100 |
||
101 |
fun mode_indent x y = #indent (get_mode ()) x y; |
|
| 23645 | 102 |
|
| 61865 | 103 |
val output_spaces = Output.output o Symbol.spaces; |
| 23645 | 104 |
val add_indent = Buffer.add o output_spaces; |
| 23617 | 105 |
|
106 |
||
|
10952
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
wenzelm
parents:
9730
diff
changeset
|
107 |
|
|
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
wenzelm
parents:
9730
diff
changeset
|
108 |
(** printing items: compound phrases, strings, and breaks **) |
|
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
wenzelm
parents:
9730
diff
changeset
|
109 |
|
| 37529 | 110 |
abstype T = |
| 61869 | 111 |
Block of (Output.output * Output.output) * bool * int * T list * int |
112 |
(*markup output, consistent, indentation, body, length*) |
|
|
61862
e2a9e46ac0fb
support pretty break indent, like underlying ML systems;
wenzelm
parents:
56334
diff
changeset
|
113 |
| Break of bool * int * int (*mandatory flag, width if not taken, extra indentation if taken*) |
| 61870 | 114 |
| Str of Output.output * int (*text, length*) |
| 37529 | 115 |
with |
|
6116
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
wenzelm
parents:
diff
changeset
|
116 |
|
| 61864 | 117 |
fun length (Block (_, _, _, _, len)) = len |
| 61870 | 118 |
| length (Break (_, wd, _)) = wd |
119 |
| length (Str (_, len)) = len; |
|
| 23645 | 120 |
|
| 61864 | 121 |
fun body_length [] len = len |
122 |
| body_length (e :: es) len = body_length es (length e + len); |
|
123 |
||
| 61869 | 124 |
fun make_block m consistent indent es = Block (m, consistent, indent, es, body_length es 0); |
| 61864 | 125 |
fun markup_block m indent es = make_block (Markup.output m) false indent es; |
126 |
||
| 9730 | 127 |
|
|
6116
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
wenzelm
parents:
diff
changeset
|
128 |
|
| 23645 | 129 |
(** derived operations to create formatting expressions **) |
130 |
||
| 61870 | 131 |
val str = Str o Output.output_width; |
| 23645 | 132 |
|
|
61862
e2a9e46ac0fb
support pretty break indent, like underlying ML systems;
wenzelm
parents:
56334
diff
changeset
|
133 |
fun brk wd = Break (false, wd, 0); |
|
e2a9e46ac0fb
support pretty break indent, like underlying ML systems;
wenzelm
parents:
56334
diff
changeset
|
134 |
fun brk_indent wd ind = Break (false, wd, ind); |
|
e2a9e46ac0fb
support pretty break indent, like underlying ML systems;
wenzelm
parents:
56334
diff
changeset
|
135 |
val fbrk = Break (true, 1, 0); |
| 23645 | 136 |
|
137 |
fun breaks prts = Library.separate (brk 1) prts; |
|
138 |
fun fbreaks prts = Library.separate fbrk prts; |
|
139 |
||
| 61864 | 140 |
fun blk (ind, es) = markup_block Markup.empty ind es; |
| 23645 | 141 |
fun block prts = blk (2, prts); |
142 |
val strs = block o breaks o map str; |
|
143 |
||
| 61864 | 144 |
fun markup m = markup_block m 0; |
| 42266 | 145 |
fun mark m prt = if m = Markup.empty then prt else markup m [prt]; |
146 |
fun mark_str (m, s) = mark m (str s); |
|
147 |
fun marks_str (ms, s) = fold_rev mark ms (str s); |
|
148 |
||
|
51570
3633828d80fc
basic support for Pretty.item, which is considered as logical markup and interpreted in Isabelle/Scala, but ignored elsewhere (TTY, latex etc.);
wenzelm
parents:
51511
diff
changeset
|
149 |
val item = markup Markup.item; |
| 52693 | 150 |
val text_fold = markup Markup.text_fold; |
|
51570
3633828d80fc
basic support for Pretty.item, which is considered as logical markup and interpreted in Isabelle/Scala, but ignored elsewhere (TTY, latex etc.);
wenzelm
parents:
51511
diff
changeset
|
151 |
|
| 55763 | 152 |
fun keyword1 name = mark_str (Markup.keyword1, name); |
153 |
fun keyword2 name = mark_str (Markup.keyword2, name); |
|
| 23645 | 154 |
|
| 50162 | 155 |
val text = breaks o map str o Symbol.explode_words; |
|
50201
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents:
50162
diff
changeset
|
156 |
val paragraph = markup Markup.paragraph; |
| 50162 | 157 |
val para = paragraph o text; |
158 |
||
| 23645 | 159 |
fun quote prt = blk (1, [str "\"", prt, str "\""]); |
160 |
fun backquote prt = blk (1, [str "`", prt, str "`"]); |
|
| 55033 | 161 |
fun cartouche prt = blk (1, [str "\\<open>", prt, str "\\<close>"]); |
| 23645 | 162 |
|
163 |
fun separate sep prts = |
|
164 |
flat (Library.separate [str sep, brk 1] (map single prts)); |
|
165 |
||
166 |
val commas = separate ","; |
|
167 |
||
168 |
fun enclose lpar rpar prts = |
|
169 |
block (str lpar :: (prts @ [str rpar])); |
|
170 |
||
171 |
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
|
172 |
|
| 30620 | 173 |
val position = |
174 |
enum "," "{" "}" o map (fn (x, y) => str (x ^ "=" ^ y)) o Position.properties_of;
|
|
175 |
||
| 23645 | 176 |
val list = enum ","; |
177 |
fun str_list lpar rpar strs = list lpar rpar (map str strs); |
|
178 |
||
179 |
fun big_list name prts = block (fbreaks (str name :: prts)); |
|
180 |
||
181 |
fun indent 0 prt = prt |
|
| 61865 | 182 |
| indent n prt = blk (0, [str (Symbol.spaces n), prt]); |
| 23645 | 183 |
|
| 61870 | 184 |
fun unbreakable (Block (m, consistent, indent, es, len)) = |
| 61869 | 185 |
Block (m, consistent, indent, map unbreakable es, len) |
| 61870 | 186 |
| unbreakable (Break (_, wd, _)) = Str (output_spaces wd, wd) |
187 |
| unbreakable (e as Str _) = e; |
|
| 23645 | 188 |
|
189 |
||
190 |
||
191 |
(** formatting **) |
|
192 |
||
193 |
(* formatted output *) |
|
194 |
||
195 |
local |
|
|
6116
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
wenzelm
parents:
diff
changeset
|
196 |
|
| 17756 | 197 |
type text = {tx: Buffer.T, ind: Buffer.T, pos: int, nl: int};
|
198 |
||
199 |
val empty: text = |
|
|
10952
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
wenzelm
parents:
9730
diff
changeset
|
200 |
{tx = Buffer.empty,
|
|
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
wenzelm
parents:
9730
diff
changeset
|
201 |
ind = Buffer.empty, |
|
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
wenzelm
parents:
9730
diff
changeset
|
202 |
pos = 0, |
|
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
wenzelm
parents:
9730
diff
changeset
|
203 |
nl = 0}; |
|
6116
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
wenzelm
parents:
diff
changeset
|
204 |
|
| 32784 | 205 |
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
|
206 |
{tx = Buffer.add (Output.output "\n") tx,
|
|
10952
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
wenzelm
parents:
9730
diff
changeset
|
207 |
ind = Buffer.empty, |
|
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
wenzelm
parents:
9730
diff
changeset
|
208 |
pos = 0, |
|
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
wenzelm
parents:
9730
diff
changeset
|
209 |
nl = nl + 1}; |
|
6116
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
wenzelm
parents:
diff
changeset
|
210 |
|
|
23628
41cdbfb9f77b
markup: emit as control information -- no indent text;
wenzelm
parents:
23617
diff
changeset
|
211 |
fun control s {tx, ind, pos: int, nl} : text =
|
|
41cdbfb9f77b
markup: emit as control information -- no indent text;
wenzelm
parents:
23617
diff
changeset
|
212 |
{tx = Buffer.add s tx,
|
|
41cdbfb9f77b
markup: emit as control information -- no indent text;
wenzelm
parents:
23617
diff
changeset
|
213 |
ind = ind, |
|
41cdbfb9f77b
markup: emit as control information -- no indent text;
wenzelm
parents:
23617
diff
changeset
|
214 |
pos = pos, |
|
41cdbfb9f77b
markup: emit as control information -- no indent text;
wenzelm
parents:
23617
diff
changeset
|
215 |
nl = nl}; |
|
41cdbfb9f77b
markup: emit as control information -- no indent text;
wenzelm
parents:
23617
diff
changeset
|
216 |
|
| 17756 | 217 |
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
|
218 |
{tx = Buffer.add s tx,
|
|
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
wenzelm
parents:
9730
diff
changeset
|
219 |
ind = Buffer.add s ind, |
|
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
wenzelm
parents:
9730
diff
changeset
|
220 |
pos = pos + len, |
|
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
wenzelm
parents:
9730
diff
changeset
|
221 |
nl = nl}; |
|
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
wenzelm
parents:
9730
diff
changeset
|
222 |
|
|
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
wenzelm
parents:
9730
diff
changeset
|
223 |
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
|
224 |
|
| 17756 | 225 |
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
|
226 |
let val s = Buffer.content buf in |
| 23617 | 227 |
{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
|
228 |
ind = Buffer.add s ind, |
|
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
wenzelm
parents:
9730
diff
changeset
|
229 |
pos = pos + len, |
|
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
wenzelm
parents:
9730
diff
changeset
|
230 |
nl = nl} |
|
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
wenzelm
parents:
9730
diff
changeset
|
231 |
end; |
|
6116
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
wenzelm
parents:
diff
changeset
|
232 |
|
|
10952
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
wenzelm
parents:
9730
diff
changeset
|
233 |
(*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
|
234 |
include "after", to account for text following this block.*) |
| 61864 | 235 |
fun break_dist (Break _ :: _, _) = 0 |
| 61870 | 236 |
| break_dist (e :: es, after) = length e + break_dist (es, after) |
| 61864 | 237 |
| break_dist ([], after) = after; |
238 |
||
239 |
val force_break = fn Break (false, wd, ind) => Break (true, wd, ind) | e => e; |
|
240 |
val force_all = map force_break; |
|
|
10952
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
wenzelm
parents:
9730
diff
changeset
|
241 |
|
|
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
wenzelm
parents:
9730
diff
changeset
|
242 |
(*Search for the next break (at this or higher levels) and force it to occur.*) |
| 61864 | 243 |
fun force_next [] = [] |
244 |
| force_next ((e as Break _) :: es) = force_break e :: es |
|
245 |
| force_next (e :: es) = e :: force_next es; |
|
|
6116
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
wenzelm
parents:
diff
changeset
|
246 |
|
| 23645 | 247 |
in |
| 19266 | 248 |
|
| 36745 | 249 |
fun formatted margin input = |
250 |
let |
|
251 |
val breakgain = margin div 20; (*minimum added space required of a break*) |
|
252 |
val emergencypos = margin div 2; (*position too far to right*) |
|
253 |
||
254 |
(*es is list of expressions to print; |
|
255 |
blockin is the indentation of the current block; |
|
256 |
after is the width of the following context until next break.*) |
|
257 |
fun format ([], _, _) text = text |
|
258 |
| format (e :: es, block as (_, blockin), after) (text as {ind, pos, nl, ...}) =
|
|
259 |
(case e of |
|
| 61869 | 260 |
Block ((bg, en), consistent, indent, bes, blen) => |
| 36745 | 261 |
let |
262 |
val pos' = pos + indent; |
|
263 |
val pos'' = pos' mod emergencypos; |
|
264 |
val block' = |
|
265 |
if pos' < emergencypos then (ind |> add_indent indent, pos') |
|
266 |
else (add_indent pos'' Buffer.empty, pos''); |
|
| 61864 | 267 |
val d = break_dist (es, after) |
268 |
val bes' = if consistent andalso pos + blen > margin - d then force_all bes else bes; |
|
| 36745 | 269 |
val btext: text = text |
270 |
|> control bg |
|
| 61864 | 271 |
|> format (bes', block', d) |
| 36745 | 272 |
|> control en; |
273 |
(*if this block was broken then force the next break*) |
|
| 61864 | 274 |
val es' = if nl < #nl btext then force_next es else es; |
| 36745 | 275 |
in format (es', block, after) btext end |
|
61862
e2a9e46ac0fb
support pretty break indent, like underlying ML systems;
wenzelm
parents:
56334
diff
changeset
|
276 |
| Break (force, wd, ind) => |
| 36745 | 277 |
(*no break if text to next break fits on this line |
278 |
or if breaking would add only breakgain to space*) |
|
279 |
format (es, block, after) |
|
280 |
(if not force andalso |
|
| 61864 | 281 |
pos + wd <= Int.max (margin - break_dist (es, after), blockin + breakgain) |
282 |
then text |> blanks wd (*just insert wd blanks*) |
|
283 |
else text |> newline |> indentation block |> blanks ind) |
|
| 61870 | 284 |
| Str str => format (es, block, after) (string str text)); |
| 36745 | 285 |
in |
|
36747
7361d5dde9ce
discontinued Pretty.setdepth, which appears to be largely unused, but can disrupt important markup if enabled accidentally;
wenzelm
parents:
36745
diff
changeset
|
286 |
#tx (format ([input], (Buffer.empty, 0), 0) empty) |
| 36745 | 287 |
end; |
|
6116
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
wenzelm
parents:
diff
changeset
|
288 |
|
| 23645 | 289 |
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
|
290 |
|
|
6116
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
wenzelm
parents:
diff
changeset
|
291 |
|
| 23645 | 292 |
(* special output *) |
| 18802 | 293 |
|
| 23645 | 294 |
(*symbolic markup -- no formatting*) |
295 |
fun symbolic prt = |
|
296 |
let |
|
| 61869 | 297 |
fun out (Block ((bg, en), _, _, [], _)) = Buffer.add bg #> Buffer.add en |
298 |
| out (Block ((bg, en), consistent, indent, prts, _)) = |
|
| 45666 | 299 |
Buffer.add bg #> |
| 61864 | 300 |
Buffer.markup (Markup.block consistent indent) (fold out prts) #> |
| 45666 | 301 |
Buffer.add en |
|
61862
e2a9e46ac0fb
support pretty break indent, like underlying ML systems;
wenzelm
parents:
56334
diff
changeset
|
302 |
| out (Break (false, wd, ind)) = |
|
e2a9e46ac0fb
support pretty break indent, like underlying ML systems;
wenzelm
parents:
56334
diff
changeset
|
303 |
Buffer.markup (Markup.break wd ind) (Buffer.add (output_spaces wd)) |
| 61870 | 304 |
| out (Break (true, _, _)) = Buffer.add (Output.output "\n") |
305 |
| out (Str (s, _)) = Buffer.add s; |
|
| 23645 | 306 |
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
|
307 |
|
| 23645 | 308 |
(*unformatted output*) |
309 |
fun unformatted prt = |
|
|
6116
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
wenzelm
parents:
diff
changeset
|
310 |
let |
| 61869 | 311 |
fun fmt (Block ((bg, en), _, _, prts, _)) = Buffer.add bg #> fold fmt prts #> Buffer.add en |
| 61870 | 312 |
| fmt (Break (_, wd, _)) = Buffer.add (output_spaces wd) |
313 |
| fmt (Str (s, _)) = Buffer.add s; |
|
|
36747
7361d5dde9ce
discontinued Pretty.setdepth, which appears to be largely unused, but can disrupt important markup if enabled accidentally;
wenzelm
parents:
36745
diff
changeset
|
314 |
in fmt prt Buffer.empty end; |
|
6116
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
wenzelm
parents:
diff
changeset
|
315 |
|
|
30624
e755b8b76365
simplified datatype ML_Pretty.pretty: model Isabelle not Poly/ML;
wenzelm
parents:
30620
diff
changeset
|
316 |
|
| 36748 | 317 |
(* output interfaces *) |
| 30620 | 318 |
|
| 36748 | 319 |
val margin_default = Unsynchronized.ref 76; (*right margin, or page width*) |
| 23645 | 320 |
|
321 |
val symbolicN = "pretty_symbolic"; |
|
322 |
||
| 36745 | 323 |
fun output_buffer margin prt = |
| 23645 | 324 |
if print_mode_active symbolicN then symbolic prt |
| 36745 | 325 |
else formatted (the_default (! margin_default) margin) prt; |
| 23645 | 326 |
|
| 36745 | 327 |
val output = Buffer.content oo output_buffer; |
328 |
fun string_of_margin margin = Output.escape o output (SOME margin); |
|
329 |
val string_of = Output.escape o output NONE; |
|
|
49656
7ff712de5747
treat wrapped markup elements as raw markup delimiters;
wenzelm
parents:
49565
diff
changeset
|
330 |
val writeln = Output.writeln o string_of; |
|
7ff712de5747
treat wrapped markup elements as raw markup delimiters;
wenzelm
parents:
49565
diff
changeset
|
331 |
|
|
7ff712de5747
treat wrapped markup elements as raw markup delimiters;
wenzelm
parents:
49565
diff
changeset
|
332 |
val symbolic_output = Buffer.content o symbolic; |
|
7ff712de5747
treat wrapped markup elements as raw markup delimiters;
wenzelm
parents:
49565
diff
changeset
|
333 |
val symbolic_string_of = Output.escape o symbolic_output; |
|
7ff712de5747
treat wrapped markup elements as raw markup delimiters;
wenzelm
parents:
49565
diff
changeset
|
334 |
|
| 23645 | 335 |
val str_of = Output.escape o Buffer.content o unformatted; |
|
6116
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
wenzelm
parents:
diff
changeset
|
336 |
|
|
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
wenzelm
parents:
diff
changeset
|
337 |
|
|
56334
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
338 |
(* chunks *) |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
339 |
|
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
340 |
fun markup_chunks m prts = markup m (fbreaks (map (text_fold o single) prts)); |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
341 |
val chunks = markup_chunks Markup.empty; |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
342 |
|
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
343 |
fun chunks2 prts = |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
344 |
(case try split_last prts of |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
345 |
NONE => blk (0, []) |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
346 |
| SOME (prefix, last) => |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
347 |
blk (0, maps (fn prt => [text_fold [prt, fbrk], fbrk]) prefix @ [text_fold [last]])); |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
348 |
|
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
349 |
fun block_enclose (prt1, prt2) prts = chunks [block (fbreaks (prt1 :: prts)), prt2]; |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
350 |
|
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
351 |
fun string_of_text_fold prt = string_of prt |> Markup.markup Markup.text_fold; |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
352 |
|
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
353 |
fun writeln_chunks prts = |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
354 |
Output.writelns (Library.separate "\n" (map string_of_text_fold prts)); |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
355 |
|
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
356 |
fun writeln_chunks2 prts = |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
357 |
(case try split_last prts of |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
358 |
NONE => () |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
359 |
| SOME (prefix, last) => |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
360 |
(map (fn prt => Markup.markup Markup.text_fold (string_of prt ^ "\n") ^ "\n") prefix @ |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
361 |
[string_of_text_fold last]) |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
362 |
|> Output.writelns); |
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
363 |
|
|
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
55918
diff
changeset
|
364 |
|
|
14832
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
wenzelm
parents:
12421
diff
changeset
|
365 |
|
| 36748 | 366 |
(** ML toplevel pretty printing **) |
367 |
||
| 61869 | 368 |
fun to_ML (Block (m, consistent, indent, prts, _)) = |
369 |
ML_Pretty.Block (m, consistent, indent, map to_ML prts) |
|
| 61870 | 370 |
| to_ML (Break b) = ML_Pretty.Break b |
371 |
| to_ML (Str s) = ML_Pretty.String s; |
|
| 36748 | 372 |
|
| 61869 | 373 |
fun from_ML (ML_Pretty.Block (m, consistent, indent, prts)) = |
| 61864 | 374 |
make_block m consistent indent (map from_ML prts) |
| 61870 | 375 |
| from_ML (ML_Pretty.Break b) = Break b |
376 |
| from_ML (ML_Pretty.String s) = Str s; |
|
| 36748 | 377 |
|
| 37529 | 378 |
end; |
379 |
||
|
6116
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
wenzelm
parents:
diff
changeset
|
380 |
end; |