97 The @{text "checkpoint"} operation produces an intermediate stepping |
97 The @{text "checkpoint"} operation produces an intermediate stepping |
98 stone that will survive the next update unscathed: both the original |
98 stone that will survive the next update unscathed: both the original |
99 and the changed theory remain valid and are related by the |
99 and the changed theory remain valid and are related by the |
100 sub-theory relation. Checkpointing essentially recovers purely |
100 sub-theory relation. Checkpointing essentially recovers purely |
101 functional theory values, at the expense of some extra internal |
101 functional theory values, at the expense of some extra internal |
102 bookeeping. |
102 bookkeeping. |
103 |
103 |
104 The @{text "copy"} operation produces an auxiliary version that has |
104 The @{text "copy"} operation produces an auxiliary version that has |
105 the same data content, but is unrelated to the original: updates of |
105 the same data content, but is unrelated to the original: updates of |
106 the copy do not affect the original, neither does the sub-theory |
106 the copy do not affect the original, neither does the sub-theory |
107 relation hold. |
107 relation hold. |
109 \medskip The example in \figref{fig:ex-theory} below shows a theory |
109 \medskip The example in \figref{fig:ex-theory} below shows a theory |
110 graph derived from @{text "Pure"}. Theory @{text "Length"} imports |
110 graph derived from @{text "Pure"}. Theory @{text "Length"} imports |
111 @{text "Nat"} and @{text "List"}. The theory body consists of a |
111 @{text "Nat"} and @{text "List"}. The theory body consists of a |
112 sequence of updates, working mostly on drafts. Intermediate |
112 sequence of updates, working mostly on drafts. Intermediate |
113 checkpoints may occur as well, due to the history mechanism provided |
113 checkpoints may occur as well, due to the history mechanism provided |
114 by the Isar toplevel, cf.\ \secref{sec:isar-toplevel}. |
114 by the Isar top-level, cf.\ \secref{sec:isar-toplevel}. |
115 |
115 |
116 \begin{figure}[htb] |
116 \begin{figure}[htb] |
117 \begin{center} |
117 \begin{center} |
118 \begin{tabular}{rcccl} |
118 \begin{tabular}{rcccl} |
119 & & @{text "Pure"} \\ |
119 & & @{text "Pure"} \\ |
254 subsection {* Generic contexts *} |
254 subsection {* Generic contexts *} |
255 |
255 |
256 text {* |
256 text {* |
257 A generic context is the disjoint sum of either a theory or proof |
257 A generic context is the disjoint sum of either a theory or proof |
258 context. Occasionally, this simplifies uniform treatment of generic |
258 context. Occasionally, this simplifies uniform treatment of generic |
259 context data, typically extralogical information. Operations on |
259 context data, typically extra-logical information. Operations on |
260 generic contexts include the usual injections, partial selections, |
260 generic contexts include the usual injections, partial selections, |
261 and combinators for lifting operations on either component of the |
261 and combinators for lifting operations on either component of the |
262 disjoint sum. |
262 disjoint sum. |
263 |
263 |
264 Moreover, there are total operations @{text "theory_of"} and @{text |
264 Moreover, there are total operations @{text "theory_of"} and @{text |
303 \paragraph{Theory data} may refer to destructive entities, which are |
303 \paragraph{Theory data} may refer to destructive entities, which are |
304 maintained in correspondence to the linear evolution of theory |
304 maintained in correspondence to the linear evolution of theory |
305 values, or explicit copies.\footnote{Most existing instances of |
305 values, or explicit copies.\footnote{Most existing instances of |
306 destructive theory data are merely historical relics (e.g.\ the |
306 destructive theory data are merely historical relics (e.g.\ the |
307 destructive theorem storage, and destructive hints for the |
307 destructive theorem storage, and destructive hints for the |
308 Simplifier and Classical rules).} A theory data declaration needs to |
308 Simplifier and Classical rules).} A theory data declaration needs |
309 provide the following information: |
309 to implement the following specification: |
310 |
310 |
311 \medskip |
311 \medskip |
312 \begin{tabular}{ll} |
312 \begin{tabular}{ll} |
313 @{text "name: string"} \\ |
313 @{text "name: string"} \\ |
314 @{text "T"} & the ML type \\ |
314 @{text "T"} & the ML type \\ |
325 "extend"} is acts like a unitary version of @{text "merge"}, both |
325 "extend"} is acts like a unitary version of @{text "merge"}, both |
326 should also include the functionality of @{text "copy"} for impure |
326 should also include the functionality of @{text "copy"} for impure |
327 data. |
327 data. |
328 |
328 |
329 \paragraph{Proof context data} is purely functional. It is declared |
329 \paragraph{Proof context data} is purely functional. It is declared |
330 by providing the following information: |
330 by implementing the following specification: |
331 |
331 |
332 \medskip |
332 \medskip |
333 \begin{tabular}{ll} |
333 \begin{tabular}{ll} |
334 @{text "name: string"} \\ |
334 @{text "name: string"} \\ |
335 @{text "T"} & the ML type \\ |
335 @{text "T"} & the ML type \\ |
369 slot within a context. By keeping these operations private, a |
369 slot within a context. By keeping these operations private, a |
370 component may maintain abstract values authentically, without other |
370 component may maintain abstract values authentically, without other |
371 components interfering. |
371 components interfering. |
372 *} |
372 *} |
373 |
373 |
|
374 text %mlref {* |
|
375 \begin{mldecls} |
|
376 @{index_ML_functor TheoryDataFun} \\ |
|
377 @{index_ML_functor ProofDataFun} \\ |
|
378 @{index_ML_functor GenericDataFun} \\ |
|
379 \end{mldecls} |
|
380 |
|
381 \begin{description} |
|
382 |
|
383 \item @{ML_functor TheoryDataFun}@{text "(spec)"} declares data for |
|
384 type @{ML_type theory} according to the specification provided as |
|
385 argument structure. The result structure provides init and access |
|
386 operations as described above. |
|
387 |
|
388 \item @{ML_functor ProofDataFun}@{text "(spec)"} is analogous for |
|
389 type @{ML_type Proof.context}. |
|
390 |
|
391 \item @{ML_functor GenericDataFun}@{text "(spec)"} is analogous for |
|
392 type @{ML_type Context.generic}. |
|
393 |
|
394 \end{description} |
|
395 *} |
|
396 |
374 |
397 |
375 section {* Named entities *} |
398 section {* Named entities *} |
376 |
399 |
377 text {* Named entities of different kinds (logical constant, type, |
400 text {* Named entities of different kinds (logical constant, type, |
378 type class, theorem, method etc.) live in separate name spaces. It is |
401 type class, theorem, method etc.) live in separate name spaces. It is |
394 |
417 |
395 |
418 |
396 subsection {* Strings of symbols *} |
419 subsection {* Strings of symbols *} |
397 |
420 |
398 text {* Isabelle strings consist of a sequence of |
421 text {* Isabelle strings consist of a sequence of |
399 symbols\glossary{Symbol}{The smalles unit of text in Isabelle, |
422 symbols\glossary{Symbol}{The smallest unit of text in Isabelle, |
400 subsumes plain ASCII characters as well as an infinite collection of |
423 subsumes plain ASCII characters as well as an infinite collection of |
401 named symbols (for greek, math etc.).}, which are either packed as an |
424 named symbols (for greek, math etc.).}, which are either packed as an |
402 actual @{text "string"}, or represented as a list. Each symbol is in |
425 actual @{text "string"}, or represented as a list. Each symbol is in |
403 itself a small string of the following form: |
426 itself a small string of the following form: |
404 |
427 |
486 *} |
509 *} |
487 |
510 |
488 |
511 |
489 subsection {* Qualified names and name spaces *} |
512 subsection {* Qualified names and name spaces *} |
490 |
513 |
491 text %FIXME {* Qualified names are constructed according to implicit naming |
514 text {* |
492 principles of the present context. |
515 FIXME |
493 |
516 |
494 |
517 Qualified names are constructed according to implicit naming |
495 The last component is called \emph{base name}; the remaining prefix of |
518 principles of the present context. |
496 qualification may be empty. |
519 |
497 |
520 |
498 Some practical conventions help to organize named entities more |
521 The last component is called \emph{base name}; the remaining prefix |
499 systematically: |
522 of qualification may be empty. |
500 |
523 |
501 \begin{itemize} |
524 Some practical conventions help to organize named entities more |
502 |
525 systematically: |
503 \item Names are qualified first by the theory name, second by an |
526 |
504 optional ``structure''. For example, a constant @{text "c"} declared |
527 \begin{itemize} |
505 as part of a certain structure @{text "b"} (say a type definition) in |
528 |
506 theory @{text "A"} will be named @{text "A.b.c"} internally. |
529 \item Names are qualified first by the theory name, second by an |
507 |
530 optional ``structure''. For example, a constant @{text "c"} |
508 \item |
531 declared as part of a certain structure @{text "b"} (say a type |
509 |
532 definition) in theory @{text "A"} will be named @{text "A.b.c"} |
510 \item |
533 internally. |
511 |
534 |
512 \item |
535 \item |
513 |
536 |
514 \item |
537 \item |
515 |
538 |
516 \end{itemize} |
539 \item |
517 |
540 |
518 Names of different kinds of entities are basically independent, but |
541 \item |
519 some practical naming conventions relate them to each other. For |
542 |
520 example, a constant @{text "foo"} may be accompanied with theorems |
543 \end{itemize} |
521 @{text "foo.intro"}, @{text "foo.elim"}, @{text "foo.simps"} etc. The |
544 |
522 same may happen for a type @{text "foo"}, which is then apt to cause |
545 Names of different kinds of entities are basically independent, but |
523 clashes in the theorem name space! To avoid this, we occasionally |
546 some practical naming conventions relate them to each other. For |
524 follow an additional convention of suffixes that determine the |
547 example, a constant @{text "foo"} may be accompanied with theorems |
525 original kind of entity that a name has been derived. For example, |
548 @{text "foo.intro"}, @{text "foo.elim"}, @{text "foo.simps"} etc. |
526 constant @{text "foo"} is associated with theorem @{text "foo.intro"}, |
549 The same may happen for a type @{text "foo"}, which is then apt to |
527 type @{text "foo"} with theorem @{text "foo_type.intro"}, and type |
550 cause clashes in the theorem name space! To avoid this, we |
528 class @{text "foo"} with @{text "foo_class.intro"}. |
551 occasionally follow an additional convention of suffixes that |
529 |
552 determine the original kind of entity that a name has been derived. |
|
553 For example, constant @{text "foo"} is associated with theorem |
|
554 @{text "foo.intro"}, type @{text "foo"} with theorem @{text |
|
555 "foo_type.intro"}, and type class @{text "foo"} with @{text |
|
556 "foo_class.intro"}. |
530 *} |
557 *} |
531 |
558 |
532 |
559 |
533 section {* Structured output *} |
560 section {* Structured output *} |
534 |
561 |