| author | wenzelm | 
| Tue, 19 Feb 2013 21:44:37 +0100 | |
| changeset 51224 | c3e99efacb67 | 
| parent 51217 | 65ab2c4f4c32 | 
| child 51228 | dff3471dd8bc | 
| permissions | -rw-r--r-- | 
| 6168 | 1 | (* Title: Pure/Thy/thy_load.ML | 
| 42002 | 2 | Author: Makarius | 
| 6168 | 3 | |
| 48888 | 4 | Loading files that contribute to a theory. Global master path for TTY loop. | 
| 6168 | 5 | *) | 
| 6 | ||
| 7 | signature THY_LOAD = | |
| 8 | sig | |
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 9 | val master_directory: theory -> Path.T | 
| 48927 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48924diff
changeset | 10 | val imports_of: theory -> (string * Position.T) list | 
| 46737 | 11 | val thy_path: Path.T -> Path.T | 
| 48906 | 12 | val parse_files: string -> (theory -> Token.file list) parser | 
| 48898 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 wenzelm parents: 
48897diff
changeset | 13 | val check_thy: Path.T -> string -> | 
| 48928 | 14 |    {master: Path.T * SHA1.digest, text: string, theory_pos: Position.T,
 | 
| 15 | imports: (string * Position.T) list, uses: (Path.T * bool) list, keywords: Thy_Header.keywords} | |
| 48906 | 16 | val provide: Path.T * SHA1.digest -> theory -> theory | 
| 17 | val provide_parse_files: string -> (theory -> Token.file list * theory) parser | |
| 43702 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43700diff
changeset | 18 | val load_file: theory -> Path.T -> (Path.T * SHA1.digest) * string | 
| 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43700diff
changeset | 19 | val use_file: Path.T -> theory -> string * theory | 
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 20 | val loaded_files: theory -> Path.T list | 
| 48886 
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
 wenzelm parents: 
48881diff
changeset | 21 | val load_current: theory -> bool | 
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 22 | val use_ml: Path.T -> unit | 
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 23 | val exec_ml: Path.T -> generic_theory -> generic_theory | 
| 46938 
cda018294515
some support for outer syntax keyword declarations within theory header;
 wenzelm parents: 
46811diff
changeset | 24 | val begin_theory: Path.T -> Thy_Header.header -> theory list -> theory | 
| 51217 
65ab2c4f4c32
support for prescient timing information within command transactions;
 wenzelm parents: 
50201diff
changeset | 25 | val load_thy: (Toplevel.transition -> Timing.timing) -> int -> Path.T -> Thy_Header.header -> | 
| 
65ab2c4f4c32
support for prescient timing information within command transactions;
 wenzelm parents: 
50201diff
changeset | 26 | Position.T -> string -> theory list -> theory * unit future | 
| 42002 | 27 | val set_master_path: Path.T -> unit | 
| 28 | val get_master_path: unit -> Path.T | |
| 6168 | 29 | end; | 
| 30 | ||
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
33221diff
changeset | 31 | structure Thy_Load: THY_LOAD = | 
| 6168 | 32 | struct | 
| 33 | ||
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 34 | (* manage source files *) | 
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 35 | |
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 36 | type files = | 
| 40741 
17d6293a1e26
moved file identification to thy_load.ML (where it is actually used);
 wenzelm parents: 
40625diff
changeset | 37 |  {master_dir: Path.T,  (*master directory of theory source*)
 | 
| 48927 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48924diff
changeset | 38 | imports: (string * Position.T) list, (*source specification of imports*) | 
| 48896 
bb1f461a7815
simplified Thy_Load.provide: do not store full path;
 wenzelm parents: 
48888diff
changeset | 39 | provided: (Path.T * SHA1.digest) list}; (*source path, digest*) | 
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 40 | |
| 48886 
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
 wenzelm parents: 
48881diff
changeset | 41 | fun make_files (master_dir, imports, provided): files = | 
| 
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
 wenzelm parents: 
