| author | wenzelm | 
| Tue, 05 Jul 2011 22:38:44 +0200 | |
| changeset 43672 | e9f26e66692d | 
| parent 43661 | 39fdbd814c7f | 
| child 43695 | 5130dfe1b7be | 
| permissions | -rw-r--r-- | 
| 28495 | 1  | 
/* Title: Pure/Thy/thy_header.scala  | 
2  | 
Author: Makarius  | 
|
3  | 
||
| 34169 | 4  | 
Theory headers -- independent of outer syntax.  | 
| 28495 | 5  | 
*/  | 
6  | 
||
7  | 
package isabelle  | 
|
8  | 
||
9  | 
||
| 43611 | 10  | 
import scala.annotation.tailrec  | 
| 34169 | 11  | 
import scala.collection.mutable  | 
| 43646 | 12  | 
import scala.util.parsing.input.{Reader, CharSequenceReader}
 | 
| 
38149
 
3c380380beac
somewhat uniform Thy_Header.split_thy_path in ML and Scala;
 
wenzelm 
parents: 
36956 
diff
changeset
 | 
13  | 
import scala.util.matching.Regex  | 
| 34169 | 14  | 
|
| 
34188
 
fbfc18be1f8c
scan: operate on file (via Scan.byte_reader), more robust exception handling;
 
wenzelm 
parents: 
34169 
diff
changeset
 | 
15  | 
import java.io.File  | 
| 
 
fbfc18be1f8c
scan: operate on file (via Scan.byte_reader), more robust exception handling;
 
wenzelm 
parents: 
34169 
diff
changeset
 | 
16  | 
|
| 34169 | 17  | 
|
| 
43661
 
39fdbd814c7f
quasi-static Isabelle_System -- reduced tendency towards "functorial style";
 
wenzelm 
parents: 
43652 
diff
changeset
 | 
