| author | huffman | 
| Wed, 10 Aug 2011 14:10:52 -0700 | |
| changeset 44137 | ac5cb4c86448 | 
| parent 43367 | 3f1469de9e47 | 
| permissions | -rw-r--r-- | 
| 30184 
37969710e61f
removed parts of the manual that are clearly obsolete, or covered by
 wenzelm parents: 
12366diff
changeset | 1 | |
| 319 | 2 | \chapter{The Classical Reasoner}\label{chap:classical}
 | 
| 286 | 3 | \index{classical reasoner|(}
 | 
| 308 | 4 | \newcommand\ainfer[2]{\begin{array}{r@{\,}l}#2\\ \hline#1\end{array}}
 | 
| 5 | ||
| 104 | 6 | \section{Classical rule sets}
 | 
| 319 | 7 | \index{classical sets}
 | 
| 104 | 8 | |
| 5550 | 9 | For elimination and destruction rules there are variants of the add operations | 
| 10 | adding a rule in a way such that it is applied only if also its second premise | |
| 11 | can be unified with an assumption of the current proof state: | |
| 5576 
dc6ee60d2be8
exchanged automatic-tactics and semi-automatic-tactics
 oheimb parents: 
5550diff
changeset | 12 | \indexbold{*addSE2}\indexbold{*addSD2}\indexbold{*addE2}\indexbold{*addD2}
 | 
| 5550 | 13 | \begin{ttbox}
 | 
| 14 | addSE2      : claset * (string * thm) -> claset           \hfill{\bf infix 4}
 | |
| 15 | addSD2      : claset * (string * thm) -> claset           \hfill{\bf infix 4}
 | |
| 16 | addE2       : claset * (string * thm) -> claset           \hfill{\bf infix 4}
 | |
| 17 | addD2       : claset * (string * thm) -> claset           \hfill{\bf infix 4}
 | |
| 18 | \end{ttbox}
 | |
| 19 | \begin{warn}
 | |
| 20 | A rule to be added in this special way must be given a name, which is used | |
| 21 |   to delete it again -- when desired -- using \texttt{delSWrappers} or 
 | |
| 22 |   \texttt{delWrappers}, respectively. This is because these add operations
 | |
| 23 |   are implemented as wrappers (see \ref{sec:modifying-search} below).
 | |
| 24 | \end{warn}
 | |
| 25 | ||
| 1099 | 26 | |
| 27 | \subsection{Modifying the search step}
 | |
| 4665 | 28 | \label{sec:modifying-search}
 | 
| 3716 | 29 | For a given classical set, the proof strategy is simple. Perform as many safe | 
| 30 | inferences as possible; or else, apply certain safe rules, allowing | |
| 31 | instantiation of unknowns; or else, apply an unsafe rule. The tactics also | |
| 32 | eliminate assumptions of the form $x=t$ by substitution if they have been set | |
| 11181 
d04f57b91166
renamed addaltern to addafter, addSaltern to addSafter
 oheimb parents: 
9695diff
changeset | 33 | up to do so (see \texttt{hyp_subst_tacs} in~{\S}\ref{sec:classical-setup} below).
 | 
| 3716 | 34 | They may perform a form of Modus Ponens: if there are assumptions $P\imp Q$ | 
| 35 | and~$P$, then replace $P\imp Q$ by~$Q$. | |
| 104 | 36 | |
| 3720 
a5b9e0ade194
Safe_tac, Safe_step_tac, and minor corrections including \tt -> \texttt
 paulson parents: 
3716diff
changeset | 37 | The classical reasoning tactics --- except \texttt{blast_tac}! --- allow
 | 
| 4649 | 38 | you to modify this basic proof strategy by applying two lists of arbitrary | 
| 39 | {\bf wrapper tacticals} to it. 
 | |
| 40 | The first wrapper list, which is considered to contain safe wrappers only, | |
| 41 | affects \ttindex{safe_step_tac} and all the tactics that call it.  
 | |
| 5550 | 42 | The second one, which may contain unsafe wrappers, affects the unsafe parts | 
| 43 | of \ttindex{step_tac}, \ttindex{slow_step_tac}, and the tactics that call them.
 | |
| 4649 | 44 | A wrapper transforms each step of the search, for example | 
| 5550 | 45 | by attempting other tactics before or after the original step tactic. | 
| 4649 | 46 | All members of a wrapper list are applied in turn to the respective step tactic. | 
| 47 | ||
| 48 | Initially the two wrapper lists are empty, which means no modification of the | |
| 49 | step tactics. Safe and unsafe wrappers are added to a claset | |
| 50 | with the functions given below, supplying them with wrapper names. | |
| 51 | These names may be used to selectively delete wrappers. | |
| 1099 | 52 | |
| 53 | \begin{ttbox} 
 | |
| 4649 | 54 | type wrapper = (int -> tactic) -> (int -> tactic); | 
| 4881 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 55 | |
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 56 | addSWrapper  : claset * (string *  wrapper       ) -> claset \hfill{\bf infix 4}
 | 
| 4649 | 57 | addSbefore   : claset * (string * (int -> tactic)) -> claset \hfill{\bf infix 4}
 | 
| 11181 
d04f57b91166
renamed addaltern to addafter, addSaltern to addSafter
 oheimb parents: 
9695diff
changeset | 58 | addSafter    : claset * (string * (int -> tactic)) -> claset \hfill{\bf infix 4}
 | 
| 4649 | 59 | delSWrapper  : claset *  string                    -> claset \hfill{\bf infix 4}
 | 
| 4881 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 60 | |
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 61 | addWrapper   : claset * (string *  wrapper       ) -> claset \hfill{\bf infix 4}
 | 
| 4649 | 62 | addbefore    : claset * (string * (int -> tactic)) -> claset \hfill{\bf infix 4}
 | 
| 11181 
d04f57b91166
renamed addaltern to addafter, addSaltern to addSafter
 oheimb parents: 
9695diff
changeset | 63 | addafter     : claset * (string * (int -> tactic)) -> claset \hfill{\bf infix 4}
 | 
| 4649 | 64 | delWrapper   : claset *  string                    -> claset \hfill{\bf infix 4}
 | 
| 65 | ||
| 4881 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 66 | addSss       : claset * simpset -> claset                 \hfill{\bf infix 4}
 | 
| 2632 | 67 | addss        : claset * simpset -> claset                 \hfill{\bf infix 4}
 | 
| 1099 | 68 | \end{ttbox}
 | 
| 69 | % | |
| 70 | ||
| 71 | \begin{ttdescription}
 | |
| 4881 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 72 | \item[$cs$ addSWrapper $(name,wrapper)$] \indexbold{*addSWrapper}
 | 
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 73 | adds a new wrapper, which should yield a safe tactic, | 
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 74 | to modify the existing safe step tactic. | 
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 75 | |
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 76 | \item[$cs$ addSbefore $(name,tac)$] \indexbold{*addSbefore}
 | 
| 5550 | 77 | adds the given tactic as a safe wrapper, such that it is tried | 
| 78 | {\em before} each safe step of the search.
 | |
| 4881 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 79 | |
| 11181 
d04f57b91166
renamed addaltern to addafter, addSaltern to addSafter
 oheimb parents: 
9695diff
changeset | 80 | \item[$cs$ addSafter $(name,tac)$] \indexbold{*addSafter}
 | 
| 5550 | 81 | adds the given tactic as a safe wrapper, such that it is tried | 
| 82 | when a safe step of the search would fail. | |
| 4881 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 83 | |
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 84 | \item[$cs$ delSWrapper $name$] \indexbold{*delSWrapper}
 | 
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 85 | deletes the safe wrapper with the given name. | 
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 86 | |
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 87 | \item[$cs$ addWrapper $(name,wrapper)$] \indexbold{*addWrapper}
 | 
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 88 | adds a new wrapper to modify the existing (unsafe) step tactic. | 
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 89 | |
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 90 | \item[$cs$ addbefore $(name,tac)$] \indexbold{*addbefore}
 | 
| 5550 | 91 | adds the given tactic as an unsafe wrapper, such that it its result is | 
| 92 | concatenated {\em before} the result of each unsafe step.
 | |
| 4881 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 93 | |
| 11181 
d04f57b91166
renamed addaltern to addafter, addSaltern to addSafter
 oheimb parents: 
9695diff
changeset | 94 | \item[$cs$ addafter $(name,tac)$] \indexbold{*addafter}
 | 
| 5550 | 95 | adds the given tactic as an unsafe wrapper, such that it its result is | 
| 96 | concatenated {\em after} the result of each unsafe step.
 | |
| 4881 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 97 | |
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 98 | \item[$cs$ delWrapper $name$] \indexbold{*delWrapper}
 | 
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 99 | deletes the unsafe wrapper with the given name. | 
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 100 | |
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 101 | \item[$cs$ addSss $ss$] \indexbold{*addss}
 | 
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 102 | adds the simpset~$ss$ to the classical set. The assumptions and goal will be | 
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 103 | simplified, in a rather safe way, after each safe step of the search. | 
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 104 | |
| 1099 | 105 | \item[$cs$ addss $ss$] \indexbold{*addss}
 | 
| 3485 
f27a30a18a17
Now there are TWO spaces after each full stop, so that the Emacs sentence
 paulson parents: 
3224diff
changeset | 106 | adds the simpset~$ss$ to the classical set. The assumptions and goal will be | 
| 4881 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 107 | simplified, before the each unsafe step of the search. | 
| 2631 
5e5c78ba955e
description of safe vs. unsafe wrapper and the functions involved
 oheimb parents: 
2479diff
changeset | 108 | |
| 4881 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 109 | \end{ttdescription}
 | 
| 2631 
5e5c78ba955e
description of safe vs. unsafe wrapper and the functions involved
 oheimb parents: 
2479diff
changeset | 110 | |
| 5550 | 111 | \index{simplification!from classical reasoner} 
 | 
| 112 | Strictly speaking, the operators \texttt{addss} and \texttt{addSss}
 | |
| 113 | are not part of the classical reasoner. | |
| 114 | , which are used as primitives | |
| 11181 
d04f57b91166
renamed addaltern to addafter, addSaltern to addSafter
 oheimb parents: 
9695diff
changeset | 115 | for the automatic tactics described in {\S}\ref{sec:automatic-tactics}, are
 | 
| 5550 | 116 | implemented as wrapper tacticals. | 
| 117 | they | |
| 4881 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 118 | \begin{warn}
 | 
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 119 | Being defined as wrappers, these operators are inappropriate for adding more | 
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 120 | than one simpset at a time: the simpset added last overwrites any earlier ones. | 
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 121 | When a simpset combined with a claset is to be augmented, this should done | 
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 122 | {\em before} combining it with the claset.
 | 
| 
d80faf83c82f
corrected and updated description of wrapper mechanism (including addss)
 oheimb parents: 
4666diff
changeset | 123 | \end{warn}
 | 
| 1099 | 124 | |
| 104 | 125 | |
| 126 | \section{The classical tactics}
 | |
| 5576 
dc6ee60d2be8
exchanged automatic-tactics and semi-automatic-tactics
 oheimb parents: 
5550diff
changeset | 127 | |
| 3224 | 128 | \subsection{Other classical tactics}
 | 
| 332 | 129 | \begin{ttbox} 
 | 
| 875 
a0b71a4bbe5e
documented slow_tac, slow_best_tac, depth_tac, deepen_tac
 lcp parents: 
332diff
changeset | 130 | slow_best_tac : claset -> int -> tactic | 
| 332 | 131 | \end{ttbox}
 | 
| 875 
a0b71a4bbe5e
documented slow_tac, slow_best_tac, depth_tac, deepen_tac
 lcp parents: 
332diff
changeset | 132 | |
| 332 | 133 | \begin{ttdescription}
 | 
| 8136 | 134 | \item[\ttindexbold{slow_best_tac} $cs$ $i$] applies \texttt{slow_step_tac} with
 | 
| 135 | best-first search to prove subgoal~$i$. | |
| 875 
a0b71a4bbe5e
documented slow_tac, slow_best_tac, depth_tac, deepen_tac
 lcp parents: 
332diff
changeset | 136 | \end{ttdescription}
 | 
| 
a0b71a4bbe5e
documented slow_tac, slow_best_tac, depth_tac, deepen_tac
 lcp parents: 
332diff
changeset | 137 | |
| 
a0b71a4bbe5e
documented slow_tac, slow_best_tac, depth_tac, deepen_tac
 lcp parents: 
332diff
changeset | 138 | |
| 104 | 139 | \subsection{Other useful tactics}
 | 
| 319 | 140 | \index{tactics!for contradiction}
 | 
| 141 | \index{tactics!for Modus Ponens}
 | |
| 104 | 142 | \begin{ttbox} 
 | 
| 143 | contr_tac : int -> tactic | |
| 144 | mp_tac : int -> tactic | |
| 145 | eq_mp_tac : int -> tactic | |
| 146 | swap_res_tac : thm list -> int -> tactic | |
| 147 | \end{ttbox}
 | |
| 148 | These can be used in the body of a specialized search. | |
| 308 | 149 | \begin{ttdescription}
 | 
| 319 | 150 | \item[\ttindexbold{contr_tac} {\it i}]\index{assumptions!contradictory}
 | 
| 151 | solves subgoal~$i$ by detecting a contradiction among two assumptions of | |
| 152 | the form $P$ and~$\neg P$, or fail. It may instantiate unknowns. The | |
| 153 | tactic can produce multiple outcomes, enumerating all possible | |
| 154 | contradictions. | |
| 104 | 155 | |
| 156 | \item[\ttindexbold{mp_tac} {\it i}] 
 | |
| 3720 
a5b9e0ade194
Safe_tac, Safe_step_tac, and minor corrections including \tt -> \texttt
 paulson parents: 
3716diff
changeset | 157 | is like \texttt{contr_tac}, but also attempts to perform Modus Ponens in
 | 
| 104 | 158 | subgoal~$i$. If there are assumptions $P\imp Q$ and~$P$, then it replaces | 
| 159 | $P\imp Q$ by~$Q$. It may instantiate unknowns. It fails if it can do | |
| 160 | nothing. | |
| 161 | ||
| 162 | \item[\ttindexbold{eq_mp_tac} {\it i}] 
 | |
| 3720 
a5b9e0ade194
Safe_tac, Safe_step_tac, and minor corrections including \tt -> \texttt
 paulson parents: 
3716diff
changeset | 163 | is like \texttt{mp_tac} {\it i}, but may not instantiate unknowns --- thus, it
 | 
| 104 | 164 | is safe. | 
| 165 | ||
| 166 | \item[\ttindexbold{swap_res_tac} {\it thms} {\it i}] refines subgoal~$i$ of
 | |
| 167 | the proof state using {\it thms}, which should be a list of introduction
 | |
| 3720 
a5b9e0ade194
Safe_tac, Safe_step_tac, and minor corrections including \tt -> \texttt
 paulson parents: 
3716diff
changeset | 168 | rules.  First, it attempts to prove the goal using \texttt{assume_tac} or
 | 
| 
a5b9e0ade194
Safe_tac, Safe_step_tac, and minor corrections including \tt -> \texttt
 paulson parents: 
3716diff
changeset | 169 | \texttt{contr_tac}.  It then attempts to apply each rule in turn, attempting
 | 
| 104 | 170 | resolution and also elim-resolution with the swapped form. | 
| 308 | 171 | \end{ttdescription}
 | 
| 104 | 172 | |
| 173 | ||
| 3716 | 174 | \section{Setting up the classical reasoner}\label{sec:classical-setup}
 | 
| 319 | 175 | \index{classical reasoner!setting up}
 | 
| 5550 | 176 | Isabelle's classical object-logics, including \texttt{FOL} and \texttt{HOL}, 
 | 
| 177 | have the classical reasoner already set up. | |
| 178 | When defining a new classical logic, you should set up the reasoner yourself. | |
| 179 | It consists of the \ML{} functor \ttindex{ClassicalFun}, which takes the 
 | |
| 180 | argument signature \texttt{CLASSICAL_DATA}:
 | |
| 104 | 181 | \begin{ttbox} 
 | 
| 182 | signature CLASSICAL_DATA = | |
| 183 | sig | |
| 184 | val mp : thm | |
| 185 | val not_elim : thm | |
| 186 | val swap : thm | |
| 187 | val sizef : thm -> int | |
| 188 | val hyp_subst_tacs : (int -> tactic) list | |
| 189 | end; | |
| 190 | \end{ttbox}
 | |
| 191 | Thus, the functor requires the following items: | |
| 308 | 192 | \begin{ttdescription}
 | 
| 319 | 193 | \item[\tdxbold{mp}] should be the Modus Ponens rule
 | 
| 104 | 194 | $\List{\Var{P}\imp\Var{Q};\; \Var{P}} \Imp \Var{Q}$.
 | 
| 195 | ||
| 319 | 196 | \item[\tdxbold{not_elim}] should be the contradiction rule
 | 
| 104 | 197 | $\List{\neg\Var{P};\; \Var{P}} \Imp \Var{R}$.
 | 
| 198 | ||
| 319 | 199 | \item[\tdxbold{swap}] should be the swap rule
 | 
| 104 | 200 | $\List{\neg \Var{P}; \; \neg \Var{R}\Imp \Var{P}} \Imp \Var{R}$.
 | 
| 201 | ||
| 202 | \item[\ttindexbold{sizef}] is the heuristic function used for best-first
 | |
| 203 | search. It should estimate the size of the remaining subgoals. A good | |
| 204 | heuristic function is \ttindex{size_of_thm}, which measures the size of the
 | |
| 205 | proof state. Another size function might ignore certain subgoals (say, | |
| 6170 | 206 | those concerned with type-checking). A heuristic function might simply | 
| 104 | 207 | count the subgoals. | 
| 208 | ||
| 319 | 209 | \item[\ttindexbold{hyp_subst_tacs}] is a list of tactics for substitution in
 | 
| 104 | 210 | the hypotheses, typically created by \ttindex{HypsubstFun} (see
 | 
| 211 | Chapter~\ref{substitution}).  This list can, of course, be empty.  The
 | |
| 212 | tactics are assumed to be safe! | |
| 308 | 213 | \end{ttdescription}
 | 
| 104 | 214 | The functor is not at all sensitive to the formalization of the | 
| 3108 | 215 | object-logic. It does not even examine the rules, but merely applies | 
| 216 | them according to its fixed strategy.  The functor resides in {\tt
 | |
| 217 | Provers/classical.ML} in the Isabelle sources. | |
| 104 | 218 | |
| 319 | 219 | \index{classical reasoner|)}
 | 
| 5371 | 220 | |
| 221 | %%% Local Variables: | |
| 222 | %%% mode: latex | |
| 223 | %%% TeX-master: "ref" | |
| 224 | %%% End: |