| author | wenzelm | 
| Sat, 19 Apr 2014 17:23:05 +0200 | |
| changeset 56618 | 874bdedb2313 | 
| parent 56393 | 22f533e6a049 | 
| child 57901 | e1abca2527da | 
| permissions | -rw-r--r-- | 
| 34166 | 1  | 
/* Title: Pure/Isar/outer_syntax.scala  | 
2  | 
Author: Makarius  | 
|
3  | 
||
4  | 
Isabelle/Isar outer syntax.  | 
|
5  | 
*/  | 
|
6  | 
||
7  | 
package isabelle  | 
|
8  | 
||
9  | 
||
10  | 
import scala.util.parsing.input.{Reader, CharSequenceReader}
 | 
|
| 
43411
 
0206466ee473
some support for partial scans with explicit context;
 
wenzelm 
parents: 
40533 
diff
changeset
 | 
11  | 
import scala.collection.mutable  | 
| 34166 | 12  | 
|
13  | 
||
| 
43774
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
14  | 
object Outer_Syntax  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
15  | 
{
 | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
16  | 
def quote_string(str: String): String =  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
17  | 
  {
 | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
18  | 
val result = new StringBuilder(str.length + 10)  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
19  | 
result += '"'  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
20  | 
    for (s <- Symbol.iterator(str)) {
 | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
21  | 
      if (s.length == 1) {
 | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
22  | 
val c = s(0)  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
23  | 
        if (c < 32 && c != YXML.X && c != YXML.Y || c == '\\' || c == '"') {
 | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
24  | 
result += '\\'  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
25  | 
if (c < 10) result += '0'  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
26  | 
if (c < 100) result += '0'  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
27  | 
result ++= (c.asInstanceOf[Int].toString)  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
28  | 
}  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
29  | 
else result += c  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
30  | 
}  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
31  | 
else result ++= s  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
32  | 
}  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
33  | 
result += '"'  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
34  | 
result.toString  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
35  | 
}  | 
| 46626 | 36  | 
|
| 46941 | 37  | 
val empty: Outer_Syntax = new Outer_Syntax()  | 
| 
48870
 
4accee106f0f
clarified initialization of Thy_Load, Thy_Info, Session;
 
wenzelm 
parents: 
48864 
diff
changeset
 | 
38  | 
|
| 46941 | 39  | 
def init(): Outer_Syntax = new Outer_Syntax(completion = Completion.init())  | 
| 
43774
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
40  | 
}  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
41  | 
|
| 46712 | 42  | 
final class Outer_Syntax private(  | 
| 
48864
 
3ee314ae1e0a
added keyword kind "thy_load" (with optional list of file extensions);
 
wenzelm 
parents: 
48708 
diff
changeset
 | 
43  | 
keywords: Map[String, (String, List[String])] = Map.empty,  | 
| 46626 | 44  | 
lexicon: Scan.Lexicon = Scan.Lexicon.empty,  | 
| 
53280
 
c63a016805b9
explicit indication of outer syntax with no tokens;
 
wenzelm 
parents: 
52439 
diff
changeset
 | 
45  | 
val completion: Completion = Completion.empty,  | 
| 55749 | 46  | 
val language_context: Completion.Language_Context = Completion.Language_Context.outer,  | 
| 
56393
 
22f533e6a049
more abstract Prover.Syntax, as proposed by Carst Tankink;
 
wenzelm 
parents: 
56314 
diff
changeset
 | 
47  | 
val has_tokens: Boolean = true) extends Prover.Syntax  | 
| 34166 | 48  | 
{
 | 
| 
48660
 
730ca503e955
static outer syntax based on session specifications;
 
wenzelm 
parents: 
47469 
diff
changeset
 | 
49  | 
override def toString: String =  | 
| 
48864
 
3ee314ae1e0a
added keyword kind "thy_load" (with optional list of file extensions);
 
wenzelm 
parents: 
48708 
diff
changeset
 | 
50  | 
    (for ((name, (kind, files)) <- keywords) yield {
 | 
| 
48660
 
730ca503e955
static outer syntax based on session specifications;
 
wenzelm 
parents: 
47469 
diff
changeset
 | 
51  | 
if (kind == Keyword.MINOR) quote(name)  | 
| 
48864
 
3ee314ae1e0a
added keyword kind "thy_load" (with optional list of file extensions);
 
wenzelm 
parents: 
48708 
diff
changeset
 | 
52  | 
else  | 
| 
 
3ee314ae1e0a
added keyword kind "thy_load" (with optional list of file extensions);
 
wenzelm 
parents: 
48708 
diff
changeset
 | 
53  | 
quote(name) + " :: " + quote(kind) +  | 
| 
 
3ee314ae1e0a
added keyword kind "thy_load" (with optional list of file extensions);
 
wenzelm 
parents: 
48708 
diff
changeset
 | 
54  | 
        (if (files.isEmpty) "" else " (" + commas_quote(files) + ")")
 | 
| 48671 | 55  | 
    }).toList.sorted.mkString("keywords\n  ", " and\n  ", "")
 | 
| 
48660
 
730ca503e955
static outer syntax based on session specifications;
 
wenzelm 
parents: 
47469 
diff
changeset
 | 
56  | 
|
| 
48864
 
3ee314ae1e0a
added keyword kind "thy_load" (with optional list of file extensions);
 
wenzelm 
parents: 
48708 
diff
changeset
 | 
57  | 
def keyword_kind_files(name: String): Option[(String, List[String])] = keywords.get(name)  | 
| 
 
3ee314ae1e0a
added keyword kind "thy_load" (with optional list of file extensions);
 
wenzelm 
parents: 
48708 
diff
changeset
 | 
58  | 
def keyword_kind(name: String): Option[String] = keyword_kind_files(name).map(_._1)  | 
| 
38471
 
0924654b8163
report command token name instead of kind, which can be retrieved later via Outer_Syntax.keyword_kind;
 
wenzelm 
parents: 
36956 
diff
changeset
 | 
59  | 
|
| 56314 | 60  | 
def load(span: List[Token]): Option[List[String]] =  | 
| 54462 | 61  | 
    keywords.get(Command.name(span)) match {
 | 
| 54513 | 62  | 
case Some((Keyword.THY_LOAD, exts)) => Some(exts)  | 
| 54462 | 63  | 
case _ => None  | 
64  | 
}  | 
|
65  | 
||
| 56314 | 66  | 
val load_commands: List[(String, List[String])] =  | 
| 48885 | 67  | 
(for ((name, (Keyword.THY_LOAD, files)) <- keywords.iterator) yield (name, files)).toList  | 
| 48872 | 68  | 
|
| 
56393
 
22f533e6a049
more abstract Prover.Syntax, as proposed by Carst Tankink;
 
wenzelm 
parents: 
56314 
diff
changeset
 | 
69  | 
def load_commands_in(text: String): Boolean =  | 
| 
 
22f533e6a049
more abstract Prover.Syntax, as proposed by Carst Tankink;
 
wenzelm 
parents: 
56314 
diff
changeset
 | 
70  | 
    load_commands.exists({ case (cmd, _) => text.containsSlice(cmd) })
 | 
| 
 
22f533e6a049
more abstract Prover.Syntax, as proposed by Carst Tankink;
 
wenzelm 
parents: 
56314 
diff
changeset
 | 
71  | 
|
| 
50128
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
48885 
diff
changeset
 | 
72  | 
def + (name: String, kind: (String, List[String]), replace: Option[String]): Outer_Syntax =  | 
| 
53280
 
c63a016805b9
explicit indication of outer syntax with no tokens;
 
wenzelm 
parents: 
52439 
diff
changeset
 | 
73  | 
  {
 | 
| 
 
c63a016805b9
explicit indication of outer syntax with no tokens;
 
wenzelm 
parents: 
52439 
diff
changeset
 | 
74  | 
val keywords1 = keywords + (name -> kind)  | 
| 
 
c63a016805b9
explicit indication of outer syntax with no tokens;
 
wenzelm 
parents: 
52439 
diff
changeset
 | 
75  | 
val lexicon1 = lexicon + name  | 
| 
 
c63a016805b9
explicit indication of outer syntax with no tokens;
 
wenzelm 
parents: 
52439 
diff
changeset
 | 
76  | 
val completion1 =  | 
| 
50128
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
48885 
diff
changeset
 | 
77  | 
      if (Keyword.control(kind._1) || replace == Some("")) completion
 | 
| 
53280
 
c63a016805b9
explicit indication of outer syntax with no tokens;
 
wenzelm 
parents: 
52439 
diff
changeset
 | 
78  | 
else completion + (name, replace getOrElse name)  | 
| 55749 | 79  | 
new Outer_Syntax(keywords1, lexicon1, completion1, language_context, true)  | 
| 
53280
 
c63a016805b9
explicit indication of outer syntax with no tokens;
 
wenzelm 
parents: 
52439 
diff
changeset
 | 
80  | 
}  | 
| 34166 | 81  | 
|
| 
53280
 
c63a016805b9
explicit indication of outer syntax with no tokens;
 
wenzelm 
parents: 
52439 
diff
changeset
 | 
82  | 
def + (name: String, kind: (String, List[String])): Outer_Syntax =  | 
| 
 
c63a016805b9
explicit indication of outer syntax with no tokens;
 
wenzelm 
parents: 
52439 
diff
changeset
 | 
83  | 
this + (name, kind, Some(name))  | 
| 
 
c63a016805b9
explicit indication of outer syntax with no tokens;
 
wenzelm 
parents: 
52439 
diff
changeset
 | 
84  | 
def + (name: String, kind: String): Outer_Syntax =  | 
| 
 
c63a016805b9
explicit indication of outer syntax with no tokens;
 
wenzelm 
parents: 
52439 
diff
changeset
 | 
85  | 
this + (name, (kind, Nil), Some(name))  | 
| 
50128
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
48885 
diff
changeset
 | 
86  | 
def + (name: String, replace: Option[String]): Outer_Syntax =  | 
| 
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
48885 
diff
changeset
 | 
87  | 
this + (name, (Keyword.MINOR, Nil), replace)  | 
| 
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
48885 
diff
changeset
 | 
88  | 
def + (name: String): Outer_Syntax = this + (name, None)  | 
| 48706 | 89  | 
|
| 48873 | 90  | 
def add_keywords(keywords: Thy_Header.Keywords): Outer_Syntax =  | 
91  | 
    (this /: keywords) {
 | 
|
| 
52439
 
4cf3f6153eb8
improved "isabelle keywords" and "isabelle update_keywords" based on Isabelle/Scala, without requiring to build sessions first;
 
wenzelm 
parents: 
52066 
diff
changeset
 | 
92  | 
case (syntax, (name, Some((kind, _)), replace)) =>  | 
| 
50128
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
48885 
diff
changeset
 | 
93  | 
syntax +  | 
| 
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
48885 
diff
changeset
 | 
94  | 
(Symbol.decode(name), kind, replace) +  | 
| 
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
48885 
diff
changeset
 | 
95  | 
(Symbol.encode(name), kind, replace)  | 
| 
52439
 
4cf3f6153eb8
improved "isabelle keywords" and "isabelle update_keywords" based on Isabelle/Scala, without requiring to build sessions first;
 
wenzelm 
parents: 
52066 
diff
changeset
 | 
96  | 
case (syntax, (name, None, replace)) =>  | 
| 
50128
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
48885 
diff
changeset
 | 
97  | 
syntax +  | 
| 
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
48885 
diff
changeset
 | 
98  | 
(Symbol.decode(name), replace) +  | 
| 
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
48885 
diff
changeset
 | 
99  | 
(Symbol.encode(name), replace)  | 
| 46940 | 100  | 
}  | 
| 34166 | 101  | 
|
102  | 
def is_command(name: String): Boolean =  | 
|
| 
40458
 
12c8c64203b3
treat main theory commands like headings, and nest anything else inside;
 
wenzelm 
parents: 
40455 
diff
changeset
 | 
103  | 
    keyword_kind(name) match {
 | 
| 36947 | 104  | 
case Some(kind) => kind != Keyword.MINOR  | 
| 34166 | 105  | 
case None => false  | 
106  | 
}  | 
|
107  | 
||
| 
40454
 
2516ea25a54b
some support for nested source structure, based on section headings;
 
wenzelm 
parents: 
38471 
diff
changeset
 | 
108  | 
def heading_level(name: String): Option[Int] =  | 
| 46969 | 109  | 
  {
 | 
110  | 
    keyword_kind(name) match {
 | 
|
111  | 
case _ if name == "header" => Some(0)  | 
|
112  | 
case Some(Keyword.THY_HEADING1) => Some(1)  | 
|
113  | 
case Some(Keyword.THY_HEADING2) | Some(Keyword.PRF_HEADING2) => Some(2)  | 
|
114  | 
case Some(Keyword.THY_HEADING3) | Some(Keyword.PRF_HEADING3) => Some(3)  | 
|
115  | 
case Some(Keyword.THY_HEADING4) | Some(Keyword.PRF_HEADING4) => Some(4)  | 
|
116  | 
case Some(kind) if Keyword.theory(kind) => Some(5)  | 
|
117  | 
case _ => None  | 
|
| 
40454
 
2516ea25a54b
some support for nested source structure, based on section headings;
 
wenzelm 
parents: 
38471 
diff
changeset
 | 
118  | 
}  | 
| 46969 | 119  | 
}  | 
| 
40454
 
2516ea25a54b
some support for nested source structure, based on section headings;
 
wenzelm 
parents: 
38471 
diff
changeset
 | 
120  | 
|
| 
 
2516ea25a54b
some support for nested source structure, based on section headings;
 
wenzelm 
parents: 
38471 
diff
changeset
 | 
121  | 
def heading_level(command: Command): Option[Int] =  | 
| 
 
2516ea25a54b
some support for nested source structure, based on section headings;
 
wenzelm 
parents: 
38471 
diff
changeset
 | 
122  | 
heading_level(command.name)  | 
| 
 
2516ea25a54b
some support for nested source structure, based on section headings;
 
wenzelm 
parents: 
38471 
diff
changeset
 | 
123  | 
|
| 34166 | 124  | 
|
| 
53280
 
c63a016805b9
explicit indication of outer syntax with no tokens;
 
wenzelm 
parents: 
52439 
diff
changeset
 | 
125  | 
/* token language */  | 
| 
 
c63a016805b9
explicit indication of outer syntax with no tokens;
 
wenzelm 
parents: 
52439 
diff
changeset
 | 
126  | 
|
| 
36956
 
21be4832c362
renamed class Outer_Lex to Token and Token_Kind to Token.Kind;
 
wenzelm 
parents: 
36947 
diff
changeset
 | 
127  | 
def scan(input: Reader[Char]): List[Token] =  | 
| 
52066
 
83b7b88770c9
discontinued odd workaround for scala-2.9.2, which is hopefully obsolete in scala-2.10.x;
 
wenzelm 
parents: 
50428 
diff
changeset
 | 
128  | 
  {
 | 
| 55616 | 129  | 
Token.Parsers.parseAll(  | 
130  | 
        Token.Parsers.rep(Token.Parsers.token(lexicon, is_command)), input) match {
 | 
|
| 55494 | 131  | 
case Token.Parsers.Success(tokens, _) => tokens  | 
| 
52066
 
83b7b88770c9
discontinued odd workaround for scala-2.9.2, which is hopefully obsolete in scala-2.10.x;
 
wenzelm 
parents: 
50428 
diff
changeset
 | 
132  | 
      case _ => error("Unexpected failure of tokenizing input:\n" + input.source.toString)
 | 
| 34166 | 133  | 
}  | 
| 
52066
 
83b7b88770c9
discontinued odd workaround for scala-2.9.2, which is hopefully obsolete in scala-2.10.x;
 
wenzelm 
parents: 
50428 
diff
changeset
 | 
134  | 
}  | 
| 34166 | 135  | 
|
| 
36956
 
21be4832c362
renamed class Outer_Lex to Token and Token_Kind to Token.Kind;
 
wenzelm 
parents: 
36947 
diff
changeset
 | 
136  | 
def scan(input: CharSequence): List[Token] =  | 
| 34166 | 137  | 
scan(new CharSequenceReader(input))  | 
| 
43411
 
0206466ee473
some support for partial scans with explicit context;
 
wenzelm 
parents: 
40533 
diff
changeset
 | 
138  | 
|
| 
55510
 
1585a65aad64
tuned signature -- emphasize line-oriented aspect;
 
wenzelm 
parents: 
55494 
diff
changeset
 | 
139  | 
def scan_line(input: CharSequence, context: Scan.Line_Context): (List[Token], Scan.Line_Context) =  | 
| 
52066
 
83b7b88770c9
discontinued odd workaround for scala-2.9.2, which is hopefully obsolete in scala-2.10.x;
 
wenzelm 
parents: 
50428 
diff
changeset
 | 
140  | 
  {
 | 
| 
 
83b7b88770c9
discontinued odd workaround for scala-2.9.2, which is hopefully obsolete in scala-2.10.x;
 
wenzelm 
parents: 
50428 
diff
changeset
 | 
141  | 
var in: Reader[Char] = new CharSequenceReader(input)  | 
| 
 
83b7b88770c9
discontinued odd workaround for scala-2.9.2, which is hopefully obsolete in scala-2.10.x;
 
wenzelm 
parents: 
50428 
diff
changeset
 | 
142  | 
val toks = new mutable.ListBuffer[Token]  | 
| 
 
83b7b88770c9
discontinued odd workaround for scala-2.9.2, which is hopefully obsolete in scala-2.10.x;
 
wenzelm 
parents: 
50428 
diff
changeset
 | 
143  | 
var ctxt = context  | 
| 
 
83b7b88770c9
discontinued odd workaround for scala-2.9.2, which is hopefully obsolete in scala-2.10.x;
 
wenzelm 
parents: 
50428 
diff
changeset
 | 
144  | 
    while (!in.atEnd) {
 | 
| 
55510
 
1585a65aad64
tuned signature -- emphasize line-oriented aspect;
 
wenzelm 
parents: 
55494 
diff
changeset
 | 
145  | 
      Token.Parsers.parse(Token.Parsers.token_line(lexicon, is_command, ctxt), in) match {
 | 
| 55494 | 146  | 
        case Token.Parsers.Success((x, c), rest) => { toks += x; ctxt = c; in = rest }
 | 
147  | 
case Token.Parsers.NoSuccess(_, rest) =>  | 
|
| 
52066
 
83b7b88770c9
discontinued odd workaround for scala-2.9.2, which is hopefully obsolete in scala-2.10.x;
 
wenzelm 
parents: 
50428 
diff
changeset
 | 
148  | 
          error("Unexpected failure of tokenizing input:\n" + rest.source.toString)
 | 
| 
43411
 
0206466ee473
some support for partial scans with explicit context;
 
wenzelm 
parents: 
40533 
diff
changeset
 | 
149  | 
}  | 
| 
 
0206466ee473
some support for partial scans with explicit context;
 
wenzelm 
parents: 
40533 
diff
changeset
 | 
150  | 
}  | 
| 
52066
 
83b7b88770c9
discontinued odd workaround for scala-2.9.2, which is hopefully obsolete in scala-2.10.x;
 
wenzelm 
parents: 
50428 
diff
changeset
 | 
151  | 
(toks.toList, ctxt)  | 
| 
 
83b7b88770c9
discontinued odd workaround for scala-2.9.2, which is hopefully obsolete in scala-2.10.x;
 
wenzelm 
parents: 
50428 
diff
changeset
 | 
152  | 
}  | 
| 55616 | 153  | 
|
154  | 
||
155  | 
/* language context */  | 
|
156  | 
||
| 55749 | 157  | 
def set_language_context(context: Completion.Language_Context): Outer_Syntax =  | 
| 55616 | 158  | 
new Outer_Syntax(keywords, lexicon, completion, context, has_tokens)  | 
159  | 
||
160  | 
def no_tokens: Outer_Syntax =  | 
|
161  | 
  {
 | 
|
162  | 
require(keywords.isEmpty && lexicon.isEmpty)  | 
|
163  | 
new Outer_Syntax(  | 
|
164  | 
completion = completion,  | 
|
| 55749 | 165  | 
language_context = language_context,  | 
| 55616 | 166  | 
has_tokens = false)  | 
167  | 
}  | 
|
| 34166 | 168  | 
}  |