| author | wenzelm | 
| Tue, 16 Aug 2011 22:48:31 +0200 | |
| changeset 44225 | a8f921e6484f | 
| parent 44222 | 9d5ef6cd4ee1 | 
| child 44578 | ca3844a3dcf7 | 
| 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  | 
|
| 
38149
 
3c380380beac
somewhat uniform Thy_Header.split_thy_path in ML and Scala;
 
wenzelm 
parents: 
36956 
diff
changeset
 | 
28  | 
|
| 
44160
 
8848867501fb
clarified document model header: master_dir (native wrt. editor, potentially URL) and node_name (full canonical path);
 
wenzelm 
parents: 
44159 
diff
changeset
 | 
29  | 
/* theory file name */  | 
| 
 
8848867501fb
clarified document model header: master_dir (native wrt. editor, potentially URL) and node_name (full canonical path);
 
wenzelm 
parents: 
44159 
diff
changeset
 | 
30  | 
|
| 
44225
 
a8f921e6484f
more robust Thy_Header.base_name, with minimal assumptions about path syntax;
 
wenzelm 
parents: 
44222 
diff
changeset
 | 
31  | 
  private val Base_Name = new Regex(""".*?([^/\\:]+)""")
 | 
| 
44222
 
9d5ef6cd4ee1
use full .thy file name as node name, which makes MiscUtilities.resolveSymlinks/File.getCanonicalPath more predictable;
 
wenzelm 
parents: 
44185 
diff
changeset
 | 
32  | 
  private val Thy_Name = new Regex(""".*?([^/\\:]+)\.thy""")
 | 
| 
44160
 
8848867501fb
clarified document model header: master_dir (native wrt. editor, potentially URL) and node_name (full canonical path);
 
wenzelm 
parents: 
44159 
diff
changeset
 | 
33  | 
|
| 
44225
 
a8f921e6484f
more robust Thy_Header.base_name, with minimal assumptions about path syntax;
 
wenzelm 
parents: 
44222 
diff
changeset
 | 
34  | 
def base_name(s: String): String =  | 
| 
 
a8f921e6484f
more robust Thy_Header.base_name, with minimal assumptions about path syntax;
 
wenzelm 
parents: 
44222 
diff
changeset
 | 
35  | 
    s match { case Base_Name(name) => name case _ => error("Malformed import: " + quote(s)) }
 | 
| 
 
a8f921e6484f
more robust Thy_Header.base_name, with minimal assumptions about path syntax;
 
wenzelm 
parents: 
44222 
diff
changeset
 | 
36  | 
|
| 
44222
 
9d5ef6cd4ee1
use full .thy file name as node name, which makes MiscUtilities.resolveSymlinks/File.getCanonicalPath more predictable;
 
wenzelm 
parents: 
44185 
diff
changeset
 | 
37  | 
def thy_name(s: String): Option[String] =  | 
| 
 
9d5ef6cd4ee1
use full .thy file name as node name, which makes MiscUtilities.resolveSymlinks/File.getCanonicalPath more predictable;
 
wenzelm 
parents: 
44185 
diff
changeset
 | 
38  | 
    s match { case Thy_Name(name) => Some(name) case _ => None }
 | 
| 
38149
 
3c380380beac
somewhat uniform Thy_Header.split_thy_path in ML and Scala;
 
wenzelm 
parents: 
36956 
diff
changeset
 | 
39  | 
|
| 
44222
 
9d5ef6cd4ee1
use full .thy file name as node name, which makes MiscUtilities.resolveSymlinks/File.getCanonicalPath more predictable;
 
wenzelm 
parents: 
44185 
diff
changeset
 | 
40  | 
  def thy_path(path: Path): Path = path.ext("thy")
 | 
| 43648 | 41  | 
  def thy_path(name: String): Path = Path.basic(name).ext("thy")
 | 
