src/Pure/Thy/thy_load.ML
author wenzelm
Sat, 16 Nov 2013 21:18:31 +0100
changeset 54456 f4b1440d9880
parent 54455 1d977436c1bf
child 54458 96ccc8972fc7
permissions -rw-r--r--
simplified HTML theory presentation;

(*  Title:      Pure/Thy/thy_load.ML
    Author:     Makarius

Loading files that contribute to a theory.
*)

signature THY_LOAD =
sig
  val master_directory: theory -> Path.T
  val imports_of: theory -> (string * Position.T) list
  val thy_path: Path.T -> Path.T
  val read_files: Path.T -> string -> Path.T * Position.T -> Token.file list
  val parse_files: string -> (theory -> Token.file list) parser
  val check_thy: Path.T -> string ->
   {master: Path.T * SHA1.digest, text: string, theory_pos: Position.T,
    imports: (string * Position.T) list, keywords: Thy_Header.keywords}
  val provide: Path.T * SHA1.digest -> theory -> theory
  val provide_parse_files: string -> (theory -> Token.file list * theory) parser
  val load_file: theory -> Path.T -> (Path.T * SHA1.digest) * string
  val loaded_files: theory -> Path.T list
  val loaded_files_current: theory -> bool
  val begin_theory: Path.T -> Thy_Header.header -> theory list -> theory
  val load_thy: (Toplevel.transition -> Time.time option) -> int -> Path.T -> Thy_Header.header ->
    Position.T -> string -> theory list -> theory * (unit -> unit) * int
end;

structure Thy_Load: THY_LOAD =
struct

(* manage source files *)

type files =
 {master_dir: Path.T,  (*master directory of theory source*)
  imports: (string * Position.T) list,  (*source specification of imports*)
  provided: (Path.T * SHA1.digest) list};  (*source path, digest*)

fun make_files (master_dir, imports, provided): files =
 {master_dir = master_dir, imports = imports, provided = provided};

structure Files = Theory_Data
(
  type T = files;
  val empty = make_files (Path.current, [], []);
  fun extend _ = empty;
  fun merge _ = empty;
);

fun map_files f =
  Files.map (fn {master_dir, imports, provided} =>
    make_files (f (master_dir, imports, provided)));


val master_directory = #master_dir o Files.get;
val imports_of = #imports o Files.get;

fun put_deps master_dir imports = map_files (fn _ => (master_dir, imports, []));


(* auxiliary files *)

fun check_file dir file = File.check_file (File.full_path dir file);

fun read_files dir cmd (path, pos) =
  let
    fun make_file file =
      let
        val _ = Position.report pos (Markup.path (Path.implode file));
        val full_path = check_file dir file;
      in {src_path = file, text = File.read full_path, pos = Path.position full_path} end;
    val paths =
      (case Keyword.command_files cmd of
        [] => [path]
      | exts => map (fn ext => Path.ext ext path) exts);
  in map make_file paths end;

fun parse_files cmd =
  Scan.ahead Parse.not_eof -- Parse.path >> (fn (tok, name) => fn thy =>
    (case Token.get_files tok of
      SOME files => files
    | NONE => read_files (master_directory thy) cmd (Path.explode name, Token.position_of tok)));


(* theory files *)

val thy_path = Path.ext "thy";

fun check_thy dir thy_name =
  let
    val path = thy_path (Path.basic thy_name);
    val master_file = check_file dir path;
    val text = File.read master_file;

    val {name = (name, pos), imports, keywords} =
      Thy_Header.read (Path.position master_file) text;
    val _ = thy_name <> name andalso
      error ("Bad file name " ^ Path.print path ^ " for theory " ^ quote name ^ Position.here pos);
  in
   {master = (master_file, SHA1.digest text), text = text, theory_pos = pos,
    imports = imports, keywords = keywords}
  end;


(* load files *)

fun provide (src_path, id) =
  map_files (fn (master_dir, imports, provided) =>
    if AList.defined (op =) provided src_path then
      error ("Duplicate use of source file: " ^ Path.print src_path)
    else (master_dir, imports, (src_path, id) :: provided));

