| author | boehmes | 
| Tue, 07 Jun 2011 10:24:16 +0200 | |
| changeset 43230 | dabf6e311213 | 
| parent 39178 | 83e9f3ccea9f | 
| child 43520 | cec9b95fa35d | 
| permissions | -rw-r--r-- | 
| 38479 
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
 wenzelm parents: 
38478diff
changeset | 1 | /* Title: Pure/PIDE/markup_tree.scala | 
| 36676 | 2 | Author: Fabian Immler, TU Munich | 
| 3 | Author: Makarius | |
| 4 | ||
| 38479 
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
 wenzelm parents: 
38478diff
changeset | 5 | Markup trees over nested / non-overlapping text ranges. | 
| 36676 | 6 | */ | 
| 34393 | 7 | |
| 34871 
e596a0b71f3c
incorporate "proofdocument" part into main Isabelle/Pure.jar -- except for html_panel.scala, which depends on external library (Lobo/Cobra browser);
 wenzelm parents: 
34760diff
changeset | 8 | package isabelle | 
| 34393 | 9 | |
| 34701 
80b0add08eef
IsabelleSideKickParser: incorporate former MarkupNode.markup2default_node, observe stopped flag;
 wenzelm parents: 
34676diff
changeset | 10 | |
| 
80b0add08eef
IsabelleSideKickParser: incorporate former MarkupNode.markup2default_node, observe stopped flag;
 wenzelm parents: 
34676diff
changeset | 11 | import javax.swing.tree.DefaultMutableTreeNode | 
| 
80b0add08eef
IsabelleSideKickParser: incorporate former MarkupNode.markup2default_node, observe stopped flag;
 wenzelm parents: 
34676diff
changeset | 12 | |
| 38479 
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
 wenzelm parents: 
38478diff
changeset | 13 | import scala.collection.immutable.SortedMap | 
| 
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
 wenzelm parents: 
38478diff
changeset | 14 | |
| 
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
 wenzelm parents: 
38478diff
changeset | 15 | |
| 
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
 wenzelm parents: 
38478diff
changeset | 16 | object Markup_Tree | 
| 
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
 wenzelm parents: 
38478diff
changeset | 17 | {
 | 
| 38566 
8176107637ce
Branches.overlapping: proper treatment of stop_range that overlaps with end;
 wenzelm parents: 
38564diff
changeset | 18 | /* branches sorted by quasi-order -- overlapping ranges appear as equivalent */ | 
| 38479 
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
 wenzelm parents: 
38478diff
changeset | 19 | |
| 
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
 wenzelm parents: 
38478diff
changeset | 20 | object Branches | 
| 
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
 wenzelm parents: 
38478diff
changeset | 21 |   {
 | 
| 38577 | 22 | type Entry = (Text.Info[Any], Markup_Tree) | 
| 38578 | 23 | type T = SortedMap[Text.Range, Entry] | 
| 38479 
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
 wenzelm parents: 
38478diff
changeset | 24 | |
| 38578 | 25 | val empty = SortedMap.empty[Text.Range, Entry](new scala.math.Ordering[Text.Range] | 
| 38479 
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
 wenzelm parents: 
38478diff
changeset | 26 |       {
 | 
| 38578 | 27 | def compare(r1: Text.Range, r2: Text.Range): Int = r1 compare r2 | 
| 38479 
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
 wenzelm parents: 
38478diff
changeset | 28 | }) | 
| 38562 | 29 | |
| 38578 | 30 | def update(branches: T, entry: Entry): T = branches + (entry._1.range -> entry) | 
| 38657 | 31 | def single(entry: Entry): T = update(empty, entry) | 
| 38562 | 32 | |
| 38659 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 33 | def overlapping(range: Text.Range, branches: T): T = // FIXME special cases!? | 
| 38566 
8176107637ce
Branches.overlapping: proper treatment of stop_range that overlaps with end;
 wenzelm parents: 
38564diff
changeset | 34 |     {
 | 
| 38578 | 35 | val start = Text.Range(range.start) | 
| 36 | val stop = Text.Range(range.stop) | |
| 38659 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 37 | val bs = branches.range(start, stop) | 
| 38566 
8176107637ce
Branches.overlapping: proper treatment of stop_range that overlaps with end;
 wenzelm parents: 
38564diff
changeset | 38 |       branches.get(stop) match {
 | 
| 38659 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 39 | case Some(end) if range overlaps end._1.range => update(bs, end) | 
| 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 40 | case _ => bs | 
| 38566 
8176107637ce
Branches.overlapping: proper treatment of stop_range that overlaps with end;
 wenzelm parents: 
38564diff
changeset | 41 | } | 
| 
8176107637ce
Branches.overlapping: proper treatment of stop_range that overlaps with end;
 wenzelm parents: 
38564diff
changeset | 42 | } | 
| 38479 
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
 wenzelm parents: 
38478diff
changeset | 43 | } | 
| 
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
 wenzelm parents: 
38478diff
changeset | 44 | |
| 
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
 wenzelm parents: 
38478diff
changeset | 45 | val empty = new Markup_Tree(Branches.empty) | 
| 39178 
83e9f3ccea9f
concentrate Isabelle specific physical rendering markup selection in isabelle_markup.scala;
 wenzelm parents: 
39177diff
changeset | 46 | |
| 
83e9f3ccea9f
concentrate Isabelle specific physical rendering markup selection in isabelle_markup.scala;
 wenzelm parents: 
39177diff
changeset | 47 | type Select[A] = PartialFunction[Text.Info[Any], A] | 
| 38479 
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
 wenzelm parents: 
38478diff
changeset | 48 | } | 
| 34554 
7dc6c231da40
abs. stops, markup nodes depend on doc-version;
 immler@in.tum.de parents: 
