| author | wenzelm | 
| Mon, 08 Apr 2013 14:18:39 +0200 | |
| changeset 51634 | 553953ad8165 | 
| parent 51570 | 3633828d80fc | 
| child 52693 | 6651abced106 | 
| 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: 
9730diff
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: 
12421diff
changeset | 23 | sig | 
| 48704 
85a3de10567d
tuned signature -- make Pretty less dependent on Symbol;
 wenzelm parents: 
46894diff
changeset | 24 | val spaces: int -> string | 
| 40131 
7cbebd636e79
explicitly qualify type Output.output, which is a slightly odd internal feature;
 wenzelm parents: 
38474diff
changeset | 25 | val default_indent: string -> int -> Output.output | 
| 
7cbebd636e79
explicitly qualify type Output.output, which is a slightly odd internal feature;
 wenzelm parents: 
38474diff
changeset | 26 | 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 | 27 | type T | 
| 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 wenzelm parents: diff
changeset | 28 | val str: string -> T | 
| 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 wenzelm parents: diff
changeset | 29 | val brk: int -> T | 
| 
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 | 
| 49656 
7ff712de5747
treat wrapped markup elements as raw markup delimiters;
 wenzelm parents: 
49565diff
changeset | 36 | val raw_markup: Output.output * Output.output -> int * T list -> T | 
| 23617 | 37 | val markup: Markup.T -> T list -> T | 
| 26703 | 38 | val mark: Markup.T -> T -> T | 
| 42266 | 39 | val mark_str: Markup.T * string -> T | 
| 40 | 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: 
51511diff
changeset | 41 | val item: T list -> T | 
| 49554 
7b7bd2d7661d
more explicit keyword1/keyword2 markup -- avoid potential conflict with input token markup produced by Token_Marker;
 wenzelm parents: 
48704diff
changeset | 42 | val command: string -> T | 
| 23617 | 43 | val keyword: string -> T | 
| 50162 | 44 | val text: string -> T list | 
| 45 | val paragraph: T list -> T | |
| 46 | val para: string -> T | |
| 23638 | 47 | val markup_chunks: Markup.T -> T list -> T | 
| 18802 | 48 | val chunks: T list -> T | 
| 19266 | 49 | val chunks2: T list -> T | 
| 23617 | 50 | val block_enclose: T * T -> T list -> T | 
| 18802 | 51 | val quote: T -> T | 
| 52 | val backquote: T -> T | |
| 53 | val separate: string -> T list -> T list | |
| 54 | 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 | 55 | val enclose: string -> string -> T list -> T | 
| 18802 | 56 | val enum: string -> string -> string -> T list -> T | 
| 30620 | 57 | val position: Position.T -> T | 
| 6116 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 wenzelm parents: diff
changeset | 58 | 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 | 59 | 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 | 60 | val big_list: string -> T list -> T | 
| 9730 | 61 | val indent: int -> T -> T | 
| 23645 | 62 | val unbreakable: T -> T | 
| 36745 | 63 | val margin_default: int Unsynchronized.ref | 
| 23645 | 64 | val symbolicN: string | 
| 36745 | 65 | val output_buffer: int option -> T -> Buffer.T | 
| 40131 
7cbebd636e79
explicitly qualify type Output.output, which is a slightly odd internal feature;
 wenzelm parents: 
38474diff
changeset | 66 | val output: int option -> T -> Output.output | 
| 36745 | 67 | val string_of_margin: int -> T -> string | 
| 23645 | 68 | val string_of: T -> string | 
| 49656 
7ff712de5747
treat wrapped markup elements as raw markup delimiters;
 wenzelm parents: 
49565diff
changeset | 69 | val writeln: T -> unit | 
| 
7ff712de5747
treat wrapped markup elements as raw markup delimiters;
 wenzelm parents: 
49565diff
changeset | 70 | val symbolic_output: T -> Output.output | 
| 49565 
ea4308b7ef0f
ML support for generic graph display, with browser and graphview backends (via print modes);
 wenzelm parents: 