48881diff
changeset | 42 |  {master_dir = master_dir, imports = imports, provided = provided};
 | 
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 43 | |
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 44 | structure Files = Theory_Data | 
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 45 | ( | 
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 46 | type T = files; | 
| 48886 
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
 wenzelm parents: 
48881diff
changeset | 47 | val empty = make_files (Path.current, [], []); | 
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 48 | fun extend _ = empty; | 
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 49 | fun merge _ = empty; | 
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 50 | ); | 
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 51 | |
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 52 | fun map_files f = | 
| 48886 
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
 wenzelm parents: 
48881diff
changeset | 53 |   Files.map (fn {master_dir, imports, provided} =>
 | 
| 
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
 wenzelm parents: 
48881diff
changeset | 54 | make_files (f (master_dir, imports, provided))); | 
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 55 | |
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 56 | |
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 57 | val master_directory = #master_dir o Files.get; | 
| 41548 
bd0bebf93fa6
Thy_Load.begin_theory: maintain source specification of imports;
 wenzelm parents: 
41414diff
changeset | 58 | val imports_of = #imports o Files.get; | 
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 59 | |
| 48888 | 60 | fun put_deps master_dir imports = map_files (fn _ => (master_dir, imports, [])); | 
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 61 | |
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 62 | |
| 48869 | 63 | (* inlined files *) | 
| 64 | ||
| 48878 | 65 | fun check_file dir file = File.check_file (File.full_path dir file); | 
| 66 | ||
| 48919 | 67 | fun read_files cmd dir (path, pos) = | 
| 48898 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 wenzelm parents: 
48897diff
changeset | 68 | let | 
| 48919 | 69 | fun make_file file = | 
| 70 | let | |
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49244diff
changeset | 71 | val _ = Position.report pos (Markup.path (Path.implode file)); | 
| 48919 | 72 | val full_path = check_file dir file; | 
| 73 |       in {src_path = file, text = File.read full_path, pos = Path.position full_path} end;
 | |