34517diff
changeset | 49 | |
| 34393 | 50 | |
| 38479 
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
 wenzelm parents: 
38478diff
changeset | 51 | case class Markup_Tree(val branches: Markup_Tree.Branches.T) | 
| 34717 
3f32e08bbb6c
sidekick root data: set buffer length to avoid crash of initial caret move;
 wenzelm parents: 
34708diff
changeset | 52 | {
 | 
| 38479 
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
 wenzelm parents: 
38478diff
changeset | 53 | import Markup_Tree._ | 
| 34557 | 54 | |
| 38563 | 55 | override def toString = | 
| 56 |     branches.toList.map(_._2) match {
 | |
| 57 | case Nil => "Empty" | |
| 58 |       case list => list.mkString("Tree(", ",", ")")
 | |
| 59 | } | |
| 60 | ||
| 38577 | 61 | def + (new_info: Text.Info[Any]): Markup_Tree = | 
| 34703 | 62 |   {
 | 
| 38578 | 63 | val new_range = new_info.range | 
| 64 |     branches.get(new_range) match {
 | |
| 38482 
7b6ee937b75f
tuned Markup_Tree, using SortedMap more carefully;
 wenzelm parents: 
38479diff
changeset | 65 | case None => | 
| 38577 | 66 | new Markup_Tree(Branches.update(branches, new_info -> empty)) | 
| 67 | case Some((info, subtree)) => | |
| 38578 | 68 | val range = info.range | 
| 38661 
f1ba2ae8e58a
Markup_Tree.+: new info tends to sink to bottom, where it is prefered by select;
 wenzelm parents: 
38659diff
changeset | 69 | if (range.contains(new_range)) | 
| 38577 | 70 | new Markup_Tree(Branches.update(branches, info -> (subtree + new_info))) | 
| 38578 | 71 | else if (new_range.contains(branches.head._1) && new_range.contains(branches.last._1)) | 
| 38657 | 72 | new Markup_Tree(Branches.single(new_info -> this)) | 
| 38482 
7b6ee937b75f
tuned Markup_Tree, using SortedMap more carefully;
 wenzelm parents: 
38479diff
changeset | 73 |         else {
 | 
| 38578 | 74 | val body = Branches.overlapping(new_range, branches) | 
| 75 |           if (body.forall(e => new_range.contains(e._1))) {
 | |
| 38659 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 76 | val rest = // branches -- body, modulo workarounds for Redblack in Scala 2.8.0 | 
| 38657 | 77 | if (body.size > 1) | 
| 78 | (Branches.empty /: branches)((rest, entry) => | |
| 79 | if (body.isDefinedAt(entry._1)) rest else rest + entry) | |
| 80 | else branches | |
| 38577 | 81 | new Markup_Tree(Branches.update(rest, new_info -> new Markup_Tree(body))) | 
| 38482 
7b6ee937b75f
tuned Markup_Tree, using SortedMap more carefully;
 wenzelm parents: 
38479diff
changeset | 82 | } | 
| 
7b6ee937b75f
tuned Markup_Tree, using SortedMap more carefully;
 wenzelm parents: 
38479diff
changeset | 83 |           else { // FIXME split markup!?
 | 
| 38577 | 84 |             System.err.println("Ignored overlapping markup information: " + new_info)
 | 
| 38482 
7b6ee937b75f
tuned Markup_Tree, using SortedMap more carefully;
 wenzelm parents: 
38479diff
changeset | 85 | this | 
| 
7b6ee937b75f
tuned Markup_Tree, using SortedMap more carefully;
 wenzelm parents: 
38479diff
changeset | 86 | } | 
| 
7b6ee937b75f
tuned Markup_Tree, using SortedMap more carefully;
 wenzelm parents: 
38479diff
changeset | 87 | } | 
| 34703 | 88 | } | 
| 89 | } | |
| 90 | ||
| 38659 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 91 | private def overlapping(range: Text.Range): Stream[(Text.Range, Branches.Entry)] = | 
| 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 92 | Branches.overlapping(range, branches).toStream | 
| 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 93 | |
| 39178 
83e9f3ccea9f
concentrate Isabelle specific physical rendering markup selection in isabelle_markup.scala;
 wenzelm parents: 
39177diff
changeset | 94 | def select[A](root_range: Text.Range)(result: Markup_Tree.Select[A]) | 
| 39177 
0468964aec11
simplified Markup_Tree.select: Stream instead of Iterator (again), explicit Option instead of default;
 wenzelm parents: 
38845diff
changeset | 95 | : Stream[Text.Info[Option[A]]] = | 
| 38486 
f5bbfc019937
Markup_Tree.select: crude version of stream-based filtering;
 wenzelm parents: 
38482diff
changeset | 96 |   {
 | 
| 39177 
0468964aec11
simplified Markup_Tree.select: Stream instead of Iterator (again), explicit Option instead of default;
 wenzelm parents: 
38845diff
changeset | 97 | def stream( | 
| 
0468964aec11
simplified Markup_Tree.select: Stream instead of Iterator (again), explicit Option instead of default;
 wenzelm parents: 
38845diff
changeset | 98 | last: Text.Offset, | 
| 
0468964aec11
simplified Markup_Tree.select: Stream instead of Iterator (again), explicit Option instead of default;
 wenzelm parents: 
38845diff
changeset | 99 | stack: List[(Text.Info[Option[A]], Stream[(Text.Range, Branches.Entry)])]) | 
| 
0468964aec11
simplified Markup_Tree.select: Stream instead of Iterator (again), explicit Option instead of default;
 wenzelm parents: 
38845diff
changeset | 100 | : Stream[Text.Info[Option[A]]] = | 
| 38486 
f5bbfc019937
Markup_Tree.select: crude version of stream-based filtering;
 wenzelm parents: 
38482diff
changeset | 101 |     {
 | 
| 38659 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 102 |       stack match {
 | 
| 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 103 | case (parent, (range, (info, tree)) #:: more) :: rest => | 
| 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 104 | val subrange = range.restrict(root_range) | 
| 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 105 | val subtree = tree.overlapping(subrange) | 
| 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 106 | val start = subrange.start | 
| 38486 
f5bbfc019937
Markup_Tree.select: crude version of stream-based filtering;
 wenzelm parents: 
38482diff
changeset | 107 | |
| 38659 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 108 |           if (result.isDefinedAt(info)) {
 | 
| 39177 
0468964aec11
simplified Markup_Tree.select: Stream instead of Iterator (again), explicit Option instead of default;
 wenzelm parents: 
38845diff
changeset | 109 | val next = Text.Info[Option[A]](subrange, Some(result(info))) | 
| 38659 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 110 | val nexts = stream(start, (next, subtree) :: (parent, more) :: rest) | 
| 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 111 | if (last < start) parent.restrict(Text.Range(last, start)) #:: nexts | 
| 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 112 | else nexts | 
| 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 113 | } | 
| 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 114 | else stream(last, (parent, subtree #::: more) :: rest) | 
| 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 115 | |
| 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 116 | case (parent, Stream.Empty) :: rest => | 
| 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 117 | val stop = parent.range.stop | 
| 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 118 | val nexts = stream(stop, rest) | 
| 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 119 | if (last < stop) parent.restrict(Text.Range(last, stop)) #:: nexts | 
| 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 120 | else nexts | 
| 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 121 | |
| 38726 
6d5f9af42eca
Markup_Tree.select: uniform treatment of root_range wrt. singularities, yielding empty result stream;
 wenzelm parents: 
38661diff
changeset | 122 | case Nil => | 
| 
6d5f9af42eca
Markup_Tree.select: uniform treatment of root_range wrt. singularities, yielding empty result stream;
 wenzelm parents: 
38661diff
changeset | 123 | val stop = root_range.stop | 
| 39177 
0468964aec11
simplified Markup_Tree.select: Stream instead of Iterator (again), explicit Option instead of default;
 wenzelm parents: 
38845diff
changeset | 124 | if (last < stop) Stream(Text.Info(Text.Range(last, stop), None)) | 
| 38726 
6d5f9af42eca
Markup_Tree.select: uniform treatment of root_range wrt. singularities, yielding empty result stream;
 wenzelm parents: 
38661diff
changeset | 125 | else Stream.empty | 
| 38659 
afac51977705
Markup_Tree.select: more straight-forward recursion producing one main stream, avoid fragmentation of parent info due to ignored subtree;
 wenzelm parents: 
38657diff
changeset | 126 | } | 
| 38486 
f5bbfc019937
Markup_Tree.select: crude version of stream-based filtering;
 wenzelm parents: 
38482diff
changeset | 127 | } | 
| 39177 
0468964aec11
simplified Markup_Tree.select: Stream instead of Iterator (again), explicit Option instead of default;
 wenzelm parents: 
38845diff
changeset | 128 | stream(root_range.start, List((Text.Info(root_range, None), overlapping(root_range)))) | 
| 38486 
f5bbfc019937
Markup_Tree.select: crude version of stream-based filtering;
 wenzelm parents: 
38482diff
changeset | 129 | } | 
| 
f5bbfc019937
Markup_Tree.select: crude version of stream-based filtering;
 wenzelm parents: 
38482diff
changeset | 130 | |
| 38577 | 131 | def swing_tree(parent: DefaultMutableTreeNode) | 
| 132 | (swing_node: Text.Info[Any] => DefaultMutableTreeNode) | |
| 38479 
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
 wenzelm parents: 
38478diff
changeset | 133 |   {
 | 
| 38577 | 134 |     for ((_, (info, subtree)) <- branches) {
 | 
| 135 | val current = swing_node(info) | |
| 38479 
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
 wenzelm parents: 
38478diff
changeset | 136 | subtree.swing_tree(current)(swing_node) | 
| 
e628da370072
more efficient Markup_Tree, based on branches sorted by quasi-order;
 wenzelm parents: 
38478diff
changeset | 137 | parent.add(current) | 
| 34514 
2104a836b415
renamed fields of MarkupNode; implemented flatten and leafs
 immler@in.tum.de parents: 
34503diff
changeset | 138 | } | 
| 
2104a836b415
renamed fields of MarkupNode; implemented flatten and leafs
 immler@in.tum.de parents: 
34503diff
changeset | 139 | } | 
| 34717 
3f32e08bbb6c
sidekick root data: set buffer length to avoid crash of initial caret move;
 wenzelm parents: 
34708diff
changeset | 140 | } | 
| 34514 
2104a836b415
renamed fields of MarkupNode; implemented flatten and leafs
 immler@in.tum.de parents: 
34503diff
changeset | 141 |