# HG changeset patch # User wenzelm # Date 1010263298 -3600 # Node ID a107eeffd5572d74f64fe81fd7041ccea6e1696a # Parent 39b93da27bc939c86433386119c5297e3eaac6db updated; diff -r 39b93da27bc9 -r a107eeffd557 doc-src/TutorialI/Documents/document/Documents.tex --- a/doc-src/TutorialI/Documents/document/Documents.tex Sat Jan 05 21:41:11 2002 +0100 +++ b/doc-src/TutorialI/Documents/document/Documents.tex Sat Jan 05 21:41:38 2002 +0100 @@ -9,11 +9,10 @@ % \begin{isamarkuptext}% Concerning Isabelle's ``inner'' language of simply-typed \isa{{\isasymlambda}}-calculus, the core concept of Isabelle's elaborate infrastructure - for concrete syntax is that of general \emph{mixfix - annotations}\index{mixfix annotations|bold}. Associated with any - kind of name and type declaration, mixfixes give rise both to - grammar productions for the parser and output templates for the - pretty printer. + for concrete syntax is that of general \bfindex{mixfix annotations}. + Associated with any kind of name and type declaration, mixfixes give + rise both to grammar productions for the parser and output templates + for the pretty printer. In full generality, the whole affair of parser and pretty printer configuration is rather subtle. Any syntax specifications given by @@ -28,7 +27,7 @@ \end{isamarkuptext}% \isamarkuptrue% % -\isamarkupsubsection{Infixes% +\isamarkupsubsection{Infix annotations% } \isamarkuptrue% % @@ -40,11 +39,11 @@ well, although this is less frequently encountered in practice (\isa{{\isacharasterisk}} and \isa{{\isacharplus}} types may come to mind). - Infix declarations\index{infix annotations|bold} provide a useful - special case of mixfixes, where users need not care about the full - details of priorities, nesting, spacing, etc. The subsequent - example of the exclusive-or operation on boolean values illustrates - typical infix declarations.% + Infix declarations\index{infix annotations} provide a useful special + case of mixfixes, where users need not care about the full details + of priorities, nesting, spacing, etc. The subsequent example of the + exclusive-or operation on boolean values illustrates typical infix + declarations.% \end{isamarkuptext}% \isamarkuptrue% \isacommand{constdefs}\isanewline @@ -120,11 +119,10 @@ proposed over decades, but none has become universally available so far, not even Unicode\index{Unicode}. - Isabelle supports a generic notion of - \emph{symbols}\index{symbols|bold} as the smallest entities of - source text, without referring to internal encodings. Such - ``generalized characters'' may be of one of the following three - kinds: + Isabelle supports a generic notion of \bfindex{symbols} as the + smallest entities of source text, without referring to internal + encodings. Such ``generalized characters'' may be of one of the + following three kinds: \begin{enumerate} @@ -172,10 +170,10 @@ immediately after continuing further input. \medskip A slightly more refined scheme is to provide alternative - syntax via the \emph{print mode}\index{print mode} concept of - Isabelle (see also \cite{isabelle-ref}). By convention, the mode - ``$xsymbols$'' is enabled whenever X-Symbol is active. Consider the - following hybrid declaration of \isa{xor}.% + syntax via the \bfindex{print mode} concept of Isabelle (see also + \cite{isabelle-ref}). By convention, the mode ``$xsymbols$'' is + enabled whenever X-Symbol is active. Consider the following hybrid + declaration of \isa{xor}.% \end{isamarkuptext}% \isamarkuptrue% \isamarkupfalse% @@ -220,15 +218,15 @@ \isacommand{syntax}\ {\isacharparenleft}xsymbols\ \isakeyword{output}{\isacharparenright}\isanewline \ \ xor\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}bool\ {\isasymRightarrow}\ bool\ {\isasymRightarrow}\ bool{\isachardoublequote}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequote}{\isasymoplus}{\isasymignore}{\isachardoublequote}\ {\isadigit{6}}{\isadigit{0}}{\isacharparenright}\isamarkupfalse% % -\isamarkupsubsection{Prefixes% +\isamarkupsubsection{Prefix annotations% } \isamarkuptrue% % \begin{isamarkuptext}% -Prefix syntax annotations\index{prefix annotation|bold} are just a - very degenerate of the general mixfix form \cite{isabelle-ref}, - without any template arguments or priorities --- just some piece of - literal syntax. +Prefix syntax annotations\index{prefix annotation} are just a very + degenerate of the general mixfix form \cite{isabelle-ref}, without + any template arguments or priorities --- just some piece of literal + syntax. The following example illustrates this idea idea by associating common symbols with the constructors of a currency datatype.% @@ -244,31 +242,25 @@ Here the degenerate mixfix annotations on the rightmost column happen to consist of a single Isabelle symbol each: \verb,\,\verb,,, \verb,\,\verb,,, - \verb,\,\verb,,, and \verb,$,. + \verb,\,\verb,,, \verb,$,. Recall that a constructor like \isa{Euro} actually is a function \isa{nat\ {\isasymRightarrow}\ currency}. An expression like \isa{Euro\ {\isadigit{1}}{\isadigit{0}}} will - be printed as \isa{{\isasymeuro}\ {\isadigit{1}}{\isadigit{0}}}. Merely the head of the application is - subject to our trivial concrete syntax; this form is sufficient to - achieve fair conformance to EU~Commission standards of currency - notation. + be printed as \isa{{\isasymeuro}\ {\isadigit{1}}{\isadigit{0}}}. Only the head of the application is + subject to our concrete syntax; this simple form already achieves + conformance with notational standards of the European Commission. \medskip Certainly, the same idea of prefix syntax also works for \isakeyword{consts}, \isakeyword{constdefs} etc. For example, we might introduce a (slightly unrealistic) function to calculate an abstract currency value, by cases on the datatype constructors and - fixed exchange rates.% + fixed exchange rates. The funny symbol used here is that of + \verb,\,.% \end{isamarkuptext}% \isamarkuptrue% \isacommand{consts}\isanewline \ \ currency\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}currency\ {\isasymRightarrow}\ nat{\isachardoublequote}\ \ \ \ {\isacharparenleft}{\isachardoublequote}{\isasymcurrency}{\isachardoublequote}{\isacharparenright}\isamarkupfalse% % -\begin{isamarkuptext}% -\noindent The funny symbol encountered here is that of - \verb,\,.% -\end{isamarkuptext}% -\isamarkuptrue% -% \isamarkupsubsection{Syntax translations \label{sec:def-translations}% } \isamarkuptrue% @@ -312,24 +304,250 @@ } \isamarkuptrue% % -\isamarkupsubsection{Batch-mode sessions% -} +\begin{isamarkuptext}% +Isabelle/Isar is centered around a certain notion of \bfindex{formal + proof documents}\index{documents|bold}: ultimately the result of the + user's theory development efforts is a human-readable record --- as + a browsable PDF file or printed on paper. The overall document + structure follows traditional mathematical articles, with + sectioning, explanations, definitions, theorem statements, and + proofs. + + The Isar proof language \cite{Wenzel-PhD}, which is not covered in + this book, admits to write formal proof texts that are acceptable + both to the machine and human readers at the same time. Thus + marginal comments and explanations may be kept at a minimum. + Nevertheless, Isabelle document output is still useful without + actual Isar proof texts; formal specifications usually deserve their + own coverage in the text, while unstructured proof scripts may be + just ignore by readers (or intentionally suppressed from the text). + + \medskip The Isabelle document preparation system essentially acts + like a formal front-end for {\LaTeX}. After checking definitions + and proofs the theory sources are turned into typesetting + instructions, so the final document is known to observe quite strong + ``soundness'' properties. This enables users to write authentic + reports on formal theory developments with little additional effort, + the most tedious consistency checks are handled by the system.% +\end{isamarkuptext}% \isamarkuptrue% % -\isamarkupsubsection{{\LaTeX} macros% -} -\isamarkuptrue% -% -\isamarkupsubsubsection{Structure markup% -} -\isamarkuptrue% -% -\isamarkupsubsubsection{Symbols and characters% +\isamarkupsubsection{Isabelle sessions% } \isamarkuptrue% % \begin{isamarkuptext}% -FIXME% +In contrast to the highly interactive mode of the formal parts of + Isabelle/Isar theory development, the document preparation stage + essentially works in batch-mode. This revolves around the concept + of a \bfindex{session}, which essentially consists of a collection + of theory source files that contribute to a single output document. + Each session is derived from a parent one (usually an object-logic + image such as \texttt{HOL}); this results in an overall tree + structure of Isabelle sessions. + + The canonical arrangement of source files of a session called + \texttt{MySession} is as follows: + + \begin{itemize} + + \item Directory \texttt{MySession} contains the required theory + files, say $A@1$\texttt{.thy}, \dots, $A@n$\texttt{.thy}. + + \item File \texttt{MySession/ROOT.ML} holds appropriate ML commands + for loading all wanted theories, usually just + \texttt{use_thy~"$A@i$"} for any $A@i$ in leaf position of the + theory dependency graph. + + \item Directory \texttt{MySession/document} contains everything + required for the {\LaTeX} stage, but only \texttt{root.tex} needs to + be provided initially. The latter file holds appropriate {\LaTeX} + code to commence a document (\verb,\documentclass, etc.), and to + include the generated files $A@i$\texttt{.tex} for each theory. The + generated file \texttt{session.tex} holds {\LaTeX} commands to + include \emph{all} theory output files in topologically sorted + order. + + \item In addition an \texttt{IsaMakefile} outside of directory + \texttt{MySession} holds appropriate dependencies and invocations of + Isabelle tools to control the batch job. The details are covered in + \cite{isabelle-sys}; \texttt{isatool usedir} is the most important + entry point. + + \end{itemize} + + With everything put in its proper place, \texttt{isatool make} + should be sufficient to process the Isabelle session completely, + with the generated document appearing in its proper place (within + \verb,~/isabelle/browser_info,). + + In practive, users may want to have \texttt{isatool mkdir} generate + an initial working setup without further ado. For example, an empty + session \texttt{MySession} derived from \texttt{HOL} may be produced + as follows: + +\begin{verbatim} + isatool mkdir HOL MySession + isatool make +\end{verbatim} + + This runs the session with sensible default options, including + verbose mode to tell the user where the result will appear. The + above dry run should produce should produce a single page of output + (with a dummy title, empty table of contents etc.). Any failure at + that stage is likely to indicate some technical problems with your + {\LaTeX} installation.\footnote{Especially make sure that + \texttt{pdflatex} is present.} + + \medskip Users may now start to populate the directory + \texttt{MySession}, and the file \texttt{MySession/ROOT.ML} + accordingly. \texttt{MySession/document/root.tex} should be also + adapted at some point; the generated version is mostly + self-explanatory. + + Realistic applications may demand additional files in + \texttt{MySession/document} for the {\LaTeX} stage, such as + \texttt{root.bib} for the bibliographic database.\footnote{Using + that particular name of \texttt{root.bib}, the Isabelle document + processor (actually \texttt{isatool document} \cite{isabelle-sys}) + will be smart enough to invoke \texttt{bibtex} accordingly.} + + \medskip Failure of the document preparation phase in an Isabelle + batch session leaves the generated sources in there target location + (as pointed out by the accompanied error message). In case of + {\LaTeX} errors, users may trace error messages at the file position + of the generated text.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Structure markup% +} +\isamarkuptrue% +% +\isamarkupsubsubsection{Sections% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The large-scale structure of Isabelle documents closely follows + {\LaTeX} convention, with chapters, sections, subsubsections etc. + The formal Isar language includes separate structure \bfindex{markup + commands}, which do not effect the formal content of a theory (or + proof), but results in corresponding {\LaTeX} elements issued to the + output. + + There are different markup commands for different formal contexts: + in header position (just before a \isakeyword{theory} command), + within the theory body, or within a proof. Note that the header + needs to be treated specially, since ordinary theory and proof + commands may only occur \emph{after} the initial \isakeyword{theory} + specification. + + \smallskip + + \begin{tabular}{llll} + header & theory & proof & default meaning \\\hline + & \commdx{chapter} & & \verb,\chapter, \\ + \commdx{header} & \commdx{section} & \commdx{sect} & \verb,\section, \\ + & \commdx{subsection} & \commdx{subsect} & \verb,\subsection, \\ + & \commdx{subsubsection} & \commdx{subsubsect} & \verb,\subsubsection, \\ + \end{tabular} + + \medskip + + From the Isabelle perspective, each markup command takes a single + text argument (delimited by \verb,",\dots\verb,", or + \verb,{,\verb,*,~\dots~\verb,*,\verb,},). After stripping + surrounding white space, the argument is passed to a {\LaTeX} macro + \verb,\isamarkupXXX, for any command \isakeyword{XXX}. The latter + are defined in \verb,isabelle.sty, according to the rightmost column + above. + + \medskip The following source fragment illustrates structure markup + of a theory. + + \begin{ttbox} + header {\ttlbrace}* Some properties of Foo Bar elements *{\ttrbrace} + + theory Foo_Bar = Main: + + subsection {\ttlbrace}* Basic definitions *{\ttrbrace} + + consts + foo :: \dots + bar :: \dots + defs \dots + + subsection {\ttlbrace}* Derived rules *{\ttrbrace} + + lemma fooI: \dots + lemma fooE: \dots + + subsection {\ttlbrace}* Main theorem *{\ttrbrace} + + theorem main: \dots + + end + \end{ttbox} + + \medskip In realistic applications, users may well want to change + the meaning of some markup commands, typically via appropriate use + of \verb,\renewcommand, in \texttt{root.tex}. The + \verb,\isamarkupheader, is a good candidate for some adaption, e.g.\ + moving it up in the hierarchy to become \verb,\chapter,. + +\begin{verbatim} + \renewcommand{\isamarkupheader}[1]{\chapter{#1}} +\end{verbatim} + + Certainly, this requires to change the default + \verb,\documentclass{article}, in \texttt{root.tex} to something + that supports the notion of chapters in the first place, e.g.\ + \texttt{report} or \texttt{book}. + + \medskip For each generated theory output the {\LaTeX} macro + \verb,\isabellecontext, holds the name of the formal context. This + is particularly useful for document headings or footings, e.g.\ like + this: + +\begin{verbatim} + \renewcommand{\isamarkupheader}[1]% + {\chapter{#1}\markright{THEORY~\isabellecontext}} +\end{verbatim} + + \noindent Make sure to include something like + \verb,\pagestyle{headings}, in \texttt{root.tex} as well. Moreover, + the document should have more than 2 pages.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Formal comments and antiquotations% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Comments of the form \verb,(*,~\dots~\verb,*),% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Symbols and characters% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +FIXME \verb,\isabellestyle,% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Suppressing output% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +FIXME \verb,no_document, + + FIXME \verb,(,\verb,*,\verb,<,\verb,*,\verb,), + \verb,(,\verb,*,\verb,>,\verb,*,\verb,),% \end{isamarkuptext}% \isamarkuptrue% \isamarkupfalse% diff -r 39b93da27bc9 -r a107eeffd557 doc-src/TutorialI/tutorial.ind --- a/doc-src/TutorialI/tutorial.ind Sat Jan 05 21:41:11 2002 +0100 +++ b/doc-src/TutorialI/tutorial.ind Sat Jan 05 21:41:38 2002 +0100 @@ -1,661 +1,661 @@ \begin{theindex} - \item \ttall, \bold{203} - \item \texttt{?}, \bold{203} - \item \isasymuniqex, \bold{203} - \item \ttuniquex, \bold{203} - \item {\texttt {\&}}, \bold{203} - \item \verb$~$, \bold{203} - \item \verb$~=$, \bold{203} - \item \ttor, \bold{203} + \item \ttall, \bold{207} + \item \texttt{?}, \bold{207} + \item \isasymuniqex, \bold{207} + \item \ttuniquex, \bold{207} + \item {\texttt {\&}}, \bold{207} + \item \verb$~$, \bold{207} + \item \verb$~=$, \bold{207} + \item \ttor, \bold{207} \item \texttt{[]}, \bold{9} \item \texttt{\#}, \bold{9} - \item \texttt{\at}, \bold{10}, \hyperpage{203} - \item \isasymnotin, \bold{203} - \item \verb$~:$, \bold{203} - \item \isasymInter, \bold{203} - \item \isasymUnion, \bold{203} - \item \isasyminverse, \bold{203} - \item \verb$^-1$, \bold{203} - \item \isactrlsup{\isacharasterisk}, \bold{203} - \item \verb$^$\texttt{*}, \bold{203} - \item \isasymAnd, \bold{12}, \bold{203} - \item \ttAnd, \bold{203} - \item \isasymrightleftharpoons, \hyperpage{57} - \item \isasymrightharpoonup, \hyperpage{57} - \item \isasymleftharpoondown, \hyperpage{57} + \item \texttt{\at}, \bold{10}, 207 + \item \isasymnotin, \bold{207} + \item \verb$~:$, \bold{207} + \item \isasymInter, \bold{207} + \item \isasymUnion, \bold{207} + \item \isasyminverse, \bold{207} + \item \verb$^-1$, \bold{207} + \item \isactrlsup{\isacharasterisk}, \bold{207} + \item \verb$^$\texttt{*}, \bold{207} + \item \isasymAnd, \bold{12}, \bold{207} + \item \ttAnd, \bold{207} + \item \isasymrightleftharpoons, 57 + \item \isasymrightharpoonup, 57 + \item \isasymleftharpoondown, 57 \item \emph {$\Rightarrow $}, \bold{5} - \item \ttlbr, \bold{203} - \item \ttrbr, \bold{203} - \item \texttt {\%}, \bold{203} + \item \ttlbr, \bold{207} + \item \ttrbr, \bold{207} + \item \texttt {\%}, \bold{207} \item \texttt {;}, \bold{7} - \item \isa {()} (constant), \hyperpage{24} - \item \isa {+} (tactical), \hyperpage{93} + \item \isa {()} (constant), 24 + \item \isa {+} (tactical), 97 \item \isa {<*lex*>}, \see{lexicographic product}{1} - \item \isa {?} (tactical), \hyperpage{93} - \item \texttt{|} (tactical), \hyperpage{93} + \item \isa {?} (tactical), 97 + \item \texttt{|} (tactical), 97 \indexspace - \item \isa {0} (constant), \hyperpage{22, 23}, \hyperpage{144} - \item \isa {1} (constant), \hyperpage{23}, \hyperpage{144, 145} + \item \isa {0} (constant), 22, 23, 148 + \item \isa {1} (constant), 23, 148, 149 \indexspace \item abandoning a proof, \bold{13} \item abandoning a theory, \bold{16} - \item \isa {abs} (constant), \hyperpage{147} - \item \texttt {abs}, \bold{203} - \item absolute value, \hyperpage{147} - \item \isa {add} (modifier), \hyperpage{29} - \item \isa {add_ac} (theorems), \hyperpage{146} - \item \isa {add_assoc} (theorem), \bold{146} - \item \isa {add_commute} (theorem), \bold{146} - \item \isa {add_mult_distrib} (theorem), \bold{145} - \item \texttt {ALL}, \bold{203} - \item \isa {All} (constant), \hyperpage{103} - \item \isa {allE} (theorem), \bold{75} - \item \isa {allI} (theorem), \bold{74} - \item append function, \hyperpage{10--14} - \item \isacommand {apply} (command), \hyperpage{15} - \item \isa {arg_cong} (theorem), \bold{90} - \item \isa {arith} (method), \hyperpage{23}, \hyperpage{143} + \item \isa {abs} (constant), 151 + \item \texttt {abs}, \bold{207} + \item absolute value, 151 + \item \isa {add} (modifier), 29 + \item \isa {add_ac} (theorems), 150 + \item \isa {add_assoc} (theorem), \bold{150} + \item \isa {add_commute} (theorem), \bold{150} + \item \isa {add_mult_distrib} (theorem), \bold{149} + \item \texttt {ALL}, \bold{207} + \item \isa {All} (constant), 107 + \item \isa {allE} (theorem), \bold{79} + \item \isa {allI} (theorem), \bold{78} + \item append function, 10--14 + \item \isacommand {apply} (command), 15 + \item \isa {arg_cong} (theorem), \bold{94} + \item \isa {arith} (method), 23, 147 \item arithmetic operations - \subitem for \protect\isa{nat}, \hyperpage{23} - \item \textsc {ascii} symbols, \bold{203} - \item Aspinall, David, \hyperpage{viii} - \item associative-commutative function, \hyperpage{170} - \item \isa {assumption} (method), \hyperpage{63} + \subitem for \protect\isa{nat}, 23 + \item \textsc {ascii} symbols, \bold{207} + \item Aspinall, David, viii + \item associative-commutative function, 174 + \item \isa {assumption} (method), 67 \item assumptions - \subitem of subgoal, \hyperpage{12} - \subitem renaming, \hyperpage{76--77} - \subitem reusing, \hyperpage{77} - \item \isa {auto} (method), \hyperpage{38}, \hyperpage{86} - \item \isa {axclass}, \hyperpage{158--164} - \item axiom of choice, \hyperpage{80} - \item axiomatic type classes, \hyperpage{158--164} + \subitem of subgoal, 12 + \subitem renaming, 80--81 + \subitem reusing, 81 + \item \isa {auto} (method), 38, 90 + \item \isa {axclass}, 162--168 + \item axiom of choice, 84 + \item axiomatic type classes, 162--168 \indexspace - \item \isacommand {back} (command), \hyperpage{72} - \item \isa {Ball} (constant), \hyperpage{103} - \item \isa {ballI} (theorem), \bold{102} - \item \isa {best} (method), \hyperpage{86} - \item \isa {Bex} (constant), \hyperpage{103} - \item \isa {bexE} (theorem), \bold{102} - \item \isa {bexI} (theorem), \bold{102} - \item \isa {bij_def} (theorem), \bold{104} - \item bijections, \hyperpage{104} - \item binary trees, \hyperpage{18} - \item binomial coefficients, \hyperpage{103} - \item bisimulations, \hyperpage{110} - \item \isa {blast} (method), \hyperpage{83--84}, \hyperpage{86} - \item \isa {bool} (type), \hyperpage{4, 5} - \item boolean expressions example, \hyperpage{20--22} - \item \isa {bspec} (theorem), \bold{102} - \item \isacommand{by} (command), \hyperpage{67} + \item \isacommand {back} (command), 76 + \item \isa {Ball} (constant), 107 + \item \isa {ballI} (theorem), \bold{106} + \item \isa {best} (method), 90 + \item \isa {Bex} (constant), 107 + \item \isa {bexE} (theorem), \bold{106} + \item \isa {bexI} (theorem), \bold{106} + \item \isa {bij_def} (theorem), \bold{108} + \item bijections, 108 + \item binary trees, 18 + \item binomial coefficients, 107 + \item bisimulations, 114 + \item \isa {blast} (method), 87--88, 90 + \item \isa {bool} (type), 4, 5 + \item boolean expressions example, 20--22 + \item \isa {bspec} (theorem), \bold{106} + \item \isacommand{by} (command), 71 \indexspace - \item \isa {card} (constant), \hyperpage{103} - \item \isa {card_Pow} (theorem), \bold{103} - \item \isa {card_Un_Int} (theorem), \bold{103} - \item cardinality, \hyperpage{103} - \item \isa {case} (symbol), \hyperpage{32, 33} - \item \isa {case} expressions, \hyperpage{5, 6}, \hyperpage{18} - \item case distinctions, \hyperpage{19} + \item \isa {card} (constant), 107 + \item \isa {card_Pow} (theorem), \bold{107} + \item \isa {card_Un_Int} (theorem), \bold{107} + \item cardinality, 107 + \item \isa {case} (symbol), 32, 33 + \item \isa {case} expressions, 5, 6, 18 + \item case distinctions, 19 \item case splits, \bold{31} - \item \isa {case_tac} (method), \hyperpage{19}, \hyperpage{95}, - \hyperpage{151} - \item \isa {cases} (method), \hyperpage{156} - \item \isa {clarify} (method), \hyperpage{85, 86} - \item \isa {clarsimp} (method), \hyperpage{85, 86} - \item \isa {classical} (theorem), \bold{67} - \item coinduction, \bold{110} - \item \isa {Collect} (constant), \hyperpage{103} - \item compiling expressions example, \hyperpage{36--38} - \item \isa {Compl_iff} (theorem), \bold{100} + \item \isa {case_tac} (method), 19, 99, 155 + \item \isa {cases} (method), 160 + \item \isacommand {chapter} (command), 59 + \item \isa {clarify} (method), 89, 90 + \item \isa {clarsimp} (method), 89, 90 + \item \isa {classical} (theorem), \bold{71} + \item coinduction, \bold{114} + \item \isa {Collect} (constant), 107 + \item compiling expressions example, 36--38 + \item \isa {Compl_iff} (theorem), \bold{104} \item complement - \subitem of a set, \hyperpage{99} + \subitem of a set, 103 \item composition - \subitem of functions, \bold{104} - \subitem of relations, \bold{106} + \subitem of functions, \bold{108} + \subitem of relations, \bold{110} \item conclusion - \subitem of subgoal, \hyperpage{12} + \subitem of subgoal, 12 \item conditional expressions, \see{\isa{if} expressions}{1} - \item conditional simplification rules, \hyperpage{31} - \item \isa {cong} (attribute), \hyperpage{170} - \item congruence rules, \bold{169} - \item \isa {conjE} (theorem), \bold{65} - \item \isa {conjI} (theorem), \bold{62} - \item \isa {Cons} (constant), \hyperpage{9} - \item \isacommand {constdefs} (command), \hyperpage{25} - \item \isacommand {consts} (command), \hyperpage{10} - \item contrapositives, \hyperpage{67} + \item conditional simplification rules, 31 + \item \isa {cong} (attribute), 174 + \item congruence rules, \bold{173} + \item \isa {conjE} (theorem), \bold{69} + \item \isa {conjI} (theorem), \bold{66} + \item \isa {Cons} (constant), 9 + \item \isacommand {constdefs} (command), 25 + \item \isacommand {consts} (command), 10 + \item contrapositives, 71 \item converse - \subitem of a relation, \bold{106} - \item \isa {converse_iff} (theorem), \bold{106} - \item CTL, \hyperpage{115--120}, \hyperpage{185--187} + \subitem of a relation, \bold{110} + \item \isa {converse_iff} (theorem), \bold{110} + \item CTL, 119--124, 189--191 \indexspace - \item \isacommand {datatype} (command), \hyperpage{9}, - \hyperpage{38--43} - \item datatypes, \hyperpage{17--22} - \subitem and nested recursion, \hyperpage{40}, \hyperpage{44} - \subitem mutually recursive, \hyperpage{38} - \subitem nested, \hyperpage{174} - \item \isacommand {defer} (command), \hyperpage{16}, \hyperpage{94} - \item Definitional Approach, \hyperpage{26} + \item \isacommand {datatype} (command), 9, 38--43 + \item datatypes, 17--22 + \subitem and nested recursion, 40, 44 + \subitem mutually recursive, 38 + \subitem nested, 178 + \item \isacommand {defer} (command), 16, 98 + \item Definitional Approach, 26 \item definitions, \bold{25} \subitem unfolding, \bold{30} - \item \isacommand {defs} (command), \hyperpage{25} - \item \isa {del} (modifier), \hyperpage{29} - \item description operators, \hyperpage{79--81} + \item \isacommand {defs} (command), 25 + \item \isa {del} (modifier), 29 + \item description operators, 83--85 \item descriptions - \subitem definite, \hyperpage{79} - \subitem indefinite, \hyperpage{80} - \item \isa {dest} (attribute), \hyperpage{96} - \item destruction rules, \hyperpage{65} - \item \isa {diff_mult_distrib} (theorem), \bold{145} + \subitem definite, 83 + \subitem indefinite, 84 + \item \isa {dest} (attribute), 100 + \item destruction rules, 69 + \item \isa {diff_mult_distrib} (theorem), \bold{149} \item difference - \subitem of sets, \bold{100} - \item \isa {disjCI} (theorem), \bold{68} - \item \isa {disjE} (theorem), \bold{64} - \item \isa {div} (symbol), \hyperpage{23} - \item divides relation, \hyperpage{78}, \hyperpage{89}, - \hyperpage{95--98}, \hyperpage{146} + \subitem of sets, \bold{104} + \item \isa {disjCI} (theorem), \bold{72} + \item \isa {disjE} (theorem), \bold{68} + \item \isa {div} (symbol), 23 + \item divides relation, 82, 93, 99--102, 150 \item division - \subitem by negative numbers, \hyperpage{147} - \subitem by zero, \hyperpage{146} - \subitem for type \protect\isa{nat}, \hyperpage{145} + \subitem by negative numbers, 151 + \subitem by zero, 150 + \subitem for type \protect\isa{nat}, 149 + \item documents, \bold{57} \item domain - \subitem of a relation, \hyperpage{106} - \item \isa {Domain_iff} (theorem), \bold{106} - \item \isacommand {done} (command), \hyperpage{13} - \item \isa {drule_tac} (method), \hyperpage{70}, \hyperpage{90} - \item \isa {dvd_add} (theorem), \bold{146} - \item \isa {dvd_anti_sym} (theorem), \bold{146} - \item \isa {dvd_def} (theorem), \bold{146} + \subitem of a relation, 110 + \item \isa {Domain_iff} (theorem), \bold{110} + \item \isacommand {done} (command), 13 + \item \isa {drule_tac} (method), 74, 94 + \item \isa {dvd_add} (theorem), \bold{150} + \item \isa {dvd_anti_sym} (theorem), \bold{150} + \item \isa {dvd_def} (theorem), \bold{150} \indexspace - \item \isa {elim!} (attribute), \hyperpage{125} - \item elimination rules, \hyperpage{63--64} - \item \isacommand {end} (command), \hyperpage{14} - \item \isa {Eps} (constant), \hyperpage{103} - \item equality, \hyperpage{5} - \subitem of functions, \bold{103} - \subitem of records, \hyperpage{155} - \subitem of sets, \bold{100} - \item \isa {equalityE} (theorem), \bold{100} - \item \isa {equalityI} (theorem), \bold{100} - \item \isa {erule} (method), \hyperpage{64} - \item \isa {erule_tac} (method), \hyperpage{70} - \item Euclid's algorithm, \hyperpage{95--98} + \item \isa {elim!} (attribute), 129 + \item elimination rules, 67--68 + \item \isacommand {end} (command), 14 + \item \isa {Eps} (constant), 107 + \item equality, 5 + \subitem of functions, \bold{107} + \subitem of records, 159 + \subitem of sets, \bold{104} + \item \isa {equalityE} (theorem), \bold{104} + \item \isa {equalityI} (theorem), \bold{104} + \item \isa {erule} (method), 68 + \item \isa {erule_tac} (method), 74 + \item Euclid's algorithm, 99--102 \item even numbers - \subitem defining inductively, \hyperpage{121--125} - \item \texttt {EX}, \bold{203} - \item \isa {Ex} (constant), \hyperpage{103} - \item \isa {exE} (theorem), \bold{76} - \item \isa {exI} (theorem), \bold{76} - \item \isa {ext} (theorem), \bold{103} - \item \isa {extend} (constant), \hyperpage{157} + \subitem defining inductively, 125--129 + \item \texttt {EX}, \bold{207} + \item \isa {Ex} (constant), 107 + \item \isa {exE} (theorem), \bold{80} + \item \isa {exI} (theorem), \bold{80} + \item \isa {ext} (theorem), \bold{107} + \item \isa {extend} (constant), 161 \item extensionality - \subitem for functions, \bold{103, 104} - \subitem for records, \hyperpage{155} - \subitem for sets, \bold{100} - \item \ttEXU, \bold{203} + \subitem for functions, \bold{107, 108} + \subitem for records, 159 + \subitem for sets, \bold{104} + \item \ttEXU, \bold{207} \indexspace - \item \isa {False} (constant), \hyperpage{5} - \item \isa {fast} (method), \hyperpage{86}, \hyperpage{118} - \item Fibonacci function, \hyperpage{47} - \item \isa {fields} (constant), \hyperpage{157} - \item \isa {finite} (symbol), \hyperpage{103} - \item \isa {Finites} (constant), \hyperpage{103} - \item fixed points, \hyperpage{110} - \item flags, \hyperpage{5, 6}, \hyperpage{33} - \subitem setting and resetting, \hyperpage{5} - \item \isa {force} (method), \hyperpage{85, 86} - \item formulae, \hyperpage{5--6} - \item forward proof, \hyperpage{86--92} - \item \isa {frule} (method), \hyperpage{77} - \item \isa {frule_tac} (method), \hyperpage{70} - \item \isa {fst} (constant), \hyperpage{24} - \item function types, \hyperpage{5} - \item functions, \hyperpage{103--105} - \subitem partial, \hyperpage{176} - \subitem total, \hyperpage{11}, \hyperpage{46--52} - \subitem underdefined, \hyperpage{177} + \item \isa {False} (constant), 5 + \item \isa {fast} (method), 90, 122 + \item Fibonacci function, 47 + \item \isa {fields} (constant), 161 + \item \isa {finite} (symbol), 107 + \item \isa {Finites} (constant), 107 + \item fixed points, 114 + \item flags, 5, 6, 33 + \subitem setting and resetting, 5 + \item \isa {force} (method), 89, 90 + \item formal proof documents, \bold{57} + \item formulae, 5--6 + \item forward proof, 90--96 + \item \isa {frule} (method), 81 + \item \isa {frule_tac} (method), 74 + \item \isa {fst} (constant), 24 + \item function types, 5 + \item functions, 107--109 + \subitem partial, 180 + \subitem total, 11, 46--52 + \subitem underdefined, 181 \indexspace - \item \isa {gcd} (constant), \hyperpage{87--88}, \hyperpage{95--98} - \item generalizing for induction, \hyperpage{123} - \item generalizing induction formulae, \hyperpage{35} - \item Girard, Jean-Yves, \fnote{65} - \item Gordon, Mike, \hyperpage{3} + \item \isa {gcd} (constant), 91--92, 99--102 + \item generalizing for induction, 127 + \item generalizing induction formulae, 35 + \item Girard, Jean-Yves, \fnote{69} + \item Gordon, Mike, 3 \item grammars - \subitem defining inductively, \hyperpage{134--139} - \item ground terms example, \hyperpage{129--134} + \subitem defining inductively, 138--143 + \item ground terms example, 133--138 \indexspace - \item \isa {hd} (constant), \hyperpage{17}, \hyperpage{37} - \item Hilbert's $\varepsilon$-operator, \hyperpage{80} - \item HOLCF, \hyperpage{43} - \item Hopcroft, J. E., \hyperpage{139} - \item \isa {hypreal} (type), \hyperpage{149} + \item \isa {hd} (constant), 17, 37 + \item \isacommand {header} (command), 59 + \item Hilbert's $\varepsilon$-operator, 84 + \item HOLCF, 43 + \item Hopcroft, J. E., 143 + \item \isa {hypreal} (type), 153 \indexspace - \item \isa {Id_def} (theorem), \bold{106} - \item \isa {id_def} (theorem), \bold{104} + \item \isa {Id_def} (theorem), \bold{110} + \item \isa {id_def} (theorem), \bold{108} \item identifiers, \bold{6} \subitem qualified, \bold{4} - \item identity function, \bold{104} - \item identity relation, \bold{106} - \item \isa {if} expressions, \hyperpage{5, 6} - \subitem simplification of, \hyperpage{33} - \subitem splitting of, \hyperpage{31}, \hyperpage{49} - \item if-and-only-if, \hyperpage{6} - \item \isa {iff} (attribute), \hyperpage{84}, \hyperpage{96}, - \hyperpage{124} - \item \isa {iffD1} (theorem), \bold{88} - \item \isa {iffD2} (theorem), \bold{88} + \item identity function, \bold{108} + \item identity relation, \bold{110} + \item \isa {if} expressions, 5, 6 + \subitem simplification of, 33 + \subitem splitting of, 31, 49 + \item if-and-only-if, 6 + \item \isa {iff} (attribute), 88, 100, 128 + \item \isa {iffD1} (theorem), \bold{92} + \item \isa {iffD2} (theorem), \bold{92} \item image - \subitem under a function, \bold{105} - \subitem under a relation, \bold{106} - \item \isa {image_def} (theorem), \bold{105} - \item \isa {Image_iff} (theorem), \bold{106} - \item \isa {impI} (theorem), \bold{66} - \item implication, \hyperpage{66--67} - \item \isa {ind_cases} (method), \hyperpage{125} - \item \isa {induct_tac} (method), \hyperpage{12}, \hyperpage{19}, - \hyperpage{52}, \hyperpage{184} - \item induction, \hyperpage{180--187} - \subitem complete, \hyperpage{182} - \subitem deriving new schemas, \hyperpage{184} - \subitem on a term, \hyperpage{181} - \subitem recursion, \hyperpage{51--52} - \subitem structural, \hyperpage{19} - \subitem well-founded, \hyperpage{109} - \item induction heuristics, \hyperpage{34--36} - \item \isacommand {inductive} (command), \hyperpage{121} + \subitem under a function, \bold{109} + \subitem under a relation, \bold{110} + \item \isa {image_def} (theorem), \bold{109} + \item \isa {Image_iff} (theorem), \bold{110} + \item \isa {impI} (theorem), \bold{70} + \item implication, 70--71 + \item \isa {ind_cases} (method), 129 + \item \isa {induct_tac} (method), 12, 19, 52, 188 + \item induction, 184--191 + \subitem complete, 186 + \subitem deriving new schemas, 188 + \subitem on a term, 185 + \subitem recursion, 51--52 + \subitem structural, 19 + \subitem well-founded, 113 + \item induction heuristics, 34--36 + \item \isacommand {inductive} (command), 125 \item inductive definition - \subitem simultaneous, \hyperpage{135} - \item inductive definitions, \hyperpage{121--139} - \item \isacommand {inductive\_cases} (command), \hyperpage{125}, - \hyperpage{133} - \item infinitely branching trees, \hyperpage{43} - \item infix annotations, \bold{54} - \item \isacommand{infixr} (annotation), \hyperpage{10} - \item \isa {inj_on_def} (theorem), \bold{104} - \item injections, \hyperpage{104} - \item \isa {insert} (constant), \hyperpage{101} - \item \isa {insert} (method), \hyperpage{91--92} - \item instance, \bold{160} - \item \texttt {INT}, \bold{203} - \item \texttt {Int}, \bold{203} - \item \isa {int} (type), \hyperpage{147--148} - \item \isa {INT_iff} (theorem), \bold{102} - \item \isa {IntD1} (theorem), \bold{99} - \item \isa {IntD2} (theorem), \bold{99} - \item integers, \hyperpage{147--148} - \item \isa {INTER} (constant), \hyperpage{103} - \item \texttt {Inter}, \bold{203} - \item \isa {Inter_iff} (theorem), \bold{102} - \item intersection, \hyperpage{99} - \subitem indexed, \hyperpage{102} - \item \isa {IntI} (theorem), \bold{99} - \item \isa {intro} (method), \hyperpage{68} - \item \isa {intro!} (attribute), \hyperpage{122} - \item \isa {intro_classes} (method), \hyperpage{160} - \item introduction rules, \hyperpage{62--63} - \item \isa {inv} (constant), \hyperpage{80} - \item \isa {inv_image_def} (theorem), \bold{109} + \subitem simultaneous, 139 + \item inductive definitions, 125--143 + \item \isacommand {inductive\_cases} (command), 129, 137 + \item infinitely branching trees, 43 + \item infix annotations, 54 + \item \isacommand{infixr} (annotation), 10 + \item \isa {inj_on_def} (theorem), \bold{108} + \item injections, 108 + \item \isa {insert} (constant), 105 + \item \isa {insert} (method), 95--96 + \item instance, \bold{164} + \item \texttt {INT}, \bold{207} + \item \texttt {Int}, \bold{207} + \item \isa {int} (type), 151--152 + \item \isa {INT_iff} (theorem), \bold{106} + \item \isa {IntD1} (theorem), \bold{103} + \item \isa {IntD2} (theorem), \bold{103} + \item integers, 151--152 + \item \isa {INTER} (constant), 107 + \item \texttt {Inter}, \bold{207} + \item \isa {Inter_iff} (theorem), \bold{106} + \item intersection, 103 + \subitem indexed, 106 + \item \isa {IntI} (theorem), \bold{103} + \item \isa {intro} (method), 72 + \item \isa {intro!} (attribute), 126 + \item \isa {intro_classes} (method), 164 + \item introduction rules, 66--67 + \item \isa {inv} (constant), 84 + \item \isa {inv_image_def} (theorem), \bold{113} \item inverse - \subitem of a function, \bold{104} - \subitem of a relation, \bold{106} + \subitem of a function, \bold{108} + \subitem of a relation, \bold{110} \item inverse image - \subitem of a function, \hyperpage{105} - \subitem of a relation, \hyperpage{108} - \item \isa {itrev} (constant), \hyperpage{34} + \subitem of a function, 109 + \subitem of a relation, 112 + \item \isa {itrev} (constant), 34 \indexspace - \item \isacommand {kill} (command), \hyperpage{16} + \item \isacommand {kill} (command), 16 \indexspace - \item $\lambda$ expressions, \hyperpage{5} - \item LCF, \hyperpage{43} - \item \isa {LEAST} (symbol), \hyperpage{23}, \hyperpage{79} - \item least number operator, \see{\protect\isa{LEAST}}{79} - \item Leibniz, Gottfried Wilhelm, \hyperpage{53} - \item \isacommand {lemma} (command), \hyperpage{13} - \item \isacommand {lemmas} (command), \hyperpage{87}, \hyperpage{96} - \item \isa {length} (symbol), \hyperpage{18} - \item \isa {length_induct}, \bold{184} - \item \isa {less_than} (constant), \hyperpage{108} - \item \isa {less_than_iff} (theorem), \bold{108} - \item \isa {let} expressions, \hyperpage{5, 6}, \hyperpage{31} - \item \isa {Let_def} (theorem), \hyperpage{31} - \item \isa {lex_prod_def} (theorem), \bold{109} - \item lexicographic product, \bold{109}, \hyperpage{172} + \item $\lambda$ expressions, 5 + \item LCF, 43 + \item \isa {LEAST} (symbol), 23, 83 + \item least number operator, \see{\protect\isa{LEAST}}{83} + \item Leibniz, Gottfried Wilhelm, 53 + \item \isacommand {lemma} (command), 13 + \item \isacommand {lemmas} (command), 91, 100 + \item \isa {length} (symbol), 18 + \item \isa {length_induct}, \bold{188} + \item \isa {less_than} (constant), 112 + \item \isa {less_than_iff} (theorem), \bold{112} + \item \isa {let} expressions, 5, 6, 31 + \item \isa {Let_def} (theorem), 31 + \item \isa {lex_prod_def} (theorem), \bold{113} + \item lexicographic product, \bold{113}, 176 \item {\texttt{lfp}} - \subitem applications of, \see{CTL}{110} - \item Library, \hyperpage{4} - \item linear arithmetic, \hyperpage{22--24}, \hyperpage{143} - \item \isa {List} (theory), \hyperpage{17} - \item \isa {list} (type), \hyperpage{5}, \hyperpage{9}, - \hyperpage{17} - \item \isa {list.split} (theorem), \hyperpage{32} - \item \isa {lists_mono} (theorem), \bold{131} - \item Lowe, Gavin, \hyperpage{190--191} + \subitem applications of, \see{CTL}{114} + \item Library, 4 + \item linear arithmetic, 22--24, 147 + \item \isa {List} (theory), 17 + \item \isa {list} (type), 5, 9, 17 + \item \isa {list.split} (theorem), 32 + \item \isa {lists_mono} (theorem), \bold{135} + \item Lowe, Gavin, 194--195 \indexspace - \item \isa {Main} (theory), \hyperpage{4} - \item major premise, \bold{69} - \item \isa {make} (constant), \hyperpage{157} - \item \isa {max} (constant), \hyperpage{23, 24} - \item measure functions, \hyperpage{47}, \hyperpage{108} - \item \isa {measure_def} (theorem), \bold{109} - \item meta-logic, \bold{74} + \item \isa {Main} (theory), 4 + \item major premise, \bold{73} + \item \isa {make} (constant), 161 + \item markup commands, \bold{59} + \item \isa {max} (constant), 23, 24 + \item measure functions, 47, 112 + \item \isa {measure_def} (theorem), \bold{113} + \item meta-logic, \bold{78} \item methods, \bold{16} - \item \isa {min} (constant), \hyperpage{23, 24} + \item \isa {min} (constant), 23, 24 \item mixfix annotations, \bold{53} - \item \isa {mod} (symbol), \hyperpage{23} - \item \isa {mod_div_equality} (theorem), \bold{145} - \item \isa {mod_mult_distrib} (theorem), \bold{145} - \item model checking example, \hyperpage{110--120} - \item \emph{modus ponens}, \hyperpage{61}, \hyperpage{66} - \item \isa {mono_def} (theorem), \bold{110} - \item monotone functions, \bold{110}, \hyperpage{133} - \subitem and inductive definitions, \hyperpage{131--132} - \item \isa {more} (constant), \hyperpage{152}, \hyperpage{154} - \item \isa {mp} (theorem), \bold{66} - \item \isa {mult_ac} (theorems), \hyperpage{146} - \item multiple inheritance, \bold{163} - \item multiset ordering, \bold{109} + \item \isa {mod} (symbol), 23 + \item \isa {mod_div_equality} (theorem), \bold{149} + \item \isa {mod_mult_distrib} (theorem), \bold{149} + \item model checking example, 114--124 + \item \emph{modus ponens}, 65, 70 + \item \isa {mono_def} (theorem), \bold{114} + \item monotone functions, \bold{114}, 137 + \subitem and inductive definitions, 135--136 + \item \isa {more} (constant), 156, 158 + \item \isa {mp} (theorem), \bold{70} + \item \isa {mult_ac} (theorems), 150 + \item multiple inheritance, \bold{167} + \item multiset ordering, \bold{113} \indexspace - \item \isa {nat} (type), \hyperpage{4}, \hyperpage{22}, - \hyperpage{145--147} - \item \isa {nat_less_induct} (theorem), \hyperpage{182} - \item natural deduction, \hyperpage{61--62} - \item natural numbers, \hyperpage{22}, \hyperpage{145--147} - \item Needham-Schroeder protocol, \hyperpage{189--191} - \item negation, \hyperpage{67--69} - \item \isa {Nil} (constant), \hyperpage{9} - \item \isa {no_asm} (modifier), \hyperpage{29} - \item \isa {no_asm_simp} (modifier), \hyperpage{30} - \item \isa {no_asm_use} (modifier), \hyperpage{30} - \item non-standard reals, \hyperpage{149} + \item \isa {nat} (type), 4, 22, 149--151 + \item \isa {nat_less_induct} (theorem), 186 + \item natural deduction, 65--66 + \item natural numbers, 22, 149--151 + \item Needham-Schroeder protocol, 193--195 + \item negation, 71--73 + \item \isa {Nil} (constant), 9 + \item \isa {no_asm} (modifier), 29 + \item \isa {no_asm_simp} (modifier), 30 + \item \isa {no_asm_use} (modifier), 30 + \item non-standard reals, 153 \item \isa {None} (constant), \bold{24} - \item \isa {notE} (theorem), \bold{67} - \item \isa {notI} (theorem), \bold{67} - \item numbers, \hyperpage{143--149} - \item numeric literals, \hyperpage{144} - \subitem for type \protect\isa{nat}, \hyperpage{145} - \subitem for type \protect\isa{real}, \hyperpage{149} + \item \isa {notE} (theorem), \bold{71} + \item \isa {notI} (theorem), \bold{71} + \item numbers, 147--153 + \item numeric literals, 148 + \subitem for type \protect\isa{nat}, 149 + \subitem for type \protect\isa{real}, 153 \indexspace - \item \isa {O} (symbol), \hyperpage{106} - \item \texttt {o}, \bold{203} - \item \isa {o_def} (theorem), \bold{104} - \item \isa {OF} (attribute), \hyperpage{89--90} - \item \isa {of} (attribute), \hyperpage{87}, \hyperpage{90} - \item \isa {only} (modifier), \hyperpage{29} - \item \isacommand {oops} (command), \hyperpage{13} + \item \isa {O} (symbol), 110 + \item \texttt {o}, \bold{207} + \item \isa {o_def} (theorem), \bold{108} + \item \isa {OF} (attribute), 93--94 + \item \isa {of} (attribute), 91, 94 + \item \isa {only} (modifier), 29 + \item \isacommand {oops} (command), 13 \item \isa {option} (type), \bold{24} - \item ordered rewriting, \bold{170} - \item overloading, \hyperpage{23}, \hyperpage{159--161} - \subitem and arithmetic, \hyperpage{144} + \item ordered rewriting, \bold{174} + \item overloading, 23, 163--165 + \subitem and arithmetic, 148 \indexspace - \item pairs and tuples, \hyperpage{24}, \hyperpage{149--152} + \item pairs and tuples, 24, 153--156 \item parent theories, \bold{4} \item pattern matching - \subitem and \isacommand{recdef}, \hyperpage{47} + \subitem and \isacommand{recdef}, 47 \item patterns - \subitem higher-order, \bold{171} - \item PDL, \hyperpage{112--114} - \item \isacommand {pr} (command), \hyperpage{16}, \hyperpage{94} - \item \isacommand {prefer} (command), \hyperpage{16}, \hyperpage{94} - \item prefix annotation, \bold{56} + \subitem higher-order, \bold{175} + \item PDL, 116--118 + \item \isacommand {pr} (command), 16, 98 + \item \isacommand {prefer} (command), 16, 98 + \item prefix annotation, 56 \item primitive recursion, \see{recursion, primitive}{1} - \item \isacommand {primrec} (command), \hyperpage{10}, \hyperpage{18}, - \hyperpage{38--43} - \item print mode, \hyperpage{55} - \item \isacommand {print\_syntax} (command), \hyperpage{54} + \item \isacommand {primrec} (command), 10, 18, 38--43 + \item print mode, \bold{55} + \item \isacommand {print\_syntax} (command), 54 \item product type, \see{pairs and tuples}{1} \item Proof General, \bold{7} - \item proof state, \hyperpage{12} + \item proof state, 12 \item proofs \subitem abandoning, \bold{13} - \subitem examples of failing, \hyperpage{81--83} + \subitem examples of failing, 85--87 \item protocols - \subitem security, \hyperpage{189--199} + \subitem security, 193--203 \indexspace - \item quantifiers, \hyperpage{6} - \subitem and inductive definitions, \hyperpage{129--131} - \subitem existential, \hyperpage{76} - \subitem for sets, \hyperpage{102} - \subitem instantiating, \hyperpage{78} - \subitem universal, \hyperpage{73--76} + \item quantifiers, 6 + \subitem and inductive definitions, 133--135 + \subitem existential, 80 + \subitem for sets, 106 + \subitem instantiating, 82 + \subitem universal, 77--80 \indexspace - \item \isa {r_into_rtrancl} (theorem), \bold{106} - \item \isa {r_into_trancl} (theorem), \bold{107} + \item \isa {r_into_rtrancl} (theorem), \bold{110} + \item \isa {r_into_trancl} (theorem), \bold{111} \item range - \subitem of a function, \hyperpage{105} - \subitem of a relation, \hyperpage{106} - \item \isa {range} (symbol), \hyperpage{105} - \item \isa {Range_iff} (theorem), \bold{106} - \item \isa {Real} (theory), \hyperpage{149} - \item \isa {real} (type), \hyperpage{148--149} - \item real numbers, \hyperpage{148--149} - \item \isacommand {recdef} (command), \hyperpage{46--52}, - \hyperpage{108}, \hyperpage{172--180} - \subitem and numeric literals, \hyperpage{144} - \item \isa {recdef_cong} (attribute), \hyperpage{176} - \item \isa {recdef_simp} (attribute), \hyperpage{49} - \item \isa {recdef_wf} (attribute), \hyperpage{174} - \item \isacommand {record} (command), \hyperpage{153} - \item records, \hyperpage{152--158} - \subitem extensible, \hyperpage{154--155} + \subitem of a function, 109 + \subitem of a relation, 110 + \item \isa {range} (symbol), 109 + \item \isa {Range_iff} (theorem), \bold{110} + \item \isa {Real} (theory), 153 + \item \isa {real} (type), 152--153 + \item real numbers, 152--153 + \item \isacommand {recdef} (command), 46--52, 112, 176--184 + \subitem and numeric literals, 148 + \item \isa {recdef_cong} (attribute), 180 + \item \isa {recdef_simp} (attribute), 49 + \item \isa {recdef_wf} (attribute), 178 + \item \isacommand {record} (command), 157 + \item records, 156--162 + \subitem extensible, 158--159 \item recursion - \subitem guarded, \hyperpage{177} - \subitem primitive, \hyperpage{18} - \subitem well-founded, \bold{173} - \item recursion induction, \hyperpage{51--52} - \item \isacommand {redo} (command), \hyperpage{16} - \item reflexive and transitive closure, \hyperpage{106--108} + \subitem guarded, 181 + \subitem primitive, 18 + \subitem well-founded, \bold{177} + \item recursion induction, 51--52 + \item \isacommand {redo} (command), 16 + \item reflexive and transitive closure, 110--112 \item reflexive transitive closure - \subitem defining inductively, \hyperpage{126--129} - \item \isa {rel_comp_def} (theorem), \bold{106} - \item relations, \hyperpage{105--108} - \subitem well-founded, \hyperpage{108--109} - \item \isa {rename_tac} (method), \hyperpage{76--77} - \item \isa {rev} (constant), \hyperpage{10--14}, \hyperpage{34} + \subitem defining inductively, 130--133 + \item \isa {rel_comp_def} (theorem), \bold{110} + \item relations, 109--112 + \subitem well-founded, 112--113 + \item \isa {rename_tac} (method), 80--81 + \item \isa {rev} (constant), 10--14, 34 \item rewrite rules, \bold{27} - \subitem permutative, \bold{170} + \subitem permutative, \bold{174} \item rewriting, \bold{27} - \item \isa {rotate_tac} (method), \hyperpage{30} - \item \isa {rtrancl_refl} (theorem), \bold{106} - \item \isa {rtrancl_trans} (theorem), \bold{106} - \item rule induction, \hyperpage{122--124} - \item rule inversion, \hyperpage{124--125}, \hyperpage{133--134} - \item \isa {rule_format} (attribute), \hyperpage{181} - \item \isa {rule_tac} (method), \hyperpage{70} - \subitem and renaming, \hyperpage{77} + \item \isa {rotate_tac} (method), 30 + \item \isa {rtrancl_refl} (theorem), \bold{110} + \item \isa {rtrancl_trans} (theorem), \bold{110} + \item rule induction, 126--128 + \item rule inversion, 128--129, 137--138 + \item \isa {rule_format} (attribute), 185 + \item \isa {rule_tac} (method), 74 + \subitem and renaming, 81 \indexspace - \item \isa {safe} (method), \hyperpage{85, 86} - \item safe rules, \bold{84} + \item \isa {safe} (method), 89, 90 + \item safe rules, \bold{88} + \item \isacommand {sect} (command), 59 + \item \isacommand {section} (command), 59 \item selector - \subitem record, \hyperpage{153} - \item \isa {set} (type), \hyperpage{5}, \hyperpage{99} - \item set comprehensions, \hyperpage{101--102} - \item \isa {set_ext} (theorem), \bold{100} - \item sets, \hyperpage{99--103} - \subitem finite, \hyperpage{103} - \subitem notation for finite, \bold{101} + \subitem record, 157 + \item session, \bold{58} + \item \isa {set} (type), 5, 103 + \item set comprehensions, 105--106 + \item \isa {set_ext} (theorem), \bold{104} + \item sets, 103--107 + \subitem finite, 107 + \subitem notation for finite, \bold{105} \item settings, \see{flags}{1} - \item \isa {show_brackets} (flag), \hyperpage{6} - \item \isa {show_types} (flag), \hyperpage{5}, \hyperpage{16} - \item \isa {simp} (attribute), \hyperpage{11}, \hyperpage{28} + \item \isa {show_brackets} (flag), 6 + \item \isa {show_types} (flag), 5, 16 + \item \isa {simp} (attribute), 11, 28 \item \isa {simp} (method), \bold{28} - \item \isa {simp} del (attribute), \hyperpage{28} - \item \isa {simp_all} (method), \hyperpage{29}, \hyperpage{38} - \item simplification, \hyperpage{27--33}, \hyperpage{169--172} - \subitem of \isa{let}-expressions, \hyperpage{31} - \subitem with definitions, \hyperpage{30} - \subitem with/of assumptions, \hyperpage{29} - \item simplification rule, \hyperpage{171--172} - \item simplification rules, \hyperpage{28} - \subitem adding and deleting, \hyperpage{29} - \item \isa {simplified} (attribute), \hyperpage{87}, \hyperpage{90} - \item \isa {size} (constant), \hyperpage{17} - \item \isa {snd} (constant), \hyperpage{24} - \item \isa {SOME} (symbol), \hyperpage{80} - \item \texttt {SOME}, \bold{203} + \item \isa {simp} del (attribute), 28 + \item \isa {simp_all} (method), 29, 38 + \item simplification, 27--33, 173--176 + \subitem of \isa{let}-expressions, 31 + \subitem with definitions, 30 + \subitem with/of assumptions, 29 + \item simplification rule, 175--176 + \item simplification rules, 28 + \subitem adding and deleting, 29 + \item \isa {simplified} (attribute), 91, 94 + \item \isa {size} (constant), 17 + \item \isa {snd} (constant), 24 + \item \isa {SOME} (symbol), 84 + \item \texttt {SOME}, \bold{207} \item \isa {Some} (constant), \bold{24} - \item \isa {some_equality} (theorem), \bold{80} - \item \isa {someI} (theorem), \bold{80} - \item \isa {someI2} (theorem), \bold{80} - \item \isa {someI_ex} (theorem), \bold{81} - \item sorts, \hyperpage{164} - \item \isa {spec} (theorem), \bold{74} - \item \isa {split} (attribute), \hyperpage{32} - \item \isa {split} (constant), \hyperpage{150} - \item \isa {split} (method), \hyperpage{31}, \hyperpage{150} - \item \isa {split} (modifier), \hyperpage{32} + \item \isa {some_equality} (theorem), \bold{84} + \item \isa {someI} (theorem), \bold{84} + \item \isa {someI2} (theorem), \bold{84} + \item \isa {someI_ex} (theorem), \bold{85} + \item sorts, 168 + \item \isa {spec} (theorem), \bold{78} + \item \isa {split} (attribute), 32 + \item \isa {split} (constant), 154 + \item \isa {split} (method), 31, 154 + \item \isa {split} (modifier), 32 \item split rule, \bold{32} - \item \isa {split_if} (theorem), \hyperpage{32} - \item \isa {split_if_asm} (theorem), \hyperpage{32} - \item \isa {ssubst} (theorem), \bold{71} + \item \isa {split_if} (theorem), 32 + \item \isa {split_if_asm} (theorem), 32 + \item \isa {ssubst} (theorem), \bold{75} \item structural induction, \see{induction, structural}{1} - \item subclasses, \hyperpage{158}, \hyperpage{163} - \item subgoal numbering, \hyperpage{46} - \item \isa {subgoal_tac} (method), \hyperpage{92} - \item subgoals, \hyperpage{12} - \item subset relation, \bold{100} - \item \isa {subsetD} (theorem), \bold{100} - \item \isa {subsetI} (theorem), \bold{100} - \item \isa {subst} (method), \hyperpage{71} - \item substitution, \hyperpage{71--73} - \item \isa {Suc} (constant), \hyperpage{22} - \item \isa {surj_def} (theorem), \bold{104} - \item surjections, \hyperpage{104} - \item \isa {sym} (theorem), \bold{88} + \item subclasses, 162, 167 + \item subgoal numbering, 46 + \item \isa {subgoal_tac} (method), 96 + \item subgoals, 12 + \item \isacommand {subsect} (command), 59 + \item \isacommand {subsection} (command), 59 + \item subset relation, \bold{104} + \item \isa {subsetD} (theorem), \bold{104} + \item \isa {subsetI} (theorem), \bold{104} + \item \isa {subst} (method), 75 + \item substitution, 75--77 + \item \isacommand {subsubsect} (command), 59 + \item \isacommand {subsubsection} (command), 59 + \item \isa {Suc} (constant), 22 + \item \isa {surj_def} (theorem), \bold{108} + \item surjections, 108 + \item \isa {sym} (theorem), \bold{92} \item symbols, \bold{55} - \item syntax, \hyperpage{6}, \hyperpage{11} - \item \isacommand {syntax} (command), \hyperpage{56} - \item syntax translations, \hyperpage{57} + \item syntax, 6, 11 + \item \isacommand {syntax} (command), 56 + \item syntax translations, 57 \indexspace - \item tacticals, \hyperpage{93} - \item tactics, \hyperpage{12} - \item \isacommand {term} (command), \hyperpage{16} + \item tacticals, 97 + \item tactics, 12 + \item \isacommand {term} (command), 16 \item term rewriting, \bold{27} \item termination, \see{functions, total}{1} - \item terms, \hyperpage{5} - \item \isa {THE} (symbol), \hyperpage{79} - \item \isa {the_equality} (theorem), \bold{79} - \item \isa {THEN} (attribute), \bold{88}, \hyperpage{90}, - \hyperpage{96} - \item \isacommand {theorem} (command), \bold{11}, \hyperpage{13} - \item theories, \hyperpage{4} + \item terms, 5 + \item \isa {THE} (symbol), 83 + \item \isa {the_equality} (theorem), \bold{83} + \item \isa {THEN} (attribute), \bold{92}, 94, 100 + \item \isacommand {theorem} (command), \bold{11}, 13 + \item theories, 4 \subitem abandoning, \bold{16} - \item \isacommand {theory} (command), \hyperpage{16} - \item theory files, \hyperpage{4} - \item \isacommand {thm} (command), \hyperpage{16} - \item \isa {tl} (constant), \hyperpage{17} - \item \isa {ToyList} example, \hyperpage{9--14} - \item \isa {trace_simp} (flag), \hyperpage{33} + \item \isacommand {theory} (command), 16 + \item theory files, 4 + \item \isacommand {thm} (command), 16 + \item \isa {tl} (constant), 17 + \item \isa {ToyList} example, 9--14 + \item \isa {trace_simp} (flag), 33 \item tracing the simplifier, \bold{33} - \item \isa {trancl_trans} (theorem), \bold{107} - \item transition systems, \hyperpage{111} - \item \isacommand {translations} (command), \hyperpage{57} - \item tries, \hyperpage{44--46} - \item \isa {True} (constant), \hyperpage{5} - \item \isa {truncate} (constant), \hyperpage{157} + \item \isa {trancl_trans} (theorem), \bold{111} + \item transition systems, 115 + \item \isacommand {translations} (command), 57 + \item tries, 44--46 + \item \isa {True} (constant), 5 + \item \isa {truncate} (constant), 161 \item tuples, \see{pairs and tuples}{1} - \item \isacommand {typ} (command), \hyperpage{16} + \item \isacommand {typ} (command), 16 \item type constraints, \bold{6} - \item type constructors, \hyperpage{5} + \item type constructors, 5 \item type inference, \bold{5} - \item type synonyms, \hyperpage{25} - \item type variables, \hyperpage{5} - \item \isacommand {typedecl} (command), \hyperpage{111}, - \hyperpage{165} - \item \isacommand {typedef} (command), \hyperpage{165--168} - \item types, \hyperpage{4--5} - \subitem declaring, \hyperpage{165} - \subitem defining, \hyperpage{165--168} - \item \isacommand {types} (command), \hyperpage{25} + \item type synonyms, 25 + \item type variables, 5 + \item \isacommand {typedecl} (command), 115, 169 + \item \isacommand {typedef} (command), 169--172 + \item types, 4--5 + \subitem declaring, 169 + \subitem defining, 169--172 + \item \isacommand {types} (command), 25 \indexspace - \item Ullman, J. D., \hyperpage{139} - \item \texttt {UN}, \bold{203} - \item \texttt {Un}, \bold{203} - \item \isa {UN_E} (theorem), \bold{102} - \item \isa {UN_I} (theorem), \bold{102} - \item \isa {UN_iff} (theorem), \bold{102} - \item \isa {Un_subset_iff} (theorem), \bold{100} - \item \isacommand {undo} (command), \hyperpage{16} + \item Ullman, J. D., 143 + \item \texttt {UN}, \bold{207} + \item \texttt {Un}, \bold{207} + \item \isa {UN_E} (theorem), \bold{106} + \item \isa {UN_I} (theorem), \bold{106} + \item \isa {UN_iff} (theorem), \bold{106} + \item \isa {Un_subset_iff} (theorem), \bold{104} + \item \isacommand {undo} (command), 16 \item \isa {unfold} (method), \bold{31} - \item Unicode, \hyperpage{55} - \item unification, \hyperpage{70--73} - \item \isa {UNION} (constant), \hyperpage{103} - \item \texttt {Union}, \bold{203} + \item Unicode, 55 + \item unification, 74--77 + \item \isa {UNION} (constant), 107 + \item \texttt {Union}, \bold{207} \item union - \subitem indexed, \hyperpage{102} - \item \isa {Union_iff} (theorem), \bold{102} - \item \isa {unit} (type), \hyperpage{24} - \item unknowns, \hyperpage{7}, \bold{62} - \item unsafe rules, \bold{84} + \subitem indexed, 106 + \item \isa {Union_iff} (theorem), \bold{106} + \item \isa {unit} (type), 24 + \item unknowns, 7, \bold{66} + \item unsafe rules, \bold{88} \item update - \subitem record, \hyperpage{153} - \item updating a function, \bold{103} + \subitem record, 157 + \item updating a function, \bold{107} \indexspace - \item variables, \hyperpage{7} - \subitem schematic, \hyperpage{7} - \subitem type, \hyperpage{5} - \item \isa {vimage_def} (theorem), \bold{105} + \item variables, 7 + \subitem schematic, 7 + \subitem type, 5 + \item \isa {vimage_def} (theorem), \bold{109} \indexspace - \item Wenzel, Markus, \hyperpage{vii} - \item \isa {wf_induct} (theorem), \bold{109} - \item \isa {wf_inv_image} (theorem), \bold{109} - \item \isa {wf_less_than} (theorem), \bold{108} - \item \isa {wf_lex_prod} (theorem), \bold{109} - \item \isa {wf_measure} (theorem), \bold{109} - \item \isa {wf_subset} (theorem), \hyperpage{174} - \item \isa {while} (constant), \hyperpage{179} - \item \isa {While_Combinator} (theory), \hyperpage{179} - \item \isa {while_rule} (theorem), \hyperpage{179} + \item Wenzel, Markus, vii + \item \isa {wf_induct} (theorem), \bold{113} + \item \isa {wf_inv_image} (theorem), \bold{113} + \item \isa {wf_less_than} (theorem), \bold{112} + \item \isa {wf_lex_prod} (theorem), \bold{113} + \item \isa {wf_measure} (theorem), \bold{113} + \item \isa {wf_subset} (theorem), 178 + \item \isa {while} (constant), 183 + \item \isa {While_Combinator} (theory), 183 + \item \isa {while_rule} (theorem), 183 \indexspace - \item \isa {zadd_ac} (theorems), \hyperpage{147} - \item \isa {zmult_ac} (theorems), \hyperpage{147} + \item \isa {zadd_ac} (theorems), 151 + \item \isa {zmult_ac} (theorems), 151 \end{theindex}