| author | wenzelm | 
| Wed, 06 Apr 2016 17:17:05 +0200 | |
| changeset 62892 | 7485507620b6 | 
| parent 62453 | b93cc7d73431 | 
| child 63424 | e4e15bbfb3e2 | 
| 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  | 
||
| 
43411
 
0206466ee473
some support for partial scans with explicit context;
 
wenzelm 
parents: 
40533 
diff
changeset
 | 
10  | 
import scala.collection.mutable  | 
| 58706 | 11  | 
import scala.annotation.tailrec  | 
| 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  | 
{
 | 
| 58706 | 16  | 
/* syntax */  | 
17  | 
||
18  | 
val empty: Outer_Syntax = new Outer_Syntax()  | 
|
19  | 
||
20  | 
def init(): Outer_Syntax = new Outer_Syntax(completion = Completion.init())  | 
|
21  | 
||
22  | 
||
23  | 
/* string literals */  | 
|
24  | 
||
| 
43774
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
25  | 
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
 | 
26  | 
  {
 | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
27  | 
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
 | 
28  | 
result += '"'  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
29  | 
    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
 | 
30  | 
      if (s.length == 1) {
 | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
31  | 
val c = s(0)  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
32  | 
        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
 | 
33  | 
result += '\\'  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
34  | 
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
 | 
35  | 
if (c < 100) result += '0'  | 
| 60215 | 36  | 
result ++= c.asInstanceOf[Int].toString  | 
| 
43774
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
37  | 
}  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
38  | 
else result += c  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
39  | 
}  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
40  | 
else result ++= s  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
41  | 
}  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
42  | 
result += '"'  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
43  | 
result.toString  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
44  | 
}  | 
| 46626 | 45  | 
|
| 58696 | 46  | 
|
| 58697 | 47  | 
/* line-oriented structure */  | 
| 58696 | 48  | 
|
| 58697 | 49  | 
object Line_Structure  | 
| 58696 | 50  | 
  {
 | 
| 58700 | 51  | 
val init = Line_Structure()  | 
| 58696 | 52  | 
}  | 
53  | 
||
| 58700 | 54  | 
sealed case class Line_Structure(  | 
55  | 
improper: Boolean = true,  | 
|
56  | 
command: Boolean = false,  | 
|
57  | 
depth: Int = 0,  | 
|
58  | 
span_depth: Int = 0,  | 
|
59  | 
after_span_depth: Int = 0)  | 
|
| 58706 | 60  | 
|
61  | 
||
62  | 
/* overall document structure */  | 
|
63  | 
||
64  | 
  sealed abstract class Document { def length: Int }
 | 
|
| 58747 | 65  | 
case class Document_Block(name: String, text: String, body: List[Document]) extends Document  | 
| 58706 | 66  | 
  {
 | 
67  | 
val length: Int = (0 /: body)(_ + _.length)  | 
|
68  | 
}  | 
|
| 58747 | 69  | 
case class Document_Atom(command: Command) extends Document  | 
| 58706 | 70  | 
  {
 | 
71  | 
def length: Int = command.length  | 
|
72  | 
}  | 
|
| 
43774
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
73  | 
}  | 
| 
 
6dfdb70496fe
added Outer_Syntax.quote_string, which is conceptually a bit different from Token.unparse;
 
wenzelm 
parents: 
43695 
diff
changeset
 | 
