| author | nipkow | 
| Mon, 24 Sep 2018 16:09:27 +0200 | |
| changeset 69065 | 440f7a575760 | 
| parent 67181 | 0da2811afd87 | 
| child 69366 | b6dacf6eabe3 | 
| 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 | 
| 60215 | 21 | private case class Root(name: String) extends Elem | 
| 22 | private case class Basic(name: String) extends Elem | |
| 23 | private case class Variable(name: String) extends Elem | |
| 43600 | 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 => | |
| 60215 | 33 | if (s.iterator.contains(c)) | 
| 52106 | 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 | ||
| 56844 | 53 | private def implode_elem(elem: Elem, short: Boolean): String = | 
| 43600 | 54 |     elem match {
 | 
| 55 |       case Root("") => ""
 | |
| 56 | case Root(s) => "//" + s | |
| 57 | case Basic(s) => s | |
| 56844 | 58 |       case Variable("USER_HOME") if short => "~"
 | 
| 59 |       case Variable("ISABELLE_HOME") if short => "~~"
 | |
| 43600 | 60 | case Variable(s) => "$" + s | 
| 61 | case Parent => ".." | |
| 62 | } | |
| 63 | ||
| 64 | ||
| 65 | /* path constructors */ | |
| 66 | ||
| 45244 | 67 | val current: Path = new Path(Nil) | 
| 68 |   val root: Path = new Path(List(Root("")))
 | |
| 69 | def named_root(s: String): Path = new Path(List(root_elem(s))) | |
| 70 | def basic(s: String): Path = new Path(List(basic_elem(s))) | |
| 71 | def variable(s: String): Path = new Path(List(variable_elem(s))) | |
| 72 | val parent: Path = new Path(List(Parent)) | |
| 43600 | 73 | |
| 74 | ||
| 75 | /* explode */ | |
| 76 | ||
| 77 | def explode(str: String): Path = | |
| 78 |   {
 | |
| 52106 | 79 | def explode_elem(s: String): Elem = | 
| 80 |       try {
 | |
| 81 | if (s == "..") Parent | |
| 82 |         else if (s == "~") Variable("USER_HOME")
 | |
| 83 |         else if (s == "~~") Variable("ISABELLE_HOME")
 | |
| 84 |         else if (s.startsWith("$")) variable_elem(s.substring(1))
 | |
| 85 | else basic_elem(s) | |
| 86 | } | |
| 87 |       catch { case ERROR(msg) => cat_error(msg, "The error(s) above occurred in " + quote(str)) }
 | |
| 88 | ||
| 43670 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43669diff
changeset | 89 |     val ss = space_explode('/', str)
 | 
| 43600 | 90 | val r = ss.takeWhile(_.isEmpty).length | 
| 91 | val es = ss.dropWhile(_.isEmpty) | |
| 92 | val (roots, raw_elems) = | |
| 93 | if (r == 0) (Nil, es) | |
| 94 |       else if (r == 1) (List(Root("")), es)
 | |
| 95 |       else if (es.isEmpty) (List(Root("")), Nil)
 | |
| 96 | else (List(root_elem(es.head)), es.tail) | |
| 52106 | 97 | val elems = raw_elems.filterNot(s => s.isEmpty || s == ".").map(explode_elem) | 
| 98 | ||
| 63866 | 99 | new Path(norm_elems(elems reverse_::: roots)) | 
| 43600 | 100 | } | 
| 43669 | 101 | |
| 55879 
ac979f750c1a
clarified path checks: avoid crash of rendering due to spurious errors;
 wenzelm parents: 
55555diff
changeset | 102 | def is_wellformed(str: String): Boolean = | 
| 48484 | 103 |     try { explode(str); true } catch { case ERROR(_) => false }
 | 
| 104 | ||
| 55879 
ac979f750c1a
clarified path checks: avoid crash of rendering due to spurious errors;
 wenzelm parents: 
55555diff
changeset | 105 | def is_valid(str: String): Boolean = | 
| 
ac979f750c1a
clarified path checks: avoid crash of rendering due to spurious errors;
 wenzelm parents: 
55555diff
changeset | 106 |     try { explode(str).expand; true } catch { case ERROR(_) => false }
 | 
| 
ac979f750c1a
clarified path checks: avoid crash of rendering due to spurious errors;
 wenzelm parents: 
55555diff
changeset | 107 | |
| 43669 | 108 | def split(str: String): List[Path] = | 
| 43670 
7f933761764b
prefer space_explode/split_lines as in Isabelle/ML;
 wenzelm parents: 
43669diff
changeset | 109 |     space_explode(':', str).filterNot(_.isEmpty).map(explode)
 | 