49554diff
changeset | 71 | 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 | 72 | val str_of: T -> string | 
| 36748 | 73 | val to_ML: T -> ML_Pretty.pretty | 
| 74 | 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: 
12421diff
changeset | 75 | end; | 
| 6116 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 wenzelm parents: diff
changeset | 76 | |
| 23617 | 77 | structure Pretty: PRETTY = | 
| 6116 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 wenzelm parents: diff
changeset | 78 | struct | 
| 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 wenzelm parents: diff
changeset | 79 | |
| 48704 
85a3de10567d
tuned signature -- make Pretty less dependent on Symbol;
 wenzelm parents: 
46894diff
changeset | 80 | (** spaces **) | 
| 
85a3de10567d
tuned signature -- make Pretty less dependent on Symbol;
 wenzelm parents: 
46894diff
changeset | 81 | |
| 
85a3de10567d
tuned signature -- make Pretty less dependent on Symbol;
 wenzelm parents: 
46894diff
changeset | 82 | local | 
| 51491 | 83 | val small_spaces = Vector.tabulate (65, fn i => replicate_string i Symbol.space); | 
| 48704 
85a3de10567d
tuned signature -- make Pretty less dependent on Symbol;
 wenzelm parents: 
46894diff
changeset | 84 | in | 
| 
85a3de10567d
tuned signature -- make Pretty less dependent on Symbol;
 wenzelm parents: 
46894diff
changeset | 85 | fun spaces k = | 
| 
85a3de10567d
tuned signature -- make Pretty less dependent on Symbol;
 wenzelm parents: 
46894diff
changeset | 86 | if k < 64 then Vector.sub (small_spaces, k) | 
| 51491 | 87 | else | 
| 88 | replicate_string (k div 64) (Vector.sub (small_spaces, 64)) ^ | |
| 89 | Vector.sub (small_spaces, k mod 64); | |
| 48704 
85a3de10567d
tuned signature -- make Pretty less dependent on Symbol;
 wenzelm parents: 
46894diff
changeset | 90 | end; | 
| 
85a3de10567d
tuned signature -- make Pretty less dependent on Symbol;
 wenzelm parents: 
46894diff
changeset | 91 | |
| 
85a3de10567d
tuned signature -- make Pretty less dependent on Symbol;
 wenzelm parents: 
46894diff
changeset | 92 | |
| 
85a3de10567d
tuned signature -- make Pretty less dependent on Symbol;
 wenzelm parents: 
46894diff
changeset | 93 | |
| 23617 | 94 | (** print mode operations **) | 
| 95 | ||
| 48704 
85a3de10567d
tuned signature -- make Pretty less dependent on Symbol;
 wenzelm parents: 
46894diff
changeset | 96 | fun default_indent (_: string) = spaces; | 
| 23617 | 97 | |
| 98 | local | |
| 23698 | 99 |   val default = {indent = default_indent};
 | 
| 43684 | 100 |   val modes = Synchronized.var "Pretty.modes" (Symtab.make [("", default)]);
 | 
| 23617 | 101 | in | 
| 43684 | 102 | 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: 
45666diff
changeset | 103 | 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: 
45666diff
changeset | 104 | (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: 
45666diff
changeset | 105 |        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: 
45666diff
changeset | 106 |        Symtab.update (name, {indent = indent}) tab));
 | 
| 23617 | 107 | fun get_mode () = | 
| 43684 | 108 | the_default default | 
| 109 | (Library.get_first (Symtab.lookup (Synchronized.value modes)) (print_mode_value ())); | |
| 23617 | 110 | end; | 
| 111 | ||
| 112 | fun mode_indent x y = #indent (get_mode ()) x y; | |
| 23645 | 113 | |
| 48704 
85a3de10567d
tuned signature -- make Pretty less dependent on Symbol;
 wenzelm parents: 
