--- a/src/Pure/General/graph.scala Sat Jan 03 15:45:01 2015 +0100
+++ b/src/Pure/General/graph.scala Sat Jan 03 20:22:27 2015 +0100
@@ -21,15 +21,13 @@
def empty[Key, A](implicit ord: Ordering[Key]): Graph[Key, A] =
new Graph[Key, A](SortedMap.empty(ord))
- def make[Key, A](entries: List[((Key, A), List[Key])], converse: Boolean = false)(
+ def make[Key, A](entries: List[((Key, A), List[Key])])(
implicit ord: Ordering[Key]): Graph[Key, A] =
{
val graph1 =
(empty[Key, A](ord) /: entries) { case (graph, ((x, info), _)) => graph.new_node(x, info) }
val graph2 =
- (graph1 /: entries) { case (graph, ((x, _), ys)) =>
- if (converse) (graph /: ys)(_.add_edge(_, x)) else (graph /: ys)(_.add_edge(x, _))
- }
+ (graph1 /: entries) { case (graph, ((x, _), ys)) => (graph /: ys)(_.add_edge(x, _)) }
graph2
}
@@ -46,11 +44,11 @@
list(pair(pair(key, info), list(key)))(graph.dest)
})
- def decode[Key, A](key: XML.Decode.T[Key], info: XML.Decode.T[A], converse: Boolean = false)(
+ def decode[Key, A](key: XML.Decode.T[Key], info: XML.Decode.T[A])(
implicit ord: Ordering[Key]): XML.Decode.T[Graph[Key, A]] =
((body: XML.Body) => {
import XML.Decode._
- make(list(pair(pair(key, info), list(key)))(body), converse)(ord)
+ make(list(pair(pair(key, info), list(key)))(body))(ord)
})
}