| author | wenzelm | 
| Sat, 10 Oct 2020 21:12:20 +0200 | |
| changeset 72425 | d0937d55eb90 | 
| parent 71781 | 3fd54f7f52b0 | 
| child 73359 | d8a0e996614b | 
| permissions | -rw-r--r-- | 
| 36683 | 1  | 
/* Title: Pure/General/pretty.scala  | 
2  | 
Author: Makarius  | 
|
3  | 
||
| 36687 | 4  | 
Generic pretty printing module.  | 
| 36683 | 5  | 
*/  | 
6  | 
||
7  | 
package isabelle  | 
|
8  | 
||
| 71601 | 9  | 
import scala.annotation.tailrec  | 
10  | 
||
| 36683 | 11  | 
|
12  | 
object Pretty  | 
|
13  | 
{
 | 
|
| 61871 | 14  | 
/* XML constructors */  | 
15  | 
||
| 61874 | 16  | 
val space: XML.Body = List(XML.Text(Symbol.space))  | 
17  | 
def spaces(n: Int): XML.Body =  | 
|
18  | 
if (n == 0) Nil  | 
|
19  | 
else if (n == 1) space  | 
|
20  | 
else List(XML.Text(Symbol.spaces(n)))  | 
|
| 
48704
 
85a3de10567d
tuned signature -- make Pretty less dependent on Symbol;
 
wenzelm 
parents: 
45666 
diff
changeset
 | 
21  | 
|
| 61871 | 22  | 
def block(consistent: Boolean, indent: Int, body: XML.Body): XML.Tree =  | 
23  | 
XML.Elem(Markup.Block(consistent, indent), body)  | 
|
24  | 
def block(indent: Int, body: XML.Body): XML.Tree = block(false, indent, body)  | 
|
25  | 
def block(body: XML.Body): XML.Tree = block(2, body)  | 
|
| 
48704
 
85a3de10567d
tuned signature -- make Pretty less dependent on Symbol;
 
wenzelm 
parents: 
45666 
diff
changeset
 | 
26  | 
|
| 61871 | 27  | 
def brk(width: Int, indent: Int = 0): XML.Tree =  | 
28  | 
XML.Elem(Markup.Break(width, indent), spaces(width))  | 
|
29  | 
||
| 69867 | 30  | 
val fbrk: XML.Tree = XML.newline  | 
| 65130 | 31  | 
def fbreaks(ts: List[XML.Tree]): XML.Body = Library.separate(fbrk, ts)  | 
| 61871 | 32  | 
|
33  | 
val Separator: XML.Body = List(XML.elem(Markup.SEPARATOR, space), fbrk)  | 
|
34  | 
def separate(ts: List[XML.Tree]): XML.Body = Library.separate(Separator, ts.map(List(_))).flatten  | 
|
| 61864 | 35  | 
|
| 
48704
 
85a3de10567d
tuned signature -- make Pretty less dependent on Symbol;
 
wenzelm 
parents: 
45666 
diff
changeset
 | 
36  | 
|
| 
51492
 
eaa1c4cc1106
more explicit Pretty.Metric, with clear distinction of unit (space width) vs. average char width (for visual adjustments) -- NB: Pretty formatting works via full space characters (despite a981a5c8a505 and 70f7483df9cb);
 
wenzelm 
parents: 
51470 
diff
changeset
 | 
37  | 
/* text metric -- standardized to width of space */  | 
| 
 
eaa1c4cc1106
more explicit Pretty.Metric, with clear distinction of unit (space width) vs. average char width (for visual adjustments) -- NB: Pretty formatting works via full space characters (despite a981a5c8a505 and 70f7483df9cb);
 
wenzelm 
parents: 
51470 
diff
changeset
 | 
38  | 
|
| 
 
eaa1c4cc1106
more explicit Pretty.Metric, with clear distinction of unit (space width) vs. average char width (for visual adjustments) -- NB: Pretty formatting works via full space characters (despite a981a5c8a505 and 70f7483df9cb);
 
wenzelm 
parents: 
51470 
diff
changeset
 | 
39  | 
abstract class Metric  | 
| 
 
eaa1c4cc1106
more explicit Pretty.Metric, with clear distinction of unit (space width) vs. average char width (for visual adjustments) -- NB: Pretty formatting works via full space characters (despite a981a5c8a505 and 70f7483df9cb);
 
wenzelm 
parents: 
51470 
diff
changeset
 | 
40  | 
  {
 | 
| 
 
eaa1c4cc1106
more explicit Pretty.Metric, with clear distinction of unit (space width) vs. average char width (for visual adjustments) -- NB: Pretty formatting works via full space characters (despite a981a5c8a505 and 70f7483df9cb);
 
wenzelm 
parents: 
51470 
diff
changeset
 | 
41  | 
val unit: Double  | 
| 
 
eaa1c4cc1106
more explicit Pretty.Metric, with clear distinction of unit (space width) vs. average char width (for visual adjustments) -- NB: Pretty formatting works via full space characters (despite a981a5c8a505 and 70f7483df9cb);
 
wenzelm 
parents: 
51470 
diff
changeset
 | 
42  | 
def apply(s: String): Double  | 
| 
 
eaa1c4cc1106
more explicit Pretty.Metric, with clear distinction of unit (space width) vs. average char width (for visual adjustments) -- NB: Pretty formatting works via full space characters (despite a981a5c8a505 and 70f7483df9cb);
 
wenzelm 
parents: 
51470 
diff
changeset
 | 
43  | 
}  | 
| 
 
eaa1c4cc1106
more explicit Pretty.Metric, with clear distinction of unit (space width) vs. average char width (for visual adjustments) -- NB: Pretty formatting works via full space characters (despite a981a5c8a505 and 70f7483df9cb);
 
wenzelm 
parents: 
51470 
diff
changeset
 | 
44  | 
|
| 67896 | 45  | 
object Default_Metric extends Metric  | 
| 
51492
 
eaa1c4cc1106
more explicit Pretty.Metric, with clear distinction of unit (space width) vs. average char width (for visual adjustments) -- NB: Pretty formatting works via full space characters (despite a981a5c8a505 and 70f7483df9cb);
 
wenzelm 
parents: 
51470 
diff
changeset
 | 
46  | 
  {
 | 
| 
 
eaa1c4cc1106
more explicit Pretty.Metric, with clear distinction of unit (space width) vs. average char width (for visual adjustments) -- NB: Pretty formatting works via full space characters (despite a981a5c8a505 and 70f7483df9cb);
 
wenzelm 
parents: 
51470 
diff
changeset
 | 
47  | 
val unit = 1.0  | 
| 
 
eaa1c4cc1106
more explicit Pretty.Metric, with clear distinction of unit (space width) vs. average char width (for visual adjustments) -- NB: Pretty formatting works via full space characters (despite a981a5c8a505 and 70f7483df9cb);
 
wenzelm 
parents: 
51470 
diff
changeset
 | 
48  | 
def apply(s: String): Double = s.length.toDouble  | 
| 
 
eaa1c4cc1106
more explicit Pretty.Metric, with clear distinction of unit (space width) vs. average char width (for visual adjustments) -- NB: Pretty formatting works via full space characters (despite a981a5c8a505 and 70f7483df9cb);
 
wenzelm 
parents: 
51470 
diff
changeset
 | 
49  | 
}  | 
| 
 
eaa1c4cc1106
more explicit Pretty.Metric, with clear distinction of unit (space width) vs. average char width (for visual adjustments) -- NB: Pretty formatting works via full space characters (despite a981a5c8a505 and 70f7483df9cb);
 
wenzelm 
parents: 
51470 
diff
changeset
 | 
50  | 
|
| 
 
eaa1c4cc1106
more explicit Pretty.Metric, with clear distinction of unit (space width) vs. average char width (for visual adjustments) -- NB: Pretty formatting works via full space characters (despite a981a5c8a505 and 70f7483df9cb);
 
wenzelm 
parents: 
51470 
diff
changeset
 | 
51  | 
|
| 36687 | 52  | 
/* markup trees with physical blocks and breaks */  | 
53  | 
||
| 62820 | 54  | 
private def force_nat(i: Int): Int = i max 0  | 
55  | 
||
| 61874 | 56  | 
  private sealed abstract class Tree { def length: Double }
 | 
57  | 
private case class Block(  | 
|
| 61871 | 58  | 
markup: Option[(Markup, Option[XML.Body])],  | 
59  | 
consistent: Boolean, indent: Int, body: List[Tree], length: Double) extends Tree  | 
|
| 61874 | 60  | 
private case class Break(force: Boolean, width: Int, indent: Int) extends Tree  | 
| 61871 | 61  | 
  { def length: Double = width.toDouble }
 | 
| 61874 | 62  | 
private case class Str(string: String, length: Double) extends Tree  | 
| 36683 | 63  | 
|
| 61874 | 64  | 
private val FBreak = Break(true, 1, 0)  | 
| 36683 | 65  | 
|
| 61874 | 66  | 
private def make_block(  | 
| 61871 | 67  | 
markup: Option[(Markup, Option[XML.Body])],  | 
68  | 
consistent: Boolean,  | 
|
69  | 
indent: Int,  | 
|
70  | 
body: List[Tree]): Tree =  | 
|
| 
61883
 
c0f34fe6aa61
clarified length of block with pre-existant forced breaks;
 
wenzelm 
parents: 
61875 
diff
changeset
 | 
71  | 
  {
 | 
| 
62785
 
70b9c7d4ed7f
more robust pretty printing: permissive treatment of bad values;
 
wenzelm 
parents: 
61883 
diff
changeset
 | 
72  | 
val indent1 = force_nat(indent)  | 
| 
 
70b9c7d4ed7f
more robust pretty printing: permissive treatment of bad values;
 
wenzelm 
parents: 
61883 
diff
changeset
 | 
73  | 
|
| 71601 | 74  | 
@tailrec def body_length(prts: List[Tree], len: Double): Double =  | 
| 
61883
 
c0f34fe6aa61
clarified length of block with pre-existant forced breaks;
 
wenzelm 
parents: 
61875 
diff
changeset
 | 
75  | 
    {
 | 
| 
 
c0f34fe6aa61
clarified length of block with pre-existant forced breaks;
 
wenzelm 
parents: 
61875 
diff
changeset
 | 
76  | 
val (line, rest) =  | 
| 
 
c0f34fe6aa61
clarified length of block with pre-existant forced breaks;
 
wenzelm 
parents: 
61875 
diff
changeset
 | 
77  | 
        Library.take_prefix[Tree]({ case Break(true, _, _) => false case _ => true }, prts)
 | 
| 
 
c0f34fe6aa61
clarified length of block with pre-existant forced breaks;
 
wenzelm 
parents: 
61875 
diff
changeset
 | 
78  | 
      val len1 = ((0.0 /: line) { case (l, t) => l + t.length }) max len
 | 
| 71781 | 79  | 
      (rest: @unchecked) match {
 | 
| 
61883
 
c0f34fe6aa61
clarified length of block with pre-existant forced breaks;
 
wenzelm 
parents: 
61875 
diff
changeset
 | 
80  | 
case Break(true, _, ind) :: rest1 =>  | 
| 
62785
 
70b9c7d4ed7f
more robust pretty printing: permissive treatment of bad values;
 
wenzelm 
parents: 
61883 
diff
changeset
 | 
81  | 
body_length(Break(false, indent1 + ind, 0) :: rest1, len1)  | 
| 
61883
 
c0f34fe6aa61
clarified length of block with pre-existant forced breaks;
 
wenzelm 
parents: 
61875 
diff
changeset
 | 
82  | 
case Nil => len1  | 
| 
 
c0f34fe6aa61
clarified length of block with pre-existant forced breaks;
 
wenzelm 
parents: 
61875 
diff
changeset
 | 
83  | 
}  | 
| 
 
c0f34fe6aa61
clarified length of block with pre-existant forced breaks;
 
wenzelm 
parents: 
61875 
diff
changeset
 | 
84  | 
}  | 
| 
62785
 
70b9c7d4ed7f
more robust pretty printing: permissive treatment of bad values;
 
wenzelm 
parents: 
61883 
diff
changeset
 | 
85  | 
Block(markup, consistent, indent1, body, body_length(body, 0.0))  | 
| 
61883
 
c0f34fe6aa61
clarified length of block with pre-existant forced breaks;
 
wenzelm 
parents: 
61875 
diff
changeset
 | 
86  | 
}  | 
| 36683 | 87  | 
|
| 61874 | 88  | 
|
89  | 
/* formatted output */  | 
|
90  | 
||
91  | 
private sealed case class Text(tx: XML.Body = Nil, pos: Double = 0.0, nl: Int = 0)  | 
|
92  | 
  {
 | 
|
93  | 
def newline: Text = copy(tx = fbrk :: tx, pos = 0.0, nl = nl + 1)  | 
|
94  | 
def string(s: String, len: Double): Text =  | 
|
95  | 
copy(tx = if (s == "") tx else XML.Text(s) :: tx, pos = pos + len)  | 
|
96  | 
def blanks(wd: Int): Text = string(Symbol.spaces(wd), wd.toDouble)  | 
|
97  | 
def content: XML.Body = tx.reverse  | 
|
98  | 
}  | 
|
99  | 
||
100  | 
private def break_dist(trees: List[Tree], after: Double): Double =  | 
|
101  | 
    trees match {
 | 
|
102  | 
case (_: Break) :: _ => 0.0  | 
|
103  | 
case t :: ts => t.length + break_dist(ts, after)  | 
|
104  | 
case Nil => after  | 
|
| 
36689
 
379f5b1e7f91
replaced slightly odd fbreak markup by plain "\n", which also coincides with regular linebreaks produced outside the ML pretty engine;
 
wenzelm 
parents: 
36687 
diff
changeset
 | 
105  | 
}  | 
| 
 
379f5b1e7f91
replaced slightly odd fbreak markup by plain "\n", which also coincides with regular linebreaks produced outside the ML pretty engine;
 
wenzelm 
parents: 
36687 
diff
changeset
 | 
106  | 
|
| 61874 | 107  | 
private def force_break(tree: Tree): Tree =  | 
108  | 
    tree match { case Break(false, wd, ind) => Break(true, wd, ind) case _ => tree }
 | 
|
| 71601 | 109  | 
private def force_all(trees: List[Tree]): List[Tree] = trees.map(force_break)  | 
| 
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: 
51569 
diff
changeset
 | 
110  | 
|
| 61874 | 111  | 
private def force_next(trees: List[Tree]): List[Tree] =  | 
112  | 
    trees match {
 | 
|
113  | 
case Nil => Nil  | 
|
114  | 
case (t: Break) :: ts => force_break(t) :: ts  | 
|
115  | 
case t :: ts => t :: force_next(ts)  | 
|
116  | 
}  | 
|
| 
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: 
51569 
diff
changeset
 | 
117  | 
|
| 71601 | 118  | 
val default_margin: Double = 76.0  | 
119  | 
val default_breakgain: Double = default_margin / 20  | 
|
| 36820 | 120  | 
|
| 
67547
 
aefe7a7b330a
clarified breakgain: keeping it constant avoids margin fluctuation in Pretty_Tooltip vs. Pretty_Text_Area;
 
wenzelm 
parents: 
65130 
diff
changeset
 | 
121  | 
def formatted(input: XML.Body,  | 
| 67896 | 122  | 
margin: Double = default_margin,  | 
123  | 
breakgain: Double = default_breakgain,  | 
|
124  | 
metric: Metric = Default_Metric): XML.Body =  | 
|
| 36687 | 125  | 
  {
 | 
| 
51569
 
4e084727faae
maintain integer indentation during formatting -- it needs to be implemented by repeated spaces eventually;
 
wenzelm 
parents: 
51568 
diff
changeset
 | 
126  | 
val emergencypos = (margin / 2).round.toInt  | 
| 36687 | 127  | 
|
| 61874 | 128  | 
def make_tree(inp: XML.Body): List[Tree] =  | 
129  | 
      inp flatMap {
 | 
|
130  | 
case XML.Wrapped_Elem(markup, body1, body2) =>  | 
|
131  | 
List(make_block(Some(markup, Some(body1)), false, 0, make_tree(body2)))  | 
|
132  | 
case XML.Elem(markup, body) =>  | 
|
133  | 
          markup match {
 | 
|
134  | 
case Markup.Block(consistent, indent) =>  | 
|
135  | 
List(make_block(None, consistent, indent, make_tree(body)))  | 
|
136  | 
case Markup.Break(width, indent) =>  | 
|
| 
62785
 
70b9c7d4ed7f
more robust pretty printing: permissive treatment of bad values;
 
wenzelm 
parents: 
61883 
diff
changeset
 | 
137  | 
List(Break(false, force_nat(width), force_nat(indent)))  | 
| 61874 | 138  | 
case Markup(Markup.ITEM, _) =>  | 
139  | 
List(make_block(None, false, 2,  | 
|
140  | 
make_tree(XML.elem(Markup.BULLET, space) :: space ::: body)))  | 
|
141  | 
case _ =>  | 
|
142  | 
List(make_block(Some((markup, None)), false, 0, make_tree(body)))  | 
|
143  | 
}  | 
|
144  | 
case XML.Text(text) =>  | 
|
145  | 
Library.separate(FBreak, split_lines(text).map(s => Str(s, metric(s))))  | 
|
| 
36817
 
ed97e877ff2d
more precise pretty printing based on actual font metrics;
 
wenzelm 
parents: 
36763 
diff
changeset
 | 
146  | 
}  | 
| 
 
ed97e877ff2d
more precise pretty printing based on actual font metrics;
 
wenzelm 
parents: 
36763 
diff
changeset
 | 
147  | 
|
| 61871 | 148  | 
def format(trees: List[Tree], blockin: Int, after: Double, text: Text): Text =  | 
| 36687 | 149  | 
      trees match {
 | 
150  | 
case Nil => text  | 
|
151  | 
||
| 61871 | 152  | 
case Block(markup, consistent, indent, body, blen) :: ts =>  | 
| 
51569
 
4e084727faae
maintain integer indentation during formatting -- it needs to be implemented by repeated spaces eventually;
 
wenzelm 
parents: 
51568 
diff
changeset
 | 
153  | 
val pos1 = (text.pos + indent).ceil.toInt  | 
| 36687 | 154  | 
val pos2 = pos1 % emergencypos  | 
| 61868 | 155  | 
val blockin1 = if (pos1 < emergencypos) pos1 else pos2  | 
| 61864 | 156  | 
val d = break_dist(ts, after)  | 
157  | 
val body1 = if (consistent && text.pos + blen > margin - d) force_all(body) else body  | 
|
| 61871 | 158  | 
val btext =  | 
159  | 
            markup match {
 | 
|
160  | 
case None => format(body1, blockin1, d, text)  | 
|
161  | 
case Some((m, markup_body)) =>  | 
|
162  | 
val btext0 = format(body1, blockin1, d, text.copy(tx = Nil))  | 
|
163  | 
val elem =  | 
|
164  | 
                  markup_body match {
 | 
|
165  | 
case None => XML.Elem(m, btext0.content)  | 
|
166  | 
case Some(b) => XML.Wrapped_Elem(m, b, btext0.content)  | 
|
167  | 
}  | 
|
168  | 
btext0.copy(tx = elem :: text.tx)  | 
|
169  | 
}  | 
|
| 61864 | 170  | 
val ts1 = if (text.nl < btext.nl) force_next(ts) else ts  | 
| 36687 | 171  | 
format(ts1, blockin, after, btext)  | 
172  | 
||
| 61871 | 173  | 
case Break(force, wd, ind) :: ts =>  | 
174  | 
if (!force &&  | 
|
175  | 
text.pos + wd <= ((margin - break_dist(ts, after)) max (blockin + breakgain)))  | 
|
| 36687 | 176  | 
format(ts, blockin, after, text.blanks(wd))  | 
| 
61862
 
e2a9e46ac0fb
support pretty break indent, like underlying ML systems;
 
wenzelm 
parents: 
55551 
diff
changeset
 | 
177  | 
else format(ts, blockin, after, text.newline.blanks(blockin + ind))  | 
| 36687 | 178  | 
|
| 61871 | 179  | 
case Str(s, len) :: ts => format(ts, blockin, after, text.string(s, len))  | 
| 36687 | 180  | 
}  | 
| 61874 | 181  | 
format(make_tree(input), 0, 0.0, Text()).content  | 
| 36687 | 182  | 
}  | 
| 
36734
 
d9b10c173330
Pretty.formatted operates directly on XML trees, treating XML.Elem like a pro-forma block of indentation 0, like the ML version;
 
wenzelm 
parents: 
36689 
diff
changeset
 | 
183  | 
|
| 
67547
 
aefe7a7b330a
clarified breakgain: keeping it constant avoids margin fluctuation in Pretty_Tooltip vs. Pretty_Text_Area;
 
wenzelm 
parents: 
65130 
diff
changeset
 | 
184  | 
def string_of(input: XML.Body,  | 
| 67896 | 185  | 
margin: Double = default_margin,  | 
186  | 
breakgain: Double = default_breakgain,  | 
|
187  | 
metric: Metric = Default_Metric): String =  | 
|
| 
67547
 
aefe7a7b330a
clarified breakgain: keeping it constant avoids margin fluctuation in Pretty_Tooltip vs. Pretty_Text_Area;
 
wenzelm 
parents: 
65130 
diff
changeset
 | 
188  | 
XML.content(formatted(input, margin = margin, breakgain = breakgain, metric = metric))  | 
| 36683 | 189  | 
}  |