| author | wenzelm |
| Sat, 28 Jun 2014 15:34:43 +0200 | |
| changeset 57421 | 94081154306d |
| parent 57343 | e0f573aca42f |
| child 57496 | 94596c573b38 |
| permissions | -rw-r--r-- |
| 57341 | 1 |
(*:wrap=hard:maxLineLen=78:*) |
2 |
||
| 29755 | 3 |
theory Integration |
4 |
imports Base |
|
5 |
begin |
|
| 18537 | 6 |
|
7 |
chapter {* System integration *}
|
|
8 |
||
| 20447 | 9 |
section {* Isar toplevel \label{sec:isar-toplevel} *}
|
| 18537 | 10 |
|
| 57343 | 11 |
text {*
|
12 |
The Isar \emph{toplevel state} represents the outermost configuration that
|
|
13 |
is transformed by a sequence of transitions (commands) within a theory body. |
|
14 |
This is a pure value with pure functions acting on it in a timeless and |
|
15 |
stateless manner. Historically, the sequence of transitions was wrapped up |
|
16 |
as sequential command loop, such that commands are applied sequentially |
|
17 |
one-by-one. In contemporary Isabelle/Isar, processing toplevel commands |
|
18 |
usually works in parallel in multi-threaded Isabelle/ML. |
|
19 |
*} |
|
| 18537 | 20 |
|
21 |
||
| 57343 | 22 |
subsection {* Toplevel state *}
|
| 18537 | 23 |
|
| 57343 | 24 |
text {*
|
25 |
The toplevel state is a disjoint sum of empty @{text toplevel}, or @{text
|
|
26 |
theory}, or @{text proof}. The initial toplevel is empty; a theory is
|
|
27 |
commenced by a @{command theory} header; within a theory we may use theory
|
|
28 |
commands such as @{command definition}, or state a @{command theorem} to be
|
|
29 |
proven. A proof state accepts a rich collection of Isar proof commands for |
|
30 |
structured proof composition, or unstructured proof scripts. When the proof |
|
31 |
is concluded we get back to the theory, which is then updated by defining |
|
32 |
the resulting fact. Further theory declarations or theorem statements with |
|
33 |
proofs may follow, until we eventually conclude the theory development by |
|
34 |
issuing @{command end} to get back to the empty toplevel.
|
|
| 18537 | 35 |
*} |
36 |
||
37 |
text %mlref {*
|
|
38 |
\begin{mldecls}
|
|
39 |
@{index_ML_type Toplevel.state} \\
|
|
| 55838 | 40 |
@{index_ML_exception Toplevel.UNDEF} \\
|
| 18537 | 41 |
@{index_ML Toplevel.is_toplevel: "Toplevel.state -> bool"} \\
|
42 |
@{index_ML Toplevel.theory_of: "Toplevel.state -> theory"} \\
|
|
43 |
@{index_ML Toplevel.proof_of: "Toplevel.state -> Proof.state"} \\
|
|
| 32833 | 44 |
@{index_ML Toplevel.timing: "bool Unsynchronized.ref"} \\
|
45 |
@{index_ML Toplevel.profiling: "int Unsynchronized.ref"} \\
|
|
| 18537 | 46 |
\end{mldecls}
|
47 |
||
48 |
\begin{description}
|
|
49 |
||
| 39864 | 50 |
\item Type @{ML_type Toplevel.state} represents Isar toplevel
|
51 |
states, which are normally manipulated through the concept of |
|
| 57343 | 52 |
toplevel transitions only (\secref{sec:toplevel-transition}).
|
| 18537 | 53 |
|
54 |
\item @{ML Toplevel.UNDEF} is raised for undefined toplevel
|
|
| 20451 | 55 |
operations. Many operations work only partially for certain cases, |
56 |
since @{ML_type Toplevel.state} is a sum type.
|
|
| 18537 | 57 |
|
| 20451 | 58 |
\item @{ML Toplevel.is_toplevel}~@{text "state"} checks for an empty
|
59 |
toplevel state. |
|
| 18537 | 60 |
|
| 34931 | 61 |
\item @{ML Toplevel.theory_of}~@{text "state"} selects the
|
62 |
background theory of @{text "state"}, raises @{ML Toplevel.UNDEF}
|
|
63 |
for an empty toplevel state. |
|
| 18537 | 64 |
|
| 20451 | 65 |
\item @{ML Toplevel.proof_of}~@{text "state"} selects the Isar proof
|
66 |
state if available, otherwise raises @{ML Toplevel.UNDEF}.
|
|
| 18537 | 67 |
|
| 32833 | 68 |
\item @{ML "Toplevel.timing := true"} makes the toplevel print timing
|
| 18537 | 69 |
information for each Isar command being executed. |
70 |
||
|
40149
4c35be108990
proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents:
39882
diff
changeset
|
71 |
\item @{ML Toplevel.profiling}~@{ML_text ":="}~@{text "n"} controls
|
|
39825
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents:
37982
diff
changeset
|
72 |
low-level profiling of the underlying ML runtime system. For |
| 20451 | 73 |
Poly/ML, @{text "n = 1"} means time and @{text "n = 2"} space
|
74 |
profiling. |
|
| 18537 | 75 |
|
76 |
\end{description}
|
|
77 |
*} |
|
78 |
||
| 39882 | 79 |
text %mlantiq {*
|
80 |
\begin{matharray}{rcl}
|
|
81 |
@{ML_antiquotation_def "Isar.state"} & : & @{text ML_antiquotation} \\
|
|
82 |
\end{matharray}
|
|
83 |
||
84 |
\begin{description}
|
|
85 |
||
86 |
\item @{text "@{Isar.state}"} refers to Isar toplevel state at that
|
|
87 |
point --- as abstract value. |
|
88 |
||
89 |
This only works for diagnostic ML commands, such as @{command
|
|
90 |
ML_val} or @{command ML_command}.
|
|
91 |
||
92 |
\end{description}
|
|
93 |
*} |
|
94 |
||
| 18537 | 95 |
|
| 20451 | 96 |
subsection {* Toplevel transitions \label{sec:toplevel-transition} *}
|
| 18537 | 97 |
|
| 20451 | 98 |
text {*
|
| 57421 | 99 |
An Isar toplevel transition consists of a partial function on the toplevel |
100 |
state, with additional information for diagnostics and error reporting: |
|
101 |
there are fields for command name, source position, and other meta-data. |
|
| 18537 | 102 |
|
| 20451 | 103 |
The operational part is represented as the sequential union of a |
104 |
list of partial functions, which are tried in turn until the first |
|
| 20475 | 105 |
one succeeds. This acts like an outer case-expression for various |
| 34931 | 106 |
alternative state transitions. For example, \isakeyword{qed} works
|
| 20475 | 107 |
differently for a local proofs vs.\ the global ending of the main |
108 |
proof. |
|
| 18537 | 109 |
|
110 |
Toplevel transitions are composed via transition transformers. |
|
111 |
Internally, Isar commands are put together from an empty transition |
|
| 34931 | 112 |
extended by name and source position. It is then left to the |
113 |
individual command parser to turn the given concrete syntax into a |
|
114 |
suitable transition transformer that adjoins actual operations on a |
|
115 |
theory or proof state etc. |
|
| 18537 | 116 |
*} |
117 |
||
118 |
text %mlref {*
|
|
119 |
\begin{mldecls}
|
|
120 |
@{index_ML Toplevel.keep: "(Toplevel.state -> unit) ->
|
|
121 |
Toplevel.transition -> Toplevel.transition"} \\ |
|
122 |
@{index_ML Toplevel.theory: "(theory -> theory) ->
|
|
123 |
Toplevel.transition -> Toplevel.transition"} \\ |
|
124 |
@{index_ML Toplevel.theory_to_proof: "(theory -> Proof.state) ->
|
|
125 |
Toplevel.transition -> Toplevel.transition"} \\ |
|
126 |
@{index_ML Toplevel.proof: "(Proof.state -> Proof.state) ->
|
|
127 |
Toplevel.transition -> Toplevel.transition"} \\ |
|
| 49864 | 128 |
@{index_ML Toplevel.proofs: "(Proof.state -> Proof.state Seq.result Seq.seq) ->
|
| 18537 | 129 |
Toplevel.transition -> Toplevel.transition"} \\ |
|
21168
0f869edd6cc1
replaced Toplevel.proof_to_theory by Toplevel.end_proof;
wenzelm
parents:
20491
diff
changeset
|
130 |
@{index_ML Toplevel.end_proof: "(bool -> Proof.state -> Proof.context) ->
|
| 18537 | 131 |
Toplevel.transition -> Toplevel.transition"} \\ |
132 |
\end{mldecls}
|
|
133 |
||
134 |
\begin{description}
|
|
135 |
||
| 20451 | 136 |
\item @{ML Toplevel.keep}~@{text "tr"} adjoins a diagnostic
|
137 |
function. |
|
| 18537 | 138 |
|
| 20451 | 139 |
\item @{ML Toplevel.theory}~@{text "tr"} adjoins a theory
|
140 |
transformer. |
|
| 18537 | 141 |
|
| 20451 | 142 |
\item @{ML Toplevel.theory_to_proof}~@{text "tr"} adjoins a global
|
143 |
goal function, which turns a theory into a proof state. The theory |
|
144 |
may be changed before entering the proof; the generic Isar goal |
|
| 57421 | 145 |
setup includes an @{verbatim after_qed} argument that specifies how to
|
146 |
apply the proven result to the enclosing context, when the proof |
|
147 |
is finished. |
|
| 18558 | 148 |
|
| 20451 | 149 |
\item @{ML Toplevel.proof}~@{text "tr"} adjoins a deterministic
|
150 |
proof command, with a singleton result. |
|
| 18537 | 151 |
|
| 20451 | 152 |
\item @{ML Toplevel.proofs}~@{text "tr"} adjoins a general proof
|
153 |
command, with zero or more result states (represented as a lazy |
|
154 |
list). |
|
| 18537 | 155 |
|
|
21168
0f869edd6cc1
replaced Toplevel.proof_to_theory by Toplevel.end_proof;
wenzelm
parents:
20491
diff
changeset
|
156 |
\item @{ML Toplevel.end_proof}~@{text "tr"} adjoins a concluding
|
|
0f869edd6cc1
replaced Toplevel.proof_to_theory by Toplevel.end_proof;
wenzelm
parents:
20491
diff
changeset
|
157 |
proof command, that returns the resulting theory, after storing the |
|
0f869edd6cc1
replaced Toplevel.proof_to_theory by Toplevel.end_proof;
wenzelm
parents:
20491
diff
changeset
|
158 |
resulting facts in the context etc. |
| 18537 | 159 |
|
160 |
\end{description}
|
|
161 |
*} |
|
162 |
||
163 |
||
| 57341 | 164 |
section {* Theory loader database *}
|
| 18537 | 165 |
|
| 20451 | 166 |
text {*
|
| 57341 | 167 |
In batch mode and within dumped logic images, the theory database maintains |
168 |
a collection of theories as a directed acyclic graph. A theory may refer to |
|
169 |
other theories as @{keyword "imports"}, or to auxiliary files via special
|
|
170 |
\emph{load commands} (e.g.\ @{command ML_file}). For each theory, the base
|
|
171 |
directory of its own theory file is called \emph{master directory}: this is
|
|
172 |
used as the relative location to refer to other files from that theory. |
|
| 18537 | 173 |
*} |
174 |
||
175 |
text %mlref {*
|
|
176 |
\begin{mldecls}
|
|
177 |
@{index_ML use_thy: "string -> unit"} \\
|
|
|
24173
4ba00a972eb8
theory loader: added use_thys, removed obsolete update_thy;
wenzelm
parents:
22095
diff
changeset
|
178 |
@{index_ML use_thys: "string list -> unit"} \\
|
| 37864 | 179 |
@{index_ML Thy_Info.get_theory: "string -> theory"} \\
|
| 57341 | 180 |
@{index_ML Thy_Info.remove_thy: "string -> unit"} \\
|
181 |
@{index_ML Thy_Info.register_thy: "theory -> unit"} \\
|
|
| 18537 | 182 |
\end{mldecls}
|
183 |
||
184 |
\begin{description}
|
|
185 |
||
|
24173
4ba00a972eb8
theory loader: added use_thys, removed obsolete update_thy;
wenzelm
parents:
22095
diff
changeset
|
186 |
\item @{ML use_thy}~@{text A} ensures that theory @{text A} is fully
|
| 57341 | 187 |
up-to-date wrt.\ the external file store, reloading outdated ancestors as |
188 |
required. |
|
| 18537 | 189 |
|
| 57341 | 190 |
\item @{ML use_thys} is similar to @{ML use_thy}, but handles several
|
191 |
theories simultaneously. Thus it acts like processing the import header of a |
|
192 |
theory, without performing the merge of the result. By loading a whole |
|
193 |
sub-graph of theories, the intrinsic parallelism can be exploited by the |
|
194 |
system to speedup loading. |
|
| 18537 | 195 |
|
| 37864 | 196 |
\item @{ML Thy_Info.get_theory}~@{text A} retrieves the theory value
|
| 57341 | 197 |
presently associated with name @{text A}. Note that the result might be
|
198 |
outdated wrt.\ the file-system content. |
|
| 37864 | 199 |
|
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
34931
diff
changeset
|
200 |
\item @{ML Thy_Info.remove_thy}~@{text A} deletes theory @{text A} and all
|
| 18537 | 201 |
descendants from the theory database. |
202 |
||
| 57341 | 203 |
\item @{ML Thy_Info.register_thy}~@{text "text thy"} registers an existing
|
204 |
theory value with the theory loader database and updates source version |
|
205 |
information according to the current file-system state. |
|
| 18537 | 206 |
|
207 |
\end{description}
|
|
208 |
*} |
|
| 30272 | 209 |
|
| 18537 | 210 |
end |