# HG changeset patch # User wenzelm # Date 1396868794 -7200 # Node ID 1e77ed11f2f78b9c7d03dd172a6568ee61a9b7de # Parent 70a13de8a154cb7a705705e653973cb077c142dc separate file_node vs. file_path, e.g. relevant on Windows for hyperlink to the latter; diff -r 70a13de8a154 -r 1e77ed11f2f7 src/Pure/PIDE/command.ML --- a/src/Pure/PIDE/command.ML Sun Apr 06 21:01:45 2014 +0200 +++ b/src/Pure/PIDE/command.ML Mon Apr 07 13:06:34 2014 +0200 @@ -6,7 +6,7 @@ signature COMMAND = sig - type blob = (string * (SHA1.digest * string list) option) Exn.result + type blob = (string * string * (SHA1.digest * string list) option) Exn.result val read_file: Path.T -> Position.T -> Path.T -> Token.file val read: (unit -> theory) -> Path.T -> blob list -> Token.T list -> Toplevel.transition type eval @@ -87,7 +87,7 @@ (* read *) type blob = - (string * (SHA1.digest * string list) option) Exn.result; (*file node name, digest, lines*) + (string * string * (SHA1.digest * string list) option) Exn.result; (*file, path, digest, lines*) fun read_file master_dir pos src_path = let @@ -101,10 +101,10 @@ local -fun blob_file src_path lines digest file = +fun blob_file src_path lines digest file_node = let val file_pos = - Position.file file (*sic!*) |> + Position.file file_node |> (case Position.get_id (Position.thread_data ()) of NONE => I | SOME exec_id => Position.put_id exec_id); @@ -115,16 +115,16 @@ [span] => span |> Thy_Syntax.resolve_files (fn cmd => fn (path, pos) => let - fun make_file src_path (Exn.Res (_, NONE)) = + fun make_file src_path (Exn.Res (_, _, NONE)) = Exn.interruptible_capture (fn () => read_file master_dir pos src_path) () - | make_file src_path (Exn.Res (file, SOME (digest, lines))) = - (Position.reports [(pos, Markup.language_path), (pos, Markup.path file)]; - Exn.Res (blob_file src_path lines digest file)) + | make_file src_path (Exn.Res (file_node, file_path, SOME (digest, lines))) = + (Position.reports [(pos, Markup.language_path), (pos, Markup.path file_path)]; + Exn.Res (blob_file src_path lines digest file_node)) | make_file _ (Exn.Exn e) = Exn.Exn e; val src_paths = Keyword.command_files cmd path; in if null blobs then - map2 make_file src_paths (map (K (Exn.Res ("", NONE))) src_paths) + map2 make_file src_paths (map (K (Exn.Res ("", "", NONE))) src_paths) else if length src_paths = length blobs then map2 make_file src_paths blobs else error ("Misalignment of inlined files" ^ Position.here pos) diff -r 70a13de8a154 -r 1e77ed11f2f7 src/Pure/PIDE/document.ML --- a/src/Pure/PIDE/document.ML Sun Apr 06 21:01:45 2014 +0200 +++ b/src/Pure/PIDE/document.ML Mon Apr 07 13:06:34 2014 +0200 @@ -18,7 +18,7 @@ type state val init_state: state val define_blob: string -> string -> state -> state - type blob_digest = (string * string option) Exn.result + type blob_digest = (string * string * string option) Exn.result val define_command: Document_ID.command -> string -> blob_digest list -> string -> state -> state val remove_versions: Document_ID.version list -> state -> state @@ -219,7 +219,8 @@ (** main state -- document structure and execution process **) -type blob_digest = (string * string option) Exn.result; (* file node name, raw digest*) +type blob_digest = + (string * string * string option) Exn.result; (* file node name, path, raw digest*) type execution = {version_id: Document_ID.version, (*static version id*) @@ -291,6 +292,10 @@ NONE => error ("Undefined blob: " ^ digest) | SOME content => content); +fun resolve_blob state (blob_digest: blob_digest) = + blob_digest |> Exn.map_result (fn (file, path, raw_digest) => + (file, path, Option.map (the_blob state) raw_digest)); + (* commands *) @@ -338,7 +343,7 @@ val blobs' = (commands', Symtab.empty) |-> Inttab.fold (fn (_, (_, blobs, _)) => blobs |> - fold (fn Exn.Res (_, SOME b) => Symtab.update (b, the_blob state b) | _ => I)); + fold (fn Exn.Res (_, _, SOME b) => Symtab.update (b, the_blob state b) | _ => I)); in (versions', blobs', commands', execution) end); @@ -505,7 +510,7 @@ val command_visible = visible_command node command_id'; val command_overlays = overlays node command_id'; val (command_name, blob_digests, span0) = the_command state command_id'; - val blobs = (map o Exn.map_result o apsnd o Option.map) (the_blob state) blob_digests; + val blobs = map (resolve_blob state) blob_digests; val span = Lazy.force span0; val eval' = diff -r 70a13de8a154 -r 1e77ed11f2f7 src/Pure/PIDE/protocol.ML --- a/src/Pure/PIDE/protocol.ML Sun Apr 06 21:01:45 2014 +0200 +++ b/src/Pure/PIDE/protocol.ML Mon Apr 07 13:06:34 2014 +0200 @@ -34,7 +34,7 @@ YXML.parse_body blobs_yxml |> let open XML.Decode in list (variant - [fn ([], a) => Exn.Res (pair string (option string) a), + [fn ([], a) => Exn.Res (triple string string (option string) a), fn ([], a) => Exn.Exn (ERROR (string a))]) end; in diff -r 70a13de8a154 -r 1e77ed11f2f7 src/Pure/PIDE/protocol.scala --- a/src/Pure/PIDE/protocol.scala Sun Apr 06 21:01:45 2014 +0200 +++ b/src/Pure/PIDE/protocol.scala Mon Apr 07 13:06:34 2014 +0200 @@ -352,7 +352,8 @@ val encode_blob: T[Command.Blob] = variant(List( { case Exn.Res((a, b)) => - (Nil, pair(string, option(string))((a.node, b.map(p => p._1.toString)))) }, + (Nil, triple(string, string, option(string))( + (a.node, Isabelle_System.posix_path(a.node), b.map(p => p._1.toString)))) }, { case Exn.Exn(e) => (Nil, string(Exn.message(e))) })) YXML.string_of_body(list(encode_blob)(command.blobs)) }