| 48457 | 110 | |
| 111 | ||
| 112 | /* encode */ | |
| 113 | ||
| 114 | val encode: XML.Encode.T[Path] = (path => XML.Encode.string(path.implode)) | |
| 43600 | 115 | } | 
| 116 | ||
| 43669 | 117 | |
| 46712 | 118 | final class Path private(private val elems: List[Path.Elem]) // reversed elements | 
| 43600 | 119 | {
 | 
| 120 | def is_current: Boolean = elems.isEmpty | |
| 59319 | 121 | def is_absolute: Boolean = elems.nonEmpty && elems.last.isInstanceOf[Path.Root] | 
| 65559 | 122 |   def is_root: Boolean = elems match { case List(Path.Root(_)) => true case _ => false }
 | 
| 43600 | 123 |   def is_basic: Boolean = elems match { case List(Path.Basic(_)) => true case _ => false }
 | 
| 124 | ||
| 45244 | 125 | def +(other: Path): Path = new Path((other.elems :\ elems)(Path.apply_elem)) | 
| 43600 | 126 | |
| 127 | ||
| 43604 | 128 | /* implode */ | 
| 43600 | 129 | |
| 56844 | 130 | private def gen_implode(short: Boolean): String = | 
| 43600 | 131 |     elems match {
 | 
| 132 | case Nil => "." | |
| 133 |       case List(Path.Root("")) => "/"
 | |
| 56844 | 134 |       case _ => elems.map(Path.implode_elem(_, short)).reverse.mkString("/")
 | 
| 43600 | 135 | } | 
| 56844 | 136 | def implode: String = gen_implode(false) | 
| 137 | def implode_short: String = gen_implode(true) | |
| 43600 | 138 | |
| 43652 | 139 | override def toString: String = quote(implode) | 
| 43600 | 140 | |
| 141 | ||
| 142 | /* base element */ | |
| 143 | ||
| 144 | private def split_path: (Path, String) = | |
| 145 |     elems match {
 | |
| 45244 | 146 | case Path.Basic(s) :: xs => (new Path(xs), s) | 
| 43604 | 147 |       case _ => error("Cannot split path into dir/base: " + toString)
 | 
| 43600 | 148 | } | 
| 149 | ||
| 150 | def dir: Path = split_path._1 | |
| 45244 | 151 | def base: Path = new Path(List(Path.Basic(split_path._2))) | 
| 65999 | 152 | def base_name: String = base.implode | 
| 43600 | 153 | |
| 154 | def ext(e: String): Path = | |
| 155 | if (e == "") this | |
| 156 |     else {
 | |
| 157 | val (prfx, s) = split_path | |
| 43604 | 158 | prfx + Path.basic(s + "." + e) | 
| 43600 | 159 | } | 
| 43604 | 160 | |
| 53336 | 161 | def backup: Path = | 
| 162 |   {
 | |
| 163 | val (prfx, s) = split_path | |
| 164 | prfx + Path.basic(s + "~") | |
| 165 | } | |
| 166 | ||
| 58610 | 167 | def backup2: Path = | 
| 168 |   {
 | |
| 169 | val (prfx, s) = split_path | |
| 170 | prfx + Path.basic(s + "~~") | |
| 171 | } | |
| 172 | ||
| 43697 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43670diff
changeset | 173 |   private val Ext = new Regex("(.*)\\.([^.]*)")
 | 
| 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43670diff
changeset | 174 | |
| 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43670diff
changeset | 175 | def split_ext: (Path, String) = | 
| 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43670diff
changeset | 176 |   {
 | 
| 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43670diff
changeset | 177 | val (prefix, base) = split_path | 
| 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43670diff
changeset | 178 |     base match {
 | 
| 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43670diff
changeset | 179 | case Ext(b, e) => (prefix + Path.basic(b), e) | 
| 56556 | 180 | case _ => (prefix + Path.basic(base), "") | 
| 43697 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43670diff
changeset | 181 | } | 
| 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43670diff
changeset | 182 | } | 
| 
77ce24aa1770
explicit Document.Node.Header, with master_dir and thy_name;
 wenzelm parents: 
43670diff
changeset | 183 | |
| 43604 | 184 | |
| 185 | /* expand */ | |
| 186 | ||
| 64228 
b46969a851a9
expand relatively to given environment, notably remote HOME;
 wenzelm parents: 
63866diff
changeset | 187 | def expand_env(env: Map[String, String]): Path = | 
| 43604 | 188 |   {
 | 
| 189 | def eval(elem: Path.Elem): List[Path.Elem] = | |
| 190 |       elem match {
 | |
| 43664 | 191 | case Path.Variable(s) => | 
| 64228 
b46969a851a9
expand relatively to given environment, notably remote HOME;
 wenzelm parents: 
63866diff
changeset | 192 | val path = Path.explode(Isabelle_System.getenv_strict(s, env)) | 
| 48658 | 193 | if (path.elems.exists(_.isInstanceOf[Path.Variable])) | 
| 53046 | 194 |             error("Illegal path variable nesting: " + s + "=" + path.toString)
 | 
| 48658 | 195 | else path.elems | 
| 43604 | 196 | case x => List(x) | 
| 197 | } | |
| 198 | ||
| 45244 | 199 | new Path(Path.norm_elems(elems.map(eval).flatten)) | 
| 43604 | 200 | } | 
| 48373 | 201 | |
| 64228 
b46969a851a9
expand relatively to given environment, notably remote HOME;
 wenzelm parents: 
63866diff
changeset | 202 | def expand: Path = expand_env(Isabelle_System.settings()) | 
| 
b46969a851a9
expand relatively to given environment, notably remote HOME;
 wenzelm parents: 
63866diff
changeset | 203 | |
| 48373 | 204 | |
| 48548 | 205 | /* source position */ | 
| 206 | ||
| 207 | def position: Position.T = Position.File(implode) | |
| 208 | ||
| 209 | ||
| 66232 | 210 | /* platform files */ | 
| 48373 | 211 | |
| 60988 | 212 | def file: JFile = File.platform_file(this) | 
| 48548 | 213 | def is_file: Boolean = file.isFile | 
| 214 | def is_dir: Boolean = file.isDirectory | |
| 65833 | 215 | |
| 66232 | 216 | def absolute_file: JFile = File.absolute(file) | 
| 217 | def canonical_file: JFile = File.canonical(file) | |
| 67181 | 218 | |
| 219 | def absolute: Path = File.path(absolute_file) | |
| 220 | def canonical: Path = File.path(canonical_file) | |
| 43600 | 221 | } |