author | wenzelm |
Sat, 21 Nov 2015 16:35:46 +0100 | |
changeset 61723 | 7feee72b5897 |
parent 61656 | cfabbc083977 |
child 61854 | 38b049cd3aad |
permissions | -rw-r--r-- |
61656 | 1 |
(*:maxLineLen=78:*) |
57341 | 2 |
|
29755 | 3 |
theory Integration |
4 |
imports Base |
|
5 |
begin |
|
18537 | 6 |
|
58618 | 7 |
chapter \<open>System integration\<close> |
18537 | 8 |
|
58618 | 9 |
section \<open>Isar toplevel \label{sec:isar-toplevel}\<close> |
18537 | 10 |
|
58618 | 11 |
text \<open> |
61477 | 12 |
The Isar \<^emph>\<open>toplevel state\<close> represents the outermost configuration that |
57343 | 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 |
|
58555 | 18 |
parallel in multi-threaded Isabelle/ML @{cite "Wenzel:2009" and |
19 |
"Wenzel:2013:ITP"}. |
|
58618 | 20 |
\<close> |
18537 | 21 |
|
22 |
||
58618 | 23 |
subsection \<open>Toplevel state\<close> |
18537 | 24 |
|
58618 | 25 |
text \<open> |
61493 | 26 |
The toplevel state is a disjoint sum of empty \<open>toplevel\<close>, or \<open>theory\<close>, or \<open>proof\<close>. The initial toplevel is empty; a theory is |
57343 | 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. |
|
58618 | 35 |
\<close> |
18537 | 36 |
|
58618 | 37 |
text %mlref \<open> |
18537 | 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 |
||
61439 | 46 |
\<^descr> Type @{ML_type Toplevel.state} represents Isar toplevel |
39864 | 47 |
states, which are normally manipulated through the concept of |
57343 | 48 |
toplevel transitions only (\secref{sec:toplevel-transition}). |
18537 | 49 |
|
61439 | 50 |
\<^descr> @{ML Toplevel.UNDEF} is raised for undefined toplevel |
20451 | 51 |
operations. Many operations work only partially for certain cases, |
52 |
since @{ML_type Toplevel.state} is a sum type. |
|
18537 | 53 |
|
61493 | 54 |
\<^descr> @{ML Toplevel.is_toplevel}~\<open>state\<close> checks for an empty |
20451 | 55 |
toplevel state. |
18537 | 56 |
|
61493 | 57 |
\<^descr> @{ML Toplevel.theory_of}~\<open>state\<close> selects the |
58 |
background theory of \<open>state\<close>, it raises @{ML Toplevel.UNDEF} |
|
34931 | 59 |
for an empty toplevel state. |
18537 | 60 |
|
61493 | 61 |
\<^descr> @{ML Toplevel.proof_of}~\<open>state\<close> selects the Isar proof |
60094 | 62 |
state if available, otherwise it raises an error. |
58618 | 63 |
\<close> |
18537 | 64 |
|
58618 | 65 |
text %mlantiq \<open> |
39882 | 66 |
\begin{matharray}{rcl} |
61493 | 67 |
@{ML_antiquotation_def "Isar.state"} & : & \<open>ML_antiquotation\<close> \\ |
39882 | 68 |
\end{matharray} |
69 |
||
61493 | 70 |
\<^descr> \<open>@{Isar.state}\<close> refers to Isar toplevel state at that |
39882 | 71 |
point --- as abstract value. |
72 |
||
73 |
This only works for diagnostic ML commands, such as @{command |
|
74 |
ML_val} or @{command ML_command}. |
|
58618 | 75 |
\<close> |
39882 | 76 |
|
18537 | 77 |
|
58618 | 78 |
subsection \<open>Toplevel transitions \label{sec:toplevel-transition}\<close> |
18537 | 79 |
|
58618 | 80 |
text \<open> |
57421 | 81 |
An Isar toplevel transition consists of a partial function on the toplevel |
82 |
state, with additional information for diagnostics and error reporting: |
|
83 |
there are fields for command name, source position, and other meta-data. |
|
18537 | 84 |
|
20451 | 85 |
The operational part is represented as the sequential union of a |
86 |
list of partial functions, which are tried in turn until the first |
|
20475 | 87 |
one succeeds. This acts like an outer case-expression for various |
34931 | 88 |
alternative state transitions. For example, \isakeyword{qed} works |
57496 | 89 |
differently for a local proofs vs.\ the global ending of an outermost |
20475 | 90 |
proof. |
18537 | 91 |
|
57496 | 92 |
Transitions are composed via transition transformers. Internally, Isar |
93 |
commands are put together from an empty transition extended by name and |
|
94 |
source position. It is then left to the individual command parser to turn |
|
95 |
the given concrete syntax into a suitable transition transformer that |
|
96 |
adjoins actual operations on a theory or proof state. |
|
58618 | 97 |
\<close> |
18537 | 98 |
|
58618 | 99 |
text %mlref \<open> |
18537 | 100 |
\begin{mldecls} |
101 |
@{index_ML Toplevel.keep: "(Toplevel.state -> unit) -> |
|
102 |
Toplevel.transition -> Toplevel.transition"} \\ |
|
103 |
@{index_ML Toplevel.theory: "(theory -> theory) -> |
|
104 |
Toplevel.transition -> Toplevel.transition"} \\ |
|
105 |
@{index_ML Toplevel.theory_to_proof: "(theory -> Proof.state) -> |
|
106 |
Toplevel.transition -> Toplevel.transition"} \\ |
|
107 |
@{index_ML Toplevel.proof: "(Proof.state -> Proof.state) -> |
|
108 |
Toplevel.transition -> Toplevel.transition"} \\ |
|
49864 | 109 |
@{index_ML Toplevel.proofs: "(Proof.state -> Proof.state Seq.result Seq.seq) -> |
18537 | 110 |
Toplevel.transition -> Toplevel.transition"} \\ |
21168
0f869edd6cc1
replaced Toplevel.proof_to_theory by Toplevel.end_proof;
wenzelm
parents:
20491
diff
changeset
|
111 |
@{index_ML Toplevel.end_proof: "(bool -> Proof.state -> Proof.context) -> |
18537 | 112 |
Toplevel.transition -> Toplevel.transition"} \\ |
113 |
\end{mldecls} |
|
114 |
||
61493 | 115 |
\<^descr> @{ML Toplevel.keep}~\<open>tr\<close> adjoins a diagnostic |
20451 | 116 |
function. |
18537 | 117 |
|
61493 | 118 |
\<^descr> @{ML Toplevel.theory}~\<open>tr\<close> adjoins a theory |
20451 | 119 |
transformer. |
18537 | 120 |
|
61493 | 121 |
\<^descr> @{ML Toplevel.theory_to_proof}~\<open>tr\<close> adjoins a global |
20451 | 122 |
goal function, which turns a theory into a proof state. The theory |
123 |
may be changed before entering the proof; the generic Isar goal |
|
61503 | 124 |
setup includes an \<^verbatim>\<open>after_qed\<close> argument that specifies how to |
57421 | 125 |
apply the proven result to the enclosing context, when the proof |
126 |
is finished. |
|
18558 | 127 |
|
61493 | 128 |
\<^descr> @{ML Toplevel.proof}~\<open>tr\<close> adjoins a deterministic |
20451 | 129 |
proof command, with a singleton result. |
18537 | 130 |
|
61493 | 131 |
\<^descr> @{ML Toplevel.proofs}~\<open>tr\<close> adjoins a general proof |
20451 | 132 |
command, with zero or more result states (represented as a lazy |
133 |
list). |
|
18537 | 134 |
|
61493 | 135 |
\<^descr> @{ML Toplevel.end_proof}~\<open>tr\<close> adjoins a concluding |
57496 | 136 |
proof command, that returns the resulting theory, after applying the |
137 |
resulting facts to the target context. |
|
58618 | 138 |
\<close> |
18537 | 139 |
|
59090 | 140 |
text %mlex \<open>The file @{"file" "~~/src/HOL/ex/Commands.thy"} shows some example |
141 |
Isar command definitions, with the all-important theory header declarations |
|
142 |
for outer syntax keywords.\<close> |
|
143 |
||
18537 | 144 |
|
58618 | 145 |
section \<open>Theory loader database\<close> |
18537 | 146 |
|
58618 | 147 |
text \<open> |
57341 | 148 |
In batch mode and within dumped logic images, the theory database maintains |
149 |
a collection of theories as a directed acyclic graph. A theory may refer to |
|
150 |
other theories as @{keyword "imports"}, or to auxiliary files via special |
|
61477 | 151 |
\<^emph>\<open>load commands\<close> (e.g.\ @{command ML_file}). For each theory, the base |
152 |
directory of its own theory file is called \<^emph>\<open>master directory\<close>: this is |
|
57341 | 153 |
used as the relative location to refer to other files from that theory. |
58618 | 154 |
\<close> |
18537 | 155 |
|
58618 | 156 |
text %mlref \<open> |
18537 | 157 |
\begin{mldecls} |
158 |
@{index_ML use_thy: "string -> unit"} \\ |
|
57496 | 159 |
@{index_ML use_thys: "string list -> unit"} \\[0.5ex] |
37864 | 160 |
@{index_ML Thy_Info.get_theory: "string -> theory"} \\ |
57341 | 161 |
@{index_ML Thy_Info.remove_thy: "string -> unit"} \\ |
162 |
@{index_ML Thy_Info.register_thy: "theory -> unit"} \\ |
|
18537 | 163 |
\end{mldecls} |
164 |
||
61493 | 165 |
\<^descr> @{ML use_thy}~\<open>A\<close> ensures that theory \<open>A\<close> is fully |
57496 | 166 |
up-to-date wrt.\ the external file store; outdated ancestors are reloaded on |
167 |
demand. |
|
18537 | 168 |
|
61439 | 169 |
\<^descr> @{ML use_thys} is similar to @{ML use_thy}, but handles several |
57341 | 170 |
theories simultaneously. Thus it acts like processing the import header of a |
171 |
theory, without performing the merge of the result. By loading a whole |
|
172 |
sub-graph of theories, the intrinsic parallelism can be exploited by the |
|
173 |
system to speedup loading. |
|
18537 | 174 |
|
60270 | 175 |
This variant is used by default in @{tool build} @{cite "isabelle-system"}. |
57496 | 176 |
|
61493 | 177 |
\<^descr> @{ML Thy_Info.get_theory}~\<open>A\<close> retrieves the theory value |
178 |
presently associated with name \<open>A\<close>. Note that the result might be |
|
57341 | 179 |
outdated wrt.\ the file-system content. |
37864 | 180 |
|
61493 | 181 |
\<^descr> @{ML Thy_Info.remove_thy}~\<open>A\<close> deletes theory \<open>A\<close> and all |
18537 | 182 |
descendants from the theory database. |
183 |
||
61493 | 184 |
\<^descr> @{ML Thy_Info.register_thy}~\<open>text thy\<close> registers an existing |
57341 | 185 |
theory value with the theory loader database and updates source version |
57496 | 186 |
information according to the file store. |
58618 | 187 |
\<close> |
30272 | 188 |
|
18537 | 189 |
end |