46894diff
changeset | 114 | val output_spaces = Output.output o spaces; | 
| 23645 | 115 | val add_indent = Buffer.add o output_spaces; | 
| 23617 | 116 | |
| 117 | ||
| 10952 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 118 | |
| 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 119 | (** printing items: compound phrases, strings, and breaks **) | 
| 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 120 | |
| 37529 | 121 | abstype T = | 
| 40131 
7cbebd636e79
explicitly qualify type Output.output, which is a slightly odd internal feature;
 wenzelm parents: 
38474diff
changeset | 122 | Block of (Output.output * Output.output) * T list * int * int | 
| 
7cbebd636e79
explicitly qualify type Output.output, which is a slightly odd internal feature;
 wenzelm parents: 
38474diff
changeset | 123 | (*markup output, body, indentation, length*) | 
| 
7cbebd636e79
explicitly qualify type Output.output, which is a slightly odd internal feature;
 wenzelm parents: 
38474diff
changeset | 124 | | String of Output.output * int (*text, length*) | 
| 
7cbebd636e79
explicitly qualify type Output.output, which is a slightly odd internal feature;
 wenzelm parents: 
38474diff
changeset | 125 | | Break of bool * int (*mandatory flag, width if not taken*) | 
| 37529 | 126 | with | 
| 6116 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 wenzelm parents: diff
changeset | 127 | |
| 23645 | 128 | fun length (Block (_, _, _, len)) = len | 
| 129 | | length (String (_, len)) = len | |
| 130 | | length (Break (_, wd)) = wd; | |
| 131 | ||
| 9730 | 132 | |
| 6116 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 wenzelm parents: diff
changeset | 133 | |
| 23645 | 134 | (** derived operations to create formatting expressions **) | 
| 135 | ||
| 136 | val str = String o Output.output_width; | |
| 137 | ||
| 138 | fun brk wd = Break (false, wd); | |
| 36690 
97d2780ad6f0
uniform treatment of length = 1 for forced breaks, also makes ML/Pretty.length coincide with Scala/XML.content_length;
 wenzelm parents: 
36689diff
changeset | 139 | val fbrk = Break (true, 1); | 
| 23645 | 140 | |
| 141 | fun breaks prts = Library.separate (brk 1) prts; | |
| 142 | fun fbreaks prts = Library.separate fbrk prts; | |
| 143 | ||
| 49656 
7ff712de5747
treat wrapped markup elements as raw markup delimiters;
 wenzelm parents: 
49565diff
changeset | 144 | fun raw_markup m (indent, es) = | 
| 23645 | 145 | let | 
| 146 | fun sum [] k = k | |
| 147 | | sum (e :: es) k = sum es (length e + k); | |
| 148 | in Block (m, es, indent, sum es 0) end; | |
| 149 | ||
| 49656 
7ff712de5747
treat wrapped markup elements as raw markup delimiters;
 wenzelm parents: 
49565diff
changeset | 150 | fun markup_block m arg = raw_markup (Markup.output m) arg; | 
| 30667 
53fbf7c679b0
Block markup: maintain output version within tree values (in accordance with String) -- changes operational behaviour wrt. print_mode;
 wenzelm parents: 
30624diff
changeset | 151 | |
| 38474 
e498dc2eb576
uniform Markup.empty/Markup.Empty in ML and Scala;
 wenzelm parents: 
37529diff
changeset | 152 | val blk = markup_block Markup.empty; | 
| 23645 | 153 | fun block prts = blk (2, prts); | 
| 154 | val strs = block o breaks o map str; | |
| 155 | ||
| 156 | fun markup m prts = markup_block m (0, prts); | |
| 42266 | 157 | fun mark m prt = if m = Markup.empty then prt else markup m [prt]; | 
| 158 | fun mark_str (m, s) = mark m (str s); | |
| 159 | fun marks_str (ms, s) = fold_rev mark ms (str s); | |
| 160 | ||
| 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: 
51511diff
changeset | 161 | val item = markup Markup.item; | 
| 
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: 
51511diff
changeset | 162 | |
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
50162diff
changeset | 163 | fun command name = mark_str (Markup.keyword1, name); | 
| 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
50162diff
changeset | 164 | fun keyword name = mark_str (Markup.keyword2, name); | 
| 23645 | 165 | |
| 50162 | 166 | 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: 
50162diff
changeset | 167 | val paragraph = markup Markup.paragraph; | 
| 50162 | 168 | val para = paragraph o text; | 
| 169 | ||
| 50545 
00bdc48c5f71
explicit text_fold markup, which is used by default in Pretty.chunks/chunks2;
 wenzelm parents: 
