doc-src/HOL/HOL.tex
author wenzelm
Mon, 23 Aug 1999 16:58:00 +0200
changeset 7328 4265615b4206
parent 7325 01bb8abb5a91
child 7490 9a74b57740d1
permissions -rw-r--r--
record_simproc;
     1 %% $Id$
     2 \chapter{Higher-Order Logic}
     3 \index{higher-order logic|(}
     4 \index{HOL system@{\sc hol} system}
     5 
     6 The theory~\thydx{HOL} implements higher-order logic.  It is based on
     7 Gordon's~{\sc hol} system~\cite{mgordon-hol}, which itself is based on
     8 Church's original paper~\cite{church40}.  Andrews's
     9 book~\cite{andrews86} is a full description of the original
    10 Church-style higher-order logic.  Experience with the {\sc hol} system
    11 has demonstrated that higher-order logic is widely applicable in many
    12 areas of mathematics and computer science, not just hardware
    13 verification, {\sc hol}'s original \textit{raison d'\^etre\/}.  It is
    14 weaker than {\ZF} set theory but for most applications this does not
    15 matter.  If you prefer {\ML} to Lisp, you will probably prefer \HOL\ 
    16 to~{\ZF}.
    17 
    18 The syntax of \HOL\footnote{Earlier versions of Isabelle's \HOL\ used a
    19 different syntax.  Ancient releases of Isabelle included still another version
    20 of~\HOL, with explicit type inference rules~\cite{paulson-COLOG}.  This
    21 version no longer exists, but \thydx{ZF} supports a similar style of
    22 reasoning.} follows $\lambda$-calculus and functional programming.  Function
    23 application is curried.  To apply the function~$f$ of type
    24 $\tau@1\To\tau@2\To\tau@3$ to the arguments~$a$ and~$b$ in \HOL, you simply
    25 write $f\,a\,b$.  There is no `apply' operator as in \thydx{ZF}.  Note that
    26 $f(a,b)$ means ``$f$ applied to the pair $(a,b)$'' in \HOL.  We write ordered
    27 pairs as $(a,b)$, not $\langle a,b\rangle$ as in {\ZF}.
    28 
    29 \HOL\ has a distinct feel, compared with {\ZF} and {\CTT}.  It
    30 identifies object-level types with meta-level types, taking advantage of
    31 Isabelle's built-in type-checker.  It identifies object-level functions
    32 with meta-level functions, so it uses Isabelle's operations for abstraction
    33 and application.
    34 
    35 These identifications allow Isabelle to support \HOL\ particularly
    36 nicely, but they also mean that \HOL\ requires more sophistication
    37 from the user --- in particular, an understanding of Isabelle's type
    38 system.  Beginners should work with \texttt{show_types} (or even
    39 \texttt{show_sorts}) set to \texttt{true}.
    40 %  Gain experience by
    41 %working in first-order logic before attempting to use higher-order logic.
    42 %This chapter assumes familiarity with~{\FOL{}}.
    43 
    44 
    45 \begin{figure}
    46 \begin{constants}
    47   \it name      &\it meta-type  & \it description \\
    48   \cdx{Trueprop}& $bool\To prop$                & coercion to $prop$\\
    49   \cdx{Not}     & $bool\To bool$                & negation ($\neg$) \\
    50   \cdx{True}    & $bool$                        & tautology ($\top$) \\
    51   \cdx{False}   & $bool$                        & absurdity ($\bot$) \\
    52   \cdx{If}      & $[bool,\alpha,\alpha]\To\alpha$ & conditional \\
    53   \cdx{Let}     & $[\alpha,\alpha\To\beta]\To\beta$ & let binder
    54 \end{constants}
    55 \subcaption{Constants}
    56 
    57 \begin{constants}
    58 \index{"@@{\tt\at} symbol}
    59 \index{*"! symbol}\index{*"? symbol}
    60 \index{*"?"! symbol}\index{*"E"X"! symbol}
    61   \it symbol &\it name     &\it meta-type & \it description \\
    62   \sdx{SOME} or \tt\at & \cdx{Eps}  & $(\alpha\To bool)\To\alpha$ & 
    63         Hilbert description ($\varepsilon$) \\
    64   \sdx{ALL} or {\tt!~} & \cdx{All}  & $(\alpha\To bool)\To bool$ & 
    65         universal quantifier ($\forall$) \\
    66   \sdx{EX} or {\tt?~}  & \cdx{Ex}   & $(\alpha\To bool)\To bool$ & 
    67         existential quantifier ($\exists$) \\
    68   \texttt{EX!} or {\tt?!} & \cdx{Ex1}  & $(\alpha\To bool)\To bool$ & 
    69         unique existence ($\exists!$)\\
    70   \texttt{LEAST}  & \cdx{Least}  & $(\alpha::ord \To bool)\To\alpha$ & 
    71         least element
    72 \end{constants}
    73 \subcaption{Binders} 
    74 
    75 \begin{constants}
    76 \index{*"= symbol}
    77 \index{&@{\tt\&} symbol}
    78 \index{*"| symbol}
    79 \index{*"-"-"> symbol}
    80   \it symbol    & \it meta-type & \it priority & \it description \\ 
    81   \sdx{o}       & $[\beta\To\gamma,\alpha\To\beta]\To (\alpha\To\gamma)$ & 
    82         Left 55 & composition ($\circ$) \\
    83   \tt =         & $[\alpha,\alpha]\To bool$ & Left 50 & equality ($=$) \\
    84   \tt <         & $[\alpha::ord,\alpha]\To bool$ & Left 50 & less than ($<$) \\
    85   \tt <=        & $[\alpha::ord,\alpha]\To bool$ & Left 50 & 
    86                 less than or equals ($\leq$)\\
    87   \tt \&        & $[bool,bool]\To bool$ & Right 35 & conjunction ($\conj$) \\
    88   \tt |         & $[bool,bool]\To bool$ & Right 30 & disjunction ($\disj$) \\
    89   \tt -->       & $[bool,bool]\To bool$ & Right 25 & implication ($\imp$)
    90 \end{constants}
    91 \subcaption{Infixes}
    92 \caption{Syntax of \texttt{HOL}} \label{hol-constants}
    93 \end{figure}
    94 
    95 
    96 \begin{figure}
    97 \index{*let symbol}
    98 \index{*in symbol}
    99 \dquotes
   100 \[\begin{array}{rclcl}
   101     term & = & \hbox{expression of class~$term$} \\
   102          & | & "SOME~" id " . " formula
   103          & | & "\at~" id " . " formula \\
   104          & | & 
   105     \multicolumn{3}{l}{"let"~id~"="~term";"\dots";"~id~"="~term~"in"~term} \\
   106          & | & 
   107     \multicolumn{3}{l}{"if"~formula~"then"~term~"else"~term} \\
   108          & | & "LEAST"~ id " . " formula \\[2ex]
   109  formula & = & \hbox{expression of type~$bool$} \\
   110          & | & term " = " term \\
   111          & | & term " \ttilde= " term \\
   112          & | & term " < " term \\
   113          & | & term " <= " term \\
   114          & | & "\ttilde\ " formula \\
   115          & | & formula " \& " formula \\
   116          & | & formula " | " formula \\
   117          & | & formula " --> " formula \\
   118          & | & "ALL~" id~id^* " . " formula
   119          & | & "!~~~" id~id^* " . " formula \\
   120          & | & "EX~~" id~id^* " . " formula 
   121          & | & "?~~~" id~id^* " . " formula \\
   122          & | & "EX!~" id~id^* " . " formula
   123          & | & "?!~~" id~id^* " . " formula \\
   124   \end{array}
   125 \]
   126 \caption{Full grammar for \HOL} \label{hol-grammar}
   127 \end{figure} 
   128 
   129 
   130 \section{Syntax}
   131 
   132 Figure~\ref{hol-constants} lists the constants (including infixes and
   133 binders), while Fig.\ts\ref{hol-grammar} presents the grammar of
   134 higher-order logic.  Note that $a$\verb|~=|$b$ is translated to
   135 $\neg(a=b)$.
   136 
   137 \begin{warn}
   138   \HOL\ has no if-and-only-if connective; logical equivalence is expressed
   139   using equality.  But equality has a high priority, as befitting a
   140   relation, while if-and-only-if typically has the lowest priority.  Thus,
   141   $\neg\neg P=P$ abbreviates $\neg\neg (P=P)$ and not $(\neg\neg P)=P$.
   142   When using $=$ to mean logical equivalence, enclose both operands in
   143   parentheses.
   144 \end{warn}
   145 
   146 \subsection{Types and classes}
   147 The universal type class of higher-order terms is called~\cldx{term}.
   148 By default, explicit type variables have class \cldx{term}.  In
   149 particular the equality symbol and quantifiers are polymorphic over
   150 class \texttt{term}.
   151 
   152 The type of formulae, \tydx{bool}, belongs to class \cldx{term}; thus,
   153 formulae are terms.  The built-in type~\tydx{fun}, which constructs
   154 function types, is overloaded with arity {\tt(term,\thinspace
   155   term)\thinspace term}.  Thus, $\sigma\To\tau$ belongs to class~{\tt
   156   term} if $\sigma$ and~$\tau$ do, allowing quantification over
   157 functions.
   158 
   159 \HOL\ offers various methods for introducing new types.
   160 See~\S\ref{sec:HOL:Types} and~\S\ref{sec:HOL:datatype}.
   161 
   162 Theory \thydx{Ord} defines the syntactic class \cldx{ord} of order
   163 signatures; the relations $<$ and $\leq$ are polymorphic over this
   164 class, as are the functions \cdx{mono}, \cdx{min} and \cdx{max}, and
   165 the \cdx{LEAST} operator. \thydx{Ord} also defines a subclass
   166 \cldx{order} of \cldx{ord} which axiomatizes partially ordered types
   167 (w.r.t.\ $\le$).
   168 
   169 Three other syntactic type classes --- \cldx{plus}, \cldx{minus} and
   170 \cldx{times} --- permit overloading of the operators {\tt+},\index{*"+
   171   symbol} {\tt-}\index{*"- symbol} and {\tt*}.\index{*"* symbol} In
   172 particular, {\tt-} is instantiated for set difference and subtraction
   173 on natural numbers.
   174 
   175 If you state a goal containing overloaded functions, you may need to include
   176 type constraints.  Type inference may otherwise make the goal more
   177 polymorphic than you intended, with confusing results.  For example, the
   178 variables $i$, $j$ and $k$ in the goal $i \le j \Imp i \le j+k$ have type
   179 $\alpha::\{ord,plus\}$, although you may have expected them to have some
   180 numeric type, e.g. $nat$.  Instead you should have stated the goal as
   181 $(i::nat) \le j \Imp i \le j+k$, which causes all three variables to have
   182 type $nat$.
   183 
   184 \begin{warn}
   185   If resolution fails for no obvious reason, try setting
   186   \ttindex{show_types} to \texttt{true}, causing Isabelle to display
   187   types of terms.  Possibly set \ttindex{show_sorts} to \texttt{true} as
   188   well, causing Isabelle to display type classes and sorts.
   189 
   190   \index{unification!incompleteness of}
   191   Where function types are involved, Isabelle's unification code does not
   192   guarantee to find instantiations for type variables automatically.  Be
   193   prepared to use \ttindex{res_inst_tac} instead of \texttt{resolve_tac},
   194   possibly instantiating type variables.  Setting
   195   \ttindex{Unify.trace_types} to \texttt{true} causes Isabelle to report
   196   omitted search paths during unification.\index{tracing!of unification}
   197 \end{warn}
   198 
   199 
   200 \subsection{Binders}
   201 
   202 Hilbert's {\bf description} operator~$\varepsilon x. P[x]$ stands for
   203 some~$x$ satisfying~$P$, if such exists.  Since all terms in \HOL\ 
   204 denote something, a description is always meaningful, but we do not
   205 know its value unless $P$ defines it uniquely.  We may write
   206 descriptions as \cdx{Eps}($\lambda x. P[x]$) or use the syntax
   207 \hbox{\tt SOME~$x$.~$P[x]$}.
   208 
   209 Existential quantification is defined by
   210 \[ \exists x. P~x \;\equiv\; P(\varepsilon x. P~x). \]
   211 The unique existence quantifier, $\exists!x. P$, is defined in terms
   212 of~$\exists$ and~$\forall$.  An Isabelle binder, it admits nested
   213 quantifications.  For instance, $\exists!x\,y. P\,x\,y$ abbreviates
   214 $\exists!x. \exists!y. P\,x\,y$; note that this does not mean that there
   215 exists a unique pair $(x,y)$ satisfying~$P\,x\,y$.
   216 
   217 \medskip
   218 
   219 \index{*"! symbol}\index{*"? symbol}\index{HOL system@{\sc hol} system} The
   220 basic Isabelle/HOL binders have two notations.  Apart from the usual
   221 \texttt{ALL} and \texttt{EX} for $\forall$ and $\exists$, Isabelle/HOL also
   222 supports the original notation of Gordon's {\sc hol} system: \texttt{!}\ 
   223 and~\texttt{?}.  In the latter case, the existential quantifier \emph{must} be
   224 followed by a space; thus {\tt?x} is an unknown, while \verb'? x. f x=y' is a
   225 quantification.  Both notations are accepted for input.  The print mode
   226 ``\ttindexbold{HOL}'' governs the output notation.  If enabled (e.g.\ by
   227 passing option \texttt{-m HOL} to the \texttt{isabelle} executable),
   228 then~{\tt!}\ and~{\tt?}\ are displayed.
   229 
   230 \medskip
   231 
   232 If $\tau$ is a type of class \cldx{ord}, $P$ a formula and $x$ a
   233 variable of type $\tau$, then the term \cdx{LEAST}~$x. P[x]$ is defined
   234 to be the least (w.r.t.\ $\le$) $x$ such that $P~x$ holds (see
   235 Fig.~\ref{hol-defs}).  The definition uses Hilbert's $\varepsilon$
   236 choice operator, so \texttt{Least} is always meaningful, but may yield
   237 nothing useful in case there is not a unique least element satisfying
   238 $P$.\footnote{Class $ord$ does not require much of its instances, so
   239   $\le$ need not be a well-ordering, not even an order at all!}
   240 
   241 \medskip All these binders have priority 10.
   242 
   243 \begin{warn}
   244 The low priority of binders means that they need to be enclosed in
   245 parenthesis when they occur in the context of other operations.  For example,
   246 instead of $P \land \forall x. Q$ you need to write $P \land (\forall x. Q)$.
   247 \end{warn}
   248 
   249 
   250 \subsection{The let and case constructions}
   251 Local abbreviations can be introduced by a \texttt{let} construct whose
   252 syntax appears in Fig.\ts\ref{hol-grammar}.  Internally it is translated into
   253 the constant~\cdx{Let}.  It can be expanded by rewriting with its
   254 definition, \tdx{Let_def}.
   255 
   256 \HOL\ also defines the basic syntax
   257 \[\dquotes"case"~e~"of"~c@1~"=>"~e@1~"|" \dots "|"~c@n~"=>"~e@n\] 
   258 as a uniform means of expressing \texttt{case} constructs.  Therefore \texttt{case}
   259 and \sdx{of} are reserved words.  Initially, this is mere syntax and has no
   260 logical meaning.  By declaring translations, you can cause instances of the
   261 \texttt{case} construct to denote applications of particular case operators.
   262 This is what happens automatically for each \texttt{datatype} definition
   263 (see~\S\ref{sec:HOL:datatype}).
   264 
   265 \begin{warn}
   266 Both \texttt{if} and \texttt{case} constructs have as low a priority as
   267 quantifiers, which requires additional enclosing parentheses in the context
   268 of most other operations.  For example, instead of $f~x = {\tt if\dots
   269 then\dots else}\dots$ you need to write $f~x = ({\tt if\dots then\dots
   270 else\dots})$.
   271 \end{warn}
   272 
   273 \section{Rules of inference}
   274 
   275 \begin{figure}
   276 \begin{ttbox}\makeatother
   277 \tdx{refl}           t = (t::'a)
   278 \tdx{subst}          [| s = t; P s |] ==> P (t::'a)
   279 \tdx{ext}            (!!x::'a. (f x :: 'b) = g x) ==> (\%x. f x) = (\%x. g x)
   280 \tdx{impI}           (P ==> Q) ==> P-->Q
   281 \tdx{mp}             [| P-->Q;  P |] ==> Q
   282 \tdx{iff}            (P-->Q) --> (Q-->P) --> (P=Q)
   283 \tdx{selectI}        P(x::'a) ==> P(@x. P x)
   284 \tdx{True_or_False}  (P=True) | (P=False)
   285 \end{ttbox}
   286 \caption{The \texttt{HOL} rules} \label{hol-rules}
   287 \end{figure}
   288 
   289 Figure~\ref{hol-rules} shows the primitive inference rules of~\HOL{},
   290 with their~{\ML} names.  Some of the rules deserve additional
   291 comments:
   292 \begin{ttdescription}
   293 \item[\tdx{ext}] expresses extensionality of functions.
   294 \item[\tdx{iff}] asserts that logically equivalent formulae are
   295   equal.
   296 \item[\tdx{selectI}] gives the defining property of the Hilbert
   297   $\varepsilon$-operator.  It is a form of the Axiom of Choice.  The derived rule
   298   \tdx{select_equality} (see below) is often easier to use.
   299 \item[\tdx{True_or_False}] makes the logic classical.\footnote{In
   300     fact, the $\varepsilon$-operator already makes the logic classical, as
   301     shown by Diaconescu; see Paulson~\cite{paulson-COLOG} for details.}
   302 \end{ttdescription}
   303 
   304 
   305 \begin{figure}\hfuzz=4pt%suppress "Overfull \hbox" message
   306 \begin{ttbox}\makeatother
   307 \tdx{True_def}   True     == ((\%x::bool. x)=(\%x. x))
   308 \tdx{All_def}    All      == (\%P. P = (\%x. True))
   309 \tdx{Ex_def}     Ex       == (\%P. P(@x. P x))
   310 \tdx{False_def}  False    == (!P. P)
   311 \tdx{not_def}    not      == (\%P. P-->False)
   312 \tdx{and_def}    op &     == (\%P Q. !R. (P-->Q-->R) --> R)
   313 \tdx{or_def}     op |     == (\%P Q. !R. (P-->R) --> (Q-->R) --> R)
   314 \tdx{Ex1_def}    Ex1      == (\%P. ? x. P x & (! y. P y --> y=x))
   315 
   316 \tdx{o_def}      op o     == (\%(f::'b=>'c) g x::'a. f(g x))
   317 \tdx{if_def}     If P x y ==
   318               (\%P x y. @z::'a.(P=True --> z=x) & (P=False --> z=y))
   319 \tdx{Let_def}    Let s f  == f s
   320 \tdx{Least_def}  Least P  == @x. P(x) & (ALL y. P(y) --> x <= y)"
   321 \end{ttbox}
   322 \caption{The \texttt{HOL} definitions} \label{hol-defs}
   323 \end{figure}
   324 
   325 
   326 \HOL{} follows standard practice in higher-order logic: only a few
   327 connectives are taken as primitive, with the remainder defined obscurely
   328 (Fig.\ts\ref{hol-defs}).  Gordon's {\sc hol} system expresses the
   329 corresponding definitions \cite[page~270]{mgordon-hol} using
   330 object-equality~({\tt=}), which is possible because equality in
   331 higher-order logic may equate formulae and even functions over formulae.
   332 But theory~\HOL{}, like all other Isabelle theories, uses
   333 meta-equality~({\tt==}) for definitions.
   334 \begin{warn}
   335 The definitions above should never be expanded and are shown for completeness
   336 only.  Instead users should reason in terms of the derived rules shown below
   337 or, better still, using high-level tactics
   338 (see~\S\ref{sec:HOL:generic-packages}).
   339 \end{warn}
   340 
   341 Some of the rules mention type variables; for example, \texttt{refl}
   342 mentions the type variable~{\tt'a}.  This allows you to instantiate
   343 type variables explicitly by calling \texttt{res_inst_tac}.
   344 
   345 
   346 \begin{figure}
   347 \begin{ttbox}
   348 \tdx{sym}         s=t ==> t=s
   349 \tdx{trans}       [| r=s; s=t |] ==> r=t
   350 \tdx{ssubst}      [| t=s; P s |] ==> P t
   351 \tdx{box_equals}  [| a=b;  a=c;  b=d |] ==> c=d  
   352 \tdx{arg_cong}    x = y ==> f x = f y
   353 \tdx{fun_cong}    f = g ==> f x = g x
   354 \tdx{cong}        [| f = g; x = y |] ==> f x = g y
   355 \tdx{not_sym}     t ~= s ==> s ~= t
   356 \subcaption{Equality}
   357 
   358 \tdx{TrueI}       True 
   359 \tdx{FalseE}      False ==> P
   360 
   361 \tdx{conjI}       [| P; Q |] ==> P&Q
   362 \tdx{conjunct1}   [| P&Q |] ==> P
   363 \tdx{conjunct2}   [| P&Q |] ==> Q 
   364 \tdx{conjE}       [| P&Q;  [| P; Q |] ==> R |] ==> R
   365 
   366 \tdx{disjI1}      P ==> P|Q
   367 \tdx{disjI2}      Q ==> P|Q
   368 \tdx{disjE}       [| P | Q; P ==> R; Q ==> R |] ==> R
   369 
   370 \tdx{notI}        (P ==> False) ==> ~ P
   371 \tdx{notE}        [| ~ P;  P |] ==> R
   372 \tdx{impE}        [| P-->Q;  P;  Q ==> R |] ==> R
   373 \subcaption{Propositional logic}
   374 
   375 \tdx{iffI}        [| P ==> Q;  Q ==> P |] ==> P=Q
   376 \tdx{iffD1}       [| P=Q; P |] ==> Q
   377 \tdx{iffD2}       [| P=Q; Q |] ==> P
   378 \tdx{iffE}        [| P=Q; [| P --> Q; Q --> P |] ==> R |] ==> R
   379 %
   380 %\tdx{eqTrueI}     P ==> P=True 
   381 %\tdx{eqTrueE}     P=True ==> P 
   382 \subcaption{Logical equivalence}
   383 
   384 \end{ttbox}
   385 \caption{Derived rules for \HOL} \label{hol-lemmas1}
   386 \end{figure}
   387 
   388 
   389 \begin{figure}
   390 \begin{ttbox}\makeatother
   391 \tdx{allI}      (!!x. P x) ==> !x. P x
   392 \tdx{spec}      !x. P x ==> P x
   393 \tdx{allE}      [| !x. P x;  P x ==> R |] ==> R
   394 \tdx{all_dupE}  [| !x. P x;  [| P x; !x. P x |] ==> R |] ==> R
   395 
   396 \tdx{exI}       P x ==> ? x. P x
   397 \tdx{exE}       [| ? x. P x; !!x. P x ==> Q |] ==> Q
   398 
   399 \tdx{ex1I}      [| P a;  !!x. P x ==> x=a |] ==> ?! x. P x
   400 \tdx{ex1E}      [| ?! x. P x;  !!x. [| P x;  ! y. P y --> y=x |] ==> R 
   401           |] ==> R
   402 
   403 \tdx{select_equality} [| P a;  !!x. P x ==> x=a |] ==> (@x. P x) = a
   404 \subcaption{Quantifiers and descriptions}
   405 
   406 \tdx{ccontr}          (~P ==> False) ==> P
   407 \tdx{classical}       (~P ==> P) ==> P
   408 \tdx{excluded_middle} ~P | P
   409 
   410 \tdx{disjCI}          (~Q ==> P) ==> P|Q
   411 \tdx{exCI}            (! x. ~ P x ==> P a) ==> ? x. P x
   412 \tdx{impCE}           [| P-->Q; ~ P ==> R; Q ==> R |] ==> R
   413 \tdx{iffCE}           [| P=Q;  [| P;Q |] ==> R;  [| ~P; ~Q |] ==> R |] ==> R
   414 \tdx{notnotD}         ~~P ==> P
   415 \tdx{swap}            ~P ==> (~Q ==> P) ==> Q
   416 \subcaption{Classical logic}
   417 
   418 %\tdx{if_True}         (if True then x else y) = x
   419 %\tdx{if_False}        (if False then x else y) = y
   420 \tdx{if_P}            P ==> (if P then x else y) = x
   421 \tdx{if_not_P}        ~ P ==> (if P then x else y) = y
   422 \tdx{split_if}        P(if Q then x else y) = ((Q --> P x) & (~Q --> P y))
   423 \subcaption{Conditionals}
   424 \end{ttbox}
   425 \caption{More derived rules} \label{hol-lemmas2}
   426 \end{figure}
   427 
   428 Some derived rules are shown in Figures~\ref{hol-lemmas1}
   429 and~\ref{hol-lemmas2}, with their {\ML} names.  These include natural rules
   430 for the logical connectives, as well as sequent-style elimination rules for
   431 conjunctions, implications, and universal quantifiers.  
   432 
   433 Note the equality rules: \tdx{ssubst} performs substitution in
   434 backward proofs, while \tdx{box_equals} supports reasoning by
   435 simplifying both sides of an equation.
   436 
   437 The following simple tactics are occasionally useful:
   438 \begin{ttdescription}
   439 \item[\ttindexbold{strip_tac} $i$] applies \texttt{allI} and \texttt{impI}
   440   repeatedly to remove all outermost universal quantifiers and implications
   441   from subgoal $i$.
   442 \item[\ttindexbold{case_tac} {\tt"}$P${\tt"} $i$] performs case distinction
   443   on $P$ for subgoal $i$: the latter is replaced by two identical subgoals
   444   with the added assumptions $P$ and $\neg P$, respectively.
   445 \end{ttdescription}
   446 
   447 
   448 \begin{figure} 
   449 \begin{center}
   450 \begin{tabular}{rrr}
   451   \it name      &\it meta-type  & \it description \\ 
   452 \index{{}@\verb'{}' symbol}
   453   \verb|{}|     & $\alpha\,set$         & the empty set \\
   454   \cdx{insert}  & $[\alpha,\alpha\,set]\To \alpha\,set$
   455         & insertion of element \\
   456   \cdx{Collect} & $(\alpha\To bool)\To\alpha\,set$
   457         & comprehension \\
   458   \cdx{Compl}   & $\alpha\,set\To\alpha\,set$
   459         & complement \\
   460   \cdx{INTER} & $[\alpha\,set,\alpha\To\beta\,set]\To\beta\,set$
   461         & intersection over a set\\
   462   \cdx{UNION} & $[\alpha\,set,\alpha\To\beta\,set]\To\beta\,set$
   463         & union over a set\\
   464   \cdx{Inter} & $(\alpha\,set)set\To\alpha\,set$
   465         &set of sets intersection \\
   466   \cdx{Union} & $(\alpha\,set)set\To\alpha\,set$
   467         &set of sets union \\
   468   \cdx{Pow}   & $\alpha\,set \To (\alpha\,set)set$
   469         & powerset \\[1ex]
   470   \cdx{range}   & $(\alpha\To\beta )\To\beta\,set$
   471         & range of a function \\[1ex]
   472   \cdx{Ball}~~\cdx{Bex} & $[\alpha\,set,\alpha\To bool]\To bool$
   473         & bounded quantifiers
   474 \end{tabular}
   475 \end{center}
   476 \subcaption{Constants}
   477 
   478 \begin{center}
   479 \begin{tabular}{llrrr} 
   480   \it symbol &\it name     &\it meta-type & \it priority & \it description \\
   481   \sdx{INT}  & \cdx{INTER1}  & $(\alpha\To\beta\,set)\To\beta\,set$ & 10 & 
   482         intersection over a type\\
   483   \sdx{UN}  & \cdx{UNION1}  & $(\alpha\To\beta\,set)\To\beta\,set$ & 10 & 
   484         union over a type
   485 \end{tabular}
   486 \end{center}
   487 \subcaption{Binders} 
   488 
   489 \begin{center}
   490 \index{*"`"` symbol}
   491 \index{*": symbol}
   492 \index{*"<"= symbol}
   493 \begin{tabular}{rrrr} 
   494   \it symbol    & \it meta-type & \it priority & \it description \\ 
   495   \tt ``        & $[\alpha\To\beta ,\alpha\,set]\To  \beta\,set$
   496         & Left 90 & image \\
   497   \sdx{Int}     & $[\alpha\,set,\alpha\,set]\To\alpha\,set$
   498         & Left 70 & intersection ($\int$) \\
   499   \sdx{Un}      & $[\alpha\,set,\alpha\,set]\To\alpha\,set$
   500         & Left 65 & union ($\un$) \\
   501   \tt:          & $[\alpha ,\alpha\,set]\To bool$       
   502         & Left 50 & membership ($\in$) \\
   503   \tt <=        & $[\alpha\,set,\alpha\,set]\To bool$
   504         & Left 50 & subset ($\subseteq$) 
   505 \end{tabular}
   506 \end{center}
   507 \subcaption{Infixes}
   508 \caption{Syntax of the theory \texttt{Set}} \label{hol-set-syntax}
   509 \end{figure} 
   510 
   511 
   512 \begin{figure} 
   513 \begin{center} \tt\frenchspacing
   514 \index{*"! symbol}
   515 \begin{tabular}{rrr} 
   516   \it external          & \it internal  & \it description \\ 
   517   $a$ \ttilde: $b$      & \ttilde($a$ : $b$)    & \rm non-membership\\
   518   {\ttlbrace}$a@1$, $\ldots${\ttrbrace}  &  insert $a@1$ $\ldots$ {\ttlbrace}{\ttrbrace} & \rm finite set \\
   519   {\ttlbrace}$x$. $P[x]${\ttrbrace}        &  Collect($\lambda x. P[x]$) &
   520         \rm comprehension \\
   521   \sdx{INT} $x$:$A$. $B[x]$      & INTER $A$ $\lambda x. B[x]$ &
   522         \rm intersection \\
   523   \sdx{UN}{\tt\ }  $x$:$A$. $B[x]$      & UNION $A$ $\lambda x. B[x]$ &
   524         \rm union \\
   525   \sdx{ALL} $x$:$A$. $P[x]$ or \sdx{!} $x$:$A$. $P[x]$ &
   526         Ball $A$ $\lambda x. P[x]$ & 
   527         \rm bounded $\forall$ \\
   528   \sdx{EX}{\tt\ } $x$:$A$. $P[x]$ or \sdx{?} $x$:$A$. $P[x]$ & 
   529         Bex $A$ $\lambda x. P[x]$ & \rm bounded $\exists$
   530 \end{tabular}
   531 \end{center}
   532 \subcaption{Translations}
   533 
   534 \dquotes
   535 \[\begin{array}{rclcl}
   536     term & = & \hbox{other terms\ldots} \\
   537          & | & "{\ttlbrace}{\ttrbrace}" \\
   538          & | & "{\ttlbrace} " term\; ("," term)^* " {\ttrbrace}" \\
   539          & | & "{\ttlbrace} " id " . " formula " {\ttrbrace}" \\
   540          & | & term " `` " term \\
   541          & | & term " Int " term \\
   542          & | & term " Un " term \\
   543          & | & "INT~~"  id ":" term " . " term \\
   544          & | & "UN~~~"  id ":" term " . " term \\
   545          & | & "INT~~"  id~id^* " . " term \\
   546          & | & "UN~~~"  id~id^* " . " term \\[2ex]
   547  formula & = & \hbox{other formulae\ldots} \\
   548          & | & term " : " term \\
   549          & | & term " \ttilde: " term \\
   550          & | & term " <= " term \\
   551          & | & "ALL " id ":" term " . " formula
   552          & | & "!~" id ":" term " . " formula \\
   553          & | & "EX~~" id ":" term " . " formula
   554          & | & "?~" id ":" term " . " formula \\
   555   \end{array}
   556 \]
   557 \subcaption{Full Grammar}
   558 \caption{Syntax of the theory \texttt{Set} (continued)} \label{hol-set-syntax2}
   559 \end{figure} 
   560 
   561 
   562 \section{A formulation of set theory}
   563 Historically, higher-order logic gives a foundation for Russell and
   564 Whitehead's theory of classes.  Let us use modern terminology and call them
   565 {\bf sets}, but note that these sets are distinct from those of {\ZF} set
   566 theory, and behave more like {\ZF} classes.
   567 \begin{itemize}
   568 \item
   569 Sets are given by predicates over some type~$\sigma$.  Types serve to
   570 define universes for sets, but type-checking is still significant.
   571 \item
   572 There is a universal set (for each type).  Thus, sets have complements, and
   573 may be defined by absolute comprehension.
   574 \item
   575 Although sets may contain other sets as elements, the containing set must
   576 have a more complex type.
   577 \end{itemize}
   578 Finite unions and intersections have the same behaviour in \HOL\ as they
   579 do in~{\ZF}.  In \HOL\ the intersection of the empty set is well-defined,
   580 denoting the universal set for the given type.
   581 
   582 \subsection{Syntax of set theory}\index{*set type}
   583 \HOL's set theory is called \thydx{Set}.  The type $\alpha\,set$ is
   584 essentially the same as $\alpha\To bool$.  The new type is defined for
   585 clarity and to avoid complications involving function types in unification.
   586 The isomorphisms between the two types are declared explicitly.  They are
   587 very natural: \texttt{Collect} maps $\alpha\To bool$ to $\alpha\,set$, while
   588 \hbox{\tt op :} maps in the other direction (ignoring argument order).
   589 
   590 Figure~\ref{hol-set-syntax} lists the constants, infixes, and syntax
   591 translations.  Figure~\ref{hol-set-syntax2} presents the grammar of the new
   592 constructs.  Infix operators include union and intersection ($A\un B$
   593 and $A\int B$), the subset and membership relations, and the image
   594 operator~{\tt``}\@.  Note that $a$\verb|~:|$b$ is translated to
   595 $\neg(a\in b)$.  
   596 
   597 The $\{a@1,\ldots\}$ notation abbreviates finite sets constructed in
   598 the obvious manner using~\texttt{insert} and~$\{\}$:
   599 \begin{eqnarray*}
   600   \{a, b, c\} & \equiv &
   601   \texttt{insert} \, a \, ({\tt insert} \, b \, ({\tt insert} \, c \, \{\}))
   602 \end{eqnarray*}
   603 
   604 The set \hbox{\tt{\ttlbrace}$x$.\ $P[x]${\ttrbrace}} consists of all $x$ (of suitable type)
   605 that satisfy~$P[x]$, where $P[x]$ is a formula that may contain free
   606 occurrences of~$x$.  This syntax expands to \cdx{Collect}$(\lambda
   607 x. P[x])$.  It defines sets by absolute comprehension, which is impossible
   608 in~{\ZF}; the type of~$x$ implicitly restricts the comprehension.
   609 
   610 The set theory defines two {\bf bounded quantifiers}:
   611 \begin{eqnarray*}
   612    \forall x\in A. P[x] &\hbox{abbreviates}& \forall x. x\in A\imp P[x] \\
   613    \exists x\in A. P[x] &\hbox{abbreviates}& \exists x. x\in A\conj P[x]
   614 \end{eqnarray*}
   615 The constants~\cdx{Ball} and~\cdx{Bex} are defined
   616 accordingly.  Instead of \texttt{Ball $A$ $P$} and \texttt{Bex $A$ $P$} we may
   617 write\index{*"! symbol}\index{*"? symbol}
   618 \index{*ALL symbol}\index{*EX symbol} 
   619 %
   620 \hbox{\tt ALL~$x$:$A$.\ $P[x]$} and \hbox{\tt EX~$x$:$A$.\ $P[x]$}.  The
   621 original notation of Gordon's {\sc hol} system is supported as well: \sdx{!}\ 
   622 and \sdx{?}.
   623 
   624 Unions and intersections over sets, namely $\bigcup@{x\in A}B[x]$ and
   625 $\bigcap@{x\in A}B[x]$, are written 
   626 \sdx{UN}~\hbox{\tt$x$:$A$.\ $B[x]$} and
   627 \sdx{INT}~\hbox{\tt$x$:$A$.\ $B[x]$}.  
   628 
   629 Unions and intersections over types, namely $\bigcup@x B[x]$ and $\bigcap@x
   630 B[x]$, are written \sdx{UN}~\hbox{\tt$x$.\ $B[x]$} and
   631 \sdx{INT}~\hbox{\tt$x$.\ $B[x]$}.  They are equivalent to the previous
   632 union and intersection operators when $A$ is the universal set.
   633 
   634 The operators $\bigcup A$ and $\bigcap A$ act upon sets of sets.  They are
   635 not binders, but are equal to $\bigcup@{x\in A}x$ and $\bigcap@{x\in A}x$,
   636 respectively.
   637 
   638 
   639 
   640 \begin{figure} \underscoreon
   641 \begin{ttbox}
   642 \tdx{mem_Collect_eq}    (a : {\ttlbrace}x. P x{\ttrbrace}) = P a
   643 \tdx{Collect_mem_eq}    {\ttlbrace}x. x:A{\ttrbrace} = A
   644 
   645 \tdx{empty_def}         {\ttlbrace}{\ttrbrace}          == {\ttlbrace}x. False{\ttrbrace}
   646 \tdx{insert_def}        insert a B  == {\ttlbrace}x. x=a{\ttrbrace} Un B
   647 \tdx{Ball_def}          Ball A P    == ! x. x:A --> P x
   648 \tdx{Bex_def}           Bex A P     == ? x. x:A & P x
   649 \tdx{subset_def}        A <= B      == ! x:A. x:B
   650 \tdx{Un_def}            A Un B      == {\ttlbrace}x. x:A | x:B{\ttrbrace}
   651 \tdx{Int_def}           A Int B     == {\ttlbrace}x. x:A & x:B{\ttrbrace}
   652 \tdx{set_diff_def}      A - B       == {\ttlbrace}x. x:A & x~:B{\ttrbrace}
   653 \tdx{Compl_def}         Compl A     == {\ttlbrace}x. ~ x:A{\ttrbrace}
   654 \tdx{INTER_def}         INTER A B   == {\ttlbrace}y. ! x:A. y: B x{\ttrbrace}
   655 \tdx{UNION_def}         UNION A B   == {\ttlbrace}y. ? x:A. y: B x{\ttrbrace}
   656 \tdx{INTER1_def}        INTER1 B    == INTER {\ttlbrace}x. True{\ttrbrace} B 
   657 \tdx{UNION1_def}        UNION1 B    == UNION {\ttlbrace}x. True{\ttrbrace} B 
   658 \tdx{Inter_def}         Inter S     == (INT x:S. x)
   659 \tdx{Union_def}         Union S     == (UN  x:S. x)
   660 \tdx{Pow_def}           Pow A       == {\ttlbrace}B. B <= A{\ttrbrace}
   661 \tdx{image_def}         f``A        == {\ttlbrace}y. ? x:A. y=f x{\ttrbrace}
   662 \tdx{range_def}         range f     == {\ttlbrace}y. ? x. y=f x{\ttrbrace}
   663 \end{ttbox}
   664 \caption{Rules of the theory \texttt{Set}} \label{hol-set-rules}
   665 \end{figure}
   666 
   667 
   668 \begin{figure} \underscoreon
   669 \begin{ttbox}
   670 \tdx{CollectI}        [| P a |] ==> a : {\ttlbrace}x. P x{\ttrbrace}
   671 \tdx{CollectD}        [| a : {\ttlbrace}x. P x{\ttrbrace} |] ==> P a
   672 \tdx{CollectE}        [| a : {\ttlbrace}x. P x{\ttrbrace};  P a ==> W |] ==> W
   673 
   674 \tdx{ballI}           [| !!x. x:A ==> P x |] ==> ! x:A. P x
   675 \tdx{bspec}           [| ! x:A. P x;  x:A |] ==> P x
   676 \tdx{ballE}           [| ! x:A. P x;  P x ==> Q;  ~ x:A ==> Q |] ==> Q
   677 
   678 \tdx{bexI}            [| P x;  x:A |] ==> ? x:A. P x
   679 \tdx{bexCI}           [| ! x:A. ~ P x ==> P a;  a:A |] ==> ? x:A. P x
   680 \tdx{bexE}            [| ? x:A. P x;  !!x. [| x:A; P x |] ==> Q  |] ==> Q
   681 \subcaption{Comprehension and Bounded quantifiers}
   682 
   683 \tdx{subsetI}         (!!x. x:A ==> x:B) ==> A <= B
   684 \tdx{subsetD}         [| A <= B;  c:A |] ==> c:B
   685 \tdx{subsetCE}        [| A <= B;  ~ (c:A) ==> P;  c:B ==> P |] ==> P
   686 
   687 \tdx{subset_refl}     A <= A
   688 \tdx{subset_trans}    [| A<=B;  B<=C |] ==> A<=C
   689 
   690 \tdx{equalityI}       [| A <= B;  B <= A |] ==> A = B
   691 \tdx{equalityD1}      A = B ==> A<=B
   692 \tdx{equalityD2}      A = B ==> B<=A
   693 \tdx{equalityE}       [| A = B;  [| A<=B; B<=A |] ==> P |]  ==>  P
   694 
   695 \tdx{equalityCE}      [| A = B;  [| c:A; c:B |] ==> P;  
   696                            [| ~ c:A; ~ c:B |] ==> P 
   697                 |]  ==>  P
   698 \subcaption{The subset and equality relations}
   699 \end{ttbox}
   700 \caption{Derived rules for set theory} \label{hol-set1}
   701 \end{figure}
   702 
   703 
   704 \begin{figure} \underscoreon
   705 \begin{ttbox}
   706 \tdx{emptyE}   a : {\ttlbrace}{\ttrbrace} ==> P
   707 
   708 \tdx{insertI1} a : insert a B
   709 \tdx{insertI2} a : B ==> a : insert b B
   710 \tdx{insertE}  [| a : insert b A;  a=b ==> P;  a:A ==> P |] ==> P
   711 
   712 \tdx{ComplI}   [| c:A ==> False |] ==> c : Compl A
   713 \tdx{ComplD}   [| c : Compl A |] ==> ~ c:A
   714 
   715 \tdx{UnI1}     c:A ==> c : A Un B
   716 \tdx{UnI2}     c:B ==> c : A Un B
   717 \tdx{UnCI}     (~c:B ==> c:A) ==> c : A Un B
   718 \tdx{UnE}      [| c : A Un B;  c:A ==> P;  c:B ==> P |] ==> P
   719 
   720 \tdx{IntI}     [| c:A;  c:B |] ==> c : A Int B
   721 \tdx{IntD1}    c : A Int B ==> c:A
   722 \tdx{IntD2}    c : A Int B ==> c:B
   723 \tdx{IntE}     [| c : A Int B;  [| c:A; c:B |] ==> P |] ==> P
   724 
   725 \tdx{UN_I}     [| a:A;  b: B a |] ==> b: (UN x:A. B x)
   726 \tdx{UN_E}     [| b: (UN x:A. B x);  !!x.[| x:A;  b:B x |] ==> R |] ==> R
   727 
   728 \tdx{INT_I}    (!!x. x:A ==> b: B x) ==> b : (INT x:A. B x)
   729 \tdx{INT_D}    [| b: (INT x:A. B x);  a:A |] ==> b: B a
   730 \tdx{INT_E}    [| b: (INT x:A. B x);  b: B a ==> R;  ~ a:A ==> R |] ==> R
   731 
   732 \tdx{UnionI}   [| X:C;  A:X |] ==> A : Union C
   733 \tdx{UnionE}   [| A : Union C;  !!X.[| A:X;  X:C |] ==> R |] ==> R
   734 
   735 \tdx{InterI}   [| !!X. X:C ==> A:X |] ==> A : Inter C
   736 \tdx{InterD}   [| A : Inter C;  X:C |] ==> A:X
   737 \tdx{InterE}   [| A : Inter C;  A:X ==> R;  ~ X:C ==> R |] ==> R
   738 
   739 \tdx{PowI}     A<=B ==> A: Pow B
   740 \tdx{PowD}     A: Pow B ==> A<=B
   741 
   742 \tdx{imageI}   [| x:A |] ==> f x : f``A
   743 \tdx{imageE}   [| b : f``A;  !!x.[| b=f x;  x:A |] ==> P |] ==> P
   744 
   745 \tdx{rangeI}   f x : range f
   746 \tdx{rangeE}   [| b : range f;  !!x.[| b=f x |] ==> P |] ==> P
   747 \end{ttbox}
   748 \caption{Further derived rules for set theory} \label{hol-set2}
   749 \end{figure}
   750 
   751 
   752 \subsection{Axioms and rules of set theory}
   753 Figure~\ref{hol-set-rules} presents the rules of theory \thydx{Set}.  The
   754 axioms \tdx{mem_Collect_eq} and \tdx{Collect_mem_eq} assert
   755 that the functions \texttt{Collect} and \hbox{\tt op :} are isomorphisms.  Of
   756 course, \hbox{\tt op :} also serves as the membership relation.
   757 
   758 All the other axioms are definitions.  They include the empty set, bounded
   759 quantifiers, unions, intersections, complements and the subset relation.
   760 They also include straightforward constructions on functions: image~({\tt``})
   761 and \texttt{range}.
   762 
   763 %The predicate \cdx{inj_on} is used for simulating type definitions.
   764 %The statement ${\tt inj_on}~f~A$ asserts that $f$ is injective on the
   765 %set~$A$, which specifies a subset of its domain type.  In a type
   766 %definition, $f$ is the abstraction function and $A$ is the set of valid
   767 %representations; we should not expect $f$ to be injective outside of~$A$.
   768 
   769 %\begin{figure} \underscoreon
   770 %\begin{ttbox}
   771 %\tdx{Inv_f_f}    inj f ==> Inv f (f x) = x
   772 %\tdx{f_Inv_f}    y : range f ==> f(Inv f y) = y
   773 %
   774 %\tdx{Inv_injective}
   775 %    [| Inv f x=Inv f y; x: range f;  y: range f |] ==> x=y
   776 %
   777 %
   778 %\tdx{monoI}      [| !!A B. A <= B ==> f A <= f B |] ==> mono f
   779 %\tdx{monoD}      [| mono f;  A <= B |] ==> f A <= f B
   780 %
   781 %\tdx{injI}       [| !! x y. f x = f y ==> x=y |] ==> inj f
   782 %\tdx{inj_inverseI}              (!!x. g(f x) = x) ==> inj f
   783 %\tdx{injD}       [| inj f; f x = f y |] ==> x=y
   784 %
   785 %\tdx{inj_onI}  (!!x y. [| f x=f y; x:A; y:A |] ==> x=y) ==> inj_on f A
   786 %\tdx{inj_onD}  [| inj_on f A;  f x=f y;  x:A;  y:A |] ==> x=y
   787 %
   788 %\tdx{inj_on_inverseI}
   789 %    (!!x. x:A ==> g(f x) = x) ==> inj_on f A
   790 %\tdx{inj_on_contraD}
   791 %    [| inj_on f A;  x~=y;  x:A;  y:A |] ==> ~ f x=f y
   792 %\end{ttbox}
   793 %\caption{Derived rules involving functions} \label{hol-fun}
   794 %\end{figure}
   795 
   796 
   797 \begin{figure} \underscoreon
   798 \begin{ttbox}
   799 \tdx{Union_upper}     B:A ==> B <= Union A
   800 \tdx{Union_least}     [| !!X. X:A ==> X<=C |] ==> Union A <= C
   801 
   802 \tdx{Inter_lower}     B:A ==> Inter A <= B
   803 \tdx{Inter_greatest}  [| !!X. X:A ==> C<=X |] ==> C <= Inter A
   804 
   805 \tdx{Un_upper1}       A <= A Un B
   806 \tdx{Un_upper2}       B <= A Un B
   807 \tdx{Un_least}        [| A<=C;  B<=C |] ==> A Un B <= C
   808 
   809 \tdx{Int_lower1}      A Int B <= A
   810 \tdx{Int_lower2}      A Int B <= B
   811 \tdx{Int_greatest}    [| C<=A;  C<=B |] ==> C <= A Int B
   812 \end{ttbox}
   813 \caption{Derived rules involving subsets} \label{hol-subset}
   814 \end{figure}
   815 
   816 
   817 \begin{figure} \underscoreon   \hfuzz=4pt%suppress "Overfull \hbox" message
   818 \begin{ttbox}
   819 \tdx{Int_absorb}        A Int A = A
   820 \tdx{Int_commute}       A Int B = B Int A
   821 \tdx{Int_assoc}         (A Int B) Int C  =  A Int (B Int C)
   822 \tdx{Int_Un_distrib}    (A Un B)  Int C  =  (A Int C) Un (B Int C)
   823 
   824 \tdx{Un_absorb}         A Un A = A
   825 \tdx{Un_commute}        A Un B = B Un A
   826 \tdx{Un_assoc}          (A Un B)  Un C  =  A Un (B Un C)
   827 \tdx{Un_Int_distrib}    (A Int B) Un C  =  (A Un C) Int (B Un C)
   828 
   829 \tdx{Compl_disjoint}    A Int (Compl A) = {\ttlbrace}x. False{\ttrbrace}
   830 \tdx{Compl_partition}   A Un  (Compl A) = {\ttlbrace}x. True{\ttrbrace}
   831 \tdx{double_complement} Compl(Compl A) = A
   832 \tdx{Compl_Un}          Compl(A Un B)  = (Compl A) Int (Compl B)
   833 \tdx{Compl_Int}         Compl(A Int B) = (Compl A) Un (Compl B)
   834 
   835 \tdx{Union_Un_distrib}  Union(A Un B) = (Union A) Un (Union B)
   836 \tdx{Int_Union}         A Int (Union B) = (UN C:B. A Int C)
   837 \tdx{Un_Union_image}    (UN x:C.(A x) Un (B x)) = Union(A``C) Un Union(B``C)
   838 
   839 \tdx{Inter_Un_distrib}  Inter(A Un B) = (Inter A) Int (Inter B)
   840 \tdx{Un_Inter}          A Un (Inter B) = (INT C:B. A Un C)
   841 \tdx{Int_Inter_image}   (INT x:C.(A x) Int (B x)) = Inter(A``C) Int Inter(B``C)
   842 \end{ttbox}
   843 \caption{Set equalities} \label{hol-equalities}
   844 \end{figure}
   845 
   846 
   847 Figures~\ref{hol-set1} and~\ref{hol-set2} present derived rules.  Most are
   848 obvious and resemble rules of Isabelle's {\ZF} set theory.  Certain rules,
   849 such as \tdx{subsetCE}, \tdx{bexCI} and \tdx{UnCI},
   850 are designed for classical reasoning; the rules \tdx{subsetD},
   851 \tdx{bexI}, \tdx{Un1} and~\tdx{Un2} are not
   852 strictly necessary but yield more natural proofs.  Similarly,
   853 \tdx{equalityCE} supports classical reasoning about extensionality,
   854 after the fashion of \tdx{iffCE}.  See the file \texttt{HOL/Set.ML} for
   855 proofs pertaining to set theory.
   856 
   857 Figure~\ref{hol-subset} presents lattice properties of the subset relation.
   858 Unions form least upper bounds; non-empty intersections form greatest lower
   859 bounds.  Reasoning directly about subsets often yields clearer proofs than
   860 reasoning about the membership relation.  See the file \texttt{HOL/subset.ML}.
   861 
   862 Figure~\ref{hol-equalities} presents many common set equalities.  They
   863 include commutative, associative and distributive laws involving unions,
   864 intersections and complements.  For a complete listing see the file {\tt
   865 HOL/equalities.ML}.
   866 
   867 \begin{warn}
   868 \texttt{Blast_tac} proves many set-theoretic theorems automatically.
   869 Hence you seldom need to refer to the theorems above.
   870 \end{warn}
   871 
   872 \begin{figure}
   873 \begin{center}
   874 \begin{tabular}{rrr}
   875   \it name      &\it meta-type  & \it description \\ 
   876   \cdx{inj}~~\cdx{surj}& $(\alpha\To\beta )\To bool$
   877         & injective/surjective \\
   878   \cdx{inj_on}        & $[\alpha\To\beta ,\alpha\,set]\To bool$
   879         & injective over subset\\
   880   \cdx{inv} & $(\alpha\To\beta)\To(\beta\To\alpha)$ & inverse function
   881 \end{tabular}
   882 \end{center}
   883 
   884 \underscoreon
   885 \begin{ttbox}
   886 \tdx{inj_def}         inj f      == ! x y. f x=f y --> x=y
   887 \tdx{surj_def}        surj f     == ! y. ? x. y=f x
   888 \tdx{inj_on_def}      inj_on f A == !x:A. !y:A. f x=f y --> x=y
   889 \tdx{inv_def}         inv f      == (\%y. @x. f(x)=y)
   890 \end{ttbox}
   891 \caption{Theory \thydx{Fun}} \label{fig:HOL:Fun}
   892 \end{figure}
   893 
   894 \subsection{Properties of functions}\nopagebreak
   895 Figure~\ref{fig:HOL:Fun} presents a theory of simple properties of functions.
   896 Note that ${\tt inv}~f$ uses Hilbert's $\varepsilon$ to yield an inverse
   897 of~$f$.  See the file \texttt{HOL/Fun.ML} for a complete listing of the derived
   898 rules.  Reasoning about function composition (the operator~\sdx{o}) and the
   899 predicate~\cdx{surj} is done simply by expanding the definitions.
   900 
   901 There is also a large collection of monotonicity theorems for constructions
   902 on sets in the file \texttt{HOL/mono.ML}.
   903 
   904 
   905 \section{Generic packages}
   906 \label{sec:HOL:generic-packages}
   907 
   908 \HOL\ instantiates most of Isabelle's generic packages, making available the
   909 simplifier and the classical reasoner.
   910 
   911 \subsection{Simplification and substitution}
   912 
   913 Simplification tactics tactics such as \texttt{Asm_simp_tac} and \texttt{Full_simp_tac} use the default simpset
   914 (\texttt{simpset()}), which works for most purposes.  A quite minimal
   915 simplification set for higher-order logic is~\ttindexbold{HOL_ss};
   916 even more frugal is \ttindexbold{HOL_basic_ss}.  Equality~($=$), which
   917 also expresses logical equivalence, may be used for rewriting.  See
   918 the file \texttt{HOL/simpdata.ML} for a complete listing of the basic
   919 simplification rules.
   920 
   921 See \iflabelundefined{chap:classical}{the {\em Reference Manual\/}}%
   922 {Chaps.\ts\ref{substitution} and~\ref{simp-chap}} for details of substitution
   923 and simplification.
   924 
   925 \begin{warn}\index{simplification!of conjunctions}%
   926   Reducing $a=b\conj P(a)$ to $a=b\conj P(b)$ is sometimes advantageous.  The
   927   left part of a conjunction helps in simplifying the right part.  This effect
   928   is not available by default: it can be slow.  It can be obtained by
   929   including \ttindex{conj_cong} in a simpset, \verb$addcongs [conj_cong]$.
   930 \end{warn}
   931 
   932 If the simplifier cannot use a certain rewrite rule --- either because
   933 of nontermination or because its left-hand side is too flexible ---
   934 then you might try \texttt{stac}:
   935 \begin{ttdescription}
   936 \item[\ttindexbold{stac} $thm$ $i,$] where $thm$ is of the form $lhs = rhs$,
   937   replaces in subgoal $i$ instances of $lhs$ by corresponding instances of
   938   $rhs$.  In case of multiple instances of $lhs$ in subgoal $i$, backtracking
   939   may be necessary to select the desired ones.
   940 
   941 If $thm$ is a conditional equality, the instantiated condition becomes an
   942 additional (first) subgoal.
   943 \end{ttdescription}
   944 
   945  \HOL{} provides the tactic \ttindex{hyp_subst_tac}, which substitutes
   946   for an equality throughout a subgoal and its hypotheses.  This tactic uses
   947   \HOL's general substitution rule.
   948 
   949 \subsubsection{Case splitting}
   950 \label{subsec:HOL:case:splitting}
   951 
   952 \HOL{} also provides convenient means for case splitting during
   953 rewriting. Goals containing a subterm of the form \texttt{if}~$b$~{\tt
   954 then\dots else\dots} often require a case distinction on $b$. This is
   955 expressed by the theorem \tdx{split_if}:
   956 $$
   957 \Var{P}(\mbox{\tt if}~\Var{b}~{\tt then}~\Var{x}~\mbox{\tt else}~\Var{y})~=~
   958 ((\Var{b} \to \Var{P}(\Var{x})) \land (\neg \Var{b} \to \Var{P}(\Var{y})))
   959 \eqno{(*)}
   960 $$
   961 For example, a simple instance of $(*)$ is
   962 \[
   963 x \in (\mbox{\tt if}~x \in A~{\tt then}~A~\mbox{\tt else}~\{x\})~=~
   964 ((x \in A \to x \in A) \land (x \notin A \to x \in \{x\}))
   965 \]
   966 Because $(*)$ is too general as a rewrite rule for the simplifier (the
   967 left-hand side is not a higher-order pattern in the sense of
   968 \iflabelundefined{chap:simplification}{the {\em Reference Manual\/}}%
   969 {Chap.\ts\ref{chap:simplification}}), there is a special infix function 
   970 \ttindexbold{addsplits} of type \texttt{simpset * thm list -> simpset}
   971 (analogous to \texttt{addsimps}) that adds rules such as $(*)$ to a
   972 simpset, as in
   973 \begin{ttbox}
   974 by(simp_tac (simpset() addsplits [split_if]) 1);
   975 \end{ttbox}
   976 The effect is that after each round of simplification, one occurrence of
   977 \texttt{if} is split acording to \texttt{split_if}, until all occurences of
   978 \texttt{if} have been eliminated.
   979 
   980 It turns out that using \texttt{split_if} is almost always the right thing to
   981 do. Hence \texttt{split_if} is already included in the default simpset. If
   982 you want to delete it from a simpset, use \ttindexbold{delsplits}, which is
   983 the inverse of \texttt{addsplits}:
   984 \begin{ttbox}
   985 by(simp_tac (simpset() delsplits [split_if]) 1);
   986 \end{ttbox}
   987 
   988 In general, \texttt{addsplits} accepts rules of the form
   989 \[
   990 \Var{P}(c~\Var{x@1}~\dots~\Var{x@n})~=~ rhs
   991 \]
   992 where $c$ is a constant and $rhs$ is arbitrary. Note that $(*)$ is of the
   993 right form because internally the left-hand side is
   994 $\Var{P}(\mathtt{If}~\Var{b}~\Var{x}~~\Var{y})$. Important further examples
   995 are splitting rules for \texttt{case} expressions (see~\S\ref{subsec:list}
   996 and~\S\ref{subsec:datatype:basics}).
   997 
   998 Analogous to \texttt{Addsimps} and \texttt{Delsimps}, there are also
   999 imperative versions of \texttt{addsplits} and \texttt{delsplits}
  1000 \begin{ttbox}
  1001 \ttindexbold{Addsplits}: thm list -> unit
  1002 \ttindexbold{Delsplits}: thm list -> unit
  1003 \end{ttbox}
  1004 for adding splitting rules to, and deleting them from the current simpset.
  1005 
  1006 \subsection{Classical reasoning}
  1007 
  1008 \HOL\ derives classical introduction rules for $\disj$ and~$\exists$, as
  1009 well as classical elimination rules for~$\imp$ and~$\bimp$, and the swap
  1010 rule; recall Fig.\ts\ref{hol-lemmas2} above.
  1011 
  1012 The classical reasoner is installed.  Tactics such as \texttt{Blast_tac} and
  1013 {\tt Best_tac} refer to the default claset (\texttt{claset()}), which works
  1014 for most purposes.  Named clasets include \ttindexbold{prop_cs}, which
  1015 includes the propositional rules, and \ttindexbold{HOL_cs}, which also
  1016 includes quantifier rules.  See the file \texttt{HOL/cladata.ML} for lists of
  1017 the classical rules,
  1018 and \iflabelundefined{chap:classical}{the {\em Reference Manual\/}}%
  1019 {Chap.\ts\ref{chap:classical}} for more discussion of classical proof methods.
  1020 
  1021 
  1022 \section{Calling the decision procedure SVC}\label{sec:HOL:SVC}
  1023 
  1024 \index{SVC decision procedure|(}
  1025 
  1026 The Stanford Validity Checker (SVC) is a tool that can check the validity of
  1027 certain types of formulae.  If it is installed on your machine, then
  1028 Isabelle/HOL can be configured to call it through the tactic
  1029 \ttindex{svc_tac}.  It is ideal for large tautologies and complex problems in
  1030 linear arithmetic.  Subexpressions that SVC cannot handle are automatically
  1031 replaced by variables, so you can call the tactic on any subgoal.  See the
  1032 file \texttt{HOL/ex/svc_test.ML} for examples.
  1033 \begin{ttbox} 
  1034 svc_tac   : int -> tactic
  1035 Svc.trace : bool ref      \hfill{\bf initially false}
  1036 \end{ttbox}
  1037 
  1038 \begin{ttdescription}
  1039 \item[\ttindexbold{svc_tac} $i$] attempts to prove subgoal~$i$ by translating
  1040   it into a formula recognized by~SVC\@.  If it succeeds then the subgoal is
  1041   removed.  It fails if SVC is unable to prove the subgoal.  It crashes with
  1042   an error message if SVC appears not to be installed.  Numeric variables may
  1043   have types \texttt{nat}, \texttt{int} or \texttt{real}.
  1044 
  1045 \item[\ttindexbold{Svc.trace}] is a flag that, if set, causes \texttt{svc_tac}
  1046   to trace its operations: abstraction of the subgoal, translation to SVC
  1047   syntax, SVC's response.
  1048 \end{ttdescription}
  1049 
  1050 Here is an example, with tracing turned on:
  1051 \begin{ttbox}
  1052 set Svc.trace;
  1053 {\out val it : bool = true}
  1054 Goal "(#3::nat)*a <= #2 + #4*b + #6*c  & #11 <= #2*a + b + #2*c & \ttback
  1055 \ttback     a + #3*b <= #5 + #2*c  --> #2 + #3*b <= #2*a + #6*c";
  1056 
  1057 by (svc_tac 1);
  1058 {\out Subgoal abstracted to}
  1059 {\out #3 * a <= #2 + #4 * b + #6 * c &}
  1060 {\out #11 <= #2 * a + b + #2 * c & a + #3 * b <= #5 + #2 * c -->}
  1061 {\out #2 + #3 * b <= #2 * a + #6 * c}
  1062 {\out Calling SVC:}
  1063 {\out (=> (<= 0  (F_c) )  (=> (<= 0  (F_b) )  (=> (<= 0  (F_a) )}
  1064 {\out   (=> (AND (<= {* 3  (F_a) }  {+ {+ 2  {* 4  (F_b) } }  }
  1065 {\out {* 6  (F_c) } } )  (AND (<= 11  {+ {+ {* 2  (F_a) }  (F_b) }}
  1066 {\out   {* 2  (F_c) } } )  (<= {+ (F_a)  {* 3  (F_b) } }  {+ 5  }
  1067 {\out {* 2  (F_c) } } ) ) )  (< {+ 2  {* 3  (F_b) } }  {+ 1  {+ }
  1068 {\out {* 2  (F_a) }  {* 6  (F_c) } } } ) ) ) ) ) }
  1069 {\out SVC Returns:}
  1070 {\out VALID}
  1071 {\out Level 1}
  1072 {\out #3 * a <= #2 + #4 * b + #6 * c &}
  1073 {\out #11 <= #2 * a + b + #2 * c & a + #3 * b <= #5 + #2 * c -->}
  1074 {\out #2 + #3 * b <= #2 * a + #6 * c}
  1075 {\out No subgoals!}
  1076 \end{ttbox}
  1077 
  1078 
  1079 \begin{warn}
  1080 Calling \ttindex{svc_tac} entails an above-average risk of
  1081 unsoundness.  Isabelle does not check SVC's result independently.  Moreover,
  1082 the tactic translates the submitted formula using code that lies outside
  1083 Isabelle's inference core.  Theorems that depend upon results proved using SVC
  1084 (and other oracles) are displayed with the annotation \texttt{[!]} attached.
  1085 You can also use \texttt{\#der (rep_thm $th$)} to examine the proof object of
  1086 theorem~$th$, as described in the \emph{Reference Manual}.  
  1087 \end{warn}
  1088 
  1089 To start, first download SVC from the Internet at URL
  1090 \begin{ttbox}
  1091 http://agamemnon.stanford.edu/~levitt/vc/index.html
  1092 \end{ttbox}
  1093 and install it using the instructions supplied.  SVC requires two environment
  1094 variables:
  1095 \begin{ttdescription}
  1096 \item[\ttindexbold{SVC_HOME}] is an absolute pathname to the SVC
  1097     distribution directory. 
  1098     
  1099   \item[\ttindexbold{SVC_MACHINE}] identifies the type of computer and
  1100     operating system.  
  1101 \end{ttdescription}
  1102 You can set these environment variables either using the Unix shell or through
  1103 an Isabelle \texttt{settings} file.  Isabelle assumes SVC to be installed if
  1104 \texttt{SVC_HOME} is defined.
  1105 
  1106 \paragraph*{Acknowledgement.}  This interface uses code supplied by S{\o}ren
  1107 Heilmann.
  1108 
  1109 
  1110 \index{SVC decision procedure|)}
  1111 
  1112 
  1113 
  1114 
  1115 \section{Types}\label{sec:HOL:Types}
  1116 This section describes \HOL's basic predefined types ($\alpha \times
  1117 \beta$, $\alpha + \beta$, $nat$ and $\alpha \; list$) and ways for
  1118 introducing new types in general.  The most important type
  1119 construction, the \texttt{datatype}, is treated separately in
  1120 \S\ref{sec:HOL:datatype}.
  1121 
  1122 
  1123 \subsection{Product and sum types}\index{*"* type}\index{*"+ type}
  1124 \label{subsec:prod-sum}
  1125 
  1126 \begin{figure}[htbp]
  1127 \begin{constants}
  1128   \it symbol    & \it meta-type &           & \it description \\ 
  1129   \cdx{Pair}    & $[\alpha,\beta]\To \alpha\times\beta$
  1130         & & ordered pairs $(a,b)$ \\
  1131   \cdx{fst}     & $\alpha\times\beta \To \alpha$        & & first projection\\
  1132   \cdx{snd}     & $\alpha\times\beta \To \beta$         & & second projection\\
  1133   \cdx{split}   & $[[\alpha,\beta]\To\gamma, \alpha\times\beta] \To \gamma$ 
  1134         & & generalized projection\\
  1135   \cdx{Sigma}  & 
  1136         $[\alpha\,set, \alpha\To\beta\,set]\To(\alpha\times\beta)set$ &
  1137         & general sum of sets
  1138 \end{constants}
  1139 \begin{ttbox}\makeatletter
  1140 %\tdx{fst_def}      fst p     == @a. ? b. p = (a,b)
  1141 %\tdx{snd_def}      snd p     == @b. ? a. p = (a,b)
  1142 %\tdx{split_def}    split c p == c (fst p) (snd p)
  1143 \tdx{Sigma_def}    Sigma A B == UN x:A. UN y:B x. {\ttlbrace}(x,y){\ttrbrace}
  1144 
  1145 \tdx{Pair_eq}      ((a,b) = (a',b')) = (a=a' & b=b')
  1146 \tdx{Pair_inject}  [| (a, b) = (a',b');  [| a=a';  b=b' |] ==> R |] ==> R
  1147 \tdx{PairE}        [| !!x y. p = (x,y) ==> Q |] ==> Q
  1148 
  1149 \tdx{fst_conv}     fst (a,b) = a
  1150 \tdx{snd_conv}     snd (a,b) = b
  1151 \tdx{surjective_pairing}  p = (fst p,snd p)
  1152 
  1153 \tdx{split}        split c (a,b) = c a b
  1154 \tdx{split_split}  R(split c p) = (! x y. p = (x,y) --> R(c x y))
  1155 
  1156 \tdx{SigmaI}    [| a:A;  b:B a |] ==> (a,b) : Sigma A B
  1157 \tdx{SigmaE}    [| c:Sigma A B; !!x y.[| x:A; y:B x; c=(x,y) |] ==> P |] ==> P
  1158 \end{ttbox}
  1159 \caption{Type $\alpha\times\beta$}\label{hol-prod}
  1160 \end{figure} 
  1161 
  1162 Theory \thydx{Prod} (Fig.\ts\ref{hol-prod}) defines the product type
  1163 $\alpha\times\beta$, with the ordered pair syntax $(a, b)$.  General
  1164 tuples are simulated by pairs nested to the right:
  1165 \begin{center}
  1166 \begin{tabular}{c|c}
  1167 external & internal \\
  1168 \hline
  1169 $\tau@1 \times \dots \times \tau@n$ & $\tau@1 \times (\dots (\tau@{n-1} \times \tau@n)\dots)$ \\
  1170 \hline
  1171 $(t@1,\dots,t@n)$ & $(t@1,(\dots,(t@{n-1},t@n)\dots)$ \\
  1172 \end{tabular}
  1173 \end{center}
  1174 In addition, it is possible to use tuples
  1175 as patterns in abstractions:
  1176 \begin{center}
  1177 {\tt\%($x$,$y$). $t$} \quad stands for\quad \texttt{split(\%$x$\thinspace$y$.\ $t$)} 
  1178 \end{center}
  1179 Nested patterns are also supported.  They are translated stepwise:
  1180 {\tt\%($x$,$y$,$z$). $t$} $\leadsto$ {\tt\%($x$,($y$,$z$)). $t$} $\leadsto$
  1181 {\tt split(\%$x$.\%($y$,$z$). $t$)} $\leadsto$ \texttt{split(\%$x$. split(\%$y$
  1182   $z$.\ $t$))}.  The reverse translation is performed upon printing.
  1183 \begin{warn}
  1184   The translation between patterns and \texttt{split} is performed automatically
  1185   by the parser and printer.  Thus the internal and external form of a term
  1186   may differ, which can affects proofs.  For example the term {\tt
  1187   (\%(x,y).(y,x))(a,b)} requires the theorem \texttt{split} (which is in the
  1188   default simpset) to rewrite to {\tt(b,a)}.
  1189 \end{warn}
  1190 In addition to explicit $\lambda$-abstractions, patterns can be used in any
  1191 variable binding construct which is internally described by a
  1192 $\lambda$-abstraction.  Some important examples are
  1193 \begin{description}
  1194 \item[Let:] \texttt{let {\it pattern} = $t$ in $u$}
  1195 \item[Quantifiers:] \texttt{!~{\it pattern}:$A$.~$P$}
  1196 \item[Choice:] {\underscoreon \tt @~{\it pattern}~.~$P$}
  1197 \item[Set operations:] \texttt{UN~{\it pattern}:$A$.~$B$}
  1198 \item[Sets:] \texttt{{\ttlbrace}~{\it pattern}~.~$P$~{\ttrbrace}}
  1199 \end{description}
  1200 
  1201 There is a simple tactic which supports reasoning about patterns:
  1202 \begin{ttdescription}
  1203 \item[\ttindexbold{split_all_tac} $i$] replaces in subgoal $i$ all
  1204   {\tt!!}-quantified variables of product type by individual variables for
  1205   each component.  A simple example:
  1206 \begin{ttbox}
  1207 {\out 1. !!p. (\%(x,y,z). (x, y, z)) p = p}
  1208 by(split_all_tac 1);
  1209 {\out 1. !!x xa ya. (\%(x,y,z). (x, y, z)) (x, xa, ya) = (x, xa, ya)}
  1210 \end{ttbox}
  1211 \end{ttdescription}
  1212 
  1213 Theory \texttt{Prod} also introduces the degenerate product type \texttt{unit}
  1214 which contains only a single element named {\tt()} with the property
  1215 \begin{ttbox}
  1216 \tdx{unit_eq}       u = ()
  1217 \end{ttbox}
  1218 \bigskip
  1219 
  1220 Theory \thydx{Sum} (Fig.~\ref{hol-sum}) defines the sum type $\alpha+\beta$
  1221 which associates to the right and has a lower priority than $*$: $\tau@1 +
  1222 \tau@2 + \tau@3*\tau@4$ means $\tau@1 + (\tau@2 + (\tau@3*\tau@4))$.
  1223 
  1224 The definition of products and sums in terms of existing types is not
  1225 shown.  The constructions are fairly standard and can be found in the
  1226 respective theory files. Although the sum and product types are
  1227 constructed manually for foundational reasons, they are represented as
  1228 actual datatypes later (see \S\ref{subsec:datatype:rep_datatype}).
  1229 Therefore, the theory \texttt{Datatype} should be used instead of
  1230 \texttt{Sum} or \texttt{Prod}.
  1231 
  1232 \begin{figure}
  1233 \begin{constants}
  1234   \it symbol    & \it meta-type &           & \it description \\ 
  1235   \cdx{Inl}     & $\alpha \To \alpha+\beta$    & & first injection\\
  1236   \cdx{Inr}     & $\beta \To \alpha+\beta$     & & second injection\\
  1237   \cdx{sum_case} & $[\alpha\To\gamma, \beta\To\gamma, \alpha+\beta] \To\gamma$
  1238         & & conditional
  1239 \end{constants}
  1240 \begin{ttbox}\makeatletter
  1241 \tdx{Inl_not_Inr}    Inl a ~= Inr b
  1242 
  1243 \tdx{inj_Inl}        inj Inl
  1244 \tdx{inj_Inr}        inj Inr
  1245 
  1246 \tdx{sumE}           [| !!x. P(Inl x);  !!y. P(Inr y) |] ==> P s
  1247 
  1248 \tdx{sum_case_Inl}   sum_case f g (Inl x) = f x
  1249 \tdx{sum_case_Inr}   sum_case f g (Inr x) = g x
  1250 
  1251 \tdx{surjective_sum} sum_case (\%x. f(Inl x)) (\%y. f(Inr y)) s = f s
  1252 \tdx{sum.split_case} R(sum_case f g s) = ((! x. s = Inl(x) --> R(f(x))) &
  1253                                      (! y. s = Inr(y) --> R(g(y))))
  1254 \end{ttbox}
  1255 \caption{Type $\alpha+\beta$}\label{hol-sum}
  1256 \end{figure}
  1257 
  1258 \begin{figure}
  1259 \index{*"< symbol}
  1260 \index{*"* symbol}
  1261 \index{*div symbol}
  1262 \index{*mod symbol}
  1263 \index{*"+ symbol}
  1264 \index{*"- symbol}
  1265 \begin{constants}
  1266   \it symbol    & \it meta-type & \it priority & \it description \\ 
  1267   \cdx{0}       & $nat$         & & zero \\
  1268   \cdx{Suc}     & $nat \To nat$ & & successor function\\
  1269 % \cdx{nat_case} & $[\alpha, nat\To\alpha, nat] \To\alpha$ & & conditional\\
  1270 % \cdx{nat_rec} & $[nat, \alpha, [nat, \alpha]\To\alpha] \To \alpha$
  1271 %        & & primitive recursor\\
  1272   \tt *         & $[nat,nat]\To nat$    &  Left 70      & multiplication \\
  1273   \tt div       & $[nat,nat]\To nat$    &  Left 70      & division\\
  1274   \tt mod       & $[nat,nat]\To nat$    &  Left 70      & modulus\\
  1275   \tt +         & $[nat,nat]\To nat$    &  Left 65      & addition\\
  1276   \tt -         & $[nat,nat]\To nat$    &  Left 65      & subtraction
  1277 \end{constants}
  1278 \subcaption{Constants and infixes}
  1279 
  1280 \begin{ttbox}\makeatother
  1281 \tdx{nat_induct}     [| P 0; !!n. P n ==> P(Suc n) |]  ==> P n
  1282 
  1283 \tdx{Suc_not_Zero}   Suc m ~= 0
  1284 \tdx{inj_Suc}        inj Suc
  1285 \tdx{n_not_Suc_n}    n~=Suc n
  1286 \subcaption{Basic properties}
  1287 \end{ttbox}
  1288 \caption{The type of natural numbers, \tydx{nat}} \label{hol-nat1}
  1289 \end{figure}
  1290 
  1291 
  1292 \begin{figure}
  1293 \begin{ttbox}\makeatother
  1294               0+n           = n
  1295               (Suc m)+n     = Suc(m+n)
  1296 
  1297               m-0           = m
  1298               0-n           = n
  1299               Suc(m)-Suc(n) = m-n
  1300 
  1301               0*n           = 0
  1302               Suc(m)*n      = n + m*n
  1303 
  1304 \tdx{mod_less}      m<n ==> m mod n = m
  1305 \tdx{mod_geq}       [| 0<n;  ~m<n |] ==> m mod n = (m-n) mod n
  1306 
  1307 \tdx{div_less}      m<n ==> m div n = 0
  1308 \tdx{div_geq}       [| 0<n;  ~m<n |] ==> m div n = Suc((m-n) div n)
  1309 \end{ttbox}
  1310 \caption{Recursion equations for the arithmetic operators} \label{hol-nat2}
  1311 \end{figure}
  1312 
  1313 \subsection{The type of natural numbers, \textit{nat}}
  1314 \index{nat@{\textit{nat}} type|(}
  1315 
  1316 The theory \thydx{NatDef} defines the natural numbers in a roundabout but
  1317 traditional way.  The axiom of infinity postulates a type~\tydx{ind} of
  1318 individuals, which is non-empty and closed under an injective operation.  The
  1319 natural numbers are inductively generated by choosing an arbitrary individual
  1320 for~0 and using the injective operation to take successors.  This is a least
  1321 fixedpoint construction.  For details see the file \texttt{NatDef.thy}.
  1322 
  1323 Type~\tydx{nat} is an instance of class~\cldx{ord}, which makes the
  1324 overloaded functions of this class (esp.\ \cdx{<} and \cdx{<=}, but also
  1325 \cdx{min}, \cdx{max} and \cdx{LEAST}) available on \tydx{nat}.  Theory
  1326 \thydx{Nat} builds on \texttt{NatDef} and shows that {\tt<=} is a partial order,
  1327 so \tydx{nat} is also an instance of class \cldx{order}.
  1328 
  1329 Theory \thydx{Arith} develops arithmetic on the natural numbers.  It defines
  1330 addition, multiplication and subtraction.  Theory \thydx{Divides} defines
  1331 division, remainder and the ``divides'' relation.  The numerous theorems
  1332 proved include commutative, associative, distributive, identity and
  1333 cancellation laws.  See Figs.\ts\ref{hol-nat1} and~\ref{hol-nat2}.  The
  1334 recursion equations for the operators \texttt{+}, \texttt{-} and \texttt{*} on
  1335 \texttt{nat} are part of the default simpset.
  1336 
  1337 Functions on \tydx{nat} can be defined by primitive or well-founded recursion;
  1338 see \S\ref{sec:HOL:recursive}.  A simple example is addition.
  1339 Here, \texttt{op +} is the name of the infix operator~\texttt{+}, following
  1340 the standard convention.
  1341 \begin{ttbox}
  1342 \sdx{primrec}
  1343       "0 + n = n"
  1344   "Suc m + n = Suc (m + n)"
  1345 \end{ttbox}
  1346 There is also a \sdx{case}-construct
  1347 of the form
  1348 \begin{ttbox}
  1349 case \(e\) of 0 => \(a\) | Suc \(m\) => \(b\)
  1350 \end{ttbox}
  1351 Note that Isabelle insists on precisely this format; you may not even change
  1352 the order of the two cases.
  1353 Both \texttt{primrec} and \texttt{case} are realized by a recursion operator
  1354 \cdx{nat_rec}, which is available because \textit{nat} is represented as
  1355 a datatype (see \S\ref{subsec:datatype:rep_datatype}).
  1356 
  1357 %The predecessor relation, \cdx{pred_nat}, is shown to be well-founded.
  1358 %Recursion along this relation resembles primitive recursion, but is
  1359 %stronger because we are in higher-order logic; using primitive recursion to
  1360 %define a higher-order function, we can easily Ackermann's function, which
  1361 %is not primitive recursive \cite[page~104]{thompson91}.
  1362 %The transitive closure of \cdx{pred_nat} is~$<$.  Many functions on the
  1363 %natural numbers are most easily expressed using recursion along~$<$.
  1364 
  1365 Tactic {\tt\ttindex{induct_tac} "$n$" $i$} performs induction on variable~$n$
  1366 in subgoal~$i$ using theorem \texttt{nat_induct}.  There is also the derived
  1367 theorem \tdx{less_induct}:
  1368 \begin{ttbox}
  1369 [| !!n. [| ! m. m<n --> P m |] ==> P n |]  ==>  P n
  1370 \end{ttbox}
  1371 
  1372 
  1373 Reasoning about arithmetic inequalities can be tedious.  Fortunately HOL
  1374 provides a decision procedure for quantifier-free linear arithmetic (i.e.\ 
  1375 only addition and subtraction). The simplifier invokes a weak version of this
  1376 decision procedure automatically. If this is not sufficent, you can invoke
  1377 the full procedure \ttindex{arith_tac} explicitly.  It copes with arbitrary
  1378 formulae involving {\tt=}, {\tt<}, {\tt<=}, {\tt+}, {\tt-}, {\tt Suc}, {\tt
  1379   min}, {\tt max} and numerical constants; other subterms are treated as
  1380 atomic; subformulae not involving type $nat$ are ignored; quantified
  1381 subformulae are ignored unless they are positive universal or negative
  1382 existential. Note that the running time is exponential in the number of
  1383 occurrences of {\tt min}, {\tt max}, and {\tt-} because they require case
  1384 distinctions. Note also that \texttt{arith_tac} is not complete: if
  1385 divisibility plays a role, it may fail to prove a valid formula, for example
  1386 $m+m \neq n+n+1$. Fortunately such examples are rare in practice.
  1387 
  1388 If \texttt{arith_tac} fails you, try to find relevant arithmetic results in
  1389 the library.  The theory \texttt{NatDef} contains theorems about {\tt<} and
  1390 {\tt<=}, the theory \texttt{Arith} contains theorems about \texttt{+},
  1391 \texttt{-} and \texttt{*}, and theory \texttt{Divides} contains theorems about
  1392 \texttt{div} and \texttt{mod}.  Use the \texttt{find}-functions to locate them
  1393 (see the {\em Reference Manual\/}).
  1394 
  1395 \begin{figure}
  1396 \index{#@{\tt[]} symbol}
  1397 \index{#@{\tt\#} symbol}
  1398 \index{"@@{\tt\at} symbol}
  1399 \index{*"! symbol}
  1400 \begin{constants}
  1401   \it symbol & \it meta-type & \it priority & \it description \\
  1402   \tt[]    & $\alpha\,list$ & & empty list\\
  1403   \tt \#   & $[\alpha,\alpha\,list]\To \alpha\,list$ & Right 65 & 
  1404         list constructor \\
  1405   \cdx{null}    & $\alpha\,list \To bool$ & & emptiness test\\
  1406   \cdx{hd}      & $\alpha\,list \To \alpha$ & & head \\
  1407   \cdx{tl}      & $\alpha\,list \To \alpha\,list$ & & tail \\
  1408   \cdx{last}    & $\alpha\,list \To \alpha$ & & last element \\
  1409   \cdx{butlast} & $\alpha\,list \To \alpha\,list$ & & drop last element \\
  1410   \tt\at  & $[\alpha\,list,\alpha\,list]\To \alpha\,list$ & Left 65 & append \\
  1411   \cdx{map}     & $(\alpha\To\beta) \To (\alpha\,list \To \beta\,list)$
  1412         & & apply to all\\
  1413   \cdx{filter}  & $(\alpha \To bool) \To (\alpha\,list \To \alpha\,list)$
  1414         & & filter functional\\
  1415   \cdx{set}& $\alpha\,list \To \alpha\,set$ & & elements\\
  1416   \sdx{mem}  & $\alpha \To \alpha\,list \To bool$  &  Left 55   & membership\\
  1417   \cdx{foldl}   & $(\beta\To\alpha\To\beta) \To \beta \To \alpha\,list \To \beta$ &
  1418   & iteration \\
  1419   \cdx{concat}   & $(\alpha\,list)list\To \alpha\,list$ & & concatenation \\
  1420   \cdx{rev}     & $\alpha\,list \To \alpha\,list$ & & reverse \\
  1421   \cdx{length}  & $\alpha\,list \To nat$ & & length \\
  1422   \tt! & $\alpha\,list \To nat \To \alpha$ & Left 100 & indexing \\
  1423   \cdx{take}, \cdx{drop} & $nat \To \alpha\,list \To \alpha\,list$ &&
  1424     take or drop a prefix \\
  1425   \cdx{takeWhile},\\
  1426   \cdx{dropWhile} &
  1427     $(\alpha \To bool) \To \alpha\,list \To \alpha\,list$ &&
  1428     take or drop a prefix
  1429 \end{constants}
  1430 \subcaption{Constants and infixes}
  1431 
  1432 \begin{center} \tt\frenchspacing
  1433 \begin{tabular}{rrr} 
  1434   \it external        & \it internal  & \it description \\{}
  1435   [$x@1$, $\dots$, $x@n$]  &  $x@1$ \# $\cdots$ \# $x@n$ \# [] &
  1436         \rm finite list \\{}
  1437   [$x$:$l$. $P$]  & filter ($\lambda x{.}P$) $l$ & 
  1438         \rm list comprehension
  1439 \end{tabular}
  1440 \end{center}
  1441 \subcaption{Translations}
  1442 \caption{The theory \thydx{List}} \label{hol-list}
  1443 \end{figure}
  1444 
  1445 
  1446 \begin{figure}
  1447 \begin{ttbox}\makeatother
  1448 null [] = True
  1449 null (x#xs) = False
  1450 
  1451 hd (x#xs) = x
  1452 tl (x#xs) = xs
  1453 tl [] = []
  1454 
  1455 [] @ ys = ys
  1456 (x#xs) @ ys = x # xs @ ys
  1457 
  1458 map f [] = []
  1459 map f (x#xs) = f x # map f xs
  1460 
  1461 filter P [] = []
  1462 filter P (x#xs) = (if P x then x#filter P xs else filter P xs)
  1463 
  1464 set [] = \ttlbrace\ttrbrace
  1465 set (x#xs) = insert x (set xs)
  1466 
  1467 x mem [] = False
  1468 x mem (y#ys) = (if y=x then True else x mem ys)
  1469 
  1470 foldl f a [] = a
  1471 foldl f a (x#xs) = foldl f (f a x) xs
  1472 
  1473 concat([]) = []
  1474 concat(x#xs) = x @ concat(xs)
  1475 
  1476 rev([]) = []
  1477 rev(x#xs) = rev(xs) @ [x]
  1478 
  1479 length([]) = 0
  1480 length(x#xs) = Suc(length(xs))
  1481 
  1482 xs!0 = hd xs
  1483 xs!(Suc n) = (tl xs)!n
  1484 
  1485 take n [] = []
  1486 take n (x#xs) = (case n of 0 => [] | Suc(m) => x # take m xs)
  1487 
  1488 drop n [] = []
  1489 drop n (x#xs) = (case n of 0 => x#xs | Suc(m) => drop m xs)
  1490 
  1491 takeWhile P [] = []
  1492 takeWhile P (x#xs) = (if P x then x#takeWhile P xs else [])
  1493 
  1494 dropWhile P [] = []
  1495 dropWhile P (x#xs) = (if P x then dropWhile P xs else xs)
  1496 \end{ttbox}
  1497 \caption{Recursions equations for list processing functions}
  1498 \label{fig:HOL:list-simps}
  1499 \end{figure}
  1500 \index{nat@{\textit{nat}} type|)}
  1501 
  1502 
  1503 \subsection{The type constructor for lists, \textit{list}}
  1504 \label{subsec:list}
  1505 \index{list@{\textit{list}} type|(}
  1506 
  1507 Figure~\ref{hol-list} presents the theory \thydx{List}: the basic list
  1508 operations with their types and syntax.  Type $\alpha \; list$ is
  1509 defined as a \texttt{datatype} with the constructors {\tt[]} and {\tt\#}.
  1510 As a result the generic structural induction and case analysis tactics
  1511 \texttt{induct\_tac} and \texttt{exhaust\_tac} also become available for
  1512 lists.  A \sdx{case} construct of the form
  1513 \begin{center}\tt
  1514 case $e$ of [] => $a$  |  \(x\)\#\(xs\) => b
  1515 \end{center}
  1516 is defined by translation.  For details see~\S\ref{sec:HOL:datatype}. There
  1517 is also a case splitting rule \tdx{split_list_case}
  1518 \[
  1519 \begin{array}{l}
  1520 P(\mathtt{case}~e~\mathtt{of}~\texttt{[] =>}~a ~\texttt{|}~
  1521                x\texttt{\#}xs~\texttt{=>}~f~x~xs) ~= \\
  1522 ((e = \texttt{[]} \to P(a)) \land
  1523  (\forall x~ xs. e = x\texttt{\#}xs \to P(f~x~xs)))
  1524 \end{array}
  1525 \]
  1526 which can be fed to \ttindex{addsplits} just like
  1527 \texttt{split_if} (see~\S\ref{subsec:HOL:case:splitting}).
  1528 
  1529 \texttt{List} provides a basic library of list processing functions defined by
  1530 primitive recursion (see~\S\ref{sec:HOL:primrec}).  The recursion equations
  1531 are shown in Fig.\ts\ref{fig:HOL:list-simps}.
  1532 
  1533 \index{list@{\textit{list}} type|)}
  1534 
  1535 
  1536 \subsection{Introducing new types} \label{sec:typedef}
  1537 
  1538 The \HOL-methodology dictates that all extensions to a theory should
  1539 be \textbf{definitional}.  The type definition mechanism that
  1540 meets this criterion is \ttindex{typedef}.  Note that \emph{type synonyms},
  1541 which are inherited from {\Pure} and described elsewhere, are just
  1542 syntactic abbreviations that have no logical meaning.
  1543 
  1544 \begin{warn}
  1545   Types in \HOL\ must be non-empty; otherwise the quantifier rules would be
  1546   unsound, because $\exists x. x=x$ is a theorem \cite[\S7]{paulson-COLOG}.
  1547 \end{warn}
  1548 A \bfindex{type definition} identifies the new type with a subset of
  1549 an existing type.  More precisely, the new type is defined by
  1550 exhibiting an existing type~$\tau$, a set~$A::\tau\,set$, and a
  1551 theorem of the form $x:A$.  Thus~$A$ is a non-empty subset of~$\tau$,
  1552 and the new type denotes this subset.  New functions are defined that
  1553 establish an isomorphism between the new type and the subset.  If
  1554 type~$\tau$ involves type variables $\alpha@1$, \ldots, $\alpha@n$,
  1555 then the type definition creates a type constructor
  1556 $(\alpha@1,\ldots,\alpha@n)ty$ rather than a particular type.
  1557 
  1558 \begin{figure}[htbp]
  1559 \begin{rail}
  1560 typedef  : 'typedef' ( () | '(' name ')') type '=' set witness;
  1561 
  1562 type    : typevarlist name ( () | '(' infix ')' );
  1563 set     : string;
  1564 witness : () | '(' id ')';
  1565 \end{rail}
  1566 \caption{Syntax of type definitions}
  1567 \label{fig:HOL:typedef}
  1568 \end{figure}
  1569 
  1570 The syntax for type definitions is shown in Fig.~\ref{fig:HOL:typedef}.  For
  1571 the definition of `typevarlist' and `infix' see
  1572 \iflabelundefined{chap:classical}
  1573 {the appendix of the {\em Reference Manual\/}}%
  1574 {Appendix~\ref{app:TheorySyntax}}.  The remaining nonterminals have the
  1575 following meaning:
  1576 \begin{description}
  1577 \item[\it type:] the new type constructor $(\alpha@1,\dots,\alpha@n)ty$ with
  1578   optional infix annotation.
  1579 \item[\it name:] an alphanumeric name $T$ for the type constructor
  1580   $ty$, in case $ty$ is a symbolic name.  Defaults to $ty$.
  1581 \item[\it set:] the representing subset $A$.
  1582 \item[\it witness:] name of a theorem of the form $a:A$ proving
  1583   non-emptiness.  It can be omitted in case Isabelle manages to prove
  1584   non-emptiness automatically.
  1585 \end{description}
  1586 If all context conditions are met (no duplicate type variables in
  1587 `typevarlist', no extra type variables in `set', and no free term variables
  1588 in `set'), the following components are added to the theory:
  1589 \begin{itemize}
  1590 \item a type $ty :: (term,\dots,term)term$
  1591 \item constants
  1592 \begin{eqnarray*}
  1593 T &::& \tau\;set \\
  1594 Rep_T &::& (\alpha@1,\dots,\alpha@n)ty \To \tau \\
  1595 Abs_T &::& \tau \To (\alpha@1,\dots,\alpha@n)ty
  1596 \end{eqnarray*}
  1597 \item a definition and three axioms
  1598 \[
  1599 \begin{array}{ll}
  1600 T{\tt_def} & T \equiv A \\
  1601 {\tt Rep_}T & Rep_T\,x \in T \\
  1602 {\tt Rep_}T{\tt_inverse} & Abs_T\,(Rep_T\,x) = x \\
  1603 {\tt Abs_}T{\tt_inverse} & y \in T \Imp Rep_T\,(Abs_T\,y) = y
  1604 \end{array}
  1605 \]
  1606 stating that $(\alpha@1,\dots,\alpha@n)ty$ is isomorphic to $A$ by $Rep_T$
  1607 and its inverse $Abs_T$.
  1608 \end{itemize}
  1609 Below are two simple examples of \HOL\ type definitions.  Non-emptiness
  1610 is proved automatically here.
  1611 \begin{ttbox}
  1612 typedef unit = "{\ttlbrace}True{\ttrbrace}"
  1613 
  1614 typedef (prod)
  1615   ('a, 'b) "*"    (infixr 20)
  1616       = "{\ttlbrace}f . EX (a::'a) (b::'b). f = (\%x y. x = a & y = b){\ttrbrace}"
  1617 \end{ttbox}
  1618 
  1619 Type definitions permit the introduction of abstract data types in a safe
  1620 way, namely by providing models based on already existing types.  Given some
  1621 abstract axiomatic description $P$ of a type, this involves two steps:
  1622 \begin{enumerate}
  1623 \item Find an appropriate type $\tau$ and subset $A$ which has the desired
  1624   properties $P$, and make a type definition based on this representation.
  1625 \item Prove that $P$ holds for $ty$ by lifting $P$ from the representation.
  1626 \end{enumerate}
  1627 You can now forget about the representation and work solely in terms of the
  1628 abstract properties $P$.
  1629 
  1630 \begin{warn}
  1631 If you introduce a new type (constructor) $ty$ axiomatically, i.e.\ by
  1632 declaring the type and its operations and by stating the desired axioms, you
  1633 should make sure the type has a non-empty model.  You must also have a clause
  1634 \par
  1635 \begin{ttbox}
  1636 arities \(ty\) :: (term,\thinspace\(\dots\),{\thinspace}term){\thinspace}term
  1637 \end{ttbox}
  1638 in your theory file to tell Isabelle that $ty$ is in class \texttt{term}, the
  1639 class of all \HOL\ types.
  1640 \end{warn}
  1641 
  1642 
  1643 \section{Records}
  1644 
  1645 At a first approximation, records are just a minor generalisation of tuples,
  1646 where components may be addressed by labels instead of just position (think of
  1647 {\ML}, for example).  The version of records offered by Isabelle/HOL is
  1648 slightly more advanced, though, supporting \emph{extensible record schemes}.
  1649 This admits operations that are polymorphic with respect to record extension,
  1650 yielding ``object-oriented'' effects like (single) inheritance.  See also
  1651 \cite{NaraschewskiW-TPHOLs98} for more details on object-oriented
  1652 verification and record subtyping in HOL.
  1653 
  1654 
  1655 \subsection{Basics}
  1656 
  1657 Isabelle/HOL supports fixed and schematic records both at the level of terms
  1658 and types.  The concrete syntax is as follows:
  1659 
  1660 \begin{center}
  1661 \begin{tabular}{l|l|l}
  1662   & record terms & record types \\ \hline
  1663   fixed & $\record{x = a\fs y = b}$ & $\record{x \ty A\fs y \ty B}$ \\
  1664   schematic & $\record{x = a\fs y = b\fs \more = m}$ &
  1665     $\record{x \ty A\fs y \ty B\fs \more \ty M}$ \\
  1666 \end{tabular}
  1667 \end{center}
  1668 
  1669 \noindent The \textsc{ascii} representation of $\record{x = a}$ is \texttt{(| x = a |)}.
  1670 
  1671 A fixed record $\record{x = a\fs y = b}$ has field $x$ of value $a$ and field
  1672 $y$ of value $b$.  The corresponding type is $\record{x \ty A\fs y \ty B}$,
  1673 assuming that $a \ty A$ and $b \ty B$.
  1674 
  1675 A record scheme like $\record{x = a\fs y = b\fs \more = m}$ contains fields
  1676 $x$ and $y$ as before, but also possibly further fields as indicated by the
  1677 ``$\more$'' notation (which is actually part of the syntax).  The improper
  1678 field ``$\more$'' of a record scheme is called the \emph{more part}.
  1679 Logically it is just a free variable, which is occasionally referred to as
  1680 \emph{row variable} in the literature.  The more part of a record scheme may
  1681 be instantiated by zero or more further components.  For example, above scheme
  1682 might get instantiated to $\record{x = a\fs y = b\fs z = c\fs \more = m'}$,
  1683 where $m'$ refers to a different more part.  Fixed records are special
  1684 instances of record schemes, where ``$\more$'' is properly terminated by the
  1685 $() :: unit$ element.  Actually, $\record{x = a\fs y = b}$ is just an
  1686 abbreviation for $\record{x = a\fs y = b\fs \more = ()}$.
  1687 
  1688 \medskip
  1689 
  1690 There are two key features that make extensible records in a simply typed
  1691 language like HOL feasible:
  1692 \begin{enumerate}
  1693 \item the more part is internalised, as a free term or type variable,
  1694 \item field names are externalised, they cannot be accessed within the logic
  1695   as first-class values.
  1696 \end{enumerate}
  1697 
  1698 \medskip
  1699 
  1700 In Isabelle/HOL record types have to be defined explicitly, fixing their field
  1701 names and types, and their (optional) parent record (see
  1702 \S\ref{sec:HOL:record-def}).  Afterwards, records may be formed using above
  1703 syntax, while obeying the canonical order of fields as given by their
  1704 declaration.  The record package also provides several operations like
  1705 selectors and updates (see \S\ref{sec:HOL:record-ops}), together with
  1706 characteristic properties (see \S\ref{sec:HOL:record-thms}).
  1707 
  1708 There is an example theory demonstrating most basic aspects of extensible
  1709 records (see theory \texttt{HOL/ex/Points} in the Isabelle sources).
  1710 
  1711 
  1712 \subsection{Defining records}\label{sec:HOL:record-def}
  1713 
  1714 The theory syntax for record type definitions is shown in
  1715 Fig.~\ref{fig:HOL:record}.  For the definition of `typevarlist' and `type' see
  1716 \iflabelundefined{chap:classical}
  1717 {the appendix of the {\em Reference Manual\/}}%
  1718 {Appendix~\ref{app:TheorySyntax}}.
  1719 
  1720 \begin{figure}[htbp]
  1721 \begin{rail}
  1722 record  : 'record' typevarlist name '=' parent (field +);
  1723 
  1724 parent  : ( () | type '+');
  1725 field   : name '::' type;
  1726 \end{rail}
  1727 \caption{Syntax of record type definitions}
  1728 \label{fig:HOL:record}
  1729 \end{figure}
  1730 
  1731 A general \ttindex{record} specification is of the following form:
  1732 \[
  1733 \mathtt{record}~(\alpha@1, \dots, \alpha@n) \, t ~=~
  1734   (\tau@1, \dots, \tau@m) \, s ~+~ c@1 :: \sigma@1 ~ \dots ~ c@l :: \sigma@l
  1735 \]
  1736 where $\vec\alpha@n$ are distinct type variables, and $\vec\tau@m$,
  1737 $\vec\sigma@l$ are types containing at most variables from $\vec\alpha@n$.
  1738 Type constructor $t$ has to be new, while $s$ has to specify an existing
  1739 record type.  Furthermore, the $\vec c@l$ have to be distinct field names.
  1740 There has to be at least one field.
  1741 
  1742 In principle, field names may never be shared with other records.  This is no
  1743 actual restriction in practice, since $\vec c@l$ are internally declared
  1744 within a separate name space qualified by the name $t$ of the record.
  1745 
  1746 \medskip
  1747 
  1748 Above definition introduces a new record type $(\vec\alpha@n) \, t$ by
  1749 extending an existing one $(\vec\tau@m) \, s$ by new fields $\vec c@l \ty
  1750 \vec\sigma@l$.  The parent record specification is optional, by omitting it
  1751 $t$ becomes a \emph{root record}.  The hierarchy of all records declared
  1752 within a theory forms a forest structure, i.e.\ a set of trees, where any of
  1753 these is rooted by some root record.
  1754 
  1755 For convenience, $(\vec\alpha@n) \, t$ is made a type abbreviation for the
  1756 fixed record type $\record{\vec c@l \ty \vec\sigma@l}$, and $(\vec\alpha@n,
  1757 \zeta) \, t_scheme$ is made an abbreviation for $\record{\vec c@l \ty
  1758   \vec\sigma@l\fs \more \ty \zeta}$.
  1759 
  1760 \medskip
  1761 
  1762 The following simple example defines a root record type $point$ with fields $x
  1763 \ty nat$ and $y \ty nat$, and record type $cpoint$ by extending $point$ with
  1764 an additional $colour$ component.
  1765 
  1766 \begin{ttbox}
  1767   record point =
  1768     x :: nat
  1769     y :: nat
  1770 
  1771   record cpoint = point +
  1772     colour :: string
  1773 \end{ttbox}
  1774 
  1775 
  1776 \subsection{Record operations}\label{sec:HOL:record-ops}
  1777 
  1778 Any record definition of the form presented above produces certain standard
  1779 operations.  Selectors and updates are provided for any field, including the
  1780 improper one ``$more$''.  There are also cumulative record constructor
  1781 functions.
  1782 
  1783 To simplify the presentation below, we first assume that $(\vec\alpha@n) \, t$
  1784 is a root record with fields $\vec c@l \ty \vec\sigma@l$.
  1785 
  1786 \medskip
  1787 
  1788 \textbf{Selectors} and \textbf{updates} are available for any field (including
  1789 ``$more$'') as follows:
  1790 \begin{matharray}{lll}
  1791   c@i & \ty & \record{\vec c@l \ty \vec \sigma@l, \more \ty \zeta} \To \sigma@i \\
  1792   c@i_update & \ty & \sigma@i \To \record{\vec c@l \ty \vec \sigma@l, \more \ty \zeta} \To
  1793     \record{\vec c@l \ty \vec \sigma@l, \more \ty \zeta}
  1794 \end{matharray}
  1795 
  1796 There is some special syntax for updates: $r \, \record{x \asn a}$ abbreviates
  1797 term $x_update \, a \, r$.  Repeated updates are also supported: $r \,
  1798 \record{x \asn a} \, \record{y \asn b} \, \record{z \asn c}$ may be written as
  1799 $r \, \record{x \asn a\fs y \asn b\fs z \asn c}$.  Note that because of
  1800 postfix notation the order of fields shown here is reverse than in the actual
  1801 term.  This might lead to confusion in conjunction with proof tools like
  1802 ordered rewriting.
  1803 
  1804 Since repeated updates are just function applications, fields may be freely
  1805 permuted in $\record{x \asn a\fs y \asn b\fs z \asn c}$, as far as the logic
  1806 is concerned.  Thus commutativity of updates can be proven within the logic
  1807 for any two fields, but not as a general theorem: fields are not first-class
  1808 values.
  1809 
  1810 \medskip
  1811 
  1812 \textbf{Make} operations provide cumulative record constructor functions:
  1813 \begin{matharray}{lll}
  1814   make & \ty & \vec\sigma@l \To \record{\vec c@l \ty \vec \sigma@l} \\
  1815   make_scheme & \ty & \vec\sigma@l \To
  1816     \zeta \To \record{\vec c@l \ty \vec \sigma@l, \more \ty \zeta} \\
  1817 \end{matharray}
  1818 \noindent
  1819 These functions are curried.  The corresponding definitions in terms of actual
  1820 record terms are part of the standard simpset.  Thus $point\dtt make\,a\,b$
  1821 rewrites to $\record{x = a\fs y = b}$.
  1822 
  1823 \medskip
  1824 
  1825 Any of above selector, update and make operations are declared within a local
  1826 name space prefixed by the name $t$ of the record.  In case that different
  1827 records share base names of fields, one has to qualify names explicitly (e.g.\ 
  1828 $t\dtt c@i_update$).  This is recommended especially for operations like
  1829 $make$ or $update_more$ that always have the same base name.  Just use $t\dtt
  1830 make$ etc.\ to avoid confusion.
  1831 
  1832 \bigskip
  1833 
  1834 We reconsider the case of non-root records, which are derived of some parent
  1835 record.  In general, the latter may depend on another parent as well,
  1836 resulting in a list of \emph{ancestor records}.  Appending the lists of fields
  1837 of all ancestors results in a certain field prefix.  The record package
  1838 automatically takes care of this by lifting operations over this context of
  1839 ancestor fields.  Assuming that $(\vec\alpha@n) \, t$ has ancestor fields
  1840 $\vec d@k \ty \vec\rho@k$, selectors will get the following types:
  1841 \begin{matharray}{lll}
  1842   c@i & \ty & \record{\vec d@k \ty \vec\rho@k, \vec c@l \ty \vec \sigma@l, \more \ty \zeta}
  1843     \To \sigma@i
  1844 \end{matharray}
  1845 \noindent
  1846 Update and make operations are analogous.
  1847 
  1848 
  1849 \subsection{Proof tools}\label{sec:HOL:record-thms}
  1850 
  1851 The record package provides the following proof rules for any record type $t$.
  1852 \begin{enumerate}
  1853   
  1854 \item Standard conversions (selectors or updates applied to record constructor
  1855   terms, make function definitions) are part of the standard simpset (via
  1856   \texttt{addsimps}).
  1857   
  1858 \item Selectors applied to updated records are automatically reduced by
  1859   simplification procedure \ttindex{record_simproc}, which is part of the
  1860   default simpset.
  1861   
  1862 \item Inject equations of a form analogous to $((x, y) = (x', y')) \equiv x=x'
  1863   \conj y=y'$ are made part of the standard simpset and claset (via
  1864   \texttt{addIffs}).
  1865   
  1866 \item A tactic for record field splitting (\ttindex{record_split_tac}) may be
  1867   made part of the claset (via \texttt{addSWrapper}).  This tactic is based on
  1868   rules analogous to $(\All x PROP~P~x) \equiv (\All{a~b} PROP~P(a, b))$ for
  1869   any field.
  1870 \end{enumerate}
  1871 
  1872 The first two kinds of rules are stored within the theory as $t\dtt simps$ and
  1873 $t\dtt iffs$, respectively.  In some situations it might be appropriate to
  1874 expand the definitions of updates: $t\dtt update_defs$.  Note that these names
  1875 are \emph{not} bound at the {\ML} level.
  1876 
  1877 \medskip
  1878 
  1879 The example theory \texttt{HOL/ex/Points} demonstrates typical proofs
  1880 concerning records.  The basic idea is to make \ttindex{record_split_tac}
  1881 expand quantified record variables and then simplify by the conversion rules.
  1882 By using a combination of the simplifier and classical prover together with
  1883 the default simpset and claset, record problems should be solved with a single
  1884 stroke of \texttt{Auto_tac} or \texttt{Force_tac}.  Most of the time, plain
  1885 \texttt{Simp_tac} should be sufficient, though.
  1886 
  1887 
  1888 \section{Datatype definitions}
  1889 \label{sec:HOL:datatype}
  1890 \index{*datatype|(}
  1891 
  1892 Inductive datatypes, similar to those of \ML, frequently appear in
  1893 applications of Isabelle/HOL.  In principle, such types could be defined by
  1894 hand via \texttt{typedef} (see \S\ref{sec:typedef}), but this would be far too
  1895 tedious.  The \ttindex{datatype} definition package of Isabelle/HOL (cf.\ 
  1896 \cite{Berghofer-Wenzel:1999:TPHOL}) automates such chores.  It generates an
  1897 appropriate \texttt{typedef} based on a least fixed-point construction, and
  1898 proves freeness theorems and induction rules, as well as theorems for
  1899 recursion and case combinators.  The user just has to give a simple
  1900 specification of new inductive types using a notation similar to {\ML} or
  1901 Haskell.
  1902 
  1903 The current datatype package can handle both mutual and indirect recursion.
  1904 It also offers to represent existing types as datatypes giving the advantage
  1905 of a more uniform view on standard theories.
  1906 
  1907 
  1908 \subsection{Basics}
  1909 \label{subsec:datatype:basics}
  1910 
  1911 A general \texttt{datatype} definition is of the following form:
  1912 \[
  1913 \begin{array}{llcl}
  1914 \mathtt{datatype} & (\alpha@1,\ldots,\alpha@h)t@1 & = &
  1915   C^1@1~\tau^1@{1,1}~\ldots~\tau^1@{1,m^1@1} ~\mid~ \ldots ~\mid~
  1916     C^1@{k@1}~\tau^1@{k@1,1}~\ldots~\tau^1@{k@1,m^1@{k@1}} \\
  1917  & & \vdots \\
  1918 \mathtt{and} & (\alpha@1,\ldots,\alpha@h)t@n & = &
  1919   C^n@1~\tau^n@{1,1}~\ldots~\tau^n@{1,m^n@1} ~\mid~ \ldots ~\mid~
  1920     C^n@{k@n}~\tau^n@{k@n,1}~\ldots~\tau^n@{k@n,m^n@{k@n}}
  1921 \end{array}
  1922 \]
  1923 where $\alpha@i$ are type variables, $C^j@i$ are distinct constructor
  1924 names and $\tau^j@{i,i'}$ are {\em admissible} types containing at
  1925 most the type variables $\alpha@1, \ldots, \alpha@h$. A type $\tau$
  1926 occurring in a \texttt{datatype} definition is {\em admissible} iff
  1927 \begin{itemize}
  1928 \item $\tau$ is non-recursive, i.e.\ $\tau$ does not contain any of the
  1929 newly defined type constructors $t@1,\ldots,t@n$, or
  1930 \item $\tau = (\alpha@1,\ldots,\alpha@h)t@{j'}$ where $1 \leq j' \leq n$, or
  1931 \item $\tau = (\tau'@1,\ldots,\tau'@{h'})t'$, where $t'$ is
  1932 the type constructor of an already existing datatype and $\tau'@1,\ldots,\tau'@{h'}$
  1933 are admissible types.
  1934 \item $\tau = \sigma \rightarrow \tau'$, where $\tau'$ is an admissible
  1935 type and $\sigma$ is non-recursive (i.e. the occurrences of the newly defined
  1936 types are {\em strictly positive})
  1937 \end{itemize}
  1938 If some $(\alpha@1,\ldots,\alpha@h)t@{j'}$ occurs in a type $\tau^j@{i,i'}$
  1939 of the form
  1940 \[
  1941 (\ldots,\ldots ~ (\alpha@1,\ldots,\alpha@h)t@{j'} ~ \ldots,\ldots)t'
  1942 \]
  1943 this is called a {\em nested} (or \emph{indirect}) occurrence. A very simple
  1944 example of a datatype is the type \texttt{list}, which can be defined by
  1945 \begin{ttbox}
  1946 datatype 'a list = Nil
  1947                  | Cons 'a ('a list)
  1948 \end{ttbox}
  1949 Arithmetic expressions \texttt{aexp} and boolean expressions \texttt{bexp} can be modelled
  1950 by the mutually recursive datatype definition
  1951 \begin{ttbox}
  1952 datatype 'a aexp = If_then_else ('a bexp) ('a aexp) ('a aexp)
  1953                  | Sum ('a aexp) ('a aexp)
  1954                  | Diff ('a aexp) ('a aexp)
  1955                  | Var 'a
  1956                  | Num nat
  1957 and      'a bexp = Less ('a aexp) ('a aexp)
  1958                  | And ('a bexp) ('a bexp)
  1959                  | Or ('a bexp) ('a bexp)
  1960 \end{ttbox}
  1961 The datatype \texttt{term}, which is defined by
  1962 \begin{ttbox}
  1963 datatype ('a, 'b) term = Var 'a
  1964                        | App 'b ((('a, 'b) term) list)
  1965 \end{ttbox}
  1966 is an example for a datatype with nested recursion. Using nested recursion
  1967 involving function spaces, we may also define infinitely branching datatypes, e.g.
  1968 \begin{ttbox}
  1969 datatype 'a tree = Atom 'a | Branch "nat => 'a tree"
  1970 \end{ttbox}
  1971 
  1972 \medskip
  1973 
  1974 Types in HOL must be non-empty. Each of the new datatypes
  1975 $(\alpha@1,\ldots,\alpha@h)t@j$ with $1 \le j \le n$ is non-empty iff it has a
  1976 constructor $C^j@i$ with the following property: for all argument types
  1977 $\tau^j@{i,i'}$ of the form $(\alpha@1,\ldots,\alpha@h)t@{j'}$ the datatype
  1978 $(\alpha@1,\ldots,\alpha@h)t@{j'}$ is non-empty.
  1979 
  1980 If there are no nested occurrences of the newly defined datatypes, obviously
  1981 at least one of the newly defined datatypes $(\alpha@1,\ldots,\alpha@h)t@j$
  1982 must have a constructor $C^j@i$ without recursive arguments, a \emph{base
  1983   case}, to ensure that the new types are non-empty. If there are nested
  1984 occurrences, a datatype can even be non-empty without having a base case
  1985 itself. Since \texttt{list} is a non-empty datatype, \texttt{datatype t = C (t
  1986   list)} is non-empty as well.
  1987 
  1988 
  1989 \subsubsection{Freeness of the constructors}
  1990 
  1991 The datatype constructors are automatically defined as functions of their
  1992 respective type:
  1993 \[ C^j@i :: [\tau^j@{i,1},\dots,\tau^j@{i,m^j@i}] \To (\alpha@1,\dots,\alpha@h)t@j \]
  1994 These functions have certain {\em freeness} properties.  They construct
  1995 distinct values:
  1996 \[
  1997 C^j@i~x@1~\dots~x@{m^j@i} \neq C^j@{i'}~y@1~\dots~y@{m^j@{i'}} \qquad
  1998 \mbox{for all}~ i \neq i'.
  1999 \]
  2000 The constructor functions are injective:
  2001 \[
  2002 (C^j@i~x@1~\dots~x@{m^j@i} = C^j@i~y@1~\dots~y@{m^j@i}) =
  2003 (x@1 = y@1 \land \dots \land x@{m^j@i} = y@{m^j@i})
  2004 \]
  2005 Since the number of distinctness inequalities is quadratic in the number of
  2006 constructors, the datatype package avoids proving them separately if there are
  2007 too many constructors. Instead, specific inequalities are proved by a suitable
  2008 simplification procedure on demand.\footnote{This procedure, which is already part
  2009 of the default simpset, may be referred to by the ML identifier
  2010 \texttt{DatatypePackage.distinct_simproc}.}
  2011 
  2012 \subsubsection{Structural induction}
  2013 
  2014 The datatype package also provides structural induction rules.  For
  2015 datatypes without nested recursion, this is of the following form:
  2016 \[
  2017 \infer{P@1~x@1 \wedge \dots \wedge P@n~x@n}
  2018   {\begin{array}{lcl}
  2019      \Forall x@1 \dots x@{m^1@1}.
  2020        \List{P@{s^1@{1,1}}~x@{r^1@{1,1}}; \dots;
  2021          P@{s^1@{1,l^1@1}}~x@{r^1@{1,l^1@1}}} & \Imp &
  2022            P@1~\left(C^1@1~x@1~\dots~x@{m^1@1}\right) \\
  2023      & \vdots \\
  2024      \Forall x@1 \dots x@{m^1@{k@1}}.
  2025        \List{P@{s^1@{k@1,1}}~x@{r^1@{k@1,1}}; \dots;
  2026          P@{s^1@{k@1,l^1@{k@1}}}~x@{r^1@{k@1,l^1@{k@1}}}} & \Imp &
  2027            P@1~\left(C^1@{k@1}~x@1~\ldots~x@{m^1@{k@1}}\right) \\
  2028      & \vdots \\
  2029      \Forall x@1 \dots x@{m^n@1}.
  2030        \List{P@{s^n@{1,1}}~x@{r^n@{1,1}}; \dots;
  2031          P@{s^n@{1,l^n@1}}~x@{r^n@{1,l^n@1}}} & \Imp &
  2032            P@n~\left(C^n@1~x@1~\ldots~x@{m^n@1}\right) \\
  2033      & \vdots \\
  2034      \Forall x@1 \dots x@{m^n@{k@n}}.
  2035        \List{P@{s^n@{k@n,1}}~x@{r^n@{k@n,1}}; \dots
  2036          P@{s^n@{k@n,l^n@{k@n}}}~x@{r^n@{k@n,l^n@{k@n}}}} & \Imp &
  2037            P@n~\left(C^n@{k@n}~x@1~\ldots~x@{m^n@{k@n}}\right)
  2038    \end{array}}
  2039 \]
  2040 where
  2041 \[
  2042 \begin{array}{rcl}
  2043 Rec^j@i & := &
  2044    \left\{\left(r^j@{i,1},s^j@{i,1}\right),\ldots,
  2045      \left(r^j@{i,l^j@i},s^j@{i,l^j@i}\right)\right\} = \\[2ex]
  2046 && \left\{(i',i'')~\left|~
  2047      1\leq i' \leq m^j@i \wedge 1 \leq i'' \leq n \wedge
  2048        \tau^j@{i,i'} = (\alpha@1,\ldots,\alpha@h)t@{i''}\right.\right\}
  2049 \end{array}
  2050 \]
  2051 i.e.\ the properties $P@j$ can be assumed for all recursive arguments.
  2052 
  2053 For datatypes with nested recursion, such as the \texttt{term} example from
  2054 above, things are a bit more complicated.  Conceptually, Isabelle/HOL unfolds
  2055 a definition like
  2056 \begin{ttbox}
  2057 datatype ('a, 'b) term = Var 'a
  2058                        | App 'b ((('a, 'b) term) list)
  2059 \end{ttbox}
  2060 to an equivalent definition without nesting:
  2061 \begin{ttbox}
  2062 datatype ('a, 'b) term      = Var
  2063                             | App 'b (('a, 'b) term_list)
  2064 and      ('a, 'b) term_list = Nil'
  2065                             | Cons' (('a,'b) term) (('a,'b) term_list)
  2066 \end{ttbox}
  2067 Note however, that the type \texttt{('a,'b) term_list} and the constructors {\tt
  2068   Nil'} and \texttt{Cons'} are not really introduced.  One can directly work with
  2069 the original (isomorphic) type \texttt{(('a, 'b) term) list} and its existing
  2070 constructors \texttt{Nil} and \texttt{Cons}. Thus, the structural induction rule for
  2071 \texttt{term} gets the form
  2072 \[
  2073 \infer{P@1~x@1 \wedge P@2~x@2}
  2074   {\begin{array}{l}
  2075      \Forall x.~P@1~(\mathtt{Var}~x) \\
  2076      \Forall x@1~x@2.~P@2~x@2 \Imp P@1~(\mathtt{App}~x@1~x@2) \\
  2077      P@2~\mathtt{Nil} \\
  2078      \Forall x@1~x@2. \List{P@1~x@1; P@2~x@2} \Imp P@2~(\mathtt{Cons}~x@1~x@2)
  2079    \end{array}}
  2080 \]
  2081 Note that there are two predicates $P@1$ and $P@2$, one for the type \texttt{('a,'b) term}
  2082 and one for the type \texttt{(('a, 'b) term) list}.
  2083 
  2084 For a datatype with function types such as \texttt{'a tree}, the induction rule
  2085 is of the form
  2086 \[
  2087 \infer{P~t}
  2088   {\Forall a.~P~(\mathtt{Atom}~a) &
  2089    \Forall ts.~(\forall x.~P~(ts~x)) \Imp P~(\mathtt{Branch}~ts)}
  2090 \]
  2091 
  2092 \medskip In principle, inductive types are already fully determined by
  2093 freeness and structural induction.  For convenience in applications,
  2094 the following derived constructions are automatically provided for any
  2095 datatype.
  2096 
  2097 \subsubsection{The \sdx{case} construct}
  2098 
  2099 The type comes with an \ML-like \texttt{case}-construct:
  2100 \[
  2101 \begin{array}{rrcl}
  2102 \mbox{\tt case}~e~\mbox{\tt of} & C^j@1~x@{1,1}~\dots~x@{1,m^j@1} & \To & e@1 \\
  2103                            \vdots \\
  2104                            \mid & C^j@{k@j}~x@{k@j,1}~\dots~x@{k@j,m^j@{k@j}} & \To & e@{k@j}
  2105 \end{array}
  2106 \]
  2107 where the $x@{i,j}$ are either identifiers or nested tuple patterns as in
  2108 \S\ref{subsec:prod-sum}.
  2109 \begin{warn}
  2110   All constructors must be present, their order is fixed, and nested patterns
  2111   are not supported (with the exception of tuples).  Violating this
  2112   restriction results in strange error messages.
  2113 \end{warn}
  2114 
  2115 To perform case distinction on a goal containing a \texttt{case}-construct,
  2116 the theorem $t@j.$\texttt{split} is provided:
  2117 \[
  2118 \begin{array}{@{}rcl@{}}
  2119 P(t@j_\mathtt{case}~f@1~\dots~f@{k@j}~e) &\!\!\!=&
  2120 \!\!\! ((\forall x@1 \dots x@{m^j@1}. e = C^j@1~x@1\dots x@{m^j@1} \to
  2121                              P(f@1~x@1\dots x@{m^j@1})) \\
  2122 &&\!\!\! ~\land~ \dots ~\land \\
  2123 &&~\!\!\! (\forall x@1 \dots x@{m^j@{k@j}}. e = C^j@{k@j}~x@1\dots x@{m^j@{k@j}} \to
  2124                              P(f@{k@j}~x@1\dots x@{m^j@{k@j}})))
  2125 \end{array}
  2126 \]
  2127 where $t@j$\texttt{_case} is the internal name of the \texttt{case}-construct.
  2128 This theorem can be added to a simpset via \ttindex{addsplits}
  2129 (see~\S\ref{subsec:HOL:case:splitting}).
  2130 
  2131 \subsubsection{The function \cdx{size}}\label{sec:HOL:size}
  2132 
  2133 Theory \texttt{Arith} declares a generic function \texttt{size} of type
  2134 $\alpha\To nat$.  Each datatype defines a particular instance of \texttt{size}
  2135 by overloading according to the following scheme:
  2136 %%% FIXME: This formula is too big and is completely unreadable
  2137 \[
  2138 size(C^j@i~x@1~\dots~x@{m^j@i}) = \!
  2139 \left\{
  2140 \begin{array}{ll}
  2141 0 & \!\mbox{if $Rec^j@i = \emptyset$} \\
  2142 1+\sum\limits@{h=1}^{l^j@i}size~x@{r^j@{i,h}} &
  2143  \!\mbox{if $Rec^j@i = \left\{\left(r^j@{i,1},s^j@{i,1}\right),\ldots,
  2144   \left(r^j@{i,l^j@i},s^j@{i,l^j@i}\right)\right\}$}
  2145 \end{array}
  2146 \right.
  2147 \]
  2148 where $Rec^j@i$ is defined above.  Viewing datatypes as generalised trees, the
  2149 size of a leaf is 0 and the size of a node is the sum of the sizes of its
  2150 subtrees ${}+1$.
  2151 
  2152 \subsection{Defining datatypes}
  2153 
  2154 The theory syntax for datatype definitions is shown in
  2155 Fig.~\ref{datatype-grammar}.  In order to be well-formed, a datatype
  2156 definition has to obey the rules stated in the previous section.  As a result
  2157 the theory is extended with the new types, the constructors, and the theorems
  2158 listed in the previous section.
  2159 
  2160 \begin{figure}
  2161 \begin{rail}
  2162 datatype : 'datatype' typedecls;
  2163 
  2164 typedecls: ( newtype '=' (cons + '|') ) + 'and'
  2165          ;
  2166 newtype  : typevarlist id ( () | '(' infix ')' )
  2167          ;
  2168 cons     : name (argtype *) ( () | ( '(' mixfix ')' ) )
  2169          ;
  2170 argtype  : id | tid | ('(' typevarlist id ')')
  2171          ;
  2172 \end{rail}
  2173 \caption{Syntax of datatype declarations}
  2174 \label{datatype-grammar}
  2175 \end{figure}
  2176 
  2177 Most of the theorems about datatypes become part of the default simpset and
  2178 you never need to see them again because the simplifier applies them
  2179 automatically.  Only induction or exhaustion are usually invoked by hand.
  2180 \begin{ttdescription}
  2181 \item[\ttindexbold{induct_tac} {\tt"}$x${\tt"} $i$]
  2182  applies structural induction on variable $x$ to subgoal $i$, provided the
  2183  type of $x$ is a datatype.
  2184 \item[\ttindexbold{mutual_induct_tac}
  2185   {\tt["}$x@1${\tt",}$\ldots${\tt,"}$x@n${\tt"]} $i$] applies simultaneous
  2186   structural induction on the variables $x@1,\ldots,x@n$ to subgoal $i$.  This
  2187   is the canonical way to prove properties of mutually recursive datatypes
  2188   such as \texttt{aexp} and \texttt{bexp}, or datatypes with nested recursion such as
  2189   \texttt{term}.
  2190 \end{ttdescription}
  2191 In some cases, induction is overkill and a case distinction over all
  2192 constructors of the datatype suffices.
  2193 \begin{ttdescription}
  2194 \item[\ttindexbold{exhaust_tac} {\tt"}$u${\tt"} $i$]
  2195  performs an exhaustive case analysis for the term $u$ whose type
  2196  must be a datatype.  If the datatype has $k@j$ constructors
  2197  $C^j@1$, \dots $C^j@{k@j}$, subgoal $i$ is replaced by $k@j$ new subgoals which
  2198  contain the additional assumption $u = C^j@{i'}~x@1~\dots~x@{m^j@{i'}}$ for
  2199  $i'=1$, $\dots$,~$k@j$.
  2200 \end{ttdescription}
  2201 
  2202 Note that induction is only allowed on free variables that should not occur
  2203 among the premises of the subgoal.  Exhaustion applies to arbitrary terms.
  2204 
  2205 \bigskip
  2206 
  2207 
  2208 For the technically minded, we exhibit some more details.  Processing the
  2209 theory file produces an \ML\ structure which, in addition to the usual
  2210 components, contains a structure named $t$ for each datatype $t$ defined in
  2211 the file.  Each structure $t$ contains the following elements:
  2212 \begin{ttbox}
  2213 val distinct : thm list
  2214 val inject : thm list
  2215 val induct : thm
  2216 val exhaust : thm
  2217 val cases : thm list
  2218 val split : thm
  2219 val split_asm : thm
  2220 val recs : thm list
  2221 val size : thm list
  2222 val simps : thm list
  2223 \end{ttbox}
  2224 \texttt{distinct}, \texttt{inject}, \texttt{induct}, \texttt{size}
  2225 and \texttt{split} contain the theorems
  2226 described above.  For user convenience, \texttt{distinct} contains
  2227 inequalities in both directions.  The reduction rules of the {\tt
  2228   case}-construct are in \texttt{cases}.  All theorems from {\tt
  2229   distinct}, \texttt{inject} and \texttt{cases} are combined in \texttt{simps}.
  2230 In case of mutually recursive datatypes, \texttt{recs}, \texttt{size}, \texttt{induct}
  2231 and \texttt{simps} are contained in a separate structure named $t@1_\ldots_t@n$.
  2232 
  2233 
  2234 \subsection{Representing existing types as datatypes}\label{subsec:datatype:rep_datatype}
  2235 
  2236 For foundational reasons, some basic types such as \texttt{nat}, \texttt{*}, {\tt
  2237   +}, \texttt{bool} and \texttt{unit} are not defined in a \texttt{datatype} section,
  2238 but by more primitive means using \texttt{typedef}. To be able to use the
  2239 tactics \texttt{induct_tac} and \texttt{exhaust_tac} and to define functions by
  2240 primitive recursion on these types, such types may be represented as actual
  2241 datatypes.  This is done by specifying an induction rule, as well as theorems
  2242 stating the distinctness and injectivity of constructors in a {\tt
  2243   rep_datatype} section.  For type \texttt{nat} this works as follows:
  2244 \begin{ttbox}
  2245 rep_datatype nat
  2246   distinct Suc_not_Zero, Zero_not_Suc
  2247   inject   Suc_Suc_eq
  2248   induct   nat_induct
  2249 \end{ttbox}
  2250 The datatype package automatically derives additional theorems for recursion
  2251 and case combinators from these rules.  Any of the basic HOL types mentioned
  2252 above are represented as datatypes.  Try an induction on \texttt{bool}
  2253 today.
  2254 
  2255 
  2256 \subsection{Examples}
  2257 
  2258 \subsubsection{The datatype $\alpha~mylist$}
  2259 
  2260 We want to define a type $\alpha~mylist$. To do this we have to build a new
  2261 theory that contains the type definition.  We start from the theory
  2262 \texttt{Datatype} instead of \texttt{Main} in order to avoid clashes with the
  2263 \texttt{List} theory of Isabelle/HOL.
  2264 \begin{ttbox}
  2265 MyList = Datatype +
  2266   datatype 'a mylist = Nil | Cons 'a ('a mylist)
  2267 end
  2268 \end{ttbox}
  2269 After loading the theory, we can prove $Cons~x~xs\neq xs$, for example.  To
  2270 ease the induction applied below, we state the goal with $x$ quantified at the
  2271 object-level.  This will be stripped later using \ttindex{qed_spec_mp}.
  2272 \begin{ttbox}
  2273 Goal "!x. Cons x xs ~= xs";
  2274 {\out Level 0}
  2275 {\out ! x. Cons x xs ~= xs}
  2276 {\out  1. ! x. Cons x xs ~= xs}
  2277 \end{ttbox}
  2278 This can be proved by the structural induction tactic:
  2279 \begin{ttbox}
  2280 by (induct_tac "xs" 1);
  2281 {\out Level 1}
  2282 {\out ! x. Cons x xs ~= xs}
  2283 {\out  1. ! x. Cons x Nil ~= Nil}
  2284 {\out  2. !!a mylist.}
  2285 {\out        ! x. Cons x mylist ~= mylist ==>}
  2286 {\out        ! x. Cons x (Cons a mylist) ~= Cons a mylist}
  2287 \end{ttbox}
  2288 The first subgoal can be proved using the simplifier.  Isabelle/HOL has
  2289 already added the freeness properties of lists to the default simplification
  2290 set.
  2291 \begin{ttbox}
  2292 by (Simp_tac 1);
  2293 {\out Level 2}
  2294 {\out ! x. Cons x xs ~= xs}
  2295 {\out  1. !!a mylist.}
  2296 {\out        ! x. Cons x mylist ~= mylist ==>}
  2297 {\out        ! x. Cons x (Cons a mylist) ~= Cons a mylist}
  2298 \end{ttbox}
  2299 Similarly, we prove the remaining goal.
  2300 \begin{ttbox}
  2301 by (Asm_simp_tac 1);
  2302 {\out Level 3}
  2303 {\out ! x. Cons x xs ~= xs}
  2304 {\out No subgoals!}
  2305 \ttbreak
  2306 qed_spec_mp "not_Cons_self";
  2307 {\out val not_Cons_self = "Cons x xs ~= xs" : thm}
  2308 \end{ttbox}
  2309 Because both subgoals could have been proved by \texttt{Asm_simp_tac}
  2310 we could have done that in one step:
  2311 \begin{ttbox}
  2312 by (ALLGOALS Asm_simp_tac);
  2313 \end{ttbox}
  2314 
  2315 
  2316 \subsubsection{The datatype $\alpha~mylist$ with mixfix syntax}
  2317 
  2318 In this example we define the type $\alpha~mylist$ again but this time
  2319 we want to write \texttt{[]} for \texttt{Nil} and we want to use infix
  2320 notation \verb|#| for \texttt{Cons}.  To do this we simply add mixfix
  2321 annotations after the constructor declarations as follows:
  2322 \begin{ttbox}
  2323 MyList = Datatype +
  2324   datatype 'a mylist =
  2325     Nil ("[]")  |
  2326     Cons 'a ('a mylist)  (infixr "#" 70)
  2327 end
  2328 \end{ttbox}
  2329 Now the theorem in the previous example can be written \verb|x#xs ~= xs|.
  2330 
  2331 
  2332 \subsubsection{A datatype for weekdays}
  2333 
  2334 This example shows a datatype that consists of 7 constructors:
  2335 \begin{ttbox}
  2336 Days = Main +
  2337   datatype days = Mon | Tue | Wed | Thu | Fri | Sat | Sun
  2338 end
  2339 \end{ttbox}
  2340 Because there are more than 6 constructors, inequality is expressed via a function
  2341 \verb|days_ord|.  The theorem \verb|Mon ~= Tue| is not directly
  2342 contained among the distinctness theorems, but the simplifier can
  2343 prove it thanks to rewrite rules inherited from theory \texttt{Arith}:
  2344 \begin{ttbox}
  2345 Goal "Mon ~= Tue";
  2346 by (Simp_tac 1);
  2347 \end{ttbox}
  2348 You need not derive such inequalities explicitly: the simplifier will dispose
  2349 of them automatically.
  2350 \index{*datatype|)}
  2351 
  2352 
  2353 \section{Recursive function definitions}\label{sec:HOL:recursive}
  2354 \index{recursive functions|see{recursion}}
  2355 
  2356 Isabelle/HOL provides two main mechanisms of defining recursive functions.
  2357 \begin{enumerate}
  2358 \item \textbf{Primitive recursion} is available only for datatypes, and it is
  2359   somewhat restrictive.  Recursive calls are only allowed on the argument's
  2360   immediate constituents.  On the other hand, it is the form of recursion most
  2361   often wanted, and it is easy to use.
  2362   
  2363 \item \textbf{Well-founded recursion} requires that you supply a well-founded
  2364   relation that governs the recursion.  Recursive calls are only allowed if
  2365   they make the argument decrease under the relation.  Complicated recursion
  2366   forms, such as nested recursion, can be dealt with.  Termination can even be
  2367   proved at a later time, though having unsolved termination conditions around
  2368   can make work difficult.%
  2369   \footnote{This facility is based on Konrad Slind's TFL
  2370     package~\cite{slind-tfl}.  Thanks are due to Konrad for implementing TFL
  2371     and assisting with its installation.}
  2372 \end{enumerate}
  2373 
  2374 Following good HOL tradition, these declarations do not assert arbitrary
  2375 axioms.  Instead, they define the function using a recursion operator.  Both
  2376 HOL and ZF derive the theory of well-founded recursion from first
  2377 principles~\cite{paulson-set-II}.  Primitive recursion over some datatype
  2378 relies on the recursion operator provided by the datatype package.  With
  2379 either form of function definition, Isabelle proves the desired recursion
  2380 equations as theorems.
  2381 
  2382 
  2383 \subsection{Primitive recursive functions}
  2384 \label{sec:HOL:primrec}
  2385 \index{recursion!primitive|(}
  2386 \index{*primrec|(}
  2387 
  2388 Datatypes come with a uniform way of defining functions, {\bf primitive
  2389   recursion}.  In principle, one could introduce primitive recursive functions
  2390 by asserting their reduction rules as new axioms, but this is not recommended:
  2391 \begin{ttbox}\slshape
  2392 Append = Main +
  2393 consts app :: ['a list, 'a list] => 'a list
  2394 rules 
  2395    app_Nil   "app [] ys = ys"
  2396    app_Cons  "app (x#xs) ys = x#app xs ys"
  2397 end
  2398 \end{ttbox}
  2399 Asserting axioms brings the danger of accidentally asserting nonsense, as
  2400 in \verb$app [] ys = us$.
  2401 
  2402 The \ttindex{primrec} declaration is a safe means of defining primitive
  2403 recursive functions on datatypes:
  2404 \begin{ttbox}
  2405 Append = Main +
  2406 consts app :: ['a list, 'a list] => 'a list
  2407 primrec
  2408    "app [] ys = ys"
  2409    "app (x#xs) ys = x#app xs ys"
  2410 end
  2411 \end{ttbox}
  2412 Isabelle will now check that the two rules do indeed form a primitive
  2413 recursive definition.  For example
  2414 \begin{ttbox}
  2415 primrec
  2416     "app [] ys = us"
  2417 \end{ttbox}
  2418 is rejected with an error message ``\texttt{Extra variables on rhs}''.
  2419 
  2420 \bigskip
  2421 
  2422 The general form of a primitive recursive definition is
  2423 \begin{ttbox}
  2424 primrec
  2425     {\it reduction rules}
  2426 \end{ttbox}
  2427 where \textit{reduction rules} specify one or more equations of the form
  2428 \[ f \, x@1 \, \dots \, x@m \, (C \, y@1 \, \dots \, y@k) \, z@1 \,
  2429 \dots \, z@n = r \] such that $C$ is a constructor of the datatype, $r$
  2430 contains only the free variables on the left-hand side, and all recursive
  2431 calls in $r$ are of the form $f \, \dots \, y@i \, \dots$ for some $i$.  There
  2432 must be at most one reduction rule for each constructor.  The order is
  2433 immaterial.  For missing constructors, the function is defined to return a
  2434 default value.  
  2435 
  2436 If you would like to refer to some rule by name, then you must prefix
  2437 the rule with an identifier.  These identifiers, like those in the
  2438 \texttt{rules} section of a theory, will be visible at the \ML\ level.
  2439 
  2440 The primitive recursive function can have infix or mixfix syntax:
  2441 \begin{ttbox}\underscoreon
  2442 consts "@"  :: ['a list, 'a list] => 'a list  (infixr 60)
  2443 primrec
  2444    "[] @ ys = ys"
  2445    "(x#xs) @ ys = x#(xs @ ys)"
  2446 \end{ttbox}
  2447 
  2448 The reduction rules become part of the default simpset, which
  2449 leads to short proof scripts:
  2450 \begin{ttbox}\underscoreon
  2451 Goal "(xs @ ys) @ zs = xs @ (ys @ zs)";
  2452 by (induct\_tac "xs" 1);
  2453 by (ALLGOALS Asm\_simp\_tac);
  2454 \end{ttbox}
  2455 
  2456 \subsubsection{Example: Evaluation of expressions}
  2457 Using mutual primitive recursion, we can define evaluation functions \texttt{evala}
  2458 and \texttt{eval_bexp} for the datatypes of arithmetic and boolean expressions mentioned in
  2459 \S\ref{subsec:datatype:basics}:
  2460 \begin{ttbox}
  2461 consts
  2462   evala :: "['a => nat, 'a aexp] => nat"
  2463   evalb :: "['a => nat, 'a bexp] => bool"
  2464 
  2465 primrec
  2466   "evala env (If_then_else b a1 a2) =
  2467      (if evalb env b then evala env a1 else evala env a2)"
  2468   "evala env (Sum a1 a2) = evala env a1 + evala env a2"
  2469   "evala env (Diff a1 a2) = evala env a1 - evala env a2"
  2470   "evala env (Var v) = env v"
  2471   "evala env (Num n) = n"
  2472 
  2473   "evalb env (Less a1 a2) = (evala env a1 < evala env a2)"
  2474   "evalb env (And b1 b2) = (evalb env b1 & evalb env b2)"
  2475   "evalb env (Or b1 b2) = (evalb env b1 & evalb env b2)"
  2476 \end{ttbox}
  2477 Since the value of an expression depends on the value of its variables,
  2478 the functions \texttt{evala} and \texttt{evalb} take an additional
  2479 parameter, an {\em environment} of type \texttt{'a => nat}, which maps
  2480 variables to their values.
  2481 
  2482 Similarly, we may define substitution functions \texttt{substa}
  2483 and \texttt{substb} for expressions: The mapping \texttt{f} of type
  2484 \texttt{'a => 'a aexp} given as a parameter is lifted canonically
  2485 on the types \texttt{'a aexp} and \texttt{'a bexp}:
  2486 \begin{ttbox}
  2487 consts
  2488   substa :: "['a => 'b aexp, 'a aexp] => 'b aexp"
  2489   substb :: "['a => 'b aexp, 'a bexp] => 'b bexp"
  2490 
  2491 primrec
  2492   "substa f (If_then_else b a1 a2) =
  2493      If_then_else (substb f b) (substa f a1) (substa f a2)"
  2494   "substa f (Sum a1 a2) = Sum (substa f a1) (substa f a2)"
  2495   "substa f (Diff a1 a2) = Diff (substa f a1) (substa f a2)"
  2496   "substa f (Var v) = f v"
  2497   "substa f (Num n) = Num n"
  2498 
  2499   "substb f (Less a1 a2) = Less (substa f a1) (substa f a2)"
  2500   "substb f (And b1 b2) = And (substb f b1) (substb f b2)"
  2501   "substb f (Or b1 b2) = Or (substb f b1) (substb f b2)"
  2502 \end{ttbox}
  2503 In textbooks about semantics one often finds {\em substitution theorems},
  2504 which express the relationship between substitution and evaluation. For
  2505 \texttt{'a aexp} and \texttt{'a bexp}, we can prove such a theorem by mutual
  2506 induction, followed by simplification:
  2507 \begin{ttbox}
  2508 Goal
  2509   "evala env (substa (Var(v := a')) a) =
  2510      evala (env(v := evala env a')) a &
  2511    evalb env (substb (Var(v := a')) b) =
  2512      evalb (env(v := evala env a')) b";
  2513 by (mutual_induct_tac ["a","b"] 1);
  2514 by (ALLGOALS Asm_full_simp_tac);
  2515 \end{ttbox}
  2516 
  2517 \subsubsection{Example: A substitution function for terms}
  2518 Functions on datatypes with nested recursion, such as the type
  2519 \texttt{term} mentioned in \S\ref{subsec:datatype:basics}, are
  2520 also defined by mutual primitive recursion. A substitution
  2521 function \texttt{subst_term} on type \texttt{term}, similar to the functions
  2522 \texttt{substa} and \texttt{substb} described above, can
  2523 be defined as follows:
  2524 \begin{ttbox}
  2525 consts
  2526   subst_term :: "['a => ('a, 'b) term, ('a, 'b) term] => ('a, 'b) term"
  2527   subst_term_list ::
  2528     "['a => ('a, 'b) term, ('a, 'b) term list] => ('a, 'b) term list"
  2529 
  2530 primrec
  2531   "subst_term f (Var a) = f a"
  2532   "subst_term f (App b ts) = App b (subst_term_list f ts)"
  2533 
  2534   "subst_term_list f [] = []"
  2535   "subst_term_list f (t # ts) =
  2536      subst_term f t # subst_term_list f ts"
  2537 \end{ttbox}
  2538 The recursion scheme follows the structure of the unfolded definition of type
  2539 \texttt{term} shown in \S\ref{subsec:datatype:basics}. To prove properties of
  2540 this substitution function, mutual induction is needed:
  2541 \begin{ttbox}
  2542 Goal
  2543   "(subst_term ((subst_term f1) o f2) t) =
  2544      (subst_term f1 (subst_term f2 t)) &
  2545    (subst_term_list ((subst_term f1) o f2) ts) =
  2546      (subst_term_list f1 (subst_term_list f2 ts))";
  2547 by (mutual_induct_tac ["t", "ts"] 1);
  2548 by (ALLGOALS Asm_full_simp_tac);
  2549 \end{ttbox}
  2550 
  2551 \subsubsection{Example: A map function for infinitely branching trees}
  2552 Defining functions on infinitely branching datatypes by primitive
  2553 recursion is just as easy. For example, we can define a function
  2554 \texttt{map_tree} on \texttt{'a tree} as follows:
  2555 \begin{ttbox}
  2556 consts
  2557   map_tree :: "('a => 'b) => 'a tree => 'b tree"
  2558 
  2559 primrec
  2560   "map_tree f (Atom a) = Atom (f a)"
  2561   "map_tree f (Branch ts) = Branch (\%x. map_tree f (ts x))"
  2562 \end{ttbox}
  2563 Note that all occurrences of functions such as \texttt{ts} in the
  2564 \texttt{primrec} clauses must be applied to an argument. In particular,
  2565 \texttt{map_tree f o ts} is not allowed.
  2566 
  2567 \index{recursion!primitive|)}
  2568 \index{*primrec|)}
  2569 
  2570 
  2571 \subsection{General recursive functions}
  2572 \label{sec:HOL:recdef}
  2573 \index{recursion!general|(}
  2574 \index{*recdef|(}
  2575 
  2576 Using \texttt{recdef}, you can declare functions involving nested recursion
  2577 and pattern-matching.  Recursion need not involve datatypes and there are few
  2578 syntactic restrictions.  Termination is proved by showing that each recursive
  2579 call makes the argument smaller in a suitable sense, which you specify by
  2580 supplying a well-founded relation.
  2581 
  2582 Here is a simple example, the Fibonacci function.  The first line declares
  2583 \texttt{fib} to be a constant.  The well-founded relation is simply~$<$ (on
  2584 the natural numbers).  Pattern-matching is used here: \texttt{1} is a
  2585 macro for \texttt{Suc~0}.
  2586 \begin{ttbox}
  2587 consts fib  :: "nat => nat"
  2588 recdef fib "less_than"
  2589     "fib 0 = 0"
  2590     "fib 1 = 1"
  2591     "fib (Suc(Suc x)) = (fib x + fib (Suc x))"
  2592 \end{ttbox}
  2593 
  2594 With \texttt{recdef}, function definitions may be incomplete, and patterns may
  2595 overlap, as in functional programming.  The \texttt{recdef} package
  2596 disambiguates overlapping patterns by taking the order of rules into account.
  2597 For missing patterns, the function is defined to return a default value.
  2598 
  2599 %For example, here is a declaration of the list function \cdx{hd}:
  2600 %\begin{ttbox}
  2601 %consts hd :: 'a list => 'a
  2602 %recdef hd "\{\}"
  2603 %    "hd (x#l) = x"
  2604 %\end{ttbox}
  2605 %Because this function is not recursive, we may supply the empty well-founded
  2606 %relation, $\{\}$.
  2607 
  2608 The well-founded relation defines a notion of ``smaller'' for the function's
  2609 argument type.  The relation $\prec$ is \textbf{well-founded} provided it
  2610 admits no infinitely decreasing chains
  2611 \[ \cdots\prec x@n\prec\cdots\prec x@1. \]
  2612 If the function's argument has type~$\tau$, then $\prec$ has to be a relation
  2613 over~$\tau$: it must have type $(\tau\times\tau)set$.
  2614 
  2615 Proving well-foundedness can be tricky, so Isabelle/HOL provides a collection
  2616 of operators for building well-founded relations.  The package recognises
  2617 these operators and automatically proves that the constructed relation is
  2618 well-founded.  Here are those operators, in order of importance:
  2619 \begin{itemize}
  2620 \item \texttt{less_than} is ``less than'' on the natural numbers.
  2621   (It has type $(nat\times nat)set$, while $<$ has type $[nat,nat]\To bool$.
  2622   
  2623 \item $\mathop{\mathtt{measure}} f$, where $f$ has type $\tau\To nat$, is the
  2624   relation~$\prec$ on type~$\tau$ such that $x\prec y$ iff $f(x)<f(y)$.
  2625   Typically, $f$ takes the recursive function's arguments (as a tuple) and
  2626   returns a result expressed in terms of the function \texttt{size}.  It is
  2627   called a \textbf{measure function}.  Recall that \texttt{size} is overloaded
  2628   and is defined on all datatypes (see \S\ref{sec:HOL:size}).
  2629                                                     
  2630 \item $\mathop{\mathtt{inv_image}} f\;R$ is a generalisation of
  2631   \texttt{measure}.  It specifies a relation such that $x\prec y$ iff $f(x)$
  2632   is less than $f(y)$ according to~$R$, which must itself be a well-founded
  2633   relation.
  2634 
  2635 \item $R@1\texttt{**}R@2$ is the lexicographic product of two relations.  It
  2636   is a relation on pairs and satisfies $(x@1,x@2)\prec(y@1,y@2)$ iff $x@1$
  2637   is less than $y@1$ according to~$R@1$ or $x@1=y@1$ and $x@2$
  2638   is less than $y@2$ according to~$R@2$.
  2639 
  2640 \item \texttt{finite_psubset} is the proper subset relation on finite sets.
  2641 \end{itemize}
  2642 
  2643 We can use \texttt{measure} to declare Euclid's algorithm for the greatest
  2644 common divisor.  The measure function, $\lambda(m,n). n$, specifies that the
  2645 recursion terminates because argument~$n$ decreases.
  2646 \begin{ttbox}
  2647 recdef gcd "measure ((\%(m,n). n) ::nat*nat=>nat)"
  2648     "gcd (m, n) = (if n=0 then m else gcd(n, m mod n))"
  2649 \end{ttbox}
  2650 
  2651 The general form of a well-founded recursive definition is
  2652 \begin{ttbox}
  2653 recdef {\it function} {\it rel}
  2654     congs   {\it congruence rules}      {\bf(optional)}
  2655     simpset {\it simplification set}      {\bf(optional)}
  2656    {\it reduction rules}
  2657 \end{ttbox}
  2658 where
  2659 \begin{itemize}
  2660 \item \textit{function} is the name of the function, either as an \textit{id}
  2661   or a \textit{string}.  
  2662   
  2663 \item \textit{rel} is a {\HOL} expression for the well-founded termination
  2664   relation.
  2665   
  2666 \item \textit{congruence rules} are required only in highly exceptional
  2667   circumstances.
  2668   
  2669 \item The \textit{simplification set} is used to prove that the supplied
  2670   relation is well-founded.  It is also used to prove the \textbf{termination
  2671     conditions}: assertions that arguments of recursive calls decrease under
  2672   \textit{rel}.  By default, simplification uses \texttt{simpset()}, which
  2673   is sufficient to prove well-foundedness for the built-in relations listed
  2674   above. 
  2675   
  2676 \item \textit{reduction rules} specify one or more recursion equations.  Each
  2677   left-hand side must have the form $f\,t$, where $f$ is the function and $t$
  2678   is a tuple of distinct variables.  If more than one equation is present then
  2679   $f$ is defined by pattern-matching on components of its argument whose type
  2680   is a \texttt{datatype}.  
  2681 
  2682   Unlike with \texttt{primrec}, the reduction rules are not added to the
  2683   default simpset, and individual rules may not be labelled with identifiers.
  2684   However, the identifier $f$\texttt{.rules} is visible at the \ML\ level
  2685   as a list of theorems.
  2686 \end{itemize}
  2687 
  2688 With the definition of \texttt{gcd} shown above, Isabelle/HOL is unable to
  2689 prove one termination condition.  It remains as a precondition of the
  2690 recursion theorems.
  2691 \begin{ttbox}
  2692 gcd.rules;
  2693 {\out ["! m n. n ~= 0 --> m mod n < n}
  2694 {\out   ==> gcd (?m, ?n) = (if ?n = 0 then ?m else gcd (?n, ?m mod ?n))"] }
  2695 {\out : thm list}
  2696 \end{ttbox}
  2697 The theory \texttt{HOL/ex/Primes} illustrates how to prove termination
  2698 conditions afterwards.  The function \texttt{Tfl.tgoalw} is like the standard
  2699 function \texttt{goalw}, which sets up a goal to prove, but its argument
  2700 should be the identifier $f$\texttt{.rules} and its effect is to set up a
  2701 proof of the termination conditions:
  2702 \begin{ttbox}
  2703 Tfl.tgoalw thy [] gcd.rules;
  2704 {\out Level 0}
  2705 {\out ! m n. n ~= 0 --> m mod n < n}
  2706 {\out  1. ! m n. n ~= 0 --> m mod n < n}
  2707 \end{ttbox}
  2708 This subgoal has a one-step proof using \texttt{simp_tac}.  Once the theorem
  2709 is proved, it can be used to eliminate the termination conditions from
  2710 elements of \texttt{gcd.rules}.  Theory \texttt{HOL/Subst/Unify} is a much
  2711 more complicated example of this process, where the termination conditions can
  2712 only be proved by complicated reasoning involving the recursive function
  2713 itself.
  2714 
  2715 Isabelle/HOL can prove the \texttt{gcd} function's termination condition
  2716 automatically if supplied with the right simpset.
  2717 \begin{ttbox}
  2718 recdef gcd "measure ((\%(m,n). n) ::nat*nat=>nat)"
  2719   simpset "simpset() addsimps [mod_less_divisor, zero_less_eq]"
  2720     "gcd (m, n) = (if n=0 then m else gcd(n, m mod n))"
  2721 \end{ttbox}
  2722 
  2723 A \texttt{recdef} definition also returns an induction rule specialised for
  2724 the recursive function.  For the \texttt{gcd} function above, the induction
  2725 rule is
  2726 \begin{ttbox}
  2727 gcd.induct;
  2728 {\out "(!!m n. n ~= 0 --> ?P n (m mod n) ==> ?P m n) ==> ?P ?u ?v" : thm}
  2729 \end{ttbox}
  2730 This rule should be used to reason inductively about the \texttt{gcd}
  2731 function.  It usually makes the induction hypothesis available at all
  2732 recursive calls, leading to very direct proofs.  If any termination conditions
  2733 remain unproved, they will become additional premises of this rule.
  2734 
  2735 \index{recursion!general|)}
  2736 \index{*recdef|)}
  2737 
  2738 
  2739 \section{Inductive and coinductive definitions}
  2740 \index{*inductive|(}
  2741 \index{*coinductive|(}
  2742 
  2743 An {\bf inductive definition} specifies the least set~$R$ closed under given
  2744 rules.  (Applying a rule to elements of~$R$ yields a result within~$R$.)  For
  2745 example, a structural operational semantics is an inductive definition of an
  2746 evaluation relation.  Dually, a {\bf coinductive definition} specifies the
  2747 greatest set~$R$ consistent with given rules.  (Every element of~$R$ can be
  2748 seen as arising by applying a rule to elements of~$R$.)  An important example
  2749 is using bisimulation relations to formalise equivalence of processes and
  2750 infinite data structures.
  2751 
  2752 A theory file may contain any number of inductive and coinductive
  2753 definitions.  They may be intermixed with other declarations; in
  2754 particular, the (co)inductive sets {\bf must} be declared separately as
  2755 constants, and may have mixfix syntax or be subject to syntax translations.
  2756 
  2757 Each (co)inductive definition adds definitions to the theory and also
  2758 proves some theorems.  Each definition creates an \ML\ structure, which is a
  2759 substructure of the main theory structure.
  2760 
  2761 This package is related to the \ZF\ one, described in a separate
  2762 paper,%
  2763 \footnote{It appeared in CADE~\cite{paulson-CADE}; a longer version is
  2764   distributed with Isabelle.}  %
  2765 which you should refer to in case of difficulties.  The package is simpler
  2766 than \ZF's thanks to \HOL's extra-logical automatic type-checking.  The types
  2767 of the (co)inductive sets determine the domain of the fixedpoint definition,
  2768 and the package does not have to use inference rules for type-checking.
  2769 
  2770 
  2771 \subsection{The result structure}
  2772 Many of the result structure's components have been discussed in the paper;
  2773 others are self-explanatory.
  2774 \begin{description}
  2775 \item[\tt defs] is the list of definitions of the recursive sets.
  2776 
  2777 \item[\tt mono] is a monotonicity theorem for the fixedpoint operator.
  2778 
  2779 \item[\tt unfold] is a fixedpoint equation for the recursive set (the union of
  2780 the recursive sets, in the case of mutual recursion).
  2781 
  2782 \item[\tt intrs] is the list of introduction rules, now proved as theorems, for
  2783 the recursive sets.  The rules are also available individually, using the
  2784 names given them in the theory file. 
  2785 
  2786 \item[\tt elims] is the list of elimination rule.
  2787 
  2788 \item[\tt elim] is the head of the list \texttt{elims}.
  2789   
  2790 \item[\tt mk_cases] is a function to create simplified instances of {\tt
  2791 elim} using freeness reasoning on underlying datatypes.
  2792 \end{description}
  2793 
  2794 For an inductive definition, the result structure contains the
  2795 rule \texttt{induct}.  For a
  2796 coinductive definition, it contains the rule \verb|coinduct|.
  2797 
  2798 Figure~\ref{def-result-fig} summarises the two result signatures,
  2799 specifying the types of all these components.
  2800 
  2801 \begin{figure}
  2802 \begin{ttbox}
  2803 sig
  2804 val defs         : thm list
  2805 val mono         : thm
  2806 val unfold       : thm
  2807 val intrs        : thm list
  2808 val elims        : thm list
  2809 val elim         : thm
  2810 val mk_cases     : string -> thm
  2811 {\it(Inductive definitions only)} 
  2812 val induct       : thm
  2813 {\it(coinductive definitions only)}
  2814 val coinduct     : thm
  2815 end
  2816 \end{ttbox}
  2817 \hrule
  2818 \caption{The {\ML} result of a (co)inductive definition} \label{def-result-fig}
  2819 \end{figure}
  2820 
  2821 \subsection{The syntax of a (co)inductive definition}
  2822 An inductive definition has the form
  2823 \begin{ttbox}
  2824 inductive    {\it inductive sets}
  2825   intrs      {\it introduction rules}
  2826   monos      {\it monotonicity theorems}
  2827   con_defs   {\it constructor definitions}
  2828 \end{ttbox}
  2829 A coinductive definition is identical, except that it starts with the keyword
  2830 \texttt{coinductive}.  
  2831 
  2832 The \texttt{monos} and \texttt{con_defs} sections are optional.  If present,
  2833 each is specified by a list of identifiers.
  2834 
  2835 \begin{itemize}
  2836 \item The \textit{inductive sets} are specified by one or more strings.
  2837 
  2838 \item The \textit{introduction rules} specify one or more introduction rules in
  2839   the form \textit{ident\/}~\textit{string}, where the identifier gives the name of
  2840   the rule in the result structure.
  2841 
  2842 \item The \textit{monotonicity theorems} are required for each operator
  2843   applied to a recursive set in the introduction rules.  There {\bf must}
  2844   be a theorem of the form $A\subseteq B\Imp M(A)\subseteq M(B)$, for each
  2845   premise $t\in M(R@i)$ in an introduction rule!
  2846 
  2847 \item The \textit{constructor definitions} contain definitions of constants
  2848   appearing in the introduction rules.  In most cases it can be omitted.
  2849 \end{itemize}
  2850 
  2851 
  2852 \subsection{Example of an inductive definition}
  2853 Two declarations, included in a theory file, define the finite powerset
  2854 operator.  First we declare the constant~\texttt{Fin}.  Then we declare it
  2855 inductively, with two introduction rules:
  2856 \begin{ttbox}
  2857 consts Fin :: 'a set => 'a set set
  2858 inductive "Fin A"
  2859   intrs
  2860     emptyI  "{\ttlbrace}{\ttrbrace} : Fin A"
  2861     insertI "[| a: A;  b: Fin A |] ==> insert a b : Fin A"
  2862 \end{ttbox}
  2863 The resulting theory structure contains a substructure, called~\texttt{Fin}.
  2864 It contains the \texttt{Fin}$~A$ introduction rules as the list \texttt{Fin.intrs},
  2865 and also individually as \texttt{Fin.emptyI} and \texttt{Fin.consI}.  The induction
  2866 rule is \texttt{Fin.induct}.
  2867 
  2868 For another example, here is a theory file defining the accessible
  2869 part of a relation.  The main thing to note is the use of~\texttt{Pow} in
  2870 the sole introduction rule, and the corresponding mention of the rule
  2871 \verb|Pow_mono| in the \texttt{monos} list.  The paper
  2872 \cite{paulson-CADE} discusses a \ZF\ version of this example in more
  2873 detail.
  2874 \begin{ttbox}
  2875 Acc = WF + 
  2876 consts pred :: "['b, ('a * 'b)set] => 'a set"   (*Set of predecessors*)
  2877        acc  :: "('a * 'a)set => 'a set"         (*Accessible part*)
  2878 defs   pred_def  "pred x r == {y. (y,x):r}"
  2879 inductive "acc r"
  2880   intrs
  2881      pred "pred a r: Pow(acc r) ==> a: acc r"
  2882   monos   Pow_mono
  2883 end
  2884 \end{ttbox}
  2885 The Isabelle distribution contains many other inductive definitions.  Simple
  2886 examples are collected on subdirectory \texttt{HOL/Induct}.  The theory
  2887 \texttt{HOL/Induct/LList} contains coinductive definitions.  Larger examples
  2888 may be found on other subdirectories of \texttt{HOL}, such as \texttt{IMP},
  2889 \texttt{Lambda} and \texttt{Auth}.
  2890 
  2891 \index{*coinductive|)} \index{*inductive|)}
  2892 
  2893 
  2894 \section{The examples directories}
  2895 
  2896 Directory \texttt{HOL/Auth} contains theories for proving the correctness of
  2897 cryptographic protocols~\cite{paulson-jcs}.  The approach is based upon
  2898 operational semantics rather than the more usual belief logics.  On the same
  2899 directory are proofs for some standard examples, such as the Needham-Schroeder
  2900 public-key authentication protocol and the Otway-Rees
  2901 protocol.
  2902 
  2903 Directory \texttt{HOL/IMP} contains a formalization of various denotational,
  2904 operational and axiomatic semantics of a simple while-language, the necessary
  2905 equivalence proofs, soundness and completeness of the Hoare rules with
  2906 respect to the denotational semantics, and soundness and completeness of a
  2907 verification condition generator.  Much of development is taken from
  2908 Winskel~\cite{winskel93}.  For details see~\cite{nipkow-IMP}.
  2909 
  2910 Directory \texttt{HOL/Hoare} contains a user friendly surface syntax for Hoare
  2911 logic, including a tactic for generating verification-conditions.
  2912 
  2913 Directory \texttt{HOL/MiniML} contains a formalization of the type system of
  2914 the core functional language Mini-ML and a correctness proof for its type
  2915 inference algorithm $\cal W$~\cite{milner78,nipkow-W}.
  2916 
  2917 Directory \texttt{HOL/Lambda} contains a formalization of untyped
  2918 $\lambda$-calculus in de~Bruijn notation and Church-Rosser proofs for $\beta$
  2919 and $\eta$ reduction~\cite{Nipkow-CR}.
  2920 
  2921 Directory \texttt{HOL/Subst} contains Martin Coen's mechanization of a theory of
  2922 substitutions and unifiers.  It is based on Paulson's previous
  2923 mechanisation in {\LCF}~\cite{paulson85} of Manna and Waldinger's
  2924 theory~\cite{mw81}.  It demonstrates a complicated use of \texttt{recdef},
  2925 with nested recursion.
  2926 
  2927 Directory \texttt{HOL/Induct} presents simple examples of (co)inductive
  2928 definitions and datatypes.
  2929 \begin{itemize}
  2930 \item Theory \texttt{PropLog} proves the soundness and completeness of
  2931   classical propositional logic, given a truth table semantics.  The only
  2932   connective is $\imp$.  A Hilbert-style axiom system is specified, and its
  2933   set of theorems defined inductively.  A similar proof in \ZF{} is
  2934   described elsewhere~\cite{paulson-set-II}.
  2935 
  2936 \item Theory \texttt{Term} defines the datatype \texttt{term}.
  2937 
  2938 \item Theory \texttt{ABexp} defines arithmetic and boolean expressions
  2939  as mutually recursive datatypes.
  2940 
  2941 \item The definition of lazy lists demonstrates methods for handling
  2942   infinite data structures and coinduction in higher-order
  2943   logic~\cite{paulson-coind}.%
  2944 \footnote{To be precise, these lists are \emph{potentially infinite} rather
  2945   than lazy.  Lazy implies a particular operational semantics.}
  2946   Theory \thydx{LList} defines an operator for
  2947   corecursion on lazy lists, which is used to define a few simple functions
  2948   such as map and append.   A coinduction principle is defined
  2949   for proving equations on lazy lists.
  2950   
  2951 \item Theory \thydx{LFilter} defines the filter functional for lazy lists.
  2952   This functional is notoriously difficult to define because finding the next
  2953   element meeting the predicate requires possibly unlimited search.  It is not
  2954   computable, but can be expressed using a combination of induction and
  2955   corecursion.  
  2956 
  2957 \item Theory \thydx{Exp} illustrates the use of iterated inductive definitions
  2958   to express a programming language semantics that appears to require mutual
  2959   induction.  Iterated induction allows greater modularity.
  2960 \end{itemize}
  2961 
  2962 Directory \texttt{HOL/ex} contains other examples and experimental proofs in
  2963 {\HOL}.  
  2964 \begin{itemize}
  2965 \item Theory \texttt{Recdef} presents many examples of using \texttt{recdef}
  2966   to define recursive functions.  Another example is \texttt{Fib}, which
  2967   defines the Fibonacci function.
  2968 
  2969 \item Theory \texttt{Primes} defines the Greatest Common Divisor of two
  2970   natural numbers and proves a key lemma of the Fundamental Theorem of
  2971   Arithmetic: if $p$ is prime and $p$ divides $m\times n$ then $p$ divides~$m$
  2972   or $p$ divides~$n$.
  2973 
  2974 \item Theory \texttt{Primrec} develops some computation theory.  It
  2975   inductively defines the set of primitive recursive functions and presents a
  2976   proof that Ackermann's function is not primitive recursive.
  2977 
  2978 \item File \texttt{cla.ML} demonstrates the classical reasoner on over sixty
  2979   predicate calculus theorems, ranging from simple tautologies to
  2980   moderately difficult problems involving equality and quantifiers.
  2981 
  2982 \item File \texttt{meson.ML} contains an experimental implementation of the {\sc
  2983     meson} proof procedure, inspired by Plaisted~\cite{plaisted90}.  It is
  2984   much more powerful than Isabelle's classical reasoner.  But it is less
  2985   useful in practice because it works only for pure logic; it does not
  2986   accept derived rules for the set theory primitives, for example.
  2987 
  2988 \item File \texttt{mesontest.ML} contains test data for the {\sc meson} proof
  2989   procedure.  These are mostly taken from Pelletier \cite{pelletier86}.
  2990 
  2991 \item File \texttt{set.ML} proves Cantor's Theorem, which is presented in
  2992   \S\ref{sec:hol-cantor} below, and the Schr\"oder-Bernstein Theorem.
  2993 
  2994 \item Theory \texttt{MT} contains Jacob Frost's formalization~\cite{frost93} of
  2995   Milner and Tofte's coinduction example~\cite{milner-coind}.  This
  2996   substantial proof concerns the soundness of a type system for a simple
  2997   functional language.  The semantics of recursion is given by a cyclic
  2998   environment, which makes a coinductive argument appropriate.
  2999 \end{itemize}
  3000 
  3001 
  3002 \goodbreak
  3003 \section{Example: Cantor's Theorem}\label{sec:hol-cantor}
  3004 Cantor's Theorem states that every set has more subsets than it has
  3005 elements.  It has become a favourite example in higher-order logic since
  3006 it is so easily expressed:
  3007 \[  \forall f::\alpha \To \alpha \To bool. \exists S::\alpha\To bool.
  3008     \forall x::\alpha. f~x \not= S 
  3009 \] 
  3010 %
  3011 Viewing types as sets, $\alpha\To bool$ represents the powerset
  3012 of~$\alpha$.  This version states that for every function from $\alpha$ to
  3013 its powerset, some subset is outside its range.  
  3014 
  3015 The Isabelle proof uses \HOL's set theory, with the type $\alpha\,set$ and
  3016 the operator \cdx{range}.
  3017 \begin{ttbox}
  3018 context Set.thy;
  3019 \end{ttbox}
  3020 The set~$S$ is given as an unknown instead of a
  3021 quantified variable so that we may inspect the subset found by the proof.
  3022 \begin{ttbox}
  3023 Goal "?S ~: range\thinspace(f :: 'a=>'a set)";
  3024 {\out Level 0}
  3025 {\out ?S ~: range f}
  3026 {\out  1. ?S ~: range f}
  3027 \end{ttbox}
  3028 The first two steps are routine.  The rule \tdx{rangeE} replaces
  3029 $\Var{S}\in \texttt{range} \, f$ by $\Var{S}=f~x$ for some~$x$.
  3030 \begin{ttbox}
  3031 by (resolve_tac [notI] 1);
  3032 {\out Level 1}
  3033 {\out ?S ~: range f}
  3034 {\out  1. ?S : range f ==> False}
  3035 \ttbreak
  3036 by (eresolve_tac [rangeE] 1);
  3037 {\out Level 2}
  3038 {\out ?S ~: range f}
  3039 {\out  1. !!x. ?S = f x ==> False}
  3040 \end{ttbox}
  3041 Next, we apply \tdx{equalityCE}, reasoning that since $\Var{S}=f~x$,
  3042 we have $\Var{c}\in \Var{S}$ if and only if $\Var{c}\in f~x$ for
  3043 any~$\Var{c}$.
  3044 \begin{ttbox}
  3045 by (eresolve_tac [equalityCE] 1);
  3046 {\out Level 3}
  3047 {\out ?S ~: range f}
  3048 {\out  1. !!x. [| ?c3 x : ?S; ?c3 x : f x |] ==> False}
  3049 {\out  2. !!x. [| ?c3 x ~: ?S; ?c3 x ~: f x |] ==> False}
  3050 \end{ttbox}
  3051 Now we use a bit of creativity.  Suppose that~$\Var{S}$ has the form of a
  3052 comprehension.  Then $\Var{c}\in\{x.\Var{P}~x\}$ implies
  3053 $\Var{P}~\Var{c}$.   Destruct-resolution using \tdx{CollectD}
  3054 instantiates~$\Var{S}$ and creates the new assumption.
  3055 \begin{ttbox}
  3056 by (dresolve_tac [CollectD] 1);
  3057 {\out Level 4}
  3058 {\out {\ttlbrace}x. ?P7 x{\ttrbrace} ~: range f}
  3059 {\out  1. !!x. [| ?c3 x : f x; ?P7(?c3 x) |] ==> False}
  3060 {\out  2. !!x. [| ?c3 x ~: {\ttlbrace}x. ?P7 x{\ttrbrace}; ?c3 x ~: f x |] ==> False}
  3061 \end{ttbox}
  3062 Forcing a contradiction between the two assumptions of subgoal~1
  3063 completes the instantiation of~$S$.  It is now the set $\{x. x\not\in
  3064 f~x\}$, which is the standard diagonal construction.
  3065 \begin{ttbox}
  3066 by (contr_tac 1);
  3067 {\out Level 5}
  3068 {\out {\ttlbrace}x. x ~: f x{\ttrbrace} ~: range f}
  3069 {\out  1. !!x. [| x ~: {\ttlbrace}x. x ~: f x{\ttrbrace}; x ~: f x |] ==> False}
  3070 \end{ttbox}
  3071 The rest should be easy.  To apply \tdx{CollectI} to the negated
  3072 assumption, we employ \ttindex{swap_res_tac}:
  3073 \begin{ttbox}
  3074 by (swap_res_tac [CollectI] 1);
  3075 {\out Level 6}
  3076 {\out {\ttlbrace}x. x ~: f x{\ttrbrace} ~: range f}
  3077 {\out  1. !!x. [| x ~: f x; ~ False |] ==> x ~: f x}
  3078 \ttbreak
  3079 by (assume_tac 1);
  3080 {\out Level 7}
  3081 {\out {\ttlbrace}x. x ~: f x{\ttrbrace} ~: range f}
  3082 {\out No subgoals!}
  3083 \end{ttbox}
  3084 How much creativity is required?  As it happens, Isabelle can prove this
  3085 theorem automatically.  The default classical set \texttt{claset()} contains rules
  3086 for most of the constructs of \HOL's set theory.  We must augment it with
  3087 \tdx{equalityCE} to break up set equalities, and then apply best-first
  3088 search.  Depth-first search would diverge, but best-first search
  3089 successfully navigates through the large search space.
  3090 \index{search!best-first}
  3091 \begin{ttbox}
  3092 choplev 0;
  3093 {\out Level 0}
  3094 {\out ?S ~: range f}
  3095 {\out  1. ?S ~: range f}
  3096 \ttbreak
  3097 by (best_tac (claset() addSEs [equalityCE]) 1);
  3098 {\out Level 1}
  3099 {\out {\ttlbrace}x. x ~: f x{\ttrbrace} ~: range f}
  3100 {\out No subgoals!}
  3101 \end{ttbox}
  3102 If you run this example interactively, make sure your current theory contains
  3103 theory \texttt{Set}, for example by executing \ttindex{context}~{\tt Set.thy}.
  3104 Otherwise the default claset may not contain the rules for set theory.
  3105 \index{higher-order logic|)}
  3106 
  3107 %%% Local Variables: 
  3108 %%% mode: latex
  3109 %%% TeX-master: "logics"
  3110 %%% End: