| 7046 |      1 | 
 | 
| 7167 |      2 | \chapter{Isabelle/HOL Tools and Packages}\label{ch:hol-tools}
 | 
| 7135 |      3 | 
 | 
| 7990 |      4 | \section{Miscellaneous attributes}
 | 
|  |      5 | 
 | 
|  |      6 | \indexisaratt{rulify}\indexisaratt{rulify-prems}
 | 
|  |      7 | \begin{matharray}{rcl}
 | 
|  |      8 |   rulify & : & \isaratt \\
 | 
|  |      9 |   rulify_prems & : & \isaratt \\
 | 
|  |     10 | \end{matharray}
 | 
|  |     11 | 
 | 
|  |     12 | \begin{descr}
 | 
|  |     13 | 
 | 
|  |     14 | \item [$rulify$] puts a theorem into object-rule form, replacing implication
 | 
|  |     15 |   and universal quantification of HOL by the corresponding meta-logical
 | 
|  |     16 |   connectives.  This is the same operation as performed by the
 | 
|  |     17 |   \texttt{qed_spec_mp} ML function.
 | 
|  |     18 |   
 | 
|  |     19 | \item [$rulify_prems$] is similar to $rulify$, but acts on the premises of a
 | 
|  |     20 |   rule.
 | 
|  |     21 | 
 | 
|  |     22 | \end{descr}
 | 
|  |     23 | 
 | 
|  |     24 | 
 | 
| 7135 |     25 | \section{Primitive types}
 | 
|  |     26 | 
 | 
| 7141 |     27 | \indexisarcmd{typedecl}\indexisarcmd{typedef}
 | 
|  |     28 | \begin{matharray}{rcl}
 | 
|  |     29 |   \isarcmd{typedecl} & : & \isartrans{theory}{theory} \\
 | 
|  |     30 |   \isarcmd{typedef} & : & \isartrans{theory}{proof(prove)} \\
 | 
|  |     31 | \end{matharray}
 | 
|  |     32 | 
 | 
|  |     33 | \begin{rail}
 | 
|  |     34 |   'typedecl' typespec infix? comment?
 | 
|  |     35 |   ;
 | 
|  |     36 |   'typedef' parname? typespec infix? \\ '=' term comment?
 | 
|  |     37 |   ;
 | 
|  |     38 | \end{rail}
 | 
|  |     39 | 
 | 
| 7167 |     40 | \begin{descr}
 | 
| 7141 |     41 | \item [$\isarkeyword{typedecl}~(\vec\alpha)t$] is similar to the original
 | 
|  |     42 |   $\isarkeyword{typedecl}$ of Isabelle/Pure (see \S\ref{sec:types-pure}), but
 | 
|  |     43 |   also declares type arity $t :: (term, \dots, term) term$, making $t$ an
 | 
|  |     44 |   actual HOL type constructor.
 | 
|  |     45 | \item [$\isarkeyword{typedef}~(\vec\alpha)t = A$] sets up a goal stating
 | 
|  |     46 |   non-emptiness of the set $A$.  After finishing the proof, the theory will be
 | 
| 7175 |     47 |   augmented by a Gordon/HOL-style type definition.  See \cite{isabelle-HOL}
 | 
| 7335 |     48 |   for more information.  Note that user-level theories usually do not directly
 | 
|  |     49 |   refer to the HOL $\isarkeyword{typedef}$ primitive, but use more advanced
 | 
|  |     50 |   packages such as $\isarkeyword{record}$ (see \S\ref{sec:record}) and
 | 
| 7175 |     51 |   $\isarkeyword{datatype}$ (see \S\ref{sec:datatype}).
 | 
| 7167 |     52 | \end{descr}
 | 
| 7141 |     53 | 
 | 
|  |     54 | 
 | 
|  |     55 | \section{Records}\label{sec:record}
 | 
|  |     56 | 
 | 
|  |     57 | %FIXME record_split method
 | 