50201diff
changeset | 170 | fun markup_chunks m prts = markup m (fbreaks (map (mark Markup.text_fold) prts)); | 
| 38474 
e498dc2eb576
uniform Markup.empty/Markup.Empty in ML and Scala;
 wenzelm parents: 
37529diff
changeset | 171 | val chunks = markup_chunks Markup.empty; | 
| 50545 
00bdc48c5f71
explicit text_fold markup, which is used by default in Pretty.chunks/chunks2;
 wenzelm parents: 
50201diff
changeset | 172 | |
| 
00bdc48c5f71
explicit text_fold markup, which is used by default in Pretty.chunks/chunks2;
 wenzelm parents: 
50201diff
changeset | 173 | fun chunks2 prts = | 
| 51511 | 174 | (case try split_last prts of | 
| 175 | NONE => blk (0, []) | |
| 176 | | SOME (prefix, last) => | |
| 177 | blk (0, | |
| 178 | maps (fn prt => [markup Markup.text_fold [prt, fbrk], fbrk]) prefix @ | |
| 179 | [mark Markup.text_fold last])); | |
| 23645 | 180 | |
| 36733 | 181 | fun block_enclose (prt1, prt2) prts = chunks [block (fbreaks (prt1 :: prts)), prt2]; | 
| 23645 | 182 | |
| 183 | fun quote prt = blk (1, [str "\"", prt, str "\""]); | |
| 184 | fun backquote prt = blk (1, [str "`", prt, str "`"]); | |
| 185 | ||
| 186 | fun separate sep prts = | |
| 187 | flat (Library.separate [str sep, brk 1] (map single prts)); | |
| 188 | ||
| 189 | val commas = separate ","; | |
| 190 | ||
| 191 | fun enclose lpar rpar prts = | |
| 192 | block (str lpar :: (prts @ [str rpar])); | |
| 193 | ||
| 194 | 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: 
9730diff
changeset | 195 | |
| 30620 | 196 | val position = | 
| 197 |   enum "," "{" "}" o map (fn (x, y) => str (x ^ "=" ^ y)) o Position.properties_of;
 | |
| 198 | ||
| 23645 | 199 | val list = enum ","; | 
| 200 | fun str_list lpar rpar strs = list lpar rpar (map str strs); | |
| 201 | ||
| 202 | fun big_list name prts = block (fbreaks (str name :: prts)); | |
| 203 | ||
| 204 | fun indent 0 prt = prt | |
| 48704 
85a3de10567d
tuned signature -- make Pretty less dependent on Symbol;
 wenzelm parents: 
46894diff
changeset | 205 | | indent n prt = blk (0, [str (spaces n), prt]); | 
| 23645 | 206 | |
| 207 | fun unbreakable (Break (_, wd)) = String (output_spaces wd, wd) | |
| 208 | | unbreakable (Block (m, es, indent, wd)) = Block (m, map unbreakable es, indent, wd) | |
| 209 | | unbreakable (e as String _) = e; | |
| 210 | ||
| 211 | ||
| 212 | ||
| 213 | (** formatting **) | |
| 214 | ||
| 215 | (* formatted output *) | |
| 216 | ||
| 217 | local | |
| 6116 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 wenzelm parents: diff
changeset | 218 | |
| 17756 | 219 | type text = {tx: Buffer.T, ind: Buffer.T, pos: int, nl: int};
 | 
| 220 | ||
| 221 | val empty: text = | |
| 10952 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 222 |  {tx = Buffer.empty,
 | 
| 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 223 | ind = Buffer.empty, | 
| 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 224 | pos = 0, | 
| 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 225 | nl = 0}; | 
| 6116 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 wenzelm parents: diff
changeset | 226 | |
| 32784 | 227 | 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: 
12421diff
changeset | 228 |  {tx = Buffer.add (Output.output "\n") tx,
 | 
| 10952 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 229 | ind = Buffer.empty, | 
| 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 230 | pos = 0, | 
| 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 231 | nl = nl + 1}; | 
| 6116 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 wenzelm parents: diff
changeset | 232 | |
| 23628 
41cdbfb9f77b
markup: emit as control information -- no indent text;
 wenzelm parents: 
23617diff
changeset | 233 | fun control s {tx, ind, pos: int, nl} : text =
 | 
| 
41cdbfb9f77b
markup: emit as control information -- no indent text;
 wenzelm parents: 
23617diff
changeset | 234 |  {tx = Buffer.add s tx,
 | 
| 
41cdbfb9f77b
markup: emit as control information -- no indent text;
 wenzelm parents: 
23617diff
changeset | 235 | ind = ind, | 
| 
41cdbfb9f77b
markup: emit as control information -- no indent text;
 wenzelm parents: 
23617diff
changeset | 236 | pos = pos, | 
| 
41cdbfb9f77b
markup: emit as control information -- no indent text;
 wenzelm parents: 
23617diff
changeset | 237 | nl = nl}; | 
| 
41cdbfb9f77b
markup: emit as control information -- no indent text;
 wenzelm parents: 
23617diff
changeset | 238 | |
| 17756 | 239 | fun string (s, len) {tx, ind, pos: int, nl} : text =
 | 
| 10952 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 240 |  {tx = Buffer.add s tx,
 | 
| 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 241 | ind = Buffer.add s ind, | 
| 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 242 | pos = pos + len, | 
| 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 243 | nl = nl}; | 
| 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 244 | |
| 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 245 | 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 | 246 | |
| 17756 | 247 | fun indentation (buf, len) {tx, ind, pos, nl} : text =
 | 
| 10952 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 248 | let val s = Buffer.content buf in | 
| 23617 | 249 |    {tx = Buffer.add (mode_indent s len) tx,
 | 
| 10952 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 250 | ind = Buffer.add s ind, | 
| 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 251 | pos = pos + len, | 
| 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 252 | nl = nl} | 
| 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 253 | end; | 
| 6116 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 wenzelm parents: diff
changeset | 254 | |
| 10952 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 255 | (*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: 
9730diff
changeset | 256 | include "after", to account for text following this block.*) | 
| 36687 | 257 | fun breakdist (Break _ :: _, _) = 0 | 
| 258 | | breakdist (Block (_, _, _, len) :: es, after) = len + breakdist (es, after) | |
| 32784 | 259 | | breakdist (String (_, len) :: es, after) = len + breakdist (es, after) | 
| 10952 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 260 | | breakdist ([], after) = after; | 
| 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 261 | |
| 
b520e4f1313b
support general indentation (e.g. for non-tt latex output);
 wenzelm parents: 
