104
|
1 |
%% $Id$
|
313
|
2 |
\chapter{First-Order Logic}
|
|
3 |
\index{first-order logic|(}
|
|
4 |
|
|
5 |
Isabelle implements Gentzen's natural deduction systems {\sc nj} and {\sc
|
|
6 |
nk}. Intuitionistic first-order logic is defined first, as theory
|
|
7 |
\thydx{IFOL}. Classical logic, theory \thydx{FOL}, is
|
104
|
8 |
obtained by adding the double negation rule. Basic proof procedures are
|
|
9 |
provided. The intuitionistic prover works with derived rules to simplify
|
313
|
10 |
implications in the assumptions. Classical~{\tt FOL} employs Isabelle's
|
|
11 |
classical reasoner, which simulates a sequent calculus.
|
104
|
12 |
|
|
13 |
\section{Syntax and rules of inference}
|
313
|
14 |
The logic is many-sorted, using Isabelle's type classes. The class of
|
|
15 |
first-order terms is called \cldx{term} and is a subclass of {\tt logic}.
|
|
16 |
No types of individuals are provided, but extensions can define types such
|
|
17 |
as {\tt nat::term} and type constructors such as {\tt list::(term)term}
|
|
18 |
(see the examples directory, {\tt FOL/ex}). Below, the type variable
|
|
19 |
$\alpha$ ranges over class {\tt term}; the equality symbol and quantifiers
|
|
20 |
are polymorphic (many-sorted). The type of formulae is~\tydx{o}, which
|
|
21 |
belongs to class~\cldx{logic}. Figure~\ref{fol-syntax} gives the syntax.
|
|
22 |
Note that $a$\verb|~=|$b$ is translated to $\neg(a=b)$.
|
104
|
23 |
|
313
|
24 |
Figure~\ref{fol-rules} shows the inference rules with their~\ML\ names.
|
|
25 |
Negation is defined in the usual way for intuitionistic logic; $\neg P$
|
|
26 |
abbreviates $P\imp\bot$. The biconditional~($\bimp$) is defined through
|
|
27 |
$\conj$ and~$\imp$; introduction and elimination rules are derived for it.
|
104
|
28 |
|
|
29 |
The unique existence quantifier, $\exists!x.P(x)$, is defined in terms
|
|
30 |
of~$\exists$ and~$\forall$. An Isabelle binder, it admits nested
|
3138
|
31 |
quantifications. For instance, $\exists!x\;y.P(x,y)$ abbreviates
|
104
|
32 |
$\exists!x. \exists!y.P(x,y)$; note that this does not mean that there
|
|
33 |
exists a unique pair $(x,y)$ satisfying~$P(x,y)$.
|
|
34 |
|
|
35 |
Some intuitionistic derived rules are shown in
|
287
|
36 |
Fig.\ts\ref{fol-int-derived}, again with their \ML\ names. These include
|
104
|
37 |
rules for the defined symbols $\neg$, $\bimp$ and $\exists!$. Natural
|
333
|
38 |
deduction typically involves a combination of forward and backward
|
104
|
39 |
reasoning, particularly with the destruction rules $(\conj E)$,
|
333
|
40 |
$({\imp}E)$, and~$(\forall E)$. Isabelle's backward style handles these
|
104
|
41 |
rules badly, so sequent-style rules are derived to eliminate conjunctions,
|
|
42 |
implications, and universal quantifiers. Used with elim-resolution,
|
313
|
43 |
\tdx{allE} eliminates a universal quantifier while \tdx{all_dupE}
|
104
|
44 |
re-inserts the quantified formula for later use. The rules {\tt
|
|
45 |
conj_impE}, etc., support the intuitionistic proof procedure
|
|
46 |
(see~\S\ref{fol-int-prover}).
|
|
47 |
|
333
|
48 |
See the files {\tt FOL/IFOL.thy}, {\tt FOL/IFOL.ML} and
|
287
|
49 |
{\tt FOL/intprover.ML} for complete listings of the rules and
|
104
|
50 |
derived rules.
|
|
51 |
|
|
52 |
\begin{figure}
|
|
53 |
\begin{center}
|
|
54 |
\begin{tabular}{rrr}
|
111
|
55 |
\it name &\it meta-type & \it description \\
|
313
|
56 |
\cdx{Trueprop}& $o\To prop$ & coercion to $prop$\\
|
|
57 |
\cdx{Not} & $o\To o$ & negation ($\neg$) \\
|
|
58 |
\cdx{True} & $o$ & tautology ($\top$) \\
|
|
59 |
\cdx{False} & $o$ & absurdity ($\bot$)
|
104
|
60 |
\end{tabular}
|
|
61 |
\end{center}
|
|
62 |
\subcaption{Constants}
|
|
63 |
|
|
64 |
\begin{center}
|
|
65 |
\begin{tabular}{llrrr}
|
313
|
66 |
\it symbol &\it name &\it meta-type & \it priority & \it description \\
|
|
67 |
\sdx{ALL} & \cdx{All} & $(\alpha\To o)\To o$ & 10 &
|
111
|
68 |
universal quantifier ($\forall$) \\
|
313
|
69 |
\sdx{EX} & \cdx{Ex} & $(\alpha\To o)\To o$ & 10 &
|
111
|
70 |
existential quantifier ($\exists$) \\
|
313
|
71 |
{\tt EX!} & \cdx{Ex1} & $(\alpha\To o)\To o$ & 10 &
|
111
|
72 |
unique existence ($\exists!$)
|
104
|
73 |
\end{tabular}
|
313
|
74 |
\index{*"E"X"! symbol}
|
104
|
75 |
\end{center}
|
|
76 |
\subcaption{Binders}
|
|
77 |
|
|
78 |
\begin{center}
|
313
|
79 |
\index{*"= symbol}
|
|
80 |
\index{&@{\tt\&} symbol}
|
|
81 |
\index{*"| symbol}
|
|
82 |
\index{*"-"-"> symbol}
|
|
83 |
\index{*"<"-"> symbol}
|
104
|
84 |
\begin{tabular}{rrrr}
|
313
|
85 |
\it symbol & \it meta-type & \it priority & \it description \\
|
111
|
86 |
\tt = & $[\alpha,\alpha]\To o$ & Left 50 & equality ($=$) \\
|
|
87 |
\tt \& & $[o,o]\To o$ & Right 35 & conjunction ($\conj$) \\
|
|
88 |
\tt | & $[o,o]\To o$ & Right 30 & disjunction ($\disj$) \\
|
|
89 |
\tt --> & $[o,o]\To o$ & Right 25 & implication ($\imp$) \\
|
|
90 |
\tt <-> & $[o,o]\To o$ & Right 25 & biconditional ($\bimp$)
|
104
|
91 |
\end{tabular}
|
|
92 |
\end{center}
|
|
93 |
\subcaption{Infixes}
|
|
94 |
|
|
95 |
\dquotes
|
|
96 |
\[\begin{array}{rcl}
|
|
97 |
formula & = & \hbox{expression of type~$o$} \\
|
111
|
98 |
& | & term " = " term \\
|
|
99 |
& | & term " \ttilde= " term \\
|
|
100 |
& | & "\ttilde\ " formula \\
|
|
101 |
& | & formula " \& " formula \\
|
|
102 |
& | & formula " | " formula \\
|
|
103 |
& | & formula " --> " formula \\
|
|
104 |
& | & formula " <-> " formula \\
|
|
105 |
& | & "ALL~" id~id^* " . " formula \\
|
|
106 |
& | & "EX~~" id~id^* " . " formula \\
|
|
107 |
& | & "EX!~" id~id^* " . " formula
|
104
|
108 |
\end{array}
|
|
109 |
\]
|
|
110 |
\subcaption{Grammar}
|
|
111 |
\caption{Syntax of {\tt FOL}} \label{fol-syntax}
|
|
112 |
\end{figure}
|
|
113 |
|
|
114 |
|
|
115 |
\begin{figure}
|
|
116 |
\begin{ttbox}
|
313
|
117 |
\tdx{refl} a=a
|
|
118 |
\tdx{subst} [| a=b; P(a) |] ==> P(b)
|
104
|
119 |
\subcaption{Equality rules}
|
|
120 |
|
313
|
121 |
\tdx{conjI} [| P; Q |] ==> P&Q
|
|
122 |
\tdx{conjunct1} P&Q ==> P
|
|
123 |
\tdx{conjunct2} P&Q ==> Q
|
104
|
124 |
|
313
|
125 |
\tdx{disjI1} P ==> P|Q
|
|
126 |
\tdx{disjI2} Q ==> P|Q
|
|
127 |
\tdx{disjE} [| P|Q; P ==> R; Q ==> R |] ==> R
|
104
|
128 |
|
313
|
129 |
\tdx{impI} (P ==> Q) ==> P-->Q
|
|
130 |
\tdx{mp} [| P-->Q; P |] ==> Q
|
104
|
131 |
|
313
|
132 |
\tdx{FalseE} False ==> P
|
104
|
133 |
\subcaption{Propositional rules}
|
|
134 |
|
313
|
135 |
\tdx{allI} (!!x. P(x)) ==> (ALL x.P(x))
|
|
136 |
\tdx{spec} (ALL x.P(x)) ==> P(x)
|
104
|
137 |
|
313
|
138 |
\tdx{exI} P(x) ==> (EX x.P(x))
|
|
139 |
\tdx{exE} [| EX x.P(x); !!x. P(x) ==> R |] ==> R
|
104
|
140 |
\subcaption{Quantifier rules}
|
|
141 |
|
313
|
142 |
\tdx{True_def} True == False-->False
|
|
143 |
\tdx{not_def} ~P == P-->False
|
|
144 |
\tdx{iff_def} P<->Q == (P-->Q) & (Q-->P)
|
|
145 |
\tdx{ex1_def} EX! x. P(x) == EX x. P(x) & (ALL y. P(y) --> y=x)
|
104
|
146 |
\subcaption{Definitions}
|
|
147 |
\end{ttbox}
|
|
148 |
|
313
|
149 |
\caption{Rules of intuitionistic logic} \label{fol-rules}
|
104
|
150 |
\end{figure}
|
|
151 |
|
|
152 |
|
|
153 |
\begin{figure}
|
|
154 |
\begin{ttbox}
|
313
|
155 |
\tdx{sym} a=b ==> b=a
|
|
156 |
\tdx{trans} [| a=b; b=c |] ==> a=c
|
|
157 |
\tdx{ssubst} [| b=a; P(a) |] ==> P(b)
|
104
|
158 |
\subcaption{Derived equality rules}
|
|
159 |
|
313
|
160 |
\tdx{TrueI} True
|
104
|
161 |
|
313
|
162 |
\tdx{notI} (P ==> False) ==> ~P
|
|
163 |
\tdx{notE} [| ~P; P |] ==> R
|
104
|
164 |
|
313
|
165 |
\tdx{iffI} [| P ==> Q; Q ==> P |] ==> P<->Q
|
|
166 |
\tdx{iffE} [| P <-> Q; [| P-->Q; Q-->P |] ==> R |] ==> R
|
|
167 |
\tdx{iffD1} [| P <-> Q; P |] ==> Q
|
|
168 |
\tdx{iffD2} [| P <-> Q; Q |] ==> P
|
104
|
169 |
|
313
|
170 |
\tdx{ex1I} [| P(a); !!x. P(x) ==> x=a |] ==> EX! x. P(x)
|
|
171 |
\tdx{ex1E} [| EX! x.P(x); !!x.[| P(x); ALL y. P(y) --> y=x |] ==> R
|
104
|
172 |
|] ==> R
|
|
173 |
\subcaption{Derived rules for \(\top\), \(\neg\), \(\bimp\) and \(\exists!\)}
|
|
174 |
|
313
|
175 |
\tdx{conjE} [| P&Q; [| P; Q |] ==> R |] ==> R
|
|
176 |
\tdx{impE} [| P-->Q; P; Q ==> R |] ==> R
|
|
177 |
\tdx{allE} [| ALL x.P(x); P(x) ==> R |] ==> R
|
|
178 |
\tdx{all_dupE} [| ALL x.P(x); [| P(x); ALL x.P(x) |] ==> R |] ==> R
|
104
|
179 |
\subcaption{Sequent-style elimination rules}
|
|
180 |
|
313
|
181 |
\tdx{conj_impE} [| (P&Q)-->S; P-->(Q-->S) ==> R |] ==> R
|
|
182 |
\tdx{disj_impE} [| (P|Q)-->S; [| P-->S; Q-->S |] ==> R |] ==> R
|
|
183 |
\tdx{imp_impE} [| (P-->Q)-->S; [| P; Q-->S |] ==> Q; S ==> R |] ==> R
|
|
184 |
\tdx{not_impE} [| ~P --> S; P ==> False; S ==> R |] ==> R
|
|
185 |
\tdx{iff_impE} [| (P<->Q)-->S; [| P; Q-->S |] ==> Q; [| Q; P-->S |] ==> P;
|
104
|
186 |
S ==> R |] ==> R
|
313
|
187 |
\tdx{all_impE} [| (ALL x.P(x))-->S; !!x.P(x); S ==> R |] ==> R
|
|
188 |
\tdx{ex_impE} [| (EX x.P(x))-->S; P(a)-->S ==> R |] ==> R
|
104
|
189 |
\end{ttbox}
|
|
190 |
\subcaption{Intuitionistic simplification of implication}
|
313
|
191 |
\caption{Derived rules for intuitionistic logic} \label{fol-int-derived}
|
104
|
192 |
\end{figure}
|
|
193 |
|
|
194 |
|
|
195 |
\section{Generic packages}
|
2495
|
196 |
\FOL{} instantiates most of Isabelle's generic packages.
|
104
|
197 |
\begin{itemize}
|
|
198 |
\item
|
2495
|
199 |
It instantiates the simplifier. Both equality ($=$) and the biconditional
|
|
200 |
($\bimp$) may be used for rewriting. Tactics such as {\tt Asm_simp_tac} and
|
|
201 |
{\tt Full_simp_tac} use the default simpset ({\tt!simpset}), which works for
|
|
202 |
most purposes. Named simplification sets include \ttindexbold{IFOL_ss},
|
|
203 |
for intuitionistic first-order logic, and \ttindexbold{FOL_ss},
|
|
204 |
for classical logic. See the file
|
287
|
205 |
{\tt FOL/simpdata.ML} for a complete listing of the simplification
|
313
|
206 |
rules%
|
|
207 |
\iflabelundefined{sec:setting-up-simp}{}%
|
|
208 |
{, and \S\ref{sec:setting-up-simp} for discussion}.
|
|
209 |
|
104
|
210 |
\item
|
313
|
211 |
It instantiates the classical reasoner. See~\S\ref{fol-cla-prover}
|
104
|
212 |
for details.
|
3133
|
213 |
|
|
214 |
\item \FOL{} provides the tactic \ttindex{hyp_subst_tac}, which substitutes
|
|
215 |
for an equality throughout a subgoal and its hypotheses. This tactic uses
|
|
216 |
\FOL's general substitution rule.
|
104
|
217 |
\end{itemize}
|
|
218 |
|
3133
|
219 |
\begin{warn}\index{simplification!of conjunctions}%
|
|
220 |
Reducing $a=b\conj P(a)$ to $a=b\conj P(b)$ is sometimes advantageous. The
|
|
221 |
left part of a conjunction helps in simplifying the right part. This effect
|
|
222 |
is not available by default: it can be slow. It can be obtained by
|
|
223 |
including \ttindex{conj_cong} in a simpset, \verb$addcongs [conj_cong]$.
|
|
224 |
\end{warn}
|
|
225 |
|
104
|
226 |
|
|
227 |
\section{Intuitionistic proof procedures} \label{fol-int-prover}
|
|
228 |
Implication elimination (the rules~{\tt mp} and~{\tt impE}) pose
|
313
|
229 |
difficulties for automated proof. In intuitionistic logic, the assumption
|
|
230 |
$P\imp Q$ cannot be treated like $\neg P\disj Q$. Given $P\imp Q$, we may
|
|
231 |
use~$Q$ provided we can prove~$P$; the proof of~$P$ may require repeated
|
|
232 |
use of $P\imp Q$. If the proof of~$P$ fails then the whole branch of the
|
|
233 |
proof must be abandoned. Thus intuitionistic propositional logic requires
|
|
234 |
backtracking.
|
|
235 |
|
|
236 |
For an elementary example, consider the intuitionistic proof of $Q$ from
|
|
237 |
$P\imp Q$ and $(P\imp Q)\imp P$. The implication $P\imp Q$ is needed
|
|
238 |
twice:
|
104
|
239 |
\[ \infer[({\imp}E)]{Q}{P\imp Q &
|
|
240 |
\infer[({\imp}E)]{P}{(P\imp Q)\imp P & P\imp Q}}
|
|
241 |
\]
|
|
242 |
The theorem prover for intuitionistic logic does not use~{\tt impE}.\@
|
|
243 |
Instead, it simplifies implications using derived rules
|
287
|
244 |
(Fig.\ts\ref{fol-int-derived}). It reduces the antecedents of implications
|
313
|
245 |
to atoms and then uses Modus Ponens: from $P\imp Q$ and~$P$ deduce~$Q$.
|
|
246 |
The rules \tdx{conj_impE} and \tdx{disj_impE} are
|
104
|
247 |
straightforward: $(P\conj Q)\imp S$ is equivalent to $P\imp (Q\imp S)$, and
|
|
248 |
$(P\disj Q)\imp S$ is equivalent to the conjunction of $P\imp S$ and $Q\imp
|
|
249 |
S$. The other \ldots{\tt_impE} rules are unsafe; the method requires
|
313
|
250 |
backtracking. All the rules are derived in the same simple manner.
|
104
|
251 |
|
|
252 |
Dyckhoff has independently discovered similar rules, and (more importantly)
|
|
253 |
has demonstrated their completeness for propositional
|
|
254 |
logic~\cite{dyckhoff}. However, the tactics given below are not complete
|
|
255 |
for first-order logic because they discard universally quantified
|
|
256 |
assumptions after a single use.
|
|
257 |
\begin{ttbox}
|
3138
|
258 |
mp_tac : int -> tactic
|
|
259 |
eq_mp_tac : int -> tactic
|
|
260 |
IntPr.safe_step_tac : int -> tactic
|
|
261 |
IntPr.safe_tac : tactic
|
|
262 |
IntPr.inst_step_tac : int -> tactic
|
|
263 |
IntPr.step_tac : int -> tactic
|
|
264 |
IntPr.fast_tac : int -> tactic
|
|
265 |
IntPr.best_tac : int -> tactic
|
104
|
266 |
\end{ttbox}
|
3138
|
267 |
Most of these belong to the structure {\tt IntPr} and resemble the
|
313
|
268 |
tactics of Isabelle's classical reasoner.
|
104
|
269 |
|
313
|
270 |
\begin{ttdescription}
|
104
|
271 |
\item[\ttindexbold{mp_tac} {\it i}]
|
313
|
272 |
attempts to use \tdx{notE} or \tdx{impE} within the assumptions in
|
104
|
273 |
subgoal $i$. For each assumption of the form $\neg P$ or $P\imp Q$, it
|
|
274 |
searches for another assumption unifiable with~$P$. By
|
|
275 |
contradiction with $\neg P$ it can solve the subgoal completely; by Modus
|
|
276 |
Ponens it can replace the assumption $P\imp Q$ by $Q$. The tactic can
|
|
277 |
produce multiple outcomes, enumerating all suitable pairs of assumptions.
|
|
278 |
|
|
279 |
\item[\ttindexbold{eq_mp_tac} {\it i}]
|
|
280 |
is like {\tt mp_tac} {\it i}, but may not instantiate unknowns --- thus, it
|
|
281 |
is safe.
|
|
282 |
|
3138
|
283 |
\item[\ttindexbold{IntPr.safe_step_tac} $i$] performs a safe step on
|
313
|
284 |
subgoal~$i$. This may include proof by assumption or Modus Ponens (taking
|
|
285 |
care not to instantiate unknowns), or {\tt hyp_subst_tac}.
|
104
|
286 |
|
3138
|
287 |
\item[\ttindexbold{IntPr.safe_tac}] repeatedly performs safe steps on all
|
104
|
288 |
subgoals. It is deterministic, with at most one outcome.
|
|
289 |
|
3138
|
290 |
\item[\ttindexbold{IntPr.inst_step_tac} $i$] is like {\tt safe_step_tac},
|
104
|
291 |
but allows unknowns to be instantiated.
|
|
292 |
|
3138
|
293 |
\item[\ttindexbold{IntPr.step_tac} $i$] tries {\tt safe_tac} or {\tt
|
313
|
294 |
inst_step_tac}, or applies an unsafe rule. This is the basic step of
|
|
295 |
the intuitionistic proof procedure.
|
104
|
296 |
|
3138
|
297 |
\item[\ttindexbold{IntPr.fast_tac} $i$] applies {\tt step_tac}, using
|
104
|
298 |
depth-first search, to solve subgoal~$i$.
|
|
299 |
|
3138
|
300 |
\item[\ttindexbold{IntPr.best_tac} $i$] applies {\tt step_tac}, using
|
104
|
301 |
best-first search (guided by the size of the proof state) to solve subgoal~$i$.
|
313
|
302 |
\end{ttdescription}
|
3138
|
303 |
Here are some of the theorems that {\tt IntPr.fast_tac} proves
|
104
|
304 |
automatically. The latter three date from {\it Principia Mathematica}
|
|
305 |
(*11.53, *11.55, *11.61)~\cite{principia}.
|
|
306 |
\begin{ttbox}
|
|
307 |
~~P & ~~(P --> Q) --> ~~Q
|
|
308 |
(ALL x y. P(x) --> Q(y)) <-> ((EX x. P(x)) --> (ALL y. Q(y)))
|
|
309 |
(EX x y. P(x) & Q(x,y)) <-> (EX x. P(x) & (EX y. Q(x,y)))
|
|
310 |
(EX y. ALL x. P(x) --> Q(x,y)) --> (ALL x. P(x) --> (EX y. Q(x,y)))
|
|
311 |
\end{ttbox}
|
|
312 |
|
|
313 |
|
|
314 |
|
|
315 |
\begin{figure}
|
|
316 |
\begin{ttbox}
|
313
|
317 |
\tdx{excluded_middle} ~P | P
|
104
|
318 |
|
313
|
319 |
\tdx{disjCI} (~Q ==> P) ==> P|Q
|
|
320 |
\tdx{exCI} (ALL x. ~P(x) ==> P(a)) ==> EX x.P(x)
|
|
321 |
\tdx{impCE} [| P-->Q; ~P ==> R; Q ==> R |] ==> R
|
|
322 |
\tdx{iffCE} [| P<->Q; [| P; Q |] ==> R; [| ~P; ~Q |] ==> R |] ==> R
|
|
323 |
\tdx{notnotD} ~~P ==> P
|
|
324 |
\tdx{swap} ~P ==> (~Q ==> P) ==> Q
|
104
|
325 |
\end{ttbox}
|
|
326 |
\caption{Derived rules for classical logic} \label{fol-cla-derived}
|
|
327 |
\end{figure}
|
|
328 |
|
|
329 |
|
|
330 |
\section{Classical proof procedures} \label{fol-cla-prover}
|
313
|
331 |
The classical theory, \thydx{FOL}, consists of intuitionistic logic plus
|
|
332 |
the rule
|
104
|
333 |
$$ \vcenter{\infer{P}{\infer*{P}{[\neg P]}}} \eqno(classical) $$
|
|
334 |
\noindent
|
|
335 |
Natural deduction in classical logic is not really all that natural.
|
|
336 |
{\FOL} derives classical introduction rules for $\disj$ and~$\exists$, as
|
|
337 |
well as classical elimination rules for~$\imp$ and~$\bimp$, and the swap
|
287
|
338 |
rule (see Fig.\ts\ref{fol-cla-derived}).
|
104
|
339 |
|
3133
|
340 |
The classical reasoner is installed. Tactics such as {\tt Blast_tac} and {\tt
|
2495
|
341 |
Best_tac} use the default claset ({\tt!claset}), which works for most
|
|
342 |
purposes. Named clasets include \ttindexbold{prop_cs}, which includes the
|
|
343 |
propositional rules, and \ttindexbold{FOL_cs}, which also includes quantifier
|
|
344 |
rules. See the file {\tt FOL/cladata.ML} for lists of the
|
313
|
345 |
classical rules, and
|
|
346 |
\iflabelundefined{chap:classical}{the {\em Reference Manual\/}}%
|
|
347 |
{Chap.\ts\ref{chap:classical}}
|
|
348 |
for more discussion of classical proof methods.
|
104
|
349 |
|
|
350 |
|
|
351 |
\section{An intuitionistic example}
|
|
352 |
Here is a session similar to one in {\em Logic and Computation}
|
|
353 |
\cite[pages~222--3]{paulson87}. Isabelle treats quantifiers differently
|
|
354 |
from {\sc lcf}-based theorem provers such as {\sc hol}. The proof begins
|
|
355 |
by entering the goal in intuitionistic logic, then applying the rule
|
|
356 |
$({\imp}I)$.
|
|
357 |
\begin{ttbox}
|
|
358 |
goal IFOL.thy "(EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))";
|
|
359 |
{\out Level 0}
|
|
360 |
{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
|
|
361 |
{\out 1. (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
|
|
362 |
\ttbreak
|
|
363 |
by (resolve_tac [impI] 1);
|
|
364 |
{\out Level 1}
|
|
365 |
{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
|
|
366 |
{\out 1. EX y. ALL x. Q(x,y) ==> ALL x. EX y. Q(x,y)}
|
|
367 |
\end{ttbox}
|
313
|
368 |
In this example, we shall never have more than one subgoal. Applying
|
104
|
369 |
$({\imp}I)$ replaces~\verb|-->| by~\verb|==>|, making
|
|
370 |
\(\ex{y}\all{x}Q(x,y)\) an assumption. We have the choice of
|
|
371 |
$({\exists}E)$ and $({\forall}I)$; let us try the latter.
|
|
372 |
\begin{ttbox}
|
|
373 |
by (resolve_tac [allI] 1);
|
|
374 |
{\out Level 2}
|
|
375 |
{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
|
|
376 |
{\out 1. !!x. EX y. ALL x. Q(x,y) ==> EX y. Q(x,y)}
|
|
377 |
\end{ttbox}
|
|
378 |
Applying $({\forall}I)$ replaces the \hbox{\tt ALL x} by \hbox{\tt!!x},
|
|
379 |
changing the universal quantifier from object~($\forall$) to
|
313
|
380 |
meta~($\Forall$). The bound variable is a {\bf parameter} of the
|
104
|
381 |
subgoal. We now must choose between $({\exists}I)$ and $({\exists}E)$. What
|
|
382 |
happens if the wrong rule is chosen?
|
|
383 |
\begin{ttbox}
|
|
384 |
by (resolve_tac [exI] 1);
|
|
385 |
{\out Level 3}
|
|
386 |
{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
|
|
387 |
{\out 1. !!x. EX y. ALL x. Q(x,y) ==> Q(x,?y2(x))}
|
|
388 |
\end{ttbox}
|
|
389 |
The new subgoal~1 contains the function variable {\tt?y2}. Instantiating
|
|
390 |
{\tt?y2} can replace~{\tt?y2(x)} by a term containing~{\tt x}, even
|
|
391 |
though~{\tt x} is a bound variable. Now we analyse the assumption
|
|
392 |
\(\exists y.\forall x. Q(x,y)\) using elimination rules:
|
|
393 |
\begin{ttbox}
|
|
394 |
by (eresolve_tac [exE] 1);
|
|
395 |
{\out Level 4}
|
|
396 |
{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
|
|
397 |
{\out 1. !!x y. ALL x. Q(x,y) ==> Q(x,?y2(x))}
|
|
398 |
\end{ttbox}
|
|
399 |
Applying $(\exists E)$ has produced the parameter {\tt y} and stripped the
|
313
|
400 |
existential quantifier from the assumption. But the subgoal is unprovable:
|
|
401 |
there is no way to unify {\tt ?y2(x)} with the bound variable~{\tt y}.
|
|
402 |
Using {\tt choplev} we can return to the critical point. This time we
|
|
403 |
apply $({\exists}E)$:
|
104
|
404 |
\begin{ttbox}
|
|
405 |
choplev 2;
|
|
406 |
{\out Level 2}
|
|
407 |
{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
|
|
408 |
{\out 1. !!x. EX y. ALL x. Q(x,y) ==> EX y. Q(x,y)}
|
|
409 |
\ttbreak
|
|
410 |
by (eresolve_tac [exE] 1);
|
|
411 |
{\out Level 3}
|
|
412 |
{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
|
|
413 |
{\out 1. !!x y. ALL x. Q(x,y) ==> EX y. Q(x,y)}
|
|
414 |
\end{ttbox}
|
313
|
415 |
We now have two parameters and no scheme variables. Applying
|
|
416 |
$({\exists}I)$ and $({\forall}E)$ produces two scheme variables, which are
|
|
417 |
applied to those parameters. Parameters should be produced early, as this
|
|
418 |
example demonstrates.
|
104
|
419 |
\begin{ttbox}
|
|
420 |
by (resolve_tac [exI] 1);
|
|
421 |
{\out Level 4}
|
|
422 |
{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
|
|
423 |
{\out 1. !!x y. ALL x. Q(x,y) ==> Q(x,?y3(x,y))}
|
|
424 |
\ttbreak
|
|
425 |
by (eresolve_tac [allE] 1);
|
|
426 |
{\out Level 5}
|
|
427 |
{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
|
|
428 |
{\out 1. !!x y. Q(?x4(x,y),y) ==> Q(x,?y3(x,y))}
|
|
429 |
\end{ttbox}
|
|
430 |
The subgoal has variables {\tt ?y3} and {\tt ?x4} applied to both
|
|
431 |
parameters. The obvious projection functions unify {\tt?x4(x,y)} with~{\tt
|
|
432 |
x} and \verb|?y3(x,y)| with~{\tt y}.
|
|
433 |
\begin{ttbox}
|
|
434 |
by (assume_tac 1);
|
|
435 |
{\out Level 6}
|
|
436 |
{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
|
|
437 |
{\out No subgoals!}
|
|
438 |
\end{ttbox}
|
|
439 |
The theorem was proved in six tactic steps, not counting the abandoned
|
3138
|
440 |
ones. But proof checking is tedious; \ttindex{IntPr.fast_tac} proves the
|
104
|
441 |
theorem in one step.
|
|
442 |
\begin{ttbox}
|
|
443 |
goal IFOL.thy "(EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))";
|
|
444 |
{\out Level 0}
|
|
445 |
{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
|
|
446 |
{\out 1. (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
|
3138
|
447 |
by (IntPr.fast_tac 1);
|
104
|
448 |
{\out Level 1}
|
|
449 |
{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
|
|
450 |
{\out No subgoals!}
|
|
451 |
\end{ttbox}
|
|
452 |
|
|
453 |
|
|
454 |
\section{An example of intuitionistic negation}
|
|
455 |
The following example demonstrates the specialized forms of implication
|
|
456 |
elimination. Even propositional formulae can be difficult to prove from
|
|
457 |
the basic rules; the specialized rules help considerably.
|
|
458 |
|
313
|
459 |
Propositional examples are easy to invent. As Dummett notes~\cite[page
|
104
|
460 |
28]{dummett}, $\neg P$ is classically provable if and only if it is
|
313
|
461 |
intuitionistically provable; therefore, $P$ is classically provable if and
|
|
462 |
only if $\neg\neg P$ is intuitionistically provable.%
|
|
463 |
\footnote{Of course this holds only for propositional logic, not if $P$ is
|
|
464 |
allowed to contain quantifiers.} Proving $\neg\neg P$ intuitionistically is
|
|
465 |
much harder than proving~$P$ classically.
|
104
|
466 |
|
313
|
467 |
Our example is the double negation of the classical tautology $(P\imp
|
|
468 |
Q)\disj (Q\imp P)$. When stating the goal, we command Isabelle to expand
|
|
469 |
negations to implications using the definition $\neg P\equiv P\imp\bot$.
|
|
470 |
This allows use of the special implication rules.
|
104
|
471 |
\begin{ttbox}
|
|
472 |
goalw IFOL.thy [not_def] "~ ~ ((P-->Q) | (Q-->P))";
|
|
473 |
{\out Level 0}
|
|
474 |
{\out ~ ~ ((P --> Q) | (Q --> P))}
|
|
475 |
{\out 1. ((P --> Q) | (Q --> P) --> False) --> False}
|
|
476 |
\end{ttbox}
|
|
477 |
The first step is trivial.
|
|
478 |
\begin{ttbox}
|
|
479 |
by (resolve_tac [impI] 1);
|
|
480 |
{\out Level 1}
|
|
481 |
{\out ~ ~ ((P --> Q) | (Q --> P))}
|
|
482 |
{\out 1. (P --> Q) | (Q --> P) --> False ==> False}
|
|
483 |
\end{ttbox}
|
313
|
484 |
By $(\imp E)$ it would suffice to prove $(P\imp Q)\disj (Q\imp P)$, but
|
|
485 |
that formula is not a theorem of intuitionistic logic. Instead we apply
|
|
486 |
the specialized implication rule \tdx{disj_impE}. It splits the
|
|
487 |
assumption into two assumptions, one for each disjunct.
|
104
|
488 |
\begin{ttbox}
|
|
489 |
by (eresolve_tac [disj_impE] 1);
|
|
490 |
{\out Level 2}
|
|
491 |
{\out ~ ~ ((P --> Q) | (Q --> P))}
|
|
492 |
{\out 1. [| (P --> Q) --> False; (Q --> P) --> False |] ==> False}
|
|
493 |
\end{ttbox}
|
|
494 |
We cannot hope to prove $P\imp Q$ or $Q\imp P$ separately, but
|
313
|
495 |
their negations are inconsistent. Applying \tdx{imp_impE} breaks down
|
104
|
496 |
the assumption $\neg(P\imp Q)$, asking to show~$Q$ while providing new
|
|
497 |
assumptions~$P$ and~$\neg Q$.
|
|
498 |
\begin{ttbox}
|
|
499 |
by (eresolve_tac [imp_impE] 1);
|
|
500 |
{\out Level 3}
|
|
501 |
{\out ~ ~ ((P --> Q) | (Q --> P))}
|
|
502 |
{\out 1. [| (Q --> P) --> False; P; Q --> False |] ==> Q}
|
|
503 |
{\out 2. [| (Q --> P) --> False; False |] ==> False}
|
|
504 |
\end{ttbox}
|
|
505 |
Subgoal~2 holds trivially; let us ignore it and continue working on
|
|
506 |
subgoal~1. Thanks to the assumption~$P$, we could prove $Q\imp P$;
|
313
|
507 |
applying \tdx{imp_impE} is simpler.
|
104
|
508 |
\begin{ttbox}
|
|
509 |
by (eresolve_tac [imp_impE] 1);
|
|
510 |
{\out Level 4}
|
|
511 |
{\out ~ ~ ((P --> Q) | (Q --> P))}
|
|
512 |
{\out 1. [| P; Q --> False; Q; P --> False |] ==> P}
|
|
513 |
{\out 2. [| P; Q --> False; False |] ==> Q}
|
|
514 |
{\out 3. [| (Q --> P) --> False; False |] ==> False}
|
|
515 |
\end{ttbox}
|
|
516 |
The three subgoals are all trivial.
|
|
517 |
\begin{ttbox}
|
|
518 |
by (REPEAT (eresolve_tac [FalseE] 2));
|
|
519 |
{\out Level 5}
|
|
520 |
{\out ~ ~ ((P --> Q) | (Q --> P))}
|
|
521 |
{\out 1. [| P; Q --> False; Q; P --> False |] ==> P}
|
287
|
522 |
\ttbreak
|
104
|
523 |
by (assume_tac 1);
|
|
524 |
{\out Level 6}
|
|
525 |
{\out ~ ~ ((P --> Q) | (Q --> P))}
|
|
526 |
{\out No subgoals!}
|
|
527 |
\end{ttbox}
|
3138
|
528 |
This proof is also trivial for {\tt IntPr.fast_tac}.
|
104
|
529 |
|
|
530 |
|
|
531 |
\section{A classical example} \label{fol-cla-example}
|
|
532 |
To illustrate classical logic, we shall prove the theorem
|
313
|
533 |
$\ex{y}\all{x}P(y)\imp P(x)$. Informally, the theorem can be proved as
|
104
|
534 |
follows. Choose~$y$ such that~$\neg P(y)$, if such exists; otherwise
|
|
535 |
$\all{x}P(x)$ is true. Either way the theorem holds.
|
|
536 |
|
|
537 |
The formal proof does not conform in any obvious way to the sketch given
|
313
|
538 |
above. The key inference is the first one, \tdx{exCI}; this classical
|
104
|
539 |
version of~$(\exists I)$ allows multiple instantiation of the quantifier.
|
|
540 |
\begin{ttbox}
|
|
541 |
goal FOL.thy "EX y. ALL x. P(y)-->P(x)";
|
|
542 |
{\out Level 0}
|
|
543 |
{\out EX y. ALL x. P(y) --> P(x)}
|
|
544 |
{\out 1. EX y. ALL x. P(y) --> P(x)}
|
|
545 |
\ttbreak
|
|
546 |
by (resolve_tac [exCI] 1);
|
|
547 |
{\out Level 1}
|
|
548 |
{\out EX y. ALL x. P(y) --> P(x)}
|
|
549 |
{\out 1. ALL y. ~ (ALL x. P(y) --> P(x)) ==> ALL x. P(?a) --> P(x)}
|
|
550 |
\end{ttbox}
|
313
|
551 |
We can either exhibit a term {\tt?a} to satisfy the conclusion of
|
104
|
552 |
subgoal~1, or produce a contradiction from the assumption. The next
|
313
|
553 |
steps are routine.
|
104
|
554 |
\begin{ttbox}
|
|
555 |
by (resolve_tac [allI] 1);
|
|
556 |
{\out Level 2}
|
|
557 |
{\out EX y. ALL x. P(y) --> P(x)}
|
|
558 |
{\out 1. !!x. ALL y. ~ (ALL x. P(y) --> P(x)) ==> P(?a) --> P(x)}
|
|
559 |
\ttbreak
|
|
560 |
by (resolve_tac [impI] 1);
|
|
561 |
{\out Level 3}
|
|
562 |
{\out EX y. ALL x. P(y) --> P(x)}
|
|
563 |
{\out 1. !!x. [| ALL y. ~ (ALL x. P(y) --> P(x)); P(?a) |] ==> P(x)}
|
313
|
564 |
\end{ttbox}
|
|
565 |
By the duality between $\exists$ and~$\forall$, applying~$(\forall E)$
|
333
|
566 |
in effect applies~$(\exists I)$ again.
|
313
|
567 |
\begin{ttbox}
|
104
|
568 |
by (eresolve_tac [allE] 1);
|
|
569 |
{\out Level 4}
|
|
570 |
{\out EX y. ALL x. P(y) --> P(x)}
|
|
571 |
{\out 1. !!x. [| P(?a); ~ (ALL xa. P(?y3(x)) --> P(xa)) |] ==> P(x)}
|
|
572 |
\end{ttbox}
|
|
573 |
In classical logic, a negated assumption is equivalent to a conclusion. To
|
313
|
574 |
get this effect, we create a swapped version of~$(\forall I)$ and apply it
|
104
|
575 |
using \ttindex{eresolve_tac}; we could equivalently have applied~$(\forall
|
|
576 |
I)$ using \ttindex{swap_res_tac}.
|
|
577 |
\begin{ttbox}
|
|
578 |
allI RSN (2,swap);
|
|
579 |
{\out val it = "[| ~ (ALL x. ?P1(x)); !!x. ~ ?Q ==> ?P1(x) |] ==> ?Q" : thm}
|
|
580 |
by (eresolve_tac [it] 1);
|
|
581 |
{\out Level 5}
|
|
582 |
{\out EX y. ALL x. P(y) --> P(x)}
|
|
583 |
{\out 1. !!x xa. [| P(?a); ~ P(x) |] ==> P(?y3(x)) --> P(xa)}
|
|
584 |
\end{ttbox}
|
313
|
585 |
The previous conclusion, {\tt P(x)}, has become a negated assumption.
|
104
|
586 |
\begin{ttbox}
|
|
587 |
by (resolve_tac [impI] 1);
|
|
588 |
{\out Level 6}
|
|
589 |
{\out EX y. ALL x. P(y) --> P(x)}
|
|
590 |
{\out 1. !!x xa. [| P(?a); ~ P(x); P(?y3(x)) |] ==> P(xa)}
|
|
591 |
\end{ttbox}
|
|
592 |
The subgoal has three assumptions. We produce a contradiction between the
|
313
|
593 |
\index{assumptions!contradictory} assumptions~\verb|~P(x)| and~{\tt
|
|
594 |
P(?y3(x))}. The proof never instantiates the unknown~{\tt?a}.
|
104
|
595 |
\begin{ttbox}
|
|
596 |
by (eresolve_tac [notE] 1);
|
|
597 |
{\out Level 7}
|
|
598 |
{\out EX y. ALL x. P(y) --> P(x)}
|
|
599 |
{\out 1. !!x xa. [| P(?a); P(?y3(x)) |] ==> P(x)}
|
|
600 |
\ttbreak
|
|
601 |
by (assume_tac 1);
|
|
602 |
{\out Level 8}
|
|
603 |
{\out EX y. ALL x. P(y) --> P(x)}
|
|
604 |
{\out No subgoals!}
|
|
605 |
\end{ttbox}
|
874
|
606 |
The civilised way to prove this theorem is through \ttindex{deepen_tac},
|
|
607 |
which automatically uses the classical version of~$(\exists I)$:
|
104
|
608 |
\begin{ttbox}
|
|
609 |
goal FOL.thy "EX y. ALL x. P(y)-->P(x)";
|
|
610 |
{\out Level 0}
|
|
611 |
{\out EX y. ALL x. P(y) --> P(x)}
|
|
612 |
{\out 1. EX y. ALL x. P(y) --> P(x)}
|
2495
|
613 |
by (Deepen_tac 0 1);
|
874
|
614 |
{\out Depth = 0}
|
|
615 |
{\out Depth = 2}
|
104
|
616 |
{\out Level 1}
|
|
617 |
{\out EX y. ALL x. P(y) --> P(x)}
|
|
618 |
{\out No subgoals!}
|
|
619 |
\end{ttbox}
|
|
620 |
If this theorem seems counterintuitive, then perhaps you are an
|
|
621 |
intuitionist. In constructive logic, proving $\ex{y}\all{x}P(y)\imp P(x)$
|
|
622 |
requires exhibiting a particular term~$t$ such that $\all{x}P(t)\imp P(x)$,
|
|
623 |
which we cannot do without further knowledge about~$P$.
|
|
624 |
|
|
625 |
|
|
626 |
\section{Derived rules and the classical tactics}
|
|
627 |
Classical first-order logic can be extended with the propositional
|
|
628 |
connective $if(P,Q,R)$, where
|
|
629 |
$$ if(P,Q,R) \equiv P\conj Q \disj \neg P \conj R. \eqno(if) $$
|
|
630 |
Theorems about $if$ can be proved by treating this as an abbreviation,
|
|
631 |
replacing $if(P,Q,R)$ by $P\conj Q \disj \neg P \conj R$ in subgoals. But
|
|
632 |
this duplicates~$P$, causing an exponential blowup and an unreadable
|
|
633 |
formula. Introducing further abbreviations makes the problem worse.
|
|
634 |
|
|
635 |
Natural deduction demands rules that introduce and eliminate $if(P,Q,R)$
|
|
636 |
directly, without reference to its definition. The simple identity
|
313
|
637 |
\[ if(P,Q,R) \,\bimp\, (P\imp Q)\conj (\neg P\imp R) \]
|
104
|
638 |
suggests that the
|
|
639 |
$if$-introduction rule should be
|
157
|
640 |
\[ \infer[({if}\,I)]{if(P,Q,R)}{\infer*{Q}{[P]} & \infer*{R}{[\neg P]}} \]
|
104
|
641 |
The $if$-elimination rule reflects the definition of $if(P,Q,R)$ and the
|
|
642 |
elimination rules for~$\disj$ and~$\conj$.
|
|
643 |
\[ \infer[({if}\,E)]{S}{if(P,Q,R) & \infer*{S}{[P,Q]}
|
|
644 |
& \infer*{S}{[\neg P,R]}}
|
|
645 |
\]
|
|
646 |
Having made these plans, we get down to work with Isabelle. The theory of
|
313
|
647 |
classical logic, {\tt FOL}, is extended with the constant
|
|
648 |
$if::[o,o,o]\To o$. The axiom \tdx{if_def} asserts the
|
104
|
649 |
equation~$(if)$.
|
|
650 |
\begin{ttbox}
|
|
651 |
If = FOL +
|
1388
|
652 |
consts if :: [o,o,o]=>o
|
104
|
653 |
rules if_def "if(P,Q,R) == P&Q | ~P&R"
|
|
654 |
end
|
|
655 |
\end{ttbox}
|
|
656 |
The derivations of the introduction and elimination rules demonstrate the
|
|
657 |
methods for rewriting with definitions. Classical reasoning is required,
|
3133
|
658 |
so we use {\tt blast_tac}.
|
104
|
659 |
|
|
660 |
|
|
661 |
\subsection{Deriving the introduction rule}
|
|
662 |
The introduction rule, given the premises $P\Imp Q$ and $\neg P\Imp R$,
|
|
663 |
concludes $if(P,Q,R)$. We propose the conclusion as the main goal
|
|
664 |
using~\ttindex{goalw}, which uses {\tt if_def} to rewrite occurrences
|
|
665 |
of $if$ in the subgoal.
|
|
666 |
\begin{ttbox}
|
|
667 |
val prems = goalw If.thy [if_def]
|
|
668 |
"[| P ==> Q; ~ P ==> R |] ==> if(P,Q,R)";
|
|
669 |
{\out Level 0}
|
|
670 |
{\out if(P,Q,R)}
|
|
671 |
{\out 1. P & Q | ~ P & R}
|
|
672 |
\end{ttbox}
|
|
673 |
The premises (bound to the {\ML} variable {\tt prems}) are passed as
|
3133
|
674 |
introduction rules to \ttindex{blast_tac}. Remember that {\tt!claset} refers
|
2495
|
675 |
to the default classical set.
|
104
|
676 |
\begin{ttbox}
|
3133
|
677 |
by (blast_tac (!claset addIs prems) 1);
|
104
|
678 |
{\out Level 1}
|
|
679 |
{\out if(P,Q,R)}
|
|
680 |
{\out No subgoals!}
|
3138
|
681 |
qed "ifI";
|
104
|
682 |
\end{ttbox}
|
|
683 |
|
|
684 |
|
|
685 |
\subsection{Deriving the elimination rule}
|
|
686 |
The elimination rule has three premises, two of which are themselves rules.
|
|
687 |
The conclusion is simply $S$.
|
|
688 |
\begin{ttbox}
|
|
689 |
val major::prems = goalw If.thy [if_def]
|
|
690 |
"[| if(P,Q,R); [| P; Q |] ==> S; [| ~ P; R |] ==> S |] ==> S";
|
|
691 |
{\out Level 0}
|
|
692 |
{\out S}
|
|
693 |
{\out 1. S}
|
|
694 |
\end{ttbox}
|
|
695 |
The major premise contains an occurrence of~$if$, but the version returned
|
|
696 |
by \ttindex{goalw} (and bound to the {\ML} variable~{\tt major}) has the
|
|
697 |
definition expanded. Now \ttindex{cut_facts_tac} inserts~{\tt major} as an
|
3133
|
698 |
assumption in the subgoal, so that \ttindex{blast_tac} can break it down.
|
104
|
699 |
\begin{ttbox}
|
|
700 |
by (cut_facts_tac [major] 1);
|
|
701 |
{\out Level 1}
|
|
702 |
{\out S}
|
|
703 |
{\out 1. P & Q | ~ P & R ==> S}
|
|
704 |
\ttbreak
|
3133
|
705 |
by (blast_tac (!claset addIs prems) 1);
|
104
|
706 |
{\out Level 2}
|
|
707 |
{\out S}
|
|
708 |
{\out No subgoals!}
|
3138
|
709 |
qed "ifE";
|
104
|
710 |
\end{ttbox}
|
313
|
711 |
As you may recall from
|
|
712 |
\iflabelundefined{definitions}{{\em Introduction to Isabelle}}%
|
|
713 |
{\S\ref{definitions}}, there are other
|
104
|
714 |
ways of treating definitions when deriving a rule. We can start the
|
313
|
715 |
proof using {\tt goal}, which does not expand definitions, instead of
|
2495
|
716 |
{\tt goalw}. We can use \ttindex{rew_tac}
|
104
|
717 |
to expand definitions in the subgoals --- perhaps after calling
|
|
718 |
\ttindex{cut_facts_tac} to insert the rule's premises. We can use
|
|
719 |
\ttindex{rewrite_rule}, which is a meta-inference rule, to expand
|
|
720 |
definitions in the premises directly.
|
|
721 |
|
|
722 |
|
|
723 |
\subsection{Using the derived rules}
|
313
|
724 |
The rules just derived have been saved with the {\ML} names \tdx{ifI}
|
|
725 |
and~\tdx{ifE}. They permit natural proofs of theorems such as the
|
104
|
726 |
following:
|
|
727 |
\begin{eqnarray*}
|
111
|
728 |
if(P, if(Q,A,B), if(Q,C,D)) & \bimp & if(Q,if(P,A,C),if(P,B,D)) \\
|
|
729 |
if(if(P,Q,R), A, B) & \bimp & if(P,if(Q,A,B),if(R,A,B))
|
104
|
730 |
\end{eqnarray*}
|
|
731 |
Proofs also require the classical reasoning rules and the $\bimp$
|
313
|
732 |
introduction rule (called~\tdx{iffI}: do not confuse with~{\tt ifI}).
|
104
|
733 |
|
|
734 |
To display the $if$-rules in action, let us analyse a proof step by step.
|
|
735 |
\begin{ttbox}
|
|
736 |
goal If.thy
|
|
737 |
"if(P, if(Q,A,B), if(Q,C,D)) <-> if(Q, if(P,A,C), if(P,B,D))";
|
|
738 |
{\out Level 0}
|
|
739 |
{\out if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))}
|
|
740 |
{\out 1. if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))}
|
|
741 |
\ttbreak
|
|
742 |
by (resolve_tac [iffI] 1);
|
|
743 |
{\out Level 1}
|
|
744 |
{\out if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))}
|
|
745 |
{\out 1. if(P,if(Q,A,B),if(Q,C,D)) ==> if(Q,if(P,A,C),if(P,B,D))}
|
|
746 |
{\out 2. if(Q,if(P,A,C),if(P,B,D)) ==> if(P,if(Q,A,B),if(Q,C,D))}
|
|
747 |
\end{ttbox}
|
|
748 |
The $if$-elimination rule can be applied twice in succession.
|
|
749 |
\begin{ttbox}
|
|
750 |
by (eresolve_tac [ifE] 1);
|
|
751 |
{\out Level 2}
|
|
752 |
{\out if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))}
|
|
753 |
{\out 1. [| P; if(Q,A,B) |] ==> if(Q,if(P,A,C),if(P,B,D))}
|
|
754 |
{\out 2. [| ~ P; if(Q,C,D) |] ==> if(Q,if(P,A,C),if(P,B,D))}
|
|
755 |
{\out 3. if(Q,if(P,A,C),if(P,B,D)) ==> if(P,if(Q,A,B),if(Q,C,D))}
|
|
756 |
\ttbreak
|
|
757 |
by (eresolve_tac [ifE] 1);
|
|
758 |
{\out Level 3}
|
|
759 |
{\out if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))}
|
|
760 |
{\out 1. [| P; Q; A |] ==> if(Q,if(P,A,C),if(P,B,D))}
|
|
761 |
{\out 2. [| P; ~ Q; B |] ==> if(Q,if(P,A,C),if(P,B,D))}
|
|
762 |
{\out 3. [| ~ P; if(Q,C,D) |] ==> if(Q,if(P,A,C),if(P,B,D))}
|
|
763 |
{\out 4. if(Q,if(P,A,C),if(P,B,D)) ==> if(P,if(Q,A,B),if(Q,C,D))}
|
|
764 |
\end{ttbox}
|
333
|
765 |
%
|
|
766 |
In the first two subgoals, all assumptions have been reduced to atoms. Now
|
104
|
767 |
$if$-introduction can be applied. Observe how the $if$-rules break down
|
|
768 |
occurrences of $if$ when they become the outermost connective.
|
|
769 |
\begin{ttbox}
|
|
770 |
by (resolve_tac [ifI] 1);
|
|
771 |
{\out Level 4}
|
|
772 |
{\out if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))}
|
|
773 |
{\out 1. [| P; Q; A; Q |] ==> if(P,A,C)}
|
|
774 |
{\out 2. [| P; Q; A; ~ Q |] ==> if(P,B,D)}
|
|
775 |
{\out 3. [| P; ~ Q; B |] ==> if(Q,if(P,A,C),if(P,B,D))}
|
|
776 |
{\out 4. [| ~ P; if(Q,C,D) |] ==> if(Q,if(P,A,C),if(P,B,D))}
|
|
777 |
{\out 5. if(Q,if(P,A,C),if(P,B,D)) ==> if(P,if(Q,A,B),if(Q,C,D))}
|
|
778 |
\ttbreak
|
|
779 |
by (resolve_tac [ifI] 1);
|
|
780 |
{\out Level 5}
|
|
781 |
{\out if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))}
|
|
782 |
{\out 1. [| P; Q; A; Q; P |] ==> A}
|
|
783 |
{\out 2. [| P; Q; A; Q; ~ P |] ==> C}
|
|
784 |
{\out 3. [| P; Q; A; ~ Q |] ==> if(P,B,D)}
|
|
785 |
{\out 4. [| P; ~ Q; B |] ==> if(Q,if(P,A,C),if(P,B,D))}
|
|
786 |
{\out 5. [| ~ P; if(Q,C,D) |] ==> if(Q,if(P,A,C),if(P,B,D))}
|
|
787 |
{\out 6. if(Q,if(P,A,C),if(P,B,D)) ==> if(P,if(Q,A,B),if(Q,C,D))}
|
|
788 |
\end{ttbox}
|
|
789 |
Where do we stand? The first subgoal holds by assumption; the second and
|
2495
|
790 |
third, by contradiction. This is getting tedious. We could use the classical
|
|
791 |
reasoner, but first let us extend the default claset with the derived rules
|
104
|
792 |
for~$if$.
|
|
793 |
\begin{ttbox}
|
2495
|
794 |
AddSIs [ifI];
|
|
795 |
AddSEs [ifE];
|
|
796 |
\end{ttbox}
|
|
797 |
Now we can revert to the
|
3133
|
798 |
initial proof state and let \ttindex{blast_tac} solve it.
|
2495
|
799 |
\begin{ttbox}
|
104
|
800 |
choplev 0;
|
|
801 |
{\out Level 0}
|
|
802 |
{\out if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))}
|
|
803 |
{\out 1. if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))}
|
3133
|
804 |
by (Blast_tac 1);
|
104
|
805 |
{\out Level 1}
|
|
806 |
{\out if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))}
|
|
807 |
{\out No subgoals!}
|
|
808 |
\end{ttbox}
|
|
809 |
This tactic also solves the other example.
|
|
810 |
\begin{ttbox}
|
|
811 |
goal If.thy "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,A,B))";
|
|
812 |
{\out Level 0}
|
|
813 |
{\out if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,A,B))}
|
|
814 |
{\out 1. if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,A,B))}
|
|
815 |
\ttbreak
|
3133
|
816 |
by (Blast_tac 1);
|
104
|
817 |
{\out Level 1}
|
|
818 |
{\out if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,A,B))}
|
|
819 |
{\out No subgoals!}
|
|
820 |
\end{ttbox}
|
|
821 |
|
|
822 |
|
|
823 |
\subsection{Derived rules versus definitions}
|
|
824 |
Dispensing with the derived rules, we can treat $if$ as an
|
3133
|
825 |
abbreviation, and let \ttindex{blast_tac} prove the expanded formula. Let
|
104
|
826 |
us redo the previous proof:
|
|
827 |
\begin{ttbox}
|
|
828 |
choplev 0;
|
|
829 |
{\out Level 0}
|
|
830 |
{\out if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,A,B))}
|
|
831 |
{\out 1. if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,A,B))}
|
287
|
832 |
\end{ttbox}
|
|
833 |
This time, simply unfold using the definition of $if$:
|
|
834 |
\begin{ttbox}
|
2495
|
835 |
by (rewtac if_def);
|
104
|
836 |
{\out Level 1}
|
|
837 |
{\out if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,A,B))}
|
|
838 |
{\out 1. (P & Q | ~ P & R) & A | ~ (P & Q | ~ P & R) & B <->}
|
|
839 |
{\out P & (Q & A | ~ Q & B) | ~ P & (R & A | ~ R & B)}
|
287
|
840 |
\end{ttbox}
|
2495
|
841 |
We are left with a subgoal in pure first-order logic, which is why the
|
|
842 |
classical reasoner can prove it given {\tt FOL_cs} alone. (We could, of
|
3133
|
843 |
course, have used {\tt Blast_tac}.)
|
287
|
844 |
\begin{ttbox}
|
3133
|
845 |
by (blast_tac FOL_cs 1);
|
104
|
846 |
{\out Level 2}
|
|
847 |
{\out if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,A,B))}
|
|
848 |
{\out No subgoals!}
|
|
849 |
\end{ttbox}
|
|
850 |
Expanding definitions reduces the extended logic to the base logic. This
|
|
851 |
approach has its merits --- especially if the prover for the base logic is
|
2495
|
852 |
good --- but can be slow. In these examples, proofs using the default
|
|
853 |
claset (which includes the derived rules) run about six times faster
|
|
854 |
than proofs using {\tt FOL_cs}.
|
104
|
855 |
|
|
856 |
Expanding definitions also complicates error diagnosis. Suppose we are having
|
|
857 |
difficulties in proving some goal. If by expanding definitions we have
|
|
858 |
made it unreadable, then we have little hope of diagnosing the problem.
|
|
859 |
|
|
860 |
Attempts at program verification often yield invalid assertions.
|
|
861 |
Let us try to prove one:
|
|
862 |
\begin{ttbox}
|
|
863 |
goal If.thy "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,B,A))";
|
|
864 |
{\out Level 0}
|
|
865 |
{\out if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,B,A))}
|
|
866 |
{\out 1. if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,B,A))}
|
3133
|
867 |
by (Blast_tac 1);
|
104
|
868 |
{\out by: tactic failed}
|
|
869 |
\end{ttbox}
|
|
870 |
This failure message is uninformative, but we can get a closer look at the
|
3138
|
871 |
situation by applying \ttindex{Step_tac}.
|
104
|
872 |
\begin{ttbox}
|
2495
|
873 |
by (REPEAT (Step_tac 1));
|
104
|
874 |
{\out Level 1}
|
|
875 |
{\out if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,B,A))}
|
|
876 |
{\out 1. [| A; ~ P; R; ~ P; R |] ==> B}
|
|
877 |
{\out 2. [| B; ~ P; ~ R; ~ P; ~ R |] ==> A}
|
|
878 |
{\out 3. [| ~ P; R; B; ~ P; R |] ==> A}
|
|
879 |
{\out 4. [| ~ P; ~ R; A; ~ B; ~ P |] ==> R}
|
|
880 |
\end{ttbox}
|
|
881 |
Subgoal~1 is unprovable and yields a countermodel: $P$ and~$B$ are false
|
|
882 |
while~$R$ and~$A$ are true. This truth assignment reduces the main goal to
|
|
883 |
$true\bimp false$, which is of course invalid.
|
|
884 |
|
|
885 |
We can repeat this analysis by expanding definitions, using just
|
|
886 |
the rules of {\FOL}:
|
|
887 |
\begin{ttbox}
|
|
888 |
choplev 0;
|
|
889 |
{\out Level 0}
|
|
890 |
{\out if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,B,A))}
|
|
891 |
{\out 1. if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,B,A))}
|
|
892 |
\ttbreak
|
2495
|
893 |
by (rewtac if_def);
|
104
|
894 |
{\out Level 1}
|
|
895 |
{\out if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,B,A))}
|
|
896 |
{\out 1. (P & Q | ~ P & R) & A | ~ (P & Q | ~ P & R) & B <->}
|
|
897 |
{\out P & (Q & A | ~ Q & B) | ~ P & (R & B | ~ R & A)}
|
3133
|
898 |
by (blast_tac FOL_cs 1);
|
104
|
899 |
{\out by: tactic failed}
|
|
900 |
\end{ttbox}
|
|
901 |
Again we apply \ttindex{step_tac}:
|
|
902 |
\begin{ttbox}
|
|
903 |
by (REPEAT (step_tac FOL_cs 1));
|
|
904 |
{\out Level 2}
|
|
905 |
{\out if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,B,A))}
|
|
906 |
{\out 1. [| A; ~ P; R; ~ P; R; ~ False |] ==> B}
|
|
907 |
{\out 2. [| A; ~ P; R; R; ~ False; ~ B; ~ B |] ==> Q}
|
|
908 |
{\out 3. [| B; ~ P; ~ R; ~ P; ~ A |] ==> R}
|
|
909 |
{\out 4. [| B; ~ P; ~ R; ~ Q; ~ A |] ==> R}
|
|
910 |
{\out 5. [| B; ~ R; ~ P; ~ A; ~ R; Q; ~ False |] ==> A}
|
|
911 |
{\out 6. [| ~ P; R; B; ~ P; R; ~ False |] ==> A}
|
|
912 |
{\out 7. [| ~ P; ~ R; A; ~ B; ~ R |] ==> P}
|
|
913 |
{\out 8. [| ~ P; ~ R; A; ~ B; ~ R |] ==> Q}
|
|
914 |
\end{ttbox}
|
|
915 |
Subgoal~1 yields the same countermodel as before. But each proof step has
|
|
916 |
taken six times as long, and the final result contains twice as many subgoals.
|
|
917 |
|
|
918 |
Expanding definitions causes a great increase in complexity. This is why
|
|
919 |
the classical prover has been designed to accept derived rules.
|
313
|
920 |
|
|
921 |
\index{first-order logic|)}
|