repaired critical proofs depending on the order inside non-confluent SimpSets,
(temporarily) removed problematic rule less_Suc_eq form simpset_of "Nat"
%% $Id$\chapter{The Classical Reasoner}\label{chap:classical}\index{classical reasoner|(}\newcommand\ainfer[2]{\begin{array}{r@{\,}l}#2\\ \hline#1\end{array}}Although Isabelle is generic, many users will be working in some extensionof classical first-order logic. Isabelle's set theory~{\tt ZF} is builtupon theory~{\tt FOL}, while higher-order logic contains first-order logicas a fragment. Theorem-proving in predicate logic is undecidable, but manyresearchers have developed strategies to assist in this task.Isabelle's classical reasoner is an \ML{} functor that accepts certaininformation about a logic and delivers a suite of automatic tactics. Eachtactic takes a collection of rules and executes a simple, non-clausal proofprocedure. They are slow and simplistic compared with resolution theoremprovers, but they can save considerable time and effort. They can provetheorems such as Pelletier's~\cite{pelletier86} problems~40 and~41 inseconds:\[ (\exists y. \forall x. J(y,x) \bimp \neg J(x,x)) \imp \neg (\forall x. \exists y. \forall z. J(z,y) \bimp \neg J(z,x)) \]\[ (\forall z. \exists y. \forall x. F(x,y) \bimp F(x,z) \conj \neg F(x,x)) \imp \neg (\exists z. \forall x. F(x,z)) \]%The tactics are generic. They are not restricted to first-order logic, andhave been heavily used in the development of Isabelle's set theory. Fewinteractive proof assistants provide this much automation. The tactics canbe traced, and their components can be called directly; in this manner,any proof can be viewed interactively.We shall first discuss the underlying principles, then consider how to usethe classical reasoner. Finally, we shall see how to instantiate it fornew logics. The logics {\tt FOL}, {\tt HOL} and {\tt ZF} have it alreadyinstalled.\section{The sequent calculus}\index{sequent calculus}Isabelle supports natural deduction, which is easy to use for interactiveproof. But natural deduction does not easily lend itself to automation,and has a bias towards intuitionism. For certain proofs in classicallogic, it can not be called natural. The {\bf sequent calculus}, ageneralization of natural deduction, is easier to automate.A {\bf sequent} has the form $\Gamma\turn\Delta$, where $\Gamma$and~$\Delta$ are sets of formulae.%\footnote{For first-order logic, sequents can equivalently be made from lists or multisets of formulae.} The sequent\[ P@1,\ldots,P@m\turn Q@1,\ldots,Q@n \]is {\bf valid} if $P@1\conj\ldots\conj P@m$ implies $Q@1\disj\ldots\disjQ@n$. Thus $P@1,\ldots,P@m$ represent assumptions, each of which is true,while $Q@1,\ldots,Q@n$ represent alternative goals. A sequent is {\bfbasic} if its left and right sides have a common formula, as in $P,Q\turnQ,R$; basic sequents are trivially valid.Sequent rules are classified as {\bf right} or {\bf left}, indicating whichside of the $\turn$~symbol they operate on. Rules that operate on theright side are analogous to natural deduction's introduction rules, andleft rules are analogous to elimination rules. Recall the natural deduction rules for first-order logic, \iflabelundefined{fol-fig}{from {\it Introduction to Isabelle}}% {Fig.\ts\ref{fol-fig}}.The sequent calculus analogue of~$({\imp}I)$ is the rule$$ \ainfer{\Gamma &\turn \Delta, P\imp Q}{P,\Gamma &\turn \Delta,Q} \eqno({\imp}R) $$This breaks down some implication on the right side of a sequent; $\Gamma$and $\Delta$ stand for the sets of formulae that are unaffected by theinference. The analogue of the pair~$({\disj}I1)$ and~$({\disj}I2)$ is thesingle rule $$ \ainfer{\Gamma &\turn \Delta, P\disj Q}{\Gamma &\turn \Delta,P,Q} \eqno({\disj}R) $$This breaks down some disjunction on the right side, replacing it by bothdisjuncts. Thus, the sequent calculus is a kind of multiple-conclusion logic.To illustrate the use of multiple formulae on the right, let us provethe classical theorem $(P\imp Q)\disj(Q\imp P)$. Working backwards, wereduce this formula to a basic sequent:\[ \infer[(\disj)R]{\turn(P\imp Q)\disj(Q\imp P)} {\infer[(\imp)R]{\turn(P\imp Q), (Q\imp P)\;} {\infer[(\imp)R]{P \turn Q, (Q\imp P)\qquad} {P, Q \turn Q, P\qquad\qquad}}}\]This example is typical of the sequent calculus: start with the desiredtheorem and apply rules backwards in a fairly arbitrary manner. This yields asurprisingly effective proof procedure. Quantifiers add few complications,since Isabelle handles parameters and schematic variables. See Chapter~10of {\em ML for the Working Programmer}~\cite{paulson91} for furtherdiscussion.\section{Simulating sequents by natural deduction}Isabelle can represent sequents directly, as in the object-logic~{\tt LK}\@.But natural deduction is easier to work with, and most object-logics employit. Fortunately, we can simulate the sequent $P@1,\ldots,P@m\turnQ@1,\ldots,Q@n$ by the Isabelle formula\[ \List{P@1;\ldots;P@m; \neg Q@2;\ldots; \neg Q@n}\Imp Q@1, \]where the order of the assumptions and the choice of~$Q@1$ are arbitrary.Elim-resolution plays a key role in simulating sequent proofs.We can easily handle reasoning on the left.As discussed in\iflabelundefined{destruct}{{\it Introduction to Isabelle}}{\S\ref{destruct}}, elim-resolution with the rules $(\disj E)$, $(\bot E)$ and $(\exists E)$achieves a similar effect as the corresponding sequent rules. For theother connectives, we use sequent-style elimination rules instead ofdestruction rules such as $({\conj}E1,2)$ and $(\forall E)$. But note thatthe rule $(\neg L)$ has no effect under our representation of sequents!$$ \ainfer{\neg P,\Gamma &\turn \Delta}{\Gamma &\turn \Delta,P} \eqno({\neg}L) $$What about reasoning on the right? Introduction rules can only affect theformula in the conclusion, namely~$Q@1$. The other right-side formulae arerepresented as negated assumptions, $\neg Q@2$, \ldots,~$\neg Q@n$. \index{assumptions!negated}In order to operate on one of these, it must first be exchanged with~$Q@1$.Elim-resolution with the {\bf swap} rule has this effect:$$ \List{\neg P; \; \neg R\Imp P} \Imp R \eqno(swap)$$To ensure that swaps occur only when necessary, each introduction rule isconverted into a swapped form: it is resolved with the second premiseof~$(swap)$. The swapped form of~$({\conj}I)$, which might becalled~$({\neg\conj}E)$, is\[ \List{\neg(P\conj Q); \; \neg R\Imp P; \; \neg R\Imp Q} \Imp R. \]Similarly, the swapped form of~$({\imp}I)$ is\[ \List{\neg(P\imp Q); \; \List{\neg R;P}\Imp Q} \Imp R \]Swapped introduction rules are applied using elim-resolution, which deletesthe negated formula. Our representation of sequents also requires the useof ordinary introduction rules. If we had no regard for readability, wecould treat the right side more uniformly by representing sequents as\[ \List{P@1;\ldots;P@m; \neg Q@1;\ldots; \neg Q@n}\Imp \bot. \]\section{Extra rules for the sequent calculus}As mentioned, destruction rules such as $({\conj}E1,2)$ and $(\forall E)$must be replaced by sequent-style elimination rules. In addition, we needrules to embody the classical equivalence between $P\imp Q$ and $\neg P\disjQ$. The introduction rules~$({\disj}I1,2)$ are replaced by a rule thatsimulates $({\disj}R)$:\[ (\neg Q\Imp P) \Imp P\disj Q \]The destruction rule $({\imp}E)$ is replaced by\[ \List{P\imp Q;\; \neg P\Imp R;\; Q\Imp R} \Imp R. \]Quantifier replication also requires special rules. In classical logic,$\exists x{.}P$ is equivalent to $\neg\forall x{.}\neg P$; the rules$(\exists R)$ and $(\forall L)$ are dual:\[ \ainfer{\Gamma &\turn \Delta, \exists x{.}P} {\Gamma &\turn \Delta, \exists x{.}P, P[t/x]} \; (\exists R) \qquad \ainfer{\forall x{.}P, \Gamma &\turn \Delta} {P[t/x], \forall x{.}P, \Gamma &\turn \Delta} \; (\forall L)\]Thus both kinds of quantifier may be replicated. Theorems requiringmultiple uses of a universal formula are easy to invent; consider \[ (\forall x.P(x)\imp P(f(x))) \conj P(a) \imp P(f^n(a)), \]for any~$n>1$. Natural examples of the multiple use of an existentialformula are rare; a standard one is $\exists x.\forall y. P(x)\imp P(y)$.Forgoing quantifier replication loses completeness, but gains decidability,since the search space becomes finite. Many useful theorems can be provedwithout replication, and the search generally delivers its verdict in areasonable time. To adopt this approach, represent the sequent rules$(\exists R)$, $(\exists L)$ and $(\forall R)$ by $(\exists I)$, $(\existsE)$ and $(\forall I)$, respectively, and put $(\forall E)$ into eliminationform:$$ \List{\forall x{.}P(x); P(t)\Imp Q} \Imp Q \eqno(\forall E@2) $$Elim-resolution with this rule will delete the universal formula after asingle use. To replicate universal quantifiers, replace the rule by$$ \List{\forall x{.}P(x);\; \List{P(t); \forall x{.}P(x)}\Imp Q} \Imp Q. \eqno(\forall E@3) $$To replicate existential quantifiers, replace $(\exists I)$ by\[ \List{\neg(\exists x{.}P(x)) \Imp P(t)} \Imp \exists x{.}P(x). \]All introduction rules mentioned above are also useful in swapped form.Replication makes the search space infinite; we must apply the rules withcare. The classical reasoner distinguishes between safe and unsaferules, applying the latter only when there is no alternative. Depth-firstsearch may well go down a blind alley; best-first search is better behavedin an infinite search space. However, quantifier replication is tooexpensive to prove any but the simplest theorems.\section{Classical rule sets}\index{classical sets}Each automatic tactic takes a {\bf classical set} --- a collection ofrules, classified as introduction or elimination and as {\bf safe} or {\bfunsafe}. In general, safe rules can be attempted blindly, while unsaferules must be used with care. A safe rule must never reduce a provablegoal to an unprovable set of subgoals. The rule~$({\disj}I1)$ is unsafe because it reduces $P\disj Q$ to~$P$. Anyrule is unsafe whose premises contain new unknowns. The eliminationrule~$(\forall E@2)$ is unsafe, since it is applied via elim-resolution,which discards the assumption $\forall x{.}P(x)$ and replaces it by theweaker assumption~$P(\Var{t})$. The rule $({\exists}I)$ is unsafe forsimilar reasons. The rule~$(\forall E@3)$ is unsafe in a different sense:since it keeps the assumption $\forall x{.}P(x)$, it is prone to looping.In classical first-order logic, all rules are safe except those mentionedabove.The safe/unsafe distinction is vague, and may be regarded merely as a wayof giving some rules priority over others. One could argue that$({\disj}E)$ is unsafe, because repeated application of it could generateexponentially many subgoals. Induction rules are unsafe because inductiveproofs are difficult to set up automatically. Any inference is unsafe thatinstantiates an unknown in the proof state --- thus \ttindex{match_tac}must be used, rather than \ttindex{resolve_tac}. Even proof by assumptionis unsafe if it instantiates unknowns shared with other subgoals --- thus\ttindex{eq_assume_tac} must be used, rather than \ttindex{assume_tac}.\subsection{Adding rules to classical sets}Classical rule sets belong to the abstract type \mltydx{claset}, whichsupports the following operations (provided the classical reasoner isinstalled!):\begin{ttbox} empty_cs : clasetprint_cs : claset -> unitaddSIs : claset * thm list -> claset \hfill{\bf infix 4}addSEs : claset * thm list -> claset \hfill{\bf infix 4}addSDs : claset * thm list -> claset \hfill{\bf infix 4}addIs : claset * thm list -> claset \hfill{\bf infix 4}addEs : claset * thm list -> claset \hfill{\bf infix 4}addDs : claset * thm list -> claset \hfill{\bf infix 4}\end{ttbox}There are no operations for deletion from a classical set. The addoperations do not check for repetitions.\begin{ttdescription}\item[\ttindexbold{empty_cs}] is the empty classical set.\item[\ttindexbold{print_cs} $cs$] prints the rules of~$cs$.\item[$cs$ addSIs $rules$] \indexbold{*addSIs}adds safe introduction~$rules$ to~$cs$.\item[$cs$ addSEs $rules$] \indexbold{*addSEs}adds safe elimination~$rules$ to~$cs$.\item[$cs$ addSDs $rules$] \indexbold{*addSDs}adds safe destruction~$rules$ to~$cs$.\item[$cs$ addIs $rules$] \indexbold{*addIs}adds unsafe introduction~$rules$ to~$cs$.\item[$cs$ addEs $rules$] \indexbold{*addEs}adds unsafe elimination~$rules$ to~$cs$.\item[$cs$ addDs $rules$] \indexbold{*addDs}adds unsafe destruction~$rules$ to~$cs$.\end{ttdescription}Introduction rules are those that can be applied using ordinary resolution.The classical set automatically generates their swapped forms, which willbe applied using elim-resolution. Elimination rules are applied usingelim-resolution. In a classical set, rules are sorted by the number of newsubgoals they will yield; rules that generate the fewest subgoals will betried first (see \S\ref{biresolve_tac}).\subsection{Modifying the search step}For a given classical set, the proof strategy is simple. Perform as manysafe inferences as possible; or else, apply certain safe rules, allowinginstantiation of unknowns; or else, apply an unsafe rule. The tactics mayalso apply {\tt hyp_subst_tac}, if they have been set up to do so (seebelow). They may perform a form of Modus Ponens: if there are assumptions$P\imp Q$ and~$P$, then replace $P\imp Q$ by~$Q$.The classical reasoner allows you to modify this basic proof strategy byapplying an arbitrary {\bf wrapper tactical} to it. This affects each step ofthe search. Usually it is the identity tactical, but it could apply anothertactic before or after the step tactic. It affects \ttindex{step_tac},\ttindex{slow_step_tac} and the tactics that call them.\begin{ttbox} addss : claset * simpset -> claset \hfill{\bf infix 4}addbefore : claset * tactic -> claset \hfill{\bf infix 4}addafter : claset * tactic -> claset \hfill{\bf infix 4}setwrapper : claset * (tactic->tactic) -> claset \hfill{\bf infix 4}compwrapper : claset * (tactic->tactic) -> claset \hfill{\bf infix 4}\end{ttbox}%\index{simplification!from classical reasoner} The wrapper tactical underlies the operator \ttindex{addss}, which precedeseach search step by simplification. Strictly speaking, {\tt addss} is notpart of the classical reasoner. It should be defined (using {\tt addbefore})when the simplifier is installed.\begin{ttdescription}\item[$cs$ addss $ss$] \indexbold{*addss}adds the simpset~$ss$ to the classical set. The assumptions and goal will besimplified before each step of the search.\item[$cs$ addbefore $tac$] \indexbold{*addbefore}changes the wrapper tactical to apply the given tactic {\em before}each step of the search.\item[$cs$ addafter $tac$] \indexbold{*addafter}changes the wrapper tactical to apply the given tactic {\em after}each step of the search.\item[$cs$ setwrapper $tactical$] \indexbold{*setwrapper}specifies a new wrapper tactical. \item[$cs$ compwrapper $tactical$] \indexbold{*compwrapper}composes the $tactical$ with the existing wrapper tactical, to combine theireffects. \end{ttdescription}\section{The classical tactics}\index{classical reasoner!tactics}If installed, the classical module provides several tactics (and otheroperations) for simulating the classical sequent calculus.\subsection{The automatic tactics}\begin{ttbox} fast_tac : claset -> int -> tacticbest_tac : claset -> int -> tacticslow_tac : claset -> int -> tacticslow_best_tac : claset -> int -> tactic\end{ttbox}These tactics work by applying {\tt step_tac} or {\tt slow_step_tac}repeatedly. Their effect is restricted (by {\tt SELECT_GOAL}) to one subgoal;they either solve this subgoal or fail. The {\tt slow_} versions are morepowerful but can be much slower. The best-first tactics are guided by a heuristic function: typically, thetotal size of the proof state. This function is supplied in the functor callthat sets up the classical reasoner.\begin{ttdescription}\item[\ttindexbold{fast_tac} $cs$ $i$] applies {\tt step_tac} usingdepth-first search, to solve subgoal~$i$.\item[\ttindexbold{best_tac} $cs$ $i$] applies {\tt step_tac} usingbest-first search, to solve subgoal~$i$.\item[\ttindexbold{slow_tac} $cs$ $i$] applies {\tt slow_step_tac} usingdepth-first search, to solve subgoal~$i$.\item[\ttindexbold{slow_best_tac} $cs$ $i$] applies {\tt slow_step_tac} usingbest-first search, to solve subgoal~$i$.\end{ttdescription}\subsection{Depth-limited tactics}\begin{ttbox} depth_tac : claset -> int -> int -> tacticdeepen_tac : claset -> int -> int -> tactic\end{ttbox}These work by exhaustive search up to a specified depth. Unsafe rules aremodified to preserve the formula they act on, so that it be used repeatedly.They can prove more goals than {\tt fast_tac} can but are muchslower, for example if the assumptions have many universal quantifiers.The depth limits the number of unsafe steps. If you can estimate the minimumnumber of unsafe steps needed, supply this value as~$m$ to save time.\begin{ttdescription}\item[\ttindexbold{depth_tac} $cs$ $m$ $i$] tries to solve subgoal~$i$ by exhaustive search up to depth~$m$.\item[\ttindexbold{deepen_tac} $cs$ $m$ $i$] tries to solve subgoal~$i$ by iterative deepening. It calls {\tt depth_tac}repeatedly with increasing depths, starting with~$m$.\end{ttdescription}\subsection{Single-step tactics}\begin{ttbox} safe_step_tac : claset -> int -> tacticsafe_tac : claset -> tacticinst_step_tac : claset -> int -> tacticstep_tac : claset -> int -> tacticslow_step_tac : claset -> int -> tactic\end{ttbox}The automatic proof procedures call these tactics. By calling themyourself, you can execute these procedures one step at a time.\begin{ttdescription}\item[\ttindexbold{safe_step_tac} $cs$ $i$] performs a safe step onsubgoal~$i$. This may include proof by assumption or Modus Ponens (takingcare not to instantiate unknowns), or {\tt hyp_subst_tac}.\item[\ttindexbold{safe_tac} $cs$] repeatedly performs safe steps on all subgoals. It is deterministic, with at most one outcome. If the automatictactics fail, try using {\tt safe_tac} to open up your formula; then youcan replicate certain quantifiers explicitly by applying appropriate rules.\item[\ttindexbold{inst_step_tac} $cs$ $i$] is like {\tt safe_step_tac},but allows unknowns to be instantiated.\item[\ttindexbold{step_tac} $cs$ $i$] is the basic step of the proof procedure. The wrapper tactical is applied to a tactic that tries {\tt safe_tac}, {\tt inst_step_tac}, or applies an unsafe rule from~$cs$.\item[\ttindexbold{slow_step_tac}] resembles {\tt step_tac}, but allows backtracking between using safe rules with instantiation ({\tt inst_step_tac}) and using unsafe rules. The resulting search space is larger.\end{ttdescription}\subsection{Other useful tactics}\index{tactics!for contradiction}\index{tactics!for Modus Ponens}\begin{ttbox} contr_tac : int -> tacticmp_tac : int -> tacticeq_mp_tac : int -> tacticswap_res_tac : thm list -> int -> tactic\end{ttbox}These can be used in the body of a specialized search.\begin{ttdescription}\item[\ttindexbold{contr_tac} {\it i}]\index{assumptions!contradictory} solves subgoal~$i$ by detecting a contradiction among two assumptions of the form $P$ and~$\neg P$, or fail. It may instantiate unknowns. The tactic can produce multiple outcomes, enumerating all possible contradictions.\item[\ttindexbold{mp_tac} {\it i}] is like {\tt contr_tac}, but also attempts to perform Modus Ponens insubgoal~$i$. If there are assumptions $P\imp Q$ and~$P$, then it replaces$P\imp Q$ by~$Q$. It may instantiate unknowns. It fails if it can donothing.\item[\ttindexbold{eq_mp_tac} {\it i}] is like {\tt mp_tac} {\it i}, but may not instantiate unknowns --- thus, itis safe.\item[\ttindexbold{swap_res_tac} {\it thms} {\it i}] refines subgoal~$i$ ofthe proof state using {\it thms}, which should be a list of introductionrules. First, it attempts to solve the goal using {\tt assume_tac} or{\tt contr_tac}. It then attempts to apply each rule in turn, attemptingresolution and also elim-resolution with the swapped form.\end{ttdescription}\subsection{Creating swapped rules}\begin{ttbox} swapify : thm list -> thm listjoinrules : thm list * thm list -> (bool * thm) list\end{ttbox}\begin{ttdescription}\item[\ttindexbold{swapify} {\it thms}] returns a list consisting of theswapped versions of~{\it thms}, regarded as introduction rules.\item[\ttindexbold{joinrules} ({\it intrs}, {\it elims})]joins introduction rules, their swapped versions, and elimination rules foruse with \ttindex{biresolve_tac}. Each rule is paired with~{\tt false}(indicating ordinary resolution) or~{\tt true} (indicatingelim-resolution).\end{ttdescription}\section{Setting up the classical reasoner}\index{classical reasoner!setting up}Isabelle's classical object-logics, including {\tt FOL} and {\tt HOL}, havethe classical reasoner already set up. When defining a new classical logic,you should set up the reasoner yourself. It consists of the \ML{} functor\ttindex{ClassicalFun}, which takes the argumentsignature {\tt CLASSICAL_DATA}:\begin{ttbox} signature CLASSICAL_DATA = sig val mp : thm val not_elim : thm val swap : thm val sizef : thm -> int val hyp_subst_tacs : (int -> tactic) list end;\end{ttbox}Thus, the functor requires the following items:\begin{ttdescription}\item[\tdxbold{mp}] should be the Modus Ponens rule$\List{\Var{P}\imp\Var{Q};\; \Var{P}} \Imp \Var{Q}$.\item[\tdxbold{not_elim}] should be the contradiction rule$\List{\neg\Var{P};\; \Var{P}} \Imp \Var{R}$.\item[\tdxbold{swap}] should be the swap rule$\List{\neg \Var{P}; \; \neg \Var{R}\Imp \Var{P}} \Imp \Var{R}$.\item[\ttindexbold{sizef}] is the heuristic function used for best-firstsearch. It should estimate the size of the remaining subgoals. A goodheuristic function is \ttindex{size_of_thm}, which measures the size of theproof state. Another size function might ignore certain subgoals (say,those concerned with type checking). A heuristic function might simplycount the subgoals.\item[\ttindexbold{hyp_subst_tacs}] is a list of tactics for substitution inthe hypotheses, typically created by \ttindex{HypsubstFun} (seeChapter~\ref{substitution}). This list can, of course, be empty. Thetactics are assumed to be safe!\end{ttdescription}The functor is not at all sensitive to the formalization of theobject-logic. It does not even examine the rules, but merely applies themaccording to its fixed strategy. The functor resides in {\ttProvers/classical.ML} in the Isabelle distribution directory.\index{classical reasoner|)}