9730diff
changeset | 262 | (*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 | 263 | fun forcenext [] = [] | 
| 36690 
97d2780ad6f0
uniform treatment of length = 1 for forced breaks, also makes ML/Pretty.length coincide with Scala/XML.content_length;
 wenzelm parents: 
36689diff
changeset | 264 | | forcenext (Break _ :: es) = fbrk :: es | 
| 6116 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 wenzelm parents: diff
changeset | 265 | | forcenext (e :: es) = e :: forcenext es; | 
| 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 wenzelm parents: diff
changeset | 266 | |
| 23645 | 267 | in | 
| 19266 | 268 | |
| 36745 | 269 | fun formatted margin input = | 
| 270 | let | |
| 271 | val breakgain = margin div 20; (*minimum added space required of a break*) | |
| 272 | val emergencypos = margin div 2; (*position too far to right*) | |
| 273 | ||
| 274 | (*es is list of expressions to print; | |
| 275 | blockin is the indentation of the current block; | |
| 276 | after is the width of the following context until next break.*) | |
| 277 | fun format ([], _, _) text = text | |
| 278 |       | format (e :: es, block as (_, blockin), after) (text as {ind, pos, nl, ...}) =
 | |
| 279 | (case e of | |
| 280 | Block ((bg, en), bes, indent, _) => | |
| 281 | let | |
| 282 | val pos' = pos + indent; | |
| 283 | val pos'' = pos' mod emergencypos; | |
| 284 | val block' = | |
| 285 | if pos' < emergencypos then (ind |> add_indent indent, pos') | |
| 286 | else (add_indent pos'' Buffer.empty, pos''); | |
| 287 | val btext: text = text | |
| 288 | |> control bg | |
| 289 | |> format (bes, block', breakdist (es, after)) | |
| 290 | |> control en; | |
| 291 | (*if this block was broken then force the next break*) | |
| 292 | val es' = if nl < #nl btext then forcenext es else es; | |
| 293 | in format (es', block, after) btext end | |
| 294 | | Break (force, wd) => | |
| 295 | (*no break if text to next break fits on this line | |
| 296 | or if breaking would add only breakgain to space*) | |
| 297 | format (es, block, after) | |
| 298 | (if not force andalso | |
| 299 | pos + wd <= Int.max (margin - breakdist (es, after), blockin + breakgain) | |
| 300 | then text |> blanks wd (*just insert wd blanks*) | |
| 301 | else text |> newline |> indentation block) | |
| 302 | | String str => format (es, block, after) (string str text)); | |
| 303 | in | |
| 36747 
7361d5dde9ce
discontinued Pretty.setdepth, which appears to be largely unused, but can disrupt important markup if enabled accidentally;
 wenzelm parents: 
36745diff
changeset | 304 | #tx (format ([input], (Buffer.empty, 0), 0) empty) | 
| 36745 | 305 | end; | 
| 6116 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 wenzelm parents: diff
changeset | 306 | |
| 23645 | 307 | end; | 
| 14832 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 wenzelm parents: 
12421diff
changeset | 308 | |
| 6116 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 wenzelm parents: diff
changeset | 309 | |
| 23645 | 310 | (* special output *) | 
| 18802 | 311 | |
| 23645 | 312 | (*symbolic markup -- no formatting*) | 
| 313 | fun symbolic prt = | |
| 314 | let | |
| 30667 
53fbf7c679b0
Block markup: maintain output version within tree values (in accordance with String) -- changes operational behaviour wrt. print_mode;
 wenzelm parents: 
30624diff
changeset | 315 | 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: 
30624diff
changeset | 316 | | out (Block ((bg, en), prts, indent, _)) = | 
| 45666 | 317 | Buffer.add bg #> | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
50162diff
changeset | 318 | Buffer.markup (Markup.block indent) (fold out prts) #> | 
| 45666 | 319 | Buffer.add en | 
| 23645 | 320 | | out (String (s, _)) = Buffer.add s | 
| 45666 | 321 | | out (Break (false, wd)) = | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
50162diff
changeset | 322 | Buffer.markup (Markup.break wd) (Buffer.add (output_spaces wd)) | 
| 36689 
379f5b1e7f91
replaced slightly odd fbreak markup by plain "\n", which also coincides with regular linebreaks produced outside the ML pretty engine;
 wenzelm parents: 
36687diff
changeset | 323 | | out (Break (true, _)) = Buffer.add (Output.output "\n"); | 
| 23645 | 324 | 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 | 325 | |
| 23645 | 326 | (*unformatted output*) | 
| 327 | fun unformatted prt = | |
| 6116 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 wenzelm parents: diff
changeset | 328 | let | 
| 30667 
53fbf7c679b0
Block markup: maintain output version within tree values (in accordance with String) -- changes operational behaviour wrt. print_mode;
 wenzelm parents: 
30624diff
changeset | 329 | fun fmt (Block ((bg, en), prts, _, _)) = Buffer.add bg #> fold fmt prts #> Buffer.add en | 
| 23617 | 330 | | fmt (String (s, _)) = Buffer.add s | 
| 36690 
97d2780ad6f0
uniform treatment of length = 1 for forced breaks, also makes ML/Pretty.length coincide with Scala/XML.content_length;
 wenzelm parents: 
36689diff
changeset | 331 | | fmt (Break (_, wd)) = Buffer.add (output_spaces wd); | 
| 36747 
7361d5dde9ce
discontinued Pretty.setdepth, which appears to be largely unused, but can disrupt important markup if enabled accidentally;
 wenzelm parents: 
36745diff
changeset | 332 | 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 | 333 | |
| 30624 
e755b8b76365
simplified datatype ML_Pretty.pretty: model Isabelle not Poly/ML;
 wenzelm parents: 
30620diff
changeset | 334 | |
| 36748 | 335 | (* output interfaces *) | 
| 30620 | 336 | |
| 36748 | 337 | val margin_default = Unsynchronized.ref 76; (*right margin, or page width*) | 
| 23645 | 338 | |
| 339 | val symbolicN = "pretty_symbolic"; | |
| 340 | ||
| 36745 | 341 | fun output_buffer margin prt = | 
| 23645 | 342 | if print_mode_active symbolicN then symbolic prt | 
| 36745 | 343 | else formatted (the_default (! margin_default) margin) prt; | 
| 23645 | 344 | |
| 36745 | 345 | val output = Buffer.content oo output_buffer; | 
| 346 | fun string_of_margin margin = Output.escape o output (SOME margin); | |
| 347 | val string_of = Output.escape o output NONE; | |
| 49656 
7ff712de5747
treat wrapped markup elements as raw markup delimiters;
 wenzelm parents: 
49565diff
changeset | 348 | val writeln = Output.writeln o string_of; | 
| 
7ff712de5747
treat wrapped markup elements as raw markup delimiters;
 wenzelm parents: 
49565diff
changeset | 349 | |
| 
7ff712de5747
treat wrapped markup elements as raw markup delimiters;
 wenzelm parents: 
49565diff
changeset | 350 | val symbolic_output = Buffer.content o symbolic; | 
| 
7ff712de5747
treat wrapped markup elements as raw markup delimiters;
 wenzelm parents: 
49565diff
changeset | 351 | val symbolic_string_of = Output.escape o symbolic_output; | 
| 
7ff712de5747
treat wrapped markup elements as raw markup delimiters;
 wenzelm parents: 
49565diff
changeset | 352 | |
| 23645 | 353 | 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 | 354 | |
| 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 wenzelm parents: diff
changeset | 355 | |
| 14832 
6589a58f57cb
pp: abstract pretty printing context; string_of/str_of: mark result as raw output; added Pretty.unbreakable;
 wenzelm parents: 
12421diff
changeset | 356 | |
| 36748 | 357 | (** ML toplevel pretty printing **) | 
| 358 | ||
| 359 | fun to_ML (Block (m, prts, ind, _)) = ML_Pretty.Block (m, map to_ML prts, ind) | |
| 360 | | to_ML (String s) = ML_Pretty.String s | |
| 361 | | to_ML (Break b) = ML_Pretty.Break b; | |
| 362 | ||
| 49656 
7ff712de5747
treat wrapped markup elements as raw markup delimiters;
 wenzelm parents: 
49565diff
changeset | 363 | fun from_ML (ML_Pretty.Block (m, prts, ind)) = raw_markup m (ind, map from_ML prts) | 
| 36748 | 364 | | from_ML (ML_Pretty.String s) = String s | 
| 365 | | from_ML (ML_Pretty.Break b) = Break b; | |
| 366 | ||
| 37529 | 367 | end; | 
| 368 | ||
| 6116 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 wenzelm parents: diff
changeset | 369 | end; |