| author | wenzelm | 
| Fri, 31 Aug 2001 22:46:23 +0200 | |
| changeset 11542 | 2afde2de26d6 | 
| parent 11052 | 1379e49c0ee9 | 
| child 11623 | 9c95b6a76e15 | 
| permissions | -rw-r--r-- | 
| 3201 | 1  | 
|
| 104 | 2  | 
%% $Id$  | 
| 145 | 3  | 
|
| 104 | 4  | 
\chapter{Theories, Terms and Types} \label{theories}
 | 
5  | 
\index{theories|(}\index{signatures|bold}
 | 
|
| 9695 | 6  | 
\index{reading!axioms|see{\texttt{assume_ax}}} Theories organize the syntax,
 | 
7  | 
declarations and axioms of a mathematical development. They are built,  | 
|
8  | 
starting from the Pure or CPure theory, by extending and merging existing  | 
|
9  | 
theories.  They have the \ML\ type \mltydx{theory}.  Theory operations signal
 | 
|
10  | 
errors by raising exception \xdx{THEORY}, returning a message and a list of
 | 
|
11  | 
theories.  | 
|
| 104 | 12  | 
|
13  | 
Signatures, which contain information about sorts, types, constants and  | 
|
| 332 | 14  | 
syntax, have the \ML\ type~\mltydx{Sign.sg}.  For identification, each
 | 
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
15  | 
signature carries a unique list of \bfindex{stamps}, which are \ML\
 | 
| 324 | 16  | 
references to strings. The strings serve as human-readable names; the  | 
| 104 | 17  | 
references serve as unique identifiers. Each primitive signature has a  | 
18  | 
single stamp. When two signatures are merged, their lists of stamps are  | 
|
19  | 
also merged. Every theory carries a unique signature.  | 
|
20  | 
||
21  | 
Terms and types are the underlying representation of logical syntax. Their  | 
|
| 275 | 22  | 
\ML\ definitions are irrelevant to naive Isabelle users. Programmers who  | 
23  | 
wish to extend Isabelle may need to know such details, say to code a tactic  | 
|
24  | 
that looks for subgoals of a particular form. Terms and types may be  | 
|
| 104 | 25  | 
`certified' to be well-formed with respect to a given signature.  | 
26  | 
||
| 324 | 27  | 
|
28  | 
\section{Defining theories}\label{sec:ref-defining-theories}
 | 
|
| 104 | 29  | 
|
| 6571 | 30  | 
Theories are defined via theory files $name$\texttt{.thy} (there are also
 | 
31  | 
\ML-level interfaces which are only intended for people building advanced  | 
|
32  | 
theory definition packages).  Appendix~\ref{app:TheorySyntax} presents the
 | 
|
33  | 
concrete syntax for theory files; here follows an explanation of the  | 
|
34  | 
constituent parts.  | 
|
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
35  | 
\begin{description}
 | 
| 6568 | 36  | 
\item[{\it theoryDef}] is the full definition.  The new theory is called $id$.
 | 
| 8136 | 37  | 
  It is the union of the named \textbf{parent
 | 
| 3108 | 38  | 
    theories}\indexbold{theories!parent}, possibly extended with new
 | 
| 6568 | 39  | 
  components.  \thydx{Pure} and \thydx{CPure} are the basic theories, which
 | 
40  | 
contain only the meta-logic. They differ just in their concrete syntax for  | 
|
41  | 
function applications.  | 
|
| 6571 | 42  | 
|
43  | 
The new theory begins as a merge of its parents.  | 
|
44  | 
  \begin{ttbox}
 | 
|
45  | 
Attempt to merge different versions of theories: "\(T@1\)", \(\ldots\), "\(T@n\)"  | 
|
46  | 
  \end{ttbox}
 | 
|
47  | 
This error may especially occur when a theory is redeclared --- say to  | 
|
48  | 
change an inappropriate definition --- and bindings to old versions persist.  | 
|
49  | 
Isabelle ensures that old and new theories of the same name are not involved  | 
|
50  | 
in a proof.  | 
|
| 324 | 51  | 
|
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
52  | 
\item[$classes$]  | 
| 
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
53  | 
  is a series of class declarations.  Declaring {\tt$id$ < $id@1$ \dots\
 | 
| 324 | 54  | 
$id@n$} makes $id$ a subclass of the existing classes $id@1\dots  | 
55  | 
id@n$. This rules out cyclic class structures. Isabelle automatically  | 
|
56  | 
computes the transitive closure of subclass hierarchies; it is not  | 
|
| 6669 | 57  | 
  necessary to declare \texttt{c < e} in addition to \texttt{c < d} and \texttt{d <
 | 
| 324 | 58  | 
e}.  | 
59  | 
||
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
60  | 
\item[$default$]  | 
| 
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
61  | 
introduces $sort$ as the new default sort for type variables. This applies  | 
| 
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
62  | 
to unconstrained type variables in an input string but not to type  | 
| 
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
63  | 
variables created internally. If omitted, the default sort is the listwise  | 
| 
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
64  | 
union of the default sorts of the parent theories (i.e.\ their logical  | 
| 
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
65  | 
intersection).  | 
| 3108 | 66  | 
|
| 7168 | 67  | 
\item[$sort$] is a finite set of classes. A single class $id$ abbreviates the  | 
68  | 
  sort $\{id\}$.
 | 
|
| 324 | 69  | 
|
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
70  | 
\item[$types$]  | 
| 324 | 71  | 
is a series of type declarations. Each declares a new type constructor  | 
72  | 
or type synonym. An $n$-place type constructor is specified by  | 
|
73  | 
$(\alpha@1,\dots,\alpha@n)name$, where the type variables serve only to  | 
|
74  | 
indicate the number~$n$.  | 
|
75  | 
||
| 8136 | 76  | 
  A \textbf{type synonym}\indexbold{type synonyms} is an abbreviation
 | 
| 1387 | 77  | 
$(\alpha@1,\dots,\alpha@n)name = \tau$, where $name$ and $\tau$ can  | 
78  | 
be strings.  | 
|
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
79  | 
|
| 
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
80  | 
\item[$infix$]  | 
| 8136 | 81  | 
declares a type or constant to be an infix operator having priority $nat$  | 
82  | 
  and associating to the left (\texttt{infixl}) or right (\texttt{infixr}).
 | 
|
83  | 
  Only 2-place type constructors can have infix status; an example is {\tt
 | 
|
| 3108 | 84  | 
  ('a,'b)~"*"~(infixr~20)}, which may express binary product types.
 | 
| 324 | 85  | 
|
| 3108 | 86  | 
\item[$arities$] is a series of type arity declarations. Each assigns  | 
87  | 
arities to type constructors. The $name$ must be an existing type  | 
|
88  | 
constructor, which is given the additional arity $arity$.  | 
|
89  | 
||
| 5369 | 90  | 
\item[$nonterminals$]\index{*nonterminal symbols} declares purely
 | 
91  | 
syntactic types to be used as nonterminal symbols of the context  | 
|
92  | 
free grammar.  | 
|
93  | 
||
| 3108 | 94  | 
\item[$consts$] is a series of constant declarations. Each new  | 
95  | 
constant $name$ is given the specified type. The optional $mixfix$  | 
|
96  | 
annotations may attach concrete syntax to the constant.  | 
|
97  | 
||
98  | 
\item[$syntax$] \index{*syntax section}\index{print mode} is a variant
 | 
|
99  | 
of $consts$ which adds just syntax without actually declaring  | 
|
100  | 
logical constants. This gives full control over a theory's context  | 
|
| 
3485
 
f27a30a18a17
Now there are TWO spaces after each full stop, so that the Emacs sentence
 
paulson 
parents: 
3201 
diff
changeset
 | 
101  | 
free grammar. The optional $mode$ specifies the print mode where the  | 
| 
 
f27a30a18a17
Now there are TWO spaces after each full stop, so that the Emacs sentence
 
paulson 
parents: 
3201 
diff
changeset
 | 
102  | 
  mixfix productions should be added.  If there is no \texttt{output}
 | 
| 3108 | 103  | 
option given, all productions are also added to the input syntax  | 
104  | 
(regardless of the print mode).  | 
|
| 324 | 105  | 
|
106  | 
\item[$mixfix$] \index{mixfix declarations}
 | 
|
107  | 
annotations can take three forms:  | 
|
| 273 | 108  | 
  \begin{itemize}
 | 
109  | 
\item A mixfix template given as a $string$ of the form  | 
|
110  | 
    {\tt"}\dots{\tt\_}\dots{\tt\_}\dots{\tt"} where the $i$-th underscore
 | 
|
| 324 | 111  | 
indicates the position where the $i$-th argument should go. The list  | 
112  | 
of numbers gives the priority of each argument. The final number gives  | 
|
113  | 
the priority of the whole construct.  | 
|
| 104 | 114  | 
|
| 324 | 115  | 
  \item A constant $f$ of type $\tau@1\To(\tau@2\To\tau)$ can be given {\bf
 | 
116  | 
infix} status.  | 
|
| 104 | 117  | 
|
| 324 | 118  | 
  \item A constant $f$ of type $(\tau@1\To\tau@2)\To\tau$ can be given {\bf
 | 
| 6669 | 119  | 
    binder} status.  The declaration \texttt{binder} $\cal Q$ $p$ causes
 | 
| 286 | 120  | 
  ${\cal Q}\,x.F(x)$ to be treated
 | 
121  | 
like $f(F)$, where $p$ is the priority.  | 
|
| 273 | 122  | 
  \end{itemize}
 | 
| 324 | 123  | 
|
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
124  | 
\item[$trans$]  | 
| 
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
125  | 
specifies syntactic translation rules (macros). There are three forms:  | 
| 6669 | 126  | 
  parse rules (\texttt{=>}), print rules (\texttt{<=}), and parse/print rules ({\tt
 | 
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
127  | 
==}).  | 
| 324 | 128  | 
|
| 1650 | 129  | 
\item[$rules$]  | 
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
130  | 
is a series of rule declarations. Each has a name $id$ and the formula is  | 
| 
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
131  | 
given by the $string$. Rule names must be distinct within any single  | 
| 3108 | 132  | 
theory.  | 
| 324 | 133  | 
|
| 1905 | 134  | 
\item[$defs$] is a series of definitions. They are just like $rules$, except  | 
135  | 
that every $string$ must be a definition (see below for details).  | 
|
| 1650 | 136  | 
|
137  | 
\item[$constdefs$] combines the declaration of constants and their  | 
|
| 
3485
 
f27a30a18a17
Now there are TWO spaces after each full stop, so that the Emacs sentence
 
paulson 
parents: 
3201 
diff
changeset
 | 
138  | 
definition. The first $string$ is the type, the second the definition.  | 
| 3108 | 139  | 
|
| 6625 | 140  | 
\item[$axclass$] \index{*axclass section} defines an \rmindex{axiomatic type
 | 
141  | 
    class} \cite{Wenzel:1997:TPHOL} as the intersection of existing classes,
 | 
|
142  | 
with additional axioms holding. Class axioms may not contain more than one  | 
|
143  | 
type variable. The class axioms (with implicit sort constraints added) are  | 
|
144  | 
bound to the given names. Furthermore a class introduction rule is  | 
|
145  | 
generated, which is automatically employed by $instance$ to prove  | 
|
146  | 
instantiations of this class.  | 
|
| 3108 | 147  | 
|
148  | 
\item[$instance$] \index{*instance section} proves class inclusions or
 | 
|
149  | 
type arities at the logical level and then transfers these to the  | 
|
| 
3485
 
f27a30a18a17
Now there are TWO spaces after each full stop, so that the Emacs sentence
 
paulson 
parents: 
3201 
diff
changeset
 | 
150  | 
type signature. The instantiation is proven and checked properly.  | 
| 3108 | 151  | 
The user has to supply sufficient witness information: theorems  | 
152  | 
  ($longident$), axioms ($string$), or even arbitrary \ML{} tactic
 | 
|
153  | 
code $verbatim$.  | 
|
| 1650 | 154  | 
|
| 1846 | 155  | 
\item[$oracle$] links the theory to a trusted external reasoner. It is  | 
156  | 
allowed to create theorems, but each theorem carries a proof object  | 
|
157  | 
  describing the oracle invocation.  See \S\ref{sec:oracles} for details.
 | 
|
| 4543 | 158  | 
|
| 5369 | 159  | 
\item[$local$, $global$] change the current name declaration mode.  | 
| 4543 | 160  | 
Initially, theories start in $local$ mode, causing all names of  | 
161  | 
types, constants, axioms etc.\ to be automatically qualified by the  | 
|
162  | 
theory name. Changing this to $global$ causes all names to be  | 
|
163  | 
declared as short base names only.  | 
|
164  | 
||
165  | 
The $local$ and $global$ declarations act like switches, affecting  | 
|
166  | 
all following theory sections until changed again explicitly. Also  | 
|
167  | 
note that the final state at the end of the theory will persist. In  | 
|
168  | 
particular, this determines how the names of theorems stored later  | 
|
169  | 
on are handled.  | 
|
| 5369 | 170  | 
|
171  | 
\item[$setup$]\index{*setup!theory} applies a list of ML functions to
 | 
|
172  | 
the theory. The argument should denote a value of type  | 
|
173  | 
  \texttt{(theory -> theory) list}.  Typically, ML packages are
 | 
|
174  | 
initialized in this way.  | 
|
| 1846 | 175  | 
|
| 324 | 176  | 
\item[$ml$] \index{*ML section}
 | 
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
177  | 
consists of \ML\ code, typically for parse and print translation functions.  | 
| 104 | 178  | 
\end{description}
 | 
| 324 | 179  | 
%  | 
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
180  | 
Chapters~\ref{Defining-Logics} and \ref{chap:syntax} explain mixfix
 | 
| 6669 | 181  | 
declarations, translation rules and the \texttt{ML} section in more detail.
 | 
| 145 | 182  | 
|
183  | 
||
| 1905 | 184  | 
\subsection{Definitions}\indexbold{definitions}
 | 
185  | 
||
| 8136 | 186  | 
\textbf{Definitions} are intended to express abbreviations.  The simplest
 | 
| 3108 | 187  | 
form of a definition is $f \equiv t$, where $f$ is a constant.  | 
188  | 
Isabelle also allows a derived forms where the arguments of~$f$ appear  | 
|
189  | 
on the left, abbreviating a string of $\lambda$-abstractions.  | 
|
| 1905 | 190  | 
|
191  | 
Isabelle makes the following checks on definitions:  | 
|
192  | 
\begin{itemize}
 | 
|
| 3108 | 193  | 
\item Arguments (on the left-hand side) must be distinct variables.  | 
| 1905 | 194  | 
\item All variables on the right-hand side must also appear on the left-hand  | 
195  | 
side.  | 
|
| 3108 | 196  | 
\item All type variables on the right-hand side must also appear on  | 
197  | 
  the left-hand side; this prohibits definitions such as {\tt
 | 
|
198  | 
(zero::nat) == length ([]::'a list)}.  | 
|
| 1905 | 199  | 
\item The definition must not be recursive. Most object-logics provide  | 
200  | 
definitional principles that can be used to express recursion safely.  | 
|
201  | 
\end{itemize}
 | 
|
202  | 
These checks are intended to catch the sort of errors that might be made  | 
|
203  | 
accidentally. Misspellings, for instance, might result in additional  | 
|
204  | 
variables appearing on the right-hand side. More elaborate checks could be  | 
|
205  | 
made, but the cost might be overly strict rules on declaration order, etc.  | 
|
206  | 
||
207  | 
||
| 275 | 208  | 
\subsection{*Classes and arities}
 | 
| 324 | 209  | 
\index{classes!context conditions}\index{arities!context conditions}
 | 
| 145 | 210  | 
|
| 286 | 211  | 
In order to guarantee principal types~\cite{nipkow-prehofer},
 | 
| 324 | 212  | 
arity declarations must obey two conditions:  | 
| 145 | 213  | 
\begin{itemize}
 | 
| 3108 | 214  | 
\item There must not be any two declarations $ty :: (\vec{r})c$ and
 | 
215  | 
  $ty :: (\vec{s})c$ with $\vec{r} \neq \vec{s}$.  For example, this
 | 
|
216  | 
excludes the following:  | 
|
| 145 | 217  | 
\begin{ttbox}
 | 
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
218  | 
arities  | 
| 8136 | 219  | 
  foo :: (\{logic{\}}) logic
 | 
220  | 
  foo :: (\{{\}})logic
 | 
|
| 145 | 221  | 
\end{ttbox}
 | 
| 286 | 222  | 
|
| 145 | 223  | 
\item If there are two declarations $ty :: (s@1,\dots,s@n)c$ and $ty ::  | 
224  | 
(s@1',\dots,s@n')c'$ such that $c' < c$ then $s@i' \preceq s@i$ must hold  | 
|
225  | 
for $i=1,\dots,n$. The relationship $\preceq$, defined as  | 
|
226  | 
\[ s' \preceq s \iff \forall c\in s. \exists c'\in s'.~ c'\le c, \]  | 
|
| 3108 | 227  | 
expresses that the set of types represented by $s'$ is a subset of the  | 
228  | 
set of types represented by $s$. Assuming $term \preceq logic$, the  | 
|
229  | 
following is forbidden:  | 
|
| 145 | 230  | 
\begin{ttbox}
 | 
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
231  | 
arities  | 
| 8136 | 232  | 
  foo :: (\{logic{\}})logic
 | 
233  | 
  foo :: (\{{\}})term
 | 
|
| 145 | 234  | 
\end{ttbox}
 | 
| 286 | 235  | 
|
| 145 | 236  | 
\end{itemize}
 | 
237  | 
||
| 104 | 238  | 
|
| 6568 | 239  | 
\section{The theory loader}\label{sec:more-theories}
 | 
240  | 
\index{theories!reading}\index{files!reading}
 | 
|
241  | 
||
242  | 
Isabelle's theory loader manages dependencies of the internal graph of theory  | 
|
243  | 
nodes (the \emph{theory database}) and the external view of the file system.
 | 
|
244  | 
See \S\ref{sec:intro-theories} for its most basic commands, such as
 | 
|
245  | 
\texttt{use_thy}.  There are a few more operations available.
 | 
|
246  | 
||
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
247  | 
\begin{ttbox}
 | 
| 6568 | 248  | 
use_thy_only : string -> unit  | 
| 7136 | 249  | 
update_thy_only : string -> unit  | 
| 6568 | 250  | 
touch_thy : string -> unit  | 
| 6658 | 251  | 
remove_thy : string -> unit  | 
| 8136 | 252  | 
delete_tmpfiles : bool ref \hfill\textbf{initially true}
 | 
| 286 | 253  | 
\end{ttbox}
 | 
254  | 
||
| 324 | 255  | 
\begin{ttdescription}
 | 
| 6569 | 256  | 
\item[\ttindexbold{use_thy_only} "$name$";] is similar to \texttt{use_thy},
 | 
257  | 
  but processes the actual theory file $name$\texttt{.thy} only, ignoring
 | 
|
| 6568 | 258  | 
  $name$\texttt{.ML}.  This might be useful in replaying proof scripts by hand
 | 
259  | 
from the very beginning, starting with the fresh theory.  | 
|
260  | 
||
| 7136 | 261  | 
\item[\ttindexbold{update_thy_only} "$name$";] is similar to
 | 
262  | 
  \texttt{update_thy}, but processes the actual theory file
 | 
|
263  | 
  $name$\texttt{.thy} only, ignoring $name$\texttt{.ML}.
 | 
|
264  | 
||
| 6569 | 265  | 
\item[\ttindexbold{touch_thy} "$name$";] marks theory node $name$ of the
 | 
| 6568 | 266  | 
internal graph as outdated. While the theory remains usable, subsequent  | 
267  | 
  operations such as \texttt{use_thy} may cause a reload.
 | 
|
268  | 
||
| 6658 | 269  | 
\item[\ttindexbold{remove_thy} "$name$";] deletes theory node $name$,
 | 
270  | 
  including \emph{all} of its descendants.  Beware!  This is a quick way to
 | 
|
271  | 
  dispose a large number of theories at once.  Note that {\ML} bindings to
 | 
|
272  | 
theorems etc.\ of removed theories may still persist.  | 
|
273  | 
||
| 6568 | 274  | 
\item[reset \ttindexbold{delete_tmpfiles};] processing theory files usually
 | 
275  | 
  involves temporary {\ML} files to be created.  By default, these are deleted
 | 
|
276  | 
  afterwards.  Resetting the \texttt{delete_tmpfiles} flag inhibits this,
 | 
|
277  | 
leaving the generated code for debugging purposes. The basic location for  | 
|
278  | 
  temporary files is determined by the \texttt{ISABELLE_TMP} environment
 | 
|
| 6571 | 279  | 
variable (which is private to the running Isabelle process and may be  | 
| 6568 | 280  | 
  retrieved by \ttindex{getenv} from {\ML}).
 | 
| 324 | 281  | 
\end{ttdescription}
 | 
| 
138
 
9ba8bff1addc
added chapter "Defining Theories" and made changes for new Readthy functions
 
clasohm 
parents: 
104 
diff
changeset
 | 
282  | 
|
| 6568 | 283  | 
\medskip Theory and {\ML} files are located by skimming through the
 | 
284  | 
directories listed in Isabelle's internal load path, which merely contains the  | 
|
285  | 
current directory ``\texttt{.}'' by default.  The load path may be accessed by
 | 
|
286  | 
the following operations.  | 
|
| 332 | 287  | 
|
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
288  | 
\begin{ttbox}
 | 
| 6568 | 289  | 
show_path: unit -> string list  | 
290  | 
add_path: string -> unit  | 
|
291  | 
del_path: string -> unit  | 
|
292  | 
reset_path: unit -> unit  | 
|
| 7440 | 293  | 
with_path: string -> ('a -> 'b) -> 'a -> 'b
 | 
| 11052 | 294  | 
no_document: ('a -> 'b) -> 'a -> 'b
 | 
| 286 | 295  | 
\end{ttbox}
 | 
296  | 
||
| 324 | 297  | 
\begin{ttdescription}
 | 
| 6568 | 298  | 
\item[\ttindexbold{show_path}();] displays the load path components in
 | 
| 6571 | 299  | 
canonical string representation (which is always according to Unix rules).  | 
| 6568 | 300  | 
|
| 6569 | 301  | 
\item[\ttindexbold{add_path} "$dir$";] adds component $dir$ to the beginning
 | 
302  | 
of the load path.  | 
|
| 6568 | 303  | 
|
| 6569 | 304  | 
\item[\ttindexbold{del_path} "$dir$";] removes any occurrences of component
 | 
| 6568 | 305  | 
$dir$ from the load path.  | 
306  | 
||
307  | 
\item[\ttindexbold{reset_path}();] resets the load path to ``\texttt{.}''
 | 
|
308  | 
(current directory) only.  | 
|
| 7440 | 309  | 
|
310  | 
\item[\ttindexbold{with_path} "$dir$" $f$ $x$;] temporarily adds component
 | 
|
| 11052 | 311  | 
$dir$ to the beginning of the load path while executing $(f~x)$.  | 
312  | 
||
313  | 
\item[\ttindexbold{no_document} $f$ $x$;] temporarily disables {\LaTeX}
 | 
|
314  | 
document generation while executing $(f~x)$.  | 
|
| 7440 | 315  | 
|
| 324 | 316  | 
\end{ttdescription}
 | 
| 286 | 317  | 
|
| 7440 | 318  | 
Furthermore, in operations referring indirectly to some file (e.g.\  | 
319  | 
\texttt{use_dir}) the argument may be prefixed by a directory that will be
 | 
|
320  | 
temporarily appended to the load path, too.  | 
|
| 104 | 321  | 
|
322  | 
||
| 6669 | 323  | 
\section{Locales}
 | 
324  | 
\label{Locales}
 | 
|
325  | 
||
326  | 
Locales \cite{kammueller-locales} are a concept of local proof contexts.  They
 | 
|
327  | 
are introduced as named syntactic objects within theories and can be  | 
|
328  | 
opened in any descendant theory.  | 
|
329  | 
||
330  | 
\subsection{Declaring Locales}
 | 
|
331  | 
||
332  | 
A locale is declared in a theory section that starts with the  | 
|
333  | 
keyword \texttt{locale}.  It consists typically of three parts, the
 | 
|
334  | 
\texttt{fixes} part, the \texttt{assumes} part, and the \texttt{defines} part.
 | 
|
335  | 
Appendix \ref{app:TheorySyntax} presents the full syntax.
 | 
|
336  | 
||
337  | 
\subsubsection{Parts of Locales}
 | 
|
338  | 
||
339  | 
The subsection introduced by the keyword \texttt{fixes} declares the locale
 | 
|
340  | 
constants in a way that closely resembles a global \texttt{consts}
 | 
|
341  | 
declaration. In particular, there may be an optional pretty printing syntax  | 
|
342  | 
for the locale constants.  | 
|
343  | 
||
344  | 
The subsequent \texttt{assumes} part specifies the locale rules.  They are
 | 
|
345  | 
defined like \texttt{rules}: by an identifier followed by the rule
 | 
|
346  | 
given as a string. Locale rules admit the statement of local assumptions  | 
|
347  | 
about the locale constants.  The \texttt{assumes} part is optional.  Non-fixed
 | 
|
348  | 
variables in locale rules are automatically bound by the universal quantifier  | 
|
349  | 
\texttt{!!} of the meta-logic.
 | 
|
350  | 
||
351  | 
Finally, the \texttt{defines} part introduces the definitions that are
 | 
|
352  | 
available in the locale.  Locale constants declared in the \texttt{fixes}
 | 
|
353  | 
section are defined using the meta-equality \texttt{==}.  If the
 | 
|
354  | 
locale constant is a functiond then its definition can (as usual) have  | 
|
355  | 
variables on the left-hand side acting as formal parameters; they are  | 
|
356  | 
considered as schematic variables and are automatically generalized by  | 
|
357  | 
universal quantification of the meta-logic. The right hand side of a  | 
|
358  | 
definition must not contain variables that are not already on the left hand  | 
|
359  | 
side. In so far locale definitions behave like theory level definitions.  | 
|
360  | 
However, the locale concept realizes \emph{dependent definitions}: any variable
 | 
|
361  | 
that is fixed as a locale constant can occur on the right hand side of  | 
|
362  | 
definitions. For an illustration of these dependent definitions see the  | 
|
363  | 
occurrence of the locale constant \texttt{G} on the right hand side of the
 | 
|
364  | 
definitions of the locale \texttt{group} below.  Naturally, definitions can
 | 
|
365  | 
already use the syntax of the locale constants in the \texttt{fixes}
 | 
|
366  | 
subsection.  The \texttt{defines} part is, as the \texttt{assumes} part,
 | 
|
367  | 
optional.  | 
|
368  | 
||
369  | 
\subsubsection{Example for Definition}
 | 
|
370  | 
The concrete syntax of locale definitions is demonstrated by example below.  | 
|
371  | 
||
372  | 
Locale \texttt{group} assumes the definition of groups in a theory
 | 
|
373  | 
file\footnote{This and other examples are from \texttt{HOL/ex}.}.  A locale
 | 
|
374  | 
defining a convenient proof environment for group related proofs may be  | 
|
375  | 
added to the theory as follows:  | 
|
376  | 
\begin{ttbox}
 | 
|
377  | 
locale group =  | 
|
378  | 
fixes  | 
|
379  | 
G :: "'a grouptype"  | 
|
380  | 
e :: "'a"  | 
|
381  | 
binop :: "'a => 'a => 'a" (infixr "#" 80)  | 
|
382  | 
      inv       :: "'a => 'a"              ("i(_)" [90] 91)
 | 
|
383  | 
assumes  | 
|
384  | 
Group_G "G: Group"  | 
|
385  | 
defines  | 
|
386  | 
e_def "e == unit G"  | 
|
387  | 
binop_def "x # y == bin_op G x y"  | 
|
388  | 
inv_def "i(x) == inverse G x"  | 
|
389  | 
\end{ttbox}
 | 
|
390  | 
||
391  | 
\subsubsection{Polymorphism}
 | 
|
392  | 
||
393  | 
In contrast to polymorphic definitions in theories, the use of the  | 
|
394  | 
same type variable for the declaration of different locale constants in the  | 
|
395  | 
fixes part means \emph{the same} type.  In other words, the scope of the
 | 
|
396  | 
polymorphic variables is extended over all constant declarations of a locale.  | 
|
397  | 
In the above example \texttt{'a} refers to the same type which is fixed inside
 | 
|
398  | 
the locale.  In an exported theorem (see \S\ref{sec:locale-export}) the
 | 
|
399  | 
constructors of locale \texttt{group} are polymorphic, yet only simultaneously
 | 
|
400  | 
instantiatable.  | 
|
401  | 
||
402  | 
\subsubsection{Nested Locales}
 | 
|
403  | 
||
404  | 
A locale can be defined as the extension of a previously defined  | 
|
405  | 
locale. This operation of extension is optional and is syntactically  | 
|
406  | 
expressed as  | 
|
407  | 
\begin{ttbox}
 | 
|
408  | 
locale foo = bar + ...  | 
|
409  | 
\end{ttbox}
 | 
|
410  | 
The locale \texttt{foo} builds on the constants and syntax of the locale {\tt
 | 
|
411  | 
bar}.  That is, all contents of the locale \texttt{bar} can be used in
 | 
|
412  | 
definitions and rules of the corresponding parts of the locale {\tt
 | 
|
413  | 
foo}.  Although locale \texttt{foo} assumes the \texttt{fixes} part of \texttt{bar} it
 | 
|
414  | 
does not automatically subsume its rules and definitions. Normally, one  | 
|
415  | 
expects to use locale \texttt{foo} only if locale \texttt{bar} is already
 | 
|
416  | 
active. These aspects of use and activation of locales are considered in the  | 
|
417  | 
subsequent section.  | 
|
418  | 
||
419  | 
||
420  | 
\subsection{Locale Scope}
 | 
|
421  | 
||
422  | 
Locales are by default inactive, but they can be invoked. The list of  | 
|
423  | 
currently active locales is called \emph{scope}.  The process of activating
 | 
|
424  | 
them is called \emph{opening}; the reverse is \emph{closing}.
 | 
|
425  | 
||
426  | 
\subsubsection{Scope}
 | 
|
427  | 
The locale scope is part of each theory. It is a dynamic stack containing  | 
|
428  | 
all active locales at a certain point in an interactive session.  | 
|
429  | 
The scope lives until all locales are explicitly closed. At one time there  | 
|
430  | 
can be more than one locale open. The contents of these various active  | 
|
431  | 
locales are all visible in the scope. In case of nested locales for example,  | 
|
432  | 
the nesting is actually reflected to the scope, which contains the nested  | 
|
433  | 
locales as layers. To check the state of the scope during a development the  | 
|
434  | 
function \texttt{Print\_scope} may be used.  It displays the names of all open
 | 
|
435  | 
locales on the scope.  The function \texttt{print\_locales} applied to a theory
 | 
|
436  | 
displays all locales contained in that theory and in addition also the  | 
|
437  | 
current scope.  | 
|
438  | 
||
439  | 
The scope is manipulated by the commands for opening and closing of locales.  | 
|
440  | 
||
441  | 
\subsubsection{Opening}
 | 
|
442  | 
Locales can be \emph{opened} at any point during a session where
 | 
|
443  | 
we want to prove theorems concerning the locale. Opening a locale means  | 
|
444  | 
making its contents visible by pushing it onto the scope of the current  | 
|
445  | 
theory. Inside a scope of opened locales, theorems can use all definitions and  | 
|
446  | 
rules contained in the locales on the scope. The rules and definitions may  | 
|
447  | 
be accessed individually using the function \ttindex{thm}.  This function is
 | 
|
448  | 
applied to the names assigned to locale rules and definitions as  | 
|
449  | 
strings.  The opening command is called \texttt{Open\_locale} and takes the 
 | 
|
450  | 
name of the locale to be opened as its argument.  | 
|
451  | 
||
452  | 
If one opens a locale \texttt{foo} that is defined by extension from locale
 | 
|
453  | 
\texttt{bar}, the function \texttt{Open\_locale} checks if locale \texttt{bar}
 | 
|
454  | 
is open.  If so, then it just opens \texttt{foo}, if not, then it prints a
 | 
|
455  | 
message and opens \texttt{bar} before opening \texttt{foo}.  Naturally, this
 | 
|
456  | 
carries on, if \texttt{bar} is again an extension.
 | 
|
457  | 
||
458  | 
\subsubsection{Closing}
 | 
|
459  | 
||
460  | 
\emph{Closing} means to cancel the last opened locale, pushing it out of the
 | 
|
461  | 
scope. Theorems proved during the life cycle of this locale will be disabled,  | 
|
462  | 
unless they have been explicitly exported, as described below. However, when  | 
|
463  | 
the same locale is opened again these theorems may be used again as well,  | 
|
464  | 
provided that they were saved as theorems in the first place, using  | 
|
465  | 
\texttt{qed} or ML assignment.  The command \texttt{Close\_locale} takes a
 | 
|
466  | 
locale name as a string and checks if this locale is actually the topmost  | 
|
467  | 
locale on the scope. If this is the case, it removes this locale, otherwise  | 
|
468  | 
it prints a warning message and does not change the scope.  | 
|
469  | 
||
470  | 
\subsubsection{Export of Theorems}
 | 
|
471  | 
\label{sec:locale-export}
 | 
|
472  | 
||
473  | 
Export of theorems transports theorems out of the scope of locales. Locale  | 
|
474  | 
rules that have been used in the proof of an exported theorem inside the  | 
|
475  | 
locale are carried by the exported form of the theorem as its individual  | 
|
476  | 
meta-assumptions. The locale constants are universally quantified variables  | 
|
477  | 
in these theorems, hence such theorems can be instantiated individually.  | 
|
478  | 
Definitions become unfolded; locale constants that were merely used for  | 
|
479  | 
definitions vanish. Logically, exporting corresponds to a combined  | 
|
480  | 
application of introduction rules for implication and universal  | 
|
481  | 
quantification. Exporting forms a kind of normalization of theorems in a  | 
|
482  | 
locale scope.  | 
|
483  | 
||
484  | 
According to the possibility of nested locales there are two different forms  | 
|
485  | 
of export.  The first one is realized by the function \texttt{export} that
 | 
|
486  | 
exports theorems through all layers of opened locales of the scope. Hence,  | 
|
487  | 
the application of export to a theorem yields a theorem of the global level,  | 
|
488  | 
that is, the current theory context without any local assumptions or  | 
|
489  | 
definitions.  | 
|
490  | 
||
491  | 
When locales are nested we might want to export a theorem, not to the global  | 
|
492  | 
level of the current theory but just to the previous level. The other export  | 
|
493  | 
function, \texttt{Export}, transports theorems one level up in the scope; the
 | 
|
494  | 
theorem still uses locale constants, definitions and rules of the locales  | 
|
495  | 
underneath.  | 
|
496  | 
||
497  | 
\subsection{Functions for Locales}
 | 
|
498  | 
\label{Syntax}
 | 
|
499  | 
\index{locales!functions}
 | 
|
500  | 
||
501  | 
Here is a quick reference list of locale functions.  | 
|
502  | 
\begin{ttbox}
 | 
|
503  | 
Open_locale : xstring -> unit  | 
|
504  | 
Close_locale : xstring -> unit  | 
|
505  | 
export : thm -> thm  | 
|
506  | 
Export : thm -> thm  | 
|
507  | 
thm : xstring -> thm  | 
|
508  | 
Print_scope : unit -> unit  | 
|
509  | 
print_locales: theory -> unit  | 
|
510  | 
\end{ttbox}
 | 
|
511  | 
\begin{ttdescription}
 | 
|
512  | 
\item[\ttindexbold{Open_locale} $xstring$] 
 | 
|
513  | 
    opens the locale {\it xstring}, adding it to the scope of the theory of the
 | 
|
514  | 
current context. If the opened locale is built by extension, the ancestors  | 
|
515  | 
are opened automatically.  | 
|
516  | 
||
517  | 
\item[\ttindexbold{Close_locale} $xstring$] eliminates the locale {\it
 | 
|
518  | 
xstring} from the scope if it is the topmost item on it, otherwise it does  | 
|
519  | 
not change the scope and produces a warning.  | 
|
520  | 
||
521  | 
\item[\ttindexbold{export} $thm$] locale definitions become expanded in {\it
 | 
|
522  | 
    thm} and locale rules that were used in the proof of {\it thm} become part
 | 
|
523  | 
of its individual assumptions. This normalization happens with respect to  | 
|
524  | 
  \emph{all open locales} on the scope.
 | 
|
525  | 
||
526  | 
\item[\ttindexbold{Export} $thm$] works like \texttt{export} but normalizes
 | 
|
527  | 
theorems only up to the previous level of locales on the scope.  | 
|
528  | 
||
529  | 
\item[\ttindexbold{thm} $xstring$] applied to the name of a locale definition
 | 
|
530  | 
or rule it returns the definition as a theorem.  | 
|
531  | 
||
532  | 
\item[\ttindexbold{Print_scope}()] prints the names of the locales in the
 | 
|
533  | 
current scope of the current theory context.  | 
|
534  | 
||
535  | 
\item[\ttindexbold{print_locale} $theory$] prints all locales that are
 | 
|
536  | 
  contained in {\it theory} directly or indirectly.  It also displays the
 | 
|
537  | 
  current scope similar to \texttt{Print\_scope}.
 | 
|
538  | 
\end{ttdescription}
 | 
|
539  | 
||
540  | 
||
| 
866
 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 
clasohm 
parents: 
864 
diff
changeset
 | 
541  | 
\section{Basic operations on theories}\label{BasicOperationsOnTheories}
 | 
| 4384 | 542  | 
|
543  | 
\subsection{*Theory inclusion}
 | 
|
544  | 
\begin{ttbox}
 | 
|
545  | 
subthy : theory * theory -> bool  | 
|
546  | 
eq_thy : theory * theory -> bool  | 
|
547  | 
transfer : theory -> thm -> thm  | 
|
548  | 
transfer_sg : Sign.sg -> thm -> thm  | 
|
549  | 
\end{ttbox}
 | 
|
550  | 
||
551  | 
Inclusion and equality of theories is determined by unique  | 
|
552  | 
identification stamps that are created when declaring new components.  | 
|
553  | 
Theorems contain a reference to the theory (actually to its signature)  | 
|
554  | 
they have been derived in. Transferring theorems to super theories  | 
|
555  | 
has no logical significance, but may affect some operations in subtle  | 
|
556  | 
ways (e.g.\ implicit merges of signatures when applying rules, or  | 
|
557  | 
pretty printing of theorems).  | 
|
558  | 
||
559  | 
\begin{ttdescription}
 | 
|
560  | 
||
561  | 
\item[\ttindexbold{subthy} ($thy@1$, $thy@2$)] determines if $thy@1$
 | 
|
562  | 
is included in $thy@2$ wrt.\ identification stamps.  | 
|
563  | 
||
564  | 
\item[\ttindexbold{eq_thy} ($thy@1$, $thy@2$)] determines if $thy@1$
 | 
|
565  | 
is exactly the same as $thy@2$.  | 
|
566  | 
||
567  | 
\item[\ttindexbold{transfer} $thy$ $thm$] transfers theorem $thm$ to
 | 
|
568  | 
theory $thy$, provided the latter includes the theory of $thm$.  | 
|
569  | 
||
570  | 
\item[\ttindexbold{transfer_sg} $sign$ $thm$] is similar to
 | 
|
571  | 
  \texttt{transfer}, but identifies the super theory via its
 | 
|
572  | 
signature.  | 
|
573  | 
||
574  | 
\end{ttdescription}
 | 
|
575  | 
||
576  | 
||
| 6571 | 577  | 
\subsection{*Primitive theories}
 | 
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
578  | 
\begin{ttbox}
 | 
| 4317 | 579  | 
ProtoPure.thy : theory  | 
| 3108 | 580  | 
Pure.thy : theory  | 
581  | 
CPure.thy : theory  | 
|
| 286 | 582  | 
\end{ttbox}
 | 
| 3108 | 583  | 
\begin{description}
 | 
| 4317 | 584  | 
\item[\ttindexbold{ProtoPure.thy}, \ttindexbold{Pure.thy},
 | 
585  | 
  \ttindexbold{CPure.thy}] contain the syntax and signature of the
 | 
|
586  | 
meta-logic. There are basically no axioms: meta-level inferences  | 
|
587  | 
  are carried out by \ML\ functions.  \texttt{Pure} and \texttt{CPure}
 | 
|
588  | 
just differ in their concrete syntax of prefix function application:  | 
|
589  | 
  $t(u@1, \ldots, u@n)$ in \texttt{Pure} vs.\ $t\,u@1,\ldots\,u@n$ in
 | 
|
590  | 
  \texttt{CPure}.  \texttt{ProtoPure} is their common parent,
 | 
|
591  | 
containing no syntax for printing prefix applications at all!  | 
|
| 6571 | 592  | 
|
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
593  | 
%% FIXME  | 
| 478 | 594  | 
%\item [\ttindexbold{extend_theory} $thy$ {\tt"}$T${\tt"} $\cdots$] extends
 | 
595  | 
% the theory $thy$ with new types, constants, etc. $T$ identifies the theory  | 
|
596  | 
% internally. When a theory is redeclared, say to change an incorrect axiom,  | 
|
597  | 
% bindings to the old axiom may persist. Isabelle ensures that the old and  | 
|
598  | 
% new theories are not involved in the same proof. Attempting to combine  | 
|
599  | 
% different theories having the same name $T$ yields the fatal error  | 
|
600  | 
%extend_theory : theory -> string -> \(\cdots\) -> theory  | 
|
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
601  | 
%\begin{ttbox}
 | 
| 
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
602  | 
%Attempt to merge different versions of theory: \(T\)  | 
| 
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
603  | 
%\end{ttbox}
 | 
| 3108 | 604  | 
\end{description}
 | 
| 286 | 605  | 
|
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
606  | 
%% FIXME  | 
| 275 | 607  | 
%\item [\ttindexbold{extend_theory} $thy$ {\tt"}$T${\tt"}
 | 
608  | 
% ($classes$, $default$, $types$, $arities$, $consts$, $sextopt$) $rules$]  | 
|
609  | 
%\hfill\break %%% include if line is just too short  | 
|
| 286 | 610  | 
%is the \ML{} equivalent of the following theory definition:
 | 
| 275 | 611  | 
%\begin{ttbox}
 | 
612  | 
%\(T\) = \(thy\) +  | 
|
613  | 
%classes \(c\) < \(c@1\),\(\dots\),\(c@m\)  | 
|
614  | 
% \dots  | 
|
615  | 
%default {\(d@1,\dots,d@r\)}
 | 
|
616  | 
%types \(tycon@1\),\dots,\(tycon@i\) \(n\)  | 
|
617  | 
% \dots  | 
|
618  | 
%arities \(tycon@1'\),\dots,\(tycon@j'\) :: (\(s@1\),\dots,\(s@n\))\(c\)  | 
|
619  | 
% \dots  | 
|
620  | 
%consts \(b@1\),\dots,\(b@k\) :: \(\tau\)  | 
|
621  | 
% \dots  | 
|
622  | 
%rules \(name\) \(rule\)  | 
|
623  | 
% \dots  | 
|
624  | 
%end  | 
|
625  | 
%\end{ttbox}
 | 
|
626  | 
%where  | 
|
627  | 
%\begin{tabular}[t]{l@{~=~}l}
 | 
|
628  | 
%$classes$ & \tt[("$c$",["$c@1$",\dots,"$c@m$"]),\dots] \\
 | 
|
629  | 
%$default$ & \tt["$d@1$",\dots,"$d@r$"]\\  | 
|
630  | 
%$types$ & \tt[([$tycon@1$,\dots,$tycon@i$], $n$),\dots] \\  | 
|
631  | 
%$arities$ & \tt[([$tycon'@1$,\dots,$tycon'@j$], ([$s@1$,\dots,$s@n$],$c$)),\dots]  | 
|
632  | 
%\\  | 
|
633  | 
%$consts$ & \tt[([$b@1$,\dots,$b@k$],$\tau$),\dots] \\  | 
|
634  | 
%$rules$   & \tt[("$name$",$rule$),\dots]
 | 
|
635  | 
%\end{tabular}
 | 
|
| 104 | 636  | 
|
637  | 
||
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
638  | 
\subsection{Inspecting a theory}\label{sec:inspct-thy}
 | 
| 104 | 639  | 
\index{theories!inspecting|bold}
 | 
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
640  | 
\begin{ttbox}
 | 
| 4317 | 641  | 
print_syntax : theory -> unit  | 
642  | 
print_theory : theory -> unit  | 
|
643  | 
parents_of : theory -> theory list  | 
|
644  | 
ancestors_of : theory -> theory list  | 
|
645  | 
sign_of : theory -> Sign.sg  | 
|
646  | 
Sign.stamp_names_of : Sign.sg -> string list  | 
|
| 104 | 647  | 
\end{ttbox}
 | 
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
648  | 
These provide means of viewing a theory's components.  | 
| 324 | 649  | 
\begin{ttdescription}
 | 
| 3108 | 650  | 
\item[\ttindexbold{print_syntax} $thy$] prints the syntax of $thy$
 | 
651  | 
(grammar, macros, translation functions etc., see  | 
|
652  | 
  page~\pageref{pg:print_syn} for more details).
 | 
|
653  | 
||
654  | 
\item[\ttindexbold{print_theory} $thy$] prints the logical parts of
 | 
|
655  | 
$thy$, excluding the syntax.  | 
|
| 4317 | 656  | 
|
657  | 
\item[\ttindexbold{parents_of} $thy$] returns the direct ancestors
 | 
|
658  | 
of~$thy$.  | 
|
659  | 
||
660  | 
\item[\ttindexbold{ancestors_of} $thy$] returns all ancestors of~$thy$
 | 
|
661  | 
(not including $thy$ itself).  | 
|
662  | 
||
663  | 
\item[\ttindexbold{sign_of} $thy$] returns the signature associated
 | 
|
664  | 
  with~$thy$.  It is useful with functions like {\tt
 | 
|
665  | 
read_instantiate_sg}, which take a signature as an argument.  | 
|
666  | 
||
667  | 
\item[\ttindexbold{Sign.stamp_names_of} $sg$]\index{signatures}
 | 
|
668  | 
  returns the names of the identification \rmindex{stamps} of ax
 | 
|
669  | 
signature. These coincide with the names of its full ancestry  | 
|
670  | 
including that of $sg$ itself.  | 
|
| 104 | 671  | 
|
| 324 | 672  | 
\end{ttdescription}
 | 
| 104 | 673  | 
|
| 1369 | 674  | 
|
| 104 | 675  | 
\section{Terms}
 | 
676  | 
\index{terms|bold}
 | 
|
| 324 | 677  | 
Terms belong to the \ML\ type \mltydx{term}, which is a concrete datatype
 | 
| 3108 | 678  | 
with six constructors:  | 
| 104 | 679  | 
\begin{ttbox}
 | 
680  | 
type indexname = string * int;  | 
|
681  | 
infix 9 $;  | 
|
682  | 
datatype term = Const of string * typ  | 
|
683  | 
| Free of string * typ  | 
|
684  | 
| Var of indexname * typ  | 
|
685  | 
| Bound of int  | 
|
686  | 
| Abs of string * typ * term  | 
|
687  | 
| op $ of term * term;  | 
|
688  | 
\end{ttbox}
 | 
|
| 324 | 689  | 
\begin{ttdescription}
 | 
| 4317 | 690  | 
\item[\ttindexbold{Const} ($a$, $T$)] \index{constants|bold}
 | 
| 8136 | 691  | 
  is the \textbf{constant} with name~$a$ and type~$T$.  Constants include
 | 
| 286 | 692  | 
connectives like $\land$ and $\forall$ as well as constants like~0  | 
693  | 
and~$Suc$. Other constants may be required to define a logic's concrete  | 
|
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
694  | 
syntax.  | 
| 104 | 695  | 
|
| 4317 | 696  | 
\item[\ttindexbold{Free} ($a$, $T$)] \index{variables!free|bold}
 | 
| 8136 | 697  | 
  is the \textbf{free variable} with name~$a$ and type~$T$.
 | 
| 104 | 698  | 
|
| 4317 | 699  | 
\item[\ttindexbold{Var} ($v$, $T$)] \index{unknowns|bold}
 | 
| 8136 | 700  | 
  is the \textbf{scheme variable} with indexname~$v$ and type~$T$.  An
 | 
| 324 | 701  | 
  \mltydx{indexname} is a string paired with a non-negative index, or
 | 
702  | 
subscript; a term's scheme variables can be systematically renamed by  | 
|
703  | 
incrementing their subscripts. Scheme variables are essentially free  | 
|
704  | 
variables, but may be instantiated during unification.  | 
|
| 104 | 705  | 
|
| 324 | 706  | 
\item[\ttindexbold{Bound} $i$] \index{variables!bound|bold}
 | 
| 8136 | 707  | 
  is the \textbf{bound variable} with de Bruijn index~$i$, which counts the
 | 
| 324 | 708  | 
number of lambdas, starting from zero, between a variable's occurrence  | 
709  | 
and its binding. The representation prevents capture of variables. For  | 
|
710  | 
  more information see de Bruijn \cite{debruijn72} or
 | 
|
| 6592 | 711  | 
  Paulson~\cite[page~376]{paulson-ml2}.
 | 
| 104 | 712  | 
|
| 4317 | 713  | 
\item[\ttindexbold{Abs} ($a$, $T$, $u$)]
 | 
| 324 | 714  | 
  \index{lambda abs@$\lambda$-abstractions|bold}
 | 
| 8136 | 715  | 
  is the $\lambda$-\textbf{abstraction} with body~$u$, and whose bound
 | 
| 324 | 716  | 
variable has name~$a$ and type~$T$. The name is used only for parsing  | 
717  | 
and printing; it has no logical significance.  | 
|
| 104 | 718  | 
|
| 324 | 719  | 
\item[$t$ \$ $u$] \index{$@{\tt\$}|bold} \index{function applications|bold}
 | 
| 8136 | 720  | 
is the \textbf{application} of~$t$ to~$u$.
 | 
| 324 | 721  | 
\end{ttdescription}
 | 
| 9695 | 722  | 
Application is written as an infix operator to aid readability. Here is an  | 
723  | 
\ML\ pattern to recognize FOL formulae of the form~$A\imp B$, binding the  | 
|
724  | 
subformulae to~$A$ and~$B$:  | 
|
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
725  | 
\begin{ttbox}
 | 
| 104 | 726  | 
Const("Trueprop",_) $ (Const("op -->",_) $ A $ B)
 | 
727  | 
\end{ttbox}
 | 
|
728  | 
||
729  | 
||
| 4317 | 730  | 
\section{*Variable binding}
 | 
| 286 | 731  | 
\begin{ttbox}
 | 
732  | 
loose_bnos : term -> int list  | 
|
733  | 
incr_boundvars : int -> term -> term  | 
|
734  | 
abstract_over : term*term -> term  | 
|
735  | 
variant_abs : string * typ * term -> string * term  | 
|
| 8136 | 736  | 
aconv          : term * term -> bool\hfill\textbf{infix}
 | 
| 286 | 737  | 
\end{ttbox}
 | 
738  | 
These functions are all concerned with the de Bruijn representation of  | 
|
739  | 
bound variables.  | 
|
| 324 | 740  | 
\begin{ttdescription}
 | 
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
741  | 
\item[\ttindexbold{loose_bnos} $t$]
 | 
| 286 | 742  | 
returns the list of all dangling bound variable references. In  | 
| 6669 | 743  | 
  particular, \texttt{Bound~0} is loose unless it is enclosed in an
 | 
744  | 
  abstraction.  Similarly \texttt{Bound~1} is loose unless it is enclosed in
 | 
|
| 286 | 745  | 
at least two abstractions; if enclosed in just one, the list will contain  | 
746  | 
the number 0. A well-formed term does not contain any loose variables.  | 
|
747  | 
||
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
748  | 
\item[\ttindexbold{incr_boundvars} $j$]
 | 
| 332 | 749  | 
increases a term's dangling bound variables by the offset~$j$. This is  | 
| 286 | 750  | 
required when moving a subterm into a context where it is enclosed by a  | 
751  | 
different number of abstractions. Bound variables with a matching  | 
|
752  | 
abstraction are unaffected.  | 
|
753  | 
||
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
754  | 
\item[\ttindexbold{abstract_over} $(v,t)$]
 | 
| 286 | 755  | 
forms the abstraction of~$t$ over~$v$, which may be any well-formed term.  | 
| 6669 | 756  | 
  It replaces every occurrence of \(v\) by a \texttt{Bound} variable with the
 | 
| 286 | 757  | 
correct index.  | 
758  | 
||
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
759  | 
\item[\ttindexbold{variant_abs} $(a,T,u)$]
 | 
| 286 | 760  | 
substitutes into $u$, which should be the body of an abstraction.  | 
761  | 
It replaces each occurrence of the outermost bound variable by a free  | 
|
762  | 
variable. The free variable has type~$T$ and its name is a variant  | 
|
| 332 | 763  | 
of~$a$ chosen to be distinct from all constants and from all variables  | 
| 286 | 764  | 
free in~$u$.  | 
765  | 
||
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
766  | 
\item[$t$ \ttindexbold{aconv} $u$]
 | 
| 286 | 767  | 
tests whether terms~$t$ and~$u$ are \(\alpha\)-convertible: identical up  | 
768  | 
to renaming of bound variables.  | 
|
769  | 
\begin{itemize}
 | 
|
770  | 
\item  | 
|
| 6669 | 771  | 
    Two constants, \texttt{Free}s, or \texttt{Var}s are \(\alpha\)-convertible
 | 
| 286 | 772  | 
if their names and types are equal.  | 
773  | 
(Variables having the same name but different types are thus distinct.  | 
|
774  | 
This confusing situation should be avoided!)  | 
|
775  | 
\item  | 
|
776  | 
Two bound variables are \(\alpha\)-convertible  | 
|
777  | 
if they have the same number.  | 
|
778  | 
\item  | 
|
779  | 
Two abstractions are \(\alpha\)-convertible  | 
|
780  | 
if their bodies are, and their bound variables have the same type.  | 
|
781  | 
\item  | 
|
782  | 
Two applications are \(\alpha\)-convertible  | 
|
783  | 
if the corresponding subterms are.  | 
|
784  | 
\end{itemize}
 | 
|
785  | 
||
| 324 | 786  | 
\end{ttdescription}
 | 
| 286 | 787  | 
|
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
788  | 
\section{Certified terms}\index{terms!certified|bold}\index{signatures}
 | 
| 8136 | 789  | 
A term $t$ can be \textbf{certified} under a signature to ensure that every type
 | 
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
790  | 
in~$t$ is well-formed and every constant in~$t$ is a type instance of a  | 
| 
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
791  | 
constant declared in the signature. The term must be well-typed and its use  | 
| 6669 | 792  | 
of bound variables must be well-formed.  Meta-rules such as \texttt{forall_elim}
 | 
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
793  | 
take certified terms as arguments.  | 
| 104 | 794  | 
|
| 324 | 795  | 
Certified terms belong to the abstract type \mltydx{cterm}.
 | 
| 104 | 796  | 
Elements of the type can only be created through the certification process.  | 
797  | 
In case of error, Isabelle raises exception~\ttindex{TERM}\@.
 | 
|
798  | 
||
799  | 
\subsection{Printing terms}
 | 
|
| 324 | 800  | 
\index{terms!printing of}
 | 
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
801  | 
\begin{ttbox}
 | 
| 275 | 802  | 
string_of_cterm : cterm -> string  | 
| 104 | 803  | 
Sign.string_of_term : Sign.sg -> term -> string  | 
804  | 
\end{ttbox}
 | 
|
| 324 | 805  | 
\begin{ttdescription}
 | 
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
806  | 
\item[\ttindexbold{string_of_cterm} $ct$]
 | 
| 104 | 807  | 
displays $ct$ as a string.  | 
808  | 
||
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
809  | 
\item[\ttindexbold{Sign.string_of_term} $sign$ $t$]
 | 
| 104 | 810  | 
displays $t$ as a string, using the syntax of~$sign$.  | 
| 324 | 811  | 
\end{ttdescription}
 | 
| 104 | 812  | 
|
813  | 
\subsection{Making and inspecting certified terms}
 | 
|
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
814  | 
\begin{ttbox}
 | 
| 8136 | 815  | 
cterm_of : Sign.sg -> term -> cterm  | 
816  | 
read_cterm : Sign.sg -> string * typ -> cterm  | 
|
817  | 
cert_axm : Sign.sg -> string * term -> string * term  | 
|
818  | 
read_axm : Sign.sg -> string * string -> string * term  | 
|
819  | 
rep_cterm      : cterm -> \{T:typ, t:term, sign:Sign.sg, maxidx:int\}
 | 
|
| 4543 | 820  | 
Sign.certify_term : Sign.sg -> term -> term * typ * int  | 
| 104 | 821  | 
\end{ttbox}
 | 
| 324 | 822  | 
\begin{ttdescription}
 | 
| 4543 | 823  | 
|
824  | 
\item[\ttindexbold{cterm_of} $sign$ $t$] \index{signatures} certifies
 | 
|
825  | 
$t$ with respect to signature~$sign$.  | 
|
826  | 
||
827  | 
\item[\ttindexbold{read_cterm} $sign$ ($s$, $T$)] reads the string~$s$
 | 
|
828  | 
using the syntax of~$sign$, creating a certified term. The term is  | 
|
829  | 
checked to have type~$T$; this type also tells the parser what kind  | 
|
830  | 
of phrase to parse.  | 
|
831  | 
||
832  | 
\item[\ttindexbold{cert_axm} $sign$ ($name$, $t$)] certifies $t$ with
 | 
|
833  | 
respect to $sign$ as a meta-proposition and converts all exceptions  | 
|
834  | 
to an error, including the final message  | 
|
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
835  | 
\begin{ttbox}
 | 
| 
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
836  | 
The error(s) above occurred in axiom "\(name\)"  | 
| 
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
837  | 
\end{ttbox}
 | 
| 
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
838  | 
|
| 4543 | 839  | 
\item[\ttindexbold{read_axm} $sign$ ($name$, $s$)] similar to {\tt
 | 
840  | 
cert_axm}, but first reads the string $s$ using the syntax of  | 
|
841  | 
$sign$.  | 
|
842  | 
||
843  | 
\item[\ttindexbold{rep_cterm} $ct$] decomposes $ct$ as a record
 | 
|
844  | 
containing its type, the term itself, its signature, and the maximum  | 
|
845  | 
subscript of its unknowns. The type and maximum subscript are  | 
|
846  | 
computed during certification.  | 
|
847  | 
||
848  | 
\item[\ttindexbold{Sign.certify_term}] is a more primitive version of
 | 
|
849  | 
  \texttt{cterm_of}, returning the internal representation instead of
 | 
|
850  | 
  an abstract \texttt{cterm}.
 | 
|
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
851  | 
|
| 324 | 852  | 
\end{ttdescription}
 | 
| 104 | 853  | 
|
854  | 
||
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
855  | 
\section{Types}\index{types|bold}
 | 
| 
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
856  | 
Types belong to the \ML\ type \mltydx{typ}, which is a concrete datatype with
 | 
| 
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
857  | 
three constructor functions. These correspond to type constructors, free  | 
| 
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
858  | 
type variables and schematic type variables. Types are classified by sorts,  | 
| 
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
859  | 
which are lists of classes (representing an intersection). A class is  | 
| 
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
860  | 
represented by a string.  | 
| 104 | 861  | 
\begin{ttbox}
 | 
862  | 
type class = string;  | 
|
863  | 
type sort = class list;  | 
|
864  | 
||
865  | 
datatype typ = Type of string * typ list  | 
|
866  | 
| TFree of string * sort  | 
|
867  | 
| TVar of indexname * sort;  | 
|
868  | 
||
869  | 
infixr 5 -->;  | 
|
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
870  | 
fun S --> T = Type ("fun", [S, T]);
 | 
| 104 | 871  | 
\end{ttbox}
 | 
| 324 | 872  | 
\begin{ttdescription}
 | 
| 4317 | 873  | 
\item[\ttindexbold{Type} ($a$, $Ts$)] \index{type constructors|bold}
 | 
| 8136 | 874  | 
  applies the \textbf{type constructor} named~$a$ to the type operand list~$Ts$.
 | 
| 324 | 875  | 
  Type constructors include~\tydx{fun}, the binary function space
 | 
876  | 
  constructor, as well as nullary type constructors such as~\tydx{prop}.
 | 
|
877  | 
Other type constructors may be introduced. In expressions, but not in  | 
|
878  | 
  patterns, \hbox{\tt$S$-->$T$} is a convenient shorthand for function
 | 
|
879  | 
types.  | 
|
| 104 | 880  | 
|
| 4317 | 881  | 
\item[\ttindexbold{TFree} ($a$, $s$)] \index{type variables|bold}
 | 
| 8136 | 882  | 
  is the \textbf{type variable} with name~$a$ and sort~$s$.
 | 
| 104 | 883  | 
|
| 4317 | 884  | 
\item[\ttindexbold{TVar} ($v$, $s$)] \index{type unknowns|bold}
 | 
| 8136 | 885  | 
  is the \textbf{type unknown} with indexname~$v$ and sort~$s$.
 | 
| 324 | 886  | 
Type unknowns are essentially free type variables, but may be  | 
887  | 
instantiated during unification.  | 
|
888  | 
\end{ttdescription}
 | 
|
| 104 | 889  | 
|
890  | 
||
891  | 
\section{Certified types}
 | 
|
892  | 
\index{types!certified|bold}
 | 
|
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
893  | 
Certified types, which are analogous to certified terms, have type  | 
| 275 | 894  | 
\ttindexbold{ctyp}.
 | 
| 104 | 895  | 
|
896  | 
\subsection{Printing types}
 | 
|
| 324 | 897  | 
\index{types!printing of}
 | 
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
898  | 
\begin{ttbox}
 | 
| 275 | 899  | 
string_of_ctyp : ctyp -> string  | 
| 104 | 900  | 
Sign.string_of_typ : Sign.sg -> typ -> string  | 
901  | 
\end{ttbox}
 | 
|
| 324 | 902  | 
\begin{ttdescription}
 | 
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
903  | 
\item[\ttindexbold{string_of_ctyp} $cT$]
 | 
| 104 | 904  | 
displays $cT$ as a string.  | 
905  | 
||
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
906  | 
\item[\ttindexbold{Sign.string_of_typ} $sign$ $T$]
 | 
| 104 | 907  | 
displays $T$ as a string, using the syntax of~$sign$.  | 
| 324 | 908  | 
\end{ttdescription}
 | 
| 104 | 909  | 
|
910  | 
||
911  | 
\subsection{Making and inspecting certified types}
 | 
|
| 
864
 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 
wenzelm 
parents: 
478 
diff
changeset
 | 
912  | 
\begin{ttbox}
 | 
| 4543 | 913  | 
ctyp_of : Sign.sg -> typ -> ctyp  | 
| 8136 | 914  | 
rep_ctyp         : ctyp -> \{T: typ, sign: Sign.sg\}
 | 
| 4543 | 915  | 
Sign.certify_typ : Sign.sg -> typ -> typ  | 
| 104 | 916  | 
\end{ttbox}
 | 
| 324 | 917  | 
\begin{ttdescription}
 | 
| 4543 | 918  | 
|
919  | 
\item[\ttindexbold{ctyp_of} $sign$ $T$] \index{signatures} certifies
 | 
|
920  | 
$T$ with respect to signature~$sign$.  | 
|
921  | 
||
922  | 
\item[\ttindexbold{rep_ctyp} $cT$] decomposes $cT$ as a record
 | 
|
923  | 
containing the type itself and its signature.  | 
|
924  | 
||
925  | 
\item[\ttindexbold{Sign.certify_typ}] is a more primitive version of
 | 
|
926  | 
  \texttt{ctyp_of}, returning the internal representation instead of
 | 
|
927  | 
  an abstract \texttt{ctyp}.
 | 
|
| 104 | 928  | 
|
| 324 | 929  | 
\end{ttdescription}
 | 
| 104 | 930  | 
|
| 1846 | 931  | 
|
| 4317 | 932  | 
\section{Oracles: calling trusted external reasoners}
 | 
| 1846 | 933  | 
\label{sec:oracles}
 | 
934  | 
\index{oracles|(}
 | 
|
935  | 
||
936  | 
Oracles allow Isabelle to take advantage of external reasoners such as  | 
|
937  | 
arithmetic decision procedures, model checkers, fast tautology checkers or  | 
|
938  | 
computer algebra systems. Invoked as an oracle, an external reasoner can  | 
|
939  | 
create arbitrary Isabelle theorems. It is your responsibility to ensure that  | 
|
940  | 
the external reasoner is as trustworthy as your application requires.  | 
|
941  | 
Isabelle's proof objects~(\S\ref{sec:proofObjects}) record how each theorem
 | 
|
942  | 
depends upon oracle calls.  | 
|
943  | 
||
944  | 
\begin{ttbox}
 | 
|
| 4317 | 945  | 
invoke_oracle : theory -> xstring -> Sign.sg * object -> thm  | 
| 
4597
 
a0bdee64194c
Fixed a lot of overfull and underfull lines (hboxes)
 
paulson 
parents: 
4543 
diff
changeset
 | 
946  | 
Theory.add_oracle : bstring * (Sign.sg * object -> term) -> theory  | 
| 
 
a0bdee64194c
Fixed a lot of overfull and underfull lines (hboxes)
 
paulson 
parents: 
4543 
diff
changeset
 | 
947  | 
-> theory  | 
| 1846 | 948  | 
\end{ttbox}
 | 
949  | 
\begin{ttdescription}
 | 
|
| 4317 | 950  | 
\item[\ttindexbold{invoke_oracle} $thy$ $name$ ($sign$, $data$)]
 | 
951  | 
invokes the oracle $name$ of theory $thy$ passing the information  | 
|
952  | 
contained in the exception value $data$ and creating a theorem  | 
|
953  | 
  having signature $sign$.  Note that type \ttindex{object} is just an
 | 
|
954  | 
  abbreviation for \texttt{exn}.  Errors arise if $thy$ does not have
 | 
|
955  | 
an oracle called $name$, if the oracle rejects its arguments or if  | 
|
956  | 
its result is ill-typed.  | 
|
957  | 
||
958  | 
\item[\ttindexbold{Theory.add_oracle} $name$ $fun$ $thy$] extends
 | 
|
959  | 
$thy$ by oracle $fun$ called $name$. It is seldom called  | 
|
960  | 
explicitly, as there is concrete syntax for oracles in theory files.  | 
|
| 1846 | 961  | 
\end{ttdescription}
 | 
962  | 
||
963  | 
A curious feature of {\ML} exceptions is that they are ordinary constructors.
 | 
|
| 6669 | 964  | 
The {\ML} type \texttt{exn} is a datatype that can be extended at any time.  (See
 | 
| 1846 | 965  | 
my {\em {ML} for the Working Programmer}~\cite{paulson-ml2}, especially
 | 
966  | 
page~136.) The oracle mechanism takes advantage of this to allow an oracle to  | 
|
967  | 
take any information whatever.  | 
|
968  | 
||
969  | 
There must be some way of invoking the external reasoner from \ML, either  | 
|
970  | 
because it is coded in {\ML} or via an operating system interface.  Isabelle
 | 
|
971  | 
expects the {\ML} function to take two arguments: a signature and an
 | 
|
| 4317 | 972  | 
exception object.  | 
| 1846 | 973  | 
\begin{itemize}
 | 
974  | 
\item The signature will typically be that of a desendant of the theory  | 
|
975  | 
declaring the oracle. The oracle will use it to distinguish constants from  | 
|
976  | 
variables, etc., and it will be attached to the generated theorems.  | 
|
977  | 
||
978  | 
\item The exception is used to pass arbitrary information to the oracle. This  | 
|
979  | 
information must contain a full description of the problem to be solved by  | 
|
980  | 
the external reasoner, including any additional information that might be  | 
|
981  | 
required. The oracle may raise the exception to indicate that it cannot  | 
|
982  | 
solve the specified problem.  | 
|
983  | 
\end{itemize}
 | 
|
984  | 
||
| 6669 | 985  | 
A trivial example is provided in theory \texttt{FOL/ex/IffOracle}.  This
 | 
| 4317 | 986  | 
oracle generates tautologies of the form $P\bimp\cdots\bimp P$, with  | 
987  | 
an even number of $P$s.  | 
|
| 1846 | 988  | 
|
| 4317 | 989  | 
The \texttt{ML} section of \texttt{IffOracle.thy} begins by declaring
 | 
990  | 
a few auxiliary functions (suppressed below) for creating the  | 
|
991  | 
tautologies. Then it declares a new exception constructor for the  | 
|
992  | 
information required by the oracle: here, just an integer. It finally  | 
|
993  | 
defines the oracle function itself.  | 
|
| 1846 | 994  | 
\begin{ttbox}
 | 
| 4317 | 995  | 
exception IffOracleExn of int;\medskip  | 
996  | 
fun mk_iff_oracle (sign, IffOracleExn n) =  | 
|
997  | 
if n > 0 andalso n mod 2 = 0  | 
|
| 6669 | 998  | 
then Trueprop \$ mk_iff n  | 
| 4317 | 999  | 
else raise IffOracleExn n;  | 
| 1846 | 1000  | 
\end{ttbox}
 | 
| 6669 | 1001  | 
Observe the function's two arguments, the signature \texttt{sign} and the
 | 
| 4317 | 1002  | 
exception given as a pattern. The function checks its argument for  | 
1003  | 
validity. If $n$ is positive and even then it creates a tautology  | 
|
1004  | 
containing $n$ occurrences of~$P$. Otherwise it signals error by  | 
|
1005  | 
raising its own exception (just by happy coincidence). Errors may be  | 
|
| 6669 | 1006  | 
signalled by other means, such as returning the theorem \texttt{True}.
 | 
| 4317 | 1007  | 
Please ensure that the oracle's result is correctly typed; Isabelle  | 
1008  | 
will reject ill-typed theorems by raising a cryptic exception at top  | 
|
1009  | 
level.  | 
|
| 1846 | 1010  | 
|
| 6669 | 1011  | 
The \texttt{oracle} section of \texttt{IffOracle.thy} installs above
 | 
| 4317 | 1012  | 
\texttt{ML} function as follows:
 | 
| 1846 | 1013  | 
\begin{ttbox}
 | 
| 4317 | 1014  | 
IffOracle = FOL +\medskip  | 
1015  | 
oracle  | 
|
1016  | 
iff = mk_iff_oracle\medskip  | 
|
| 1846 | 1017  | 
end  | 
1018  | 
\end{ttbox}
 | 
|
1019  | 
||
| 4317 | 1020  | 
Now in \texttt{IffOracle.ML} we first define a wrapper for invoking
 | 
1021  | 
the oracle:  | 
|
| 1846 | 1022  | 
\begin{ttbox}
 | 
| 
4597
 
a0bdee64194c
Fixed a lot of overfull and underfull lines (hboxes)
 
paulson 
parents: 
4543 
diff
changeset
 | 
1023  | 
fun iff_oracle n = invoke_oracle IffOracle.thy "iff"  | 
| 
 
a0bdee64194c
Fixed a lot of overfull and underfull lines (hboxes)
 
paulson 
parents: 
4543 
diff
changeset
 | 
1024  | 
(sign_of IffOracle.thy, IffOracleExn n);  | 
| 4317 | 1025  | 
\end{ttbox}
 | 
1026  | 
||
1027  | 
Here are some example applications of the \texttt{iff} oracle.  An
 | 
|
1028  | 
argument of 10 is allowed, but one of 5 is forbidden:  | 
|
1029  | 
\begin{ttbox}
 | 
|
1030  | 
iff_oracle 10;  | 
|
| 1846 | 1031  | 
{\out  "P <-> P <-> P <-> P <-> P <-> P <-> P <-> P <-> P <-> P" : thm}
 | 
| 4317 | 1032  | 
iff_oracle 5;  | 
| 1846 | 1033  | 
{\out Exception- IffOracleExn 5 raised}
 | 
1034  | 
\end{ttbox}
 | 
|
1035  | 
||
1036  | 
\index{oracles|)}
 | 
|
| 104 | 1037  | 
\index{theories|)}
 | 
| 5369 | 1038  | 
|
1039  | 
||
1040  | 
%%% Local Variables:  | 
|
1041  | 
%%% mode: latex  | 
|
1042  | 
%%% TeX-master: "ref"  | 
|
1043  | 
%%% End:  |