7 |
7 |
8 section {* Isar toplevel \label{sec:isar-toplevel} *} |
8 section {* Isar toplevel \label{sec:isar-toplevel} *} |
9 |
9 |
10 text {* The Isar toplevel may be considered the centeral hub of the |
10 text {* The Isar toplevel may be considered the centeral hub of the |
11 Isabelle/Isar system, where all key components and sub-systems are |
11 Isabelle/Isar system, where all key components and sub-systems are |
12 integrated into a single read-eval-print loop of Isar commands. |
12 integrated into a single read-eval-print loop of Isar commands. We |
13 Here we even incorporate the existing {\ML} toplevel of the compiler |
13 shall even incorporate the existing {\ML} toplevel of the compiler |
14 and run-time system (cf.\ \secref{sec:ML-toplevel}). |
14 and run-time system (cf.\ \secref{sec:ML-toplevel}). |
15 |
15 |
16 Isabelle/Isar departs from original ``LCF system architecture'' |
16 Isabelle/Isar departs from the original ``LCF system architecture'' |
17 where {\ML} was really The Meta Language for defining theories and |
17 where {\ML} was really The Meta Language for defining theories and |
18 conducting proofs. Instead, {\ML} merely serves as the |
18 conducting proofs. Instead, {\ML} now only serves as the |
19 implementation language for the system (and user extensions), while |
19 implementation language for the system (and user extensions), while |
20 our specific Isar toplevel supports particular notions of |
20 the specific Isar toplevel supports the concepts of theory and proof |
21 incremental theory and proof development more directly. This |
21 development natively. This includes the graph structure of theories |
22 includes the graph structure of theories and the block structure of |
22 and the block structure of proofs, support for unlimited undo, |
23 proofs, support for unlimited undo, facilities for tracing, |
23 facilities for tracing, debugging, timing, profiling etc. |
24 debugging, timing, profiling. |
|
25 |
24 |
26 \medskip The toplevel maintains an implicit state, which is |
25 \medskip The toplevel maintains an implicit state, which is |
27 transformed by a sequence of transitions -- either interactively or |
26 transformed by a sequence of transitions -- either interactively or |
28 in batch-mode. In interactive mode, Isar state transitions are |
27 in batch-mode. In interactive mode, Isar state transitions are |
29 encapsulated as safe transactions, such that both failure and undo |
28 encapsulated as safe transactions, such that both failure and undo |
30 are handled conveniently without destroying the underlying draft |
29 are handled conveniently without destroying the underlying draft |
31 theory (cf.~\secref{sec:context-theory}). In batch mode, |
30 theory (cf.~\secref{sec:context-theory}). In batch mode, |
32 transitions operate in a strictly linear (destructive) fashion, such |
31 transitions operate in a linear (destructive) fashion, such that |
33 that error conditions abort the present attempt to construct a |
32 error conditions abort the present attempt to construct a theory or |
34 theory altogether. |
33 proof altogether. |
35 |
34 |
36 The toplevel state is a disjoint sum of empty @{text toplevel}, or |
35 The toplevel state is a disjoint sum of empty @{text toplevel}, or |
37 @{text theory}, or @{text proof}. On entering the main Isar loop we |
36 @{text theory}, or @{text proof}. On entering the main Isar loop we |
38 start with an empty toplevel. A theory is commenced by giving a |
37 start with an empty toplevel. A theory is commenced by giving a |
39 @{text \<THEORY>} header; within a theory we may issue theory |
38 @{text \<THEORY>} header; within a theory we may issue theory |
67 \end{mldecls} |
66 \end{mldecls} |
68 |
67 |
69 \begin{description} |
68 \begin{description} |
70 |
69 |
71 \item @{ML_type Toplevel.state} represents Isar toplevel states, |
70 \item @{ML_type Toplevel.state} represents Isar toplevel states, |
72 which are normally only manipulated through the toplevel transition |
71 which are normally manipulated through the concept of toplevel |
73 concept (\secref{sec:toplevel-transition}). Also note that a |
72 transitions only (\secref{sec:toplevel-transition}). Also note that |
74 toplevel state is subject to the same linerarity restrictions as a |
73 a raw toplevel state is subject to the same linearity restrictions |
75 theory context (cf.~\secref{sec:context-theory}). |
74 as a theory context (cf.~\secref{sec:context-theory}). |
76 |
75 |
77 \item @{ML Toplevel.UNDEF} is raised for undefined toplevel |
76 \item @{ML Toplevel.UNDEF} is raised for undefined toplevel |
78 operations: @{ML_type Toplevel.state} is a sum type, many operations |
77 operations. Many operations work only partially for certain cases, |
79 work only partially for certain cases. |
78 since @{ML_type Toplevel.state} is a sum type. |
80 |
79 |
81 \item @{ML Toplevel.is_toplevel} checks for an empty toplevel state. |
80 \item @{ML Toplevel.is_toplevel}~@{text "state"} checks for an empty |
82 |
81 toplevel state. |
83 \item @{ML Toplevel.theory_of} gets the theory of a theory or proof |
82 |
84 (!), otherwise raises @{ML Toplevel.UNDEF}. |
83 \item @{ML Toplevel.theory_of}~@{text "state"} selects the theory of |
85 |
84 a theory or proof (!), otherwise raises @{ML Toplevel.UNDEF}. |
86 \item @{ML Toplevel.proof_of} gets the Isar proof state if |
85 |
87 available, otherwise raises @{ML Toplevel.UNDEF}. |
86 \item @{ML Toplevel.proof_of}~@{text "state"} selects the Isar proof |
|
87 state if available, otherwise raises @{ML Toplevel.UNDEF}. |
88 |
88 |
89 \item @{ML "set Toplevel.debug"} makes the toplevel print further |
89 \item @{ML "set Toplevel.debug"} makes the toplevel print further |
90 details about internal error conditions, exceptions being raised |
90 details about internal error conditions, exceptions being raised |
91 etc. |
91 etc. |
92 |
92 |
93 \item @{ML "set Toplevel.timing"} makes the toplevel print timing |
93 \item @{ML "set Toplevel.timing"} makes the toplevel print timing |
94 information for each Isar command being executed. |
94 information for each Isar command being executed. |
95 |
95 |
96 \item @{ML Toplevel.profiling} controls low-level profiling of the |
96 \item @{ML Toplevel.profiling}~@{verbatim ":="}~@{text "n"} controls |
97 underlying {\ML} runtime system.\footnote{For Poly/ML, 1 means time |
97 low-level profiling of the underlying {\ML} runtime system. For |
98 and 2 space profiling.} |
98 Poly/ML, @{text "n = 1"} means time and @{text "n = 2"} space |
99 |
99 profiling. |
100 \end{description} |
100 |
101 *} |
101 \end{description} |
102 |
102 *} |
103 |
103 |
104 subsection {* Toplevel transitions *} |
104 |
105 |
105 subsection {* Toplevel transitions \label{sec:toplevel-transition} *} |
106 text {* An Isar toplevel transition consists of a partial |
106 |
107 function on the toplevel state, with additional information for |
107 text {* |
108 diagnostics and error reporting: there are fields for command name, |
108 An Isar toplevel transition consists of a partial function on the |
109 source position, optional source text, as well as flags for |
109 toplevel state, with additional information for diagnostics and |
110 interactive-only commands (which issue a warning in batch-mode), |
110 error reporting: there are fields for command name, source position, |
111 printing of result state, etc. |
111 optional source text, as well as flags for interactive-only commands |
112 |
112 (which issue a warning in batch-mode), printing of result state, |
113 The operational part is represented as a sequential union of a list |
113 etc. |
114 of partial functions, which are tried in turn until the first one |
114 |
115 succeeds (i.e.\ does not raise @{ML Toplevel.UNDEF}). For example, |
115 The operational part is represented as the sequential union of a |
116 a single Isar command like \isacommand{qed} consists of the union of |
116 list of partial functions, which are tried in turn until the first |
117 some function @{ML_type "Proof.state -> Proof.state"} for proofs |
117 one succeeds (i.e.\ does not raise @{ML Toplevel.UNDEF}). This acts |
118 within proofs, plus @{ML_type "Proof.state -> theory"} for proofs at |
118 like an outer case-expression for various alternative state |
119 the outer theory level. |
119 transitions. For example, \isakeyword{qed} acts differently for a |
|
120 local proofs vs.\ the global ending of the main proof. |
120 |
121 |
121 Toplevel transitions are composed via transition transformers. |
122 Toplevel transitions are composed via transition transformers. |
122 Internally, Isar commands are put together from an empty transition |
123 Internally, Isar commands are put together from an empty transition |
123 extended by name and source position (and optional source text). It |
124 extended by name and source position (and optional source text). It |
124 is then left to the individual command parser to turn the given |
125 is then left to the individual command parser to turn the given |
125 syntax body into a suitable transition transformer that adjoin |
126 concrete syntax into a suitable transition transformer that adjoin |
126 actual operations on a theory or proof state etc. |
127 actual operations on a theory or proof state etc. |
127 *} |
128 *} |
128 |
129 |
129 text %mlref {* |
130 text %mlref {* |
130 \begin{mldecls} |
131 \begin{mldecls} |
144 Toplevel.transition -> Toplevel.transition"} \\ |
145 Toplevel.transition -> Toplevel.transition"} \\ |
145 \end{mldecls} |
146 \end{mldecls} |
146 |
147 |
147 \begin{description} |
148 \begin{description} |
148 |
149 |
149 \item @{ML Toplevel.print} sets the print flag, which causes the |
150 \item @{ML Toplevel.print}~@{text "tr"} sets the print flag, which |
150 resulting state of the transition to be echoed in interactive mode. |
151 causes the toplevel loop to echo the result state (in interactive |
151 |
152 mode). |
152 \item @{ML Toplevel.no_timing} indicates that the transition should |
153 |
153 never show timing information, e.g.\ because it is merely a |
154 \item @{ML Toplevel.no_timing}~@{text "tr"} indicates that the |
154 diagnostic command. |
155 transition should never show timing information, e.g.\ because it is |
155 |
156 a diagnostic command. |
156 \item @{ML Toplevel.keep} adjoins a diagnostic function. |
157 |
157 |
158 \item @{ML Toplevel.keep}~@{text "tr"} adjoins a diagnostic |
158 \item @{ML Toplevel.theory} adjoins a theory transformer. |
159 function. |
159 |
160 |
160 \item @{ML Toplevel.theory_to_proof} adjoins a global goal function, |
161 \item @{ML Toplevel.theory}~@{text "tr"} adjoins a theory |
161 which turns a theory into a proof state. The theory may be changed |
162 transformer. |
162 before entering the proof; the generic Isar goal setup includes an |
163 |
163 argument that specifies how to apply the proven result to the |
164 \item @{ML Toplevel.theory_to_proof}~@{text "tr"} adjoins a global |
164 theory, when the proof is finished. |
165 goal function, which turns a theory into a proof state. The theory |
165 |
166 may be changed before entering the proof; the generic Isar goal |
166 \item @{ML Toplevel.proof} adjoins a deterministic proof command, |
167 setup includes an argument that specifies how to apply the proven |
167 with a singleton result state. |
168 result to the theory, when the proof is finished. |
168 |
169 |
169 \item @{ML Toplevel.proofs} adjoins a general proof command, with |
170 \item @{ML Toplevel.proof}~@{text "tr"} adjoins a deterministic |
170 zero or more result states (represented as a lazy list). |
171 proof command, with a singleton result. |
171 |
172 |
172 \item @{ML Toplevel.proof_to_theory} adjoins a concluding proof |
173 \item @{ML Toplevel.proofs}~@{text "tr"} adjoins a general proof |
173 command, that returns the resulting theory, after storing the |
174 command, with zero or more result states (represented as a lazy |
174 resulting facts etc. |
175 list). |
|
176 |
|
177 \item @{ML Toplevel.proof_to_theory}~@{text "tr"} adjoins a |
|
178 concluding proof command, that returns the resulting theory, after |
|
179 storing the resulting facts etc. |
175 |
180 |
176 \end{description} |
181 \end{description} |
177 *} |
182 *} |
178 |
183 |
179 |
184 |
180 subsection {* Toplevel control *} |
185 subsection {* Toplevel control *} |
181 |
186 |
182 text {* Apart from regular toplevel transactions there are a few |
187 text {* |
183 special control commands that modify the behavior the toplevel |
188 There are a few special control commands that modify the behavior |
184 itself, and only make sense in interactive mode. Under normal |
189 the toplevel itself, and only make sense in interactive mode. Under |
185 circumstances, the user encounters these only implicitly as part of |
190 normal circumstances, the user encounters these only implicitly as |
186 the protocol between the Isabelle/Isar system and a user-interface |
191 part of the protocol between the Isabelle/Isar system and a |
187 such as ProofGeneral. |
192 user-interface such as ProofGeneral. |
188 |
193 |
189 \begin{description} |
194 \begin{description} |
190 |
195 |
191 \item \isacommand{undo} follows the three-level hierarchy of empty |
196 \item \isacommand{undo} follows the three-level hierarchy of empty |
192 toplevel vs.\ theory vs.\ proof: undo within a proof reverts to the |
197 toplevel vs.\ theory vs.\ proof: undo within a proof reverts to the |
219 operate on the global system state, which consists of the compiler |
224 operate on the global system state, which consists of the compiler |
220 environment plus the values of {\ML} reference variables. There is |
225 environment plus the values of {\ML} reference variables. There is |
221 no clean way to undo {\ML} declarations, except for reverting to a |
226 no clean way to undo {\ML} declarations, except for reverting to a |
222 previously saved state of the whole Isabelle process. {\ML} input |
227 previously saved state of the whole Isabelle process. {\ML} input |
223 is either read interactively from a TTY, or from a string (usually |
228 is either read interactively from a TTY, or from a string (usually |
224 within a theory text), or from a source file (usually associated |
229 within a theory text), or from a source file (usually loaded from a |
225 with a theory). |
230 theory). |
226 |
231 |
227 Whenever the {\ML} toplevel is active, the current Isabelle theory |
232 Whenever the {\ML} toplevel is active, the current Isabelle theory |
228 context is passed as an internal reference variable. Thus {\ML} |
233 context is passed as an internal reference variable. Thus {\ML} |
229 code may access the theory context during compilation, it may even |
234 code may access the theory context during compilation, it may even |
230 change the value of a theory being under construction --- following |
235 change the value of a theory being under construction --- while |
231 the usual linearity restrictions (cf.~\secref{sec:context-theory}). |
236 observing the usual linearity restrictions |
|
237 (cf.~\secref{sec:context-theory}). |
232 *} |
238 *} |
233 |
239 |
234 text %mlref {* |
240 text %mlref {* |
235 \begin{mldecls} |
241 \begin{mldecls} |
236 @{index_ML context: "theory -> unit"} \\ |
242 @{index_ML context: "theory -> unit"} \\ |
245 when dropping out of the interactive Isar toplevel into {\ML}, or |
251 when dropping out of the interactive Isar toplevel into {\ML}, or |
246 when Isar invokes {\ML} to process code from a string or a file. |
252 when Isar invokes {\ML} to process code from a string or a file. |
247 |
253 |
248 \item @{ML "the_context ()"} refers to the theory context of the |
254 \item @{ML "the_context ()"} refers to the theory context of the |
249 {\ML} toplevel --- at compile time! {\ML} code needs to take care |
255 {\ML} toplevel --- at compile time! {\ML} code needs to take care |
250 to refer to @{ML "the_context ()"} correctly, recall that evaluation |
256 to refer to @{ML "the_context ()"} correctly. Recall that |
251 of a function body is delayed until actual runtime. Moreover, |
257 evaluation of a function body is delayed until actual runtime. |
252 persistent {\ML} toplevel bindings to an unfinished theory should be |
258 Moreover, persistent {\ML} toplevel bindings to an unfinished theory |
253 avoided: code should either project out the desired information |
259 should be avoided: code should either project out the desired |
254 immediately, or produce an explicit @{ML_type theory_ref} (cf.\ |
260 information immediately, or produce an explicit @{ML_type |
255 \secref{sec:context-theory}). |
261 theory_ref} (cf.\ \secref{sec:context-theory}). |
256 |
262 |
257 \item @{ML "Context.>>"}~@{text f} applies theory transformation |
263 \item @{ML "Context.>>"}~@{text f} applies theory transformation |
258 @{text f} to the current theory of the {\ML} toplevel. In order to |
264 @{text f} to the current theory of the {\ML} toplevel. In order to |
259 work as expected, the theory should be still under construction, and |
265 work as expected, the theory should be still under construction, and |
260 the Isar language element that invoked the {\ML} compiler in the |
266 the Isar language element that invoked the {\ML} compiler in the |
261 first place shoule be ready to accept the changed theory value |
267 first place should be ready to accept the changed theory value |
262 (e.g.\ \isakeyword{ML-setup}, but not plain \isakeyword{ML}). |
268 (e.g.\ \isakeyword{ML-setup}, but not plain \isakeyword{ML}). |
263 Otherwise the theory may get destroyed! |
269 Otherwise the theory becomes stale! |
264 |
270 |
265 \end{description} |
271 \end{description} |
266 |
272 |
267 It is very important to note that the above functions are really |
273 It is very important to note that the above functions are really |
268 restricted to the compile time, even though the {\ML} compiler is |
274 restricted to the compile time, even though the {\ML} compiler is |
269 invoked at runtime! The majority of {\ML} code uses explicit |
275 invoked at runtime! The majority of {\ML} code uses explicit |
270 functional arguments of a theory or proof context, as required. |
276 functional arguments of a theory or proof context instead. Thus it |
271 Thus it may get run in an arbitrary context later on. |
277 may be invoked for an arbitrary context later on, without having to |
|
278 worry about any operational details. |
272 |
279 |
273 \bigskip |
280 \bigskip |
274 |
281 |
275 \begin{mldecls} |
282 \begin{mldecls} |
276 @{index_ML Isar.main: "unit -> unit"} \\ |
283 @{index_ML Isar.main: "unit -> unit"} \\ |
281 \end{mldecls} |
288 \end{mldecls} |
282 |
289 |
283 \begin{description} |
290 \begin{description} |
284 |
291 |
285 \item @{ML "Isar.main ()"} invokes the Isar toplevel from {\ML}, |
292 \item @{ML "Isar.main ()"} invokes the Isar toplevel from {\ML}, |
286 initializing the state to empty toplevel state. |
293 initializing an empty toplevel state. |
287 |
294 |
288 \item @{ML "Isar.loop ()"} continues the Isar toplevel with the |
295 \item @{ML "Isar.loop ()"} continues the Isar toplevel with the |
289 current state, after dropping out of the Isar toplevel loop. |
296 current state, after having dropped out of the Isar toplevel loop. |
290 |
297 |
291 \item @{ML "Isar.state ()"} and @{ML "Isar.exn ()"} get current |
298 \item @{ML "Isar.state ()"} and @{ML "Isar.exn ()"} get current |
292 toplevel state and optional error condition, respectively. This |
299 toplevel state and error condition, respectively. This only works |
293 only works after dropping out of the Isar toplevel loop. |
300 after having dropped out of the Isar toplevel loop. |
294 |
301 |
295 \item @{ML "Isar.context ()"} produces the proof context from @{ML |
302 \item @{ML "Isar.context ()"} produces the proof context from @{ML |
296 "Isar.state ()"}. |
303 "Isar.state ()"}, analogous to @{ML Context.proof_of} |
297 |
304 (\secref{sec:generic-context}). |
298 \end{description} |
305 |
299 *} |
306 \end{description} |
300 |
307 *} |
301 |
308 |
302 section {* Theory database *} |
309 |
303 |
310 section {* Theory database \label{sec:theory-database} *} |
304 text {* The theory database maintains a collection of theories, |
311 |
305 together with some administrative information about the original |
312 text {* |
306 sources, which are held in an external store (i.e.\ a collection of |
313 The theory database maintains a collection of theories, together |
307 directories within the regular file system of the underlying |
314 with some administrative information about their original sources, |
308 platform). |
315 which are held in an external store (i.e.\ some directory within the |
309 |
316 regular file system). |
310 The theory database is organized as a directed acyclic graph, with |
317 |
311 entries referenced by theory name. Although some external |
318 The theory database is organized as a directed acyclic graph; |
312 interfaces allow to include a directory specification, this is only |
319 entries are referenced by theory name. Although some additional |
313 a hint to the underlying theory loader mechanism: the internal |
320 interfaces allow to include a directory specification as well, this |
314 theory name space is flat. |
321 is only a hint to the underlying theory loader. The internal theory |
|
322 name space is flat! |
315 |
323 |
316 Theory @{text A} is associated with the main theory file @{text |
324 Theory @{text A} is associated with the main theory file @{text |
317 A}\verb,.thy,, which needs to be accessible through the theory |
325 A}\verb,.thy,, which needs to be accessible through the theory |
318 loader path. A number of optional {\ML} source files may be |
326 loader path. Any number of additional {\ML} source files may be |
319 associated with each theory, by declaring these dependencies in the |
327 associated with each theory, by declaring these dependencies in the |
320 theory header as @{text \<USES>}, and loading them consecutively |
328 theory header as @{text \<USES>}, and loading them consecutively |
321 within the theory context. The system keeps track of incoming {\ML} |
329 within the theory context. The system keeps track of incoming {\ML} |
322 sources and associates them with the current theory. The special |
330 sources and associates them with the current theory. The file |
323 theory {\ML} file @{text A}\verb,.ML, is loaded after a theory has |
331 @{text A}\verb,.ML, is loaded after a theory has been concluded, in |
324 been concluded, in order to support legacy proof {\ML} proof |
332 order to support legacy proof {\ML} proof scripts. |
325 scripts. |
|
326 |
333 |
327 The basic internal actions of the theory database are @{text |
334 The basic internal actions of the theory database are @{text |
328 "update"}, @{text "outdate"}, and @{text "remove"}: |
335 "update"}, @{text "outdate"}, and @{text "remove"}: |
329 |
336 |
330 \begin{itemize} |
337 \begin{itemize} |
331 |
338 |
332 \item @{text "update A"} introduces a link of @{text "A"} with a |
339 \item @{text "update A"} introduces a link of @{text "A"} with a |
333 @{text "theory"} value of the same name; it asserts that the theory |
340 @{text "theory"} value of the same name; it asserts that the theory |
334 sources are consistent with that value. |
341 sources are now consistent with that value; |
335 |
342 |
336 \item @{text "outdate A"} invalidates the link of a theory database |
343 \item @{text "outdate A"} invalidates the link of a theory database |
337 entry to its sources, but retains the present theory value. |
344 entry to its sources, but retains the present theory value; |
338 |
345 |
339 \item @{text "remove A"} removes entry @{text "A"} from the theory |
346 \item @{text "remove A"} deletes entry @{text "A"} from the theory |
340 database. |
347 database. |
341 |
348 |
342 \end{itemize} |
349 \end{itemize} |
343 |
350 |
344 These actions are propagated to sub- or super-graphs of a theory |
351 These actions are propagated to sub- or super-graphs of a theory |
345 entry in the usual way, in order to preserve global consistency of |
352 entry as expected, in order to preserve global consistency of the |
346 the state of all loaded theories with the sources of the external |
353 state of all loaded theories with the sources of the external store. |
347 store. This implies causal dependencies of certain actions: @{text |
354 This implies certain causalities between actions: @{text "update"} |
348 "update"} or @{text "outdate"} of an entry will @{text "outdate"} |
355 or @{text "outdate"} of an entry will @{text "outdate"} all |
349 all descendants; @{text "remove"} will @{text "remove"} all |
356 descendants; @{text "remove"} will @{text "remove"} all descendants. |
350 descendants. |
|
351 |
357 |
352 \medskip There are separate user-level interfaces to operate on the |
358 \medskip There are separate user-level interfaces to operate on the |
353 theory database directly or indirectly. The primitive actions then |
359 theory database directly or indirectly. The primitive actions then |
354 just happen automatically while working with the system. In |
360 just happen automatically while working with the system. In |
355 particular, processing a theory header @{text "\<THEORY> A |
361 particular, processing a theory header @{text "\<THEORY> A |
356 \<IMPORTS> B\<^sub>1 \<dots> B\<^sub>n \<BEGIN>"} ensure that the |
362 \<IMPORTS> B\<^sub>1 \<dots> B\<^sub>n \<BEGIN>"} ensures that the |
357 sub-graph of the collective imports @{text "B\<^sub>1 \<dots> B\<^sub>n"} |
363 sub-graph of the collective imports @{text "B\<^sub>1 \<dots> B\<^sub>n"} |
358 is up-to-date. Earlier theories are reloaded as required, with |
364 is up-to-date, too. Earlier theories are reloaded as required, with |
359 @{text update} actions proceeding in topological order according to |
365 @{text update} actions proceeding in topological order according to |
360 theory dependencies. There may be also a wave of implied @{text |
366 theory dependencies. There may be also a wave of implied @{text |
361 outdate} actions for derived theory nodes until a stable situation |
367 outdate} actions for derived theory nodes until a stable situation |
362 is achieved eventually. |
368 is achieved eventually. |
363 *} |
369 *} |
379 \end{mldecls} |
383 \end{mldecls} |
380 |
384 |
381 \begin{description} |
385 \begin{description} |
382 |
386 |
383 \item @{ML theory}~@{text A} retrieves the theory value presently |
387 \item @{ML theory}~@{text A} retrieves the theory value presently |
384 associated with @{text A}. The result is not necessarily |
388 associated with name @{text A}. Note that the result might be |
385 up-to-date! |
389 outdated. |
386 |
390 |
387 \item @{ML use_thy}~@{text A} loads theory @{text A} if it is absent |
391 \item @{ML use_thy}~@{text A} loads theory @{text A} if it is absent |
388 or out-of-date. It ensures that all parent theories are available |
392 or out-of-date. It ensures that all parent theories are available |
389 as well, but does not reload them if older versions are already |
393 as well, but does not reload them if older versions are already |
390 present. |
394 present. |
391 |
395 |
392 \item @{ML update_thy} is similar to @{ML use_thy}, but ensures that |
396 \item @{ML update_thy} is similar to @{ML use_thy}, but ensures that |
393 the @{text A} and all of its ancestors are fully up-to-date. |
397 theory @{text A} and all ancestors are fully up-to-date. |
394 |
398 |
395 \item @{ML use_thy_only}~@{text A} is like @{ML use_thy}~@{text A}, |
399 \item @{ML touch_thy}~@{text A} performs and @{text outdate} action |
396 but refrains from loading the attached @{text A}\verb,.ML, file. |
400 on theory @{text A} and all descendants. |
397 This is occasionally useful in replaying legacy {\ML} proof scripts |
401 |
398 by hand. |
402 \item @{ML remove_thy}~@{text A} deletes theory @{text A} and all |
399 |
|
400 \item @{ML update_thy_only} is analogous to @{ML use_thy_only}, but |
|
401 proceeds like @{ML update_thy} for ancestors. |
|
402 |
|
403 \item @{ML touch_thy}~@{text A} performs @{text outdate} action on |
|
404 theory @{text A} and all of its descendants. |
|
405 |
|
406 \item @{ML remove_thy}~@{text A} removes @{text A} and all of its |
|
407 descendants from the theory database. |
403 descendants from the theory database. |
408 |
404 |
409 \item @{ML ThyInfo.begin_theory} is the basic operation behind a |
405 \item @{ML ThyInfo.begin_theory} is the basic operation behind a |
410 @{text \<THEORY>} header declaration. The boolean argument |
406 @{text \<THEORY>} header declaration. The boolean argument |
411 indicates the strictness of treating ancestors: for @{ML true} (as |
407 indicates the strictness of treating ancestors: for @{ML true} (as |
412 in interactive mode) like @{ML update_thy}, and for @{ML false} (as |
408 in interactive mode) like @{ML update_thy}, and for @{ML false} (as |
413 in batch mode) like @{ML use_thy}. This is {\ML} functions is |
409 in batch mode) like @{ML use_thy}. This is {\ML} functions is |
414 normally not invoked directly. |
410 normally not invoked directly. |
415 |
411 |
416 \item @{ML ThyInfo.end_theory} concludes the loading of a theory |
412 \item @{ML ThyInfo.end_theory} concludes the loading of a theory |
417 proper; an attached theory {\ML} file may be still loaded later on. |
413 proper. An attached theory {\ML} file may be still loaded later on. |
418 This is {\ML} functions is normally not invoked directly. |
414 This is function is normally not invoked directly. |
419 |
415 |
420 \item @{ML ThyInfo.register_theory}~{text thy} registers an existing |
416 \item @{ML ThyInfo.register_theory}~@{text "text thy"} registers an |
421 theory value with the theory loader database. There is no |
417 existing theory value with the theory loader database. There is no |
422 management of associated sources; this is mainly for bootstrapping. |
418 management of associated sources. |
423 |
419 |
424 \item @{ML "ThyInfo.add_hook"}~@{text f} registers function @{text |
420 \item @{ML "ThyInfo.add_hook"}~@{text f} registers function @{text |
425 f} as a hook for theory database actions. The function will be |
421 f} as a hook for theory database actions. The function will be |
426 invoked with the action and theory name being involved; thus derived |
422 invoked with the action and theory name being involved; thus derived |
427 actions may be performed in associated system components, e.g.\ |
423 actions may be performed in associated system components, e.g.\ |
428 maintaining the state of an editor for theory sources. |
424 maintaining the state of an editor for the theory sources. |
429 |
425 |
430 The kind and order of actions occurring in practice depends both on |
426 The kind and order of actions occurring in practice depends both on |
431 user interactions and the internal process of resolving theory |
427 user interactions and the internal process of resolving theory |
432 imports. Hooks should not rely on a particular policy here! Any |
428 imports. Hooks should not rely on a particular policy here! Any |
433 exceptions raised by the hook are ignored by the theory database. |
429 exceptions raised by the hook are ignored. |
434 |
430 |
435 \end{description} |
431 \end{description} |
436 *} |
432 *} |
437 |
433 |
438 |
434 |
439 (* FIXME section {* Sessions and document preparation *} *) |
435 section {* Sessions and document preparation *} |
|
436 |
|
437 text FIXME |
440 |
438 |
441 end |
439 end |