--- a/src/Pure/Isar/outer_syntax.scala Sat Oct 07 12:50:05 2017 +0200
+++ b/src/Pure/Isar/outer_syntax.scala Sat Oct 07 13:13:46 2017 +0200
@@ -18,9 +18,7 @@
def init(): Outer_Syntax = new Outer_Syntax(completion = Completion.init())
- def merge(syns: List[Outer_Syntax]): Outer_Syntax =
- if (syns.isEmpty) Thy_Header.bootstrap_syntax
- else (syns.head /: syns.tail)(_ ++ _)
+ def merge(syns: List[Outer_Syntax]): Outer_Syntax = (empty /: syns)(_ ++ _)
/* string literals */
@@ -109,6 +107,7 @@
def ++ (other: Outer_Syntax): Outer_Syntax =
if (this eq other) this
+ else if (this eq Outer_Syntax.empty) other
else {
val keywords1 = keywords ++ other.keywords
val completion1 = completion ++ other.completion
--- a/src/Pure/Thy/thy_info.scala Sat Oct 07 12:50:05 2017 +0200
+++ b/src/Pure/Thy/thy_info.scala Sat Oct 07 13:13:46 2017 +0200
@@ -49,12 +49,13 @@
val name = entry.name.theory
val imports = entry.header.imports.map(p => p._1.theory)
- val graph1 = (graph /: (name :: imports))(_.default_node(_, Thy_Header.bootstrap_syntax))
+ val graph1 = (graph /: (name :: imports))(_.default_node(_, Outer_Syntax.empty))
val graph2 = (graph1 /: imports)(_.add_edge(_, name))
- val syntax =
- Outer_Syntax.merge((name :: graph2.imm_preds(name).toList).map(graph2.get_node(_))) +
- entry.header
+ val syntax0 = if (name == Thy_Header.PURE) List(Thy_Header.bootstrap_syntax) else Nil
+ val syntax1 = (name :: graph2.imm_preds(name).toList).map(graph2.get_node(_))
+ val syntax = Outer_Syntax.merge(syntax0 ::: syntax1) + entry.header
+
graph2.map_node(name, _ => syntax)
})