|  |     58 | \indexisarcmd{record}
 | 
|  |     59 | \begin{matharray}{rcl}
 | 
|  |     60 |   \isarcmd{record} & : & \isartrans{theory}{theory} \\
 | 
|  |     61 | \end{matharray}
 | 
|  |     62 | 
 | 
|  |     63 | \begin{rail}
 | 
|  |     64 |   'record' typespec '=' (type '+')? (field +)
 | 
|  |     65 |   ;
 | 
| 7135 |     66 | 
 | 
| 7141 |     67 |   field: name '::' type comment?
 | 
|  |     68 |   ;
 | 
|  |     69 | \end{rail}
 | 
|  |     70 | 
 | 
| 7167 |     71 | \begin{descr}
 | 
| 7141 |     72 | \item [$\isarkeyword{record}~(\vec\alpha)t = \tau + \vec c :: \vec\sigma$]
 | 
|  |     73 |   defines extensible record type $(\vec\alpha)t$, derived from the optional
 | 
|  |     74 |   parent record $\tau$ by adding new field components $\vec c :: \vec\sigma$.
 | 
| 7335 |     75 |   See \cite{isabelle-HOL,NaraschewskiW-TPHOLs98} for more information only
 | 
|  |     76 |   simply-typed extensible records.
 | 
| 7167 |     77 | \end{descr}
 | 
| 7141 |     78 | 
 | 
|  |     79 | 
 | 
|  |     80 | \section{Datatypes}\label{sec:datatype}
 | 
|  |     81 | 
 | 
| 7167 |     82 | \indexisarcmd{datatype}\indexisarcmd{rep-datatype}
 | 
| 7141 |     83 | \begin{matharray}{rcl}
 | 
|  |     84 |   \isarcmd{datatype} & : & \isartrans{theory}{theory} \\
 | 
|  |     85 |   \isarcmd{rep_datatype} & : & \isartrans{theory}{theory} \\
 | 
|  |     86 | \end{matharray}
 | 
|  |     87 | 
 | 
|  |     88 | \railalias{repdatatype}{rep\_datatype}
 | 
|  |     89 | \railterm{repdatatype}
 | 
|  |     90 | 
 | 
|  |     91 | \begin{rail}
 | 
| 7175 |     92 |   'datatype' (parname? typespec infix? \\ '=' (constructor + '|') + 'and')
 | 
| 7141 |     93 |   ;
 | 
|  |     94 |   repdatatype (name * ) \\ 'distinct' thmrefs 'inject' thmrefs 'induction' thmrefs
 | 
|  |     95 |   ;
 | 
|  |     96 | 
 | 
| 7175 |     97 |   constructor: name (type * ) mixfix? comment?
 | 
| 7141 |     98 |   ;
 | 
|  |     99 | \end{rail}
 | 
|  |    100 | 
 | 
| 7167 |    101 | \begin{descr}
 | 
| 7319 |    102 | \item [$\isarkeyword{datatype}$] defines inductive datatypes in HOL.
 | 
|  |    103 | \item [$\isarkeyword{rep_datatype}$] represents existing types as inductive
 | 
|  |    104 |   ones, generating the standard infrastructure of derived concepts (primitive
 | 
|  |    105 |   recursion etc.).
 | 
| 7167 |    106 | \end{descr}
 | 
| 7141 |    107 | 
 | 
| 8449 |    108 | The induction and exhaustion theorems generated provide case names according
 | 
|  |    109 | to the constructors involved, while parameters are named after the types (see
 | 
|  |    110 | also \S\ref{sec:induct-method}).
 | 
|  |    111 | 
 | 
| 7319 |    112 | See \cite{isabelle-HOL} for more details on datatypes.  Note that the theory
 | 
| 7335 |    113 | syntax above has been slightly simplified over the old version, usually
 | 
| 8531 |    114 | requiring more quotes and less parentheses.  Apart from proper proof methods
 | 