| 48898 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 wenzelm parents: 
48897diff
changeset | 74 | val paths = | 
| 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 wenzelm parents: 
48897diff
changeset | 75 | (case Keyword.command_files cmd of | 
| 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 wenzelm parents: 
48897diff
changeset | 76 | [] => [path] | 
| 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 wenzelm parents: 
48897diff
changeset | 77 | | exts => map (fn ext => Path.ext ext path) exts); | 
| 48919 | 78 | in map make_file paths end; | 
| 48898 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 wenzelm parents: 
48897diff
changeset | 79 | |
| 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 wenzelm parents: 
48897diff
changeset | 80 | fun parse_files cmd = | 
| 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 wenzelm parents: 
48897diff
changeset | 81 | Scan.ahead Parse.not_eof -- Parse.path >> (fn (tok, name) => fn thy => | 
| 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 wenzelm parents: 
48897diff
changeset | 82 | (case Token.get_files tok of | 
| 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 wenzelm parents: 
48897diff
changeset | 83 | SOME files => files | 
| 48919 | 84 | | NONE => read_files cmd (master_directory thy) (Path.explode name, Token.position_of tok))); | 
| 48898 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 wenzelm parents: 
48897diff
changeset | 85 | |
| 48874 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48869diff
changeset | 86 | local | 
| 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48869diff
changeset | 87 | |
| 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48869diff
changeset | 88 | fun clean ((i1, t1) :: (i2, t2) :: toks) = | 
| 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48869diff
changeset | 89 | if Token.keyword_with (fn s => s = "%" orelse s = "--") t1 then clean toks | 
| 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48869diff
changeset | 90 | else (i1, t1) :: clean ((i2, t2) :: toks) | 
| 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48869diff
changeset | 91 | | clean toks = toks; | 
| 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48869diff
changeset | 92 | |
| 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48869diff
changeset | 93 | fun clean_tokens toks = | 
| 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48869diff
changeset | 94 | ((0 upto length toks - 1) ~~ toks) | 
| 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48869diff
changeset | 95 | |> filter (fn (_, tok) => Token.is_proper tok) | 
| 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48869diff
changeset | 96 | |> clean; | 
| 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48869diff
changeset | 97 | |
| 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48869diff
changeset | 98 | fun find_file toks = | 
| 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48869diff
changeset | 99 | rev (clean_tokens toks) |> get_first (fn (i, tok) => | 
| 48881 
46e053eda5dd
clarified Parse.path vs. Parse.explode -- prefer errors in proper transaction context;
 wenzelm parents: 
48880diff
changeset | 100 | if Token.is_name tok then | 
| 48919 | 101 | SOME (i, (Path.explode (Token.content_of tok), Token.position_of tok)) | 
| 48886 
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
 wenzelm parents: 
48881diff
changeset | 102 | handle ERROR msg => error (msg ^ Token.pos_of tok) | 
| 48874 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48869diff
changeset | 103 | else NONE); | 
| 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48869diff
changeset | 104 | |
| 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48869diff
changeset | 105 | in | 
| 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48869diff
changeset | 106 | |
| 48888 | 107 | fun resolve_files master_dir span = | 
| 48874 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48869diff
changeset | 108 | (case span of | 
| 48878 | 109 | Thy_Syntax.Span (Thy_Syntax.Command (cmd, pos), toks) => | 
| 48869 | 110 | if Keyword.is_theory_load cmd then | 
| 48880 | 111 | (case find_file toks of | 
| 48992 | 112 |           NONE => error ("Bad file argument of command " ^ quote cmd ^ Position.here pos)
 | 
| 48874 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48869diff
changeset | 113 | | SOME (i, path) => | 
| 48869 | 114 | let | 
| 115 | val toks' = toks |> map_index (fn (j, tok) => | |
| 48888 | 116 | if i = j then Token.put_files (read_files cmd master_dir path) tok | 
| 48869 | 117 | else tok); | 
| 48878 | 118 | in Thy_Syntax.Span (Thy_Syntax.Command (cmd, pos), toks') end) | 
| 48869 | 119 | else span | 
| 48874 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48869diff
changeset | 120 | | span => span); | 
| 48869 | 121 | |
| 122 | end; | |
| 123 | ||
| 124 | ||
| 23872 | 125 | (* check files *) | 
| 6232 | 126 | |
| 46737 | 127 | val thy_path = Path.ext "thy"; | 
| 128 | ||
| 48898 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 wenzelm parents: 
48897diff
changeset | 129 | fun check_thy dir thy_name = | 
| 23872 | 130 | let | 
| 48876 
157dd47032e0
more standard Thy_Load.check_thy for Pure.thy, relying on its header;
 wenzelm parents: 
48874diff
changeset | 131 | val path = thy_path (Path.basic thy_name); | 
| 46737 | 132 | val master_file = check_file dir path; | 
| 42003 
6e45dc518ebb
replaced File.check by specific File.check_file, File.check_dir;
 wenzelm parents: 
42002diff
changeset | 133 | val text = File.read master_file; | 
| 48876 
157dd47032e0
more standard Thy_Load.check_thy for Pure.thy, relying on its header;
 wenzelm parents: 
48874diff
changeset | 134 | |
| 48927 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48924diff
changeset | 135 |     val {name = (name, pos), imports, uses, keywords} =
 | 
| 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48924diff
changeset | 136 | Thy_Header.read (Path.position master_file) text; | 
| 48876 
157dd47032e0
more standard Thy_Load.check_thy for Pure.thy, relying on its header;
 wenzelm parents: 
48874diff
changeset | 137 | val _ = thy_name <> name andalso | 
| 48992 | 138 |       error ("Bad file name " ^ Path.print path ^ " for theory " ^ quote name ^ Position.here pos);
 | 
| 48874 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48869diff
changeset | 139 | in | 
| 48928 | 140 |    {master = (master_file, SHA1.digest text), text = text, theory_pos = pos,
 | 
| 48898 
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
 wenzelm parents: 
48897diff
changeset | 141 | imports = imports, uses = uses, keywords = keywords} | 
| 48874 
ff9cd47be39b
refined Thy_Load.check_thy: find more uses in body text, based on keywords;
 wenzelm parents: 
48869diff
changeset | 142 | end; | 
| 6168 | 143 | |
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 144 | |
| 43702 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43700diff
changeset | 145 | (* load files *) | 
| 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43700diff
changeset | 146 | |
| 48906 | 147 | fun provide (src_path, id) = | 
| 148 | map_files (fn (master_dir, imports, provided) => | |
| 149 | if AList.defined (op =) provided src_path then | |
| 150 |       error ("Duplicate use of source file: " ^ Path.print src_path)
 | |
| 151 | else (master_dir, imports, (src_path, id) :: provided)); | |
| 152 | ||
| 153 | fun provide_parse_files cmd = | |
| 154 | parse_files cmd >> (fn files => fn thy => | |
| 155 | let | |
| 156 | val fs = files thy; | |
| 157 |       val thy' = fold (fn {src_path, text, ...} => provide (src_path, SHA1.digest text)) fs thy;
 | |
| 158 | in (fs, thy') end); | |
| 159 | ||
| 43702 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43700diff
changeset | 160 | fun load_file thy src_path = | 
| 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43700diff
changeset | 161 | let | 
| 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43700diff
changeset | 162 | val full_path = check_file (master_directory thy) src_path; | 
| 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43700diff
changeset | 163 | val text = File.read full_path; | 
| 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43700diff
changeset | 164 | val id = SHA1.digest text; | 
| 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43700diff
changeset | 165 | in ((full_path, id), text) end; | 
| 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43700diff
changeset | 166 | |
| 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43700diff
changeset | 167 | fun use_file src_path thy = | 
| 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43700diff
changeset | 168 | let | 
| 48896 
bb1f461a7815
simplified Thy_Load.provide: do not store full path;
 wenzelm parents: 
48888diff
changeset | 169 | val ((_, id), text) = load_file thy src_path; | 
| 
bb1f461a7815
simplified Thy_Load.provide: do not store full path;
 wenzelm parents: 
48888diff
changeset | 170 | val thy' = provide (src_path, id) thy; | 
| 43702 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43700diff
changeset | 171 | in (text, thy') end; | 
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 172 | |
| 48896 
bb1f461a7815
simplified Thy_Load.provide: do not store full path;
 wenzelm parents: 
48888diff
changeset | 173 | fun loaded_files thy = | 
| 
bb1f461a7815
simplified Thy_Load.provide: do not store full path;
 wenzelm parents: 
48888diff
changeset | 174 |   let val {master_dir, provided, ...} = Files.get thy
 | 
| 
bb1f461a7815
simplified Thy_Load.provide: do not store full path;
 wenzelm parents: 
48888diff
changeset | 175 | in map (File.full_path master_dir o #1) provided end; | 
| 6168 | 176 | |
| 48886 
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
 wenzelm parents: 
48881diff
changeset | 177 | fun load_current thy = | 
| 
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
 wenzelm parents: 
48881diff
changeset | 178 | #provided (Files.get thy) |> | 
| 48896 
bb1f461a7815
simplified Thy_Load.provide: do not store full path;
 wenzelm parents: 
48888diff
changeset | 179 | forall (fn (src_path, id) => | 
| 43702 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43700diff
changeset | 180 | (case try (load_file thy) src_path of | 
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 181 | NONE => false | 
| 48886 
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
 wenzelm parents: 
48881diff
changeset | 182 | | SOME ((_, id'), _) => id = id')); | 
| 37952 | 183 | |
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 184 | |
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 185 | (* provide files *) | 
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 186 | |
| 43700 
e4ece46a9ca7
clarified Thy_Load.digest_file -- read ML files only once;
 wenzelm parents: 
42003diff
changeset | 187 | fun eval_file path text = ML_Context.eval_text true (Path.position path) text; | 
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 188 | |
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 189 | fun use_ml src_path = | 
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 190 | if is_none (Context.thread_data ()) then | 
| 43700 
e4ece46a9ca7
clarified Thy_Load.digest_file -- read ML files only once;
 wenzelm parents: 
42003diff
changeset | 191 | let val path = check_file Path.current src_path | 
| 
e4ece46a9ca7
clarified Thy_Load.digest_file -- read ML files only once;
 wenzelm parents: 
42003diff
changeset | 192 | in eval_file path (File.read path) end | 
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 193 | else | 
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 194 | let | 
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 195 | val thy = ML_Context.the_global_context (); | 
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 196 | |
| 43702 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 wenzelm parents: 
43700diff
changeset | 197 | val ((path, id), text) = load_file thy src_path; | 
| 43700 
e4ece46a9ca7
clarified Thy_Load.digest_file -- read ML files only once;
 wenzelm parents: 
42003diff
changeset | 198 | val _ = eval_file path text; | 
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 199 | val _ = Context.>> Local_Theory.propagate_ml_env; | 
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 200 | |
| 48896 
bb1f461a7815
simplified Thy_Load.provide: do not store full path;
 wenzelm parents: 
48888diff
changeset | 201 | val provide = provide (src_path, id); | 
| 38757 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 wenzelm parents: 
38133diff
changeset | 202 | val _ = Context.>> (Context.mapping provide (Local_Theory.background_theory provide)); | 
| 37978 
548f3f165d05
simplified Thy_Header.read -- include Source.of_string_limited here;
 wenzelm parents: 
37977diff
changeset | 203 | in () end; | 
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 204 | |
| 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
37943diff
changeset | 205 | fun exec_ml src_path = ML_Context.exec (fn () => use_ml src_path); | 
| 6168 | 206 | |
| 37977 
3ceccd415145
simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
 wenzelm parents: 
37952diff
changeset | 207 | |
| 43712 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 208 | (* load_thy *) | 
| 37977 
3ceccd415145
simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
 wenzelm parents: 
37952diff
changeset | 209 | |
| 48888 | 210 | fun begin_theory master_dir {name, imports, keywords, uses} parents =
 | 
| 37977 
3ceccd415145
simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
 wenzelm parents: 
37952diff
changeset | 211 | Theory.begin_theory name parents | 
| 48888 | 212 | |> put_deps master_dir imports | 
| 46938 
cda018294515
some support for outer syntax keyword declarations within theory header;
 wenzelm parents: 
46811diff
changeset | 213 | |> fold Thy_Header.declare_keyword keywords | 
| 37977 
3ceccd415145
simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
 wenzelm parents: 
37952diff
changeset | 214 | |> fold (fn (path, true) => Context.theory_map (exec_ml path) o Theory.checkpoint | _ => I) uses | 
| 
3ceccd415145
simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
 wenzelm parents: 
37952diff
changeset | 215 | |> Theory.checkpoint; | 
| 
3ceccd415145
simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
 wenzelm parents: 
37952diff
changeset | 216 | |
| 51217 
65ab2c4f4c32
support for prescient timing information within command transactions;
 wenzelm parents: 
50201diff
changeset | 217 | fun excursion last_timing init elements = | 
| 46959 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 218 | let | 
| 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 219 | val immediate = not (Goal.future_enabled ()); | 
| 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 220 | |
| 51217 
65ab2c4f4c32
support for prescient timing information within command transactions;
 wenzelm parents: 
50201diff
changeset | 221 | fun put_timing tr = Toplevel.put_timing (last_timing tr) tr; | 
| 
65ab2c4f4c32
support for prescient timing information within command transactions;
 wenzelm parents: 
50201diff
changeset | 222 | fun put_timings (tr, trs) = (put_timing tr, map put_timing trs); | 
| 
65ab2c4f4c32
support for prescient timing information within command transactions;
 wenzelm parents: 
50201diff
changeset | 223 | |
| 46959 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 224 | fun proof_result (tr, trs) (st, _) = | 
| 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 225 | let | 
| 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 226 | val (result, st') = Toplevel.proof_result immediate (tr, trs) st; | 
| 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 227 | val pos' = Toplevel.pos_of (List.last (tr :: trs)); | 
| 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 228 | in (result, (st', pos')) end; | 
| 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 229 | |
| 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 230 | fun element_result elem x = | 
| 51217 
65ab2c4f4c32
support for prescient timing information within command transactions;
 wenzelm parents: 
50201diff
changeset | 231 | fold_map (proof_result o put_timings) | 
| 46959 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 232 | (Outer_Syntax.read_element (#2 (Outer_Syntax.get_syntax ())) init elem) x; | 
| 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 233 | |
| 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 234 | val (results, (end_state, end_pos)) = | 
| 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 235 | fold_map element_result elements (Toplevel.toplevel, Position.none); | 
| 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 236 | |
| 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 237 | val thy = Toplevel.end_theory end_pos end_state; | 
| 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 238 | in (flat results, thy) end; | 
| 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 239 | |
| 51217 
65ab2c4f4c32
support for prescient timing information within command transactions;
 wenzelm parents: 
50201diff
changeset | 240 | fun load_thy last_timing update_time master_dir header text_pos text parents = | 
| 43712 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 241 | let | 
| 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 242 | val time = ! Toplevel.timing; | 
| 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 243 | |
| 48927 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48924diff
changeset | 244 |     val {name = (name, _), uses, ...} = header;
 | 
| 46958 
0ec8f04e753a
define keywords early when processing the theory header, before running the body commands;
 wenzelm parents: 
46938diff
changeset | 245 | val _ = Thy_Header.define_keywords header; | 
| 43712 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 246 | val _ = Present.init_theory name; | 
| 44186 
806f0ec1a43d
simplified Toplevel.init_theory: discontinued special master argument;
 wenzelm parents: 
44113diff
changeset | 247 | fun init () = | 
| 48888 | 248 | begin_theory master_dir header parents | 
| 249 | |> Present.begin_theory update_time master_dir uses; | |
| 43712 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 250 | |
| 46959 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 251 | val lexs = Keyword.get_lexicons (); | 
| 46958 
0ec8f04e753a
define keywords early when processing the theory header, before running the body commands;
 wenzelm parents: 
46938diff
changeset | 252 | |
| 48927 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48924diff
changeset | 253 | val toks = Thy_Syntax.parse_tokens lexs text_pos text; | 
| 48888 | 254 | val spans = map (resolve_files master_dir) (Thy_Syntax.parse_spans toks); | 
| 46959 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 255 | val elements = Thy_Syntax.parse_elements spans; | 
| 43712 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 256 | |
| 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 257 | val _ = Present.theory_source name | 
| 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 258 | (fn () => HTML.html_mode (implode o map Thy_Syntax.present_span) spans); | 
| 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 259 | |
| 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 260 |     val _ = if time then writeln ("\n**** Starting theory " ^ quote name ^ " ****") else ();
 | 
| 51217 
65ab2c4f4c32
support for prescient timing information within command transactions;
 wenzelm parents: 
50201diff
changeset | 261 | val (results, thy) = cond_timeit time "" (fn () => excursion last_timing init elements); | 
| 43712 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 262 |     val _ = if time then writeln ("**** Finished theory " ^ quote name ^ " ****\n") else ();
 | 
| 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 263 | |
| 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 264 | val present = | 
| 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 265 |       singleton (Future.cond_forks {name = "Outer_Syntax.present:" ^ name, group = NONE,
 | 
| 44113 | 266 | deps = map Future.task_of results, pri = 0, interrupts = true}) | 
| 43712 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 267 | (fn () => | 
| 46959 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 268 | let val ((minor, _), outer_syntax) = Outer_Syntax.get_syntax () in | 
| 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 269 | Thy_Output.present_thy minor Keyword.command_tags | 
| 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 270 | (Outer_Syntax.is_markup outer_syntax) | 
| 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 271 | (maps Future.join results) toks | 
| 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 272 | |> Buffer.content | 
| 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 273 | |> Present.theory_output name | 
| 
cdc791910460
defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
 wenzelm parents: 
46958diff
changeset | 274 | end); | 
| 43712 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 275 | |
| 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 276 | in (thy, present) end; | 
| 
3c2c912af2ef
moved Outer_Syntax.load_thy to Thy_Load.load_thy;
 wenzelm parents: 
43702diff
changeset | 277 | |
| 42002 | 278 | |
| 49244 
fb669aff821e
formal markup for @{file} (for hyperlinks etc.) -- interpret path wrt. master directory as usual;
 wenzelm parents: 
48992diff
changeset | 279 | (* document antiquotation *) | 
| 
fb669aff821e
formal markup for @{file} (for hyperlinks etc.) -- interpret path wrt. master directory as usual;
 wenzelm parents: 
48992diff
changeset | 280 | |
| 
fb669aff821e
formal markup for @{file} (for hyperlinks etc.) -- interpret path wrt. master directory as usual;
 wenzelm parents: 
48992diff
changeset | 281 | val _ = | 
| 
fb669aff821e
formal markup for @{file} (for hyperlinks etc.) -- interpret path wrt. master directory as usual;
 wenzelm parents: 
48992diff
changeset | 282 | Context.>> (Context.map_theory | 
| 
fb669aff821e
formal markup for @{file} (for hyperlinks etc.) -- interpret path wrt. master directory as usual;
 wenzelm parents: 
48992diff
changeset | 283 | (Thy_Output.antiquotation (Binding.name "file") (Scan.lift (Parse.position Parse.path)) | 
| 
fb669aff821e
formal markup for @{file} (for hyperlinks etc.) -- interpret path wrt. master directory as usual;
 wenzelm parents: 
48992diff
changeset | 284 |     (fn {context = ctxt, ...} => fn (name, pos) =>
 | 
| 
fb669aff821e
formal markup for @{file} (for hyperlinks etc.) -- interpret path wrt. master directory as usual;
 wenzelm parents: 
48992diff
changeset | 285 | let | 
| 
fb669aff821e
formal markup for @{file} (for hyperlinks etc.) -- interpret path wrt. master directory as usual;
 wenzelm parents: 
48992diff
changeset | 286 | val dir = master_directory (Proof_Context.theory_of ctxt); | 
| 
fb669aff821e
formal markup for @{file} (for hyperlinks etc.) -- interpret path wrt. master directory as usual;
 wenzelm parents: 
48992diff
changeset | 287 | val path = Path.append dir (Path.explode name); | 
| 
fb669aff821e
formal markup for @{file} (for hyperlinks etc.) -- interpret path wrt. master directory as usual;
 wenzelm parents: 
48992diff
changeset | 288 | val _ = | 
| 
fb669aff821e
formal markup for @{file} (for hyperlinks etc.) -- interpret path wrt. master directory as usual;
 wenzelm parents: 
48992diff
changeset | 289 | if File.exists path then () | 
| 
fb669aff821e
formal markup for @{file} (for hyperlinks etc.) -- interpret path wrt. master directory as usual;
 wenzelm parents: 
48992diff
changeset | 290 |           else error ("Bad file: " ^ Path.print (Path.expand path) ^ Position.here pos);
 | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49244diff
changeset | 291 | val _ = Position.report pos (Markup.path name); | 
| 49244 
fb669aff821e
formal markup for @{file} (for hyperlinks etc.) -- interpret path wrt. master directory as usual;
 wenzelm parents: 
48992diff
changeset | 292 | in Thy_Output.verb_text name end))); | 
| 
fb669aff821e
formal markup for @{file} (for hyperlinks etc.) -- interpret path wrt. master directory as usual;
 wenzelm parents: 
48992diff
changeset | 293 | |
| 
fb669aff821e
formal markup for @{file} (for hyperlinks etc.) -- interpret path wrt. master directory as usual;
 wenzelm parents: 
48992diff
changeset | 294 | |
| 42002 | 295 | (* global master path *) | 
| 296 | ||
| 297 | local | |
| 298 | val master_path = Unsynchronized.ref Path.current; | |
| 299 | in | |
| 300 | ||
| 301 | fun set_master_path path = master_path := path; | |
| 302 | fun get_master_path () = ! master_path; | |
| 303 | ||
| 6168 | 304 | end; | 
| 305 | ||
| 42002 | 306 | end; |