| author | wenzelm | 
| Wed, 18 Sep 2013 16:09:38 +0200 | |
| changeset 53713 | bb15972a644d | 
| parent 53336 | b3bf6d72fea5 | 
| child 55555 | 9c16317c91d1 | 
| permissions | -rw-r--r-- | 
| 43600 | 1  | 
/* Title: Pure/General/path.scala  | 
2  | 
Author: Makarius  | 
|
3  | 
||
| 43601 | 4  | 
Algebra of file-system paths: basic POSIX notation, extended by named  | 
5  | 
roots (e.g. //foo) and variables (e.g. $BAR).  | 
|
| 43600 | 6  | 
*/  | 
7  | 
||
8  | 
package isabelle  | 
|
9  | 
||
10  | 
||
| 48409 | 11  | 
import java.io.{File => JFile}
 | 
| 48373 | 12  | 
|
| 
43697
 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 
wenzelm 
parents: 
43670 
diff
changeset
 | 
13  | 
import scala.util.matching.Regex  | 
| 
 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 
wenzelm 
parents: 
43670 
diff
changeset
 | 
14  | 
|
| 
 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 
wenzelm 
parents: 
43670 
diff
changeset
 | 
15  | 
|
| 43600 | 16  | 
object Path  | 
17  | 
{
 | 
|
18  | 
/* path elements */  | 
|
19  | 
||
| 45244 | 20  | 
sealed abstract class Elem  | 
| 43600 | 21  | 
private case class Root(val name: String) extends Elem  | 
22  | 
private case class Basic(val name: String) extends Elem  | 
|
23  | 
private case class Variable(val name: String) extends Elem  | 
|
24  | 
private case object Parent extends Elem  | 
|
25  | 
||
26  | 
private def err_elem(msg: String, s: String): Nothing =  | 
|
| 53046 | 27  | 
error(msg + " path element specification " + quote(s))  | 
| 43600 | 28  | 
|
29  | 
private def check_elem(s: String): String =  | 
|
30  | 
    if (s == "" || s == "~" || s == "~~") err_elem("Illegal", s)
 | 
|
| 52106 | 31  | 
    else {
 | 
32  | 
"/\\$:\"'".iterator.foreach(c =>  | 
|
33  | 
if (s.iterator.exists(_ == c))  | 
|
34  | 
          err_elem("Illegal character " + quote(c.toString) + " in", s))
 | 
|
35  | 
s  | 
|
36  | 
}  | 
|
| 43600 | 37  | 
|
38  | 
private def root_elem(s: String): Elem = Root(check_elem(s))  | 
|
39  | 
private def basic_elem(s: String): Elem = Basic(check_elem(s))  | 
|
40  | 
private def variable_elem(s: String): Elem = Variable(check_elem(s))  | 
|
41  | 
||
42  | 
private def apply_elem(y: Elem, xs: List[Elem]): List[Elem] =  | 
|
43  | 
    (y, xs) match {
 | 
|
44  | 
case (Root(_), _) => List(y)  | 
|
45  | 
case (Parent, Root(_) :: _) => xs  | 
|
46  | 
case (Parent, Basic(_) :: rest) => rest  | 
|
47  | 
case _ => y :: xs  | 
|
48  | 
}  | 
|
49  | 
||
50  | 
private def norm_elems(elems: List[Elem]): List[Elem] =  | 
|
51  | 
(elems :\ (Nil: List[Elem]))(apply_elem)  | 
|
52  | 
||
53  | 
private def implode_elem(elem: Elem): String =  | 
|
54  | 
    elem match {
 | 
|
55  | 
      case Root("") => ""
 | 
|
56  | 
case Root(s) => "//" + s  | 
|
57  | 
case Basic(s) => s  | 
|
58  | 
case Variable(s) => "$" + s  | 
|
59  | 
case Parent => ".."  | 
|
60  | 
}  | 
|
61  | 
||
62  | 
||
63  | 
/* path constructors */  | 
|
64  | 
||
| 45244 | 65  | 
val current: Path = new Path(Nil)  | 
66  | 
  val root: Path = new Path(List(Root("")))
 | 
|
67  | 
def named_root(s: String): Path = new Path(List(root_elem(s)))  | 
|
68  | 
def basic(s: String): Path = new Path(List(basic_elem(s)))  | 
|
69  | 
def variable(s: String): Path = new Path(List(variable_elem(s)))  | 
|
70  | 
val parent: Path = new Path(List(Parent))  | 
|
| 43600 | 71  | 
|
72  | 
||
73  | 
/* explode */  | 
|
74  | 
||
75  | 
def explode(str: String): Path =  | 
|
76  | 
  {
 | 
|
| 52106 | 77  | 
def explode_elem(s: String): Elem =  | 
78  | 
      try {
 | 
|
79  | 
if (s == "..") Parent  | 
|
80  | 
        else if (s == "~") Variable("USER_HOME")
 | 
|
81  | 
        else if (s == "~~") Variable("ISABELLE_HOME")
 | 
|
82  | 
        else if (s.startsWith("$")) variable_elem(s.substring(1))
 | 
|
83  | 
else basic_elem(s)  | 
|
84  | 
}  | 
|
85  | 
      catch { case ERROR(msg) => cat_error(msg, "The error(s) above occurred in " + quote(str)) }
 | 
|
86  | 
||
| 
43670
 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 
wenzelm 
parents: 
43669 
diff
changeset
 | 
87  | 
    val ss = space_explode('/', str)
 | 
| 43600 | 88  | 
val r = ss.takeWhile(_.isEmpty).length  | 
89  | 
val es = ss.dropWhile(_.isEmpty)  | 
|
90  | 
val (roots, raw_elems) =  | 
|
91  | 
if (r == 0) (Nil, es)  | 
|
92  | 
      else if (r == 1) (List(Root("")), es)
 | 
|
93  | 
      else if (es.isEmpty) (List(Root("")), Nil)
 | 
|
94  | 
else (List(root_elem(es.head)), es.tail)  | 
|
| 52106 | 95  | 
val elems = raw_elems.filterNot(s => s.isEmpty || s == ".").map(explode_elem)  | 
96  | 
||
97  | 
new Path(norm_elems(elems.reverse ++ roots))  | 
|
| 43600 | 98  | 
}  | 
| 43669 | 99  | 
|
| 48484 | 100  | 
def is_ok(str: String): Boolean =  | 
101  | 
    try { explode(str); true } catch { case ERROR(_) => false }
 | 
|
102  | 
||
| 43669 | 103  | 
def split(str: String): List[Path] =  | 
| 
43670
 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 
wenzelm 
parents: 
43669 
diff
changeset
 | 
104  | 
    space_explode(':', str).filterNot(_.isEmpty).map(explode)
 | 
| 48457 | 105  | 
|
106  | 
||
107  | 
/* encode */  | 
|
108  | 
||
109  | 
val encode: XML.Encode.T[Path] = (path => XML.Encode.string(path.implode))  | 
|
| 43600 | 110  | 
}  | 
111  | 
||
| 43669 | 112  | 
|
| 46712 | 113  | 
final class Path private(private val elems: List[Path.Elem]) // reversed elements  | 
| 43600 | 114  | 
{
 | 
115  | 
def is_current: Boolean = elems.isEmpty  | 
|
116  | 
def is_absolute: Boolean = !elems.isEmpty && elems.last.isInstanceOf[Path.Root]  | 
|
117  | 
  def is_basic: Boolean = elems match { case List(Path.Basic(_)) => true case _ => false }
 | 
|
118  | 
||
| 45244 | 119  | 
def +(other: Path): Path = new Path((other.elems :\ elems)(Path.apply_elem))  | 
| 43600 | 120  | 
|
121  | 
||
| 43604 | 122  | 
/* implode */  | 
| 43600 | 123  | 
|
| 43604 | 124  | 
def implode: String =  | 
| 43600 | 125  | 
    elems match {
 | 
126  | 
case Nil => "."  | 
|
127  | 
      case List(Path.Root("")) => "/"
 | 
|
| 43604 | 128  | 
      case _ => elems.map(Path.implode_elem).reverse.mkString("/")
 | 
| 43600 | 129  | 
}  | 
130  | 
||
| 43652 | 131  | 
override def toString: String = quote(implode)  | 
| 43600 | 132  | 
|
133  | 
||
134  | 
/* base element */  | 
|
135  | 
||
136  | 
private def split_path: (Path, String) =  | 
|
137  | 
    elems match {
 | 
|
| 45244 | 138  | 
case Path.Basic(s) :: xs => (new Path(xs), s)  | 
| 43604 | 139  | 
      case _ => error("Cannot split path into dir/base: " + toString)
 | 
| 43600 | 140  | 
}  | 
141  | 
||
142  | 
def dir: Path = split_path._1  | 
|
| 45244 | 143  | 
def base: Path = new Path(List(Path.Basic(split_path._2)))  | 
| 43600 | 144  | 
|
145  | 
def ext(e: String): Path =  | 
|
146  | 
if (e == "") this  | 
|
147  | 
    else {
 | 
|
148  | 
val (prfx, s) = split_path  | 
|
| 43604 | 149  | 
prfx + Path.basic(s + "." + e)  | 
| 43600 | 150  | 
}  | 
| 43604 | 151  | 
|
| 53336 | 152  | 
def backup: Path =  | 
153  | 
  {
 | 
|
154  | 
val (prfx, s) = split_path  | 
|
155  | 
prfx + Path.basic(s + "~")  | 
|
156  | 
}  | 
|
157  | 
||
| 
43697
 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 
wenzelm 
parents: 
43670 
diff
changeset
 | 
158  | 
  private val Ext = new Regex("(.*)\\.([^.]*)")
 | 
| 
 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 
wenzelm 
parents: 
43670 
diff
changeset
 | 
159  | 
|
| 
 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 
wenzelm 
parents: 
43670 
diff
changeset
 | 
160  | 
def split_ext: (Path, String) =  | 
| 
 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 
wenzelm 
parents: 
43670 
diff
changeset
 | 
161  | 
  {
 | 
| 
 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 
wenzelm 
parents: 
43670 
diff
changeset
 | 
162  | 
val (prefix, base) = split_path  | 
| 
 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 
wenzelm 
parents: 
43670 
diff
changeset
 | 
163  | 
    base match {
 | 
| 
 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 
wenzelm 
parents: 
43670 
diff
changeset
 | 
164  | 
case Ext(b, e) => (prefix + Path.basic(b), e)  | 
| 
 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 
wenzelm 
parents: 
43670 
diff
changeset
 | 
165  | 
case _ => (Path.basic(base), "")  | 
| 
 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 
wenzelm 
parents: 
43670 
diff
changeset
 | 
166  | 
}  | 
| 
 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 
wenzelm 
parents: 
43670 
diff
changeset
 | 
167  | 
}  | 
| 
 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 
wenzelm 
parents: 
43670 
diff
changeset
 | 
168  | 
|
| 43604 | 169  | 
|
170  | 
/* expand */  | 
|
171  | 
||
| 43664 | 172  | 
def expand: Path =  | 
| 43604 | 173  | 
  {
 | 
174  | 
def eval(elem: Path.Elem): List[Path.Elem] =  | 
|
175  | 
      elem match {
 | 
|
| 43664 | 176  | 
case Path.Variable(s) =>  | 
| 48658 | 177  | 
val path = Path.explode(Isabelle_System.getenv_strict(s))  | 
178  | 
if (path.elems.exists(_.isInstanceOf[Path.Variable]))  | 
|
| 53046 | 179  | 
            error("Illegal path variable nesting: " + s + "=" + path.toString)
 | 
| 48658 | 180  | 
else path.elems  | 
| 43604 | 181  | 
case x => List(x)  | 
182  | 
}  | 
|
183  | 
||
| 45244 | 184  | 
new Path(Path.norm_elems(elems.map(eval).flatten))  | 
| 43604 | 185  | 
}  | 
| 48373 | 186  | 
|
187  | 
||
| 48548 | 188  | 
/* source position */  | 
189  | 
||
190  | 
def position: Position.T = Position.File(implode)  | 
|
191  | 
||
192  | 
||
| 48373 | 193  | 
/* platform file */  | 
194  | 
||
| 48409 | 195  | 
def file: JFile = Isabelle_System.platform_file(this)  | 
| 48548 | 196  | 
def is_file: Boolean = file.isFile  | 
197  | 
def is_dir: Boolean = file.isDirectory  | 
|
| 43600 | 198  | 
}  |