| author | wenzelm | 
| Tue, 07 Jun 2022 12:32:53 +0200 | |
| changeset 75523 | 0dcaf0e5107b | 
| parent 75407 | c7051638a38c | 
| child 75906 | 2167b9e3157a | 
| permissions | -rw-r--r-- | 
| 28495 | 1  | 
/* Title: Pure/Thy/thy_header.scala  | 
2  | 
Author: Makarius  | 
|
3  | 
||
| 46939 | 4  | 
Static theory header information.  | 
| 28495 | 5  | 
*/  | 
6  | 
||
7  | 
package isabelle  | 
|
8  | 
||
9  | 
||
| 64824 | 10  | 
import scala.util.parsing.input.Reader  | 
| 
38149
 
3c380380beac
somewhat uniform Thy_Header.split_thy_path in ML and Scala;
 
wenzelm 
parents: 
36956 
diff
changeset
 | 
11  | 
import scala.util.matching.Regex  | 
| 34169 | 12  | 
|
13  | 
||
| 75393 | 14  | 
object Thy_Header {
 | 
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58908 
diff
changeset
 | 
15  | 
/* bootstrap keywords */  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58908 
diff
changeset
 | 
16  | 
|
| 63579 | 17  | 
type Keywords = List[(String, Keyword.Spec)]  | 
18  | 
type Abbrevs = List[(String, String)]  | 
|
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58908 
diff
changeset
 | 
19  | 
|
| 
58868
 
c5e1cce7ace3
uniform heading commands work in any context, even in theory header;
 
wenzelm 
parents: 
58861 
diff
changeset
 | 
20  | 
val CHAPTER = "chapter"  | 
| 
 
c5e1cce7ace3
uniform heading commands work in any context, even in theory header;
 
wenzelm 
parents: 
58861 
diff
changeset
 | 
21  | 
val SECTION = "section"  | 
| 
 
c5e1cce7ace3
uniform heading commands work in any context, even in theory header;
 
wenzelm 
parents: 
58861 
diff
changeset
 | 
22  | 
val SUBSECTION = "subsection"  | 
| 
 
c5e1cce7ace3
uniform heading commands work in any context, even in theory header;
 
wenzelm 
parents: 
58861 
diff
changeset
 | 
23  | 
val SUBSUBSECTION = "subsubsection"  | 
| 61463 | 24  | 
val PARAGRAPH = "paragraph"  | 
25  | 
val SUBPARAGRAPH = "subparagraph"  | 
|
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58928 
diff
changeset
 | 
26  | 
val TEXT = "text"  | 
| 
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58928 
diff
changeset
 | 
27  | 
val TXT = "txt"  | 
| 
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58928 
diff
changeset
 | 
28  | 
val TEXT_RAW = "text_raw"  | 
| 
58868
 
c5e1cce7ace3
uniform heading commands work in any context, even in theory header;
 
wenzelm 
parents: 
58861 
diff
changeset
 | 
29  | 
|
| 28495 | 30  | 
val THEORY = "theory"  | 
31  | 
val IMPORTS = "imports"  | 
|
| 
46938
 
cda018294515
some support for outer syntax keyword declarations within theory header;
 
wenzelm 
parents: 
46737 
diff
changeset
 | 
32  | 
val KEYWORDS = "keywords"  | 
| 63579 | 33  | 
val ABBREVS = "abbrevs"  | 
| 
46938
 
cda018294515
some support for outer syntax keyword declarations within theory header;
 
wenzelm 
parents: 
46737 
diff
changeset
 | 
34  | 
val AND = "and"  | 
| 28495 | 35  | 
val BEGIN = "begin"  | 
36  | 
||
| 64854 | 37  | 
val bootstrap_header: Keywords =  | 
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58908 
diff
changeset
 | 
38  | 
List(  | 
| 72764 | 39  | 
      ("%", Keyword.Spec()),
 | 
40  | 
      ("(", Keyword.Spec()),
 | 
|
41  | 
      (")", Keyword.Spec()),
 | 
|
42  | 
      (",", Keyword.Spec()),
 | 
|
43  | 
      ("::", Keyword.Spec()),
 | 
|
44  | 
      ("=", Keyword.Spec()),
 | 
|
45  | 
(AND, Keyword.Spec()),  | 
|
46  | 
(BEGIN, Keyword.Spec(kind = Keyword.QUASI_COMMAND)),  | 
|
47  | 
(IMPORTS, Keyword.Spec(kind = Keyword.QUASI_COMMAND)),  | 
|
48  | 
(KEYWORDS, Keyword.Spec(kind = Keyword.QUASI_COMMAND)),  | 
|
49  | 
(ABBREVS, Keyword.Spec(kind = Keyword.QUASI_COMMAND)),  | 
|
50  | 
(CHAPTER, Keyword.Spec(kind = Keyword.DOCUMENT_HEADING)),  | 
|
51  | 
(SECTION, Keyword.Spec(kind = Keyword.DOCUMENT_HEADING)),  | 
|
52  | 
(SUBSECTION, Keyword.Spec(kind = Keyword.DOCUMENT_HEADING)),  | 
|
53  | 
(SUBSUBSECTION, Keyword.Spec(kind = Keyword.DOCUMENT_HEADING)),  | 
|
54  | 
(PARAGRAPH, Keyword.Spec(kind = Keyword.DOCUMENT_HEADING)),  | 
|
55  | 
(SUBPARAGRAPH, Keyword.Spec(kind = Keyword.DOCUMENT_HEADING)),  | 
|
56  | 
(TEXT, Keyword.Spec(kind = Keyword.DOCUMENT_BODY)),  | 
|
57  | 
(TXT, Keyword.Spec(kind = Keyword.DOCUMENT_BODY)),  | 
|
58  | 
(TEXT_RAW, Keyword.Spec(kind = Keyword.DOCUMENT_RAW)),  | 
|
59  | 
      (THEORY, Keyword.Spec(kind = Keyword.THY_BEGIN, tags = List("theory"))),
 | 
|
60  | 
      ("ML", Keyword.Spec(kind = Keyword.THY_DECL, tags = List("ML"))))
 | 
|
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58908 
diff
changeset
 | 
61  | 
|
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58908 
diff
changeset
 | 
62  | 
private val bootstrap_keywords =  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58908 
diff
changeset
 | 
63  | 
Keyword.Keywords.empty.add_keywords(bootstrap_header)  | 
| 
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58908 
diff
changeset
 | 
64  | 
|
| 
67004
 
af72fa58f71b
clarified lazy Completion within Outer_Syntax: measurable speedup of Sessions.deps;
 
wenzelm 
parents: 
66984 
diff
changeset
 | 
65  | 
val bootstrap_syntax: Outer_Syntax =  | 
| 
 
af72fa58f71b
clarified lazy Completion within Outer_Syntax: measurable speedup of Sessions.deps;
 
wenzelm 
parents: 
66984 
diff
changeset
 | 
66  | 
Outer_Syntax.empty.add_keywords(bootstrap_header)  | 
| 34190 | 67  | 
|
| 
38149
 
3c380380beac
somewhat uniform Thy_Header.split_thy_path in ML and Scala;
 
wenzelm 
parents: 
36956 
diff
changeset
 | 
68  | 
|
| 
64673
 
b5965890e54d
more uniform treatment of file name vs. theory name and special header;
 
wenzelm 
parents: 
64671 
diff
changeset
 | 
69  | 
/* file name vs. theory name */  | 
| 
62895
 
54c2abe7e9a4
treat ROOT.ML as theory with header "theory ML_Root imports ML_Bootstrap begin";
 
wenzelm 
parents: 
62849 
diff
changeset
 | 
70  | 
|
| 63022 | 71  | 
val PURE = "Pure"  | 
| 
62895
 
54c2abe7e9a4
treat ROOT.ML as theory with header "theory ML_Root imports ML_Bootstrap begin";
 
wenzelm 
parents: 
62849 
diff
changeset
 | 
72  | 
val ML_BOOTSTRAP = "ML_Bootstrap"  | 
| 65490 | 73  | 
  val ml_roots = List("ROOT0.ML" -> "ML_Root0", "ROOT.ML" -> "ML_Root")
 | 
| 63022 | 74  | 
  val bootstrap_thys = List(PURE, ML_BOOTSTRAP).map(a => a -> ("Bootstrap_" + a))
 | 
| 
44160
 
8848867501fb
clarified document model header: master_dir (native wrt. editor, potentially URL) and node_name (full canonical path);
 
wenzelm 
parents: 
44159 
diff
changeset
 | 
75  | 
|
| 67215 | 76  | 
val bootstrap_global_theories =  | 
77  | 
(Sessions.root_name :: (ml_roots ::: bootstrap_thys).map(_._2)).map(_ -> PURE)  | 
|
| 65490 | 78  | 
|
| 65452 | 79  | 
  private val Thy_File_Name = new Regex(""".*?([^/\\:]+)\.thy""")
 | 
| 
69255
 
800b1ce96fce
more general support for Isabelle/PIDE file formats -- less hardwired Bibtex operations;
 
wenzelm 
parents: 
68841 
diff
changeset
 | 
80  | 
  private val Split_File_Name = new Regex("""(.*?)[/\\]*([^/\\:]+)""")
 | 
| 67212 | 81  | 
  private val File_Name = new Regex(""".*?([^/\\:]+)""")
 | 
| 
44160
 
8848867501fb
clarified document model header: master_dir (native wrt. editor, potentially URL) and node_name (full canonical path);
 
wenzelm 
parents: 
44159 
diff
changeset
 | 
82  | 
|
| 65526 | 83  | 
def is_base_name(s: String): Boolean =  | 
84  | 
    s != "" && !s.exists("/\\:".contains(_))
 | 
|
85  | 
||
| 
69255
 
800b1ce96fce
more general support for Isabelle/PIDE file formats -- less hardwired Bibtex operations;
 
wenzelm 
parents: 
68841 
diff
changeset
 | 
86  | 
def split_file_name(s: String): Option[(String, String)] =  | 
| 67290 | 87  | 
    s match {
 | 
| 
69255
 
800b1ce96fce
more general support for Isabelle/PIDE file formats -- less hardwired Bibtex operations;
 
wenzelm 
parents: 
68841 
diff
changeset
 | 
88  | 
case Split_File_Name(s1, s2) => Some((s1, s2))  | 
| 67290 | 89  | 
case _ => None  | 
90  | 
}  | 
|
91  | 
||
| 65452 | 92  | 
def import_name(s: String): String =  | 
| 67164 | 93  | 
    s match {
 | 
| 67212 | 94  | 
      case File_Name(name) if !name.endsWith(".thy") => name
 | 
| 67164 | 95  | 
      case _ => error("Malformed theory import: " + quote(s))
 | 
96  | 
}  | 
|
| 
44225
 
a8f921e6484f
more robust Thy_Header.base_name, with minimal assumptions about path syntax;
 
wenzelm 
parents: 
44222 
diff
changeset
 | 
97  | 
|
| 65452 | 98  | 
def theory_name(s: String): String =  | 
| 
62895
 
54c2abe7e9a4
treat ROOT.ML as theory with header "theory ML_Root imports ML_Bootstrap begin";
 
wenzelm 
parents: 
62849 
diff
changeset
 | 
99  | 
    s match {
 | 
| 66195 | 100  | 
case Thy_File_Name(name) => bootstrap_name(name)  | 
| 67212 | 101  | 
case File_Name(name) =>  | 
| 67215 | 102  | 
if (name == Sessions.root_name) name  | 
103  | 
        else ml_roots.collectFirst({ case (a, b) if a == name => b }).getOrElse("")
 | 
|
| 65452 | 104  | 
case _ => ""  | 
| 
62895
 
54c2abe7e9a4
treat ROOT.ML as theory with header "theory ML_Root imports ML_Bootstrap begin";
 
wenzelm 
parents: 
62849 
diff
changeset
 | 
105  | 
}  | 
| 
38149
 
3c380380beac
somewhat uniform Thy_Header.split_thy_path in ML and Scala;
 
wenzelm 
parents: 
36956 
diff
changeset
 | 
106  | 
|
| 
64673
 
b5965890e54d
more uniform treatment of file name vs. theory name and special header;
 
wenzelm 
parents: 
64671 
diff
changeset
 | 
107  | 
def is_ml_root(theory: String): Boolean =  | 
| 
 
b5965890e54d
more uniform treatment of file name vs. theory name and special header;
 
wenzelm 
parents: 
64671 
diff
changeset
 | 
108  | 
    ml_roots.exists({ case (_, b) => b == theory })
 | 
| 
 
b5965890e54d
more uniform treatment of file name vs. theory name and special header;
 
wenzelm 
parents: 
64671 
diff
changeset
 | 
109  | 
|
| 
 
b5965890e54d
more uniform treatment of file name vs. theory name and special header;
 
wenzelm 
parents: 
64671 
diff
changeset
 | 
110  | 
def is_bootstrap(theory: String): Boolean =  | 
| 
 
b5965890e54d
more uniform treatment of file name vs. theory name and special header;
 
wenzelm 
parents: 
64671 
diff
changeset
 | 
111  | 
    bootstrap_thys.exists({ case (_, b) => b == theory })
 | 
| 
 
b5965890e54d
more uniform treatment of file name vs. theory name and special header;
 
wenzelm 
parents: 
64671 
diff
changeset
 | 
112  | 
|
| 66195 | 113  | 
def bootstrap_name(theory: String): String =  | 
114  | 
    bootstrap_thys.collectFirst({ case (a, b) if a == theory => b }).getOrElse(theory)
 | 
|
115  | 
||
| 75393 | 116  | 
  def try_read_dir(dir: Path): List[String] = {
 | 
| 
70713
 
fd188463066e
clarified theory imports completion, based on session directories and current master directory (no support for local session-subdirectories);
 
wenzelm 
parents: 
70638 
diff
changeset
 | 
117  | 
val files =  | 
| 
 
fd188463066e
clarified theory imports completion, based on session directories and current master directory (no support for local session-subdirectories);
 
wenzelm 
parents: 
70638 
diff
changeset
 | 
118  | 
      try { File.read_dir(dir) }
 | 
| 
 
fd188463066e
clarified theory imports completion, based on session directories and current master directory (no support for local session-subdirectories);
 
wenzelm 
parents: 
70638 
diff
changeset
 | 
119  | 
      catch { case ERROR(_) => Nil }
 | 
| 
 
fd188463066e
clarified theory imports completion, based on session directories and current master directory (no support for local session-subdirectories);
 
wenzelm 
parents: 
70638 
diff
changeset
 | 
120  | 
    for { Thy_File_Name(name) <- files } yield name
 | 
| 
 
fd188463066e
clarified theory imports completion, based on session directories and current master directory (no support for local session-subdirectories);
 
wenzelm 
parents: 
70638 
diff
changeset
 | 
121  | 
}  | 
| 
 
fd188463066e
clarified theory imports completion, based on session directories and current master directory (no support for local session-subdirectories);
 
wenzelm 
parents: 
70638 
diff
changeset
 | 
122  | 
|
| 34169 | 123  | 
|
| 68841 | 124  | 
/* parser */  | 
| 34169 | 125  | 
|
| 75405 | 126  | 
  trait Parsers extends Parse.Parsers {
 | 
| 75393 | 127  | 
    val header: Parser[Thy_Header] = {
 | 
| 72765 | 128  | 
val load_command =  | 
129  | 
        ($$$("(") ~! (position(name) <~ $$$(")")) ^^ { case _ ~ x => x }) |
 | 
|
130  | 
          success(("", Position.none))
 | 
|
| 
72748
 
04d5f6d769a7
more flexible syntax for theory load commands via Isabelle/Scala;
 
wenzelm 
parents: 
72747 
diff
changeset
 | 
131  | 
|
| 72765 | 132  | 
      val keyword_kind = atom("outer syntax keyword specification", _.is_name)
 | 
| 68841 | 133  | 
val keyword_spec =  | 
| 72765 | 134  | 
position(keyword_kind) ~ load_command ~ tags ^^  | 
135  | 
          { case (a, b) ~ c ~ d =>
 | 
|
136  | 
Keyword.Spec(kind = a, kind_pos = b,  | 
|
137  | 
load_command = c._1, load_command_pos = c._2, tags = d)  | 
|
138  | 
}  | 
|
| 59694 | 139  | 
|
| 68841 | 140  | 
val keyword_decl =  | 
141  | 
rep1(string) ~  | 
|
142  | 
        opt($$$("::") ~! keyword_spec ^^ { case _ ~ x => x }) ^^
 | 
|
| 72764 | 143  | 
        { case xs ~ y => xs.map((_, y.getOrElse(Keyword.Spec()))) }
 | 
| 
46938
 
cda018294515
some support for outer syntax keyword declarations within theory header;
 
wenzelm 
parents: 
46737 
diff
changeset
 | 
144  | 
|
| 68841 | 145  | 
val keyword_decls =  | 
146  | 
        keyword_decl ~ rep($$$(AND) ~! keyword_decl ^^ { case _ ~ x => x }) ^^
 | 
|
147  | 
        { case xs ~ yss => (xs :: yss).flatten }
 | 
|
148  | 
||
149  | 
val abbrevs =  | 
|
150  | 
        rep1sep(rep1(text) ~ ($$$("=") ~! rep1(text)), $$$("and")) ^^
 | 
|
151  | 
          { case res => for ((as ~ (_ ~ bs)) <- res; a <- as; b <- bs) yield (a, b) }
 | 
|
| 63579 | 152  | 
|
| 68841 | 153  | 
val args =  | 
| 75407 | 154  | 
position(this.theory_name) ~  | 
155  | 
(opt($$$(IMPORTS) ~! rep1(position(this.theory_name))) ^^  | 
|
| 68841 | 156  | 
          { case None => Nil case Some(_ ~ xs) => xs }) ~
 | 
157  | 
(opt($$$(KEYWORDS) ~! keyword_decls) ^^  | 
|
158  | 
          { case None => Nil case Some(_ ~ xs) => xs }) ~
 | 
|
159  | 
(opt($$$(ABBREVS) ~! abbrevs) ^^  | 
|
160  | 
          { case None => Nil case Some(_ ~ xs) => xs }) ~
 | 
|
| 72779 | 161  | 
        $$$(BEGIN) ^^ { case a ~ b ~ c ~ d ~ _ => Thy_Header(a._1, a._2, b, c, d) }
 | 
| 34169 | 162  | 
|
| 68841 | 163  | 
val heading =  | 
164  | 
(command(CHAPTER) |  | 
|
165  | 
command(SECTION) |  | 
|
166  | 
command(SUBSECTION) |  | 
|
167  | 
command(SUBSUBSECTION) |  | 
|
168  | 
command(PARAGRAPH) |  | 
|
169  | 
command(SUBPARAGRAPH) |  | 
|
170  | 
command(TEXT) |  | 
|
171  | 
command(TXT) |  | 
|
172  | 
command(TEXT_RAW)) ~  | 
|
| 69887 | 173  | 
annotation ~! document_source  | 
| 
58868
 
c5e1cce7ace3
uniform heading commands work in any context, even in theory header;
 
wenzelm 
parents: 
58861 
diff
changeset
 | 
174  | 
|
| 69887 | 175  | 
      (rep(heading) ~ command(THEORY) ~ annotation) ~! args ^^ { case _ ~ x => x }
 | 
| 68841 | 176  | 
}  | 
| 34169 | 177  | 
}  | 
178  | 
||
179  | 
||
| 34190 | 180  | 
/* read -- lazy scanning */  | 
| 34169 | 181  | 
|
| 75393 | 182  | 
  private def read_tokens(reader: Reader[Char], strict: Boolean): (List[Token], List[Token]) = {
 | 
| 
58928
 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 
wenzelm 
parents: 
58908 
diff
changeset
 | 
183  | 
val token = Token.Parsers.token(bootstrap_keywords)  | 
| 73368 | 184  | 
def make_tokens(in: Reader[Char]): LazyList[Token] =  | 
| 34169 | 185  | 
      token(in) match {
 | 
| 64825 | 186  | 
case Token.Parsers.Success(tok, rest) => tok #:: make_tokens(rest)  | 
| 73368 | 187  | 
case _ => LazyList.empty  | 
| 34169 | 188  | 
}  | 
| 64825 | 189  | 
|
| 65539 | 190  | 
val all_tokens = make_tokens(reader)  | 
191  | 
val drop_tokens =  | 
|
192  | 
if (strict) Nil  | 
|
193  | 
else all_tokens.takeWhile(tok => !tok.is_command(Thy_Header.THEORY)).toList  | 
|
| 34190 | 194  | 
|
| 65539 | 195  | 
val tokens = all_tokens.drop(drop_tokens.length)  | 
| 64825 | 196  | 
val tokens1 = tokens.takeWhile(tok => !tok.is_begin).toList  | 
197  | 
val tokens2 = tokens.dropWhile(tok => !tok.is_begin).headOption.toList  | 
|
198  | 
||
| 65539 | 199  | 
(drop_tokens, tokens1 ::: tokens2)  | 
200  | 
}  | 
|
201  | 
||
| 75405 | 202  | 
  private object Parsers extends Parsers {
 | 
| 68841 | 203  | 
def parse_header(tokens: List[Token], pos: Token.Pos): Thy_Header =  | 
204  | 
      parse(commit(header), Token.reader(tokens, pos)) match {
 | 
|
205  | 
case Success(result, _) => result  | 
|
206  | 
case bad => error(bad.toString)  | 
|
207  | 
}  | 
|
208  | 
}  | 
|
209  | 
||
| 72775 | 210  | 
def read(node_name: Document.Node.Name, reader: Reader[Char],  | 
| 
72777
 
164cb0806d0a
proper positions for inlined command messages, e.g. for completion within theory header;
 
wenzelm 
parents: 
72776 
diff
changeset
 | 
211  | 
command: Boolean = true,  | 
| 75393 | 212  | 
strict: Boolean = true  | 
213  | 
  ): Thy_Header = {
 | 
|
| 65539 | 214  | 
val (_, tokens0) = read_tokens(reader, true)  | 
| 
66918
 
ec2b50aeb0dd
more robust treatment of UTF8 in raw byte sources;
 
wenzelm 
parents: 
66713 
diff
changeset
 | 
215  | 
val text = Scan.reader_decode_utf8(reader, Token.implode(tokens0))  | 
| 65539 | 216  | 
|
| 
72777
 
164cb0806d0a
proper positions for inlined command messages, e.g. for completion within theory header;
 
wenzelm 
parents: 
72776 
diff
changeset
 | 
217  | 
val (skip_tokens, tokens) = read_tokens(Scan.char_reader(text), strict)  | 
| 
 
164cb0806d0a
proper positions for inlined command messages, e.g. for completion within theory header;
 
wenzelm 
parents: 
72776 
diff
changeset
 | 
218  | 
val pos =  | 
| 
 
164cb0806d0a
proper positions for inlined command messages, e.g. for completion within theory header;
 
wenzelm 
parents: 
72776 
diff
changeset
 | 
219  | 
if (command) Token.Pos.command  | 
| 73359 | 220  | 
else skip_tokens.foldLeft(Token.Pos.file(node_name.node))(_ advance _)  | 
| 65539 | 221  | 
|
| 75405 | 222  | 
Parsers.parse_header(tokens, pos).map(Symbol.decode).check(node_name)  | 
| 34169 | 223  | 
}  | 
224  | 
}  | 
|
| 44159 | 225  | 
|
| 44185 | 226  | 
sealed case class Thy_Header(  | 
| 72778 | 227  | 
name: String,  | 
228  | 
pos: Position.T,  | 
|
229  | 
imports: List[(String, Position.T)],  | 
|
| 63579 | 230  | 
keywords: Thy_Header.Keywords,  | 
| 75393 | 231  | 
abbrevs: Thy_Header.Abbrevs  | 
232  | 
) {
 | 
|
| 72764 | 233  | 
def map(f: String => String): Thy_Header =  | 
| 72778 | 234  | 
Thy_Header(f(name), pos,  | 
235  | 
      imports.map({ case (a, b) => (f(a), b) }),
 | 
|
| 72764 | 236  | 
      keywords.map({ case (a, spec) => (f(a), spec.map(f)) }),
 | 
237  | 
      abbrevs.map({ case (a, b) => (f(a), f(b)) }))
 | 
|
| 72765 | 238  | 
|
| 75393 | 239  | 
  def check(node_name: Document.Node.Name): Thy_Header = {
 | 
| 72775 | 240  | 
val base_name = node_name.theory_base_name  | 
241  | 
    if (Long_Name.is_qualified(name)) {
 | 
|
242  | 
      error("Bad theory name " + quote(name) + " with qualification" + Position.here(pos))
 | 
|
243  | 
}  | 
|
244  | 
    if (base_name != name) {
 | 
|
| 72776 | 245  | 
      error("Bad theory name " + quote(name) + " for file " + Path.basic(base_name).thy +
 | 
246  | 
Position.here(pos) + Completion.report_theories(pos, List(base_name)))  | 
|
| 72775 | 247  | 
}  | 
248  | 
||
| 72765 | 249  | 
    for ((_, spec) <- keywords) {
 | 
250  | 
      if (spec.kind != Keyword.THY_LOAD && spec.load_command.nonEmpty) {
 | 
|
251  | 
        error("Illegal load command specification for kind: " + quote(spec.kind) +
 | 
|
252  | 
Position.here(spec.kind_pos))  | 
|
253  | 
}  | 
|
254  | 
      if (!Command_Span.load_commands.exists(_.name == spec.load_command)) {
 | 
|
| 72782 | 255  | 
val completion =  | 
256  | 
          for {
 | 
|
257  | 
load_command <- Command_Span.load_commands  | 
|
258  | 
name = load_command.name  | 
|
259  | 
if name.startsWith(spec.load_command)  | 
|
260  | 
} yield (name, (Markup.LOAD_COMMAND, name))  | 
|
| 72765 | 261  | 
        error("Unknown load command specification: " + quote(spec.load_command) +
 | 
| 72782 | 262  | 
Position.here(spec.load_command_pos) +  | 
263  | 
Completion.report_names(spec.load_command_pos, completion))  | 
|
| 72765 | 264  | 
}  | 
265  | 
}  | 
|
266  | 
this  | 
|
267  | 
}  | 
|
| 
70638
 
f164cec7ac22
clarified signature: prefer operations without position;
 
wenzelm 
parents: 
69887 
diff
changeset
 | 
268  | 
}  |