18  | 
object Thy_Header extends Parse.Parser  | 
| 32450 | 19  | 
{
 | 
| 28495 | 20  | 
val HEADER = "header"  | 
21  | 
val THEORY = "theory"  | 
|
22  | 
val IMPORTS = "imports"  | 
|
23  | 
val USES = "uses"  | 
|
24  | 
val BEGIN = "begin"  | 
|
25  | 
||
| 34169 | 26  | 
  val lexicon = Scan.Lexicon("%", "(", ")", ";", BEGIN, HEADER, IMPORTS, THEORY, USES)
 | 
| 34190 | 27  | 
|
| 34300 | 28  | 
final case class Header(val name: String, val imports: List[String], val uses: List[String])  | 
| 43611 | 29  | 
  {
 | 
30  | 
def decode_permissive_utf8: Header =  | 
|
31  | 
Header(Standard_System.decode_permissive_utf8(name),  | 
|
32  | 
imports.map(Standard_System.decode_permissive_utf8),  | 
|
33  | 
uses.map(Standard_System.decode_permissive_utf8))  | 
|
34  | 
}  | 
|
| 
38149
 
3c380380beac
somewhat uniform Thy_Header.split_thy_path in ML and Scala;
 
wenzelm 
parents: 
36956 
diff
changeset
 | 
35  | 
|
| 
 
3c380380beac
somewhat uniform Thy_Header.split_thy_path in ML and Scala;
 
wenzelm 
parents: 
36956 
diff
changeset
 | 
36  | 
|
| 
 
3c380380beac
somewhat uniform Thy_Header.split_thy_path in ML and Scala;
 
wenzelm 
parents: 
36956 
diff
changeset
 | 
37  | 
/* file name */  | 
| 
 
3c380380beac
somewhat uniform Thy_Header.split_thy_path in ML and Scala;
 
wenzelm 
parents: 
36956 
diff
changeset
 | 
38  | 
|
| 43648 | 39  | 
  def thy_path(name: String): Path = Path.basic(name).ext("thy")
 | 
40  | 
||
41  | 
  private val Thy_Path1 = new Regex("([^/]*)\\.thy")
 | 
|
42  | 
  private val Thy_Path2 = new Regex("(.*)/([^/]*)\\.thy")
 | 
|
| 
38149
 
3c380380beac
somewhat uniform Thy_Header.split_thy_path in ML and Scala;
 
wenzelm 
parents: 
36956 
diff
changeset
 | 
43  | 
|
| 
39630
 
44181423183a
explicit Session.Phase indication with associated event bus;
 
wenzelm 
parents: 
38149 
diff
changeset
 | 
44  | 
def split_thy_path(path: String): Option[(String, String)] =  | 
| 
38149
 
3c380380beac
somewhat uniform Thy_Header.split_thy_path in ML and Scala;
 
wenzelm 
parents: 
36956 
diff
changeset
 | 
45  | 
    path match {
 | 
| 
39630
 
44181423183a
explicit Session.Phase indication with associated event bus;
 
wenzelm 
parents: 
38149 
diff
changeset
 | 
46  | 
      case Thy_Path1(name) => Some(("", name))
 | 
| 
 
44181423183a
explicit Session.Phase indication with associated event bus;
 
wenzelm 
parents: 
38149 
diff
changeset
 | 
47  | 
case Thy_Path2(dir, name) => Some((dir, name))  | 
| 
 
44181423183a
explicit Session.Phase indication with associated event bus;
 
wenzelm 
parents: 
38149 
diff
changeset
 | 
48  | 
case _ => None  | 
| 
38149
 
3c380380beac
somewhat uniform Thy_Header.split_thy_path in ML and Scala;
 
wenzelm 
parents: 
36956 
diff
changeset
 | 
49  | 
}  | 
| 34169 | 50  | 
|
51  | 
||
52  | 
/* header */  | 
|
53  | 
||
| 34190 | 54  | 
val header: Parser[Header] =  | 
| 34169 | 55  | 
  {
 | 
| 43611 | 56  | 
    val file_name = atom("file name", _.is_name)
 | 
57  | 
    val theory_name = atom("theory name", _.is_name)
 | 
|
| 34169 | 58  | 
|
59  | 
val file =  | 
|
60  | 
      keyword("(") ~! (file_name ~ keyword(")")) ^^ { case _ ~ (x ~ _) => x } | file_name
 | 
|
61  | 
||
62  | 
    val uses = opt(keyword(USES) ~! (rep1(file))) ^^ { case None => Nil case Some(_ ~ xs) => xs }
 | 
|
63  | 
||
64  | 
val args =  | 
|
65  | 
theory_name ~ (keyword(IMPORTS) ~! (rep1(theory_name) ~ uses ~ keyword(BEGIN))) ^^  | 
|
| 34190 | 66  | 
        { case x ~ (_ ~ (ys ~ zs ~ _)) => Header(x, ys, zs) }
 | 
| 34169 | 67  | 
|
68  | 
(keyword(HEADER) ~ tags) ~!  | 
|
69  | 
      ((doc_source ~ rep(keyword(";")) ~ keyword(THEORY) ~ tags) ~> args) ^^ { case _ ~ x => x } |
 | 
|
70  | 
    (keyword(THEORY) ~ tags) ~! args ^^ { case _ ~ x => x }
 | 
|
71  | 
}  | 
|
72  | 
||
73  | 
||
| 34190 | 74  | 
/* read -- lazy scanning */  | 
| 34169 | 75  | 
|
| 43611 | 76  | 
def read(reader: Reader[Char]): Header =  | 
| 34169 | 77  | 
  {
 | 
| 
43661
 
39fdbd814c7f
quasi-static Isabelle_System -- reduced tendency towards "functorial style";
 
wenzelm 
parents: 
43652 
diff
changeset
 | 
78  | 
val token = lexicon.token(Isabelle_System.symbols, _ => false)  | 
| 
36956
 
21be4832c362
renamed class Outer_Lex to Token and Token_Kind to Token.Kind;
 
wenzelm 
parents: 
36948 
diff
changeset
 | 
79  | 
val toks = new mutable.ListBuffer[Token]  | 
| 
34188
 
fbfc18be1f8c
scan: operate on file (via Scan.byte_reader), more robust exception handling;
 
wenzelm 
parents: 
34169 
diff
changeset
 | 
80  | 
|
| 43611 | 81  | 
@tailrec def scan_to_begin(in: Reader[Char])  | 
| 34169 | 82  | 
    {
 | 
83  | 
      token(in) match {
 | 
|
84  | 
case lexicon.Success(tok, rest) =>  | 
|
85  | 
toks += tok  | 
|
| 43611 | 86  | 
if (!tok.is_begin) scan_to_begin(rest)  | 
| 34169 | 87  | 
case _ =>  | 
88  | 
}  | 
|
89  | 
}  | 
|
| 43611 | 90  | 
scan_to_begin(reader)  | 
| 34190 | 91  | 
|
| 
36956
 
21be4832c362
renamed class Outer_Lex to Token and Token_Kind to Token.Kind;
 
wenzelm 
parents: 
36948 
diff
changeset
 | 
92  | 
    parse(commit(header), Token.reader(toks.toList)) match {
 | 
| 34190 | 93  | 
case Success(result, _) => result  | 
94  | 
case bad => error(bad.toString)  | 
|
95  | 
}  | 
|
| 34169 | 96  | 
}  | 
| 43611 | 97  | 
|
| 43646 | 98  | 
def read(source: CharSequence): Header =  | 
99  | 
read(new CharSequenceReader(source))  | 
|
100  | 
||
| 43611 | 101  | 
def read(file: File): Header =  | 
102  | 
  {
 | 
|
103  | 
val reader = Scan.byte_reader(file)  | 
|
104  | 
    try { read(reader).decode_permissive_utf8 }
 | 
|
105  | 
    finally { reader.close }
 | 
|
106  | 
}  | 
|
| 43648 | 107  | 
|
108  | 
||
109  | 
/* check */  | 
|
110  | 
||
111  | 
def check(name: String, source: CharSequence): Header =  | 
|
112  | 
  {
 | 
|
113  | 
val header = read(source)  | 
|
114  | 
val name1 = header.name  | 
|
| 43672 | 115  | 
    if (name != name1) error("Bad file name " + thy_path(name) + " for theory " + quote(name1))
 | 
116  | 
Path.explode(name)  | 
|
117  | 
header  | 
|
| 43648 | 118  | 
}  | 
| 34169 | 119  | 
}  |