author | wenzelm |
Thu, 03 Jul 2014 12:17:55 +0200 | |
changeset 57496 | 94596c573b38 |
parent 57421 | 94081154306d |
child 58555 | 7975676c08c0 |
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 |
|
57496 | 16 |
as sequential command loop, such that commands are applied one-by-one. In |
17 |
contemporary Isabelle/Isar, processing toplevel commands usually works in |
|
18 |
parallel in multi-threaded Isabelle/ML \cite{Wenzel:2009,Wenzel:2013:ITP}. |
|
57343 | 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 |
|
57496 | 31 |
is concluded we get back to the (local) theory, which is then updated by |
32 |
defining the resulting fact. Further theory declarations or theorem |
|
33 |
statements with proofs may follow, until we eventually conclude the theory |
|
34 |
development by 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"} \\ |
|
44 |
\end{mldecls} |
|
45 |
||
46 |
\begin{description} |
|
47 |
||
39864 | 48 |
\item Type @{ML_type Toplevel.state} represents Isar toplevel |
49 |
states, which are normally manipulated through the concept of |
|
57343 | 50 |
toplevel transitions only (\secref{sec:toplevel-transition}). |
18537 | 51 |
|
52 |
\item @{ML Toplevel.UNDEF} is raised for undefined toplevel |
|
20451 | 53 |
operations. Many operations work only partially for certain cases, |
54 |
since @{ML_type Toplevel.state} is a sum type. |
|
18537 | 55 |
|
20451 | 56 |
\item @{ML Toplevel.is_toplevel}~@{text "state"} checks for an empty |
57 |
toplevel state. |
|
18537 | 58 |
|
34931 | 59 |
\item @{ML Toplevel.theory_of}~@{text "state"} selects the |
57496 | 60 |
background theory of @{text "state"}, it raises @{ML Toplevel.UNDEF} |
34931 | 61 |
for an empty toplevel state. |
18537 | 62 |
|
20451 | 63 |
\item @{ML Toplevel.proof_of}~@{text "state"} selects the Isar proof |
57496 | 64 |
state if available, otherwise it raises @{ML Toplevel.UNDEF}. |
18537 | 65 |
|
66 |
\end{description} |
|
67 |
*} |
|
68 |
||
39882 | 69 |
text %mlantiq {* |
70 |
\begin{matharray}{rcl} |
|
71 |
@{ML_antiquotation_def "Isar.state"} & : & @{text ML_antiquotation} \\ |
|
72 |
\end{matharray} |
|
73 |
||
74 |
\begin{description} |
|
75 |
||
76 |
\item @{text "@{Isar.state}"} refers to Isar toplevel state at that |
|
77 |
point --- as abstract value. |
|
78 |
||
79 |
This only works for diagnostic ML commands, such as @{command |
|
80 |
ML_val} or @{command ML_command}. |
|
81 |
||
82 |
\end{description} |
|
83 |
*} |
|
84 |
||
18537 | 85 |
|
20451 | 86 |
subsection {* Toplevel transitions \label{sec:toplevel-transition} *} |
18537 | 87 |
|
20451 | 88 |
text {* |
57421 | 89 |
An Isar toplevel transition consists of a partial function on the toplevel |
90 |
state, with additional information for diagnostics and error reporting: |
|
91 |
there are fields for command name, source position, and other meta-data. |
|
18537 | 92 |
|
20451 | 93 |
The operational part is represented as the sequential union of a |
94 |
list of partial functions, which are tried in turn until the first |
|
20475 | 95 |
one succeeds. This acts like an outer case-expression for various |
34931 | 96 |
alternative state transitions. For example, \isakeyword{qed} works |
57496 | 97 |
differently for a local proofs vs.\ the global ending of an outermost |
20475 | 98 |
proof. |
18537 | 99 |
|
57496 | 100 |
Transitions are composed via transition transformers. Internally, Isar |
101 |
commands are put together from an empty transition extended by name and |
|
102 |
source position. It is then left to the individual command parser to turn |
|
103 |
the given concrete syntax into a suitable transition transformer that |
|
104 |
adjoins actual operations on a theory or proof state. |
|
18537 | 105 |
*} |
106 |
||
107 |
text %mlref {* |
|
108 |
\begin{mldecls} |
|
109 |
@{index_ML Toplevel.keep: "(Toplevel.state -> unit) -> |
|
110 |
Toplevel.transition -> Toplevel.transition"} \\ |
|
111 |
@{index_ML Toplevel.theory: "(theory -> theory) -> |
|
112 |
Toplevel.transition -> Toplevel.transition"} \\ |
|
113 |
@{index_ML Toplevel.theory_to_proof: "(theory -> Proof.state) -> |
|
114 |
Toplevel.transition -> Toplevel.transition"} \\ |
|
115 |
@{index_ML Toplevel.proof: "(Proof.state -> Proof.state) -> |
|
116 |
Toplevel.transition -> Toplevel.transition"} \\ |
|
49864 | 117 |
@{index_ML Toplevel.proofs: "(Proof.state -> Proof.state Seq.result Seq.seq) -> |
18537 | 118 |
Toplevel.transition -> Toplevel.transition"} \\ |
21168
0f869edd6cc1
replaced Toplevel.proof_to_theory by Toplevel.end_proof;
wenzelm
parents:
20491
diff
changeset
|
119 |
@{index_ML Toplevel.end_proof: "(bool -> Proof.state -> Proof.context) -> |
18537 | 120 |
Toplevel.transition -> Toplevel.transition"} \\ |
121 |
\end{mldecls} |
|
122 |
||
123 |
\begin{description} |
|
124 |
||
20451 | 125 |
\item @{ML Toplevel.keep}~@{text "tr"} adjoins a diagnostic |
126 |
function. |
|
18537 | 127 |
|
20451 | 128 |
\item @{ML Toplevel.theory}~@{text "tr"} adjoins a theory |
129 |
transformer. |
|
18537 | 130 |
|
20451 | 131 |
\item @{ML Toplevel.theory_to_proof}~@{text "tr"} adjoins a global |
132 |
goal function, which turns a theory into a proof state. The theory |
|
133 |
may be changed before entering the proof; the generic Isar goal |
|
57421 | 134 |
setup includes an @{verbatim after_qed} argument that specifies how to |
135 |
apply the proven result to the enclosing context, when the proof |
|
136 |
is finished. |
|
18558 | 137 |
|
20451 | 138 |
\item @{ML Toplevel.proof}~@{text "tr"} adjoins a deterministic |
139 |
proof command, with a singleton result. |
|
18537 | 140 |
|
20451 | 141 |
\item @{ML Toplevel.proofs}~@{text "tr"} adjoins a general proof |
142 |
command, with zero or more result states (represented as a lazy |
|
143 |
list). |
|
18537 | 144 |
|
21168
0f869edd6cc1
replaced Toplevel.proof_to_theory by Toplevel.end_proof;
wenzelm
parents:
20491
diff
changeset
|
145 |
\item @{ML Toplevel.end_proof}~@{text "tr"} adjoins a concluding |
57496 | 146 |
proof command, that returns the resulting theory, after applying the |
147 |
resulting facts to the target context. |
|
18537 | 148 |
|
149 |
\end{description} |
|
150 |
*} |
|
151 |
||
152 |
||
57341 | 153 |
section {* Theory loader database *} |
18537 | 154 |
|
20451 | 155 |
text {* |
57341 | 156 |
In batch mode and within dumped logic images, the theory database maintains |
157 |
a collection of theories as a directed acyclic graph. A theory may refer to |
|
158 |
other theories as @{keyword "imports"}, or to auxiliary files via special |
|
159 |
\emph{load commands} (e.g.\ @{command ML_file}). For each theory, the base |
|
160 |
directory of its own theory file is called \emph{master directory}: this is |
|
161 |
used as the relative location to refer to other files from that theory. |
|
18537 | 162 |
*} |
163 |
||
164 |
text %mlref {* |
|
165 |
\begin{mldecls} |
|
166 |
@{index_ML use_thy: "string -> unit"} \\ |
|
57496 | 167 |
@{index_ML use_thys: "string list -> unit"} \\[0.5ex] |
37864 | 168 |
@{index_ML Thy_Info.get_theory: "string -> theory"} \\ |
57341 | 169 |
@{index_ML Thy_Info.remove_thy: "string -> unit"} \\ |
170 |
@{index_ML Thy_Info.register_thy: "theory -> unit"} \\ |
|
18537 | 171 |
\end{mldecls} |
172 |
||
173 |
\begin{description} |
|
174 |
||
24173
4ba00a972eb8
theory loader: added use_thys, removed obsolete update_thy;
wenzelm
parents:
22095
diff
changeset
|
175 |
\item @{ML use_thy}~@{text A} ensures that theory @{text A} is fully |
57496 | 176 |
up-to-date wrt.\ the external file store; outdated ancestors are reloaded on |
177 |
demand. |
|
18537 | 178 |
|
57341 | 179 |
\item @{ML use_thys} is similar to @{ML use_thy}, but handles several |
180 |
theories simultaneously. Thus it acts like processing the import header of a |
|
181 |
theory, without performing the merge of the result. By loading a whole |
|
182 |
sub-graph of theories, the intrinsic parallelism can be exploited by the |
|
183 |
system to speedup loading. |
|
18537 | 184 |
|
57496 | 185 |
This variant is used by default in @{tool build} \cite{isabelle-sys}. |
186 |
||
37864 | 187 |
\item @{ML Thy_Info.get_theory}~@{text A} retrieves the theory value |
57341 | 188 |
presently associated with name @{text A}. Note that the result might be |
189 |
outdated wrt.\ the file-system content. |
|
37864 | 190 |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
34931
diff
changeset
|
191 |
\item @{ML Thy_Info.remove_thy}~@{text A} deletes theory @{text A} and all |
18537 | 192 |
descendants from the theory database. |
193 |
||
57341 | 194 |
\item @{ML Thy_Info.register_thy}~@{text "text thy"} registers an existing |
195 |
theory value with the theory loader database and updates source version |
|
57496 | 196 |
information according to the file store. |
18537 | 197 |
|
198 |
\end{description} |
|
199 |
*} |
|
30272 | 200 |
|
18537 | 201 |
end |