42  | 
||
| 34169 | 43  | 
|
44  | 
/* header */  | 
|
45  | 
||
| 44159 | 46  | 
val header: Parser[Thy_Header] =  | 
| 34169 | 47  | 
  {
 | 
| 43611 | 48  | 
    val file_name = atom("file name", _.is_name)
 | 
49  | 
    val theory_name = atom("theory name", _.is_name)
 | 
|
| 34169 | 50  | 
|
51  | 
val file =  | 
|
| 44185 | 52  | 
      keyword("(") ~! (file_name ~ keyword(")")) ^^ { case _ ~ (x ~ _) => (x, false) } |
 | 
53  | 
file_name ^^ (x => (x, true))  | 
|
| 34169 | 54  | 
|
55  | 
    val uses = opt(keyword(USES) ~! (rep1(file))) ^^ { case None => Nil case Some(_ ~ xs) => xs }
 | 
|
56  | 
||
57  | 
val args =  | 
|
58  | 
theory_name ~ (keyword(IMPORTS) ~! (rep1(theory_name) ~ uses ~ keyword(BEGIN))) ^^  | 
|
| 44159 | 59  | 
        { case x ~ (_ ~ (ys ~ zs ~ _)) => Thy_Header(x, ys, zs) }
 | 
| 34169 | 60  | 
|
61  | 
(keyword(HEADER) ~ tags) ~!  | 
|
62  | 
      ((doc_source ~ rep(keyword(";")) ~ keyword(THEORY) ~ tags) ~> args) ^^ { case _ ~ x => x } |
 | 
|
63  | 
    (keyword(THEORY) ~ tags) ~! args ^^ { case _ ~ x => x }
 | 
|
64  | 
}  | 
|
65  | 
||
66  | 
||
| 34190 | 67  | 
/* read -- lazy scanning */  | 
| 34169 | 68  | 
|
| 44159 | 69  | 
def read(reader: Reader[Char]): Thy_Header =  | 
| 34169 | 70  | 
  {
 | 
| 
43695
 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 
wenzelm 
parents: 
43672 
diff
changeset
 | 
71  | 
val token = lexicon.token(_ => false)  | 
| 
36956
 
21be4832c362
renamed class Outer_Lex to Token and Token_Kind to Token.Kind;
 
wenzelm 
parents: 
36948 
diff
changeset
 | 
72  | 
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
 | 
73  | 
|
| 43611 | 74  | 
@tailrec def scan_to_begin(in: Reader[Char])  | 
| 34169 | 75  | 
    {
 | 
76  | 
      token(in) match {
 | 
|
77  | 
case lexicon.Success(tok, rest) =>  | 
|
78  | 
toks += tok  | 
|
| 43611 | 79  | 
if (!tok.is_begin) scan_to_begin(rest)  | 
| 34169 | 80  | 
case _ =>  | 
81  | 
}  | 
|
82  | 
}  | 
|
| 43611 | 83  | 
scan_to_begin(reader)  | 
| 34190 | 84  | 
|
| 
36956
 
21be4832c362
renamed class Outer_Lex to Token and Token_Kind to Token.Kind;
 
wenzelm 
parents: 
36948 
diff
changeset
 | 
85  | 
    parse(commit(header), Token.reader(toks.toList)) match {
 | 
| 34190 | 86  | 
case Success(result, _) => result  | 
87  | 
case bad => error(bad.toString)  | 
|
88  | 
}  | 
|
| 34169 | 89  | 
}  | 
| 43611 | 90  | 
|
| 44159 | 91  | 
def read(source: CharSequence): Thy_Header =  | 
| 43646 | 92  | 
read(new CharSequenceReader(source))  | 
93  | 
||
| 44159 | 94  | 
def read(file: File): Thy_Header =  | 
| 43611 | 95  | 
  {
 | 
96  | 
val reader = Scan.byte_reader(file)  | 
|
| 43699 | 97  | 
    try { read(reader).map(Standard_System.decode_permissive_utf8) }
 | 
| 43611 | 98  | 
    finally { reader.close }
 | 
99  | 
}  | 
|
| 43648 | 100  | 
|
101  | 
||
102  | 
/* check */  | 
|
103  | 
||
| 44159 | 104  | 
def check(name: String, source: CharSequence): Thy_Header =  | 
| 43648 | 105  | 
  {
 | 
106  | 
val header = read(source)  | 
|
107  | 
val name1 = header.name  | 
|
| 43672 | 108  | 
    if (name != name1) error("Bad file name " + thy_path(name) + " for theory " + quote(name1))
 | 
109  | 
Path.explode(name)  | 
|
110  | 
header  | 
|
| 43648 | 111  | 
}  | 
| 34169 | 112  | 
}  | 
| 44159 | 113  | 
|
114  | 
||
| 44185 | 115  | 
sealed case class Thy_Header(  | 
116  | 
val name: String, val imports: List[String], val uses: List[(String, Boolean)])  | 
|
| 44159 | 117  | 
{
 | 
118  | 
def map(f: String => String): Thy_Header =  | 
|
| 44185 | 119  | 
Thy_Header(f(name), imports.map(f), uses.map(p => (f(p._1), p._2)))  | 
| 44163 | 120  | 
|
| 
44222
 
9d5ef6cd4ee1
use full .thy file name as node name, which makes MiscUtilities.resolveSymlinks/File.getCanonicalPath more predictable;
 
wenzelm 
parents: 
44185 
diff
changeset
 | 
121  | 
def norm_deps(f: String => String, g: String => String): Thy_Header =  | 
| 
 
9d5ef6cd4ee1
use full .thy file name as node name, which makes MiscUtilities.resolveSymlinks/File.getCanonicalPath more predictable;
 
wenzelm 
parents: 
44185 
diff
changeset
 | 
122  | 
copy(imports = imports.map(name => f(name)), uses = uses.map(p => (g(p._1), p._2)))  | 
| 44159 | 123  | 
}  | 
124  |