|  |    115 | for case-analysis and induction, there are also emulations of ML tactics
 | 
| 8945 |    116 | \texttt{case_tac} and \texttt{induct_tac} available, see
 | 
| 8665 |    117 | \S\ref{sec:induct_tac}.
 | 
| 7319 |    118 | 
 | 
| 7135 |    119 | 
 | 
|  |    120 | \section{Recursive functions}
 | 
|  |    121 | 
 | 
| 7141 |    122 | \indexisarcmd{primrec}\indexisarcmd{recdef}
 | 
|  |    123 | \begin{matharray}{rcl}
 | 
|  |    124 |   \isarcmd{primrec} & : & \isartrans{theory}{theory} \\
 | 
|  |    125 |   \isarcmd{recdef} & : & \isartrans{theory}{theory} \\
 | 
|  |    126 | %FIXME
 | 
|  |    127 | %  \isarcmd{defer_recdef} & : & \isartrans{theory}{theory} \\
 | 
|  |    128 | \end{matharray}
 | 
|  |    129 | 
 | 
|  |    130 | \begin{rail}
 | 
| 8657 |    131 |   'primrec' parname? (equation + )
 | 
|  |    132 |   ;
 | 
|  |    133 |   'recdef' name term (equation +) hints
 | 
| 7141 |    134 |   ;
 | 
| 8657 |    135 | 
 | 
|  |    136 |   equation: thmdecl? prop comment?
 | 
|  |    137 |   ;
 | 
| 8710 |    138 |   hints: ('congs' thmrefs)?
 | 
| 7141 |    139 |   ;
 | 
|  |    140 | \end{rail}
 | 
|  |    141 | 
 | 
| 7167 |    142 | \begin{descr}
 | 
| 7319 |    143 | \item [$\isarkeyword{primrec}$] defines primitive recursive functions over
 | 
|  |    144 |   datatypes.
 | 
|  |    145 | \item [$\isarkeyword{recdef}$] defines general well-founded recursive
 | 
|  |    146 |   functions (using the TFL package).
 | 
| 7167 |    147 | \end{descr}
 | 
| 7141 |    148 | 
 | 
| 8449 |    149 | Both definitions accommodate reasoning proof by induction (cf.\ 
 | 
|  |    150 | \S\ref{sec:induct-method}): rule $c\mathord{.}induct$ (where $c$ is the name
 | 
|  |    151 | of the function definition) refers to a specific induction rule, with
 | 
|  |    152 | parameters named according to the user-specified equations.  Case names of
 | 
|  |    153 | $\isarkeyword{primrec}$ are that of the datatypes involved, while those of
 | 
|  |    154 | $\isarkeyword{recdef}$ are numbered (starting from $1$).
 | 
|  |    155 | 
 | 
| 8657 |    156 | The equations provided by these packages may be referred later as theorem list
 | 
|  |    157 | $f\mathord.simps$, where $f$ is the (collective) name of the functions
 | 
|  |    158 | defined.  Individual equations may be named explicitly as well; note that for
 | 
|  |    159 | $\isarkeyword{recdef}$ each specification given by the user may result in
 | 
|  |    160 | several theorems.
 | 
|  |    161 | 
 | 
| 8449 |    162 | See \cite{isabelle-HOL} for further information on recursive function
 | 
|  |    163 | definitions in HOL.
 | 
| 7319 |    164 | 
 | 
| 7141 |    165 | 
 | 
| 7135 |    166 | \section{(Co)Inductive sets}
 | 
|  |    167 | 
 | 
| 7167 |    168 | \indexisarcmd{inductive}\indexisarcmd{coinductive}\indexisarcmd{inductive-cases}
 | 
| 7990 |    169 | \indexisaratt{mono}
 | 
| 7141 |    170 | \begin{matharray}{rcl}
 | 
