| author | wenzelm | 
| Fri, 27 Jul 2012 12:43:58 +0200 | |
| changeset 48542 | 0a5f598cacec | 
| parent 48484 | 70898d016538 | 
| child 48548 | 49afe0e92163 | 
| 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: 
43670diff
changeset | 13 | import scala.util.matching.Regex | 
| 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43670diff
changeset | 14 | |
| 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43670diff
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 = | |
| 43652 | 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)
 | |
| 31 | else | |
| 48420 
a8ed41b6280b
disallow quotes in path specifications -- extra paranoia;
 wenzelm parents: 
48409diff
changeset | 32 | s.iterator.filter(c => | 
| 
a8ed41b6280b
disallow quotes in path specifications -- extra paranoia;
 wenzelm parents: 
48409diff
changeset | 33 |           c == '/' || c == '\\' || c == '$' || c == ':' || c == '"' || c == '\'').toList match {
 | 
| 43600 | 34 | case Nil => s | 
| 35 | case bads => | |
| 43652 | 36 |           err_elem ("Illegal character(s) " + commas_quote(bads.map(_.toString)) + " in", s)
 | 
| 43600 | 37 | } | 
| 38 | ||
| 39 | private def root_elem(s: String): Elem = Root(check_elem(s)) | |
| 40 | private def basic_elem(s: String): Elem = Basic(check_elem(s)) | |
| 41 | private def variable_elem(s: String): Elem = Variable(check_elem(s)) | |
| 42 | ||
| 43 | private def apply_elem(y: Elem, xs: List[Elem]): List[Elem] = | |
| 44 |     (y, xs) match {
 | |
| 45 | case (Root(_), _) => List(y) | |
| 46 | case (Parent, Root(_) :: _) => xs | |
| 47 | case (Parent, Basic(_) :: rest) => rest | |
| 48 | case _ => y :: xs | |
| 49 | } | |
| 50 | ||
| 51 | private def norm_elems(elems: List[Elem]): List[Elem] = | |
| 52 | (elems :\ (Nil: List[Elem]))(apply_elem) | |
| 53 | ||
| 54 | private def implode_elem(elem: Elem): String = | |
| 55 |     elem match {
 | |
| 56 |       case Root("") => ""
 | |
| 57 | case Root(s) => "//" + s | |
| 58 | case Basic(s) => s | |
| 59 | case Variable(s) => "$" + s | |
| 60 | case Parent => ".." | |
| 61 | } | |
| 62 | ||
| 63 | ||
| 64 | /* path constructors */ | |
| 65 | ||
| 45244 | 66 | val current: Path = new Path(Nil) | 
| 67 |   val root: Path = new Path(List(Root("")))
 | |
| 68 | def named_root(s: String): Path = new Path(List(root_elem(s))) | |
| 69 | def basic(s: String): Path = new Path(List(basic_elem(s))) | |
| 70 | def variable(s: String): Path = new Path(List(variable_elem(s))) | |
| 71 | val parent: Path = new Path(List(Parent)) | |
| 43600 | 72 | |
| 73 | ||
| 74 | /* explode */ | |
| 75 | ||
| 76 | private def explode_elem(s: String): Elem = | |
| 77 | if (s == "..") Parent | |
| 47661 
012a887997f3
USER_HOME settings variable points to cross-platform user home directory;
 wenzelm parents: 
46712diff
changeset | 78 |     else if (s == "~") Variable("USER_HOME")
 | 
| 43600 | 79 |     else if (s == "~~") Variable("ISABELLE_HOME")
 | 
| 80 |     else if (s.startsWith("$")) variable_elem(s.substring(1))
 | |
| 81 | else basic_elem(s) | |
| 82 | ||
| 83 | private def explode_elems(ss: List[String]): List[Elem] = | |
| 84 | ss.filterNot(s => s.isEmpty || s == ".").map(explode_elem).reverse | |
| 85 | ||
| 86 | def explode(str: String): Path = | |
| 87 |   {
 | |
| 43670 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43669diff
changeset | 88 |     val ss = space_explode('/', str)
 | 
| 43600 | 89 | val r = ss.takeWhile(_.isEmpty).length | 
| 90 | val es = ss.dropWhile(_.isEmpty) | |
| 91 | val (roots, raw_elems) = | |
| 92 | if (r == 0) (Nil, es) | |
| 93 |       else if (r == 1) (List(Root("")), es)
 | |
| 94 |       else if (es.isEmpty) (List(Root("")), Nil)
 | |
| 95 | else (List(root_elem(es.head)), es.tail) | |
| 45244 | 96 | new Path(norm_elems(explode_elems(raw_elems) ++ roots)) | 
| 43600 | 97 | } | 
| 43669 | 98 | |
| 48484 | 99 | def is_ok(str: String): Boolean = | 
| 100 |     try { explode(str); true } catch { case ERROR(_) => false }
 | |
