author | wenzelm |
Fri, 08 Nov 2024 22:52:29 +0100 | |
changeset 81409 | 07c802837a8c |
parent 79502 | c7a98469c0e7 |
permissions | -rw-r--r-- |
79502 | 1 |
(* Title: Pure/Build/resources.ML |
42002 | 2 |
Author: Makarius |
6168 | 3 |
|
56208 | 4 |
Resources for theories and auxiliary files. |
6168 | 5 |
*) |
6 |
||
56208 | 7 |
signature RESOURCES = |
6168 | 8 |
sig |
65532 | 9 |
val default_qualifier: string |
71875 | 10 |
val init_session: |
72669 | 11 |
{session_positions: (string * Properties.T) list, |
70683
8c7706b053c7
find theory files via session structure: much faster Prover IDE startup;
wenzelm
parents:
70049
diff
changeset
|
12 |
session_directories: (string * string) list, |
72638
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
13 |
command_timings: Properties.T list, |
74671 | 14 |
load_commands: (string * Position.T) list, |
75586 | 15 |
scala_functions: (string * ((bool * bool) * Position.T)) list, |
67219 | 16 |
global_theories: (string * string) list, |
70712
a3cfe859d915
find theories via session directories only -- ignore known_theories;
wenzelm
parents:
70683
diff
changeset
|
17 |
loaded_theories: string list} -> unit |
75626 | 18 |
val init_session_yxml: Bytes.T -> unit |
75590
99b7638d9177
clarified session resources for bootstrap, notably for Scala functions;
wenzelm
parents:
75586
diff
changeset
|
19 |
val init_session_env: unit -> unit |
65478
7c40477e0a87
clarified init_session_base / finish_session_base: retain some information for plain "isabelle process", without rechecking dependencies as in "isabelle console";
wenzelm
parents:
65476
diff
changeset
|
20 |
val finish_session_base: unit -> unit |
65457 | 21 |
val global_theory: string -> string option |
66712 | 22 |
val loaded_theory: string -> bool |
67219 | 23 |
val check_session: Proof.context -> string * Position.T -> string |
72638
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
24 |
val last_timing: Toplevel.transition -> Time.time |
74671 | 25 |
val check_load_command: Proof.context -> string * Position.T -> string |
75586 | 26 |
val check_scala_function: Proof.context -> string * Position.T -> string * (bool * bool) |
37949
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents:
37943
diff
changeset
|
27 |
val master_directory: theory -> Path.T |
48927
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
wenzelm
parents:
48924
diff
changeset
|
28 |
val imports_of: theory -> (string * Position.T) list |
65505 | 29 |
val begin_theory: Path.T -> Thy_Header.header -> theory list -> theory |
46737 | 30 |
val thy_path: Path.T -> Path.T |
65445
e9e7f5f5794c
more qualifier treatment, but in the end it is still ignored;
wenzelm
parents:
65443
diff
changeset
|
31 |
val theory_qualifier: string -> string |
70683
8c7706b053c7
find theory files via session structure: much faster Prover IDE startup;
wenzelm
parents:
70049
diff
changeset
|
32 |
val find_theory_file: string -> Path.T option |
67215 | 33 |
val import_name: string -> Path.T -> string -> |
34 |
{node_name: Path.T, master_dir: Path.T, theory_name: string} |
|
48898
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
wenzelm
parents:
48897
diff
changeset
|
35 |
val check_thy: Path.T -> string -> |
48928 | 36 |
{master: Path.T * SHA1.digest, text: string, theory_pos: Position.T, |
51293
05b1bbae748d
discontinued obsolete 'uses' within theory header;
wenzelm
parents:
51273
diff
changeset
|
37 |
imports: (string * Position.T) list, keywords: Thy_Header.keywords} |
76806 | 38 |
val read_file_node: string -> Path.T -> Path.T * Position.T -> Token.file |
76818 | 39 |
val read_file: Path.T -> Path.T * Position.T -> Token.file |
76807 | 40 |
val parsed_files: (Path.T -> Path.T list) -> |
41 |
Token.file Exn.result list * Input.source -> theory -> Token.file list |
|
72747
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
wenzelm
parents:
72669
diff
changeset
|
42 |
val parse_files: (Path.T -> Path.T list) -> (theory -> Token.file list) parser |
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
wenzelm
parents:
72669
diff
changeset
|
43 |
val parse_file: (theory -> Token.file) parser |
48906 | 44 |
val provide: Path.T * SHA1.digest -> theory -> theory |
69851 | 45 |
val provide_file: Token.file -> theory -> theory |
76801 | 46 |
val provide_file': Token.file -> theory -> Token.file * theory |
72747
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
wenzelm
parents:
72669
diff
changeset
|
47 |
val provide_parse_files: (Path.T -> Path.T list) -> (theory -> Token.file list * theory) parser |
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
wenzelm
parents:
72669
diff
changeset
|
48 |
val provide_parse_file: (theory -> Token.file * theory) parser |
54446 | 49 |
val loaded_files_current: theory -> bool |
72841 | 50 |
val check_path: Proof.context -> Path.T option -> Input.source -> Path.T |
51 |
val check_file: Proof.context -> Path.T option -> Input.source -> Path.T |
|
52 |
val check_dir: Proof.context -> Path.T option -> Input.source -> Path.T |
|
73312 | 53 |
val check_session_dir: Proof.context -> Path.T option -> Input.source -> Path.T |
6168 | 54 |
end; |
55 |
||
56208 | 56 |
structure Resources: RESOURCES = |
6168 | 57 |
struct |
58 |
||
72638
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
59 |
(* command timings *) |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
60 |
|
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
61 |
type timings = ((string * Time.time) Inttab.table) Symtab.table; (*file -> offset -> name, time*) |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
62 |
|
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
63 |
val empty_timings: timings = Symtab.empty; |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
64 |
|
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
65 |
fun update_timings props = |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
66 |
(case Markup.parse_command_timing_properties props of |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
67 |
SOME ({file, offset, name}, time) => |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
68 |
Symtab.map_default (file, Inttab.empty) |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
69 |
(Inttab.map_default (offset, (name, time)) (fn (_, t) => (name, t + time))) |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
70 |
| NONE => I); |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
71 |
|
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
72 |
fun make_timings command_timings = |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
73 |
fold update_timings command_timings empty_timings; |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
74 |
|
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
75 |
fun approximative_id name pos = |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
76 |
(case (Position.file_of pos, Position.offset_of pos) of |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
77 |
(SOME file, SOME offset) => |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
78 |
if name = "" then NONE else SOME {file = file, offset = offset, name = name} |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
79 |
| _ => NONE); |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
80 |
|
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
81 |
fun get_timings timings tr = |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
82 |
(case approximative_id (Toplevel.name_of tr) (Toplevel.pos_of tr) of |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
83 |
SOME {file, offset, name} => |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
84 |
(case Symtab.lookup timings file of |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
85 |
SOME offsets => |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
86 |
(case Inttab.lookup offsets offset of |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
87 |
SOME (name', time) => if name = name' then SOME time else NONE |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
88 |
| NONE => NONE) |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
89 |
| NONE => NONE) |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
90 |
| NONE => NONE) |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
91 |
|> the_default Time.zeroTime; |
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
92 |
|
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
93 |
|
65431
4a3e6cda3b94
provide session base for "isabelle build" and "isabelle console" ML process;
wenzelm
parents:
65058
diff
changeset
|
94 |
(* session base *) |
4a3e6cda3b94
provide session base for "isabelle build" and "isabelle console" ML process;
wenzelm
parents:
65058
diff
changeset
|
95 |
|
65532 | 96 |
val default_qualifier = "Draft"; |
97 |
||
67493
c4e9e0c50487
treat sessions as entities with defining position;
wenzelm
parents:
67473
diff
changeset
|
98 |
type entry = {pos: Position.T, serial: serial}; |
c4e9e0c50487
treat sessions as entities with defining position;
wenzelm
parents:
67473
diff
changeset
|
99 |
|
c4e9e0c50487
treat sessions as entities with defining position;
wenzelm
parents:
67473
diff
changeset
|
100 |
fun make_entry props : entry = |
c4e9e0c50487
treat sessions as entities with defining position;
wenzelm
parents:
67473
diff
changeset
|
101 |
{pos = Position.of_properties props, serial = serial ()}; |
c4e9e0c50487
treat sessions as entities with defining position;
wenzelm
parents:
67473
diff
changeset
|
102 |
|
65478
7c40477e0a87
clarified init_session_base / finish_session_base: retain some information for plain "isabelle process", without rechecking dependencies as in "isabelle console";
wenzelm
parents:
65476
diff
changeset
|
103 |
val empty_session_base = |
72758 | 104 |
({session_positions = []: (string * entry) list, |
105 |
session_directories = Symtab.empty: Path.T list Symtab.table, |
|
106 |
timings = empty_timings, |
|
74671 | 107 |
load_commands = []: (string * Position.T) list, |
75586 | 108 |
scala_functions = Symtab.empty: ((bool * bool) * Position.T) Symtab.table}, |
72758 | 109 |
{global_theories = Symtab.empty: string Symtab.table, |
77723
b761c91c2447
performance tuning: prefer functor Set() over Table();
wenzelm
parents:
77028
diff
changeset
|
110 |
loaded_theories = Symset.empty: Symset.T}); |
65431
4a3e6cda3b94
provide session base for "isabelle build" and "isabelle console" ML process;
wenzelm
parents:
65058
diff
changeset
|
111 |
|
65441 | 112 |
val global_session_base = |
65478
7c40477e0a87
clarified init_session_base / finish_session_base: retain some information for plain "isabelle process", without rechecking dependencies as in "isabelle console";
wenzelm
parents:
65476
diff
changeset
|
113 |
Synchronized.var "Sessions.base" empty_session_base; |
65441 | 114 |
|
71875 | 115 |
fun init_session |
77028
f5896dea6fce
more direct check of bibtex entries via Isabelle/Scala;
wenzelm
parents:
76884
diff
changeset
|
116 |
{session_positions, session_directories, command_timings, load_commands, |
74696 | 117 |
scala_functions, global_theories, loaded_theories} = |
65431
4a3e6cda3b94
provide session base for "isabelle build" and "isabelle console" ML process;
wenzelm
parents:
65058
diff
changeset
|
118 |
Synchronized.change global_session_base |
65441 | 119 |
(fn _ => |
72758 | 120 |
({session_positions = sort_by #1 (map (apsnd make_entry) session_positions), |
121 |
session_directories = |
|
74232 | 122 |
Symtab.build (session_directories |> fold_rev (fn (dir, name) => |
123 |
Symtab.cons_list (name, Path.explode dir))), |
|
72758 | 124 |
timings = make_timings command_timings, |
74671 | 125 |
load_commands = load_commands, |
72758 | 126 |
scala_functions = Symtab.make scala_functions}, |
127 |
{global_theories = Symtab.make global_theories, |
|
77723
b761c91c2447
performance tuning: prefer functor Set() over Table();
wenzelm
parents:
77028
diff
changeset
|
128 |
loaded_theories = Symset.make loaded_theories})); |
65431
4a3e6cda3b94
provide session base for "isabelle build" and "isabelle console" ML process;
wenzelm
parents:
65058
diff
changeset
|
129 |
|
72637 | 130 |
fun init_session_yxml yxml = |
131 |
let |
|
77028
f5896dea6fce
more direct check of bibtex entries via Isabelle/Scala;
wenzelm
parents:
76884
diff
changeset
|
132 |
val (session_positions, (session_directories, (command_timings, |
f5896dea6fce
more direct check of bibtex entries via Isabelle/Scala;
wenzelm
parents:
76884
diff
changeset
|
133 |
(load_commands, (scala_functions, (global_theories, loaded_theories)))))) = |
75626 | 134 |
YXML.parse_body_bytes yxml |> |
72637 | 135 |
let open XML.Decode in |
72669 | 136 |
(pair (list (pair string properties)) |
74696 | 137 |
(pair (list (pair string string)) |
77028
f5896dea6fce
more direct check of bibtex entries via Isabelle/Scala;
wenzelm
parents:
76884
diff
changeset
|
138 |
(pair (list properties) |
74671 | 139 |
(pair (list (pair string properties)) |
75586 | 140 |
(pair (list (pair string (pair (pair bool bool) properties))) |
77028
f5896dea6fce
more direct check of bibtex entries via Isabelle/Scala;
wenzelm
parents:
76884
diff
changeset
|
141 |
(pair (list (pair string string)) (list string))))))) |
72637 | 142 |
end; |
143 |
in |
|
144 |
init_session |
|
72669 | 145 |
{session_positions = session_positions, |
72637 | 146 |
session_directories = session_directories, |
72638
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
147 |
command_timings = command_timings, |
74671 | 148 |
load_commands = (map o apsnd) Position.of_properties load_commands, |
73565 | 149 |
scala_functions = (map o apsnd o apsnd) Position.of_properties scala_functions, |
72637 | 150 |
global_theories = global_theories, |
151 |
loaded_theories = loaded_theories} |
|
152 |
end; |
|
153 |
||
75590
99b7638d9177
clarified session resources for bootstrap, notably for Scala functions;
wenzelm
parents:
75586
diff
changeset
|
154 |
fun init_session_env () = |
99b7638d9177
clarified session resources for bootstrap, notably for Scala functions;
wenzelm
parents:
75586
diff
changeset
|
155 |
(case getenv "ISABELLE_INIT_SESSION" of |
99b7638d9177
clarified session resources for bootstrap, notably for Scala functions;
wenzelm
parents:
75586
diff
changeset
|
156 |
"" => () |
99b7638d9177
clarified session resources for bootstrap, notably for Scala functions;
wenzelm
parents:
75586
diff
changeset
|
157 |
| name => |
75626 | 158 |
try Bytes.read (Path.explode name) |
75590
99b7638d9177
clarified session resources for bootstrap, notably for Scala functions;
wenzelm
parents:
75586
diff
changeset
|
159 |
|> Option.app init_session_yxml); |
99b7638d9177
clarified session resources for bootstrap, notably for Scala functions;
wenzelm
parents:
75586
diff
changeset
|
160 |
|
99b7638d9177
clarified session resources for bootstrap, notably for Scala functions;
wenzelm
parents:
75586
diff
changeset
|
161 |
val _ = init_session_env (); |
72637 | 162 |
|
65478
7c40477e0a87
clarified init_session_base / finish_session_base: retain some information for plain "isabelle process", without rechecking dependencies as in "isabelle console";
wenzelm
parents:
65476
diff
changeset
|
163 |
fun finish_session_base () = |
7c40477e0a87
clarified init_session_base / finish_session_base: retain some information for plain "isabelle process", without rechecking dependencies as in "isabelle console";
wenzelm
parents:
65476
diff
changeset
|
164 |
Synchronized.change global_session_base |
72758 | 165 |
(apfst (K (#1 empty_session_base))); |
65441 | 166 |
|
167 |
fun get_session_base f = f (Synchronized.value global_session_base); |
|
72758 | 168 |
fun get_session_base1 f = get_session_base (f o #1); |
169 |
fun get_session_base2 f = get_session_base (f o #2); |
|
65431
4a3e6cda3b94
provide session base for "isabelle build" and "isabelle console" ML process;
wenzelm
parents:
65058
diff
changeset
|
170 |
|
72758 | 171 |
fun global_theory a = Symtab.lookup (get_session_base2 #global_theories) a; |
77723
b761c91c2447
performance tuning: prefer functor Set() over Table();
wenzelm
parents:
77028
diff
changeset
|
172 |
fun loaded_theory a = Symset.member (get_session_base2 #loaded_theories) a; |
65431
4a3e6cda3b94
provide session base for "isabelle build" and "isabelle console" ML process;
wenzelm
parents:
65058
diff
changeset
|
173 |
|
72760
042180540068
clarified protocol: Doc.check at run-time via Scala function;
wenzelm
parents:
72758
diff
changeset
|
174 |
fun check_session ctxt arg = |
042180540068
clarified protocol: Doc.check at run-time via Scala function;
wenzelm
parents:
72758
diff
changeset
|
175 |
Completion.check_item "session" |
71911 | 176 |
(fn (name, {pos, serial}) => |
74262 | 177 |
Position.make_entity_markup {def = false} serial Markup.sessionN (name, pos)) |
72760
042180540068
clarified protocol: Doc.check at run-time via Scala function;
wenzelm
parents:
72758
diff
changeset
|
178 |
(get_session_base1 #session_positions) ctxt arg; |
67493
c4e9e0c50487
treat sessions as entities with defining position;
wenzelm
parents:
67473
diff
changeset
|
179 |
|
72758 | 180 |
fun last_timing tr = get_timings (get_session_base1 #timings) tr; |
72638
2a7fc87495e0
refer to command_timings/last_timing via resources;
wenzelm
parents:
72637
diff
changeset
|
181 |
|
74671 | 182 |
fun check_load_command ctxt arg = |
183 |
Completion.check_entity Markup.load_commandN (get_session_base1 #load_commands) ctxt arg; |
|
184 |
||
73226
4c8edf348c4e
clarified modules: allow early invocation of Scala functions;
wenzelm
parents:
72841
diff
changeset
|
185 |
|
4c8edf348c4e
clarified modules: allow early invocation of Scala functions;
wenzelm
parents:
72841
diff
changeset
|
186 |
(* Scala functions *) |
4c8edf348c4e
clarified modules: allow early invocation of Scala functions;
wenzelm
parents:
72841
diff
changeset
|
187 |
|
4c8edf348c4e
clarified modules: allow early invocation of Scala functions;
wenzelm
parents:
72841
diff
changeset
|
188 |
fun check_scala_function ctxt arg = |
73565 | 189 |
let |
75590
99b7638d9177
clarified session resources for bootstrap, notably for Scala functions;
wenzelm
parents:
75586
diff
changeset
|
190 |
val table = get_session_base1 #scala_functions; |
99b7638d9177
clarified session resources for bootstrap, notably for Scala functions;
wenzelm
parents:
75586
diff
changeset
|
191 |
val funs = Symtab.fold (fn (a, (_, pos)) => cons (a, pos)) table [] |> sort_by #1; |
99b7638d9177
clarified session resources for bootstrap, notably for Scala functions;
wenzelm
parents:
75586
diff
changeset
|
192 |
val name = Completion.check_entity Markup.scala_functionN funs ctxt arg; |
99b7638d9177
clarified session resources for bootstrap, notably for Scala functions;
wenzelm
parents:
75586
diff
changeset
|
193 |
val flags = #1 (the (Symtab.lookup table name)); |
75586 | 194 |
in (name, flags) end; |
73226
4c8edf348c4e
clarified modules: allow early invocation of Scala functions;
wenzelm
parents:
72841
diff
changeset
|
195 |
|
4c8edf348c4e
clarified modules: allow early invocation of Scala functions;
wenzelm
parents:
72841
diff
changeset
|
196 |
val _ = Theory.setup |
73761 | 197 |
(Document_Output.antiquotation_verbatim_embedded \<^binding>\<open>scala_function\<close> |
73565 | 198 |
(Scan.lift Parse.embedded_position) (#1 oo check_scala_function) #> |
73226
4c8edf348c4e
clarified modules: allow early invocation of Scala functions;
wenzelm
parents:
72841
diff
changeset
|
199 |
ML_Antiquotation.inline_embedded \<^binding>\<open>scala_function\<close> |
4c8edf348c4e
clarified modules: allow early invocation of Scala functions;
wenzelm
parents:
72841
diff
changeset
|
200 |
(Args.context -- Scan.lift Parse.embedded_position |
73565 | 201 |
>> (uncurry check_scala_function #> #1 #> ML_Syntax.print_string)) #> |
73226
4c8edf348c4e
clarified modules: allow early invocation of Scala functions;
wenzelm
parents:
72841
diff
changeset
|
202 |
ML_Antiquotation.value_embedded \<^binding>\<open>scala\<close> |
74563 | 203 |
(Args.context -- Scan.lift Parse.embedded_position >> (fn (ctxt, arg) => |
73565 | 204 |
let |
75586 | 205 |
val (name, (single, bytes)) = check_scala_function ctxt arg; |
206 |
val func = |
|
207 |
(if single then "Scala.function1" else "Scala.function") ^ |
|
208 |
(if bytes then "_bytes" else ""); |
|
73565 | 209 |
in ML_Syntax.atomic (func ^ " " ^ ML_Syntax.print_string name) end))); |
72756 | 210 |
|
67471 | 211 |
|
37949
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents:
37943
diff
changeset
|
212 |
(* manage source files *) |
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents:
37943
diff
changeset
|
213 |
|
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents:
37943
diff
changeset
|
214 |
type files = |
40741
17d6293a1e26
moved file identification to thy_load.ML (where it is actually used);
wenzelm
parents:
40625
diff
changeset
|
215 |
{master_dir: Path.T, (*master directory of theory source*) |
48927
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
wenzelm
parents:
48924
diff
changeset
|
216 |
imports: (string * Position.T) list, (*source specification of imports*) |
48896
bb1f461a7815
simplified Thy_Load.provide: do not store full path;
wenzelm
parents:
48888
diff
changeset
|
217 |
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:
37943
diff
changeset
|
218 |
|
48886
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
wenzelm
parents:
48881
diff
changeset
|
219 |
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:
48881
diff
changeset
|
220 |
{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:
37943
diff
changeset
|
221 |
|
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents:
37943
diff
changeset
|
222 |
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:
37943
diff
changeset
|
223 |
( |
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents:
37943
diff
changeset
|
224 |
type T = files; |
48886
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
wenzelm
parents:
48881
diff
changeset
|
225 |
val empty = make_files (Path.current, [], []); |
72045
2c7cfd2f9b6c
more thorough extend/merge, notably for master_dir across Theory.join_theory (e.g. for @{file} antiquotation);
wenzelm
parents:
71912
diff
changeset
|
226 |
fun merge ({master_dir, imports, provided = provided1}, {provided = provided2, ...}) = |
2c7cfd2f9b6c
more thorough extend/merge, notably for master_dir across Theory.join_theory (e.g. for @{file} antiquotation);
wenzelm
parents:
71912
diff
changeset
|
227 |
let val provided' = Library.merge (op =) (provided1, provided2) |
2c7cfd2f9b6c
more thorough extend/merge, notably for master_dir across Theory.join_theory (e.g. for @{file} antiquotation);
wenzelm
parents:
71912
diff
changeset
|
228 |
in make_files (master_dir, imports, provided') end |
37949
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents:
37943
diff
changeset
|
229 |
); |
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents:
37943
diff
changeset
|
230 |
|
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents:
37943
diff
changeset
|
231 |
fun map_files f = |
48886
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
wenzelm
parents:
48881
diff
changeset
|
232 |
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:
48881
diff
changeset
|
233 |
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:
37943
diff
changeset
|
234 |
|
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents:
37943
diff
changeset
|
235 |
|
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents:
37943
diff
changeset
|
236 |
val master_directory = #master_dir o Files.get; |
41548
bd0bebf93fa6
Thy_Load.begin_theory: maintain source specification of imports;
wenzelm
parents:
41414
diff
changeset
|
237 |
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:
37943
diff
changeset
|
238 |
|
65505 | 239 |
fun begin_theory master_dir {name, imports, keywords} parents = |
74671 | 240 |
let |
241 |
val thy = |
|
242 |
Theory.begin_theory name parents |
|
76884 | 243 |
|> map_files (fn _ => (Path.explode (File.symbolic_path master_dir), imports, [])) |
74671 | 244 |
|> Thy_Header.add_keywords keywords; |
245 |
val ctxt = Proof_Context.init_global thy; |
|
246 |
val _ = List.app (ignore o check_load_command ctxt o #load_command o #2) keywords; |
|
247 |
in thy end; |
|
37949
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents:
37943
diff
changeset
|
248 |
|
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents:
37943
diff
changeset
|
249 |
|
54525 | 250 |
(* theory files *) |
251 |
||
252 |
val thy_path = Path.ext "thy"; |
|
48869 | 253 |
|
65445
e9e7f5f5794c
more qualifier treatment, but in the end it is still ignored;
wenzelm
parents:
65443
diff
changeset
|
254 |
fun theory_qualifier theory = |
65457 | 255 |
(case global_theory theory of |
256 |
SOME qualifier => qualifier |
|
257 |
| NONE => Long_Name.qualifier theory); |
|
65445
e9e7f5f5794c
more qualifier treatment, but in the end it is still ignored;
wenzelm
parents:
65443
diff
changeset
|
258 |
|
76049 | 259 |
fun literal_theory theory = |
260 |
Long_Name.is_qualified theory orelse is_some (global_theory theory); |
|
261 |
||
66712 | 262 |
fun theory_name qualifier theory = |
76049 | 263 |
if literal_theory theory then theory |
67104
a2fa0c6a7aff
clarified theory_name vs. loaded_theory: proper import_name for already loaded theories from other sessions (amending 4c98c929a12a);
wenzelm
parents:
66771
diff
changeset
|
264 |
else Long_Name.qualify qualifier theory; |
65471
05e5bffcf1d8
clarified loaded_theories: map to qualified theory name;
wenzelm
parents:
65457
diff
changeset
|
265 |
|
70683
8c7706b053c7
find theory files via session structure: much faster Prover IDE startup;
wenzelm
parents:
70049
diff
changeset
|
266 |
fun find_theory_file thy_name = |
70712
a3cfe859d915
find theories via session directories only -- ignore known_theories;
wenzelm
parents:
70683
diff
changeset
|
267 |
let |
a3cfe859d915
find theories via session directories only -- ignore known_theories;
wenzelm
parents:
70683
diff
changeset
|
268 |
val thy_file = thy_path (Path.basic (Long_Name.base_name thy_name)); |
a3cfe859d915
find theories via session directories only -- ignore known_theories;
wenzelm
parents:
70683
diff
changeset
|
269 |
val session = theory_qualifier thy_name; |
72758 | 270 |
val dirs = Symtab.lookup_list (get_session_base1 #session_directories) session; |
70712
a3cfe859d915
find theories via session directories only -- ignore known_theories;
wenzelm
parents:
70683
diff
changeset
|
271 |
in |
a3cfe859d915
find theories via session directories only -- ignore known_theories;
wenzelm
parents:
70683
diff
changeset
|
272 |
dirs |> get_first (fn dir => |
72511
460d743010bc
clarified signature: overloaded "+" for Path.append;
wenzelm
parents:
72162
diff
changeset
|
273 |
let val path = dir + thy_file |
70712
a3cfe859d915
find theories via session directories only -- ignore known_theories;
wenzelm
parents:
70683
diff
changeset
|
274 |
in if File.is_file path then SOME path else NONE end) |
a3cfe859d915
find theories via session directories only -- ignore known_theories;
wenzelm
parents:
70683
diff
changeset
|
275 |
end; |
70683
8c7706b053c7
find theory files via session structure: much faster Prover IDE startup;
wenzelm
parents:
70049
diff
changeset
|
276 |
|
70717
cceb10dcc9f9
clarified import_name: observe directory notation more strictly;
wenzelm
parents:
70712
diff
changeset
|
277 |
fun make_theory_node node_name theory = |
cceb10dcc9f9
clarified import_name: observe directory notation more strictly;
wenzelm
parents:
70712
diff
changeset
|
278 |
{node_name = node_name, master_dir = Path.dir node_name, theory_name = theory}; |
cceb10dcc9f9
clarified import_name: observe directory notation more strictly;
wenzelm
parents:
70712
diff
changeset
|
279 |
|
cceb10dcc9f9
clarified import_name: observe directory notation more strictly;
wenzelm
parents:
70712
diff
changeset
|
280 |
fun loaded_theory_node theory = |
cceb10dcc9f9
clarified import_name: observe directory notation more strictly;
wenzelm
parents:
70712
diff
changeset
|
281 |
{node_name = Path.basic theory, master_dir = Path.current, theory_name = theory}; |
cceb10dcc9f9
clarified import_name: observe directory notation more strictly;
wenzelm
parents:
70712
diff
changeset
|
282 |
|
65454 | 283 |
fun import_name qualifier dir s = |
70717
cceb10dcc9f9
clarified import_name: observe directory notation more strictly;
wenzelm
parents:
70712
diff
changeset
|
284 |
let |
cceb10dcc9f9
clarified import_name: observe directory notation more strictly;
wenzelm
parents:
70712
diff
changeset
|
285 |
val theory = theory_name qualifier (Thy_Header.import_name s); |
76046
507c65cc4332
back to more traditional import_name (reverting cceb10dcc9f9), e.g. relevant for "isabelle jedit -l CTT src/CTT/ex/Elimination.thy" to produce proper error "Cannot update finished theory CTT.Elimination";
wenzelm
parents:
76015
diff
changeset
|
286 |
fun theory_node path = make_theory_node path theory; |
76050
f1dc3d9d5164
check imports more strictly, e.g. reject ".../Pure" or ".../HOL-Library.Multiset";
wenzelm
parents:
76049
diff
changeset
|
287 |
val literal_import = literal_theory theory andalso qualifier <> theory_qualifier theory; |
f1dc3d9d5164
check imports more strictly, e.g. reject ".../Pure" or ".../HOL-Library.Multiset";
wenzelm
parents:
76049
diff
changeset
|
288 |
val _ = |
f1dc3d9d5164
check imports more strictly, e.g. reject ".../Pure" or ".../HOL-Library.Multiset";
wenzelm
parents:
76049
diff
changeset
|
289 |
if literal_import andalso not (Thy_Header.is_base_name s) then |
f1dc3d9d5164
check imports more strictly, e.g. reject ".../Pure" or ".../HOL-Library.Multiset";
wenzelm
parents:
76049
diff
changeset
|
290 |
error ("Bad import of theory from other session via file-path: " ^ quote s) |
f1dc3d9d5164
check imports more strictly, e.g. reject ".../Pure" or ".../HOL-Library.Multiset";
wenzelm
parents:
76049
diff
changeset
|
291 |
else (); |
70717
cceb10dcc9f9
clarified import_name: observe directory notation more strictly;
wenzelm
parents:
70712
diff
changeset
|
292 |
in |
76046
507c65cc4332
back to more traditional import_name (reverting cceb10dcc9f9), e.g. relevant for "isabelle jedit -l CTT src/CTT/ex/Elimination.thy" to produce proper error "Cannot update finished theory CTT.Elimination";
wenzelm
parents:
76015
diff
changeset
|
293 |
if loaded_theory theory then loaded_theory_node theory |
67104
a2fa0c6a7aff
clarified theory_name vs. loaded_theory: proper import_name for already loaded theories from other sessions (amending 4c98c929a12a);
wenzelm
parents:
66771
diff
changeset
|
294 |
else |
70717
cceb10dcc9f9
clarified import_name: observe directory notation more strictly;
wenzelm
parents:
70712
diff
changeset
|
295 |
(case find_theory_file theory of |
76046
507c65cc4332
back to more traditional import_name (reverting cceb10dcc9f9), e.g. relevant for "isabelle jedit -l CTT src/CTT/ex/Elimination.thy" to produce proper error "Cannot update finished theory CTT.Elimination";
wenzelm
parents:
76015
diff
changeset
|
296 |
SOME node_name => theory_node node_name |
507c65cc4332
back to more traditional import_name (reverting cceb10dcc9f9), e.g. relevant for "isabelle jedit -l CTT src/CTT/ex/Elimination.thy" to produce proper error "Cannot update finished theory CTT.Elimination";
wenzelm
parents:
76015
diff
changeset
|
297 |
| NONE => |
507c65cc4332
back to more traditional import_name (reverting cceb10dcc9f9), e.g. relevant for "isabelle jedit -l CTT src/CTT/ex/Elimination.thy" to produce proper error "Cannot update finished theory CTT.Elimination";
wenzelm
parents:
76015
diff
changeset
|
298 |
if Thy_Header.is_base_name s andalso Long_Name.is_qualified s |
507c65cc4332
back to more traditional import_name (reverting cceb10dcc9f9), e.g. relevant for "isabelle jedit -l CTT src/CTT/ex/Elimination.thy" to produce proper error "Cannot update finished theory CTT.Elimination";
wenzelm
parents:
76015
diff
changeset
|
299 |
then loaded_theory_node theory |
507c65cc4332
back to more traditional import_name (reverting cceb10dcc9f9), e.g. relevant for "isabelle jedit -l CTT src/CTT/ex/Elimination.thy" to produce proper error "Cannot update finished theory CTT.Elimination";
wenzelm
parents:
76015
diff
changeset
|
300 |
else theory_node (File.full_path dir (thy_path (Path.expand (Path.explode s))))) |
67104
a2fa0c6a7aff
clarified theory_name vs. loaded_theory: proper import_name for already loaded theories from other sessions (amending 4c98c929a12a);
wenzelm
parents:
66771
diff
changeset
|
301 |
end; |
65443
dccbfc715904
provide Resources.import_name in ML, similar to Scala version;
wenzelm
parents:
65442
diff
changeset
|
302 |
|
48878 | 303 |
fun check_file dir file = File.check_file (File.full_path dir file); |
304 |
||
54525 | 305 |
fun check_thy dir thy_name = |
306 |
let |
|
65433 | 307 |
val thy_base_name = Long_Name.base_name thy_name; |
308 |
val master_file = |
|
70683
8c7706b053c7
find theory files via session structure: much faster Prover IDE startup;
wenzelm
parents:
70049
diff
changeset
|
309 |
(case find_theory_file thy_name of |
70712
a3cfe859d915
find theories via session directories only -- ignore known_theories;
wenzelm
parents:
70683
diff
changeset
|
310 |
SOME path => check_file Path.current path |
65442 | 311 |
| NONE => check_file dir (thy_path (Path.basic thy_base_name))); |
54525 | 312 |
val text = File.read master_file; |
313 |
||
314 |
val {name = (name, pos), imports, keywords} = |
|
76884 | 315 |
Thy_Header.read (Position.file (File.symbolic_path master_file)) text; |
65433 | 316 |
val _ = |
317 |
thy_base_name <> name andalso |
|
318 |
error ("Bad theory name " ^ quote name ^ |
|
65442 | 319 |
" for file " ^ Path.print (Path.base master_file) ^ Position.here pos); |
54525 | 320 |
in |
321 |
{master = (master_file, SHA1.digest text), text = text, theory_pos = pos, |
|
322 |
imports = imports, keywords = keywords} |
|
323 |
end; |
|
324 |
||
325 |
||
76818 | 326 |
(* read_file *) |
76803 | 327 |
|
76806 | 328 |
fun read_file_node file_node master_dir (src_path, pos) = |
76803 | 329 |
let |
330 |
fun read_local () = |
|
331 |
let |
|
332 |
val path = File.check_file (File.full_path master_dir src_path); |
|
333 |
val text = File.read path; |
|
76884 | 334 |
val file_pos = Position.file (File.symbolic_path path); |
76803 | 335 |
in (text, file_pos) end; |
336 |
||
337 |
fun read_remote () = |
|
338 |
let |
|
339 |
val text = Bytes.content (Isabelle_System.download file_node); |
|
340 |
val file_pos = Position.file file_node; |
|
341 |
in (text, file_pos) end; |
|
342 |
||
343 |
val (text, file_pos) = |
|
344 |
(case try Url.explode file_node of |
|
345 |
NONE => read_local () |
|
346 |
| SOME (Url.File _) => read_local () |
|
347 |
| _ => read_remote ()); |
|
348 |
||
349 |
val lines = split_lines text; |
|
350 |
val digest = SHA1.digest text; |
|
351 |
in {src_path = src_path, lines = lines, digest = digest, pos = Position.copy_id pos file_pos} end |
|
352 |
handle ERROR msg => error (msg ^ Position.here pos); |
|
353 |
||
76818 | 354 |
val read_file = read_file_node ""; |
355 |
||
76803 | 356 |
|
54526 | 357 |
(* load files *) |
48898
9fc880720663
simplified Thy_Load.check_thy (again) -- no need to pass keywords nor find files in body text;
wenzelm
parents:
48897
diff
changeset
|
358 |
|
76807 | 359 |
fun parsed_files make_paths (files, source) thy = |
360 |
if null files then |
|
361 |
let |
|
362 |
val master_dir = master_directory thy; |
|
363 |
val name = Input.string_of source; |
|
364 |
val pos = Input.pos_of source; |
|
365 |
val delimited = Input.is_delimited source; |
|
366 |
val src_paths = make_paths (Path.explode name); |
|
367 |
val reports = |
|
368 |
src_paths |> maps (fn src_path => |
|
76884 | 369 |
[(pos, Markup.path (File.symbolic_path (master_dir + src_path))), |
76807 | 370 |
(pos, Markup.language_path delimited)]); |
371 |
val _ = Position.reports reports; |
|
76818 | 372 |
in map (read_file master_dir o rpair pos) src_paths end |
76807 | 373 |
else map Exn.release files; |
374 |
||
72747
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
wenzelm
parents:
72669
diff
changeset
|
375 |
fun parse_files make_paths = |
76807 | 376 |
(Scan.ahead Parse.not_eof >> Token.get_files) -- Parse.path_input >> parsed_files make_paths; |
48869 | 377 |
|
72747
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
wenzelm
parents:
72669
diff
changeset
|
378 |
val parse_file = parse_files single >> (fn f => f #> the_single); |
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
wenzelm
parents:
72669
diff
changeset
|
379 |
|
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
wenzelm
parents:
72669
diff
changeset
|
380 |
|
48906 | 381 |
fun provide (src_path, id) = |
382 |
map_files (fn (master_dir, imports, provided) => |
|
383 |
if AList.defined (op =) provided src_path then |
|
384 |
error ("Duplicate use of source file: " ^ Path.print src_path) |
|
385 |
else (master_dir, imports, (src_path, id) :: provided)); |
|
386 |
||
69851 | 387 |
fun provide_file (file: Token.file) = provide (#src_path file, #digest file); |
76801 | 388 |
fun provide_file' file thy = (file, provide_file file thy); |
69851 | 389 |
|
72747
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
wenzelm
parents:
72669
diff
changeset
|
390 |
fun provide_parse_files make_paths = |
76801 | 391 |
parse_files make_paths >> (fn files => fn thy => fold_map provide_file' (files thy) thy); |
48906 | 392 |
|
72747
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
wenzelm
parents:
72669
diff
changeset
|
393 |
val provide_parse_file = provide_parse_files single >> (fn f => f #>> the_single); |
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
wenzelm
parents:
72669
diff
changeset
|
394 |
|
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
wenzelm
parents:
72669
diff
changeset
|
395 |
|
43702
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
wenzelm
parents:
43700
diff
changeset
|
396 |
fun load_file thy src_path = |
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
wenzelm
parents:
43700
diff
changeset
|
397 |
let |
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
wenzelm
parents:
43700
diff
changeset
|
398 |
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:
43700
diff
changeset
|
399 |
val text = File.read full_path; |
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
wenzelm
parents:
43700
diff
changeset
|
400 |
val id = SHA1.digest text; |
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
wenzelm
parents:
43700
diff
changeset
|
401 |
in ((full_path, id), text) end; |
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
wenzelm
parents:
43700
diff
changeset
|
402 |
|
54446 | 403 |
fun loaded_files_current thy = |
48886
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
wenzelm
parents:
48881
diff
changeset
|
404 |
#provided (Files.get thy) |> |
48896
bb1f461a7815
simplified Thy_Load.provide: do not store full path;
wenzelm
parents:
48888
diff
changeset
|
405 |
forall (fn (src_path, id) => |
43702
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
wenzelm
parents:
43700
diff
changeset
|
406 |
(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:
37943
diff
changeset
|
407 |
NONE => false |
48886
9604c6563226
discontinued separate list of required files -- maintain only provided files as they occur at runtime;
wenzelm
parents:
48881
diff
changeset
|
408 |
| SOME ((_, id'), _) => id = id')); |
37952 | 409 |
|
37949
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents:
37943
diff
changeset
|
410 |
|
67217 | 411 |
(* formal check *) |
49244
fb669aff821e
formal markup for @{file} (for hyperlinks etc.) -- interpret path wrt. master directory as usual;
wenzelm
parents:
48992
diff
changeset
|
412 |
|
76800 | 413 |
fun formal_check (check: Path.T -> Path.T) ctxt opt_dir source = |
54705 | 414 |
let |
72841 | 415 |
val name = Input.string_of source; |
416 |
val pos = Input.pos_of source; |
|
417 |
val delimited = Input.is_delimited source; |
|
418 |
||
419 |
val _ = Context_Position.report ctxt pos (Markup.language_path delimited); |
|
420 |
||
67214 | 421 |
fun err msg = error (msg ^ Position.here pos); |
70049 | 422 |
val dir = |
423 |
(case opt_dir of |
|
424 |
SOME dir => dir |
|
425 |
| NONE => master_directory (Proof_Context.theory_of ctxt)); |
|
72511
460d743010bc
clarified signature: overloaded "+" for Path.append;
wenzelm
parents:
72162
diff
changeset
|
426 |
val path = dir + Path.explode name handle ERROR msg => err msg; |
67214 | 427 |
val _ = Path.expand path handle ERROR msg => err msg; |
76884 | 428 |
val _ = Context_Position.report ctxt pos (Markup.path (File.symbolic_path path)); |
76800 | 429 |
val _ = check path handle ERROR msg => err msg; |
63675 | 430 |
in path end; |
431 |
||
67217 | 432 |
val check_path = formal_check I; |
433 |
val check_file = formal_check File.check_file; |
|
434 |
val check_dir = formal_check File.check_dir; |
|
435 |
||
73312 | 436 |
fun check_session_dir ctxt opt_dir s = |
437 |
let |
|
438 |
val dir = Path.expand (check_dir ctxt opt_dir s); |
|
439 |
val ok = |
|
440 |
File.is_file (dir + Path.explode("ROOT")) orelse |
|
441 |
File.is_file (dir + Path.explode("ROOTS")); |
|
442 |
in |
|
443 |
if ok then dir |
|
444 |
else |
|
445 |
error ("Bad session root directory (missing ROOT or ROOTS): " ^ |
|
446 |
Path.print dir ^ Position.here (Input.pos_of s)) |
|
447 |
end; |
|
448 |
||
67217 | 449 |
|
450 |
(* antiquotations *) |
|
451 |
||
452 |
local |
|
453 |
||
72841 | 454 |
fun document_antiq (check: Proof.context -> Path.T option -> Input.source -> Path.T) = |
455 |
Args.context -- Scan.lift Parse.path_input >> (fn (ctxt, source) => |
|
73780
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
wenzelm
parents:
73761
diff
changeset
|
456 |
(check ctxt NONE source; |
466fae6bf22e
compose Latex text as XML, output exported YXML in Isabelle/Scala;
wenzelm
parents:
73761
diff
changeset
|
457 |
Latex.string (Latex.output_ascii_breakable "/" (Input.string_of source)) |
74790 | 458 |
|> Latex.macro "isatt")); |
54705 | 459 |
|
69281 | 460 |
fun ML_antiq check = |
72841 | 461 |
Args.context -- Scan.lift Parse.path_input >> (fn (ctxt, source) => |
462 |
check ctxt (SOME Path.current) source |> ML_Syntax.print_path); |
|
63675 | 463 |
|
56135 | 464 |
in |
465 |
||
53171 | 466 |
val _ = Theory.setup |
73761 | 467 |
(Document_Output.antiquotation_verbatim_embedded \<^binding>\<open>session\<close> |
69349
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
wenzelm
parents:
69289
diff
changeset
|
468 |
(Scan.lift Parse.embedded_position) check_session #> |
73761 | 469 |
Document_Output.antiquotation_raw_embedded \<^binding>\<open>path\<close> (document_antiq check_path) (K I) #> |
470 |
Document_Output.antiquotation_raw_embedded \<^binding>\<open>file\<close> (document_antiq check_file) (K I) #> |
|
471 |
Document_Output.antiquotation_raw_embedded \<^binding>\<open>dir\<close> (document_antiq check_dir) (K I) #> |
|
69592
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
472 |
ML_Antiquotation.value_embedded \<^binding>\<open>path\<close> (ML_antiq check_path) #> |
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
473 |
ML_Antiquotation.value_embedded \<^binding>\<open>file\<close> (ML_antiq check_file) #> |
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
474 |
ML_Antiquotation.value_embedded \<^binding>\<open>dir\<close> (ML_antiq check_dir) #> |
70015
c8e08d8ffb93
clarified signature: more explicit type Path.binding;
wenzelm
parents:
69851
diff
changeset
|
475 |
ML_Antiquotation.value_embedded \<^binding>\<open>path_binding\<close> |
c8e08d8ffb93
clarified signature: more explicit type Path.binding;
wenzelm
parents:
69851
diff
changeset
|
476 |
(Scan.lift (Parse.position Parse.path) >> |
c8e08d8ffb93
clarified signature: more explicit type Path.binding;
wenzelm
parents:
69851
diff
changeset
|
477 |
(ML_Syntax.print_path_binding o Path.explode_binding)) #> |
69282 | 478 |
ML_Antiquotation.value \<^binding>\<open>master_dir\<close> |
479 |
(Args.theory >> (ML_Syntax.print_path o master_directory))); |
|
49244
fb669aff821e
formal markup for @{file} (for hyperlinks etc.) -- interpret path wrt. master directory as usual;
wenzelm
parents:
48992
diff
changeset
|
480 |
|
6168 | 481 |
end; |
56135 | 482 |
|
483 |
end; |