|  |    171 |   \isarcmd{inductive} & : & \isartrans{theory}{theory} \\
 | 
|  |    172 |   \isarcmd{coinductive} & : & \isartrans{theory}{theory} \\
 | 
| 7990 |    173 |   mono & : & \isaratt \\
 | 
| 7141 |    174 |   \isarcmd{inductive_cases} & : & \isartrans{theory}{theory} \\
 | 
|  |    175 | \end{matharray}
 | 
|  |    176 | 
 | 
|  |    177 | \railalias{condefs}{con\_defs}
 | 
|  |    178 | \railalias{indcases}{inductive\_cases}
 | 
|  |    179 | \railterm{condefs,indcases}
 | 
|  |    180 | 
 | 
|  |    181 | \begin{rail}
 | 
|  |    182 |   ('inductive' | 'coinductive') (term comment? +) \\
 | 
|  |    183 |     'intrs' attributes? (thmdecl? prop comment? +) \\
 | 
|  |    184 |     'monos' thmrefs comment? \\ condefs thmrefs comment?
 | 
|  |    185 |   ;
 | 
|  |    186 |   indcases thmdef? nameref ':' \\ (prop +) comment?
 | 
|  |    187 |   ;
 | 
| 7990 |    188 |   'mono' (() | 'add' | 'del')
 | 
|  |    189 |   ;
 | 
| 7141 |    190 | \end{rail}
 | 
|  |    191 | 
 | 
| 7167 |    192 | \begin{descr}
 | 
| 7319 |    193 | \item [$\isarkeyword{inductive}$ and $\isarkeyword{coinductive}$] define
 | 
|  |    194 |   (co)inductive sets from the given introduction rules.
 | 
| 8547 |    195 | \item [$mono$] declares monotonicity rules.  These rule are involved in the
 | 
|  |    196 |   automated monotonicity proof of $\isarkeyword{inductive}$.
 | 
| 8449 |    197 | \item [$\isarkeyword{inductive_cases}$] creates instances of elimination rules
 | 
|  |    198 |   of (co)inductive sets, solving obvious cases by simplification.
 | 
|  |    199 |   
 | 
|  |    200 |   The $cases$ proof method (see \S\ref{sec:induct-method}) provides a more
 | 
|  |    201 |   direct way for reasoning by cases (including optional simplification).
 | 
|  |    202 |   
 | 
|  |    203 |   Unlike the \texttt{mk_cases} ML function exported with any inductive
 | 
|  |    204 |   definition \cite{isabelle-HOL}, $\isarkeyword{inductive_cases}$ it does
 | 
| 8547 |    205 |   \emph{not} modify cases by simplification that are not solved completely
 | 
|  |    206 |   anyway (e.g.\ due to contradictory assumptions).  Thus
 | 
|  |    207 |   $\isarkeyword{inductive_cases}$ conforms to the way Isar proofs are
 | 
|  |    208 |   conducted, rather than old-style tactic scripts.
 | 
| 7167 |    209 | \end{descr}
 | 
| 7141 |    210 | 
 | 
| 8449 |    211 | See \cite{isabelle-HOL} for further information on inductive definitions in
 | 
|  |    212 | HOL.
 | 
| 7319 |    213 | 
 | 
| 7141 |    214 | 
 | 
| 8449 |    215 | \section{Proof by cases and induction}\label{sec:induct-method}
 | 
|  |    216 | 
 | 
| 8666 |    217 | \subsection{Proof methods}\label{sec:induct-method-proper}
 | 
| 7141 |    218 | 
 | 
| 8449 |    219 | \indexisarmeth{cases}\indexisarmeth{induct}
 | 
| 7319 |    220 | \begin{matharray}{rcl}
 | 
| 8449 |    221 |   cases & : & \isarmeth \\
 | 
| 7319 |    222 |   induct & : & \isarmeth \\
 | 
|  |    223 | \end{matharray}
 | 
|  |    224 | 
 | 