| 101 | ||
| 43669 | 102 | def split(str: String): List[Path] = | 
| 43670 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43669diff
changeset | 103 |     space_explode(':', str).filterNot(_.isEmpty).map(explode)
 | 
| 48457 | 104 | |
| 105 | ||
| 106 | /* encode */ | |
| 107 | ||
| 108 | val encode: XML.Encode.T[Path] = (path => XML.Encode.string(path.implode)) | |
| 43600 | 109 | } | 
| 110 | ||
| 43669 | 111 | |
| 46712 | 112 | final class Path private(private val elems: List[Path.Elem]) // reversed elements | 
| 43600 | 113 | {
 | 
| 114 | def is_current: Boolean = elems.isEmpty | |
| 115 | def is_absolute: Boolean = !elems.isEmpty && elems.last.isInstanceOf[Path.Root] | |
| 116 |   def is_basic: Boolean = elems match { case List(Path.Basic(_)) => true case _ => false }
 | |
| 117 | ||
| 45244 | 118 | def +(other: Path): Path = new Path((other.elems :\ elems)(Path.apply_elem)) | 
| 43600 | 119 | |
| 120 | ||
| 43604 | 121 | /* implode */ | 
| 43600 | 122 | |
| 43604 | 123 | def implode: String = | 
| 43600 | 124 |     elems match {
 | 
| 125 | case Nil => "." | |
| 126 |       case List(Path.Root("")) => "/"
 | |
| 43604 | 127 |       case _ => elems.map(Path.implode_elem).reverse.mkString("/")
 | 
| 43600 | 128 | } | 
| 129 | ||
| 43652 | 130 | override def toString: String = quote(implode) | 
| 43600 | 131 | |
| 132 | ||
| 133 | /* base element */ | |
| 134 | ||
| 135 | private def split_path: (Path, String) = | |
| 136 |     elems match {
 | |
| 45244 | 137 | case Path.Basic(s) :: xs => (new Path(xs), s) | 
| 43604 | 138 |       case _ => error("Cannot split path into dir/base: " + toString)
 | 
| 43600 | 139 | } | 
| 140 | ||
| 141 | def dir: Path = split_path._1 | |
| 45244 | 142 | def base: Path = new Path(List(Path.Basic(split_path._2))) | 
| 43600 | 143 | |
| 144 | def ext(e: String): Path = | |
| 145 | if (e == "") this | |
| 146 |     else {
 | |
| 147 | val (prfx, s) = split_path | |
| 43604 | 148 | prfx + Path.basic(s + "." + e) | 
| 43600 | 149 | } | 
| 43604 | 150 | |
| 43697 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43670diff
changeset | 151 |   private val Ext = new Regex("(.*)\\.([^.]*)")
 | 
| 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43670diff
changeset | 152 | |
| 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43670diff
changeset | 153 | def split_ext: (Path, String) = | 
| 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43670diff
changeset | 154 |   {
 | 
| 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43670diff
changeset | 155 | val (prefix, base) = split_path | 
| 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43670diff
changeset | 156 |     base match {
 | 
| 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43670diff
changeset | 157 | case Ext(b, e) => (prefix + Path.basic(b), e) | 
| 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43670diff
changeset | 158 | case _ => (Path.basic(base), "") | 
| 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43670diff
changeset | 159 | } | 
| 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43670diff
changeset | 160 | } | 
| 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43670diff
changeset | 161 | |
| 43604 | 162 | |
| 163 | /* expand */ | |
| 164 | ||
| 43664 | 165 | def expand: Path = | 
| 43604 | 166 |   {
 | 
| 167 | def eval(elem: Path.Elem): List[Path.Elem] = | |
| 168 |       elem match {
 | |
| 43664 | 169 | case Path.Variable(s) => | 
| 170 | Path.explode(Isabelle_System.getenv_strict(s)).elems | |
| 43604 | 171 | case x => List(x) | 
| 172 | } | |
| 173 | ||
| 45244 | 174 | new Path(Path.norm_elems(elems.map(eval).flatten)) | 
| 43604 | 175 | } | 
| 48373 | 176 | |
| 177 | ||
| 178 | /* platform file */ | |
| 179 | ||
| 48409 | 180 | def file: JFile = Isabelle_System.platform_file(this) | 
| 43600 | 181 | } |