author | wenzelm |
Tue, 06 May 1997 12:50:16 +0200 | |
changeset 3108 | 335efc3f5632 |
parent 3107 | 492a3d5d2b17 |
child 3109 | d95748813188 |
--- a/doc-src/Ref/classical.tex Mon May 05 21:18:01 1997 +0200 +++ b/doc-src/Ref/classical.tex Tue May 06 12:50:16 1997 +0200 @@ -3,10 +3,11 @@ \index{classical reasoner|(} \newcommand\ainfer[2]{\begin{array}{r@{\,}l}#2\\ \hline#1\end{array}} -Although Isabelle is generic, many users will be working in some extension -of classical first-order logic. Isabelle's set theory~{\tt ZF} is built -upon theory~{\tt FOL}, while higher-order logic contains first-order logic -as a fragment. Theorem-proving in predicate logic is undecidable, but many +Although Isabelle is generic, many users will be working in some +extension of classical first-order logic. Isabelle's set theory~{\tt + ZF} is built upon theory~{\tt FOL}, while higher-order logic +conceptually contains first-order logic as a fragment. +Theorem-proving in predicate logic is undecidable, but many researchers have developed strategies to assist in this task. Isabelle's classical reasoner is an \ML{} functor that accepts certain @@ -38,9 +39,10 @@ effectively, you need to know how it works and how to choose among the many tactics available, including {\tt Fast_tac} and {\tt Best_tac}. -We shall first discuss the underlying principles, then present the classical -reasoner. Finally, we shall see how to instantiate it for new logics. The -logics {\tt FOL}, {\tt HOL} and {\tt ZF} have it already installed. +We shall first discuss the underlying principles, then present the +classical reasoner. Finally, we shall see how to instantiate it for +new logics. The logics \FOL, \ZF, {\HOL} and {\HOLCF} have it already +installed. \section{The sequent calculus} @@ -71,14 +73,18 @@ \iflabelundefined{fol-fig}{from {\it Introduction to Isabelle}}% {Fig.\ts\ref{fol-fig}}. The sequent calculus analogue of~$({\imp}I)$ is the rule -$$ \ainfer{\Gamma &\turn \Delta, P\imp Q}{P,\Gamma &\turn \Delta,Q} - \eqno({\imp}R) $$ +$$ +\ainfer{\Gamma &\turn \Delta, P\imp Q}{P,\Gamma &\turn \Delta,Q} +\eqno({\imp}R) +$$ This breaks down some implication on the right side of a sequent; $\Gamma$ and $\Delta$ stand for the sets of formulae that are unaffected by the inference. The analogue of the pair~$({\disj}I1)$ and~$({\disj}I2)$ is the single rule -$$ \ainfer{\Gamma &\turn \Delta, P\disj Q}{\Gamma &\turn \Delta,P,Q} - \eqno({\disj}R) $$ +$$ +\ainfer{\Gamma &\turn \Delta, P\disj Q}{\Gamma &\turn \Delta,P,Q} +\eqno({\disj}R) +$$ This breaks down some disjunction on the right side, replacing it by both disjuncts. Thus, the sequent calculus is a kind of multiple-conclusion logic. @@ -115,15 +121,16 @@ other connectives, we use sequent-style elimination rules instead of destruction rules such as $({\conj}E1,2)$ and $(\forall E)$. But note that the rule $(\neg L)$ has no effect under our representation of sequents! -$$ \ainfer{\neg P,\Gamma &\turn \Delta}{\Gamma &\turn \Delta,P} - \eqno({\neg}L) $$ +$$ +\ainfer{\neg P,\Gamma &\turn \Delta}{\Gamma &\turn \Delta,P}\eqno({\neg}L) +$$ What about reasoning on the right? Introduction rules can only affect the formula in the conclusion, namely~$Q@1$. The other right-side formulae are represented as negated assumptions, $\neg Q@2$, \ldots,~$\neg Q@n$. \index{assumptions!negated} In order to operate on one of these, it must first be exchanged with~$Q@1$. Elim-resolution with the {\bf swap} rule has this effect: -$$ \List{\neg P; \; \neg R\Imp P} \Imp R \eqno(swap)$$ +$$ \List{\neg P; \; \neg R\Imp P} \Imp R \eqno(swap) $$ To ensure that swaps occur only when necessary, each introduction rule is converted into a swapped form: it is resolved with the second premise of~$(swap)$. The swapped form of~$({\conj}I)$, which might be @@ -172,8 +179,10 @@ $$ \List{\forall x{.}P(x); P(t)\Imp Q} \Imp Q \eqno(\forall E@2) $$ Elim-resolution with this rule will delete the universal formula after a single use. To replicate universal quantifiers, replace the rule by -$$ \List{\forall x{.}P(x);\; \List{P(t); \forall x{.}P(x)}\Imp Q} \Imp Q. - \eqno(\forall E@3) $$ +$$ +\List{\forall x{.}P(x);\; \List{P(t); \forall x{.}P(x)}\Imp Q} \Imp Q. +\eqno(\forall E@3) +$$ To replicate existential quantifiers, replace $(\exists I)$ by \[ \List{\neg(\exists x{.}P(x)) \Imp P(t)} \Imp \exists x{.}P(x). \] All introduction rules mentioned above are also useful in swapped form. @@ -290,14 +299,15 @@ below). They may perform a form of Modus Ponens: if there are assumptions $P\imp Q$ and~$P$, then replace $P\imp Q$ by~$Q$. -The classical reasoning tactics---except {\tt blast_tac}!---allow you to -modify this basic proof strategy by -applying two arbitrary {\bf wrapper tacticals} to it. This affects each step of -the search. Usually they are the identity tacticals, but they could apply -another tactic before or after the step tactic. The first one, which is -considered to be safe, affects \ttindex{safe_step_tac} and all the tactics that -call it. The the second one, which may be unsafe, affects -\ttindex{step_tac}, \ttindex{slow_step_tac} and the tactics that call them. +The classical reasoning tactics --- except {\tt blast_tac}! --- allow +you to modify this basic proof strategy by applying two arbitrary {\bf + wrapper tacticals} to it. This affects each step of the search. +Usually they are the identity tacticals, but they could apply another +tactic before or after the step tactic. The first one, which is +considered to be safe, affects \ttindex{safe_step_tac} and all the +tactics that call it. The the second one, which may be unsafe, affects +\ttindex{step_tac}, \ttindex{slow_step_tac} and the tactics that call +them. \begin{ttbox} addss : claset * simpset -> claset \hfill{\bf infix 4} @@ -315,10 +325,12 @@ (int -> tactic)) -> claset \hfill{\bf infix 4} \end{ttbox} % -\index{simplification!from classical reasoner} -The wrapper tacticals underly the operator \ttindex{addss}, which combines -each search step by simplification. Strictly speaking, {\tt addss} is not -part of the classical reasoner. It should be defined (using {\tt addSaltern (CHANGED o (safe_asm_more_full_simp_tac ss)}) when the simplifier is installed. +\index{simplification!from classical reasoner} The wrapper tacticals +underly the operator addss, which combines each search step by +simplification. Strictly speaking, {\tt addss} is not part of the +classical reasoner. It should be defined (using {\tt addSaltern + (CHANGED o (safe_asm_more_full_simp_tac ss)}) when the simplifier is +installed. \begin{ttdescription} \item[$cs$ addss $ss$] \indexbold{*addss} @@ -501,7 +513,7 @@ warnings. Just like simpsets, clasets can be associated with theories. The tactics \begin{ttbox} -Blast_tac : int -> tactic +Blast_tac : int -> tactic Fast_tac : int -> tactic Best_tac : int -> tactic Deepen_tac : int -> int -> tactic @@ -620,8 +632,8 @@ tactics are assumed to be safe! \end{ttdescription} The functor is not at all sensitive to the formalization of the -object-logic. It does not even examine the rules, but merely applies them -according to its fixed strategy. The functor resides in {\tt -Provers/classical.ML} in the Isabelle distribution directory. +object-logic. It does not even examine the rules, but merely applies +them according to its fixed strategy. The functor resides in {\tt + Provers/classical.ML} in the Isabelle sources. \index{classical reasoner|)}
--- a/doc-src/Ref/defining.tex Mon May 05 21:18:01 1997 +0200 +++ b/doc-src/Ref/defining.tex Tue May 06 12:50:16 1997 +0200 @@ -266,7 +266,7 @@ abstract syntax tree. -\subsection{*Inspecting the syntax} +\subsection{*Inspecting the syntax} \label{pg:print_syn} \begin{ttbox} syn_of : theory -> Syntax.syntax print_syntax : theory -> unit @@ -311,6 +311,7 @@ {\out type = tvar "::" sort[0] => "_ofsort" (1000)} {\out \vdots} \ttbreak +{\out print modes: "symbols" "xterm"} {\out consts: "_K" "_appl" "_aprop" "_args" "_asms" "_bigimpl" \dots} {\out parse_ast_translation: "_appl" "_bigimpl" "_bracket"} {\out "_idtyp" "_lambda" "_tapp" "_tappl"} @@ -353,6 +354,9 @@ conceptually, they are removed from the grammar by adding new productions. Priority information attached to chain productions is ignored; only the dummy value $-1$ is displayed. + + \item[\ttindex{print_modes}] lists the alternative print modes + provided by this syntax (see \S\ref{sec:prmodes}). \item[{\tt consts}, {\tt parse_rules}, {\tt print_rules}] relate to macros (see \S\ref{sec:macros}). @@ -483,7 +487,7 @@ declarations encoding the priority grammar from \S\ref{sec:priority_grammars}: \begin{ttbox} -EXP = Pure + +ExpSyntax = Pure + types exp syntax @@ -493,10 +497,10 @@ "-" :: exp => exp ("- _" [3] 3) end \end{ttbox} -If you put this into a file {\tt EXP.thy} and load it via {\tt use_thy"EXP"}, -you can run some tests: +If you put this into a file {\tt ExpSyntax.thy} and load it via {\tt + use_thy"ExpSyntax"}, you can run some tests: \begin{ttbox} -val read_exp = Syntax.test_read (syn_of EXP.thy) "exp"; +val read_exp = Syntax.test_read (syn_of ExpSyntax.thy) "exp"; {\out val it = fn : string -> unit} read_exp "0 * 0 * 0 * 0 + 0 + 0 + 0"; {\out tokens: "0" "*" "0" "*" "0" "*" "0" "+" "0" "+" "0" "+" "0"} @@ -515,18 +519,19 @@ Executing {\tt Syntax.print_gram} reveals the productions derived from the above mixfix declarations (lots of additional information deleted): \begin{ttbox} -Syntax.print_gram (syn_of EXP.thy); +Syntax.print_gram (syn_of ExpSyntax.thy); {\out exp = "0" => "0" (9)} {\out exp = exp[0] "+" exp[1] => "+" (0)} {\out exp = exp[3] "*" exp[2] => "*" (2)} {\out exp = "-" exp[3] => "-" (3)} \end{ttbox} -Note that because {\tt exp} is not of class {\tt logic}, it has been retained -as a separate nonterminal. This also entails that the syntax does not provide -for identifiers or paranthesized expressions. Normally you would also want to -add the declaration {\tt arities exp :: logic} and use {\tt consts} instead -of {\tt syntax}. Try this as an exercise and study the changes in the +Note that because {\tt exp} is not of class {\tt logic}, it has been +retained as a separate nonterminal. This also entails that the syntax +does not provide for identifiers or paranthesized expressions. +Normally you would also want to add the declaration {\tt arities + exp::logic} after {\tt types} and use {\tt consts} instead of {\tt + syntax}. Try this as an exercise and study the changes in the grammar. \subsection{The mixfix template} @@ -576,10 +581,9 @@ \subsection{Infixes} \indexbold{infixes} -Infix operators associating to the left or right can be declared -using {\tt infixl} or {\tt infixr}. -Roughly speaking, the form {\tt $c$ ::\ $\sigma$ (infixl $p$)} -abbreviates the mixfix declarations +Infix operators associating to the left or right can be declared using +{\tt infixl} or {\tt infixr}. Basically, the form {\tt $c$ ::\ + $\sigma$ (infixl $p$)} abbreviates the mixfix declarations \begin{ttbox} "op \(c\)" :: \(\sigma\) ("(_ \(c\)/ _)" [\(p\), \(p+1\)] \(p\)) "op \(c\)" :: \(\sigma\) ("op \(c\)") @@ -594,6 +598,16 @@ function symbols, as in \ML. Special characters occurring in~$c$ must be escaped, as in delimiters, using a single quote. +A slightly more general form of infix declarations allows constant +names to be independent from their concrete syntax, namely \texttt{$c$ + ::\ $\sigma$\ (infixl "$sy$" $p$)}, the same for \texttt{infixr}. As +an example consider: +\begin{ttbox} +and :: [bool, bool] => bool (infixr "&" 35) +\end{ttbox} +The internal constant name will then be just \texttt{and}, without any +\texttt{op} prefixed. + \subsection{Binders} \indexbold{binders} @@ -642,6 +656,45 @@ \index{mixfix declarations|)} + +\section{*Alternative print modes} \label{sec:prmodes} +\index{print modes|(} +% +Isabelle's pretty printer supports alternative output syntaxes. These +may be used independently or in cooperation. The currently active +print modes (with precedence from left to right) are determined by a +reference variable. +\begin{ttbox}\index{*print_mode} +print_mode: string list ref +\end{ttbox} +Initially this may already contain some print mode identifiers, +depending on how Isabelle has been invoked (e.g.\ by some user +interface). So changes should be incremental --- adding or deleting +modes relative to the current value. + +Any \ML{} string is a legal print mode identifier, without any +predeclaration required. The following names should be considered +reserved, though: \texttt{""} (yes, the empty string), +\texttt{symbols}, \texttt{latex}, \texttt{xterm}. + +There is a separate table of mixfix productions for pretty printing +associated with each print mode. The currently active ones are +conceptually just concatenated from left to right, with the standard +syntax output table always coming last as default. Thus mixfix +productions of preceding modes in the list may override those of later +ones. Also note that token translations are always relative to some +print mode (see \S\ref{sec:tok_tr}). + +\medskip The canonical application of print modes is optional printing +of mathematical symbols from a special screen font instead of {\sc + ascii}. Another example is to re-use Isabelle's advanced +$\lambda$-term printing mechanisms to generate completely different +output, say for interfacing external tools like model checkers (e.g.\ +see \texttt{HOL/ex/EindhovenMC.thy}). + +\index{print modes|)} + + \section{Ambiguity of parsed expressions} \label{sec:ambiguity} \index{ambiguity!of parsed expressions} @@ -810,12 +863,7 @@ end \end{ttbox} And if we want to have all three connectives together, we create and load a -theory file consisting of a single line:\footnote{We can combine the - theories without creating a theory file using the ML declaration -\begin{ttbox} -val MinIFC_thy = merge_theories(MinIF,MinC) -\end{ttbox} -\index{*merge_theories|fnote}} +theory file consisting of a single line: \begin{ttbox} MinIFC = MinIF + MinC \end{ttbox}
--- a/doc-src/Ref/goals.tex Mon May 05 21:18:01 1997 +0200 +++ b/doc-src/Ref/goals.tex Tue May 06 12:50:16 1997 +0200 @@ -25,14 +25,14 @@ \section{Basic commands} Most proofs begin with {\tt goal} or {\tt goalw} and require no other -commands than {\tt by}, {\tt chop} and {\tt undo}. They typically end with -a call to {\tt result}. +commands than {\tt by}, {\tt chop} and {\tt undo}. They typically end +with a call to {\tt qed}. \subsection{Starting a backward proof} \index{proofs!starting} \begin{ttbox} goal : theory -> string -> thm list goalw : theory -> thm list -> string -> thm list -goalw_cterm : thm list -> Sign.cterm -> thm list +goalw_cterm : thm list -> cterm -> thm list premises : unit -> thm list \end{ttbox} The {\tt goal} commands start a new proof by setting the goal. They @@ -46,10 +46,11 @@ \begin{ttbox} val prems = goal{\it theory\/ formula}; \end{ttbox}\index{assumptions!of main goal} -These assumptions serve as the premises when you are deriving a rule. They -are also stored internally and can be retrieved later by the function {\tt - premises}. When the proof is finished, {\tt result} compares the -stored assumptions with the actual assumptions in the proof state. +These assumptions serve as the premises when you are deriving a rule. +They are also stored internally and can be retrieved later by the +function {\tt premises}. When the proof is finished, {\tt qed} +compares the stored assumptions with the actual assumptions in the +proof state. \index{definitions!unfolding} Some of the commands unfold definitions using meta-rewrite rules. This @@ -74,11 +75,11 @@ meta-rewrite rules to the first subgoal and the premises. \index{meta-rewriting} -\item[\ttindexbold{goalw_cterm} {\it theory} {\it defs} {\it ct};] -is a version of {\tt goalw} for programming applications. The main goal is -supplied as a cterm, not as a string. Typically, the cterm is created from -a term~$t$ by \hbox{\tt Sign.cterm_of (sign_of thy) $t$}. -\index{*Sign.cterm_of}\index{*sign_of} +\item[\ttindexbold{goalw_cterm} {\it theory} {\it defs} {\it ct};] is + a version of {\tt goalw} for programming applications. The main + goal is supplied as a cterm, not as a string. Typically, the cterm + is created from a term~$t$ by \hbox{\tt cterm_of (sign_of thy) $t$}. + \index{*cterm_of}\index{*sign_of} \item[\ttindexbold{premises}()] returns the list of meta-hypotheses associated with the current proof (in @@ -398,7 +399,7 @@ qed_goal : string->theory-> string->(thm list->tactic list)->unit prove_goalw: theory->thm list->string->(thm list->tactic list)->thm qed_goalw : string->theory->thm list->string->(thm list->tactic list)->unit -prove_goalw_cterm: thm list->Sign.cterm->(thm list->tactic list)->thm +prove_goalw_cterm: thm list->cterm->(thm list->tactic list)->thm \end{ttbox} These batch functions create an initial proof state, then apply a tactic to it, yielding a sequence of final proof states. The head of the sequence is @@ -413,11 +414,11 @@ \begin{ttbox} val prems = goal {\it theory} {\it formula}; by \(tac@1\); \ldots by \(tac@n\); -val my_thm = result(); +qed "my_thm"; \end{ttbox} can be transformed to an expression as follows: \begin{ttbox} -val my_thm = prove_goal {\it theory} {\it formula} +qed_goal "my_thm" {\it theory} {\it formula} (fn prems=> [ \(tac@1\), \ldots, \(tac@n\) ]); \end{ttbox} The methods perform identical processing of the initial {\it formula} and @@ -449,9 +450,9 @@ goal is supplied as a cterm, not as a string. Typically, the cterm is created from a term~$t$ as follows: \begin{ttbox} -Sign.cterm_of (sign_of thy) \(t\) +cterm_of (sign_of thy) \(t\) \end{ttbox} -\index{*Sign.cterm_of}\index{*sign_of} +\index{*cterm_of}\index{*sign_of} \end{ttdescription}
--- a/doc-src/Ref/introduction.tex Mon May 05 21:18:01 1997 +0200 +++ b/doc-src/Ref/introduction.tex Tue May 06 12:50:16 1997 +0200 @@ -1,9 +1,11 @@ %% $Id$ + \chapter{Basic Use of Isabelle}\index{sessions|(} -The Reference Manual is a comprehensive description of Isabelle, including -all commands, functions and packages. It really is intended for reference, -perhaps for browsing, but not for reading through. It is not a tutorial, -but assumes familiarity with the basic concepts of Isabelle. +The Reference Manual is a comprehensive description of Isabelle +proper, including all \ML{} commands, functions and packages. It +really is intended for reference, perhaps for browsing, but not for +reading through. It is not a tutorial, but assumes familiarity with +the basic logical concepts of Isabelle. When you are looking for a way of performing some task, scan the Table of Contents for a relevant heading. Functions are organized by their purpose, @@ -19,76 +21,98 @@ \section{Basic interaction with Isabelle} \index{starting up|bold}\nobreak % -First, read the instructions on file {\tt README} in the top-level -distribution directory. Follow them to create the object-logics you need -on a directory you have created to hold binary images. Suppose the -environment variable {\tt ISABELLEBIN} refers to this directory. The -instructions for starting up a logic (say {\tt HOL}) depend upon which -Standard ML compiler you are using. -\begin{itemize} -\item With Standard ML of New Jersey, type \verb|$ISABELLEBIN/HOL|. -\item With Poly/ML, type \verb|poly $ISABELLEBIN/HOL|, possibly prefixing the - command with the directory where {\tt poly} is kept. -\end{itemize} -Either way, you will find yourself at the \ML{} top level. All Isabelle -commands are bound to \ML{} identifiers. +We assume that your local Isabelle administrator (this might be you!) +has already installed the \Pure\ system and several object-logics +properly --- otherwise see the {\tt INSTALL} file in the top-level +directory of the distribution on how to build it. + +\medskip Let $\langle isabellehome \rangle$ denote the location where +the distribution has been installed. To run Isabelle from a the shell +prompt within an ordinary text terminal session, simply type: +\begin{ttbox} +\({\langle}isabellehome{\rangle}\)/bin/isabelle +\end{ttbox} +This should start an interactive \ML{} session with the default +object-logic already preloaded. All Isabelle commands are bound to +\ML{} identifiers. + +Subsequently we assume that {\tt \(\langle isabellehome \rangle\)/bin} +has been added to your shell's search path, in order to avoid typing +full path specifications of the executable files. + +The object-logic image to load may be also specified explicitly as an +argument to the {\tt isabelle} command, e.g.: +\begin{ttbox} +isabelle FOL +\end{ttbox} +This should put you into the world of polymorphic first-order logic +(assuming that {\FOL} has been pre-built). -\index{saving your work|bold} -Isabelle provides no means of storing theorems or proofs on files. -Theorems are simply part of the \ML{} state and are named by \ML{} -identifiers. To save your work between sessions, you must save a copy of -the \ML{} image. The procedure for doing so is compiler-dependent: -\begin{itemize}\index{Poly/{\ML} compiler} -\item At the end of a session, Poly/\ML{} saves the state, provided you - have created a database for your own use. You can create a database by - copying an existing one, or by calling the Poly/\ML{} function {\tt - make_database}; the latter course uses much less disc space. A - Poly/\ML{} database {\em does not\/} save the contents of references, - such as the current state of a backward proof. +\index{saving your work|bold} Isabelle provides no means of storing +theorems or proofs on files. Theorems are simply part of the \ML{} +state and are named by \ML{} identifiers. To save your work between +sessions, you must dump the \ML{} system state to a file. This is done +automatically when ending the session normally (e.g.\ by typing +control-D), provided that the image has been opened \emph{writable} in +the first place. The standard object-logics are usually read-only, so +you probably have to create a private working copy first. For example, +the following shell command puts you into a writable Isabelle session +of name \texttt{Foo} that initially contains just \FOL: +\begin{ttbox} +isabelle FOL Foo +\end{ttbox} +Ending the \texttt{Foo} session with control-D will cause the complete +\ML{} world to be saved somewhere in your home directory\footnote{The + default location is in \texttt{\~\relax/isabelle/heaps}, but this + depends on your local configuration.}. Make sure there is enough +space available! Then one may later continue at exactly the same point +by running +\begin{ttbox} +isabelle Foo +\end{ttbox} -\item With New Jersey \ML{} you must save the state explicitly before -ending the session. While a Poly/\ML{} database can be small, a New Jersey -image occupies several megabytes. -\end{itemize} -See your \ML{} compiler's documentation for full instructions on saving the -state. +%FIXME not yet +%More details about \texttt{isabelle} may be found in the \emph{System +% Manual}. + +\medskip Saving the state is not enough. Record, on a file, the +top-level commands that generate your theories and proofs. Such a +record allows you to replay the proofs whenever required, for instance +after making minor changes to the axioms. Ideally, your record will +be somewhat intelligible to others as a formal description of your +work. -Saving the state is not enough. Record, on a file, the top-level commands -that generate your theories and proofs. Such a record allows you to replay -the proofs whenever required, for instance after making minor changes to -the axioms. Ideally, your record will be intelligible to others as a -formal description of your work. +\medskip There are more comfortable user interfaces than the +bare-bones \ML{} top-level run from a text terminal. The +\texttt{Isabelle} executable (note the capital I) runs one such +interface, depending on your local configuration. Furthermore there +are a number of external utilities available. These are started +uniformly via the \texttt{isatool} wrapper. -Since Isabelle's user interface is the \ML{} top level, some kind of window -support is essential. One window displays the Isabelle session, while the -other displays a file --- your proof record --- being edited. The Emacs -editor supports windows and can manage interactive sessions. +%FIXME not yet +%Again, see the \emph{System Manual} for more information user +%interfaces and utilities. \section{Ending a session} \begin{ttbox} -quit : unit -> unit -commit : unit -> unit \hfill{\bf Poly/ML only} -exportML : string -> bool \hfill{\bf New Jersey ML only} +quit : unit -> unit +exit : int -> unit +commit : unit -> unit \end{ttbox} \begin{ttdescription} -\item[\ttindexbold{quit}();] -aborts the Isabelle session, without saving the state. +\item[\ttindexbold{quit}();] ends the Isabelle session, without saving + the state. -\item[\ttindexbold{commit}();] - saves the current state in your Poly/\ML{} database without ending the - session. The contents of references are lost, so never do this during an - interactive proof!\index{Poly/{\ML} compiler} +\item[\ttindexbold{exit}();] same as {\tt quit}, passing a return code + to the operating system. -\item[\ttindexbold{exportML} "{\it file}";] saves an -image of your session to the given {\it file}. +\item[\ttindexbold{commit}();] saves the current state without ending + the session, provided that the logic image is opened read-write. \end{ttdescription} -\begin{warn} -Typing control-D also finishes the session, but its effect is -compiler-dependent. Poly/\ML{} will then save the state, if you have a -private database. New Jersey \ML{} will discard the state! -\end{warn} +Typing control-D also finishes the session in essentially the same way +as the sequence {\tt commit(); quit();} would. \section{Reading ML files} @@ -105,7 +129,7 @@ changes the current directory to {\it dir}. This is the default directory for reading files and for writing temporary files. -\item[\ttindexbold{pwd} ();] returns the path of the current directory. +\item[\ttindexbold{pwd}();] returns the path of the current directory. \item[\ttindexbold{use} "$file$";] reads the given {\it file} as input to the \ML{} session. Reading a file @@ -116,6 +140,16 @@ \end{ttdescription} +\section{Setting flags} +\begin{ttbox} +set : bool ref -> bool +reset : bool ref -> bool +toggle : bool ref -> bool +\end{ttbox}\index{*set}\index{*reset}\index{*toggle} +These are some shorthands for manipulating boolean references. The new +value is returned. + + \section{Printing of terms and theorems}\label{sec:printing-control} \index{printing control|(} Isabelle's pretty printer is controlled by a number of parameters. @@ -236,35 +270,36 @@ theory used in the last interactive proof. \end{warn} -\section{Shell scripts}\label{sec:shell-scripts} -\index{shell scripts|bold} The following files are distributed with -Isabelle, and work under Unix$^{\rm TM}$. They can be executed as commands -to the Unix shell. Some of them depend upon shell environment variables. -\begin{ttdescription} -\item[make-all $switches$] \index{*make-all shell script} - compiles the Isabelle system, when executed on the source directory. - Environment variables specify which \ML{} compiler to use. These - variables, and the {\it switches}, are documented on the file itself. - -\item[teeinput $program$] \index{*teeinput shell script} - executes the {\it program}, while piping the standard input to a log file - designated by the \verb|$LISTEN| environment variable. Normally the - program is Isabelle, and the log file receives a copy of all the Isabelle - commands. - -\item[xlisten $program$] \index{*xlisten shell script} - is a trivial `user interface' for the X Window System. It creates two - windows using {\tt xterm}. One executes an interactive session via - \hbox{\tt teeinput $program$}, while the other displays the log file. To - make a proof record, simply paste lines from the log file into an editor - window. - -\item[expandshort $files$] \index{*expandshort shell script} - reads the {\it files\/} and replaces all occurrences of the shorthand - commands {\tt br}, {\tt be}, {\tt brs}, {\tt bes}, etc., with the - corresponding full commands. Shorthand commands should appear one - per line. The old versions of the files - are renamed to have the suffix~\verb'~~'. -\end{ttdescription} +%FIXME remove +%\section{Shell scripts}\label{sec:shell-scripts} +%\index{shell scripts|bold} The following files are distributed with +%Isabelle, and work under Unix$^{\rm TM}$. They can be executed as commands +%to the Unix shell. Some of them depend upon shell environment variables. +%\begin{ttdescription} +%\item[make-all $switches$] \index{*make-all shell script} +% compiles the Isabelle system, when executed on the source directory. +% Environment variables specify which \ML{} compiler to use. These +% variables, and the {\it switches}, are documented on the file itself. +% +%\item[teeinput $program$] \index{*teeinput shell script} +% executes the {\it program}, while piping the standard input to a log file +% designated by the \verb|$LISTEN| environment variable. Normally the +% program is Isabelle, and the log file receives a copy of all the Isabelle +% commands. +% +%\item[xlisten $program$] \index{*xlisten shell script} +% is a trivial `user interface' for the X Window System. It creates two +% windows using {\tt xterm}. One executes an interactive session via +% \hbox{\tt teeinput $program$}, while the other displays the log file. To +% make a proof record, simply paste lines from the log file into an editor +% window. +% +%\item[expandshort $files$] \index{*expandshort shell script} +% reads the {\it files\/} and replaces all occurrences of the shorthand +% commands {\tt br}, {\tt be}, {\tt brs}, {\tt bes}, etc., with the +% corresponding full commands. Shorthand commands should appear one +% per line. The old versions of the files +% are renamed to have the suffix~\verb'~~'. +%\end{ttdescription} \index{sessions|)}
--- a/doc-src/Ref/ref.ind Mon May 05 21:18:01 1997 +0200 +++ b/doc-src/Ref/ref.ind Tue May 06 12:50:16 1997 +0200 @@ -1,806 +1,811 @@ \begin{theindex} - \item {\ptt !!} symbol, 64 + \item {\tt !!} symbol, 65 \subitem in main goal, 7 - \item {\tt\$}, \bold{56}, 79 - \item {\tt\%} symbol, 64 + \item {\tt\$}, \bold{57}, 82 + \item {\tt\%} symbol, 65 \item * - \subitem claset, 121 - \subitem simpset, 99 - \item {\ptt ::} symbol, 64, 65 - \item {\ptt ==} symbol, 64 - \item {\ptt ==>} symbol, 64 - \item {\ptt =>} symbol, 64 - \item {\ptt =?=} symbol, 64 - \item {\tt\at Enum} constant, 85 - \item {\tt\at Finset} constant, 85 - \item {\ptt [} symbol, 64 - \item {\ptt [|} symbol, 64 - \item {\ptt ]} symbol, 64 - \item {\ptt _K} constant, 86, 88 - \item \verb'{}' symbol, 85 - \item {\tt\ttlbrace} symbol, 64 - \item {\tt\ttrbrace} symbol, 64 - \item {\ptt |]} symbol, 64 + \subitem claset, 127 + \subitem simpset, 106 + \item {\tt ::} symbol, 65, 66 + \item {\tt ==} symbol, 65 + \item {\tt ==>} symbol, 65 + \item {\tt =>} symbol, 65 + \item {\tt =?=} symbol, 65 + \item {\tt\at Enum} constant, 88 + \item {\tt\at Finset} constant, 88, 89 + \item {\tt [} symbol, 65 + \item {\tt [|} symbol, 65 + \item {\tt ]} symbol, 65 + \item {\tt _K} constant, 90, 92 + \item \verb'{}' symbol, 88 + \item {\tt\ttlbrace} symbol, 65 + \item {\tt\ttrbrace} symbol, 65 + \item {\tt |]} symbol, 65 \indexspace - \item {\ptt Abs}, \bold{56}, 79 - \item {\ptt abstract_over}, \bold{57} - \item {\ptt abstract_rule}, \bold{41} - \item {\ptt aconv}, \bold{57} - \item {\ptt addaltern}, \bold{118} - \item {\ptt addbefore}, \bold{118} - \item {\ptt addcongs}, 97, 110, \bold{110} - \item {\ptt AddDs}, \bold{121} - \item {\ptt addDs}, \bold{116} - \item {\ptt addeqcongs}, \bold{97}, 100, 110 - \item {\ptt AddEs}, \bold{121} - \item {\ptt addEs}, \bold{116} - \item {\ptt AddIs}, \bold{121} - \item {\ptt addIs}, \bold{116} - \item {\ptt addloop}, 99, 100 - \item {\ptt addSaltern}, \bold{118} - \item {\ptt addSbefore}, \bold{118} - \item {\ptt AddSDs}, \bold{121} - \item {\ptt addSDs}, \bold{116} - \item {\ptt AddSEs}, \bold{121} - \item {\ptt addSEs}, \bold{116} - \item {\ptt Addsimps}, \bold{95}, 99, 100 - \item {\ptt addsimps}, 96, 99, 100, 110 - \item {\ptt AddSIs}, \bold{121} - \item {\ptt addSIs}, \bold{116} - \item {\ptt addSolver}, 98, 100 - \item {\ptt addss}, 117, \bold{117}, 118 - \item {\ptt addSSolver}, 98, 100 - \item {\ptt all_tac}, \bold{27} - \item {\ptt ALLGOALS}, \bold{32}, 103, 106 + \item {\tt Abs}, \bold{57}, 82 + \item {\tt abstract_over}, \bold{58} + \item {\tt abstract_rule}, \bold{44} + \item {\tt aconv}, \bold{58} + \item {\tt addaltern}, \bold{124} + \item {\tt addbefore}, \bold{124} + \item {\tt addcongs}, 103, \bold{115}, 116 + \item {\tt AddDs}, \bold{127} + \item {\tt addDs}, \bold{123} + \item {\tt addeqcongs}, \bold{103}, 105, 115 + \item {\tt AddEs}, \bold{127} + \item {\tt addEs}, \bold{123} + \item {\tt AddIs}, \bold{127} + \item {\tt addIs}, \bold{123} + \item {\tt addloop}, 104, 105 + \item {\tt addSaltern}, \bold{124} + \item {\tt addSbefore}, \bold{124} + \item {\tt AddSDs}, \bold{127} + \item {\tt addSDs}, \bold{123} + \item {\tt AddSEs}, \bold{127} + \item {\tt addSEs}, \bold{123} + \item {\tt Addsimps}, \bold{100}, 105, 106 + \item {\tt addsimps}, 102, 105, 106, 116 + \item {\tt AddSIs}, \bold{127} + \item {\tt addSIs}, \bold{123} + \item {\tt addSolver}, 104, 105 + \item {\tt addss}, \bold{124}, 125 + \item {\tt addSSolver}, 104, 105 + \item {\tt all_tac}, \bold{30} + \item {\tt ALLGOALS}, \bold{34}, 108, 112 \item ambiguity - \subitem of parsed expressions, 72 - \item {\ptt any} nonterminal, \bold{63} - \item {\ptt APPEND}, \bold{26}, 27 - \item {\ptt APPEND'}, 33 - \item {\ptt Appl}, 76 - \item {\ptt aprop} nonterminal, \bold{63} - \item {\ptt ares_tac}, \bold{18} - \item {\ptt args} nonterminal, 85 - \item {\ptt Arith} theory, 105 + \subitem of parsed expressions, 75 + \item {\tt any} nonterminal, \bold{64} + \item {\tt APPEND}, \bold{28}, 30 + \item {\tt APPEND'}, 35 + \item {\tt Appl}, 79 + \item {\tt aprop} nonterminal, \bold{66} + \item {\tt ares_tac}, \bold{19} + \item {\tt args} nonterminal, 89 + \item {\tt Arith} theory, 111 \item arities - \subitem context conditions, 48 - \item {\ptt Asm_full_simp_tac}, \bold{95}, 100 - \item {\ptt asm_full_simp_tac}, 20, \bold{99}, 100, 104 - \item {\ptt asm_rl} theorem, 19 - \item {\ptt Asm_simp_tac}, \bold{94}, 100 - \item {\ptt asm_simp_tac}, \bold{99}, 100, 102, 110 - \item associative-commutative operators, 105 - \item {\ptt assume}, \bold{40} - \item {\ptt assume_ax}, 8, \bold{52} - \item {\ptt assume_tac}, \bold{16}, 116 - \item {\ptt assumption}, \bold{43} + \subitem context conditions, 52 + \item {\tt Asm_full_simp_tac}, \bold{100}, 105 + \item {\tt asm_full_simp_tac}, 22, 105, \bold{106}, 109 + \item {\tt asm_rl} theorem, 21 + \item {\tt Asm_simp_tac}, \bold{100}, 105 + \item {\tt asm_simp_tac}, 105, \bold{106}, 107, 116 + \item associative-commutative operators, 110 + \item {\tt assume}, \bold{42} + \item {\tt assume_ax}, 8, \bold{55} + \item {\tt assume_tac}, \bold{17}, 122 + \item {\tt assumption}, \bold{45} \item assumptions - \subitem contradictory, 121 - \subitem deleting, 20 - \subitem in simplification, 94, 98 - \subitem inserting, 18 - \subitem negated, 114 - \subitem of main goal, 6, 8, 13 - \subitem reordering, 104 - \subitem rotating, 20 - \subitem substitution in, 91 - \subitem tactics for, 16 - \item ASTs, 76--80 - \subitem made from parse trees, 77 - \subitem made from terms, 79 - \item {\ptt atac}, \bold{17} + \subitem contradictory, 128 + \subitem deleting, 22 + \subitem in simplification, 99, 104 + \subitem inserting, 19 + \subitem negated, 120 + \subitem of main goal, 7--9, 14 + \subitem reordering, 109 + \subitem rotating, 22 + \subitem substitution in, 96 + \subitem tactics for, 17 + \item ASTs, 79--84 + \subitem made from parse trees, 80 + \subitem made from terms, 82 + \item {\tt atac}, \bold{19} + \item {\tt axclass} section, 51 + \item axiomatic type class, 51 \item axioms - \subitem extracting, 51 - \item {\ptt axioms_of}, \bold{52} + \subitem extracting, 55 + \item {\tt axioms_of}, \bold{56} \indexspace - \item {\ptt ba}, \bold{10} - \item {\ptt back}, \bold{9} - \item batch execution, 11 - \item {\ptt bd}, \bold{10} - \item {\ptt bds}, \bold{10} - \item {\ptt be}, \bold{10} - \item {\ptt bes}, \bold{10} - \item {\ptt BEST_FIRST}, \bold{29}, 30 - \item {\ptt Best_tac}, \bold{121} - \item {\ptt best_tac}, \bold{119} - \item {\ptt beta_conversion}, \bold{41} - \item {\ptt bicompose}, \bold{43} - \item {\ptt bimatch_tac}, \bold{22} - \item {\ptt bind_thm}, 8, \bold{8}, 35 - \item binders, \bold{72} - \item {\ptt biresolution}, \bold{43} - \item {\ptt biresolve_tac}, \bold{22}, 121 - \item {\ptt Blast.depth_tac}, \bold{119} - \item {\ptt Blast.trace}, \bold{119} - \item {\ptt Blast_tac}, \bold{121} - \item {\ptt blast_tac}, \bold{119} - \item {\ptt Bound}, \bold{56}, 77, 79, 80 - \item {\ptt bound_hyp_subst_tac}, \bold{91} - \item {\ptt br}, \bold{10} - \item {\ptt BREADTH_FIRST}, \bold{29} - \item {\ptt brs}, \bold{10} - \item {\ptt bw}, \bold{11} - \item {\ptt bws}, \bold{11} - \item {\ptt by}, \bold{7}, 9, 10, 14 - \item {\ptt byev}, \bold{7} + \item {\tt ba}, \bold{11} + \item {\tt back}, \bold{9} + \item batch execution, 12 + \item {\tt bd}, \bold{11} + \item {\tt bds}, \bold{11} + \item {\tt be}, \bold{11} + \item {\tt bes}, \bold{11} + \item {\tt BEST_FIRST}, \bold{31}, 32 + \item {\tt Best_tac}, \bold{127} + \item {\tt best_tac}, \bold{126} + \item {\tt beta_conversion}, \bold{43} + \item {\tt bicompose}, \bold{46} + \item {\tt bimatch_tac}, \bold{23} + \item {\tt bind_thm}, \bold{8}, 9, 37 + \item binders, \bold{74} + \item {\tt biresolution}, \bold{45} + \item {\tt biresolve_tac}, \bold{23}, 128 + \item {\tt Blast.depth_tac}, \bold{125} + \item {\tt Blast.trace}, \bold{125} + \item {\tt Blast_tac}, \bold{127} + \item {\tt blast_tac}, \bold{125} + \item {\tt Bound}, \bold{57}, 80, 82, 83 + \item {\tt bound_hyp_subst_tac}, \bold{96} + \item {\tt br}, \bold{11} + \item {\tt BREADTH_FIRST}, \bold{31} + \item {\tt brs}, \bold{11} + \item {\tt bw}, \bold{12} + \item {\tt bws}, \bold{12} + \item {\tt by}, \bold{7}, 9, 10, 15 + \item {\tt byev}, \bold{7} \indexspace - \item {\ptt cd}, \bold{2}, 50 - \item {\ptt cert_axm}, \bold{58} - \item {\ptt CHANGED}, \bold{28} - \item {\ptt chop}, \bold{9}, 13 - \item {\ptt choplev}, \bold{9} + \item {\tt cd}, \bold{3}, 53 + \item {\tt cert_axm}, \bold{59} + \item {\tt CHANGED}, \bold{30} + \item {\tt chop}, \bold{9}, 13 + \item {\tt choplev}, \bold{9} \item claset - \subitem current, 120 - \item {\ptt claset} ML type, 116 + \subitem current, 127 + \item {\tt claset} ML type, 122 \item classes - \subitem context conditions, 48 - \item classical reasoner, 112--122 - \subitem setting up, 122 - \subitem tactics, 118 - \item classical sets, 115 - \item {\ptt ClassicalFun}, 122 - \item {\ptt combination}, \bold{41} - \item {\ptt commit}, \bold{2} - \item {\ptt COMP}, \bold{43} - \item {\ptt compose}, \bold{43} - \item {\ptt compose_tac}, \bold{21} - \item {\ptt compSWrapper}, \bold{118} - \item {\ptt compWrapper}, \bold{118} - \item {\ptt concl_of}, \bold{37} - \item {\ptt COND}, \bold{30} - \item congruence rules, 97 - \item {\ptt Const}, \bold{56}, 79, 88 - \item {\ptt Constant}, 76, 88 - \item constants, \bold{56} - \subitem for translations, 68 - \subitem syntactic, 81 - \item {\ptt contr_tac}, \bold{121} - \item {\ptt could_unify}, \bold{23} - \item {\ptt cterm} ML type, 58 - \item {\ptt cterm_instantiate}, \bold{36} - \item {\ptt cterm_of}, \bold{58} - \item {\ptt ctyp}, \bold{59} - \item {\ptt ctyp_of}, \bold{59} - \item {\ptt cut_facts_tac}, 7, \bold{18}, 92 - \item {\ptt cut_inst_tac}, \bold{18} - \item {\ptt cut_rl} theorem, 19 + \subitem context conditions, 52 + \item classical reasoner, 118--129 + \subitem setting up, 128 + \subitem tactics, 124 + \item classical sets, 122 + \item {\tt ClassicalFun}, 129 + \item {\tt combination}, \bold{44} + \item {\tt commit}, \bold{2} + \item {\tt COMP}, \bold{46} + \item {\tt compose}, \bold{46} + \item {\tt compose_tac}, \bold{22} + \item {\tt compSWrapper}, \bold{124} + \item {\tt compWrapper}, \bold{124} + \item {\tt concl_of}, \bold{39} + \item {\tt COND}, \bold{32} + \item congruence rules, 102 + \item {\tt Const}, \bold{57}, 82, 92 + \item {\tt Constant}, 79, 92 + \item constants, \bold{57} + \subitem for translations, 69 + \subitem syntactic, 84 + \item {\tt contr_tac}, \bold{128} + \item {\tt could_unify}, \bold{24} + \item {\tt CPure} theory, 49 + \item {\tt CPure.thy}, \bold{55} + \item {\tt cterm} ML type, 59 + \item {\tt cterm_instantiate}, \bold{38} + \item {\tt cterm_of}, 7, 13, \bold{59} + \item {\tt ctyp}, \bold{60} + \item {\tt ctyp_of}, \bold{61} + \item {\tt cut_facts_tac}, 7, \bold{19}, 97 + \item {\tt cut_inst_tac}, \bold{19} + \item {\tt cut_rl} theorem, 21 \indexspace - \item {\ptt datatype}, 95 - \item {\ptt Deepen_tac}, \bold{121} - \item {\ptt deepen_tac}, \bold{120} - \item {\ptt defer_tac}, \bold{18} - \item definitions, \see{rewriting, meta-level}{0}, 18, \bold{48} + \item {\tt datatype}, 100 + \item {\tt Deepen_tac}, \bold{127} + \item {\tt deepen_tac}, \bold{126} + \item {\tt defer_tac}, \bold{20} + \item definitions, \see{rewriting, meta-level}{1}, 20, \bold{51} \subitem unfolding, 7, 8 - \item {\ptt delcongs}, 97 - \item {\ptt deleqcongs}, 97, 100 - \item {\ptt delete_tmpfiles}, \bold{49} - \item delimiters, \bold{65}, 67, 68, 70 - \item {\ptt delrules}, \bold{116} - \item {\ptt Delsimps}, \bold{95}, 99, 100 - \item {\ptt delsimps}, 99, 100 - \item {\ptt dependent_tr'}, 86, \bold{88} - \item {\ptt DEPTH_FIRST}, \bold{29} - \item {\ptt DEPTH_SOLVE}, \bold{29} - \item {\ptt DEPTH_SOLVE_1}, \bold{29} - \item {\ptt depth_tac}, \bold{120} - \item {\ptt Deriv.drop}, \bold{45} - \item {\ptt Deriv.linear}, \bold{45} - \item {\ptt Deriv.size}, \bold{45} - \item {\ptt Deriv.tree}, \bold{45} - \item {\ptt dest_eq}, \bold{92} - \item {\ptt dest_state}, \bold{37} - \item destruct-resolution, 15 - \item {\ptt DETERM}, \bold{30} - \item discrimination nets, \bold{22} - \item {\ptt dmatch_tac}, \bold{16} - \item {\ptt dres_inst_tac}, \bold{17} - \item {\ptt dresolve_tac}, \bold{15} - \item {\ptt dtac}, \bold{17} - \item {\ptt dummyT}, 79, 80, 89 + \item {\tt delcongs}, 103 + \item {\tt deleqcongs}, 103, 105 + \item {\tt delete_tmpfiles}, \bold{53} + \item delimiters, \bold{66}, 69, 70, 72 + \item {\tt delrules}, \bold{123} + \item {\tt Delsimps}, \bold{100}, 105, 106 + \item {\tt delsimps}, 105, 106 + \item {\tt dependent_tr'}, 90, \bold{92} + \item {\tt DEPTH_FIRST}, \bold{31} + \item {\tt DEPTH_SOLVE}, \bold{31} + \item {\tt DEPTH_SOLVE_1}, \bold{31} + \item {\tt depth_tac}, \bold{126} + \item {\tt Deriv.drop}, \bold{48} + \item {\tt Deriv.linear}, \bold{48} + \item {\tt Deriv.size}, \bold{48} + \item {\tt Deriv.tree}, \bold{48} + \item {\tt dest_eq}, \bold{97} + \item {\tt dest_state}, \bold{39} + \item destruct-resolution, 17 + \item {\tt DETERM}, \bold{32} + \item discrimination nets, \bold{24} + \item {\tt dmatch_tac}, \bold{17} + \item {\tt dres_inst_tac}, \bold{18} + \item {\tt dresolve_tac}, \bold{17} + \item {\tt dtac}, \bold{19} + \item {\tt dummyT}, 82, 83, 93 \indexspace - \item elim-resolution, 15 - \item {\ptt ematch_tac}, \bold{16} - \item {\ptt empty} constant, 85 - \item {\ptt empty_cs}, \bold{116} - \item {\ptt empty_ss}, 100, 110 - \item {\ptt eq_assume_tac}, \bold{16}, 116 - \item {\ptt eq_assumption}, \bold{43} - \item {\ptt eq_mp_tac}, \bold{121} - \item {\ptt eq_reflection} theorem, \bold{92}, 107 - \item {\ptt eq_thm}, \bold{30} - \item {\ptt equal_elim}, \bold{41} - \item {\ptt equal_intr}, \bold{40} - \item equality, 90--93 - \item {\ptt eres_inst_tac}, \bold{17} - \item {\ptt eresolve_tac}, \bold{15} - \item {\ptt eta_contract}, \bold{4}, 83 - \item {\ptt etac}, \bold{17} - \item {\ptt EVERY}, \bold{27} - \item {\ptt EVERY'}, 33 - \item {\ptt EVERY1}, \bold{33} + \item elim-resolution, 16 + \item {\tt ematch_tac}, \bold{17} + \item {\tt empty} constant, 88 + \item {\tt empty_cs}, \bold{122} + \item {\tt empty_ss}, 105, 116 + \item {\tt eq_assume_tac}, \bold{17}, 122 + \item {\tt eq_assumption}, \bold{45} + \item {\tt eq_mp_tac}, \bold{128} + \item {\tt eq_reflection} theorem, \bold{97}, 113 + \item {\tt eq_thm}, \bold{32} + \item {\tt equal_elim}, \bold{43} + \item {\tt equal_intr}, \bold{43} + \item equality, 95--98 + \item {\tt eres_inst_tac}, \bold{18} + \item {\tt eresolve_tac}, \bold{16} + \item {\tt eta_contract}, \bold{4}, 86 + \item {\tt etac}, \bold{19} + \item {\tt EVERY}, \bold{29} + \item {\tt EVERY'}, 35 + \item {\tt EVERY1}, \bold{35} \item examples - \subitem of logic definitions, 73 - \subitem of macros, 85, 86 - \subitem of mixfix declarations, 70 - \subitem of simplification, 101 - \subitem of translations, 88 + \subitem of logic definitions, 76 + \subitem of macros, 88, 89 + \subitem of mixfix declarations, 71 + \subitem of simplification, 106 + \subitem of translations, 92 \item exceptions \subitem printing of, 4 - \item {\ptt expandshort} shell script, 5 - \item {\ptt exportML}, \bold{2} - \item {\ptt extensional}, \bold{41} + \item {\tt exit}, \bold{2} + \item {\tt extensional}, \bold{44} \indexspace - \item {\ptt fa}, \bold{11} - \item {\ptt Fast_tac}, \bold{121} - \item {\ptt fast_tac}, \bold{119} - \item {\ptt fd}, \bold{11} - \item {\ptt fds}, \bold{11} - \item {\ptt fe}, \bold{11} - \item {\ptt fes}, \bold{11} + \item {\tt fa}, \bold{11} + \item {\tt Fast_tac}, \bold{127} + \item {\tt fast_tac}, \bold{126} + \item {\tt fd}, \bold{11} + \item {\tt fds}, \bold{11} + \item {\tt fe}, \bold{11} + \item {\tt fes}, \bold{11} \item files - \subitem reading, 2, 49 - \item {\ptt filt_resolve_tac}, \bold{23} - \item {\ptt FILTER}, \bold{28} - \item {\ptt filter_goal}, \bold{14} - \item {\ptt filter_thms}, \bold{23} - \item {\ptt findE}, \bold{9} - \item {\ptt findEs}, \bold{9} - \item {\ptt findI}, \bold{8} - \item {\ptt finish_html}, \bold{54} - \item {\ptt FIRST}, \bold{27} - \item {\ptt FIRST'}, 33 - \item {\ptt FIRST1}, \bold{33} - \item {\ptt FIRSTGOAL}, \bold{32} - \item flex-flex constraints, 20, 37, 44 - \item {\ptt flexflex_rule}, \bold{44} - \item {\ptt flexflex_tac}, \bold{21} - \item {\ptt fold_goals_tac}, \bold{19} - \item {\ptt fold_tac}, \bold{19} - \item {\ptt forall_elim}, \bold{42} - \item {\ptt forall_elim_list}, \bold{42} - \item {\ptt forall_elim_var}, \bold{42} - \item {\ptt forall_elim_vars}, \bold{42} - \item {\ptt forall_intr}, \bold{41} - \item {\ptt forall_intr_frees}, \bold{42} - \item {\ptt forall_intr_list}, \bold{42} - \item {\ptt force_strip_shyps}, \bold{37} - \item {\ptt forw_inst_tac}, \bold{17} - \item forward proof, 15, 16, 35 - \item {\ptt forward_tac}, \bold{16} - \item {\ptt fr}, \bold{11} - \item {\ptt Free}, \bold{56}, 79 - \item {\ptt freezeT}, \bold{42} - \item {\ptt frs}, \bold{11} - \item {\ptt Full_simp_tac}, \bold{94}, 100 - \item {\ptt full_simp_tac}, \bold{99}, 100 - \item {\ptt fun} type, 59 - \item function applications, \bold{56} + \subitem reading, 3, 52 + \item {\tt filt_resolve_tac}, \bold{24} + \item {\tt FILTER}, \bold{30} + \item {\tt filter_goal}, \bold{15} + \item {\tt filter_thms}, \bold{25} + \item {\tt findE}, \bold{9} + \item {\tt findEs}, \bold{9} + \item {\tt findI}, \bold{9} + \item {\tt FIRST}, \bold{29} + \item {\tt FIRST'}, 35 + \item {\tt FIRST1}, \bold{35} + \item {\tt FIRSTGOAL}, \bold{34} + \item flex-flex constraints, 22, 39, 47 + \item {\tt flexflex_rule}, \bold{47} + \item {\tt flexflex_tac}, \bold{22} + \item {\tt fold_goals_tac}, \bold{20} + \item {\tt fold_tac}, \bold{20} + \item {\tt forall_elim}, \bold{44} + \item {\tt forall_elim_list}, \bold{44} + \item {\tt forall_elim_var}, \bold{44} + \item {\tt forall_elim_vars}, \bold{45} + \item {\tt forall_intr}, \bold{44} + \item {\tt forall_intr_frees}, \bold{44} + \item {\tt forall_intr_list}, \bold{44} + \item {\tt force_strip_shyps}, \bold{40} + \item {\tt forw_inst_tac}, \bold{18} + \item forward proof, 17, 37 + \item {\tt forward_tac}, \bold{17} + \item {\tt fr}, \bold{11} + \item {\tt Free}, \bold{57}, 82 + \item {\tt freezeT}, \bold{45} + \item {\tt frs}, \bold{11} + \item {\tt Full_simp_tac}, \bold{100}, 105 + \item {\tt full_simp_tac}, 105, \bold{106} + \item {\tt fun} type, 60 + \item function applications, \bold{57} \indexspace - \item {\ptt get_axiom}, \bold{51} - \item {\ptt get_thm}, \bold{51} - \item {\ptt getgoal}, \bold{14} - \item {\ptt gethyps}, \bold{14}, 31 - \item {\ptt goal}, \bold{7}, 13 - \item {\ptt goals_limit}, \bold{10} - \item {\ptt goalw}, \bold{7} - \item {\ptt goalw_cterm}, \bold{7} + \item {\tt get_axiom}, \bold{55} + \item {\tt get_thm}, \bold{55} + \item {\tt getgoal}, \bold{15} + \item {\tt gethyps}, \bold{15}, 33 + \item {\tt goal}, \bold{7}, 13 + \item {\tt goals_limit}, \bold{10} + \item {\tt goalw}, \bold{7} + \item {\tt goalw_cterm}, \bold{7} \indexspace - \item {\ptt has_fewer_prems}, \bold{30} - \item HTML, \bold{53} - \item {\ptt hyp_subst_tac}, \bold{91} - \item {\ptt hyp_subst_tacs}, \bold{122} - \item {\ptt HypsubstFun}, 92, 122 + \item {\tt has_fewer_prems}, \bold{32} + \item {\tt hyp_subst_tac}, \bold{96} + \item {\tt hyp_subst_tacs}, \bold{129} + \item {\tt HypsubstFun}, 97, 129 \indexspace - \item {\ptt id} nonterminal, \bold{65}, 77, 84 - \item identifiers, 65 - \item {\ptt idt} nonterminal, 83 - \item {\ptt idts} nonterminal, \bold{65}, 72 - \item {\ptt IF_UNSOLVED}, \bold{30} - \item {\ptt iff_reflection} theorem, 107 - \item {\ptt imp_intr} theorem, \bold{92} - \item {\ptt implies_elim}, \bold{40} - \item {\ptt implies_elim_list}, \bold{40} - \item {\ptt implies_intr}, \bold{40} - \item {\ptt implies_intr_hyps}, \bold{40} - \item {\ptt implies_intr_list}, \bold{40} - \item {\ptt incr_boundvars}, \bold{57}, 88 - \item {\ptt indexname} ML type, 56, 66 - \item infixes, \bold{71} - \item {\ptt init_html}, \bold{54} - \item {\ptt insert} constant, 85 - \item {\ptt inst_step_tac}, \bold{120} - \item {\ptt instantiate}, \bold{42} - \item instantiation, 16, 36, 42 - \item {\ptt INTLEAVE}, \bold{26}, 27 - \item {\ptt INTLEAVE'}, 33 - \item {\ptt invoke_oracle}, \bold{60} - \item {\ptt is} nonterminal, 85 + \item {\tt id} nonterminal, \bold{66}, 80, 87 + \item identifiers, 66 + \item {\tt idt} nonterminal, 86 + \item {\tt idts} nonterminal, \bold{66}, 74 + \item {\tt IF_UNSOLVED}, \bold{32} + \item {\tt iff_reflection} theorem, 113 + \item {\tt imp_intr} theorem, \bold{97} + \item {\tt implies_elim}, \bold{43} + \item {\tt implies_elim_list}, \bold{43} + \item {\tt implies_intr}, \bold{42} + \item {\tt implies_intr_hyps}, \bold{43} + \item {\tt implies_intr_list}, \bold{42} + \item {\tt incr_boundvars}, \bold{58}, 92 + \item {\tt indexname} ML type, 57, 67 + \item infixes, \bold{73} + \item {\tt insert} constant, 88 + \item {\tt inst_step_tac}, \bold{127} + \item {\tt instance} section, 51 + \item {\tt instantiate}, \bold{45} + \item instantiation, 17, 38, 45 + \item {\tt INTLEAVE}, \bold{28}, 30 + \item {\tt INTLEAVE'}, 35 + \item {\tt invoke_oracle}, \bold{61} + \item {\tt is} nonterminal, 88 \indexspace - \item {\ptt joinrules}, \bold{121} + \item {\tt joinrules}, \bold{128} \indexspace - \item {\ptt keep_derivs}, \bold{45} + \item {\tt keep_derivs}, \bold{48} \indexspace - \item $\lambda$-abstractions, 22, \bold{56} - \item $\lambda$-calculus, 39, 41, 65 - \item {\ptt lessb}, \bold{22} - \item lexer, \bold{65} - \item {\ptt lift_rule}, \bold{44} - \item lifting, 44 - \item {\ptt loadpath}, \bold{49} - \item {\ptt logic} class, 63, 65, 69 - \item {\ptt logic} nonterminal, \bold{63} - \item {\ptt Logic.auto_rename}, \bold{20} - \item {\ptt Logic.set_rename_prefix}, \bold{20} - \item {\ptt loose_bnos}, \bold{57}, 89 + \item $\lambda$-abstractions, 24, \bold{57} + \item $\lambda$-calculus, 42, 43, 66 + \item {\tt lessb}, \bold{23} + \item lexer, \bold{66} + \item {\tt lift_rule}, \bold{46} + \item lifting, 46 + \item {\tt loadpath}, \bold{53} + \item {\tt logic} class, 66, 71 + \item {\tt logic} nonterminal, \bold{66} + \item {\tt Logic.auto_rename}, \bold{21} + \item {\tt Logic.set_rename_prefix}, \bold{21} + \item {\tt loose_bnos}, \bold{58}, 93 \indexspace - \item macros, 81--86 - \item {\ptt make-all} shell script, 5 - \item {\ptt make_elim}, \bold{37}, 117 - \item {\ptt make_html}, \bold{54} - \item {\ptt Match} exception, 88, 92 - \item {\ptt match_tac}, \bold{16}, 116 - \item {\ptt max_pri}, 63, \bold{69} - \item {\ptt merge_ss}, 100 - \item {\ptt merge_theories}, \bold{52}, \fnote{75} - \item meta-assumptions, 31, 39, 40, 43 - \subitem printing of, 3 - \item meta-equality, 39--41 - \item meta-implication, 39, 40 - \item meta-quantifiers, 40, 41 - \item meta-rewriting, 7, 11, 12, \bold{18}, - \seealso{tactics, theorems}{123} - \subitem in terms, 44 - \subitem in theorems, 35 - \item meta-rules, \see{meta-rules}{0}, 39--44 - \item {\ptt METAHYPS}, 14, \bold{31} - \item mixfix declarations, 47, 68--72 - \item {\ptt mk_case_split_tac}, \bold{111} - \item {\ptt ML} section, 48, 87, 89 - \item {\ptt mp} theorem, \bold{122} - \item {\ptt mp_tac}, \bold{121} - \item {\ptt MRL}, \bold{35} - \item {\ptt MRS}, \bold{35} + \item macros, 84--90 + \item {\tt make_elim}, \bold{39}, 123 + \item {\tt Match} exception, 91, 97 + \item {\tt match_tac}, \bold{17}, 122 + \item {\tt max_pri}, 64, \bold{70} + \item {\tt merge_ss}, 105 + \item {\tt merge_theories}, \bold{55} + \item meta-assumptions, 33, 41, 42, 45 + \subitem printing of, 4 + \item meta-equality, 41, 43 + \item meta-implication, 41, 42 + \item meta-quantifiers, 42, 44 + \item meta-rewriting, 7, 12, 13, \bold{20}, + \seealso{tactics, theorems}{130} + \subitem in terms, 47 + \subitem in theorems, 38 + \item meta-rules, \see{meta-rules}{1}, 41--47 + \item {\tt METAHYPS}, 15, \bold{33} + \item mixfix declarations, 50, 69--74 + \item {\tt mk_case_split_tac}, \bold{116} + \item {\tt ML} section, 51, 91, 93 + \item {\tt mp} theorem, \bold{129} + \item {\tt mp_tac}, \bold{128} + \item {\tt MRL}, \bold{37} + \item {\tt MRS}, \bold{37} \indexspace - \item name tokens, \bold{65} - \item {\ptt net_bimatch_tac}, \bold{23} - \item {\ptt net_biresolve_tac}, \bold{23} - \item {\ptt net_match_tac}, \bold{23} - \item {\ptt net_resolve_tac}, \bold{22} - \item {\ptt no_tac}, \bold{27} - \item {\ptt None}, \bold{24} - \item {\ptt not_elim} theorem, \bold{122} - \item {\ptt nprems_of}, \bold{37} - \item numerals, 65 + \item name tokens, \bold{66} + \item {\tt net_bimatch_tac}, \bold{24} + \item {\tt net_biresolve_tac}, \bold{24} + \item {\tt net_match_tac}, \bold{24} + \item {\tt net_resolve_tac}, \bold{24} + \item {\tt no_tac}, \bold{30} + \item {\tt None}, \bold{26} + \item {\tt not_elim} theorem, \bold{129} + \item {\tt nprems_of}, \bold{39} + \item numerals, 66 \indexspace - \item {\ptt o} type, 73 - \item {\ptt op} symbol, 71 - \item {\ptt option} ML type, 24 - \item oracles, 60--61 - \item {\ptt ORELSE}, \bold{26}, 27, 28, 32 - \item {\ptt ORELSE'}, 33 + \item {\tt o} type, 76 + \item {\tt op} symbol, 73 + \item {\tt option} ML type, 26 + \item oracles, 61--62 + \item {\tt ORELSE}, \bold{28}, 30, 34 + \item {\tt ORELSE'}, 35 \indexspace \item parameters - \subitem removing unused, 20 - \subitem renaming, 11, 19, 44 - \item {\ptt parents_of}, \bold{52} - \item parse trees, 76 - \item {\ptt parse_ast_translation}, 87 - \item {\ptt parse_rules}, 83 - \item {\ptt parse_translation}, 87 - \item {\ptt pause_tac}, \bold{24} - \item Poly/{\ML} compiler, 1, 2, 4, 50 - \item {\ptt pop_proof}, \bold{13} - \item {\ptt pr}, \bold{10} - \item {\ptt premises}, \bold{7}, 13 - \item {\ptt prems_of}, \bold{37} - \item {\ptt prems_of_ss}, 100 - \item pretty printing, 68, 71, 85 - \item {\ptt Pretty.setdepth}, \bold{3} - \item {\ptt Pretty.setmargin}, \bold{3} - \item {\ptt PRIMITIVE}, \bold{24} - \item {\ptt primrec}, 95 - \item {\ptt prin}, 4, \bold{14} - \item {\ptt print_ast_translation}, 87 - \item {\ptt print_cs}, \bold{116} - \item {\ptt print_depth}, \bold{3} - \item {\ptt print_exn}, \bold{4}, 34 - \item {\ptt print_goals}, \bold{35} - \item {\ptt print_rules}, 83 - \item {\ptt print_syntax}, \bold{66} - \item {\ptt print_tac}, \bold{24} - \item {\ptt print_theory}, \bold{52} - \item {\ptt print_thm}, \bold{35} - \item {\ptt print_translation}, 87 + \subitem removing unused, 22 + \subitem renaming, 12, 21, 47 + \item {\tt parents_of}, \bold{56} + \item parse trees, 79 + \item {\tt parse_ast_translation}, 91 + \item {\tt parse_rules}, 86 + \item {\tt parse_translation}, 91 + \item {\tt pause_tac}, \bold{26} + \item Poly/{\ML} compiler, 5 + \item {\tt pop_proof}, \bold{14} + \item {\tt pr}, \bold{10} + \item {\tt premises}, \bold{7}, 14 + \item {\tt prems_of}, \bold{39} + \item {\tt prems_of_ss}, 105 + \item pretty printing, 70, 72--73, 89 + \item {\tt Pretty.setdepth}, \bold{3} + \item {\tt Pretty.setmargin}, \bold{3} + \item {\tt PRIMITIVE}, \bold{25} + \item {\tt primrec}, 100 + \item {\tt prin}, 5, \bold{14} + \item print mode, 50, 93 + \item print modes, 74--75 + \item {\tt print_ast_translation}, 91 + \item {\tt print_cs}, \bold{122} + \item {\tt print_depth}, \bold{4} + \item {\tt print_exn}, \bold{5}, 36 + \item {\tt print_goals}, \bold{37} + \item {\tt print_mode}, 74 + \item {\tt print_modes}, 69 + \item {\tt print_rules}, 86 + \item {\tt print_syntax}, \bold{56}, \bold{68} + \item {\tt print_tac}, \bold{26} + \item {\tt print_theory}, \bold{56} + \item {\tt print_thm}, \bold{37} + \item {\tt print_translation}, 91 \item printing control, 3--4 - \item {\ptt printyp}, \bold{14} - \item priorities, 62, \bold{69} - \item priority grammars, 62--63 - \item {\ptt prlev}, \bold{10} - \item {\ptt prlim}, \bold{10} - \item productions, 62, 67, 68 - \subitem copy, \bold{67}, 68, 77 - \item {\ptt proof} ML type, 13 - \item proof objects, 44--45 + \item {\tt printyp}, \bold{14} + \item priorities, 63, \bold{70} + \item priority grammars, 63--64 + \item {\tt prlev}, \bold{10} + \item {\tt prlim}, \bold{10} + \item productions, 63, 69, 70 + \subitem copy, \bold{69}, 70, 81 + \item {\tt proof} ML type, 14 + \item proof objects, 47--48 \item proof state, 6 - \subitem printing of, 9 - \item {\ptt proof_timing}, \bold{10} - \item proofs, 6--14 - \subitem inspecting the state, 14 - \subitem managing multiple, 12 - \subitem saving and restoring, 13 + \subitem printing of, 10 + \item {\tt proof_timing}, \bold{10} + \item proofs, 6--15 + \subitem inspecting the state, 15 + \subitem managing multiple, 13 + \subitem saving and restoring, 14 \subitem stacking, 13 \subitem starting, 6 \subitem timing, 10 - \item {\ptt PROP} symbol, 64 - \item {\ptt prop} nonterminal, \bold{63}, 73 - \item {\ptt prop} type, 59, 63 - \item {\ptt prove_goal}, 10, \bold{12} - \item {\ptt prove_goalw}, \bold{12} - \item {\ptt prove_goalw_cterm}, \bold{12} - \item {\ptt prth}, \bold{34} - \item {\ptt prthq}, \bold{35} - \item {\ptt prths}, \bold{34} - \item {\ptt prune_params_tac}, \bold{21} - \item {\ptt Pure} theory, 46, 63, 67 - \item {\ptt pure_thy}, \bold{52} - \item {\ptt push_proof}, \bold{13} - \item {\ptt pwd}, \bold{2} + \item {\tt PROP} symbol, 65 + \item {\tt prop} nonterminal, \bold{64}, 76 + \item {\tt prop} type, 60, 66 + \item {\tt prove_goal}, 10, \bold{12} + \item {\tt prove_goalw}, \bold{13} + \item {\tt prove_goalw_cterm}, \bold{13} + \item {\tt prth}, \bold{36} + \item {\tt prthq}, \bold{37} + \item {\tt prths}, \bold{37} + \item {\tt prune_params_tac}, \bold{22} + \item {\tt Pure} theory, 49, 64, 68 + \item {\tt Pure.thy}, \bold{55} + \item {\tt push_proof}, \bold{14} + \item {\tt pwd}, \bold{3} \indexspace - \item {\ptt qed}, 8, \bold{8} - \item {\ptt qed_goal}, \bold{12} - \item {\ptt qed_goalw}, \bold{12} - \item quantifiers, 72 - \item {\ptt quit}, \bold{2} + \item {\tt qed}, \bold{8}, 9 + \item {\tt qed_goal}, \bold{13} + \item {\tt qed_goalw}, \bold{13} + \item quantifiers, 74 + \item {\tt quit}, \bold{2} \indexspace - \item {\ptt read}, \bold{14} - \item {\ptt read_axm}, \bold{58} - \item {\ptt read_cterm}, \bold{58} - \item {\ptt read_instantiate}, \bold{36} - \item {\ptt read_instantiate_sg}, \bold{36} + \item {\tt read}, \bold{14} + \item {\tt read_axm}, \bold{59} + \item {\tt read_cterm}, \bold{59} + \item {\tt read_instantiate}, \bold{38} + \item {\tt read_instantiate_sg}, \bold{38} \item reading - \subitem axioms, \see{{\tt assume_ax}}{46} + \subitem axioms, \see{{\tt assume_ax}}{49} \subitem goals, \see{proofs, starting}{6} - \item {\ptt reflexive}, \bold{41} - \item {\ptt ren}, \bold{11} - \item {\ptt rename_last_tac}, \bold{20} - \item {\ptt rename_params_rule}, \bold{44} - \item {\ptt rename_tac}, \bold{20} - \item {\ptt rep_cterm}, \bold{58} - \item {\ptt rep_ctyp}, \bold{59} - \item {\ptt rep_ss}, 100 - \item {\ptt rep_thm}, \bold{37} - \item {\ptt REPEAT}, \bold{27, 28} - \item {\ptt REPEAT1}, \bold{27} - \item {\ptt REPEAT_DETERM}, \bold{27} - \item {\ptt REPEAT_FIRST}, \bold{32} - \item {\ptt REPEAT_SOME}, \bold{32} - \item {\ptt res_inst_tac}, \bold{17}, 21 - \item reserved words, 65, 86 - \item resolution, 35, 43 - \subitem tactics, 15 - \subitem without lifting, 43 - \item {\ptt resolve_tac}, \bold{15}, 116 - \item {\ptt restore_proof}, \bold{13} - \item {\ptt result}, \bold{8}, 14, 52 - \item {\ptt rev_mp} theorem, \bold{92} - \item rewrite rules, 96--97 - \subitem permutative, 104--107 - \item {\ptt rewrite_cterm}, \bold{44} - \item {\ptt rewrite_goals_rule}, \bold{36} - \item {\ptt rewrite_goals_tac}, \bold{19}, 36 - \item {\ptt rewrite_rule}, \bold{36} - \item {\ptt rewrite_tac}, 8, \bold{19} + \item {\tt reflexive}, \bold{43} + \item {\tt ren}, \bold{12} + \item {\tt rename_last_tac}, \bold{21} + \item {\tt rename_params_rule}, \bold{46} + \item {\tt rename_tac}, \bold{21} + \item {\tt rep_cterm}, \bold{59} + \item {\tt rep_ctyp}, \bold{61} + \item {\tt rep_ss}, 105 + \item {\tt rep_thm}, \bold{40} + \item {\tt REPEAT}, \bold{29, 30} + \item {\tt REPEAT1}, \bold{29} + \item {\tt REPEAT_DETERM}, \bold{29} + \item {\tt REPEAT_FIRST}, \bold{34} + \item {\tt REPEAT_SOME}, \bold{34} + \item {\tt res_inst_tac}, \bold{18}, 22 + \item reserved words, 66, 89 + \item {\tt reset}, 3 + \item resolution, 37, 45 + \subitem tactics, 16 + \subitem without lifting, 46 + \item {\tt resolve_tac}, \bold{16}, 122 + \item {\tt restore_proof}, \bold{14} + \item {\tt result}, \bold{8}, 15, 55 + \item {\tt rev_mp} theorem, \bold{97} + \item rewrite rules, 101--102 + \subitem permutative, 110--113 + \item {\tt rewrite_cterm}, \bold{47} + \item {\tt rewrite_goals_rule}, \bold{38} + \item {\tt rewrite_goals_tac}, \bold{20}, 38 + \item {\tt rewrite_rule}, \bold{38} + \item {\tt rewrite_tac}, 8, \bold{20} \item rewriting - \subitem object-level, \see{simplification}{0} - \subitem ordered, 104 - \subitem syntactic, 81--86 - \item {\ptt rewtac}, \bold{18} - \item {\ptt RL}, \bold{35} - \item {\ptt RLN}, \bold{35} - \item {\ptt rotate_proof}, \bold{13} - \item {\ptt rotate_tac}, \bold{20} - \item {\ptt RS}, \bold{35} - \item {\ptt RSN}, \bold{35} - \item {\ptt rtac}, \bold{17} - \item {\ptt rule_by_tactic}, 21, \bold{37} + \subitem object-level, \see{simplification}{1} + \subitem ordered, 110 + \subitem syntactic, 84--90 + \item {\tt rewtac}, \bold{19} + \item {\tt RL}, \bold{37} + \item {\tt RLN}, \bold{37} + \item {\tt rotate_proof}, \bold{14} + \item {\tt rotate_tac}, \bold{22} + \item {\tt RS}, \bold{37} + \item {\tt RSN}, \bold{37} + \item {\tt rtac}, \bold{19} + \item {\tt rule_by_tactic}, 22, \bold{39} \item rules - \subitem converting destruction to elimination, 37 + \subitem converting destruction to elimination, 39 \indexspace - \item {\ptt safe_asm_full_simp_tac}, \bold{98}, 100 - \item {\ptt safe_step_tac}, 117, \bold{120} - \item {\ptt safe_tac}, \bold{120} - \item {\ptt save_proof}, \bold{13} + \item {\tt safe_asm_full_simp_tac}, \bold{104}, 105 + \item {\tt safe_step_tac}, 123, \bold{126} + \item {\tt safe_tac}, \bold{127} + \item {\tt save_proof}, \bold{14} \item saving your work, \bold{1} - \item search, 26 - \subitem tacticals, 28--30 - \item {\ptt SELECT_GOAL}, 19, \bold{30} - \item {\ptt Sequence.seq} ML type, 23 - \item sequences (lazy lists), \bold{24} - \item sequent calculus, 112 + \item search, 28 + \subitem tacticals, 30--32 + \item {\tt SELECT_GOAL}, 20, \bold{33} + \item {\tt Sequence.seq} ML type, 25 + \item sequences (lazy lists), \bold{26} + \item sequent calculus, 118 \item sessions, 1--5 - \item {\ptt set_current_thy}, 96 - \item {\ptt set_oracle}, \bold{60} - \item {\ptt setloop}, 99, 100 - \item {\ptt setmksimps}, 96, 100, 110 - \item {\ptt setSolver}, 98, 100, 110 - \item {\ptt setSSolver}, 98, 100, 110 - \item {\ptt setsubgoaler}, 98, 100, 110 - \item {\ptt setSWrapper}, \bold{118} - \item {\ptt setWrapper}, \bold{118} - \item shell scripts, \bold{4} + \item {\tt set}, 3 + \item {\tt set_current_thy}, 101 + \item {\tt set_oracle}, \bold{61} + \item {\tt setloop}, 104, 105 + \item {\tt setmksimps}, 102, 105, 116 + \item {\tt setSolver}, 104, 105, 116 + \item {\tt setSSolver}, 104, 105, 116 + \item {\tt setsubgoaler}, 103, 105, 116 + \item {\tt setSWrapper}, \bold{124} + \item {\tt setWrapper}, \bold{124} \item shortcuts - \subitem for tactics, 17 + \subitem for tactics, 18 \subitem for {\tt by} commands, 10 - \item {\ptt show_brackets}, \bold{3} - \item {\ptt show_hyps}, \bold{3} - \item {\ptt show_sorts}, \bold{3}, 80 - \item {\ptt show_types}, \bold{3}, 80, 83, 89 - \item {\ptt Sign.cterm_of}, 7, 12 - \item {\ptt Sign.sg} ML type, 46 - \item {\ptt Sign.string_of_term}, \bold{58} - \item {\ptt Sign.string_of_typ}, \bold{59} - \item {\ptt sign_of}, 7, 12, \bold{52} - \item signatures, \bold{46}, 53, 57--59 - \item {\ptt Simp_tac}, \bold{94}, 100 - \item {\ptt simp_tac}, \bold{99}, 100 - \item simplification, 94--111 - \subitem from classical reasoner, 117 - \subitem setting up, 107 - \subitem tactics, 99 - \item simplification sets, 96 + \item {\tt show_brackets}, \bold{4} + \item {\tt show_hyps}, \bold{4} + \item {\tt show_sorts}, \bold{4}, 83 + \item {\tt show_types}, \bold{4}, 83, 86, 93 + \item {\tt Sign.sg} ML type, 49 + \item {\tt Sign.string_of_term}, \bold{59} + \item {\tt Sign.string_of_typ}, \bold{60} + \item {\tt sign_of}, 7, 13, \bold{56} + \item signatures, \bold{49}, 56, 58, 59, 61 + \item {\tt Simp_tac}, \bold{99}, 105 + \item {\tt simp_tac}, 105, \bold{106} + \item simplification, 99--117 + \subitem from classical reasoner, 124 + \subitem setting up, 113 + \subitem tactics, 106 + \item simplification sets, 101 \item simpset - \subitem current, 94, 96, 99 - \item {\ptt simpset}, 100 - \item {\ptt simpset} ML type, 100 - \item {\tt simpset} ML value, 96 - \item {\ptt simpset_of}, 101 - \item {\ptt size_of_thm}, 29, \bold{30}, 122 - \item {\ptt sizef}, \bold{122} - \item {\ptt slow_best_tac}, \bold{119} - \item {\ptt slow_step_tac}, 117, \bold{120} - \item {\ptt slow_tac}, \bold{119} - \item {\ptt Some}, \bold{24} - \item {\ptt SOMEGOAL}, \bold{32} - \item {\ptt sort} nonterminal, \bold{65} - \item sort constraints, 64 - \item sort hypotheses, 37 + \subitem current, 99, 101, 106 + \item {\tt simpset}, 105 + \item {\tt simpset} ML type, 105 + \item {\tt simpset} ML value, 101 + \item {\tt simpset_of}, 106 + \item {\tt size_of_thm}, 31, \bold{32}, 129 + \item {\tt sizef}, \bold{129} + \item {\tt slow_best_tac}, \bold{126} + \item {\tt slow_step_tac}, 123, \bold{127} + \item {\tt slow_tac}, \bold{126} + \item {\tt Some}, \bold{26} + \item {\tt SOMEGOAL}, \bold{34} + \item {\tt sort} nonterminal, \bold{66} + \item sort constraints, 65 + \item sort hypotheses, 40 \item sorts - \subitem printing of, 3 - \item {\ptt split_tac}, \bold{111} - \item {\ptt ssubst} theorem, \bold{90} - \item {\ptt stac}, \bold{91} - \item stamps, 37, \bold{46}, 53 - \item {\ptt stamps_of_thm}, \bold{37} - \item {\ptt stamps_of_thy}, \bold{53} - \item {\ptt standard}, \bold{36} + \subitem printing of, 4 + \item {\tt split_tac}, \bold{116} + \item {\tt ssubst} theorem, \bold{95} + \item {\tt stac}, \bold{96} + \item stamps, 39, \bold{49}, 56 + \item {\tt stamps_of_thm}, \bold{39} + \item {\tt stamps_of_thy}, \bold{56} + \item {\tt standard}, \bold{39} \item starting up, \bold{1} - \item {\ptt STATE}, \bold{24} - \item {\ptt Step_tac}, \bold{121} - \item {\ptt step_tac}, 117, \bold{120} - \item {\ptt store_thm}, \bold{8} - \item {\ptt string_of_cterm}, \bold{58} - \item {\ptt string_of_ctyp}, \bold{59} - \item {\ptt string_of_thm}, \bold{35} - \item strings, 65 - \item {\ptt SUBGOAL}, \bold{24} - \item subgoal module, 6--14 - \item {\ptt subgoal_tac}, \bold{18} - \item {\ptt subgoals_of_brl}, \bold{22} - \item {\ptt subgoals_tac}, \bold{18} - \item {\ptt subst} theorem, 90, \bold{92} + \item {\tt STATE}, \bold{25} + \item {\tt Step_tac}, \bold{127} + \item {\tt step_tac}, 123, \bold{127} + \item {\tt store_thm}, \bold{8} + \item {\tt string_of_cterm}, \bold{59} + \item {\tt string_of_ctyp}, \bold{60} + \item {\tt string_of_thm}, \bold{37} + \item strings, 66 + \item {\tt SUBGOAL}, \bold{25} + \item subgoal module, 6--15 + \item {\tt subgoal_tac}, \bold{19} + \item {\tt subgoals_of_brl}, \bold{23} + \item {\tt subgoals_tac}, \bold{19} + \item {\tt subst} theorem, 95, \bold{97} \item substitution - \subitem rules, 90 - \item {\ptt swap} theorem, \bold{122} - \item {\ptt swap_res_tac}, \bold{121} - \item {\ptt swapify}, \bold{121} - \item {\ptt sym} theorem, 91, \bold{92} - \item {\ptt symmetric}, \bold{41} - \item {\ptt syn_of}, \bold{66} + \subitem rules, 95 + \item {\tt swap} theorem, \bold{129} + \item {\tt swap_res_tac}, \bold{128} + \item {\tt swapify}, \bold{128} + \item {\tt sym} theorem, 96, \bold{97} + \item {\tt symmetric}, \bold{43} + \item {\tt syn_of}, \bold{68} \item syntax - \subitem Pure, 63--68 - \subitem transformations, 76--89 - \item {\ptt syntax} section, 47 - \item {\ptt Syntax.ast} ML type, 76 - \item {\ptt Syntax.print_gram}, \bold{67} - \item {\ptt Syntax.print_syntax}, \bold{66} - \item {\ptt Syntax.print_trans}, \bold{67} - \item {\ptt Syntax.stat_norm_ast}, 84 - \item {\ptt Syntax.syntax} ML type, 66 - \item {\ptt Syntax.test_read}, 70, 84 - \item {\ptt Syntax.trace_norm_ast}, 84 + \subitem Pure, 64--69 + \subitem transformations, 79--93 + \item {\tt syntax} section, 50 + \item {\tt Syntax.ast} ML type, 79 + \item {\tt Syntax.mark_boundT}, 93 + \item {\tt Syntax.print_gram}, \bold{68} + \item {\tt Syntax.print_syntax}, \bold{68} + \item {\tt Syntax.print_trans}, \bold{68} + \item {\tt Syntax.stat_norm_ast}, 88 + \item {\tt Syntax.syntax} ML type, 68 + \item {\tt Syntax.test_read}, 72, 88 + \item {\tt Syntax.trace_norm_ast}, 88 + \item {\tt Syntax.variant_abs'}, 93 \indexspace - \item {\ptt Tactic}, \bold{24} - \item {\ptt tactic} ML type, 15 - \item tacticals, 26--33 - \subitem conditional, 29 - \subitem deterministic, 29 - \subitem for filtering, 28 - \subitem for restriction to a subgoal, 30 - \subitem identities for, 27 - \subitem joining a list of tactics, 26 - \subitem joining tactic functions, 32, 33 - \subitem joining two tactics, 26 - \subitem repetition, 27 - \subitem scanning for subgoals, 31 - \subitem searching, 28, 29 - \item tactics, 15--25 - \subitem assumption, \bold{16}, 17 + \item {\tt tactic} ML type, 16 + \item tacticals, 28--35 + \subitem conditional, 32 + \subitem deterministic, 32 + \subitem for filtering, 30 + \subitem for restriction to a subgoal, 33 + \subitem identities for, 29 + \subitem joining a list of tactics, 29 + \subitem joining tactic functions, 35 + \subitem joining two tactics, 28 + \subitem repetition, 29 + \subitem scanning for subgoals, 34 + \subitem searching, 31 + \item tactics, 16--27 + \subitem assumption, \bold{17}, 18 \subitem commands for applying, 7 - \subitem debugging, 13 - \subitem filtering results of, 28 - \subitem for composition, 21 - \subitem for contradiction, 121 - \subitem for inserting facts, 18 - \subitem for Modus Ponens, 121 - \subitem instantiation, 16 - \subitem matching, 16 - \subitem meta-rewriting, 17, \bold{18} - \subitem primitives for coding, 23 - \subitem resolution, \bold{15}, 17, 21, 22 - \subitem restricting to a subgoal, 30 - \subitem simplification, 99 - \subitem substitution, 90--93 - \subitem tracing, 24 - \item {\ptt tapply}, \bold{23} - \item {\ptt teeinput} shell script, 5 - \item {\ptt TERM}, 58 - \item {\ptt term} ML type, 56, 79 - \item terms, \bold{56} - \subitem certified, \bold{57} - \subitem made from ASTs, 79 - \subitem printing of, 14, 58 + \subitem debugging, 14 + \subitem filtering results of, 30 + \subitem for composition, 22 + \subitem for contradiction, 128 + \subitem for inserting facts, 19 + \subitem for Modus Ponens, 128 + \subitem instantiation, 17 + \subitem matching, 17 + \subitem meta-rewriting, 18, \bold{20} + \subitem primitives for coding, 25 + \subitem resolution, \bold{16}, 18, 23, 24 + \subitem restricting to a subgoal, 33 + \subitem simplification, 106 + \subitem substitution, 95--98 + \subitem tracing, 25 + \item {\tt TERM}, 59 + \item {\tt term} ML type, 57, 82 + \item terms, \bold{57} + \subitem certified, \bold{58} + \subitem made from ASTs, 82 + \subitem printing of, 14, 59 \subitem reading of, 14 - \item {\ptt TFree}, \bold{59} - \item {\ptt THEN}, \bold{26}, 27, 32 - \item {\ptt THEN'}, 33 - \item {\ptt THEN_BEST_FIRST}, \bold{29} - \item theorems, 34--45 - \subitem equality of, 30 - \subitem extracting, 51 + \item {\tt TFree}, \bold{60} + \item {\tt THEN}, \bold{28}, 30, 34 + \item {\tt THEN'}, 35 + \item {\tt THEN_BEST_FIRST}, \bold{31} + \item theorems, 36--48 + \subitem equality of, 32 + \subitem extracting, 55 \subitem extracting proved, 8 - \subitem joining by resolution, 35 - \subitem of pure theory, 19 - \subitem printing of, 34 - \subitem retrieving, 8 - \subitem size of, 30 - \subitem standardizing, 36 + \subitem joining by resolution, 37 + \subitem of pure theory, 20 + \subitem printing of, 36 + \subitem retrieving, 9 + \subitem size of, 32 + \subitem standardizing, 38 \subitem storing, 8 - \subitem taking apart, 37 - \item theories, 46--61 - \subitem axioms of, 51 - \subitem constructing, \bold{52} - \subitem inspecting, \bold{52} - \subitem loading, 49 - \subitem parent, \bold{46} - \subitem pseudo, \bold{50} - \subitem reloading, \bold{50} - \subitem removing, \bold{50} - \subitem theorems of, 51 - \item {\ptt THEORY} exception, 46, 51 - \item {\ptt theory} ML type, 46 - \item {\ptt theory_of_thm}, \bold{37} - \item {\ptt thin_tac}, \bold{20} - \item {\ptt THM} exception, 34, 35, 39, 43 - \item {\ptt thm} ML type, 34 - \item {\ptt thms_containing}, \bold{9} - \item {\ptt thms_of}, \bold{52} - \item {\ptt tid} nonterminal, \bold{65}, 77, 84 - \item {\ptt time_use}, \bold{2} - \item {\ptt time_use_thy}, \bold{49} + \subitem taking apart, 39 + \item theories, 49--62 + \subitem axioms of, 55 + \subitem constructing, \bold{55} + \subitem inspecting, \bold{56} + \subitem loading, 52 + \subitem parent, \bold{49} + \subitem pseudo, \bold{54} + \subitem reloading, \bold{53} + \subitem removing, \bold{54} + \subitem theorems of, 55 + \item {\tt THEORY} exception, 49, 55 + \item {\tt theory} ML type, 49 + \item {\tt theory_of_thm}, \bold{39} + \item {\tt thin_tac}, \bold{22} + \item {\tt THM} exception, 36, 37, 41, 46 + \item {\tt thm} ML type, 36 + \item {\tt thms_containing}, \bold{9} + \item {\tt thms_of}, \bold{56} + \item {\tt tid} nonterminal, \bold{66}, 80, 87 + \item {\tt time_use}, \bold{3} + \item {\tt time_use_thy}, \bold{53} \item timing statistics, 10 - \item {\ptt topthm}, \bold{14} - \item {\ptt tpairs_of}, \bold{37} - \item {\ptt trace_BEST_FIRST}, \bold{29} - \item {\ptt trace_DEPTH_FIRST}, \bold{29} - \item {\ptt trace_goalno_tac}, \bold{32} - \item {\ptt trace_REPEAT}, \bold{27} - \item {\ptt trace_simp}, 95, 102--103 + \item {\tt toggle}, 3 + \item token class, 93 + \item token translations, 93--94 + \item token_translation, 94 + \item {\tt token_translation}, 94 + \item {\tt topthm}, \bold{15} + \item {\tt tpairs_of}, \bold{39} + \item {\tt trace_BEST_FIRST}, \bold{31} + \item {\tt trace_DEPTH_FIRST}, \bold{31} + \item {\tt trace_goalno_tac}, \bold{34} + \item {\tt trace_REPEAT}, \bold{29} + \item {\tt trace_simp}, 100, 108 \item tracing - \subitem of classical prover, 119 - \subitem of macros, 84 - \subitem of searching tacticals, 28, 29 - \subitem of simplification, 95, 102--103 - \subitem of tactics, 24 - \subitem of unification, 38 - \item {\ptt transitive}, \bold{41} - \item translations, 86--89 - \subitem parse, 72, 79 - \subitem parse AST, \bold{77}, 78 - \subitem print, 72 - \subitem print AST, \bold{80} - \item {\ptt translations} section, 82 - \item {\ptt trivial}, \bold{44} - \item {\ptt Trueprop} constant, 73 - \item {\ptt TRY}, \bold{27, 28} - \item {\ptt TRYALL}, \bold{32} - \item {\ptt TVar}, \bold{59} - \item {\ptt tvar} nonterminal, \bold{65, 66}, 77, 84 - \item {\ptt typ} ML type, 58 - \item {\ptt Type}, \bold{59} - \item {\ptt type} nonterminal, \bold{65} - \item {\ptt type} type, 69 - \item type constraints, 65, 72, 80 - \item type constructors, \bold{59} - \item type identifiers, 65 - \item type synonyms, \bold{47} - \item type unknowns, \bold{59}, 65 - \subitem freezing/thawing of, 42 - \item type variables, \bold{59} - \item types, \bold{58} - \subitem certified, \bold{59} - \subitem printing of, 3, 14, 59 + \subitem of classical prover, 125 + \subitem of macros, 88 + \subitem of searching tacticals, 31 + \subitem of simplification, 100, 108--109 + \subitem of tactics, 25 + \subitem of unification, 40 + \item {\tt transitive}, \bold{43} + \item translations, 90--93 + \subitem parse, 74, 82 + \subitem parse AST, \bold{80}, 81 + \subitem print, 74 + \subitem print AST, \bold{83} + \item {\tt translations} section, 85 + \item {\tt trivial}, \bold{46} + \item {\tt Trueprop} constant, 76 + \item {\tt TRY}, \bold{29, 30} + \item {\tt TRYALL}, \bold{34} + \item {\tt TVar}, \bold{60} + \item {\tt tvar} nonterminal, \bold{66, 67}, 80, 87 + \item {\tt typ} ML type, 60 + \item {\tt Type}, \bold{60} + \item {\tt type} nonterminal, \bold{66} + \item {\tt type} type, 71 + \item type constraints, 66, 74, 83 + \item type constructors, \bold{60} + \item type identifiers, 66 + \item type synonyms, \bold{50} + \item type unknowns, \bold{60}, 66 + \subitem freezing/thawing of, 45 + \item type variables, \bold{60} + \item types, \bold{60} + \subitem certified, \bold{60} + \subitem printing of, 4, 14, 60 \indexspace - \item {\ptt undo}, 6, \bold{9}, 13 - \item unknowns, \bold{56}, 65 - \item {\ptt unlink_thy}, \bold{50} - \item {\ptt update}, \bold{50} - \item {\ptt uresult}, \bold{8}, 14, 52 - \item {\ptt use}, \bold{2} - \item use_dir, 54, 55 - \item {\ptt use_thy}, \bold{49}, 50 + \item {\tt undo}, 6, \bold{9}, 13 + \item unknowns, \bold{57}, 66 + \item {\tt unlink_thy}, \bold{54} + \item {\tt update}, \bold{54} + \item {\tt uresult}, \bold{8}, 15, 55 + \item {\tt use}, \bold{3} + \item {\tt use_thy}, \bold{52}, 53 \indexspace - \item {\ptt Var}, \bold{56}, 79 - \item {\ptt var} nonterminal, \bold{65, 66}, 77, 84 - \item {\ptt Variable}, 76 + \item {\tt Var}, \bold{57}, 82 + \item {\tt var} nonterminal, \bold{66, 67}, 80, 87 + \item {\tt Variable}, 79 \item variables - \subitem bound, \bold{56} - \subitem free, \bold{56} - \item {\ptt variant_abs}, \bold{57}, 89 - \item {\ptt varifyT}, \bold{42} + \subitem bound, \bold{57} + \subitem free, \bold{57} + \item {\tt variant_abs}, \bold{58} + \item {\tt varifyT}, \bold{45} \indexspace - \item {\ptt xlisten} shell script, 5 - \item {\ptt xnum} nonterminal, \bold{65}, 77, 84 - \item {\ptt xstr} nonterminal, \bold{65}, 77, 84 + \item {\tt xnum} nonterminal, \bold{66}, 80, 87 + \item {\tt xstr} nonterminal, \bold{66}, 80, 87 \indexspace - \item {\ptt zero_var_indexes}, \bold{36} + \item {\tt zero_var_indexes}, \bold{39} \end{theindex}
--- a/doc-src/Ref/ref.rao Mon May 05 21:18:01 1997 +0200 +++ b/doc-src/Ref/ref.rao Tue May 06 12:50:16 1997 +0200 @@ -1,7 +1,7 @@ -% This file was generated by '/usr/stud/berghofe/latex/rail/rail' from 'ref.rai' +% This file was generated by 'rail' from 'ref.rai' \rail@t {lbrace} \rail@t {rbrace} -\rail@i {1}{ \par theoryDef : id '=' (name + '+') ('+' extension | ()); \par name: id | string; \par extension : (section +) 'end' ( () | ml ); \par section : classes | default | types | arities | consts | constdefs | trans | defs | rules | oracle ; \par classes : 'classes' ( ( id ( () | '<' (id + ',') ) ) + ) ; \par default : 'default' sort ; \par sort : id | lbrace (id * ',') rbrace ; \par types : 'types' ( ( typeDecl ( () | '(' infix ')' ) ) + ) ; \par infix : ( 'infixr' | 'infixl' ) nat; \par typeDecl : typevarlist name ( () | '=' ( string | type ) ); \par typevarlist : () | tid | '(' ( tid + ',' ) ')'; \par type : simpleType | '(' type ')' | type '=>' type | '[' ( type + "," ) ']' '=>' type; \par simpleType: id | ( tid ( () | '::' id ) ) | '(' ( type + "," ) ')' id | simpleType id; \par \par arities : 'arities' ((( name + ',' ) '::' arity ) + ) ; \par arity : ( () | '(' (sort + ',') ')' ) id ; \par \par consts : 'consts' ( ( constDecl ( () | ( '(' mixfix ')' ) ) ) + ) ; \par constDecl : ( name + ',') '::' (string | type); \par mixfix : string ( () | ( () | ('[' (nat + ',') ']')) nat ) | infix | 'binder' string nat ; \par constdefs : 'constdefs' (id '::' (string | type) string +) ; \par trans : 'translations' ( pat ( '==' | '=>' | '<=' ) pat + ) ; \par pat : ( () | ( '(' id ')' ) ) string; \par rules : 'rules' (( id string ) + ) ; \par defs : 'defs' (( id string ) + ) ; \par oracle : 'oracle' name ; \par ml : 'ML' text ; \par } +\rail@i {1}{ \par theoryDef : id '=' (name + '+') ('+' extension | ()) ; \par name: id | string ; \par extension : (section +) 'end' ( () | ml ) ; \par section : classes | default | types | arities | consts | syntax | trans | defs | constdefs | rules | axclass | instance | oracle ; \par classes : 'classes' ( classDecl + ) ; \par classDecl : (id (() | '<' (id + ','))) ; \par default : 'default' sort ; \par sort : id | lbrace (id * ',') rbrace ; \par types : 'types' ( ( typeDecl ( () | '(' infix ')' ) ) + ) ; \par infix : ( 'infixr' | 'infixl' ) nat ; \par typeDecl : typevarlist name ( () | '=' ( string | type ) ); \par typevarlist : () | tid | '(' ( tid + ',' ) ')'; \par type : simpleType | '(' type ')' | type '=>' type | '[' ( type + "," ) ']' '=>' type; \par simpleType: id | ( tid ( () | '::' id ) ) | '(' ( type + "," ) ')' id | simpleType id ; \par arities : 'arities' ((name + ',') '::' arity +) ; \par arity : ( () | '(' (sort + ',') ')' ) sort ; \par consts : 'consts' ( mixfixConstDecl + ) ; \par syntax : 'syntax' (() | mode) ( mixfixConstDecl + ); \par mode : '(' name (() | 'output') ')' ; \par mixfixConstDecl : constDecl (() | ( '(' mixfix ')' )) ; \par constDecl : ( name + ',') '::' (string | type); \par mixfix : string ( () | ( () | ('[' (nat + ',') ']')) nat ) | ( 'infixr' | 'infixl' ) (() | string) nat | 'binder' string nat ; \par trans : 'translations' ( pat ( '==' | '=>' | '<=' ) pat + ) ; \par pat : ( () | ( '(' id ')' ) ) string; \par rules : 'rules' (( id string ) + ) ; \par defs : 'defs' (( id string ) + ) ; \par constdefs : 'constdefs' (id '::' (string | type) string +) ; \par axclass : 'axclass' classDecl (() | ( id string ) +) ; \par instance : 'instance' ( name '<' name | name '::' arity) witness ; \par witness : (() | ((string | longident) + ',')) (() | verbatim) ; \par oracle : 'oracle' name ; \par ml : 'ML' text ; \par } \rail@o {1}{ \rail@begin{2}{theoryDef} \rail@nont{id}[] @@ -35,7 +35,7 @@ \rail@nont{ml}[] \rail@endbar \rail@end -\rail@begin{10}{section} +\rail@begin{13}{section} \rail@bar \rail@nont{classes}[] \rail@nextbar{1} @@ -47,20 +47,31 @@ \rail@nextbar{4} \rail@nont{consts}[] \rail@nextbar{5} -\rail@nont{constdefs}[] +\rail@nont{syntax}[] \rail@nextbar{6} \rail@nont{trans}[] \rail@nextbar{7} \rail@nont{defs}[] \rail@nextbar{8} +\rail@nont{constdefs}[] +\rail@nextbar{9} \rail@nont{rules}[] -\rail@nextbar{9} +\rail@nextbar{10} +\rail@nont{axclass}[] +\rail@nextbar{11} +\rail@nont{instance}[] +\rail@nextbar{12} \rail@nont{oracle}[] \rail@endbar \rail@end -\rail@begin{4}{classes} +\rail@begin{2}{classes} \rail@term{classes}[] \rail@plus +\rail@nont{classDecl}[] +\rail@nextplus{1} +\rail@endplus +\rail@end +\rail@begin{3}{classDecl} \rail@nont{id}[] \rail@bar \rail@nextbar{1} @@ -71,8 +82,6 @@ \rail@cterm{,}[] \rail@endplus \rail@endbar -\rail@nextplus{3} -\rail@endplus \rail@end \rail@begin{1}{default} \rail@term{default}[] @@ -213,11 +222,36 @@ \rail@endplus \rail@term{)}[] \rail@endbar -\rail@nont{id}[] +\rail@nont{sort}[] \rail@end -\rail@begin{3}{consts} +\rail@begin{2}{consts} \rail@term{consts}[] \rail@plus +\rail@nont{mixfixConstDecl}[] +\rail@nextplus{1} +\rail@endplus +\rail@end +\rail@begin{2}{syntax} +\rail@term{syntax}[] +\rail@bar +\rail@nextbar{1} +\rail@nont{mode}[] +\rail@endbar +\rail@plus +\rail@nont{mixfixConstDecl}[] +\rail@nextplus{1} +\rail@endplus +\rail@end +\rail@begin{2}{mode} +\rail@term{(}[] +\rail@nont{name}[] +\rail@bar +\rail@nextbar{1} +\rail@term{output}[] +\rail@endbar +\rail@term{)}[] +\rail@end +\rail@begin{2}{mixfixConstDecl} \rail@nont{constDecl}[] \rail@bar \rail@nextbar{1} @@ -225,8 +259,6 @@ \rail@nont{mixfix}[] \rail@term{)}[] \rail@endbar -\rail@nextplus{2} -\rail@endplus \rail@end \rail@begin{2}{constDecl} \rail@plus @@ -241,7 +273,7 @@ \rail@nont{type}[] \rail@endbar \rail@end -\rail@begin{6}{mixfix} +\rail@begin{7}{mixfix} \rail@bar \rail@nont{string}[] \rail@bar @@ -259,27 +291,22 @@ \rail@nont{nat}[] \rail@endbar \rail@nextbar{4} -\rail@nont{infix}[] +\rail@bar +\rail@term{infixr}[] \rail@nextbar{5} +\rail@term{infixl}[] +\rail@endbar +\rail@bar +\rail@nextbar{5} +\rail@nont{string}[] +\rail@endbar +\rail@nont{nat}[] +\rail@nextbar{6} \rail@term{binder}[] \rail@nont{string}[] \rail@nont{nat}[] \rail@endbar \rail@end -\rail@begin{3}{constdefs} -\rail@term{constdefs}[] -\rail@plus -\rail@nont{id}[] -\rail@term{::}[] -\rail@bar -\rail@nont{string}[] -\rail@nextbar{1} -\rail@nont{type}[] -\rail@endbar -\rail@nont{string}[] -\rail@nextplus{2} -\rail@endplus -\rail@end \rail@begin{4}{trans} \rail@term{translations}[] \rail@plus @@ -320,6 +347,63 @@ \rail@nextplus{1} \rail@endplus \rail@end +\rail@begin{3}{constdefs} +\rail@term{constdefs}[] +\rail@plus +\rail@nont{id}[] +\rail@term{::}[] +\rail@bar +\rail@nont{string}[] +\rail@nextbar{1} +\rail@nont{type}[] +\rail@endbar +\rail@nont{string}[] +\rail@nextplus{2} +\rail@endplus +\rail@end +\rail@begin{3}{axclass} +\rail@term{axclass}[] +\rail@nont{classDecl}[] +\rail@bar +\rail@nextbar{1} +\rail@plus +\rail@nont{id}[] +\rail@nont{string}[] +\rail@nextplus{2} +\rail@endplus +\rail@endbar +\rail@end +\rail@begin{2}{instance} +\rail@term{instance}[] +\rail@bar +\rail@nont{name}[] +\rail@term{<}[] +\rail@nont{name}[] +\rail@nextbar{1} +\rail@nont{name}[] +\rail@term{::}[] +\rail@nont{arity}[] +\rail@endbar +\rail@nont{witness}[] +\rail@end +\rail@begin{4}{witness} +\rail@bar +\rail@nextbar{1} +\rail@plus +\rail@bar +\rail@nont{string}[] +\rail@nextbar{2} +\rail@nont{longident}[] +\rail@endbar +\rail@nextplus{3} +\rail@cterm{,}[] +\rail@endplus +\rail@endbar +\rail@bar +\rail@nextbar{1} +\rail@nont{verbatim}[] +\rail@endbar +\rail@end \rail@begin{1}{oracle} \rail@term{oracle}[] \rail@nont{name}[]
--- a/doc-src/Ref/ref.tex Mon May 05 21:18:01 1997 +0200 +++ b/doc-src/Ref/ref.tex Tue May 06 12:50:16 1997 +0200 @@ -32,19 +32,19 @@ \makeindex -\underscoreoff - \setcounter{secnumdepth}{1} \setcounter{tocdepth}{2} \pagestyle{headings} \sloppy \binperiod %%%treat . like a binary operator -\railalias{lbrace}{\{} -\railalias{rbrace}{\}} +\railalias{lbrace}{\ttlbrace} +\railalias{rbrace}{\ttrbrace} \railterm{lbrace,rbrace} \begin{document} +\underscoreoff + \index{definitions|see{rewriting, meta-level}} \index{rewriting!object-level|see{simplification}} \index{meta-rules|see{meta-rules}}
--- a/doc-src/Ref/simplifier.tex Mon May 05 21:18:01 1997 +0200 +++ b/doc-src/Ref/simplifier.tex Tue May 06 12:50:16 1997 +0200 @@ -83,9 +83,9 @@ \S\ref{sec:reordering-asms} for ways around this problem. \end{warn} -Using the simplifier effectively may take a bit of experimentation. -\index{tracing!of simplification|)}\ttindex{trace_simp} -The tactics can be traced with the ML command \verb$trace_simp := true$. +Using the simplifier effectively may take a bit of experimentation. +\index{tracing!of simplification}\index{*trace_simp} The tactics can +be traced by setting \verb$trace_simp := true$. There is not just one global current simpset, but one associated with each theory as well. How are these simpset built up? @@ -145,10 +145,9 @@ defaults so that most simplifier calls specify only rewrite rules. Experienced users can exploit the other components to streamline proofs. -Logics like \HOL, which support a current simpset\index{simpset!current}, -store its value in an ML reference variable usually called {\tt - simpset}\index{simpset@{\tt simpset} ML value}. It can be accessed via -{\tt!simpset} and updated via {\tt simpset := \dots}. +Logics like \HOL, which support a current +simpset\index{simpset!current}, store its value in an ML reference +variable called {\tt simpset}\index{simpset@{\tt simpset} ML value}. \subsection{Rewrite rules} \index{rewrite rules|(} @@ -159,7 +158,7 @@ \Var{A} \un \Var{B} &\equiv& \{x.x\in \Var{A} \disj x\in \Var{B}\} \end{eqnarray*} Conditional rewrites such as $\Var{m}<\Var{n} \Imp \Var{m}/\Var{n} = -0$ are permitted; the conditions can be arbitrary terms. +0$ are permitted; the conditions can be arbitrary formulas. Internally, all rewrite rules are translated into meta-equalities, theorems with conclusion $lhs \equiv rhs$. Each simpset contains a function for @@ -200,7 +199,7 @@ \subsection{*Congruence rules}\index{congruence rules} Congruence rules are meta-equalities of the form -\[ \List{\dots} \Imp +\[ \dots \Imp f(\Var{x@1},\ldots,\Var{x@n}) \equiv f(\Var{y@1},\ldots,\Var{y@n}). \] This governs the simplification of the arguments of~$f$. For @@ -272,26 +271,27 @@ True} and $t=t$. It could use sophisticated means such as {\tt fast_tac}, though that could make simplification expensive. -Rewriting does not instantiate unknowns. For example, rewriting cannot -prove $a\in \Var{A}$ since this requires instantiating~$\Var{A}$. The -solver, however, is an arbitrary tactic and may instantiate unknowns as it -pleases. This is the only way the simplifier can handle a conditional -rewrite rule whose condition contains extra variables. When a simplification -tactic is to be combined with other provers, especially with the classical -reasoner, it is important whether it can be considered safe or not. Therefore, -the solver is split into a safe and anunsafe part. Both parts can be set -independently of each other using -\ttindex{setSSolver} (with a safe tactic as argument) and \ttindex{setSolver} -(with an unsafe tactic) . Additional solvers, which are tried after the already -set solvers, may be added using \ttindex{addSSolver} and \ttindex{addSolver}. +Rewriting does not instantiate unknowns. For example, rewriting +cannot prove $a\in \Var{A}$ since this requires +instantiating~$\Var{A}$. The solver, however, is an arbitrary tactic +and may instantiate unknowns as it pleases. This is the only way the +simplifier can handle a conditional rewrite rule whose condition +contains extra variables. When a simplification tactic is to be +combined with other provers, especially with the classical reasoner, +it is important whether it can be considered safe or not. Therefore, +the solver is split into a safe and an unsafe part. Both parts can be +set independently, using either \ttindex{setSSolver} with a safe +tactic as argument, or \ttindex{setSolver} with an unsafe tactic. +Additional solvers, which are tried after the already set solvers, may +be added using \ttindex{addSSolver} and \ttindex{addSolver}. -The standard simplification procedures uses solely the unsafe solver, which is -appropriate in most cases. But for special applications, where the simplification -process is not allowed to instantiate unknowns within the goal, the tactic -\ttindexbold{safe_asm_full_simp_tac} is provided. Like {\tt asm_full_simp_tac}, -it uses the unsafe solver for any embedded simplification steps -(i.e. for solving subgoals created by the subgoaler), -but for the current subgoal, it applies the safe solver. +The standard simplification strategy solely uses the unsafe solver, +which is appropriate in most cases. But for special applications where +the simplification process is not allowed to instantiate unknowns +within the goal, the tactic \ttindexbold{safe_asm_full_simp_tac} is +provided. It uses the \emph{safe} solver for the current subgoal, but +applies ordinary unsafe {\tt asm_full_simp_tac} for recursive internal +simplifications (for conditional rules or congruences). \index{assumptions!in simplification} The tactic is presented with the full goal, including the asssumptions. @@ -308,16 +308,17 @@ \Var{x}$, usually by reflexivity. In particular, reflexivity should be tried before any of the fancy tactics like {\tt fast_tac}. -It may even happen that, due to simplification, the subgoal is no longer an -equality. For example $False \bimp \Var{Q}$ could be rewritten to -$\neg\Var{Q}$. To cover this case, the solver could try resolving with the -theorem $\neg False$. +It may even happen that due to simplification the subgoal is no longer +an equality. For example $False \bimp \Var{Q}$ could be rewritten to +$\neg\Var{Q}$. To cover this case, the solver could try resolving +with the theorem $\neg False$. \begin{warn} - If the simplifier aborts with the message {\tt Failed congruence proof!}, - then the subgoaler or solver has failed to prove a premise of a - congruence rule. This should never occur; it indicates that some - congruence rule, or possibly the subgoaler or solver, is faulty. + If the simplifier aborts with the message {\tt Failed congruence + proof!}, then the subgoaler or solver has failed to prove a + premise of a congruence rule. This should never occur under normal + circumstances; it indicates that some congruence rule, or possibly + the subgoaler or solver, is faulty. \end{warn} @@ -430,9 +431,9 @@ in conditional rewrite rules are solved recursively before the rewrite rule is applied. -The infix operations \ttindex{addsimps}/\ttindex{delsimps} add/delete rewrite -rules to/from a simpset. They are used to implement \ttindex{Addsimps} and -\ttindex{Delsimps}, e.g. +The infix operation \ttindex{addsimps} adds rewrite rules to a +simpset, while \ttindex{delsimps} deletes them. They are used to +implement \ttindex{Addsimps} and \ttindex{Delsimps}, e.g. \begin{ttbox} fun Addsimps thms = (simpset := (!simpset addsimps thms)); \end{ttbox} @@ -515,7 +516,7 @@ \end{ttbox} \subsection{An example of tracing} -\index{tracing!of simplification|(}\ttindex{trace_simp|(} +\index{tracing!of simplification|(}\index{*trace_simp} Let us prove a similar result involving more complex terms. The two equations together can be used to prove that addition is commutative. \begin{ttbox} @@ -572,7 +573,7 @@ {\out m + Suc(n) = Suc(m + n)} {\out No subgoals!} \end{ttbox} -\index{tracing!of simplification|)}\ttindex{trace_simp|)} +\index{tracing!of simplification|)} \subsection{Free variables and simplification} Here is a conjecture to be proved for an arbitrary function~$f$ satisfying @@ -663,7 +664,7 @@ Because ordinary rewriting loops given such rules, the simplifier employs a special strategy, called {\bf ordered rewriting}\index{rewriting!ordered}. -There is a built-in lexicographic ordering on terms. A permutative rewrite +There is a standard lexicographic ordering on terms. A permutative rewrite rule is applied only if it decreases the given term with respect to this ordering. For example, commutativity rewrites~$b+a$ to $a+b$, but then stops because $a+b$ is strictly less than $b+a$. The Boyer-Moore theorem @@ -692,74 +693,73 @@ examples; other algebraic structures are amenable to ordered rewriting, such as boolean rings. -\subsection{Example: sums of integers} -This example is set in Isabelle's higher-order logic. Theory -\thydx{Arith} contains the theory of arithmetic. The simpset {\tt - arith_ss} contains many arithmetic laws including distributivity -of~$\times$ over~$+$, while {\tt add_ac} is a list consisting of the A, C -and LC laws for~$+$. Let us prove the theorem +\subsection{Example: sums of natural numbers} +This example is set in \HOL\ (see \texttt{HOL/ex/NatSum}). Theory +\thydx{Arith} contains natural numbers arithmetic. Its associated +simpset contains many arithmetic laws including distributivity +of~$\times$ over~$+$, while {\tt add_ac} is a list consisting of the +A, C and LC laws for~$+$ on type \texttt{nat}. Let us prove the +theorem \[ \sum@{i=1}^n i = n\times(n+1)/2. \] % A functional~{\tt sum} represents the summation operator under the -interpretation ${\tt sum}(f,n+1) = \sum@{i=0}^n f(i)$. We extend {\tt - Arith} using a theory file: +interpretation ${\tt sum} \, f \, (n + 1) = \sum@{i=0}^n f\,i$. We +extend {\tt Arith} using a theory file: \begin{ttbox} NatSum = Arith + consts sum :: [nat=>nat, nat] => nat -rules sum_0 "sum(f,0) = 0" - sum_Suc "sum(f,Suc(n)) = f(n) + sum(f,n)" +rules sum_0 "sum f 0 = 0" + sum_Suc "sum f (Suc n) = f n + sum f n" end \end{ttbox} -After declaring {\tt open NatSum}, we make the required simpset by adding -the AC-rules for~$+$ and the axioms for~{\tt sum}: +We make the required simpset by adding the AC-rules for~$+$ and the +axioms for~{\tt sum}: \begin{ttbox} -val natsum_ss = arith_ss addsimps ([sum_0,sum_Suc] \at add_ac); -{\out val natsum_ss = SS \{\ldots\} : simpset} +val natsum_ss = simpset_of "Arith" addsimps ([sum_0,sum_Suc] \at add_ac); +{\out val natsum_ss = \ldots : simpset} \end{ttbox} -Our desired theorem now reads ${\tt sum}(\lambda i.i,n+1) = +Our desired theorem now reads ${\tt sum} \, (\lambda i.i) \, (n+1) = n\times(n+1)/2$. The Isabelle goal has both sides multiplied by~$2$: \begin{ttbox} -goal NatSum.thy "Suc(Suc(0))*sum(\%i.i,Suc(n)) = n*Suc(n)"; +goal NatSum.thy "2 * sum (\%i.i) (Suc n) = n * Suc n"; {\out Level 0} -{\out Suc(Suc(0)) * sum(\%i. i, Suc(n)) = n * Suc(n)} -{\out 1. Suc(Suc(0)) * sum(\%i. i, Suc(n)) = n * Suc(n)} +{\out 2 * sum (\%i. i) (Suc n) = n * Suc n} +{\out 1. 2 * sum (\%i. i) (Suc n) = n * Suc n} \end{ttbox} -Induction should not be applied until the goal is in the simplest form: +Induction should not be applied until the goal is in the simplest +form: \begin{ttbox} by (simp_tac natsum_ss 1); {\out Level 1} -{\out Suc(Suc(0)) * sum(\%i. i, Suc(n)) = n * Suc(n)} -{\out 1. n + (n + (sum(\%i. i, n) + sum(\%i. i, n))) = n + n * n} +{\out 2 * sum (\%i. i) (Suc n) = n * Suc n} +{\out 1. n + (sum (\%i. i) n + sum (\%i. i) n) = n * n} \end{ttbox} -Ordered rewriting has sorted the terms in the left-hand side. -The subgoal is now ready for induction: +Ordered rewriting has sorted the terms in the left-hand side. The +subgoal is now ready for induction: \begin{ttbox} by (nat_ind_tac "n" 1); {\out Level 2} -{\out Suc(Suc(0)) * sum(\%i. i, Suc(n)) = n * Suc(n)} -{\out 1. 0 + (0 + (sum(\%i. i, 0) + sum(\%i. i, 0))) = 0 + 0 * 0} +{\out 2 * sum (\%i. i) (Suc n) = n * Suc n} +{\out 1. 0 + (sum (\%i. i) 0 + sum (\%i. i) 0) = 0 * 0} \ttbreak -{\out 2. !!n1. n1 + (n1 + (sum(\%i. i, n1) + sum(\%i. i, n1))) =} -{\out n1 + n1 * n1 ==>} -{\out Suc(n1) +} -{\out (Suc(n1) + (sum(\%i. i, Suc(n1)) + sum(\%i. i, Suc(n1)))) =} -{\out Suc(n1) + Suc(n1) * Suc(n1)} +{\out 2. !!n1. n1 + (sum (\%i. i) n1 + sum (\%i. i) n1) = n1 * n1} +{\out ==> Suc n1 + (sum (\%i. i) (Suc n1) + sum (\%i. i) (Suc n1)) =} +{\out Suc n1 * Suc n1} \end{ttbox} Simplification proves both subgoals immediately:\index{*ALLGOALS} \begin{ttbox} by (ALLGOALS (asm_simp_tac natsum_ss)); {\out Level 3} -{\out Suc(Suc(0)) * sum(\%i. i, Suc(n)) = n * Suc(n)} +{\out 2 * sum (\%i. i) (Suc n) = n * Suc n} {\out No subgoals!} \end{ttbox} If we had omitted {\tt add_ac} from the simpset, simplification would have failed to prove the induction step: \begin{ttbox} -Suc(Suc(0)) * sum(\%i. i, Suc(n)) = n * Suc(n) - 1. !!n1. n1 + (n1 + (sum(\%i. i, n1) + sum(\%i. i, n1))) = - n1 + n1 * n1 ==> - n1 + (n1 + (n1 + sum(\%i. i, n1) + (n1 + sum(\%i. i, n1)))) = - n1 + (n1 + (n1 + n1 * n1)) +2 * sum (\%i. i) (Suc n) = n * Suc n + 1. !!n1. n1 + sum (\%i. i) n1 + (n1 + sum (\%i. i) n1) = n1 + n1 * n1 + ==> n1 + (n1 + sum (\%i. i) n1) + (n1 + (n1 + sum (\%i. i) n1)) = + n1 + (n1 + (n1 + n1 * n1)) \end{ttbox} Ordered rewriting proves this by sorting the left-hand side. Proving arithmetic theorems without ordered rewriting requires explicit use of @@ -811,7 +811,7 @@ Simplification works by reducing various object-equalities to meta-equality. It requires rules stating that equal terms and equivalent formulae are also equal at the meta-level. The rule declaration part of -the file {\tt FOL/ifol.thy} contains the two lines +the file {\tt FOL/IFOL.thy} contains the two lines \begin{ttbox}\index{*eq_reflection theorem}\index{*iff_reflection theorem} eq_reflection "(x=y) ==> (x==y)" iff_reflection "(P<->Q) ==> (P==Q)"
--- a/doc-src/Ref/substitution.tex Mon May 05 21:18:01 1997 +0200 +++ b/doc-src/Ref/substitution.tex Tue May 06 12:50:16 1997 +0200 @@ -21,8 +21,10 @@ \section{Substitution rules} \index{substitution!rules}\index{*subst theorem} Many logics include a substitution rule of the form -$$ \List{\Var{a}=\Var{b}; \Var{P}(\Var{a})} \Imp - \Var{P}(\Var{b}) \eqno(subst)$$ +$$ +\List{\Var{a}=\Var{b}; \Var{P}(\Var{a})} \Imp +\Var{P}(\Var{b}) \eqno(subst) +$$ In backward proof, this may seem difficult to use: the conclusion $\Var{P}(\Var{b})$ admits far too many unifiers. But, if the theorem {\tt eqth} asserts $t=u$, then \hbox{\tt eqth RS subst} is the derived rule @@ -42,8 +44,10 @@ resolve_tac [eqth RS ssubst] \(i\){\it,} \end{ttbox} where \tdxbold{ssubst} is the `swapped' substitution rule -$$ \List{\Var{a}=\Var{b}; \Var{P}(\Var{b})} \Imp - \Var{P}(\Var{a}). \eqno(ssubst)$$ +$$ +\List{\Var{a}=\Var{b}; \Var{P}(\Var{b})} \Imp +\Var{P}(\Var{a}). \eqno(ssubst) +$$ If \tdx{sym} denotes the symmetry rule \(\Var{a}=\Var{b}\Imp\Var{b}=\Var{a}\), then {\tt ssubst} is just \hbox{\tt sym RS subst}. Many logics with equality include the rules {\tt @@ -57,7 +61,7 @@ eresolve_tac [subst] \(i\) {\rm or} eresolve_tac [ssubst] \(i\){\it.} \end{ttbox} -Logics HOL, FOL and ZF define the tactic \ttindexbold{stac} by +Logics \HOL, {\FOL} and {\ZF} define the tactic \ttindexbold{stac} by \begin{ttbox} fun stac eqth = CHANGED o rtac (eqth RS ssubst); \end{ttbox}
--- a/doc-src/Ref/syntax.tex Mon May 05 21:18:01 1997 +0200 +++ b/doc-src/Ref/syntax.tex Tue May 06 12:50:16 1997 +0200 @@ -31,7 +31,7 @@ \begin{center} \begin{tabular}{cl} string & \\ -$\downarrow$ & parser \\ +$\downarrow$ & lexer, parser \\ parse tree & \\ $\downarrow$ & parse \AST{} translation \\ \AST{} & \\ @@ -43,7 +43,7 @@ \AST{} & \\ $\downarrow$ & \AST{} rewriting (macros) \\ \AST{} & \\ -$\downarrow$ & print \AST{} translation, printer \\ +$\downarrow$ & print \AST{} translation, token translation \\ string & \end{tabular} @@ -304,11 +304,13 @@ and their syntactic sugar (denoted by \dots{} above) are joined to make a single string. -If an application $(\mtt"c\mtt"~ x@1 \ldots x@m)$ has more arguments than the -corresponding production, it is first split into $((\mtt"c\mtt"~ x@1 \ldots -x@n) ~ x@{n+1} \ldots x@m)$. Applications with too few arguments or with -non-constant head or without a corresponding production are printed as -$f(x@1, \ldots, x@l)$ or $(\alpha@1, \ldots, \alpha@l) ty$. An occurrence of +If an application $(\mtt"c\mtt"~ x@1 \ldots x@m)$ has more arguments +than the corresponding production, it is first split into +$((\mtt"c\mtt"~ x@1 \ldots x@n) ~ x@{n+1} \ldots x@m)$. Applications +with too few arguments or with non-constant head or without a +corresponding production are printed as $f(x@1, \ldots, x@l)$ or +$(\alpha@1, \ldots, \alpha@l) ty$. Multiple productions associated +with some name $c$ are tried in order of appearance. An occurrence of $\Variable x$ is simply printed as~$x$. Blanks are {\em not\/} inserted automatically. If blanks are required to @@ -318,7 +320,7 @@ -\section{Macros: Syntactic rewriting} \label{sec:macros} +\section{Macros: syntactic rewriting} \label{sec:macros} \index{macros|(}\index{rewriting!syntactic|(} Mixfix declarations alone can handle situations where there is a direct @@ -338,19 +340,20 @@ on abstract syntax trees. They are usually easy to read and write, and can express all but the most obscure translations. -Figure~\ref{fig:set_trans} defines a fragment of first-order logic and set -theory.\footnote{This and the following theories are complete working - examples, though they specify only syntax, no axioms. The file {\tt - ZF/ZF.thy} presents a full set theory definition, including many - macro rules.} Theory {\tt SET} defines constants for set comprehension -({\tt Collect}), replacement ({\tt Replace}) and bounded universal -quantification ({\tt Ball}). Each of these binds some variables. Without -additional syntax we should have to write $\forall x \in A. P$ as {\tt - Ball(A,\%x.P)}, and similarly for the others. +Figure~\ref{fig:set_trans} defines a fragment of first-order logic and +set theory.\footnote{This and the following theories are complete + working examples, though they specify only syntax, no axioms. The + file {\tt ZF/ZF.thy} presents a full set theory definition, + including many macro rules.} Theory {\tt SetSyntax} declares +constants for set comprehension ({\tt Collect}), replacement ({\tt + Replace}) and bounded universal quantification ({\tt Ball}). Each +of these binds some variables. Without additional syntax we should +have to write $\forall x \in A. P$ as {\tt Ball(A,\%x.P)}, and +similarly for the others. \begin{figure} \begin{ttbox} -SET = Pure + +SetSyntax = Pure + types i o arities @@ -425,8 +428,8 @@ \item Every variable in~$r$ must also occur in~$l$. \end{itemize} -Macro rules may refer to any syntax from the parent theories. They may -also refer to anything defined before the {\tt .thy} file's {\tt +Macro rules may refer to any syntax from the parent theories. They +may also refer to anything defined before the current {\tt translations} section --- including any mixfix declarations. Upon declaration, both sides of the macro rule undergo parsing and parse @@ -440,10 +443,11 @@ matching. These are all names that have been declared as classes, types or constants (logical and syntactic). -The result of this preprocessing is two lists of macro rules, each stored -as a pair of \AST{}s. They can be viewed using {\tt print_syntax} -(sections \ttindex{parse_rules} and \ttindex{print_rules}). For -theory~{\tt SET} of Fig.~\ref{fig:set_trans} these are +The result of this preprocessing is two lists of macro rules, each +stored as a pair of \AST{}s. They can be viewed using {\tt + print_syntax} (sections \ttindex{parse_rules} and +\ttindex{print_rules}). For theory~{\tt SetSyntax} of +Fig.~\ref{fig:set_trans} these are \begin{ttbox} parse_rules: ("{\at}Collect" x A P) -> ("Collect" A ("_abs" x P)) @@ -484,9 +488,7 @@ To allow such constraints to be re-read, your syntax should specify bound variables using the nonterminal~\ndx{idt}. This is the case in our -example. Choosing {\tt id} instead of {\tt idt} is a common error, -especially since it appears in former versions of most of Isabelle's -object-logics. +example. Choosing {\tt id} instead of {\tt idt} is a common error. \end{warn} @@ -552,17 +554,17 @@ The term \verb|%Nil.t| will be printed as \verb|%[].t|, which might not be expected! Guess how type~{\tt Nil} is printed? -Normalizing an \AST{} involves repeatedly applying macro rules until none -are applicable. Macro rules are chosen in the order that they appear in the -{\tt translations} section. You can watch the normalization of \AST{}s -during parsing and printing by setting \ttindex{Syntax.trace_norm_ast} to -{\tt true}.\index{tracing!of macros} Alternatively, use +Normalizing an \AST{} involves repeatedly applying macro rules until +none are applicable. Macro rules are chosen in order of appearance in +the theory definitions. You can watch the normalization of \AST{}s +during parsing and printing by setting \ttindex{Syntax.trace_norm_ast} +to {\tt true}.\index{tracing!of macros} Alternatively, use \ttindex{Syntax.test_read}. The information displayed when tracing -includes the \AST{} before normalization ({\tt pre}), redexes with results -({\tt rewrote}), the normal form finally reached ({\tt post}) and some -statistics ({\tt normalize}). If tracing is off, -\ttindex{Syntax.stat_norm_ast} can be set to {\tt true} in order to enable -printing of the normal form and statistics only. +includes the \AST{} before normalization ({\tt pre}), redexes with +results ({\tt rewrote}), the normal form finally reached ({\tt post}) +and some statistics ({\tt normalize}). If tracing is off, +\ttindex{Syntax.stat_norm_ast} can be set to {\tt true} in order to +enable printing of the normal form and statistics only. \subsection{Example: the syntax of finite sets} @@ -574,7 +576,7 @@ \index{"@Enum@{\tt\at Enum} constant} \index{"@Finset@{\tt\at Finset} constant} \begin{ttbox} -FINSET = SET + +FinSyntax = SetSyntax + types is syntax @@ -654,7 +656,7 @@ readable in some cases: {\em calling\/} translation functions by parse macros: \begin{ttbox} -PROD = FINSET + +ProdSyntax = FinSyntax + consts Pi :: [i, i => i] => i syntax @@ -693,10 +695,10 @@ \index{translations|(} % This section describes the translation function mechanism. By writing -\ML{} functions, you can do almost everything with terms or \AST{}s during -parsing and printing. The logic \LK\ is a good example of sophisticated -transformations between internal and external representations of sequents; -here, macros would be useless. +\ML{} functions, you can do almost everything to terms or \AST{}s +during parsing and printing. The logic \LK\ is a good example of +sophisticated transformations between internal and external +representations of sequents; here, macros would be useless. A full understanding of translations requires some familiarity with Isabelle's internals, especially the datatypes {\tt term}, {\tt typ}, @@ -705,50 +707,55 @@ use translation functions. \subsection{Declaring translation functions} -There are four kinds of translation functions. Each such function is -associated with a name, which triggers calls to it. Such names can be -constants (logical or syntactic) or type constructors. +There are four kinds of translation functions, with one of these +coming in two variants. Each such function is associated with a name, +which triggers calls to it. Such names can be constants (logical or +syntactic) or type constructors. {\tt print_syntax} displays the sets of names associated with the -translation functions of a {\tt Syntax.syntax} under +translation functions of a theory under \ttindex{parse_ast_translation}, \ttindex{parse_translation}, -\ttindex{print_translation} and \ttindex{print_ast_translation}. You can -add new ones via the {\tt ML} section\index{*ML section} of -a {\tt .thy} file. There may never be more than one function of the same -kind per name. Even though the {\tt ML} section is the very last part of a -{\tt .thy} file, newly installed translation functions are effective when -processing the preceding section. +\ttindex{print_translation} and \ttindex{print_ast_translation}. You +can add new ones via the {\tt ML} section\index{*ML section} of a +theory definition file. There may never be more than one function of +the same kind per name. Even though the {\tt ML} section is the very +last part of the file, newly installed translation functions are +already effective when processing all of the preceding sections. -The {\tt ML} section is copied verbatim near the beginning of the \ML\ file -generated from a {\tt .thy} file. Definitions made here are accessible as -components of an \ML\ structure; to make some definitions private, use an -\ML{} {\tt local} declaration. The {\tt ML} section may install translation -functions by declaring any of the following identifiers: +The {\tt ML} section's contents are simply copied verbatim near the +beginning of the \ML\ file generated from a theory definition file. +Definitions made here are accessible as components of an \ML\ +structure; to make some parts private, use an \ML{} {\tt local} +declaration. The {\ML} code may install translation functions by +declaring any of the following identifiers: \begin{ttbox} -val parse_ast_translation : (string * (ast list -> ast)) list -val print_ast_translation : (string * (ast list -> ast)) list -val parse_translation : (string * (term list -> term)) list -val print_translation : (string * (term list -> term)) list +val parse_ast_translation : (string * (ast list -> ast)) list +val print_ast_translation : (string * (ast list -> ast)) list +val parse_translation : (string * (term list -> term)) list +val print_translation : (string * (term list -> term)) list +val typed_print_translation : (string * (typ -> term list -> term)) list \end{ttbox} \subsection{The translation strategy} -All four kinds of translation functions are treated similarly. They are -called during the transformations between parse trees, \AST{}s and terms -(recall Fig.\ts\ref{fig:parse_print}). Whenever a combination of the form -$(\mtt"c\mtt"~x@1 \ldots x@n)$ is encountered, and a translation function -$f$ of appropriate kind exists for $c$, the result is computed by the \ML{} -function call $f \mtt[ x@1, \ldots, x@n \mtt]$. +The different kinds of translation functions are called during the +transformations between parse trees, \AST{}s and terms (recall +Fig.\ts\ref{fig:parse_print}). Whenever a combination of the form +$(\mtt"c\mtt"~x@1 \ldots x@n)$ is encountered, and a translation +function $f$ of appropriate kind exists for $c$, the result is +computed by the \ML{} function call $f \mtt[ x@1, \ldots, x@n \mtt]$. -For \AST{} translations, the arguments $x@1, \ldots, x@n$ are \AST{}s. A -combination has the form $\Constant c$ or $\Appl{\Constant c, x@1, \ldots, - x@n}$. For term translations, the arguments are terms and a combination -has the form $\ttfct{Const} (c, \tau)$ or $\ttfct{Const} (c, \tau) \ttapp -x@1 \ttapp \ldots \ttapp x@n$. Terms allow more sophisticated -transformations than \AST{}s do, typically involving abstractions and bound -variables. +For \AST{} translations, the arguments $x@1, \ldots, x@n$ are \AST{}s. +A combination has the form $\Constant c$ or $\Appl{\Constant c, x@1, + \ldots, x@n}$. For term translations, the arguments are terms and a +combination has the form $\ttfct{Const} (c, \tau)$ or $\ttfct{Const} +(c, \tau) \ttapp x@1 \ttapp \ldots \ttapp x@n$. Terms allow more +sophisticated transformations than \AST{}s do, typically involving +abstractions and bound variables. {\em Typed} print translations may +even peek at the type $\tau$ of the constant they are invoked on. -Regardless of whether they act on terms or \AST{}s, parse translations differ -from print translations in their overall behaviour fundamentally: +Regardless of whether they act on terms or \AST{}s, translation +functions called during the parsing process differ from those for +printing more fundamentally in their overall behaviour: \begin{description} \item[Parse translations] are applied bottom-up. The arguments are already in translated form. The translations must not fail; exceptions trigger @@ -804,40 +811,93 @@ \begin{ttbox} fun dependent_tr' (q, r) (A :: Abs (x, T, B) :: ts) = if 0 mem (loose_bnos B) then - let val (x', B') = variant_abs (x, dummyT, B); - in list_comb (Const (q, dummyT) $ Free (x', T) $ A $ B', ts) + let val (x', B') = Syntax.variant_abs' (x, dummyT, B) in + list_comb + (Const (q, dummyT) $ Syntax.mark_boundT (x', T) $ A $ B', ts) end else list_comb (Const (r, dummyT) $ A $ B, ts) | dependent_tr' _ _ = raise Match; \end{ttbox} -The argument {\tt (q,r)} is supplied to the curried function {\tt - dependent_tr'} by a partial application during its installation. We -can set up print translations for both {\tt Pi} and {\tt Sigma} by -including +The argument {\tt (q, r)} is supplied to the curried function {\tt + dependent_tr'} by a partial application during its installation. +For example, we could set up print translations for both {\tt Pi} and +{\tt Sigma} by including \begin{ttbox}\index{*ML section} val print_translation = [("Pi", dependent_tr' ("{\at}PROD", "{\at}->")), ("Sigma", dependent_tr' ("{\at}SUM", "{\at}*"))]; \end{ttbox} -within the {\tt ML} section. The first of these transforms ${\tt Pi}(A, -\mtt{Abs}(x, T, B))$ into $\hbox{\tt{\at}PROD}(x', A, B')$ or -$\hbox{\tt{\at}->}(A, B)$, choosing the latter form if $B$ does not depend -on~$x$. It checks this using \ttindex{loose_bnos}, yet another function -from {\tt Pure/term.ML}. Note that $x'$ is a version of $x$ renamed away -from all names in $B$, and $B'$ is the body $B$ with {\tt Bound} nodes -referring to the {\tt Abs} node replaced by $\ttfct{Free} (x', -\mtt{dummyT})$. +within the {\tt ML} section. The first of these transforms ${\tt + Pi}(A, \mtt{Abs}(x, T, B))$ into $\hbox{\tt{\at}PROD}(x', A, B')$ or +$\hbox{\tt{\at}->}(A, B)$, choosing the latter form if $B$ does not +depend on~$x$. It checks this using \ttindex{loose_bnos}, yet another +function from {\tt Pure/term.ML}. Note that $x'$ is a version of $x$ +renamed away from all names in $B$, and $B'$ is the body $B$ with {\tt + Bound} nodes referring to the {\tt Abs} node replaced by +$\ttfct{Free} (x', \mtt{dummyT})$ (but marked as representing a bound +variable). We must be careful with types here. While types of {\tt Const}s are ignored, type constraints may be printed for some {\tt Free}s and {\tt Var}s if \ttindex{show_types} is set to {\tt true}. Variables of type \ttindex{dummyT} are never printed with constraint, though. The line \begin{ttbox} - let val (x', B') = variant_abs (x, dummyT, B); -\end{ttbox}\index{*variant_abs} + let val (x', B') = Syntax.variant_abs' (x, dummyT, B); +\end{ttbox}\index{*Syntax.variant_abs'} replaces bound variable occurrences in~$B$ by the free variable $x'$ with type {\tt dummyT}. Only the binding occurrence of~$x'$ is given the correct type~{\tt T}, so this is the only place where a type constraint might appear. -\index{translations|)} -\index{syntax!transformations|)} + +Also note that we are responsible to mark free identifiers that +actually represent bound variables. This is achieved by +\ttindex{Syntax.variant_abs'} and \ttindex{Syntax.mark_boundT} above. +Failing to do so may cause these names to be printed in the wrong +style. \index{translations|)} \index{syntax!transformations|)} + + +\section{Token translations} \label{sec:tok_tr} +\index{token translations|(} +% +Isabelle's meta-logic features quite a lot of different kinds of +identifiers, namely {\em class}, {\em tfree}, {\em tvar}, {\em free}, +{\em bound}, {\em var}. One might want to have these printed in +different styles, e.g.\ in bold or italic, or even transcribed into +something more readable like $\alpha, \alpha', \beta$ instead of {\tt + 'a}, {\tt 'aa}, {\tt 'b} for type variables. Token translations +provide a means to such ends, enabling the user to install certain +\ML{} functions associated with any logical \rmindex{token class} and +depending on some \rmindex{print mode}. + +The logical class of identifiers can not necessarily be determined by +its syntactic category, though. For example, consider free vs.\ bound +variables. So Isabelle's pretty printing mechanism, starting from +fully typed terms, has to be careful to preserve this additional +information\footnote{This is done by marking atoms in abstract syntax + trees appropriately. The marks are actually visible by print + translation functions -- they are just special constants applied to + atomic asts, for example \texttt{("_bound" x)}.}. In particular, +user-supplied print translation functions operating on terms have to +be well-behaved in this respect. Free identifiers introduced to +represent bound variables have to be marked appropriately (cf.\ the +example at the end of \S\ref{sec:tr_funs}). + +\medskip Token translations may be installed by declaring the +\ttindex{token_translation} value within the \texttt{ML} section of a +theory definition file: +\begin{ttbox} +val token_translation: (string * string * (string -> string * int)) list +\end{ttbox}\index{token_translation} +The elements of this list are of the form $(m, c, f)$, where $m$ is a +print mode identifier, $c$ a token class, and $f\colon string \to +string \times int$ the actual translation function. Assuming that $x$ +is of identifier class $c$, and print mode $m$ is the first one of all +currently active modes that provide some translation for $c$, then $x$ +is output according to $f(x) = (x', len)$. Thereby $x'$ is the +modified identifier name and $len$ its visual length approximated in +terms of whole characters. Thus $x'$ may include non-printing parts +like control sequences or markup information for typesetting systems. + +%FIXME example (?) + +\index{token translations|)}
--- a/doc-src/Ref/tactic.tex Mon May 05 21:18:01 1997 +0200 +++ b/doc-src/Ref/tactic.tex Tue May 06 12:50:16 1997 +0200 @@ -1,10 +1,10 @@ %% $Id$ \chapter{Tactics} \label{tactics} -\index{tactics|(} -Tactics have type \mltydx{tactic}. They are essentially -functions from theorems to theorem sequences, where the theorems represent -states of a backward proof. Tactics seldom need to be coded from scratch, -as functions; instead they are expressed using basic tactics and tacticals. +\index{tactics|(} Tactics have type \mltydx{tactic}. This is just an +abbreviation for functions from theorems to theorem sequences, where +the theorems represent states of a backward proof. Tactics seldom +need to be coded from scratch, as functions; instead they are +expressed using basic tactics and tacticals. This chapter only presents the primitive tactics. Substantial proofs require the power of simplification (Chapter~\ref{simp-chap}) and classical reasoning @@ -561,26 +561,17 @@ \index{tactics!primitives for coding} A tactic maps theorems to theorem sequences (lazy lists). The type constructor for sequences is called \mltydx{Sequence.seq}. To simplify the -types of tactics and tacticals, Isabelle defines a type of tactics: +types of tactics and tacticals, Isabelle defines a type abbreviations: \begin{ttbox} -datatype tactic = Tactic of thm -> thm Sequence.seq +type tactic = thm -> thm Sequence.seq \end{ttbox} -{\tt Tactic} and {\tt tapply} convert between tactics and functions. The -other operations provide means for coding tactics in a clean style. +The following operations provide means for coding tactics in a clean style. \begin{ttbox} -tapply : tactic * thm -> thm Sequence.seq -Tactic : (thm -> thm Sequence.seq) -> tactic PRIMITIVE : (thm -> thm) -> tactic STATE : (thm -> tactic) -> tactic SUBGOAL : ((term*int) -> tactic) -> int -> tactic \end{ttbox} \begin{ttdescription} -\item[\ttindexbold{tapply}({\it tac}, {\it thm})] -returns the result of applying the tactic, as a function, to {\it thm}. - -\item[\ttindexbold{Tactic} {\it f}] -packages {\it f} as a tactic. - \item[\ttindexbold{PRIMITIVE} $f$] applies $f$ to the proof state and returns the result as a one-element sequence. This packages the meta-rule~$f$ as a tactic.
--- a/doc-src/Ref/tctical.tex Mon May 05 21:18:01 1997 +0200 +++ b/doc-src/Ref/tctical.tex Tue May 06 12:50:16 1997 +0200 @@ -126,9 +126,8 @@ \begin{ttbox} fun TRY tac = tac ORELSE all_tac; -fun REPEAT tac = Tactic - (fn state => tapply((tac THEN REPEAT tac) ORELSE all_tac, - state)); +fun REPEAT tac = + (fn state => ((tac THEN REPEAT tac) ORELSE all_tac) state); \end{ttbox} If $tac$ can return multiple outcomes then so can \hbox{\tt REPEAT $tac$}. Since {\tt REPEAT} uses \ttindex{ORELSE} and not {\tt APPEND} or {\tt
--- a/doc-src/Ref/theories.tex Mon May 05 21:18:01 1997 +0200 +++ b/doc-src/Ref/theories.tex Tue May 06 12:50:16 1997 +0200 @@ -2,12 +2,12 @@ \chapter{Theories, Terms and Types} \label{theories} \index{theories|(}\index{signatures|bold} -\index{reading!axioms|see{{\tt assume_ax}}} -Theories organize the syntax, declarations and axioms of a mathematical -development. They are built, starting from the Pure theory, by extending -and merging existing theories. They have the \ML\ type \mltydx{theory}. -Theory operations signal errors by raising exception \xdx{THEORY}, -returning a message and a list of theories. +\index{reading!axioms|see{{\tt assume_ax}}} Theories organize the +syntax, declarations and axioms of a mathematical development. They +are built, starting from the {\Pure} or {\CPure} theory, by extending +and merging existing theories. They have the \ML\ type +\mltydx{theory}. Theory operations signal errors by raising exception +\xdx{THEORY}, returning a message and a list of theories. Signatures, which contain information about sorts, types, constants and syntax, have the \ML\ type~\mltydx{Sign.sg}. For identification, each @@ -32,11 +32,12 @@ Appendix~\ref{app:TheorySyntax} presents the concrete syntax for theory definitions; here is an explanation of the constituent parts: \begin{description} -\item[{\it theoryDef}] - is the full definition. The new theory is called $id$. It is the union - of the named {\bf parent theories}\indexbold{theories!parent}, possibly - extended with new classes, etc. The basic theory, which contains only - the meta-logic, is called \thydx{Pure}. +\item[{\it theoryDef}] is the full definition. The new theory is + called $id$. It is the union of the named {\bf parent + theories}\indexbold{theories!parent}, possibly extended with new + components. \thydx{Pure} and \thydx{CPure} are the basic theories, + which contain only the meta-logic. They differ just in their + concrete syntax for function applications. Normally each {\it name\/} is an identifier, the name of the parent theory. Quoted strings can be used to document additional file dependencies; see @@ -56,10 +57,9 @@ variables created internally. If omitted, the default sort is the listwise union of the default sorts of the parent theories (i.e.\ their logical intersection). - -\item[$sort$] - is a finite set of classes. A single class $id$ abbreviates the singleton - set {\tt\{}$id${\tt\}}. + +\item[$sort$] is a finite set of classes. A single class $id$ + abbreviates the sort $\ttlbrace id\ttrbrace$. \item[$types$] is a series of type declarations. Each declares a new type constructor @@ -75,19 +75,23 @@ declares a type or constant to be an infix operator of priority $nat$ associating to the left ({\tt infixl}) or right ({\tt infixr}). Only 2-place type constructors can have infix status; an example is {\tt - ('a,'b)~"*"~(infixr~20)}, which expresses binary product types. - -\item[$arities$] - is a series of arity declarations. Each assigns arities to type - constructors. The $name$ must be an existing type constructor, which is - given the additional arity $arity$. + ('a,'b)~"*"~(infixr~20)}, which may express binary product types. -\item[$constDecl$] - is a series of constant declarations. Each new constant $name$ is given - the specified type. The optional $mixfix$ annotations may - attach concrete syntax to the constant. A variant of {\tt consts} is the - {\tt syntax} section\index{*syntax section}, which adds just syntax without - declaring logical constants. +\item[$arities$] is a series of type arity declarations. Each assigns + arities to type constructors. The $name$ must be an existing type + constructor, which is given the additional arity $arity$. + +\item[$consts$] is a series of constant declarations. Each new + constant $name$ is given the specified type. The optional $mixfix$ + annotations may attach concrete syntax to the constant. + +\item[$syntax$] \index{*syntax section}\index{print mode} is a variant + of $consts$ which adds just syntax without actually declaring + logical constants. This gives full control over a theory's context + free grammar. The optional $mode$ specifies the print mode where the + mixfix productions should be added. If there is no \texttt{output} + option given, all productions are also added to the input syntax + (regardless of the print mode). \item[$mixfix$] \index{mixfix declarations} annotations can take three forms: @@ -115,13 +119,28 @@ \item[$rules$] is a series of rule declarations. Each has a name $id$ and the formula is given by the $string$. Rule names must be distinct within any single - theory file. + theory. \item[$defs$] is a series of definitions. They are just like $rules$, except that every $string$ must be a definition (see below for details). \item[$constdefs$] combines the declaration of constants and their definition. The first $string$ is the type, the second the definition. + +\item[$axclass$] \index{*axclass section} defines an + \rmindex{axiomatic type class} as the intersection of existing + classes, with additional axioms holding. Class axioms may not + contain more than one type variable. The class axioms (with implicit + sort constraints added) are bound to the given names. Furthermore a + class introduction rule is generated, which is automatically + employed by $instance$ to prove instantiations of this class. + +\item[$instance$] \index{*instance section} proves class inclusions or + type arities at the logical level and then transfers these to the + type signature. The instantiation is proven and checked properly. + The user has to supply sufficient witness information: theorems + ($longident$), axioms ($string$), or even arbitrary \ML{} tactic + code $verbatim$. \item[$oracle$] links the theory to a trusted external reasoner. It is allowed to create theorems, but each theorem carries a proof object @@ -137,18 +156,19 @@ \subsection{Definitions}\indexbold{definitions} -{\bf Definitions} are intended to express abbreviations. The simplest form of -a definition is $f \equiv t$, where $f$ is a constant. Isabelle also allows a -derived form where the arguments of~$f$ appear on the left, abbreviating a -string of $\lambda$-abstractions. +{\bf Definitions} are intended to express abbreviations. The simplest +form of a definition is $f \equiv t$, where $f$ is a constant. +Isabelle also allows a derived forms where the arguments of~$f$ appear +on the left, abbreviating a string of $\lambda$-abstractions. Isabelle makes the following checks on definitions: \begin{itemize} -\item Arguments (on the left-hand side) must be distinct variables +\item Arguments (on the left-hand side) must be distinct variables. \item All variables on the right-hand side must also appear on the left-hand side. -\item All type variables on the right-hand side must also appear on the - left-hand side; this prohibits definitions such as {\tt zero == length []}. +\item All type variables on the right-hand side must also appear on + the left-hand side; this prohibits definitions such as {\tt + (zero::nat) == length ([]::'a list)}. \item The definition must not be recursive. Most object-logics provide definitional principles that can be used to express recursion safely. \end{itemize} @@ -164,31 +184,26 @@ In order to guarantee principal types~\cite{nipkow-prehofer}, arity declarations must obey two conditions: \begin{itemize} -\item There must be no two declarations $ty :: (\vec{r})c$ and $ty :: - (\vec{s})c$ with $\vec{r} \neq \vec{s}$. For example, the following is - forbidden: +\item There must not be any two declarations $ty :: (\vec{r})c$ and + $ty :: (\vec{s})c$ with $\vec{r} \neq \vec{s}$. For example, this + excludes the following: \begin{ttbox} -types - 'a ty arities - ty :: ({\ttlbrace}logic{\ttrbrace}) logic - ty :: ({\ttlbrace}{\ttrbrace})logic + foo :: ({\ttlbrace}logic{\ttrbrace}) logic + foo :: ({\ttlbrace}{\ttrbrace})logic \end{ttbox} \item If there are two declarations $ty :: (s@1,\dots,s@n)c$ and $ty :: (s@1',\dots,s@n')c'$ such that $c' < c$ then $s@i' \preceq s@i$ must hold for $i=1,\dots,n$. The relationship $\preceq$, defined as \[ s' \preceq s \iff \forall c\in s. \exists c'\in s'.~ c'\le c, \] -expresses that the set of types represented by $s'$ is a subset of the set of -types represented by $s$. For example, the following is forbidden: +expresses that the set of types represented by $s'$ is a subset of the +set of types represented by $s$. Assuming $term \preceq logic$, the +following is forbidden: \begin{ttbox} -classes - term < logic -types - 'a ty arities - ty :: ({\ttlbrace}logic{\ttrbrace})logic - ty :: ({\ttlbrace}{\ttrbrace})term + foo :: ({\ttlbrace}logic{\ttrbrace})logic + foo :: ({\ttlbrace}{\ttrbrace})term \end{ttbox} \end{itemize} @@ -219,14 +234,15 @@ suppresses the deletion of temporary files. \end{ttdescription} % -Each theory definition must reside in a separate file. Let the file {\it - T}{\tt.thy} contain the definition of a theory called~$T$, whose parent -theories are $TB@1$ \dots $TB@n$. Calling \ttindexbold{use_thy}~{\tt"{\it - T\/}"} reads the file {\it T}{\tt.thy}, writes a temporary \ML{} -file {\tt.{\it T}.thy.ML}, and reads the latter file. Recursive {\tt - use_thy} calls load those parent theories that have not been loaded -previously; the recursive calls may continue to any depth. One {\tt use_thy} -call can read an entire logic provided all theories are linked appropriately. +Each theory definition must reside in a separate file. Let the file +{\it T}{\tt.thy} contain the definition of a theory called~$T$, whose +parent theories are $TB@1$ \dots $TB@n$. Calling +\ttindex{use_thy}~{\tt"{\it T\/}"} reads the file {\it T}{\tt.thy}, +writes a temporary \ML{} file {\tt.{\it T}.thy.ML}, and reads the +latter file. Recursive {\tt use_thy} calls load those parent theories +that have not been loaded previously; the recursive calls may continue +to any depth. One {\tt use_thy} call can read an entire logic +provided all theories are linked appropriately. The result is an \ML\ structure~$T$ containing at least a component {\tt thy} for the new theory and components for each of the rules. The structure also @@ -234,9 +250,9 @@ {\tt.{\it T}.thy.ML} is then deleted if {\tt delete_tmpfiles} is set to {\tt true} and no errors occurred. -Finally the file {\it T}{\tt.ML} is read, if it exists. This file normally -begins with the declaration {\tt open~$T$} and contains proofs involving -the new theory. +Finally the file {\it T}{\tt.ML} is read, if it exists. Since the +structure $T$ is automatically open in this context, proof scripts may +(or even should) refer to its components by unqualified names. Some applications construct theories directly by calling \ML\ functions. In this situation there is no {\tt.thy} file, only an {\tt.ML} file. The @@ -284,17 +300,18 @@ \end{ttdescription} -\goodbreak -\subsection{Important note for Poly/ML users}\index{Poly/{\ML} compiler} -The theory mechanism depends upon reference variables. At the end of a -Poly/\ML{} session, the contents of references are lost unless they are -declared in the current database. In particular, assignments to references -of the {\tt Pure} database are lost, including all information about loaded -theories. To avoid losing this information simply call -\begin{ttbox} -init_thy_reader(); -\end{ttbox} -when building the new database. +%FIXME remove +%\goodbreak +%\subsection{Important note for Poly/ML users}\index{Poly/{\ML} compiler} +%The theory mechanism depends upon reference variables. At the end of a +%Poly/\ML{} session, the contents of references are lost unless they are +%declared in the current database. In particular, assignments to references +%of the {\tt Pure} database are lost, including all information about loaded +%theories. To avoid losing this information simply call +%\begin{ttbox} +%init_thy_reader(); +%\end{ttbox} +%when building the new database. \subsection{*Pseudo theories}\label{sec:pseudo-theories} @@ -346,9 +363,9 @@ The resulting theory ensures that {\tt update} reloads {\tt orphan} whenever it reloads one of the $A@i$. -For an extensive example of how this technique can be used to link lots of -theory files and load them by just a few {\tt use_thy} calls, consult the -sources of \ZF{} set theory. +For an extensive example of how this technique can be used to link +lots of theory files and load them by just a few {\tt use_thy} calls +see the sources of one of the major object-logics (e.g.\ \ZF). @@ -384,17 +401,21 @@ \hbox{\verb'?a=?b ==> ?b=?a [!!a b. a=b ==> b=a]'} \end{ttdescription} -\subsection{Building a theory} +\subsection{*Building a theory} \label{BuildingATheory} \index{theories!constructing|bold} \begin{ttbox} -pure_thy : theory +Pure.thy : theory +CPure.thy : theory merge_theories : theory * theory -> theory \end{ttbox} -\begin{ttdescription} -\item[\ttindexbold{pure_thy}] contains just the syntax and signature - of the meta-logic. There are no axioms: meta-level inferences are carried - out by \ML\ functions. +\begin{description} +\item[\ttindexbold{Pure.thy}, \ttindexbold{CPure.thy}] contain the + syntax and signature of the meta-logic. There are no axioms: + meta-level inferences are carried out by \ML\ functions. The two + \Pure s just differ in their concrete syntax of function + application: $t(u@1, \ldots, u@n)$ vs.\ $t\,u@1,\ldots\,u@n$. + \item[\ttindexbold{merge_theories} ($thy@1$, $thy@2$)] merges the two theories $thy@1$ and $thy@2$. The resulting theory contains all of the syntax, signature and axioms of the constituent theories. Merging theories @@ -419,7 +440,7 @@ %\begin{ttbox} %Attempt to merge different versions of theory: \(T\) %\end{ttbox} -\end{ttdescription} +\end{description} %% FIXME %\item [\ttindexbold{extend_theory} $thy$ {\tt"}$T${\tt"} @@ -456,6 +477,7 @@ \subsection{Inspecting a theory}\label{sec:inspct-thy} \index{theories!inspecting|bold} \begin{ttbox} +print_syntax : theory -> unit print_theory : theory -> unit axioms_of : theory -> (string * thm) list thms_of : theory -> (string * thm) list @@ -465,9 +487,12 @@ \end{ttbox} These provide means of viewing a theory's components. \begin{ttdescription} -\item[\ttindexbold{print_theory} $thy$] -prints the contents of $thy$ excluding the syntax related parts (which are -shown by {\tt print_syntax}). The output is quite verbose. +\item[\ttindexbold{print_syntax} $thy$] prints the syntax of $thy$ + (grammar, macros, translation functions etc., see + page~\pageref{pg:print_syn} for more details). + +\item[\ttindexbold{print_theory} $thy$] prints the logical parts of + $thy$, excluding the syntax. \item[\ttindexbold{axioms_of} $thy$] returns the additional axioms of the most recent extend node of~$thy$. @@ -487,223 +512,224 @@ with~$thy$. \end{ttdescription} - -\section{Generating HTML documents} -\index{HTML|bold} - -Isabelle is able to make HTML documents that show a theory's -definition, the theorems proved in its ML file and the relationship -with its ancestors and descendants. HTML stands for Hypertext Markup -Language and is used in the World Wide Web to represent text -containing images and links to other documents. Web browsers like -{\tt xmosaic} or {\tt netscape} are used to view these documents. - -Besides the three HTML files that are made for every theory -(definition and theorems, ancestors, descendants), Isabelle stores -links to all theories in an index file. These indexes are themself -linked with other indexes to represent the hierarchic structure of -Isabelle's logics. - -To make HTML files for logics that are part of the Isabelle -distribution, simply set the shell environment variable {\tt -MAKE_HTML} before compiling a logic. This works for single logics as -well as for the shell script {\tt make-all} (see -\ref{sec:shell-scripts}). To make HTML files for {\tt FOL} using a -{\tt csh} style shell, the following commands can be used: - -\begin{ttbox} -cd FOL -setenv MAKE_HTML -make -\end{ttbox} - -The databases made this way do not differ from the ones made with an -unset {\tt MAKE_HTML}; in particular no HTML files are written if the -database is used to manually load a theory. - -As you will see below, the HTML generation is controlled by a boolean -reference variable. If you want to make databases which define this -variable's value as {\tt true} and where therefore HTML files are -written each time {\tt use_thy} is invoked, you have to set {\tt -MAKE_HTML} to ``{\tt true}'': - -\begin{ttbox} -cd FOL -setenv MAKE_HTML true -make -\end{ttbox} - -All theories loaded from within the {\tt FOL} database and all -databases derived from it will now cause HTML files to be written. -This behaviour can be changed by assigning a value of {\tt false} to -the boolean reference variable {\tt make_html}. Be careful when making -such databases publicly available since it means that your users will -generate HTML files though they might not intend to do so. - -As some of Isabelle's logics are based on others (e.g. {\tt ZF} on -{\tt FOL}) and because the HTML files list a theory's ancestors, you -should not make HTML files for a logic if the HTML files for the base -logic do not exist. Otherwise some of the hypertext links might point -to non-existing documents. - -The entry point to all logics is the {\tt index.html} file located in -Isabelle's main directory. You can also access a HTML version of the -distribution package at - -\begin{ttbox} -http://www4.informatik.tu-muenchen.de/~nipkow/isabelle -\end{ttbox} - - -\subsection*{Manual HTML generation} - -To manually control the generation of HTML files the following -commands and reference variables are used: - -\begin{ttbox} -init_html : unit -> unit -make_html : bool ref -finish_html : unit -> unit -\end{ttbox} - -\begin{ttdescription} -\item[\ttindexbold{init_html}] -activates the HTML facility. It stores the current working directory -as the place where the {\tt index.html} file for all theories loaded -afterwards will be stored. - -\item[\ttindexbold{make_html}] -is a boolean reference variable read by {\tt use_thy} and other -functions to decide whether HTML files should be made. After you have -used {\tt init_html} you can manually change {\tt make_html}'s value -to temporarily disable HTML generation. - -\item[\ttindexbold{finish_html}] -has to be called after all theories have been read that should be -listed in the current {\tt index.html} file. It reads a temporary -file with information about the theories read since the last use of -{\tt init_html} and makes the {\tt index.html} file. If {\tt -make_html} is {\tt false} nothing is done. - -The indexes made by this function also contain a link to the {\tt -README} file if there exists one in the directory where the index is -stored. If there's a {\tt README.html} file it is used instead of -{\tt README}. - -\end{ttdescription} +%FIXME move to sysman! +%\section{Generating HTML documents} +%\index{HTML|bold} +% +%Isabelle is able to make HTML documents that show a theory's +%definition, the theorems proved in its ML file and the relationship +%with its ancestors and descendants. HTML stands for Hypertext Markup +%Language and is used in the World Wide Web to represent text +%containing images and links to other documents. Web browsers like +%{\tt xmosaic} or {\tt netscape} are used to view these documents. +% +%Besides the three HTML files that are made for every theory +%(definition and theorems, ancestors, descendants), Isabelle stores +%links to all theories in an index file. These indexes are themself +%linked with other indexes to represent the hierarchic structure of +%Isabelle's logics. +% +%To make HTML files for logics that are part of the Isabelle +%distribution, simply set the shell environment variable {\tt +%MAKE_HTML} before compiling a logic. This works for single logics as +%well as for the shell script {\tt make-all} (see +%\ref{sec:shell-scripts}). To make HTML files for {\tt FOL} using a +%{\tt csh} style shell, the following commands can be used: +% +%\begin{ttbox} +%cd FOL +%setenv MAKE_HTML +%make +%\end{ttbox} +% +%The databases made this way do not differ from the ones made with an +%unset {\tt MAKE_HTML}; in particular no HTML files are written if the +%database is used to manually load a theory. +% +%As you will see below, the HTML generation is controlled by a boolean +%reference variable. If you want to make databases which define this +%variable's value as {\tt true} and where therefore HTML files are +%written each time {\tt use_thy} is invoked, you have to set {\tt +%MAKE_HTML} to ``{\tt true}'': +% +%\begin{ttbox} +%cd FOL +%setenv MAKE_HTML true +%make +%\end{ttbox} +% +%All theories loaded from within the {\tt FOL} database and all +%databases derived from it will now cause HTML files to be written. +%This behaviour can be changed by assigning a value of {\tt false} to +%the boolean reference variable {\tt make_html}. Be careful when making +%such databases publicly available since it means that your users will +%generate HTML files though they might not intend to do so. +% +%As some of Isabelle's logics are based on others (e.g. {\tt ZF} on +%{\tt FOL}) and because the HTML files list a theory's ancestors, you +%should not make HTML files for a logic if the HTML files for the base +%logic do not exist. Otherwise some of the hypertext links might point +%to non-existing documents. +% +%The entry point to all logics is the {\tt index.html} file located in +%Isabelle's main directory. You can also access a HTML version of the +%distribution package at +% +%\begin{ttbox} +%http://www4.informatik.tu-muenchen.de/~nipkow/isabelle/ +%\end{ttbox} +% +% +%\subsection*{Manual HTML generation} +% +%To manually control the generation of HTML files the following +%commands and reference variables are used: +% +%\begin{ttbox} +%init_html : unit -> unit +%make_html : bool ref +%finish_html : unit -> unit +%\end{ttbox} +% +%\begin{ttdescription} +%\item[\ttindexbold{init_html}] +%activates the HTML facility. It stores the current working directory +%as the place where the {\tt index.html} file for all theories loaded +%afterwards will be stored. +% +%\item[\ttindexbold{make_html}] +%is a boolean reference variable read by {\tt use_thy} and other +%functions to decide whether HTML files should be made. After you have +%used {\tt init_html} you can manually change {\tt make_html}'s value +%to temporarily disable HTML generation. +% +%\item[\ttindexbold{finish_html}] +%has to be called after all theories have been read that should be +%listed in the current {\tt index.html} file. It reads a temporary +%file with information about the theories read since the last use of +%{\tt init_html} and makes the {\tt index.html} file. If {\tt +%make_html} is {\tt false} nothing is done. +% +%The indexes made by this function also contain a link to the {\tt +%README} file if there exists one in the directory where the index is +%stored. If there's a {\tt README.html} file it is used instead of +%{\tt README}. +% +%\end{ttdescription} +% +%The above functions could be used in the following way: +% +%\begin{ttbox} +%init_html(); +%{\out Setting path for index.html to "/home/clasohm/isabelle/HOL"} +%use_thy "List"; +%\dots +%finish_html(); +%\end{ttbox} +% +%Please note that HTML files are made only for those theories that are +%read while {\tt make_html} is {\tt true}. These files may contain +%links to theories that were read with a {\tt false} {\tt make_html} +%and therefore point to non-existing files. +% +% +%\subsection*{Extending or adding a logic} +% +%If you add a new subdirectory to Isabelle's logics (or add a completly +%new logic), you would have to call {\tt init_html} at the start of every +%directory's {\tt ROOT.ML} file and {\tt finish_html} at the end of +%it. This is automatically done if you use +% +%\begin{ttbox}\index{use_dir} +%use_dir : string -> unit +%\end{ttbox} +% +%This function takes a path as its parameter, changes the working +%directory, calls {\tt init_html} if {\tt make_html} is {\tt true}, +%executes {\tt ROOT.ML}, and calls {\tt finish_html}. The {\tt +%index.html} file written in this directory will be automatically +%linked to the first index found in the (recursively searched) +%superdirectories. +% +%Instead of adding something like +% +%\begin{ttbox} +%use"ex/ROOT.ML"; +%\end{ttbox} +% +%to the logic's makefile you have to use this: +% +%\begin{ttbox} +%use_dir"ex"; +%\end{ttbox} +% +%Since {\tt use_dir} calls {\tt init_html} only if {\tt make_html} is +%{\tt true} the generation of HTML files depends on the value this +%reference variable has. It can either be inherited from the used \ML{} +%database or set in the makefile before {\tt use_dir} is invoked, +%e.g. to set it's value according to the environment variable {\tt +%MAKE_HTML}. +% +%The generated HTML files contain all theorems that were proved in the +%theory's \ML{} file with {\tt qed}, {\tt qed_goal} and {\tt qed_goalw}, +%or stored with {\tt bind_thm} and {\tt store_thm}. Additionally there +%is a hypertext link to the whole \ML{} file. +% +%You can add section headings to the list of theorems by using +% +%\begin{ttbox}\index{use_dir} +%section: string -> unit +%\end{ttbox} +% +%in a theory's ML file, which converts a plain string to a HTML +%heading and inserts it before the next theorem proved or stored with +%one of the above functions. If {\tt make_html} is {\tt false} nothing +%is done. +% +% +%\subsection*{Using someone else's database} +% +%To make them independent from their storage place, the HTML files only +%contain relative paths which are derived from absolute ones like the +%current working directory, {\tt gif_path} or {\tt base_path}. The +%latter two are reference variables which are initialized at the time +%when the {\tt Pure} database is made. Because you need write access +%for the current directory to make HTML files and therefore (probably) +%generate them in your home directory, the absolute {\tt base_path} is +%not correct if you use someone else's database or a database derived +%from it. +% +%In that case you first should set {\tt base_path} to the value of {\em +%your} Isabelle main directory, i.e. the directory that contains the +%subdirectories where standard logics like {\tt FOL} and {\tt HOL} or +%your own logics are stored. If you do not do this, the generated HTML +%files will still be usable but may contain incomplete titles and lack +%some hypertext links. +% +%It's also a good idea to set {\tt gif_path} which points to the +%directory containing two GIF images used in the HTML +%documents. Normally this is the {\tt Tools} subdirectory of Isabelle's +%main directory. While its value in general is still valid, your HTML +%files would depend on files not owned by you. This prevents you from +%changing the location of the HTML files (as they contain relative +%paths) and also causes trouble if the database's maker (re)moves the +%GIFs. +% +%Here's what you should do before invoking {\tt init_html} using +%someone else's \ML{} database: +% +%\begin{ttbox} +%base_path := "/home/smith/isabelle"; +%gif_path := "/home/smith/isabelle/Tools"; +%init_html(); +%\dots +%\end{ttbox} -The above functions could be used in the following way: - -\begin{ttbox} -init_html(); -{\out Setting path for index.html to "/home/clasohm/isabelle/HOL"} -use_thy "List"; -\dots -finish_html(); -\end{ttbox} - -Please note that HTML files are made only for those theories that are -read while {\tt make_html} is {\tt true}. These files may contain -links to theories that were read with a {\tt false} {\tt make_html} -and therefore point to non-existing files. - - -\subsection*{Extending or adding a logic} - -If you add a new subdirectory to Isabelle's logics (or add a completly -new logic), you would have to call {\tt init_html} at the start of every -directory's {\tt ROOT.ML} file and {\tt finish_html} at the end of -it. This is automatically done if you use - -\begin{ttbox}\index{use_dir} -use_dir : string -> unit -\end{ttbox} - -This function takes a path as its parameter, changes the working -directory, calls {\tt init_html} if {\tt make_html} is {\tt true}, -executes {\tt ROOT.ML}, and calls {\tt finish_html}. The {\tt -index.html} file written in this directory will be automatically -linked to the first index found in the (recursively searched) -superdirectories. - -Instead of adding something like - -\begin{ttbox} -use"ex/ROOT.ML"; -\end{ttbox} - -to the logic's makefile you have to use this: - -\begin{ttbox} -use_dir"ex"; -\end{ttbox} - -Since {\tt use_dir} calls {\tt init_html} only if {\tt make_html} is -{\tt true} the generation of HTML files depends on the value this -reference variable has. It can either be inherited from the used \ML{} -database or set in the makefile before {\tt use_dir} is invoked, -e.g. to set it's value according to the environment variable {\tt -MAKE_HTML}. - -The generated HTML files contain all theorems that were proved in the -theory's \ML{} file with {\tt qed}, {\tt qed_goal} and {\tt qed_goalw}, -or stored with {\tt bind_thm} and {\tt store_thm}. Additionally there -is a hypertext link to the whole \ML{} file. - -You can add section headings to the list of theorems by using - -\begin{ttbox}\index{use_dir} -section: string -> unit -\end{ttbox} - -in a theory's ML file, which converts a plain string to a HTML -heading and inserts it before the next theorem proved or stored with -one of the above functions. If {\tt make_html} is {\tt false} nothing -is done. - - -\subsection*{Using someone else's database} - -To make them independent from their storage place, the HTML files only -contain relative paths which are derived from absolute ones like the -current working directory, {\tt gif_path} or {\tt base_path}. The -latter two are reference variables which are initialized at the time -when the {\tt Pure} database is made. Because you need write access -for the current directory to make HTML files and therefore (probably) -generate them in your home directory, the absolute {\tt base_path} is -not correct if you use someone else's database or a database derived -from it. - -In that case you first should set {\tt base_path} to the value of {\em -your} Isabelle main directory, i.e. the directory that contains the -subdirectories where standard logics like {\tt FOL} and {\tt HOL} or -your own logics are stored. If you do not do this, the generated HTML -files will still be usable but may contain incomplete titles and lack -some hypertext links. - -It's also a good idea to set {\tt gif_path} which points to the -directory containing two GIF images used in the HTML -documents. Normally this is the {\tt Tools} subdirectory of Isabelle's -main directory. While its value in general is still valid, your HTML -files would depend on files not owned by you. This prevents you from -changing the location of the HTML files (as they contain relative -paths) and also causes trouble if the database's maker (re)moves the -GIFs. - -Here's what you should do before invoking {\tt init_html} using -someone else's \ML{} database: - -\begin{ttbox} -base_path := "/home/smith/isabelle"; -gif_path := "/home/smith/isabelle/Tools"; -init_html(); -\dots -\end{ttbox} \section{Terms} \index{terms|bold} Terms belong to the \ML\ type \mltydx{term}, which is a concrete datatype -with six constructors: there are six kinds of term. +with six constructors: \begin{ttbox} type indexname = string * int; infix 9 $; @@ -844,7 +870,7 @@ read_cterm : Sign.sg -> string * typ -> cterm cert_axm : Sign.sg -> string * term -> string * term read_axm : Sign.sg -> string * string -> string * term -rep_cterm : cterm -> \{T:typ, t:term, sign:Sign.sg, maxidx:int\} +rep_cterm : cterm -> {\ttlbrace}T:typ, t:term, sign:Sign.sg, maxidx:int\ttrbrace \end{ttbox} \begin{ttdescription} \item[\ttindexbold{cterm_of} $sign$ $t$] \index{signatures} @@ -932,7 +958,7 @@ \subsection{Making and inspecting certified types} \begin{ttbox} ctyp_of : Sign.sg -> typ -> ctyp -rep_ctyp : ctyp -> \{T: typ, sign: Sign.sg\} +rep_ctyp : ctyp -> {\ttlbrace}T: typ, sign: Sign.sg\ttrbrace \end{ttbox} \begin{ttdescription} \item[\ttindexbold{ctyp_of} $sign$ $T$] \index{signatures} @@ -968,7 +994,7 @@ \item[\ttindexbold{set_oracle} $fn$ $thy$] sets the oracle of theory $thy$ to be $fn$. It is seldom called explicitly, as there is syntax for oracles in - theory files. A theory can have at most one oracle. + theory files. Any theory node can have at most one oracle. \end{ttdescription} A curious feature of {\ML} exceptions is that they are ordinary constructors.
--- a/doc-src/Ref/theory-syntax.tex Mon May 05 21:18:01 1997 +0200 +++ b/doc-src/Ref/theory-syntax.tex Tue May 06 12:50:16 1997 +0200 @@ -4,45 +4,59 @@ \newlinechar=-1 %mathsing.sty sets \newlinechar=`\|, which would cause mayhem \chapter{Syntax of Isabelle Theories}\label{app:TheorySyntax} -Chapter~\ref{sec:ref-defining-theories} explains the meanings of these -constructs. The syntax obeys the following conventions: + +Below we present the full syntax of theory definition files as +provided by {\Pure} Isabelle --- object-logics may add their own +sections. \S\ref{sec:ref-defining-theories} explains the meanings of +these constructs. The syntax obeys the following conventions: \begin{itemize} \item {\tt Typewriter font} denotes terminal symbols. -\item $id$, $tid$, $nat$, $string$ and $text$ are the lexical classes of - identifiers, type identifiers, natural numbers, \ML\ strings (with their - quotation marks, but without the need for \verb$\$ at the end of a line and - the beginning of the next line) and arbitrary \ML\ text. The first three - are fully defined in \iflabelundefined{Defining-Logics}% - {{\it The Isabelle Reference Manual}, chapter `Defining Logics'}% - {Chap.\ts\ref{Defining-Logics}}. + +\item $id$, $tid$, $nat$, $string$ and $longident$ are the lexical + classes of identifiers, type identifiers, natural numbers, quoted + strings (without the need for \verb$\$\dots\verb$\$ between lines) + and long qualified \ML{} identifiers. + The categories $id$, $tid$, $nat$ are fully defined in \iflabelundefined{Defining-Logics}% + {{\it The Isabelle Reference Manual}, chapter `Defining Logics'}% + {\S\ref{Defining-Logics}}. + +\item $text$ is all text from the current position to the end of file, + $verbatim$ is any text enclosed in \verb.{|.\dots\verb.|}. + +\item Comments in theories take the form {\tt (*}\dots{\tt*)} and may + be nested, just as in \ML. \end{itemize} -Comments in theories take the form {\tt (*} {\it text\/} {\tt*)}, where -{\it text\/} should not contain the character sequence {\tt*)}. \begin{rail} -theoryDef : id '=' (name + '+') ('+' extension | ()); +theoryDef : id '=' (name + '+') ('+' extension | ()) + ; -name: id | string; +name: id | string + ; -extension : (section +) 'end' ( () | ml ); +extension : (section +) 'end' ( () | ml ) + ; section : classes | default | types | arities | consts - | constdefs + | syntax | trans | defs + | constdefs | rules + | axclass + | instance | oracle ; -classes : 'classes' ( ( id ( () - | '<' (id + ',') - ) - ) + ) +classes : 'classes' ( classDecl + ) + ; + +classDecl : (id (() | '<' (id + ','))) ; default : 'default' sort @@ -55,7 +69,8 @@ types : 'types' ( ( typeDecl ( () | '(' infix ')' ) ) + ) ; -infix : ( 'infixr' | 'infixl' ) nat; +infix : ( 'infixr' | 'infixl' ) nat + ; typeDecl : typevarlist name ( () | '=' ( string | type ) ); @@ -66,30 +81,32 @@ '[' ( type + "," ) ']' '=>' type; simpleType: id | ( tid ( () | '::' id ) ) | - '(' ( type + "," ) ')' id | simpleType id; + '(' ( type + "," ) ')' id | simpleType id + ; - -arities : 'arities' ((( name + ',' ) '::' arity ) + ) +arities : 'arities' ((name + ',') '::' arity +) ; -arity : ( () - | '(' (sort + ',') ')' - ) id - ; +arity : ( () | '(' (sort + ',') ')' ) sort + ; + +consts : 'consts' ( mixfixConstDecl + ) + ; +syntax : 'syntax' (() | mode) ( mixfixConstDecl + ); -consts : 'consts' ( ( constDecl ( () | ( '(' mixfix ')' ) ) ) + ) - ; +mode : '(' name (() | 'output') ')' + ; + +mixfixConstDecl : constDecl (() | ( '(' mixfix ')' )) + ; constDecl : ( name + ',') '::' (string | type); mixfix : string ( () | ( () | ('[' (nat + ',') ']')) nat ) - | infix + | ( 'infixr' | 'infixl' ) (() | string) nat | 'binder' string nat ; -constdefs : 'constdefs' (id '::' (string | type) string +) - ; - trans : 'translations' ( pat ( '==' | '=>' | '<=' ) pat + ) ; @@ -101,6 +118,18 @@ defs : 'defs' (( id string ) + ) ; +constdefs : 'constdefs' (id '::' (string | type) string +) + ; + +axclass : 'axclass' classDecl (() | ( id string ) +) + ; + +instance : 'instance' ( name '<' name | name '::' arity) witness + ; + +witness : (() | ((string | longident) + ',')) (() | verbatim) + ; + oracle : 'oracle' name ;
--- a/doc-src/Ref/thm.tex Mon May 05 21:18:01 1997 +0200 +++ b/doc-src/Ref/thm.tex Tue May 06 12:50:16 1997 +0200 @@ -2,15 +2,16 @@ \chapter{Theorems and Forward Proof} \index{theorems|(} -Theorems, which represent the axioms, theorems and rules of object-logics, -have type \mltydx{thm}. This chapter begins by describing operations that -print theorems and that join them in forward proof. Most theorem -operations are intended for advanced applications, such as programming new -proof procedures. Many of these operations refer to signatures, certified -terms and certified types, which have the \ML{} types {\tt Sign.sg}, {\tt - Sign.cterm} and {\tt Sign.ctyp} and are discussed in -Chapter~\ref{theories}. Beginning users should ignore such complexities ---- and skip all but the first section of this chapter. +Theorems, which represent the axioms, theorems and rules of +object-logics, have type \mltydx{thm}. This chapter begins by +describing operations that print theorems and that join them in +forward proof. Most theorem operations are intended for advanced +applications, such as programming new proof procedures. Many of these +operations refer to signatures, certified terms and certified types, +which have the \ML{} types {\tt Sign.sg}, {\tt cterm} and {\tt ctyp} +and are discussed in Chapter~\ref{theories}. Beginning users should +ignore such complexities --- and skip all but the first section of +this chapter. The theorem operations do not print error messages. Instead, they raise exception~\xdx{THM}\@. Use \ttindex{print_exn} to display @@ -134,9 +135,9 @@ \subsection{Instantiating a theorem} \index{instantiation} \begin{ttbox} -read_instantiate : (string*string)list -> thm -> thm -read_instantiate_sg : Sign.sg -> (string*string)list -> thm -> thm -cterm_instantiate : (Sign.cterm*Sign.cterm)list -> thm -> thm +read_instantiate : (string * string) list -> thm -> thm +read_instantiate_sg : Sign.sg -> (string * string) list -> thm -> thm +cterm_instantiate : (cterm * cterm) list -> thm -> thm \end{ttbox} These meta-rules instantiate type and term unknowns in a theorem. They are occasionally useful. They can prevent difficulties with higher-order @@ -177,10 +178,11 @@ rule_by_tactic : tactic -> thm -> thm \end{ttbox} \begin{ttdescription} -\item[\ttindexbold{standard} $thm$] -puts $thm$ into the standard form of object-rules. It discharges all -meta-assumptions, replaces free variables by schematic variables, and -renames schematic variables to have subscript zero. +\item[\ttindexbold{standard} $thm$] puts $thm$ into the standard form + of object-rules. It discharges all meta-assumptions, replaces free + variables by schematic variables, renames schematic variables to + have subscript zero, also strips outer (meta) quantifiers and + removes dangling sort hypotheses. \item[\ttindexbold{zero_var_indexes} $thm$] makes all schematic variables have subscript zero, renaming them to avoid @@ -213,8 +215,8 @@ stamps_of_thy : thm -> string ref list theory_of_thm : thm -> theory dest_state : thm*int -> (term*term)list*term list*term*term -rep_thm : thm -> \{prop: term, hyps: term list, der: deriv, - maxidx: int, sign: Sign.sg, shyps: sort list\} +rep_thm : thm -> {\ttlbrace}prop: term, hyps: term list, der: deriv, + maxidx: int, sign: Sign.sg, shyps: sort list\ttrbrace \end{ttbox} \begin{ttdescription} \item[\ttindexbold{concl_of} $thm$] @@ -267,10 +269,11 @@ this may result in some sorts becoming {\em empty\/}: where one cannot exhibit a type belonging to it because certain axioms are unsatisfiable. -If a theorem contain a type variable whose sort is empty, then that theorem -has no instances. In effect, it asserts nothing. But what if it is used to -prove another theorem that no longer involves that sort? The latter theorem -holds only if the sort is non-empty. +If a theorem contains a type variable that is constrained by an empty +sort, then that theorem has no instances. It is basically an instance +of {\em ex falso quodlibet}. But what if it is used to prove another +theorem that no longer involves that sort? The latter theorem holds +only if under an additional non-emptiness assumption. Therefore, Isabelle's theorems carry around sort hypotheses. The {\tt shyps} field is a list of sorts occurring in type variables in the current @@ -334,9 +337,11 @@ of the assumptions of the premises. Formally, the system works with assertions of the form \[ \phi \quad [\phi@1,\ldots,\phi@n], \] -where $\phi@1$,~\ldots,~$\phi@n$ are the assumptions. Do not confuse -meta-level assumptions with the object-level assumptions in a subgoal, -which are represented in the meta-logic using~$\Imp$. +where $\phi@1$,~\ldots,~$\phi@n$ are the assumptions. This can be +also read as a single conclusion sequent $\phi@1,\ldots,\phi@n \vdash +\phi$. Do not confuse meta-level assumptions with the object-level +assumptions in a subgoal, which are represented in the meta-logic +using~$\Imp$. Each theorem has a signature. Certified terms have a signature. When a rule takes several premises and certified terms, it merges the signatures @@ -385,7 +390,7 @@ \subsection{Assumption rule} \index{meta-assumptions} \begin{ttbox} -assume: Sign.cterm -> thm +assume: cterm -> thm \end{ttbox} \begin{ttdescription} \item[\ttindexbold{assume} $ct$] @@ -397,8 +402,8 @@ \subsection{Implication rules} \index{meta-implication} \begin{ttbox} -implies_intr : Sign.cterm -> thm -> thm -implies_intr_list : Sign.cterm list -> thm -> thm +implies_intr : cterm -> thm -> thm +implies_intr_list : cterm list -> thm -> thm implies_intr_hyps : thm -> thm implies_elim : thm -> thm -> thm implies_elim_list : thm -> thm list -> thm @@ -450,7 +455,7 @@ \subsection{Equality rules} \index{meta-equality} \begin{ttbox} -reflexive : Sign.cterm -> thm +reflexive : cterm -> thm symmetric : thm -> thm transitive : thm -> thm -> thm \end{ttbox} @@ -469,9 +474,9 @@ \subsection{The $\lambda$-conversion rules} \index{lambda calc@$\lambda$-calculus} \begin{ttbox} -beta_conversion : Sign.cterm -> thm +beta_conversion : cterm -> thm extensional : thm -> thm -abstract_rule : string -> Sign.cterm -> thm -> thm +abstract_rule : string -> cterm -> thm -> thm combination : thm -> thm -> thm \end{ttbox} There is no rule for $\alpha$-conversion because Isabelle regards @@ -503,9 +508,9 @@ \subsection{Forall introduction rules} \index{meta-quantifiers} \begin{ttbox} -forall_intr : Sign.cterm -> thm -> thm -forall_intr_list : Sign.cterm list -> thm -> thm -forall_intr_frees : thm -> thm +forall_intr : cterm -> thm -> thm +forall_intr_list : cterm list -> thm -> thm +forall_intr_frees : thm -> thm \end{ttbox} \begin{ttdescription} @@ -526,10 +531,10 @@ \subsection{Forall elimination rules} \begin{ttbox} -forall_elim : Sign.cterm -> thm -> thm -forall_elim_list : Sign.cterm list -> thm -> thm -forall_elim_var : int -> thm -> thm -forall_elim_vars : int -> thm -> thm +forall_elim : cterm -> thm -> thm +forall_elim_list : cterm list -> thm -> thm +forall_elim_var : int -> thm -> thm +forall_elim_vars : int -> thm -> thm \end{ttbox} \begin{ttdescription} @@ -552,8 +557,7 @@ \subsection{Instantiation of unknowns} \index{instantiation} \begin{ttbox} -instantiate: (indexname*Sign.ctyp)list * - (Sign.cterm*Sign.cterm)list -> thm -> thm +instantiate: (indexname * ctyp) list * (cterm * cterm) list -> thm -> thm \end{ttbox} \begin{ttdescription} \item[\ttindexbold{instantiate} ($tyinsts$, $insts$) $thm$] @@ -656,10 +660,10 @@ \subsection{Other meta-rules} \begin{ttbox} -trivial : Sign.cterm -> thm +trivial : cterm -> thm lift_rule : (thm * int) -> thm -> thm rename_params_rule : string list * int -> thm -> thm -rewrite_cterm : thm list -> Sign.cterm -> thm +rewrite_cterm : thm list -> cterm -> thm flexflex_rule : thm -> thm Sequence.seq \end{ttbox} \begin{ttdescription} @@ -730,8 +734,7 @@ record: \begin{ttbox} #der (rep_thm conjI); -{\out Join (Theorem ({ProtoPure, CPure, HOL},"conjI"),} -{\out [Join (MinProof,[])]) : deriv} +{\out Join (Theorem "conjI", [Join (MinProof,[])]) : deriv} \end{ttbox} This proof object identifies a labelled theorem, {\tt conjI}, whose underlying proof has not been recorded; all we have is {\tt MinProof}.