| 8449 |    225 | The $cases$ and $induct$ methods provide a uniform interface to case analysis
 | 
|  |    226 | and induction over datatypes, inductive sets, and recursive functions.  The
 | 
|  |    227 | corresponding rules may be specified and instantiated in a casual manner.
 | 
|  |    228 | Furthermore, these methods provide named local contexts that may be invoked
 | 
|  |    229 | via the $\CASENAME$ proof command within the subsequent proof text (cf.\ 
 | 
| 8484 |    230 | \S\ref{sec:cases}).  This accommodates compact proof texts even when reasoning
 | 
|  |    231 | about large specifications.
 | 
| 7319 |    232 | 
 | 
|  |    233 | \begin{rail}
 | 
| 9307 |    234 |   'cases' ('(simplified)')? ('(opaque)')? \\ term? rule?  ;
 | 
| 8449 |    235 | 
 | 
| 9307 |    236 |   'induct' ('(stripped)')? ('(opaque)')? \\ (insts * 'and') rule?
 | 
| 7319 |    237 |   ;
 | 
|  |    238 | 
 | 
| 8449 |    239 |   rule: ('type' | 'set') ':' nameref | 'rule' ':' thmref
 | 
| 7319 |    240 |   ;
 | 
|  |    241 | \end{rail}
 | 
|  |    242 | 
 | 
|  |    243 | \begin{descr}
 | 
| 8449 |    244 | \item [$cases~t~R$] applies method $rule$ with an appropriate case distinction
 | 
|  |    245 |   theorem, instantiated to the subject $t$.  Symbolic case names are bound
 | 
|  |    246 |   according to the rule's local contexts.
 | 
|  |    247 |   
 | 
|  |    248 |   The rule is determined as follows, according to the facts and arguments
 | 
|  |    249 |   passed to the $cases$ method:
 | 
|  |    250 |   \begin{matharray}{llll}
 | 
|  |    251 |     \text{facts}    &       & \text{arguments} & \text{rule} \\\hline
 | 
|  |    252 |                     & cases &           & \text{classical case split} \\
 | 
|  |    253 |                     & cases & t         & \text{datatype exhaustion (type of $t$)} \\
 | 
|  |    254 |     \edrv a \in A   & cases & \dots     & \text{inductive set elimination (of $A$)} \\
 | 
|  |    255 |     \dots           & cases & \dots ~ R & \text{explicit rule $R$} \\
 | 
|  |    256 |   \end{matharray}
 | 
|  |    257 | 
 | 
|  |    258 |   The $simplified$ option causes ``obvious cases'' of the rule to be solved
 | 
|  |    259 |   beforehand, while the others are left unscathed.
 | 
|  |    260 |   
 | 
| 9307 |    261 |   The $opaque$ option causes the parameters of the new local contexts to be
 | 
|  |    262 |   turned into ``internal'' names.  This results in quasi-existentially bound
 | 
|  |    263 |   elements to be introduced when individual cases are invoked later.  Thus
 | 
|  |    264 |   both unwanted hiding of existing local variables and references to
 | 
|  |    265 |   implicitly bound variables (stemming from cases) are avoided.  So by using
 | 
|  |    266 |   the $opaque$ option, a proof text may become slightly more readable and
 | 
|  |    267 |   robust.
 | 
|  |    268 |   
 | 
| 8449 |    269 | \item [$induct~insts~R$] is analogous to the $cases$ method, but refers to
 | 
|  |    270 |   induction rules, which are determined as follows:
 | 
|  |    271 |   \begin{matharray}{llll}
 | 
|  |    272 |     \text{facts}    &        & \text{arguments} & \text{rule} \\\hline
 | 
|  |    273 |                     & induct & P ~ x ~ \dots & \text{datatype induction (type of $x$)} \\
 | 
|  |    274 |     \edrv x \in A   & induct & \dots         & \text{set induction (of $A$)} \\
 | 
