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