74  | 
|
| 46712 | 75  | 
final class Outer_Syntax private(  | 
| 58900 | 76  | 
val keywords: Keyword.Keywords = Keyword.Keywords.empty,  | 
| 
53280
 
c63a016805b9
explicit indication of outer syntax with no tokens;
 
wenzelm 
parents: 
52439 
diff
changeset
 | 
77  | 
val completion: Completion = Completion.empty,  | 
| 55749 | 78  | 
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
 | 
79  | 
val has_tokens: Boolean = true) extends Prover.Syntax  | 
| 34166 | 80  | 
{
 | 
| 58706 | 81  | 
/** syntax content **/  | 
82  | 
||
| 58900 | 83  | 
override def toString: String = keywords.toString  | 
| 
56393
 
22f533e6a049
more abstract Prover.Syntax, as proposed by Carst Tankink;
 
wenzelm 
parents: 
56314 
diff
changeset
 | 
84  | 
|
| 58695 | 85  | 
|
86  | 
/* add keywords */  | 
|
87  | 
||
| 58907 | 88  | 
def + (name: String): Outer_Syntax = this + (name, None, None)  | 
89  | 
def + (name: String, kind: String): Outer_Syntax = this + (name, Some((kind, Nil)), None)  | 
|
| 58901 | 90  | 
def + (name: String, opt_kind: Option[(String, List[String])], replace: Option[String])  | 
91  | 
: Outer_Syntax =  | 
|
| 
53280
 
c63a016805b9
explicit indication of outer syntax with no tokens;
 
wenzelm 
parents: 
52439 
diff
changeset
 | 
92  | 
  {
 | 
| 58901 | 93  | 
val keywords1 =  | 
94  | 
      opt_kind match {
 | 
|
95  | 
case None => keywords + name  | 
|
96  | 
case Some(kind) => keywords + (name, kind)  | 
|
97  | 
}  | 
|
| 
53280
 
c63a016805b9
explicit indication of outer syntax with no tokens;
 
wenzelm 
parents: 
52439 
diff
changeset
 | 
98  | 
val completion1 =  | 
| 58853 | 99  | 
      if (replace == Some("")) completion
 | 
| 
53280
 
c63a016805b9
explicit indication of outer syntax with no tokens;
 
wenzelm 
parents: 
52439 
diff
changeset
 | 
100  | 
else completion + (name, replace getOrElse name)  | 
| 58900 | 101  | 
new Outer_Syntax(keywords1, completion1, language_context, true)  | 
| 
53280
 
c63a016805b9
explicit indication of outer syntax with no tokens;
 
wenzelm 
parents: 
52439 
diff
changeset
 | 
102  | 
}  | 
| 48706 | 103  | 
|
| 48873 | 104  | 
def add_keywords(keywords: Thy_Header.Keywords): Outer_Syntax =  | 
105  | 
    (this /: keywords) {
 | 
|
| 58901 | 106  | 
case (syntax, (name, opt_spec, replace)) =>  | 
107  | 
val opt_kind = opt_spec.map(_._1)  | 
|
| 
50128
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
48885 
diff
changeset
 | 
108  | 
syntax +  | 
| 58901 | 109  | 
(Symbol.decode(name), opt_kind, replace) +  | 
110  | 
(Symbol.encode(name), opt_kind, replace)  | 
|
| 46940 | 111  | 
}  | 
| 34166 | 112  | 
|
| 58695 | 113  | 
|
| 59073 | 114  | 
/* merge */  | 
115  | 
||
| 
59077
 
7e0d3da6e6d8
node-specific syntax, with base_syntax as default;
 
wenzelm 
parents: 
59073 
diff
changeset
 | 
116  | 
def ++ (other: Prover.Syntax): Prover.Syntax =  | 
| 59073 | 117  | 
if (this eq other) this  | 
118  | 
    else {
 | 
|
| 
59077
 
7e0d3da6e6d8
node-specific syntax, with base_syntax as default;
 
wenzelm 
parents: 
59073 
diff
changeset
 | 
119  | 
val keywords1 = keywords ++ other.asInstanceOf[Outer_Syntax].keywords  | 
| 
 
7e0d3da6e6d8
node-specific syntax, with base_syntax as default;
 
wenzelm 
parents: 
59073 
diff
changeset
 | 
120  | 
val completion1 = completion ++ other.asInstanceOf[Outer_Syntax].completion  | 
| 
 
7e0d3da6e6d8
node-specific syntax, with base_syntax as default;
 
wenzelm 
parents: 
59073 
diff
changeset
 | 
121  | 
if ((keywords eq keywords1) && (completion eq completion1)) this  | 
| 
 
7e0d3da6e6d8
node-specific syntax, with base_syntax as default;
 
wenzelm 
parents: 
59073 
diff
changeset
 | 
122  | 
else new Outer_Syntax(keywords1, completion1, language_context, has_tokens)  | 
| 59073 | 123  | 
}  | 
124  | 
||
125  | 
||
| 59735 | 126  | 
/* load commands */  | 
| 58900 | 127  | 
|
128  | 
def load_command(name: String): Option[List[String]] = keywords.load_command(name)  | 
|
129  | 
def load_commands_in(text: String): Boolean = keywords.load_commands_in(text)  | 
|
130  | 
||
131  | 
||
| 58706 | 132  | 
/* language context */  | 
| 34166 | 133  | 
|
| 58706 | 134  | 
def set_language_context(context: Completion.Language_Context): Outer_Syntax =  | 
| 58900 | 135  | 
new Outer_Syntax(keywords, completion, context, has_tokens)  | 
| 58706 | 136  | 
|
137  | 
def no_tokens: Outer_Syntax =  | 
|
| 46969 | 138  | 
  {
 | 
| 58900 | 139  | 
require(keywords.is_empty)  | 
| 58706 | 140  | 
new Outer_Syntax(  | 
141  | 
completion = completion,  | 
|
142  | 
language_context = language_context,  | 
|
143  | 
has_tokens = false)  | 
|
| 46969 | 144  | 
}  | 
| 
40454
 
2516ea25a54b
some support for nested source structure, based on section headings;
 
wenzelm 
parents: 
38471 
diff
changeset
 | 
145  | 
|
| 58706 | 146  | 
|
| 
40454
 
2516ea25a54b
some support for nested source structure, based on section headings;
 
wenzelm 
parents: 
38471 
diff
changeset
 | 
147  | 
|
| 58706 | 148  | 
/** parsing **/  | 
| 34166 | 149  | 
|
| 58697 | 150  | 
/* line-oriented structure */  | 
| 58696 | 151  | 
|
| 
59924
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59735 
diff
changeset
 | 
152  | 
def line_structure(tokens: List[Token], structure: Outer_Syntax.Line_Structure)  | 
| 58700 | 153  | 
: Outer_Syntax.Line_Structure =  | 
| 58696 | 154  | 
  {
 | 
| 58700 | 155  | 
val improper1 = tokens.forall(_.is_improper)  | 
156  | 
val command1 = tokens.exists(_.is_command)  | 
|
157  | 
||
| 58696 | 158  | 
val depth1 =  | 
| 59122 | 159  | 
if (tokens.exists(tok => tok.is_command_kind(keywords, Keyword.theory))) 0  | 
| 
59924
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59735 
diff
changeset
 | 
160  | 
else if (command1) structure.after_span_depth  | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59735 
diff
changeset
 | 
161  | 
else structure.span_depth  | 
| 58700 | 162  | 
|
163  | 
val (span_depth1, after_span_depth1) =  | 
|
| 
59924
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59735 
diff
changeset
 | 
164  | 
      ((structure.span_depth, structure.after_span_depth) /: tokens) {
 | 
| 58703 | 165  | 
case ((x, y), tok) =>  | 
166  | 
          if (tok.is_command) {
 | 
|
| 60692 | 167  | 
if (tok.is_command_kind(keywords, Keyword.theory_goal)) (2, 1)  | 
168  | 
else if (tok.is_command_kind(keywords, Keyword.theory)) (1, 0)  | 
|
| 
62244
 
5d513565749e
proper nesting: 'qed' needs to close the corresponding 'proof' and goal statement;
 
wenzelm 
parents: 
61463 
diff
changeset
 | 
169  | 
else if (tok.is_command_kind(keywords, Keyword.proof_open)) (y + 2, y + 1)  | 
| 
60694
 
b3fa4a8cdb5f
clarified text folds: proof ... qed counts as extra block;
 
wenzelm 
parents: 
60692 
diff
changeset
 | 
170  | 
else if (tok.is_command_kind(keywords, Keyword.PRF_BLOCK == _)) (y + 2, y + 1)  | 
| 
62244
 
5d513565749e
proper nesting: 'qed' needs to close the corresponding 'proof' and goal statement;
 
wenzelm 
parents: 
61463 
diff
changeset
 | 
171  | 
else if (tok.is_command_kind(keywords, Keyword.QED_BLOCK == _)) (y + 1, y - 2)  | 
| 60692 | 172  | 
else if (tok.is_command_kind(keywords, Keyword.proof_close)) (y + 1, y - 1)  | 
173  | 
else if (tok.is_command_kind(keywords, Keyword.qed_global)) (1, 0)  | 
|
| 58703 | 174  | 
else (x, y)  | 
175  | 
}  | 
|
176  | 
else (x, y)  | 
|
| 58696 | 177  | 
}  | 
| 58700 | 178  | 
|
179  | 
Outer_Syntax.Line_Structure(improper1, command1, depth1, span_depth1, after_span_depth1)  | 
|
| 58696 | 180  | 
}  | 
181  | 
||
182  | 
||
| 58706 | 183  | 
/* command spans */  | 
| 
57905
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
184  | 
|
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
185  | 
def parse_spans(toks: List[Token]): List[Command_Span.Span] =  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
186  | 
  {
 | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
187  | 
val result = new mutable.ListBuffer[Command_Span.Span]  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
188  | 
val content = new mutable.ListBuffer[Token]  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
189  | 
val improper = new mutable.ListBuffer[Token]  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
190  | 
|
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
191  | 
def ship(span: List[Token])  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
192  | 
    {
 | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
193  | 
val kind =  | 
| 
59924
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59735 
diff
changeset
 | 
194  | 
if (span.forall(_.is_improper)) Command_Span.Ignored_Span  | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59735 
diff
changeset
 | 
195  | 
else if (span.exists(_.is_error)) Command_Span.Malformed_Span  | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59735 
diff
changeset
 | 
196  | 
else  | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59735 
diff
changeset
 | 
197  | 
          span.find(_.is_command) match {
 | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59735 
diff
changeset
 | 
198  | 
case None => Command_Span.Malformed_Span  | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59735 
diff
changeset
 | 
199  | 
case Some(cmd) =>  | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59735 
diff
changeset
 | 
200  | 
val name = cmd.source  | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59735 
diff
changeset
 | 
201  | 
val offset =  | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59735 
diff
changeset
 | 
202  | 
                (0 /: span.takeWhile(_ != cmd)) {
 | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59735 
diff
changeset
 | 
203  | 
case (i, tok) => i + Symbol.iterator(tok.source).length }  | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59735 
diff
changeset
 | 
204  | 
val end_offset = offset + Symbol.iterator(name).length  | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59735 
diff
changeset
 | 
205  | 
val pos = Position.Range(Text.Range(offset, end_offset) + 1)  | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59735 
diff
changeset
 | 
206  | 
Command_Span.Command_Span(name, pos)  | 
| 
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59735 
diff
changeset
 | 
207  | 
}  | 
| 
57905
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
208  | 
result += Command_Span.Span(kind, span)  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
209  | 
}  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
210  | 
|
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
211  | 
def flush()  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
212  | 
    {
 | 
| 59319 | 213  | 
      if (content.nonEmpty) { ship(content.toList); content.clear }
 | 
214  | 
      if (improper.nonEmpty) { ship(improper.toList); improper.clear }
 | 
|
| 
57905
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
215  | 
}  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
216  | 
|
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
217  | 
    for (tok <- toks) {
 | 
| 
59924
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59735 
diff
changeset
 | 
218  | 
if (tok.is_improper) improper += tok  | 
| 
59939
 
7d46aa03696e
support for 'restricted' modifier: only qualified accesses outside the local scope;
 
wenzelm 
parents: 
59924 
diff
changeset
 | 
219  | 
else if (tok.is_command_modifier ||  | 
| 
 
7d46aa03696e
support for 'restricted' modifier: only qualified accesses outside the local scope;
 
wenzelm 
parents: 
59924 
diff
changeset
 | 
220  | 
tok.is_command && (!content.exists(_.is_command_modifier) || content.exists(_.is_command)))  | 
| 
59924
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59735 
diff
changeset
 | 
221  | 
      { flush(); content += tok }
 | 
| 
57905
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
222  | 
      else { content ++= improper; improper.clear; content += tok }
 | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
223  | 
}  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
224  | 
flush()  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
225  | 
|
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
226  | 
result.toList  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
227  | 
}  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
228  | 
|
| 57906 | 229  | 
def parse_spans(input: CharSequence): List[Command_Span.Span] =  | 
| 59083 | 230  | 
parse_spans(Token.explode(keywords, input))  | 
| 57906 | 231  | 
|
| 
57905
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents: 
57901 
diff
changeset
 | 
232  | 
|
| 58706 | 233  | 
/* overall document structure */  | 
| 55616 | 234  | 
|
| 58706 | 235  | 
def heading_level(command: Command): Option[Int] =  | 
236  | 
  {
 | 
|
| 59735 | 237  | 
val name = command.span.name  | 
238  | 
    name match {
 | 
|
239  | 
case Thy_Header.CHAPTER => Some(0)  | 
|
| 62453 | 240  | 
case Thy_Header.SECTION => Some(1)  | 
| 59735 | 241  | 
case Thy_Header.SUBSECTION => Some(2)  | 
242  | 
case Thy_Header.SUBSUBSECTION => Some(3)  | 
|
| 61463 | 243  | 
case Thy_Header.PARAGRAPH => Some(4)  | 
244  | 
case Thy_Header.SUBPARAGRAPH => Some(5)  | 
|
| 
58868
 
c5e1cce7ace3
uniform heading commands work in any context, even in theory header;
 
wenzelm 
parents: 
58853 
diff
changeset
 | 
245  | 
case _ =>  | 
| 59735 | 246  | 
        keywords.command_kind(name) match {
 | 
| 61463 | 247  | 
case Some(kind) if Keyword.theory(kind) && !Keyword.theory_end(kind) => Some(6)  | 
| 
58868
 
c5e1cce7ace3
uniform heading commands work in any context, even in theory header;
 
wenzelm 
parents: 
58853 
diff
changeset
 | 
248  | 
case _ => None  | 
| 
 
c5e1cce7ace3
uniform heading commands work in any context, even in theory header;
 
wenzelm 
parents: 
58853 
diff
changeset
 | 
249  | 
}  | 
| 58706 | 250  | 
}  | 
251  | 
}  | 
|
252  | 
||
| 58743 | 253  | 
def parse_document(node_name: Document.Node.Name, text: CharSequence):  | 
254  | 
List[Outer_Syntax.Document] =  | 
|
| 58706 | 255  | 
  {
 | 
256  | 
/* stack operations */  | 
|
257  | 
||
258  | 
def buffer(): mutable.ListBuffer[Outer_Syntax.Document] =  | 
|
259  | 
new mutable.ListBuffer[Outer_Syntax.Document]  | 
|
260  | 
||
| 58747 | 261  | 
var stack: List[(Int, Command, mutable.ListBuffer[Outer_Syntax.Document])] =  | 
262  | 
List((0, Command.empty, buffer()))  | 
|
| 55616 | 263  | 
|
| 58706 | 264  | 
@tailrec def close(level: Int => Boolean)  | 
265  | 
    {
 | 
|
266  | 
      stack match {
 | 
|
| 58747 | 267  | 
case (lev, command, body) :: (_, _, body2) :: rest if level(lev) =>  | 
| 59735 | 268  | 
body2 += Outer_Syntax.Document_Block(command.span.name, command.source, body.toList)  | 
| 58706 | 269  | 
stack = stack.tail  | 
270  | 
close(level)  | 
|
271  | 
case _ =>  | 
|
272  | 
}  | 
|
273  | 
}  | 
|
274  | 
||
| 58743 | 275  | 
def result(): List[Outer_Syntax.Document] =  | 
| 58706 | 276  | 
    {
 | 
277  | 
close(_ => true)  | 
|
| 58743 | 278  | 
stack.head._3.toList  | 
| 58706 | 279  | 
}  | 
280  | 
||
281  | 
def add(command: Command)  | 
|
282  | 
    {
 | 
|
283  | 
      heading_level(command) match {
 | 
|
284  | 
case Some(i) =>  | 
|
285  | 
close(_ > i)  | 
|
| 58747 | 286  | 
stack = (i + 1, command, buffer()) :: stack  | 
| 58706 | 287  | 
case None =>  | 
288  | 
}  | 
|
289  | 
stack.head._3 += Outer_Syntax.Document_Atom(command)  | 
|
290  | 
}  | 
|
291  | 
||
292  | 
||
293  | 
/* result structure */  | 
|
294  | 
||
295  | 
val spans = parse_spans(text)  | 
|
| 
59702
 
58dfaa369c11
hybrid use of command blobs: inlined errors and auxiliary files;
 
wenzelm 
parents: 
59700 
diff
changeset
 | 
296  | 
spans.foreach(span => add(Command(Document_ID.none, node_name, Command.no_blobs, span)))  | 
| 58706 | 297  | 
result()  | 
| 55616 | 298  | 
}  | 
| 34166 | 299  | 
}  |