|  |    275 |     \dots           & induct & \dots ~ R     & \text{explicit rule $R$} \\
 | 
|  |    276 |   \end{matharray}
 | 
|  |    277 |   
 | 
|  |    278 |   Several instantiations may be given, each referring to some part of a mutual
 | 
|  |    279 |   inductive definition or datatype --- only related partial induction rules
 | 
|  |    280 |   may be used together, though.  Any of the lists of terms $P, x, \dots$
 | 
|  |    281 |   refers to the \emph{suffix} of variables present in the induction rule.
 | 
|  |    282 |   This enables the writer to specify only induction variables, or both
 | 
|  |    283 |   predicates and variables, for example.
 | 
| 7507 |    284 |   
 | 
| 8449 |    285 |   The $stripped$ option causes implications and (bounded) universal
 | 
|  |    286 |   quantifiers to be removed from each new subgoal emerging from the
 | 
| 8547 |    287 |   application of the induction rule.  This accommodates typical
 | 
|  |    288 |   ``strengthening of induction'' predicates.
 | 
| 9307 |    289 |   
 | 
|  |    290 |   The $opaque$ option has the same effect as for the $cases$ method, see
 | 
|  |    291 |   above.
 | 
| 7319 |    292 | \end{descr}
 | 
| 7141 |    293 | 
 | 
| 8484 |    294 | Above methods produce named local contexts (cf.\ \S\ref{sec:cases}), as
 | 
|  |    295 | determined by the instantiated rule \emph{before} it has been applied to the
 | 
|  |    296 | internal proof state.\footnote{As a general principle, Isar proof text may
 | 
| 8449 |    297 |   never refer to parts of proof states directly.} Thus proper use of symbolic
 | 
|  |    298 | cases usually require the rule to be instantiated fully, as far as the
 | 
|  |    299 | emerging local contexts and subgoals are concerned.  In particular, for
 | 
|  |    300 | induction both the predicates and variables have to be specified.  Otherwise
 | 
| 8547 |    301 | the $\CASENAME$ command would refuse to invoke cases containing schematic
 | 
| 8449 |    302 | variables.
 | 
|  |    303 | 
 | 
|  |    304 | The $\isarkeyword{print_cases}$ command (\S\ref{sec:diag}) prints all named
 | 
| 8547 |    305 | cases present in the current proof state.
 | 
| 8449 |    306 | 
 | 
|  |    307 | 
 | 
| 8484 |    308 | \subsection{Declaring rules}
 | 
| 8449 |    309 | 
 | 
|  |    310 | \indexisaratt{cases}\indexisaratt{induct}
 | 
|  |    311 | \begin{matharray}{rcl}
 | 
|  |    312 |   cases & : & \isaratt \\
 | 
|  |    313 |   induct & : & \isaratt \\
 | 
|  |    314 | \end{matharray}
 | 
|  |    315 | 
 | 
|  |    316 | \begin{rail}
 | 
|  |    317 |   'cases' spec
 | 
|  |    318 |   ;
 | 
|  |    319 |   'induct' spec
 | 
|  |    320 |   ;
 | 
|  |    321 | 
 | 
|  |    322 |   spec: ('type' | 'set') ':' nameref
 | 
|  |    323 |   ;
 | 
|  |    324 | \end{rail}
 | 
|  |    325 | 
 | 
|  |    326 | The $cases$ and $induct$ attributes augment the corresponding context of rules
 | 
|  |    327 | for reasoning about inductive sets and types.  The standard rules are already
 | 
|  |    328 | declared by HOL definitional packages.  For special applications, these may be
 | 
|  |    329 | replaced manually by variant versions.
 | 
|  |    330 | 
 | 
| 8484 |    331 | Refer to the $case_names$ and $params$ attributes (see \S\ref{sec:cases}) to
 | 
|  |    332 | adjust names of cases and parameters of a rule.
 | 
|  |    333 | 
 | 