fun provide_parse_files cmd =
  parse_files cmd >> (fn files => fn thy =>
    let
      val fs = files thy;
      val thy' = fold (fn {src_path, text, ...} => provide (src_path, SHA1.digest text)) fs thy;
    in (fs, thy') end);

fun load_file thy src_path =
  let
    val full_path = check_file (master_directory thy) src_path;
    val text = File.read full_path;
    val id = SHA1.digest text;
  in ((full_path, id), text) end;

fun loaded_files_current thy =
  #provided (Files.get thy) |>
    forall (fn (src_path, id) =>
      (case try (load_file thy) src_path of
        NONE => false
      | SOME ((_, id'), _) => id = id'));

(*Proof General legacy*)
fun loaded_files thy =
  let val {master_dir, provided, ...} = Files.get thy
  in map (File.full_path master_dir o #1) provided end;


(* load_thy *)

fun begin_theory master_dir {name, imports, keywords} parents =
  Theory.begin_theory name parents
  |> put_deps master_dir imports
  |> fold Thy_Header.declare_keyword keywords;

fun excursion last_timing init elements =
  let
    fun prepare_span span =
      Thy_Syntax.span_content span
      |> Command.read init
      |> (fn tr => Toplevel.put_timing (last_timing tr) tr);

    fun element_result span_elem (st, _) =
      let
        val elem = Thy_Syntax.map_element prepare_span span_elem;
        val (results, st') = Toplevel.element_result elem st;
        val pos' = Toplevel.pos_of (Thy_Syntax.last_element elem);
      in (results, (st', pos')) end;

    val (results, (end_state, end_pos)) =
      fold_map element_result elements (Toplevel.toplevel, Position.none);

    val thy = Toplevel.end_theory end_pos end_state;
  in (results, thy) end;

fun load_thy last_timing update_time master_dir header text_pos text parents =
  let
    val time = ! Toplevel.timing;

    val {name = (name, _), ...} = header;
    val _ = Thy_Header.define_keywords header;

    val lexs = Keyword.get_lexicons ();
    val toks = Thy_Syntax.parse_tokens lexs text_pos text;
    val spans = map (Thy_Syntax.resolve_files (read_files master_dir)) (Thy_Syntax.parse_spans toks);
    val elements = Thy_Syntax.parse_elements spans;

    fun init () =
      begin_theory master_dir header parents
      |> Present.begin_theory update_time
          (fn () => HTML.html_mode (implode o map Thy_Syntax.present_span) spans);

    val _ = if time then writeln ("\n**** Starting theory " ^ quote name ^ " ****") else ();
    val (results, thy) = cond_timeit time "" (fn () => excursion last_timing init elements);
    val _ = if time then writeln ("**** Finished theory " ^ quote name ^ " ****\n") else ();

    fun present () =
      let
        val res = filter_out (Toplevel.is_ignored o #1) (maps Toplevel.join_results results);
        val ((minor, _), outer_syntax) = Outer_Syntax.get_syntax ();
      in
        if exists (Toplevel.is_skipped_proof o #2) res then
          warning ("Cannot present theory with skipped proofs: " ^ quote name)
        else
          Thy_Output.present_thy minor Keyword.command_tags
            (Outer_Syntax.is_markup outer_syntax) res toks
          |> Buffer.content
          |> Present.theory_output name
      end;

  in (thy, present, size text) end;


(* document antiquotation *)

val _ = Theory.setup
  (Thy_Output.antiquotation (Binding.name "file") (Scan.lift (Parse.position Parse.path))
    (fn {context = ctxt, ...} => fn (name, pos) =>
      let
        val dir = master_directory (Proof_Context.theory_of ctxt);
        val path = Path.append dir (Path.explode name);
        val _ =
          if File.exists path then ()
          else error ("Bad file: " ^ Path.print (Path.expand path) ^ Position.here pos);
        val _ = Position.report pos (Markup.path name);
      in
        space_explode "/" name
        |> map Thy_Output.verb_text
        |> space_implode (Thy_Output.verb_text "/" ^ "\\discretionary{}{}{}")
      end));

end;