| 7046 |    334 | 
 | 
| 8665 |    335 | \subsection{Emulating tactic scripts}\label{sec:induct_tac}
 | 
|  |    336 | 
 | 
|  |    337 | \indexisarmeth{case-tac}\indexisarmeth{induct-tac}
 | 
|  |    338 | \begin{matharray}{rcl}
 | 
|  |    339 |   case_tac & : & \isarmeth \\
 | 
|  |    340 |   induct_tac & : & \isarmeth \\
 | 
|  |    341 | \end{matharray}
 | 
|  |    342 | 
 | 
|  |    343 | These proof methods directly correspond to the ML tactics of the same name
 | 
|  |    344 | \cite{isabelle-HOL}.  In particular, the instantiation given refers to the
 | 
|  |    345 | \emph{dynamic} proof state, rather than the current proof text.  This enables
 | 
|  |    346 | proof scripts to refer to parameters of some subgoal, for example.
 | 
|  |    347 | 
 | 
|  |    348 | \railalias{casetac}{case\_tac}
 | 
|  |    349 | \railterm{casetac}
 | 
|  |    350 | \railalias{inducttac}{induct\_tac}
 | 
|  |    351 | \railterm{inducttac}
 | 
|  |    352 | 
 | 
|  |    353 | \begin{rail}
 | 
| 8666 |    354 |   casetac goalspec? term rule?
 | 
| 8665 |    355 |   ;
 | 
| 8692 |    356 |   inducttac goalspec? (insts * 'and') rule?
 | 
| 8666 |    357 |   ;
 | 
|  |    358 | 
 | 
|  |    359 |   rule: ('rule' ':' thmref)
 | 
| 8665 |    360 |   ;
 | 
|  |    361 | \end{rail}
 | 
|  |    362 | 
 | 
| 8980 |    363 | By default, $case_tac$ and $induct_tac$ admit to reason about inductive
 | 
|  |    364 | datatypes only, unless an alternative rule is given explicitly.  Furthermore,
 | 
|  |    365 | $case_tac$ does a classical case split on booleans; $induct_tac$ allows only
 | 
|  |    366 | variables to be given as instantiation.  Also note that named local contexts
 | 
|  |    367 | (see \S\ref{sec:cases}) are not provided as would be by the proper $induct$
 | 
|  |    368 | and $cases$ proof methods (see \S\ref{sec:induct-method-proper}).
 | 
| 8665 |    369 | 
 | 
|  |    370 | 
 | 
| 7390 |    371 | \section{Arithmetic}
 | 
|  |    372 | 
 | 
|  |    373 | \indexisarmeth{arith}
 | 
|  |    374 | \begin{matharray}{rcl}
 | 
|  |    375 |   arith & : & \isarmeth \\
 | 
|  |    376 | \end{matharray}
 | 
|  |    377 | 
 | 
| 8506 |    378 | \begin{rail}
 | 
|  |    379 |   'arith' '!'?
 | 
|  |    380 |   ;
 | 
|  |    381 | \end{rail}
 | 
|  |    382 | 
 | 
| 7390 |    383 | The $arith$ method decides linear arithmetic problems (on types $nat$, $int$,
 | 
| 8506 |    384 | $real$).  Any current facts are inserted into the goal before running the
 | 
|  |    385 | procedure.  The ``!''~argument causes the full context of assumptions to be
 | 
| 8665 |    386 | included.
 | 
| 8506 |    387 | 
 | 
|  |    388 | Note that a simpler (but faster) version of arithmetic reasoning is already
 | 
|  |    389 | performed by the Simplifier.
 | 
| 7390 |    390 | 
 | 
|  |    391 | 
 | 
| 7046 |    392 | %%% Local Variables: 
 | 
|  |    393 | %%% mode: latex
 | 
|  |    394 | %%% TeX-master: "isar-ref"
 | 
|  |    395 | %%% End: 
 |