author | wenzelm |
Fri, 23 Jun 2017 14:38:32 +0200 | |
changeset 66176 | b51a40281016 |
parent 63821 | 52235c27538c |
child 67399 | eab6ce8368fa |
permissions | -rw-r--r-- |
61656 | 1 |
(*:maxLineLen=78:*) |
2 |
||
26782 | 3 |
theory Generic |
63531 | 4 |
imports Main Base |
26782 | 5 |
begin |
6 |
||
58618 | 7 |
chapter \<open>Generic tools and packages \label{ch:gen-tools}\<close> |
26782 | 8 |
|
58618 | 9 |
section \<open>Configuration options \label{sec:config}\<close> |
26782 | 10 |
|
63531 | 11 |
text \<open> |
12 |
Isabelle/Pure maintains a record of named configuration options within the |
|
13 |
theory or proof context, with values of type @{ML_type bool}, @{ML_type |
|
14 |
int}, @{ML_type real}, or @{ML_type string}. Tools may declare options in |
|
15 |
ML, and then refer to these values (relative to the context). Thus global |
|
16 |
reference variables are easily avoided. The user may change the value of a |
|
17 |
configuration option by means of an associated attribute of the same name. |
|
18 |
This form of context declaration works particularly well with commands such |
|
19 |
as @{command "declare"} or @{command "using"} like this: |
|
58618 | 20 |
\<close> |
42655 | 21 |
|
59905 | 22 |
(*<*)experiment begin(*>*) |
42655 | 23 |
declare [[show_main_goal = false]] |
26782 | 24 |
|
42655 | 25 |
notepad |
26 |
begin |
|
27 |
note [[show_main_goal = true]] |
|
28 |
end |
|
59905 | 29 |
(*<*)end(*>*) |
42655 | 30 |
|
59921 | 31 |
text \<open> |
26782 | 32 |
\begin{matharray}{rcll} |
61493 | 33 |
@{command_def "print_options"} & : & \<open>context \<rightarrow>\<close> \\ |
26782 | 34 |
\end{matharray} |
35 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
36 |
@{rail \<open> |
59917
9830c944670f
more uniform "verbose" option to print name space;
wenzelm
parents:
59905
diff
changeset
|
37 |
@@{command print_options} ('!'?) |
9830c944670f
more uniform "verbose" option to print name space;
wenzelm
parents:
59905
diff
changeset
|
38 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
39 |
@{syntax name} ('=' ('true' | 'false' | @{syntax int} | @{syntax float} | @{syntax name}))? |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
40 |
\<close>} |
26782 | 41 |
|
63531 | 42 |
\<^descr> @{command "print_options"} prints the available configuration options, |
43 |
with names, types, and current values; the ``\<open>!\<close>'' option indicates extra |
|
44 |
verbosity. |
|
26782 | 45 |
|
63531 | 46 |
\<^descr> \<open>name = value\<close> as an attribute expression modifies the named option, with |
47 |
the syntax of the value depending on the option's type. For @{ML_type bool} |
|
48 |
the default value is \<open>true\<close>. Any attempt to change a global option in a |
|
49 |
local context is ignored. |
|
58618 | 50 |
\<close> |
26782 | 51 |
|
52 |
||
58618 | 53 |
section \<open>Basic proof tools\<close> |
26782 | 54 |
|
58618 | 55 |
subsection \<open>Miscellaneous methods and attributes \label{sec:misc-meth-att}\<close> |
26782 | 56 |
|
58618 | 57 |
text \<open> |
26782 | 58 |
\begin{matharray}{rcl} |
61493 | 59 |
@{method_def unfold} & : & \<open>method\<close> \\ |
60 |
@{method_def fold} & : & \<open>method\<close> \\ |
|
61 |
@{method_def insert} & : & \<open>method\<close> \\[0.5ex] |
|
62 |
@{method_def erule}\<open>\<^sup>*\<close> & : & \<open>method\<close> \\ |
|
63 |
@{method_def drule}\<open>\<^sup>*\<close> & : & \<open>method\<close> \\ |
|
64 |
@{method_def frule}\<open>\<^sup>*\<close> & : & \<open>method\<close> \\ |
|
65 |
@{method_def intro} & : & \<open>method\<close> \\ |
|
66 |
@{method_def elim} & : & \<open>method\<close> \\ |
|
67 |
@{method_def fail} & : & \<open>method\<close> \\ |
|
68 |
@{method_def succeed} & : & \<open>method\<close> \\ |
|
69 |
@{method_def sleep} & : & \<open>method\<close> \\ |
|
26782 | 70 |
\end{matharray} |
71 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
72 |
@{rail \<open> |
62969 | 73 |
(@@{method fold} | @@{method unfold} | @@{method insert}) @{syntax thms} |
26782 | 74 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
75 |
(@@{method erule} | @@{method drule} | @@{method frule}) |
62969 | 76 |
('(' @{syntax nat} ')')? @{syntax thms} |
43365 | 77 |
; |
62969 | 78 |
(@@{method intro} | @@{method elim}) @{syntax thms}? |
60556 | 79 |
; |
80 |
@@{method sleep} @{syntax real} |
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
81 |
\<close>} |
26782 | 82 |
|
63531 | 83 |
\<^descr> @{method unfold}~\<open>a\<^sub>1 \<dots> a\<^sub>n\<close> and @{method fold}~\<open>a\<^sub>1 \<dots> a\<^sub>n\<close> expand (or |
84 |
fold back) the given definitions throughout all goals; any chained facts |
|
85 |
provided are inserted into the goal and subject to rewriting as well. |
|
26782 | 86 |
|
63821 | 87 |
Unfolding works in two stages: first, the given equations are used directly |
88 |
for rewriting; second, the equations are passed through the attribute |
|
89 |
@{attribute_ref abs_def} before rewriting --- to ensure that definitions are |
|
90 |
fully expanded, regardless of the actual parameters that are provided. |
|
91 |
||
63531 | 92 |
\<^descr> @{method insert}~\<open>a\<^sub>1 \<dots> a\<^sub>n\<close> inserts theorems as facts into all goals of |
93 |
the proof state. Note that current facts indicated for forward chaining are |
|
94 |
ignored. |
|
95 |
||
96 |
\<^descr> @{method erule}~\<open>a\<^sub>1 \<dots> a\<^sub>n\<close>, @{method drule}~\<open>a\<^sub>1 \<dots> a\<^sub>n\<close>, and @{method |
|
97 |
frule}~\<open>a\<^sub>1 \<dots> a\<^sub>n\<close> are similar to the basic @{method rule} method (see |
|
98 |
\secref{sec:pure-meth-att}), but apply rules by elim-resolution, |
|
99 |
destruct-resolution, and forward-resolution, respectively @{cite |
|
100 |
"isabelle-implementation"}. The optional natural number argument (default 0) |
|
101 |
specifies additional assumption steps to be performed here. |
|
26782 | 102 |
|
103 |
Note that these methods are improper ones, mainly serving for |
|
63531 | 104 |
experimentation and tactic script emulation. Different modes of basic rule |
105 |
application are usually expressed in Isar at the proof language level, |
|
106 |
rather than via implicit proof state manipulations. For example, a proper |
|
107 |
single-step elimination would be done using the plain @{method rule} method, |
|
108 |
with forward chaining of current facts. |
|
26782 | 109 |
|
63531 | 110 |
\<^descr> @{method intro} and @{method elim} repeatedly refine some goal by intro- |
111 |
or elim-resolution, after having inserted any chained facts. Exactly the |
|
112 |
rules given as arguments are taken into account; this allows fine-tuned |
|
113 |
decomposition of a proof problem, in contrast to common automated tools. |
|
43365 | 114 |
|
63531 | 115 |
\<^descr> @{method fail} yields an empty result sequence; it is the identity of the |
116 |
``\<open>|\<close>'' method combinator (cf.\ \secref{sec:proof-meth}). |
|
60554 | 117 |
|
63531 | 118 |
\<^descr> @{method succeed} yields a single (unchanged) result; it is the identity |
119 |
of the ``\<open>,\<close>'' method combinator (cf.\ \secref{sec:proof-meth}). |
|
26782 | 120 |
|
63531 | 121 |
\<^descr> @{method sleep}~\<open>s\<close> succeeds after a real-time delay of \<open>s\<close> seconds. This |
122 |
is occasionally useful for demonstration and testing purposes. |
|
26782 | 123 |
|
124 |
||
125 |
\begin{matharray}{rcl} |
|
61493 | 126 |
@{attribute_def tagged} & : & \<open>attribute\<close> \\ |
127 |
@{attribute_def untagged} & : & \<open>attribute\<close> \\[0.5ex] |
|
128 |
@{attribute_def THEN} & : & \<open>attribute\<close> \\ |
|
129 |
@{attribute_def unfolded} & : & \<open>attribute\<close> \\ |
|
130 |
@{attribute_def folded} & : & \<open>attribute\<close> \\ |
|
131 |
@{attribute_def abs_def} & : & \<open>attribute\<close> \\[0.5ex] |
|
132 |
@{attribute_def rotated} & : & \<open>attribute\<close> \\ |
|
133 |
@{attribute_def (Pure) elim_format} & : & \<open>attribute\<close> \\ |
|
134 |
@{attribute_def no_vars}\<open>\<^sup>*\<close> & : & \<open>attribute\<close> \\ |
|
26782 | 135 |
\end{matharray} |
136 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
137 |
@{rail \<open> |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
138 |
@@{attribute tagged} @{syntax name} @{syntax name} |
26782 | 139 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
140 |
@@{attribute untagged} @{syntax name} |
26782 | 141 |
; |
62969 | 142 |
@@{attribute THEN} ('[' @{syntax nat} ']')? @{syntax thm} |
26782 | 143 |
; |
62969 | 144 |
(@@{attribute unfolded} | @@{attribute folded}) @{syntax thms} |
26782 | 145 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
146 |
@@{attribute rotated} @{syntax int}? |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
147 |
\<close>} |
26782 | 148 |
|
63531 | 149 |
\<^descr> @{attribute tagged}~\<open>name value\<close> and @{attribute untagged}~\<open>name\<close> add and |
150 |
remove \<^emph>\<open>tags\<close> of some theorem. Tags may be any list of string pairs that |
|
151 |
serve as formal comment. The first string is considered the tag name, the |
|
152 |
second its value. Note that @{attribute untagged} removes any tags of the |
|
153 |
same name. |
|
26782 | 154 |
|
63531 | 155 |
\<^descr> @{attribute THEN}~\<open>a\<close> composes rules by resolution; it resolves with the |
156 |
first premise of \<open>a\<close> (an alternative position may be also specified). See |
|
157 |
also @{ML_op "RS"} in @{cite "isabelle-implementation"}. |
|
26782 | 158 |
|
63531 | 159 |
\<^descr> @{attribute unfolded}~\<open>a\<^sub>1 \<dots> a\<^sub>n\<close> and @{attribute folded}~\<open>a\<^sub>1 \<dots> a\<^sub>n\<close> |
160 |
expand and fold back again the given definitions throughout a rule. |
|
26782 | 161 |
|
63531 | 162 |
\<^descr> @{attribute abs_def} turns an equation of the form @{prop "f x y \<equiv> t"} |
63821 | 163 |
into @{prop "f \<equiv> \<lambda>x y. t"}, which ensures that @{method simp} steps always |
164 |
expand it. This also works for object-logic equality. |
|
47497 | 165 |
|
63531 | 166 |
\<^descr> @{attribute rotated}~\<open>n\<close> rotate the premises of a theorem by \<open>n\<close> (default |
167 |
1). |
|
26782 | 168 |
|
63531 | 169 |
\<^descr> @{attribute (Pure) elim_format} turns a destruction rule into elimination |
170 |
rule format, by resolving with the rule @{prop "PROP A \<Longrightarrow> (PROP A \<Longrightarrow> PROP B) \<Longrightarrow> |
|
171 |
PROP B"}. |
|
26782 | 172 |
|
63531 | 173 |
Note that the Classical Reasoner (\secref{sec:classical}) provides its own |
174 |
version of this operation. |
|
26782 | 175 |
|
63531 | 176 |
\<^descr> @{attribute no_vars} replaces schematic variables by free ones; this is |
177 |
mainly for tuning output of pretty printed theorems. |
|
58618 | 178 |
\<close> |
26782 | 179 |
|
180 |
||
58618 | 181 |
subsection \<open>Low-level equational reasoning\<close> |
27044 | 182 |
|
58618 | 183 |
text \<open> |
27044 | 184 |
\begin{matharray}{rcl} |
61493 | 185 |
@{method_def subst} & : & \<open>method\<close> \\ |
186 |
@{method_def hypsubst} & : & \<open>method\<close> \\ |
|
187 |
@{method_def split} & : & \<open>method\<close> \\ |
|
27044 | 188 |
\end{matharray} |
189 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
190 |
@{rail \<open> |
62969 | 191 |
@@{method subst} ('(' 'asm' ')')? \<newline> ('(' (@{syntax nat}+) ')')? @{syntax thm} |
27044 | 192 |
; |
62969 | 193 |
@@{method split} @{syntax thms} |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
194 |
\<close>} |
27044 | 195 |
|
63531 | 196 |
These methods provide low-level facilities for equational reasoning that are |
197 |
intended for specialized applications only. Normally, single step |
|
198 |
calculations would be performed in a structured text (see also |
|
199 |
\secref{sec:calculation}), while the Simplifier methods provide the |
|
200 |
canonical way for automated normalization (see \secref{sec:simplifier}). |
|
201 |
||
202 |
\<^descr> @{method subst}~\<open>eq\<close> performs a single substitution step using rule \<open>eq\<close>, |
|
203 |
which may be either a meta or object equality. |
|
27044 | 204 |
|
63531 | 205 |
\<^descr> @{method subst}~\<open>(asm) eq\<close> substitutes in an assumption. |
27044 | 206 |
|
63531 | 207 |
\<^descr> @{method subst}~\<open>(i \<dots> j) eq\<close> performs several substitutions in the |
208 |
conclusion. The numbers \<open>i\<close> to \<open>j\<close> indicate the positions to substitute at. |
|
209 |
Positions are ordered from the top of the term tree moving down from left to |
|
210 |
right. For example, in \<open>(a + b) + (c + d)\<close> there are three positions where |
|
211 |
commutativity of \<open>+\<close> is applicable: 1 refers to \<open>a + b\<close>, 2 to the whole |
|
212 |
term, and 3 to \<open>c + d\<close>. |
|
27044 | 213 |
|
63531 | 214 |
If the positions in the list \<open>(i \<dots> j)\<close> are non-overlapping (e.g.\ \<open>(2 3)\<close> in |
215 |
\<open>(a + b) + (c + d)\<close>) you may assume all substitutions are performed |
|
216 |
simultaneously. Otherwise the behaviour of \<open>subst\<close> is not specified. |
|
27044 | 217 |
|
63531 | 218 |
\<^descr> @{method subst}~\<open>(asm) (i \<dots> j) eq\<close> performs the substitutions in the |
219 |
assumptions. The positions refer to the assumptions in order from left to |
|
220 |
right. For example, given in a goal of the form \<open>P (a + b) \<Longrightarrow> P (c + d) \<Longrightarrow> \<dots>\<close>, |
|
221 |
position 1 of commutativity of \<open>+\<close> is the subterm \<open>a + b\<close> and position 2 is |
|
222 |
the subterm \<open>c + d\<close>. |
|
27044 | 223 |
|
63531 | 224 |
\<^descr> @{method hypsubst} performs substitution using some assumption; this only |
225 |
works for equations of the form \<open>x = t\<close> where \<open>x\<close> is a free or bound |
|
226 |
variable. |
|
27044 | 227 |
|
63531 | 228 |
\<^descr> @{method split}~\<open>a\<^sub>1 \<dots> a\<^sub>n\<close> performs single-step case splitting using the |
229 |
given rules. Splitting is performed in the conclusion or some assumption of |
|
230 |
the subgoal, depending of the structure of the rule. |
|
27044 | 231 |
|
63531 | 232 |
Note that the @{method simp} method already involves repeated application of |
233 |
split rules as declared in the current context, using @{attribute split}, |
|
234 |
for example. |
|
58618 | 235 |
\<close> |
27044 | 236 |
|
237 |
||
58618 | 238 |
section \<open>The Simplifier \label{sec:simplifier}\<close> |
26782 | 239 |
|
62655 | 240 |
text \<open> |
241 |
The Simplifier performs conditional and unconditional rewriting and uses |
|
242 |
contextual information: rule declarations in the background theory or local |
|
243 |
proof context are taken into account, as well as chained facts and subgoal |
|
244 |
premises (``local assumptions''). There are several general hooks that allow |
|
245 |
to modify the simplification strategy, or incorporate other proof tools that |
|
246 |
solve sub-problems, produce rewrite rules on demand etc. |
|
50063 | 247 |
|
62655 | 248 |
The rewriting strategy is always strictly bottom up, except for congruence |
249 |
rules, which are applied while descending into a term. Conditions in |
|
250 |
conditional rewrite rules are solved recursively before the rewrite rule is |
|
251 |
applied. |
|
50075 | 252 |
|
62655 | 253 |
The default Simplifier setup of major object logics (HOL, HOLCF, FOL, ZF) |
254 |
makes the Simplifier ready for immediate use, without engaging into the |
|
255 |
internal structures. Thus it serves as general-purpose proof tool with the |
|
256 |
main focus on equational reasoning, and a bit more than that. |
|
58618 | 257 |
\<close> |
50063 | 258 |
|
259 |
||
58618 | 260 |
subsection \<open>Simplification methods \label{sec:simp-meth}\<close> |
26782 | 261 |
|
58618 | 262 |
text \<open> |
57591
8c095aef6769
clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents:
57590
diff
changeset
|
263 |
\begin{tabular}{rcll} |
61493 | 264 |
@{method_def simp} & : & \<open>method\<close> \\ |
265 |
@{method_def simp_all} & : & \<open>method\<close> \\ |
|
63532
b01154b74314
provide Pure.simp/simp_all, which only know about meta-equality;
wenzelm
parents:
63531
diff
changeset
|
266 |
\<open>Pure.\<close>@{method_def (Pure) simp} & : & \<open>method\<close> \\ |
b01154b74314
provide Pure.simp/simp_all, which only know about meta-equality;
wenzelm
parents:
63531
diff
changeset
|
267 |
\<open>Pure.\<close>@{method_def (Pure) simp_all} & : & \<open>method\<close> \\ |
61493 | 268 |
@{attribute_def simp_depth_limit} & : & \<open>attribute\<close> & default \<open>100\<close> \\ |
57591
8c095aef6769
clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents:
57590
diff
changeset
|
269 |
\end{tabular} |
61421 | 270 |
\<^medskip> |
26782 | 271 |
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
272 |
@{rail \<open> |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
273 |
(@@{method simp} | @@{method simp_all}) opt? (@{syntax simpmod} * ) |
26782 | 274 |
; |
275 |
||
40255
9ffbc25e1606
eliminated obsolete \_ escapes in rail environments;
wenzelm
parents:
35613
diff
changeset
|
276 |
opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use' | 'asm_lr' ) ')' |
26782 | 277 |
; |
63650 | 278 |
@{syntax_def simpmod}: ('add' | 'del' | 'only' | 'split' (() | '!' | 'del') | |
62969 | 279 |
'cong' (() | 'add' | 'del')) ':' @{syntax thms} |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
280 |
\<close>} |
26782 | 281 |
|
62655 | 282 |
\<^descr> @{method simp} invokes the Simplifier on the first subgoal, after |
283 |
inserting chained facts as additional goal premises; further rule |
|
284 |
declarations may be included via \<open>(simp add: facts)\<close>. The proof method fails |
|
285 |
if the subgoal remains unchanged after simplification. |
|
26782 | 286 |
|
62655 | 287 |
Note that the original goal premises and chained facts are subject to |
288 |
simplification themselves, while declarations via \<open>add\<close>/\<open>del\<close> merely follow |
|
289 |
the policies of the object-logic to extract rewrite rules from theorems, |
|
290 |
without further simplification. This may lead to slightly different behavior |
|
291 |
in either case, which might be required precisely like that in some boundary |
|
292 |
situations to perform the intended simplification step! |
|
50063 | 293 |
|
61421 | 294 |
\<^medskip> |
62655 | 295 |
The \<open>only\<close> modifier first removes all other rewrite rules, looper tactics |
296 |
(including split rules), congruence rules, and then behaves like \<open>add\<close>. |
|
297 |
Implicit solvers remain, which means that trivial rules like reflexivity or |
|
298 |
introduction of \<open>True\<close> are available to solve the simplified subgoals, but |
|
299 |
also non-trivial tools like linear arithmetic in HOL. The latter may lead to |
|
300 |
some surprise of the meaning of ``only'' in Isabelle/HOL compared to |
|
301 |
English! |
|
26782 | 302 |
|
61421 | 303 |
\<^medskip> |
62655 | 304 |
The \<open>split\<close> modifiers add or delete rules for the Splitter (see also |
305 |
\secref{sec:simp-strategies} on the looper). This works only if the |
|
306 |
Simplifier method has been properly setup to include the Splitter (all major |
|
307 |
object logics such HOL, HOLCF, FOL, ZF do this already). |
|
63650 | 308 |
The \<open>!\<close> option causes the split rules to be used aggressively: |
309 |
after each application of a split rule in the conclusion, the \<open>safe\<close> |
|
310 |
tactic of the classical reasoner (see \secref{sec:classical:partial}) |
|
311 |
is applied to the new goal. The net effect is that the goal is split into |
|
312 |
the different cases. This option can speed up simplification of goals |
|
313 |
with many nested conditional or case expressions significantly. |
|
26782 | 314 |
|
50065 | 315 |
There is also a separate @{method_ref split} method available for |
62655 | 316 |
single-step case splitting. The effect of repeatedly applying \<open>(split thms)\<close> |
317 |
can be imitated by ``\<open>(simp only: split: thms)\<close>''. |
|
50065 | 318 |
|
61421 | 319 |
\<^medskip> |
62655 | 320 |
The \<open>cong\<close> modifiers add or delete Simplifier congruence rules (see also |
321 |
\secref{sec:simp-rules}); the default is to add. |
|
50063 | 322 |
|
62655 | 323 |
\<^descr> @{method simp_all} is similar to @{method simp}, but acts on all goals, |
324 |
working backwards from the last to the first one as usual in Isabelle.\<^footnote>\<open>The |
|
325 |
order is irrelevant for goals without schematic variables, so simplification |
|
326 |
might actually be performed in parallel here.\<close> |
|
50063 | 327 |
|
62655 | 328 |
Chained facts are inserted into all subgoals, before the simplification |
329 |
process starts. Further rule declarations are the same as for @{method |
|
330 |
simp}. |
|
50063 | 331 |
|
332 |
The proof method fails if all subgoals remain unchanged after |
|
333 |
simplification. |
|
26782 | 334 |
|
62655 | 335 |
\<^descr> @{attribute simp_depth_limit} limits the number of recursive invocations |
336 |
of the Simplifier during conditional rewriting. |
|
57591
8c095aef6769
clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents:
57590
diff
changeset
|
337 |
|
26782 | 338 |
|
62655 | 339 |
By default the Simplifier methods above take local assumptions fully into |
340 |
account, using equational assumptions in the subsequent normalization |
|
341 |
process, or simplifying assumptions themselves. Further options allow to |
|
342 |
fine-tune the behavior of the Simplifier in this respect, corresponding to a |
|
343 |
variety of ML tactics as follows.\<^footnote>\<open>Unlike the corresponding Isar proof |
|
344 |
methods, the ML tactics do not insist in changing the goal state.\<close> |
|
50063 | 345 |
|
346 |
\begin{center} |
|
347 |
\small |
|
59782 | 348 |
\begin{tabular}{|l|l|p{0.3\textwidth}|} |
50063 | 349 |
\hline |
350 |
Isar method & ML tactic & behavior \\\hline |
|
351 |
||
62655 | 352 |
\<open>(simp (no_asm))\<close> & @{ML simp_tac} & assumptions are ignored completely |
353 |
\\\hline |
|
26782 | 354 |
|
62655 | 355 |
\<open>(simp (no_asm_simp))\<close> & @{ML asm_simp_tac} & assumptions are used in the |
356 |
simplification of the conclusion but are not themselves simplified \\\hline |
|
50063 | 357 |
|
62655 | 358 |
\<open>(simp (no_asm_use))\<close> & @{ML full_simp_tac} & assumptions are simplified but |
359 |
are not used in the simplification of each other or the conclusion \\\hline |
|
26782 | 360 |
|
62655 | 361 |
\<open>(simp)\<close> & @{ML asm_full_simp_tac} & assumptions are used in the |
362 |
simplification of the conclusion and to simplify other assumptions \\\hline |
|
50063 | 363 |
|
62655 | 364 |
\<open>(simp (asm_lr))\<close> & @{ML asm_lr_simp_tac} & compatibility mode: an |
365 |
assumption is only used for simplifying assumptions which are to the right |
|
366 |
of it \\\hline |
|
50063 | 367 |
|
59782 | 368 |
\end{tabular} |
50063 | 369 |
\end{center} |
63532
b01154b74314
provide Pure.simp/simp_all, which only know about meta-equality;
wenzelm
parents:
63531
diff
changeset
|
370 |
|
b01154b74314
provide Pure.simp/simp_all, which only know about meta-equality;
wenzelm
parents:
63531
diff
changeset
|
371 |
\<^medskip> |
b01154b74314
provide Pure.simp/simp_all, which only know about meta-equality;
wenzelm
parents:
63531
diff
changeset
|
372 |
In Isabelle/Pure, proof methods @{method (Pure) simp} and @{method (Pure) |
b01154b74314
provide Pure.simp/simp_all, which only know about meta-equality;
wenzelm
parents:
63531
diff
changeset
|
373 |
simp_all} only know about meta-equality \<open>\<equiv>\<close>. Any new object-logic needs to |
b01154b74314
provide Pure.simp/simp_all, which only know about meta-equality;
wenzelm
parents:
63531
diff
changeset
|
374 |
re-define these methods via @{ML Simplifier.method_setup} in ML: |
b01154b74314
provide Pure.simp/simp_all, which only know about meta-equality;
wenzelm
parents:
63531
diff
changeset
|
375 |
Isabelle/FOL or Isabelle/HOL may serve as blue-prints. |
58618 | 376 |
\<close> |
26782 | 377 |
|
378 |
||
58618 | 379 |
subsubsection \<open>Examples\<close> |
50064 | 380 |
|
62655 | 381 |
text \<open> |
382 |
We consider basic algebraic simplifications in Isabelle/HOL. The rather |
|
383 |
trivial goal @{prop "0 + (x + 0) = x + 0 + 0"} looks like a good candidate |
|
384 |
to be solved by a single call of @{method simp}: |
|
58618 | 385 |
\<close> |
50064 | 386 |
|
387 |
lemma "0 + (x + 0) = x + 0 + 0" apply simp? oops |
|
388 |
||
62655 | 389 |
text \<open> |
390 |
The above attempt \<^emph>\<open>fails\<close>, because @{term "0"} and @{term "op +"} in the |
|
391 |
HOL library are declared as generic type class operations, without stating |
|
392 |
any algebraic laws yet. More specific types are required to get access to |
|
393 |
certain standard simplifications of the theory context, e.g.\ like this:\<close> |
|
50064 | 394 |
|
395 |
lemma fixes x :: nat shows "0 + (x + 0) = x + 0 + 0" by simp |
|
396 |
lemma fixes x :: int shows "0 + (x + 0) = x + 0 + 0" by simp |
|
397 |
lemma fixes x :: "'a :: monoid_add" shows "0 + (x + 0) = x + 0 + 0" by simp |
|
398 |
||
58618 | 399 |
text \<open> |
61421 | 400 |
\<^medskip> |
62655 | 401 |
In many cases, assumptions of a subgoal are also needed in the |
402 |
simplification process. For example: |
|
58618 | 403 |
\<close> |
50064 | 404 |
|
405 |
lemma fixes x :: nat shows "x = 0 \<Longrightarrow> x + x = 0" by simp |
|
406 |
lemma fixes x :: nat assumes "x = 0" shows "x + x = 0" apply simp oops |
|
407 |
lemma fixes x :: nat assumes "x = 0" shows "x + x = 0" using assms by simp |
|
408 |
||
62655 | 409 |
text \<open> |
410 |
As seen above, local assumptions that shall contribute to simplification |
|
411 |
need to be part of the subgoal already, or indicated explicitly for use by |
|
412 |
the subsequent method invocation. Both too little or too much information |
|
413 |
can make simplification fail, for different reasons. |
|
50064 | 414 |
|
62655 | 415 |
In the next example the malicious assumption @{prop "\<And>x::nat. f x = g (f (g |
416 |
x))"} does not contribute to solve the problem, but makes the default |
|
417 |
@{method simp} method loop: the rewrite rule \<open>f ?x \<equiv> g (f (g ?x))\<close> extracted |
|
418 |
from the assumption does not terminate. The Simplifier notices certain |
|
419 |
simple forms of nontermination, but not this one. The problem can be solved |
|
420 |
nonetheless, by ignoring assumptions via special options as explained |
|
421 |
before: |
|
58618 | 422 |
\<close> |
50064 | 423 |
|
424 |
lemma "(\<And>x::nat. f x = g (f (g x))) \<Longrightarrow> f 0 = f 0 + 0" |
|
425 |
by (simp (no_asm)) |
|
426 |
||
62655 | 427 |
text \<open> |
428 |
The latter form is typical for long unstructured proof scripts, where the |
|
429 |
control over the goal content is limited. In structured proofs it is usually |
|
430 |
better to avoid pushing too many facts into the goal state in the first |
|
431 |
place. Assumptions in the Isar proof context do not intrude the reasoning if |
|
432 |
not used explicitly. This is illustrated for a toplevel statement and a |
|
50064 | 433 |
local proof body as follows: |
58618 | 434 |
\<close> |
50064 | 435 |
|
436 |
lemma |
|
437 |
assumes "\<And>x::nat. f x = g (f (g x))" |
|
438 |
shows "f 0 = f 0 + 0" by simp |
|
439 |
||
440 |
notepad |
|
441 |
begin |
|
442 |
assume "\<And>x::nat. f x = g (f (g x))" |
|
443 |
have "f 0 = f 0 + 0" by simp |
|
444 |
end |
|
445 |
||
61421 | 446 |
text \<open> |
447 |
\<^medskip> |
|
62655 | 448 |
Because assumptions may simplify each other, there can be very subtle cases |
449 |
of nontermination. For example, the regular @{method simp} method applied to |
|
450 |
@{prop "P (f x) \<Longrightarrow> y = x \<Longrightarrow> f x = f y \<Longrightarrow> Q"} gives rise to the infinite |
|
451 |
reduction sequence |
|
50064 | 452 |
\[ |
61493 | 453 |
\<open>P (f x)\<close> \stackrel{\<open>f x \<equiv> f y\<close>}{\longmapsto} |
454 |
\<open>P (f y)\<close> \stackrel{\<open>y \<equiv> x\<close>}{\longmapsto} |
|
455 |
\<open>P (f x)\<close> \stackrel{\<open>f x \<equiv> f y\<close>}{\longmapsto} \cdots |
|
50064 | 456 |
\] |
62655 | 457 |
whereas applying the same to @{prop "y = x \<Longrightarrow> f x = f y \<Longrightarrow> P (f x) \<Longrightarrow> Q"} |
458 |
terminates (without solving the goal): |
|
58618 | 459 |
\<close> |
50064 | 460 |
|
461 |
lemma "y = x \<Longrightarrow> f x = f y \<Longrightarrow> P (f x) \<Longrightarrow> Q" |
|
462 |
apply simp |
|
463 |
oops |
|
464 |
||
62655 | 465 |
text \<open> |
466 |
See also \secref{sec:simp-trace} for options to enable Simplifier trace |
|
467 |
mode, which often helps to diagnose problems with rewrite systems. |
|
58618 | 468 |
\<close> |
50064 | 469 |
|
470 |
||
58618 | 471 |
subsection \<open>Declaring rules \label{sec:simp-rules}\<close> |
26782 | 472 |
|
58618 | 473 |
text \<open> |
26782 | 474 |
\begin{matharray}{rcl} |
61493 | 475 |
@{attribute_def simp} & : & \<open>attribute\<close> \\ |
476 |
@{attribute_def split} & : & \<open>attribute\<close> \\ |
|
477 |
@{attribute_def cong} & : & \<open>attribute\<close> \\ |
|
478 |
@{command_def "print_simpset"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\ |
|
26782 | 479 |
\end{matharray} |
480 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
481 |
@{rail \<open> |
63650 | 482 |
(@@{attribute simp} | @@{attribute cong}) (() | 'add' | 'del') | |
483 |
@@{attribute split} (() | '!' | 'del') |
|
59917
9830c944670f
more uniform "verbose" option to print name space;
wenzelm
parents:
59905
diff
changeset
|
484 |
; |
9830c944670f
more uniform "verbose" option to print name space;
wenzelm
parents:
59905
diff
changeset
|
485 |
@@{command print_simpset} ('!'?) |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
486 |
\<close>} |
26782 | 487 |
|
62655 | 488 |
\<^descr> @{attribute simp} declares rewrite rules, by adding or deleting them from |
489 |
the simpset within the theory or proof context. Rewrite rules are theorems |
|
490 |
expressing some form of equality, for example: |
|
50076 | 491 |
|
61493 | 492 |
\<open>Suc ?m + ?n = ?m + Suc ?n\<close> \\ |
493 |
\<open>?P \<and> ?P \<longleftrightarrow> ?P\<close> \\ |
|
494 |
\<open>?A \<union> ?B \<equiv> {x. x \<in> ?A \<or> x \<in> ?B}\<close> |
|
50076 | 495 |
|
61421 | 496 |
\<^medskip> |
62655 | 497 |
Conditional rewrites such as \<open>?m < ?n \<Longrightarrow> ?m div ?n = 0\<close> are also permitted; |
498 |
the conditions can be arbitrary formulas. |
|
50076 | 499 |
|
61421 | 500 |
\<^medskip> |
62655 | 501 |
Internally, all rewrite rules are translated into Pure equalities, theorems |
502 |
with conclusion \<open>lhs \<equiv> rhs\<close>. The simpset contains a function for extracting |
|
503 |
equalities from arbitrary theorems, which is usually installed when the |
|
504 |
object-logic is configured initially. For example, \<open>\<not> ?x \<in> {}\<close> could be |
|
505 |
turned into \<open>?x \<in> {} \<equiv> False\<close>. Theorems that are declared as @{attribute |
|
506 |
simp} and local assumptions within a goal are treated uniformly in this |
|
507 |
respect. |
|
50076 | 508 |
|
62655 | 509 |
The Simplifier accepts the following formats for the \<open>lhs\<close> term: |
50076 | 510 |
|
62655 | 511 |
\<^enum> First-order patterns, considering the sublanguage of application of |
512 |
constant operators to variable operands, without \<open>\<lambda>\<close>-abstractions or |
|
513 |
functional variables. For example: |
|
50076 | 514 |
|
61493 | 515 |
\<open>(?x + ?y) + ?z \<equiv> ?x + (?y + ?z)\<close> \\ |
516 |
\<open>f (f ?x ?y) ?z \<equiv> f ?x (f ?y ?z)\<close> |
|
50076 | 517 |
|
62655 | 518 |
\<^enum> Higher-order patterns in the sense of @{cite "nipkow-patterns"}. These |
519 |
are terms in \<open>\<beta>\<close>-normal form (this will always be the case unless you have |
|
520 |
done something strange) where each occurrence of an unknown is of the form |
|
521 |
\<open>?F x\<^sub>1 \<dots> x\<^sub>n\<close>, where the \<open>x\<^sub>i\<close> are distinct bound variables. |
|
50076 | 522 |
|
62655 | 523 |
For example, \<open>(\<forall>x. ?P x \<and> ?Q x) \<equiv> (\<forall>x. ?P x) \<and> (\<forall>x. ?Q x)\<close> or its |
524 |
symmetric form, since the \<open>rhs\<close> is also a higher-order pattern. |
|
50076 | 525 |
|
62655 | 526 |
\<^enum> Physical first-order patterns over raw \<open>\<lambda>\<close>-term structure without |
527 |
\<open>\<alpha>\<beta>\<eta>\<close>-equality; abstractions and bound variables are treated like |
|
528 |
quasi-constant term material. |
|
50076 | 529 |
|
62655 | 530 |
For example, the rule \<open>?f ?x \<in> range ?f = True\<close> rewrites the term \<open>g a \<in> |
531 |
range g\<close> to \<open>True\<close>, but will fail to match \<open>g (h b) \<in> range (\<lambda>x. g (h |
|
532 |
x))\<close>. However, offending subterms (in our case \<open>?f ?x\<close>, which is not a |
|
533 |
pattern) can be replaced by adding new variables and conditions like this: |
|
534 |
\<open>?y = ?f ?x \<Longrightarrow> ?y \<in> range ?f = True\<close> is acceptable as a conditional rewrite |
|
535 |
rule of the second category since conditions can be arbitrary terms. |
|
50076 | 536 |
|
61439 | 537 |
\<^descr> @{attribute split} declares case split rules. |
26782 | 538 |
|
62655 | 539 |
\<^descr> @{attribute cong} declares congruence rules to the Simplifier context. |
45645 | 540 |
|
541 |
Congruence rules are equalities of the form @{text [display] |
|
542 |
"\<dots> \<Longrightarrow> f ?x\<^sub>1 \<dots> ?x\<^sub>n = f ?y\<^sub>1 \<dots> ?y\<^sub>n"} |
|
543 |
||
62655 | 544 |
This controls the simplification of the arguments of \<open>f\<close>. For example, some |
545 |
arguments can be simplified under additional assumptions: |
|
546 |
@{text [display] |
|
547 |
"?P\<^sub>1 \<longleftrightarrow> ?Q\<^sub>1 \<Longrightarrow> |
|
548 |
(?Q\<^sub>1 \<Longrightarrow> ?P\<^sub>2 \<longleftrightarrow> ?Q\<^sub>2) \<Longrightarrow> |
|
549 |
(?P\<^sub>1 \<longrightarrow> ?P\<^sub>2) \<longleftrightarrow> (?Q\<^sub>1 \<longrightarrow> ?Q\<^sub>2)"} |
|
45645 | 550 |
|
62655 | 551 |
Given this rule, the Simplifier assumes \<open>?Q\<^sub>1\<close> and extracts rewrite rules |
552 |
from it when simplifying \<open>?P\<^sub>2\<close>. Such local assumptions are effective for |
|
553 |
rewriting formulae such as \<open>x = 0 \<longrightarrow> y + x = y\<close>. |
|
45645 | 554 |
|
555 |
%FIXME |
|
556 |
%The local assumptions are also provided as theorems to the solver; |
|
557 |
%see \secref{sec:simp-solver} below. |
|
558 |
||
61421 | 559 |
\<^medskip> |
62655 | 560 |
The following congruence rule for bounded quantifiers also supplies |
561 |
contextual information --- about the bound variable: @{text [display] |
|
562 |
"(?A = ?B) \<Longrightarrow> |
|
563 |
(\<And>x. x \<in> ?B \<Longrightarrow> ?P x \<longleftrightarrow> ?Q x) \<Longrightarrow> |
|
45645 | 564 |
(\<forall>x \<in> ?A. ?P x) \<longleftrightarrow> (\<forall>x \<in> ?B. ?Q x)"} |
565 |
||
61421 | 566 |
\<^medskip> |
62655 | 567 |
This congruence rule for conditional expressions can supply contextual |
568 |
information for simplifying the arms: @{text [display] |
|
569 |
"?p = ?q \<Longrightarrow> |
|
570 |
(?q \<Longrightarrow> ?a = ?c) \<Longrightarrow> |
|
571 |
(\<not> ?q \<Longrightarrow> ?b = ?d) \<Longrightarrow> |
|
45645 | 572 |
(if ?p then ?a else ?b) = (if ?q then ?c else ?d)"} |
573 |
||
62655 | 574 |
A congruence rule can also \<^emph>\<open>prevent\<close> simplification of some arguments. Here |
575 |
is an alternative congruence rule for conditional expressions that conforms |
|
576 |
to non-strict functional evaluation: @{text [display] |
|
577 |
"?p = ?q \<Longrightarrow> |
|
578 |
(if ?p then ?a else ?b) = (if ?q then ?a else ?b)"} |
|
45645 | 579 |
|
62655 | 580 |
Only the first argument is simplified; the others remain unchanged. This can |
581 |
make simplification much faster, but may require an extra case split over |
|
582 |
the condition \<open>?q\<close> to prove the goal. |
|
50063 | 583 |
|
62655 | 584 |
\<^descr> @{command "print_simpset"} prints the collection of rules declared to the |
585 |
Simplifier, which is also known as ``simpset'' internally; the ``\<open>!\<close>'' |
|
586 |
option indicates extra verbosity. |
|
50077 | 587 |
|
62655 | 588 |
The implicit simpset of the theory context is propagated monotonically |
589 |
through the theory hierarchy: forming a new theory, the union of the |
|
590 |
simpsets of its imports are taken as starting point. Also note that |
|
591 |
definitional packages like @{command "datatype"}, @{command "primrec"}, |
|
592 |
@{command "fun"} routinely declare Simplifier rules to the target context, |
|
593 |
while plain @{command "definition"} is an exception in \<^emph>\<open>not\<close> declaring |
|
50065 | 594 |
anything. |
595 |
||
61421 | 596 |
\<^medskip> |
62655 | 597 |
It is up the user to manipulate the current simpset further by explicitly |
598 |
adding or deleting theorems as simplification rules, or installing other |
|
599 |
tools via simplification procedures (\secref{sec:simproc}). Good simpsets |
|
600 |
are hard to design. Rules that obviously simplify, like \<open>?n + 0 \<equiv> ?n\<close> are |
|
601 |
good candidates for the implicit simpset, unless a special non-normalizing |
|
602 |
behavior of certain operations is intended. More specific rules (such as |
|
603 |
distributive laws, which duplicate subterms) should be added only for |
|
604 |
specific proof steps. Conversely, sometimes a rule needs to be deleted just |
|
605 |
for some part of a proof. The need of frequent additions or deletions may |
|
606 |
indicate a poorly designed simpset. |
|
50065 | 607 |
|
608 |
\begin{warn} |
|
62655 | 609 |
The union of simpsets from theory imports (as described above) is not always |
610 |
a good starting point for the new theory. If some ancestors have deleted |
|
611 |
simplification rules because they are no longer wanted, while others have |
|
612 |
left those rules in, then the union will contain the unwanted rules, and |
|
613 |
thus have to be deleted again in the theory body. |
|
50065 | 614 |
\end{warn} |
58618 | 615 |
\<close> |
45645 | 616 |
|
617 |
||
58618 | 618 |
subsection \<open>Ordered rewriting with permutative rules\<close> |
50080 | 619 |
|
62655 | 620 |
text \<open> |
621 |
A rewrite rule is \<^emph>\<open>permutative\<close> if the left-hand side and right-hand side |
|
622 |
are the equal up to renaming of variables. The most common permutative rule |
|
623 |
is commutativity: \<open>?x + ?y = ?y + ?x\<close>. Other examples include \<open>(?x - ?y) - |
|
624 |
?z = (?x - ?z) - ?y\<close> in arithmetic and \<open>insert ?x (insert ?y ?A) = insert ?y |
|
625 |
(insert ?x ?A)\<close> for sets. Such rules are common enough to merit special |
|
626 |
attention. |
|
50080 | 627 |
|
62655 | 628 |
Because ordinary rewriting loops given such rules, the Simplifier employs a |
629 |
special strategy, called \<^emph>\<open>ordered rewriting\<close>. Permutative rules are |
|
630 |
detected and only applied if the rewriting step decreases the redex wrt.\ a |
|
631 |
given term ordering. For example, commutativity rewrites \<open>b + a\<close> to \<open>a + b\<close>, |
|
632 |
but then stops, because the redex cannot be decreased further in the sense |
|
633 |
of the term ordering. |
|
50080 | 634 |
|
62655 | 635 |
The default is lexicographic ordering of term structure, but this could be |
636 |
also changed locally for special applications via @{index_ML |
|
637 |
Simplifier.set_termless} in Isabelle/ML. |
|
50080 | 638 |
|
61421 | 639 |
\<^medskip> |
62655 | 640 |
Permutative rewrite rules are declared to the Simplifier just like other |
641 |
rewrite rules. Their special status is recognized automatically, and their |
|
642 |
application is guarded by the term ordering accordingly. |
|
643 |
\<close> |
|
50080 | 644 |
|
645 |
||
58618 | 646 |
subsubsection \<open>Rewriting with AC operators\<close> |
50080 | 647 |
|
62655 | 648 |
text \<open> |
649 |
Ordered rewriting is particularly effective in the case of |
|
650 |
associative-commutative operators. (Associativity by itself is not |
|
651 |
permutative.) When dealing with an AC-operator \<open>f\<close>, keep the following |
|
652 |
points in mind: |
|
50080 | 653 |
|
62655 | 654 |
\<^item> The associative law must always be oriented from left to right, namely |
655 |
\<open>f (f x y) z = f x (f y z)\<close>. The opposite orientation, if used with |
|
656 |
commutativity, leads to looping in conjunction with the standard term |
|
657 |
order. |
|
50080 | 658 |
|
62655 | 659 |
\<^item> To complete your set of rewrite rules, you must add not just |
660 |
associativity (A) and commutativity (C) but also a derived rule |
|
661 |
\<^emph>\<open>left-commutativity\<close> (LC): \<open>f x (f y z) = f y (f x z)\<close>. |
|
50080 | 662 |
|
663 |
Ordered rewriting with the combination of A, C, and LC sorts a term |
|
664 |
lexicographically --- the rewriting engine imitates bubble-sort. |
|
58618 | 665 |
\<close> |
50080 | 666 |
|
59905 | 667 |
experiment |
50080 | 668 |
fixes f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infix "\<bullet>" 60) |
669 |
assumes assoc: "(x \<bullet> y) \<bullet> z = x \<bullet> (y \<bullet> z)" |
|
670 |
assumes commute: "x \<bullet> y = y \<bullet> x" |
|
671 |
begin |
|
672 |
||
673 |
lemma left_commute: "x \<bullet> (y \<bullet> z) = y \<bullet> (x \<bullet> z)" |
|
674 |
proof - |
|
675 |
have "(x \<bullet> y) \<bullet> z = (y \<bullet> x) \<bullet> z" by (simp only: commute) |
|
676 |
then show ?thesis by (simp only: assoc) |
|
677 |
qed |
|
678 |
||
679 |
lemmas AC_rules = assoc commute left_commute |
|
680 |
||
62655 | 681 |
text \<open> |
682 |
Thus the Simplifier is able to establish equalities with arbitrary |
|
683 |
permutations of subterms, by normalizing to a common standard form. For |
|
684 |
example: |
|
685 |
\<close> |
|
50080 | 686 |
|
687 |
lemma "(b \<bullet> c) \<bullet> a = xxx" |
|
688 |
apply (simp only: AC_rules) |
|
58618 | 689 |
txt \<open>@{subgoals}\<close> |
50080 | 690 |
oops |
691 |
||
692 |
lemma "(b \<bullet> c) \<bullet> a = a \<bullet> (b \<bullet> c)" by (simp only: AC_rules) |
|
693 |
lemma "(b \<bullet> c) \<bullet> a = c \<bullet> (b \<bullet> a)" by (simp only: AC_rules) |
|
694 |
lemma "(b \<bullet> c) \<bullet> a = (c \<bullet> b) \<bullet> a" by (simp only: AC_rules) |
|
695 |
||
696 |
end |
|
697 |
||
62655 | 698 |
text \<open> |
699 |
Martin and Nipkow @{cite "martin-nipkow"} discuss the theory and give many |
|
700 |
examples; other algebraic structures are amenable to ordered rewriting, such |
|
701 |
as Boolean rings. The Boyer-Moore theorem prover @{cite bm88book} also |
|
702 |
employs ordered rewriting. |
|
58618 | 703 |
\<close> |
50080 | 704 |
|
705 |
||
58618 | 706 |
subsubsection \<open>Re-orienting equalities\<close> |
50080 | 707 |
|
58618 | 708 |
text \<open>Another application of ordered rewriting uses the derived rule |
50080 | 709 |
@{thm [source] eq_commute}: @{thm [source = false] eq_commute} to |
710 |
reverse equations. |
|
711 |
||
712 |
This is occasionally useful to re-orient local assumptions according |
|
713 |
to the term ordering, when other built-in mechanisms of |
|
58618 | 714 |
reorientation and mutual simplification fail to apply.\<close> |
50080 | 715 |
|
716 |
||
58618 | 717 |
subsection \<open>Simplifier tracing and debugging \label{sec:simp-trace}\<close> |
50063 | 718 |
|
58618 | 719 |
text \<open> |
50063 | 720 |
\begin{tabular}{rcll} |
61493 | 721 |
@{attribute_def simp_trace} & : & \<open>attribute\<close> & default \<open>false\<close> \\ |
722 |
@{attribute_def simp_trace_depth_limit} & : & \<open>attribute\<close> & default \<open>1\<close> \\ |
|
723 |
@{attribute_def simp_debug} & : & \<open>attribute\<close> & default \<open>false\<close> \\ |
|
724 |
@{attribute_def simp_trace_new} & : & \<open>attribute\<close> \\ |
|
725 |
@{attribute_def simp_break} & : & \<open>attribute\<close> \\ |
|
50063 | 726 |
\end{tabular} |
61421 | 727 |
\<^medskip> |
50063 | 728 |
|
57591
8c095aef6769
clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents:
57590
diff
changeset
|
729 |
@{rail \<open> |
8c095aef6769
clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents:
57590
diff
changeset
|
730 |
@@{attribute simp_trace_new} ('interactive')? \<newline> |
8c095aef6769
clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents:
57590
diff
changeset
|
731 |
('mode' '=' ('full' | 'normal'))? \<newline> |
8c095aef6769
clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents:
57590
diff
changeset
|
732 |
('depth' '=' @{syntax nat})? |
8c095aef6769
clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents:
57590
diff
changeset
|
733 |
; |
8c095aef6769
clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents:
57590
diff
changeset
|
734 |
|
8c095aef6769
clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents:
57590
diff
changeset
|
735 |
@@{attribute simp_break} (@{syntax term}*) |
8c095aef6769
clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents:
57590
diff
changeset
|
736 |
\<close>} |
8c095aef6769
clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents:
57590
diff
changeset
|
737 |
|
8c095aef6769
clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents:
57590
diff
changeset
|
738 |
These attributes and configurations options control various aspects of |
8c095aef6769
clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents:
57590
diff
changeset
|
739 |
Simplifier tracing and debugging. |
50063 | 740 |
|
62655 | 741 |
\<^descr> @{attribute simp_trace} makes the Simplifier output internal operations. |
742 |
This includes rewrite steps, but also bookkeeping like modifications of the |
|
743 |
simpset. |
|
50063 | 744 |
|
62655 | 745 |
\<^descr> @{attribute simp_trace_depth_limit} limits the effect of @{attribute |
746 |
simp_trace} to the given depth of recursive Simplifier invocations (when |
|
747 |
solving conditions of rewrite rules). |
|
50063 | 748 |
|
62655 | 749 |
\<^descr> @{attribute simp_debug} makes the Simplifier output some extra information |
750 |
about internal operations. This includes any attempted invocation of |
|
751 |
simplification procedures. |
|
50063 | 752 |
|
61439 | 753 |
\<^descr> @{attribute simp_trace_new} controls Simplifier tracing within |
58552 | 754 |
Isabelle/PIDE applications, notably Isabelle/jEdit @{cite "isabelle-jedit"}. |
62655 | 755 |
This provides a hierarchical representation of the rewriting steps performed |
756 |
by the Simplifier. |
|
57591
8c095aef6769
clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents:
57590
diff
changeset
|
757 |
|
8c095aef6769
clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents:
57590
diff
changeset
|
758 |
Users can configure the behaviour by specifying breakpoints, verbosity and |
8c095aef6769
clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents:
57590
diff
changeset
|
759 |
enabling or disabling the interactive mode. In normal verbosity (the |
62655 | 760 |
default), only rule applications matching a breakpoint will be shown to the |
761 |
user. In full verbosity, all rule applications will be logged. Interactive |
|
762 |
mode interrupts the normal flow of the Simplifier and defers the decision |
|
763 |
how to continue to the user via some GUI dialog. |
|
57591
8c095aef6769
clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents:
57590
diff
changeset
|
764 |
|
61439 | 765 |
\<^descr> @{attribute simp_break} declares term or theorem breakpoints for |
57591
8c095aef6769
clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents:
57590
diff
changeset
|
766 |
@{attribute simp_trace_new} as described above. Term breakpoints are |
8c095aef6769
clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents:
57590
diff
changeset
|
767 |
patterns which are checked for matches on the redex of a rule application. |
8c095aef6769
clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents:
57590
diff
changeset
|
768 |
Theorem breakpoints trigger when the corresponding theorem is applied in a |
8c095aef6769
clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents:
57590
diff
changeset
|
769 |
rewrite step. For example: |
58618 | 770 |
\<close> |
50063 | 771 |
|
59905 | 772 |
(*<*)experiment begin(*>*) |
57591
8c095aef6769
clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents:
57590
diff
changeset
|
773 |
declare conjI [simp_break] |
57590 | 774 |
declare [[simp_break "?x \<and> ?y"]] |
59905 | 775 |
(*<*)end(*>*) |
57590 | 776 |
|
50063 | 777 |
|
58618 | 778 |
subsection \<open>Simplification procedures \label{sec:simproc}\<close> |
26782 | 779 |
|
62655 | 780 |
text \<open> |
781 |
Simplification procedures are ML functions that produce proven rewrite rules |
|
782 |
on demand. They are associated with higher-order patterns that approximate |
|
783 |
the left-hand sides of equations. The Simplifier first matches the current |
|
784 |
redex against one of the LHS patterns; if this succeeds, the corresponding |
|
785 |
ML function is invoked, passing the Simplifier context and redex term. Thus |
|
786 |
rules may be specifically fashioned for particular situations, resulting in |
|
787 |
a more powerful mechanism than term rewriting by a fixed set of rules. |
|
42925 | 788 |
|
62655 | 789 |
Any successful result needs to be a (possibly conditional) rewrite rule \<open>t \<equiv> |
790 |
u\<close> that is applicable to the current redex. The rule will be applied just as |
|
791 |
any ordinary rewrite rule. It is expected to be already in \<^emph>\<open>internal form\<close>, |
|
792 |
bypassing the automatic preprocessing of object-level equivalences. |
|
42925 | 793 |
|
26782 | 794 |
\begin{matharray}{rcl} |
61493 | 795 |
@{command_def "simproc_setup"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\ |
796 |
simproc & : & \<open>attribute\<close> \\ |
|
26782 | 797 |
\end{matharray} |
798 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
799 |
@{rail \<open> |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
800 |
@@{command simproc_setup} @{syntax name} '(' (@{syntax term} + '|') ')' '=' |
62913 | 801 |
@{syntax text}; |
26782 | 802 |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
803 |
@@{attribute simproc} (('add' ':')? | 'del' ':') (@{syntax name}+) |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
804 |
\<close>} |
26782 | 805 |
|
62655 | 806 |
\<^descr> @{command "simproc_setup"} defines a named simplification procedure that |
807 |
is invoked by the Simplifier whenever any of the given term patterns match |
|
808 |
the current redex. The implementation, which is provided as ML source text, |
|
809 |
needs to be of type |
|
810 |
@{ML_type "morphism -> Proof.context -> cterm -> thm option"}, where the |
|
811 |
@{ML_type cterm} represents the current redex \<open>r\<close> and the result is supposed |
|
812 |
to be some proven rewrite rule \<open>r \<equiv> r'\<close> (or a generalized version), or @{ML |
|
813 |
NONE} to indicate failure. The @{ML_type Proof.context} argument holds the |
|
814 |
full context of the current Simplifier invocation. The @{ML_type morphism} |
|
815 |
informs about the difference of the original compilation context wrt.\ the |
|
62913 | 816 |
one of the actual application later on. |
26782 | 817 |
|
62913 | 818 |
Morphisms are only relevant for simprocs that are defined within a local |
819 |
target context, e.g.\ in a locale. |
|
26782 | 820 |
|
62655 | 821 |
\<^descr> \<open>simproc add: name\<close> and \<open>simproc del: name\<close> add or delete named simprocs |
822 |
to the current Simplifier context. The default is to add a simproc. Note |
|
823 |
that @{command "simproc_setup"} already adds the new simproc to the |
|
824 |
subsequent context. |
|
58618 | 825 |
\<close> |
26782 | 826 |
|
827 |
||
58618 | 828 |
subsubsection \<open>Example\<close> |
42925 | 829 |
|
62655 | 830 |
text \<open> |
831 |
The following simplification procedure for @{thm [source = false, |
|
832 |
show_types] unit_eq} in HOL performs fine-grained control over rule |
|
833 |
application, beyond higher-order pattern matching. Declaring @{thm unit_eq} |
|
834 |
as @{attribute simp} directly would make the Simplifier loop! Note that a |
|
835 |
version of this simplification procedure is already active in Isabelle/HOL. |
|
836 |
\<close> |
|
42925 | 837 |
|
59905 | 838 |
(*<*)experiment begin(*>*) |
59782 | 839 |
simproc_setup unit ("x::unit") = |
840 |
\<open>fn _ => fn _ => fn ct => |
|
59582 | 841 |
if HOLogic.is_unit (Thm.term_of ct) then NONE |
59782 | 842 |
else SOME (mk_meta_eq @{thm unit_eq})\<close> |
59905 | 843 |
(*<*)end(*>*) |
42925 | 844 |
|
62655 | 845 |
text \<open> |
846 |
Since the Simplifier applies simplification procedures frequently, it is |
|
847 |
important to make the failure check in ML reasonably fast.\<close> |
|
42925 | 848 |
|
849 |
||
58618 | 850 |
subsection \<open>Configurable Simplifier strategies \label{sec:simp-strategies}\<close> |
50079 | 851 |
|
62655 | 852 |
text \<open> |
853 |
The core term-rewriting engine of the Simplifier is normally used in |
|
854 |
combination with some add-on components that modify the strategy and allow |
|
855 |
to integrate other non-Simplifier proof tools. These may be reconfigured in |
|
856 |
ML as explained below. Even if the default strategies of object-logics like |
|
857 |
Isabelle/HOL are used unchanged, it helps to understand how the standard |
|
858 |
Simplifier strategies work.\<close> |
|
50079 | 859 |
|
860 |
||
58618 | 861 |
subsubsection \<open>The subgoaler\<close> |
50079 | 862 |
|
58618 | 863 |
text \<open> |
50079 | 864 |
\begin{mldecls} |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
865 |
@{index_ML Simplifier.set_subgoaler: "(Proof.context -> int -> tactic) -> |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
866 |
Proof.context -> Proof.context"} \\ |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
867 |
@{index_ML Simplifier.prems_of: "Proof.context -> thm list"} \\ |
50079 | 868 |
\end{mldecls} |
869 |
||
870 |
The subgoaler is the tactic used to solve subgoals arising out of |
|
62655 | 871 |
conditional rewrite rules or congruence rules. The default should be |
872 |
simplification itself. In rare situations, this strategy may need to be |
|
873 |
changed. For example, if the premise of a conditional rule is an instance of |
|
874 |
its conclusion, as in \<open>Suc ?m < ?n \<Longrightarrow> ?m < ?n\<close>, the default strategy could |
|
875 |
loop. % FIXME !?? |
|
50079 | 876 |
|
62655 | 877 |
\<^descr> @{ML Simplifier.set_subgoaler}~\<open>tac ctxt\<close> sets the subgoaler of the |
878 |
context to \<open>tac\<close>. The tactic will be applied to the context of the running |
|
879 |
Simplifier instance. |
|
50079 | 880 |
|
62655 | 881 |
\<^descr> @{ML Simplifier.prems_of}~\<open>ctxt\<close> retrieves the current set of premises |
882 |
from the context. This may be non-empty only if the Simplifier has been |
|
883 |
told to utilize local assumptions in the first place (cf.\ the options in |
|
884 |
\secref{sec:simp-meth}). |
|
50079 | 885 |
|
886 |
As an example, consider the following alternative subgoaler: |
|
58618 | 887 |
\<close> |
50079 | 888 |
|
59905 | 889 |
ML_val \<open> |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
890 |
fun subgoaler_tac ctxt = |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
58618
diff
changeset
|
891 |
assume_tac ctxt ORELSE' |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58963
diff
changeset
|
892 |
resolve_tac ctxt (Simplifier.prems_of ctxt) ORELSE' |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
893 |
asm_simp_tac ctxt |
58618 | 894 |
\<close> |
50079 | 895 |
|
62655 | 896 |
text \<open> |
897 |
This tactic first tries to solve the subgoal by assumption or by resolving |
|
898 |
with with one of the premises, calling simplification only if that fails.\<close> |
|
50079 | 899 |
|
900 |
||
58618 | 901 |
subsubsection \<open>The solver\<close> |
50079 | 902 |
|
58618 | 903 |
text \<open> |
50079 | 904 |
\begin{mldecls} |
905 |
@{index_ML_type solver} \\ |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
906 |
@{index_ML Simplifier.mk_solver: "string -> |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
907 |
(Proof.context -> int -> tactic) -> solver"} \\ |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
908 |
@{index_ML_op setSolver: "Proof.context * solver -> Proof.context"} \\ |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
909 |
@{index_ML_op addSolver: "Proof.context * solver -> Proof.context"} \\ |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
910 |
@{index_ML_op setSSolver: "Proof.context * solver -> Proof.context"} \\ |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
911 |
@{index_ML_op addSSolver: "Proof.context * solver -> Proof.context"} \\ |
50079 | 912 |
\end{mldecls} |
913 |
||
62655 | 914 |
A solver is a tactic that attempts to solve a subgoal after simplification. |
915 |
Its core functionality is to prove trivial subgoals such as @{prop "True"} |
|
916 |
and \<open>t = t\<close>, but object-logics might be more ambitious. For example, |
|
917 |
Isabelle/HOL performs a restricted version of linear arithmetic here. |
|
50079 | 918 |
|
62655 | 919 |
Solvers are packaged up in abstract type @{ML_type solver}, with @{ML |
920 |
Simplifier.mk_solver} as the only operation to create a solver. |
|
50079 | 921 |
|
61421 | 922 |
\<^medskip> |
62655 | 923 |
Rewriting does not instantiate unknowns. For example, rewriting alone cannot |
924 |
prove \<open>a \<in> ?A\<close> since this requires instantiating \<open>?A\<close>. The solver, however, |
|
925 |
is an arbitrary tactic and may instantiate unknowns as it pleases. This is |
|
926 |
the only way the Simplifier can handle a conditional rewrite rule whose |
|
927 |
condition contains extra variables. When a simplification tactic is to be |
|
928 |
combined with other provers, especially with the Classical Reasoner, it is |
|
929 |
important whether it can be considered safe or not. For this reason a |
|
930 |
simpset contains two solvers: safe and unsafe. |
|
50079 | 931 |
|
62655 | 932 |
The standard simplification strategy solely uses the unsafe solver, which is |
933 |
appropriate in most cases. For special applications where the simplification |
|
934 |
process is not allowed to instantiate unknowns within the goal, |
|
935 |
simplification starts with the safe solver, but may still apply the ordinary |
|
936 |
unsafe one in nested simplifications for conditional rules or congruences. |
|
937 |
Note that in this way the overall tactic is not totally safe: it may |
|
938 |
instantiate unknowns that appear also in other subgoals. |
|
50079 | 939 |
|
62655 | 940 |
\<^descr> @{ML Simplifier.mk_solver}~\<open>name tac\<close> turns \<open>tac\<close> into a solver; the |
941 |
\<open>name\<close> is only attached as a comment and has no further significance. |
|
50079 | 942 |
|
62655 | 943 |
\<^descr> \<open>ctxt setSSolver solver\<close> installs \<open>solver\<close> as the safe solver of \<open>ctxt\<close>. |
50079 | 944 |
|
62655 | 945 |
\<^descr> \<open>ctxt addSSolver solver\<close> adds \<open>solver\<close> as an additional safe solver; it |
946 |
will be tried after the solvers which had already been present in \<open>ctxt\<close>. |
|
50079 | 947 |
|
62655 | 948 |
\<^descr> \<open>ctxt setSolver solver\<close> installs \<open>solver\<close> as the unsafe solver of \<open>ctxt\<close>. |
949 |
||
950 |
\<^descr> \<open>ctxt addSolver solver\<close> adds \<open>solver\<close> as an additional unsafe solver; it |
|
951 |
will be tried after the solvers which had already been present in \<open>ctxt\<close>. |
|
50079 | 952 |
|
953 |
||
61421 | 954 |
\<^medskip> |
62655 | 955 |
The solver tactic is invoked with the context of the running Simplifier. |
956 |
Further operations may be used to retrieve relevant information, such as the |
|
957 |
list of local Simplifier premises via @{ML Simplifier.prems_of} --- this |
|
958 |
list may be non-empty only if the Simplifier runs in a mode that utilizes |
|
959 |
local assumptions (see also \secref{sec:simp-meth}). The solver is also |
|
960 |
presented the full goal including its assumptions in any case. Thus it can |
|
961 |
use these (e.g.\ by calling @{ML assume_tac}), even if the Simplifier proper |
|
962 |
happens to ignore local premises at the moment. |
|
50079 | 963 |
|
61421 | 964 |
\<^medskip> |
62655 | 965 |
As explained before, the subgoaler is also used to solve the premises of |
966 |
congruence rules. These are usually of the form \<open>s = ?x\<close>, where \<open>s\<close> needs to |
|
967 |
be simplified and \<open>?x\<close> needs to be instantiated with the result. Typically, |
|
50079 | 968 |
the subgoaler will invoke the Simplifier at some point, which will |
62655 | 969 |
eventually call the solver. For this reason, solver tactics must be prepared |
970 |
to solve goals of the form \<open>t = ?x\<close>, usually by reflexivity. In particular, |
|
971 |
reflexivity should be tried before any of the fancy automated proof tools. |
|
50079 | 972 |
|
62655 | 973 |
It may even happen that due to simplification the subgoal is no longer an |
974 |
equality. For example, \<open>False \<longleftrightarrow> ?Q\<close> could be rewritten to \<open>\<not> ?Q\<close>. To cover |
|
975 |
this case, the solver could try resolving with the theorem \<open>\<not> False\<close> of the |
|
50079 | 976 |
object-logic. |
977 |
||
61421 | 978 |
\<^medskip> |
50079 | 979 |
\begin{warn} |
62655 | 980 |
If a premise of a congruence rule cannot be proved, then the congruence is |
981 |
ignored. This should only happen if the rule is \<^emph>\<open>conditional\<close> --- that is, |
|
982 |
contains premises not of the form \<open>t = ?x\<close>. Otherwise it indicates that some |
|
983 |
congruence rule, or possibly the subgoaler or solver, is faulty. |
|
50079 | 984 |
\end{warn} |
58618 | 985 |
\<close> |
50079 | 986 |
|
987 |
||
58618 | 988 |
subsubsection \<open>The looper\<close> |
50079 | 989 |
|
58618 | 990 |
text \<open> |
50079 | 991 |
\begin{mldecls} |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
992 |
@{index_ML_op setloop: "Proof.context * |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
993 |
(Proof.context -> int -> tactic) -> Proof.context"} \\ |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
994 |
@{index_ML_op addloop: "Proof.context * |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
995 |
(string * (Proof.context -> int -> tactic)) |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
996 |
-> Proof.context"} \\ |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
997 |
@{index_ML_op delloop: "Proof.context * string -> Proof.context"} \\ |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
998 |
@{index_ML Splitter.add_split: "thm -> Proof.context -> Proof.context"} \\ |
63650 | 999 |
@{index_ML Splitter.add_split: "thm -> Proof.context -> Proof.context"} \\ |
1000 |
@{index_ML Splitter.add_split_bang: " |
|
1001 |
thm -> Proof.context -> Proof.context"} \\ |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
1002 |
@{index_ML Splitter.del_split: "thm -> Proof.context -> Proof.context"} \\ |
50079 | 1003 |
\end{mldecls} |
1004 |
||
62655 | 1005 |
The looper is a list of tactics that are applied after simplification, in |
1006 |
case the solver failed to solve the simplified goal. If the looper succeeds, |
|
1007 |
the simplification process is started all over again. Each of the subgoals |
|
1008 |
generated by the looper is attacked in turn, in reverse order. |
|
50079 | 1009 |
|
62655 | 1010 |
A typical looper is \<^emph>\<open>case splitting\<close>: the expansion of a conditional. |
1011 |
Another possibility is to apply an elimination rule on the assumptions. More |
|
1012 |
adventurous loopers could start an induction. |
|
50079 | 1013 |
|
62655 | 1014 |
\<^descr> \<open>ctxt setloop tac\<close> installs \<open>tac\<close> as the only looper tactic of \<open>ctxt\<close>. |
50079 | 1015 |
|
62655 | 1016 |
\<^descr> \<open>ctxt addloop (name, tac)\<close> adds \<open>tac\<close> as an additional looper tactic |
1017 |
with name \<open>name\<close>, which is significant for managing the collection of |
|
1018 |
loopers. The tactic will be tried after the looper tactics that had |
|
1019 |
already been present in \<open>ctxt\<close>. |
|
50079 | 1020 |
|
62655 | 1021 |
\<^descr> \<open>ctxt delloop name\<close> deletes the looper tactic that was associated with |
1022 |
\<open>name\<close> from \<open>ctxt\<close>. |
|
50079 | 1023 |
|
63650 | 1024 |
\<^descr> @{ML Splitter.add_split}~\<open>thm ctxt\<close> adds split tactic |
1025 |
for \<open>thm\<close> as additional looper tactic of \<open>ctxt\<close> |
|
1026 |
(overwriting previous split tactic for the same constant). |
|
1027 |
||
1028 |
\<^descr> @{ML Splitter.add_split_bang}~\<open>thm ctxt\<close> adds aggressive |
|
1029 |
(see \S\ref{sec:simp-meth}) |
|
1030 |
split tactic for \<open>thm\<close> as additional looper tactic of \<open>ctxt\<close> |
|
1031 |
(overwriting previous split tactic for the same constant). |
|
50079 | 1032 |
|
62655 | 1033 |
\<^descr> @{ML Splitter.del_split}~\<open>thm ctxt\<close> deletes the split tactic |
1034 |
corresponding to \<open>thm\<close> from the looper tactics of \<open>ctxt\<close>. |
|
50079 | 1035 |
|
62655 | 1036 |
The splitter replaces applications of a given function; the right-hand side |
1037 |
of the replacement can be anything. For example, here is a splitting rule |
|
1038 |
for conditional expressions: |
|
50079 | 1039 |
|
1040 |
@{text [display] "?P (if ?Q ?x ?y) \<longleftrightarrow> (?Q \<longrightarrow> ?P ?x) \<and> (\<not> ?Q \<longrightarrow> ?P ?y)"} |
|
1041 |
||
62655 | 1042 |
Another example is the elimination operator for Cartesian products (which |
1043 |
happens to be called @{const case_prod} in Isabelle/HOL: |
|
50079 | 1044 |
|
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
61421
diff
changeset
|
1045 |
@{text [display] "?P (case_prod ?f ?p) \<longleftrightarrow> (\<forall>a b. ?p = (a, b) \<longrightarrow> ?P (f a b))"} |
50079 | 1046 |
|
62655 | 1047 |
For technical reasons, there is a distinction between case splitting in the |
1048 |
conclusion and in the premises of a subgoal. The former is done by @{ML |
|
1049 |
Splitter.split_tac} with rules like @{thm [source] if_split} or @{thm |
|
1050 |
[source] option.split}, which do not split the subgoal, while the latter is |
|
1051 |
done by @{ML Splitter.split_asm_tac} with rules like @{thm [source] |
|
1052 |
if_split_asm} or @{thm [source] option.split_asm}, which split the subgoal. |
|
1053 |
The function @{ML Splitter.add_split} automatically takes care of which |
|
1054 |
tactic to call, analyzing the form of the rules given as argument; it is the |
|
1055 |
same operation behind \<open>split\<close> attribute or method modifier syntax in the |
|
1056 |
Isar source language. |
|
50079 | 1057 |
|
62655 | 1058 |
Case splits should be allowed only when necessary; they are expensive and |
1059 |
hard to control. Case-splitting on if-expressions in the conclusion is |
|
1060 |
usually beneficial, so it is enabled by default in Isabelle/HOL and |
|
1061 |
Isabelle/FOL/ZF. |
|
50079 | 1062 |
|
1063 |
\begin{warn} |
|
62655 | 1064 |
With @{ML Splitter.split_asm_tac} as looper component, the Simplifier may |
1065 |
split subgoals! This might cause unexpected problems in tactic expressions |
|
1066 |
that silently assume 0 or 1 subgoals after simplification. |
|
50079 | 1067 |
\end{warn} |
58618 | 1068 |
\<close> |
50079 | 1069 |
|
1070 |
||
58618 | 1071 |
subsection \<open>Forward simplification \label{sec:simp-forward}\<close> |
26782 | 1072 |
|
58618 | 1073 |
text \<open> |
26782 | 1074 |
\begin{matharray}{rcl} |
61493 | 1075 |
@{attribute_def simplified} & : & \<open>attribute\<close> \\ |
26782 | 1076 |
\end{matharray} |
1077 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1078 |
@{rail \<open> |
62969 | 1079 |
@@{attribute simplified} opt? @{syntax thms}? |
26782 | 1080 |
; |
1081 |
||
40255
9ffbc25e1606
eliminated obsolete \_ escapes in rail environments;
wenzelm
parents:
35613
diff
changeset
|
1082 |
opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use') ')' |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1083 |
\<close>} |
26782 | 1084 |
|
62655 | 1085 |
\<^descr> @{attribute simplified}~\<open>a\<^sub>1 \<dots> a\<^sub>n\<close> causes a theorem to be simplified, |
1086 |
either by exactly the specified rules \<open>a\<^sub>1, \<dots>, a\<^sub>n\<close>, or the implicit |
|
1087 |
Simplifier context if no arguments are given. The result is fully simplified |
|
1088 |
by default, including assumptions and conclusion; the options \<open>no_asm\<close> etc.\ |
|
1089 |
tune the Simplifier in the same way as the for the \<open>simp\<close> method. |
|
26782 | 1090 |
|
62655 | 1091 |
Note that forward simplification restricts the Simplifier to its most basic |
1092 |
operation of term rewriting; solver and looper tactics |
|
1093 |
(\secref{sec:simp-strategies}) are \<^emph>\<open>not\<close> involved here. The @{attribute |
|
1094 |
simplified} attribute should be only rarely required under normal |
|
1095 |
circumstances. |
|
58618 | 1096 |
\<close> |
26782 | 1097 |
|
1098 |
||
58618 | 1099 |
section \<open>The Classical Reasoner \label{sec:classical}\<close> |
26782 | 1100 |
|
58618 | 1101 |
subsection \<open>Basic concepts\<close> |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1102 |
|
58618 | 1103 |
text \<open>Although Isabelle is generic, many users will be working in |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1104 |
some extension of classical first-order logic. Isabelle/ZF is built |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1105 |
upon theory FOL, while Isabelle/HOL conceptually contains |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1106 |
first-order logic as a fragment. Theorem-proving in predicate logic |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1107 |
is undecidable, but many automated strategies have been developed to |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1108 |
assist in this task. |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1109 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1110 |
Isabelle's classical reasoner is a generic package that accepts |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1111 |
certain information about a logic and delivers a suite of automatic |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1112 |
proof tools, based on rules that are classified and declared in the |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1113 |
context. These proof procedures are slow and simplistic compared |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1114 |
with high-end automated theorem provers, but they can save |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1115 |
considerable time and effort in practice. They can prove theorems |
58552 | 1116 |
such as Pelletier's @{cite pelletier86} problems 40 and 41 in a few |
58618 | 1117 |
milliseconds (including full proof reconstruction):\<close> |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1118 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1119 |
lemma "(\<exists>y. \<forall>x. F x y \<longleftrightarrow> F x x) \<longrightarrow> \<not> (\<forall>x. \<exists>y. \<forall>z. F z y \<longleftrightarrow> \<not> F z x)" |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1120 |
by blast |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1121 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1122 |
lemma "(\<forall>z. \<exists>y. \<forall>x. f x y \<longleftrightarrow> f x z \<and> \<not> f x x) \<longrightarrow> \<not> (\<exists>z. \<forall>x. f x z)" |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1123 |
by blast |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1124 |
|
58618 | 1125 |
text \<open>The proof tools are generic. They are not restricted to |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1126 |
first-order logic, and have been heavily used in the development of |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1127 |
the Isabelle/HOL library and applications. The tactics can be |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1128 |
traced, and their components can be called directly; in this manner, |
58618 | 1129 |
any proof can be viewed interactively.\<close> |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1130 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1131 |
|
58618 | 1132 |
subsubsection \<open>The sequent calculus\<close> |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1133 |
|
58618 | 1134 |
text \<open>Isabelle supports natural deduction, which is easy to use for |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1135 |
interactive proof. But natural deduction does not easily lend |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1136 |
itself to automation, and has a bias towards intuitionism. For |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1137 |
certain proofs in classical logic, it can not be called natural. |
61477 | 1138 |
The \<^emph>\<open>sequent calculus\<close>, a generalization of natural deduction, |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1139 |
is easier to automate. |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1140 |
|
61493 | 1141 |
A \<^bold>\<open>sequent\<close> has the form \<open>\<Gamma> \<turnstile> \<Delta>\<close>, where \<open>\<Gamma>\<close> |
61572 | 1142 |
and \<open>\<Delta>\<close> are sets of formulae.\<^footnote>\<open>For first-order |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1143 |
logic, sequents can equivalently be made from lists or multisets of |
61572 | 1144 |
formulae.\<close> The sequent \<open>P\<^sub>1, \<dots>, P\<^sub>m \<turnstile> Q\<^sub>1, \<dots>, Q\<^sub>n\<close> is |
61493 | 1145 |
\<^bold>\<open>valid\<close> if \<open>P\<^sub>1 \<and> \<dots> \<and> P\<^sub>m\<close> implies \<open>Q\<^sub>1 \<or> \<dots> \<or> |
1146 |
Q\<^sub>n\<close>. Thus \<open>P\<^sub>1, \<dots>, P\<^sub>m\<close> represent assumptions, each of which |
|
1147 |
is true, while \<open>Q\<^sub>1, \<dots>, Q\<^sub>n\<close> represent alternative goals. A |
|
61477 | 1148 |
sequent is \<^bold>\<open>basic\<close> if its left and right sides have a common |
61493 | 1149 |
formula, as in \<open>P, Q \<turnstile> Q, R\<close>; basic sequents are trivially |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1150 |
valid. |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1151 |
|
61477 | 1152 |
Sequent rules are classified as \<^bold>\<open>right\<close> or \<^bold>\<open>left\<close>, |
61493 | 1153 |
indicating which side of the \<open>\<turnstile>\<close> symbol they operate on. |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1154 |
Rules that operate on the right side are analogous to natural |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1155 |
deduction's introduction rules, and left rules are analogous to |
61493 | 1156 |
elimination rules. The sequent calculus analogue of \<open>(\<longrightarrow>I)\<close> |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1157 |
is the rule |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1158 |
\[ |
61493 | 1159 |
\infer[\<open>(\<longrightarrow>R)\<close>]{\<open>\<Gamma> \<turnstile> \<Delta>, P \<longrightarrow> Q\<close>}{\<open>P, \<Gamma> \<turnstile> \<Delta>, Q\<close>} |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1160 |
\] |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1161 |
Applying the rule backwards, this breaks down some implication on |
61493 | 1162 |
the right side of a sequent; \<open>\<Gamma>\<close> and \<open>\<Delta>\<close> stand for |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1163 |
the sets of formulae that are unaffected by the inference. The |
61493 | 1164 |
analogue of the pair \<open>(\<or>I1)\<close> and \<open>(\<or>I2)\<close> is the |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1165 |
single rule |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1166 |
\[ |
61493 | 1167 |
\infer[\<open>(\<or>R)\<close>]{\<open>\<Gamma> \<turnstile> \<Delta>, P \<or> Q\<close>}{\<open>\<Gamma> \<turnstile> \<Delta>, P, Q\<close>} |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1168 |
\] |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1169 |
This breaks down some disjunction on the right side, replacing it by |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1170 |
both disjuncts. Thus, the sequent calculus is a kind of |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1171 |
multiple-conclusion logic. |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1172 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1173 |
To illustrate the use of multiple formulae on the right, let us |
61493 | 1174 |
prove the classical theorem \<open>(P \<longrightarrow> Q) \<or> (Q \<longrightarrow> P)\<close>. Working |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1175 |
backwards, we reduce this formula to a basic sequent: |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1176 |
\[ |
61493 | 1177 |
\infer[\<open>(\<or>R)\<close>]{\<open>\<turnstile> (P \<longrightarrow> Q) \<or> (Q \<longrightarrow> P)\<close>} |
1178 |
{\infer[\<open>(\<longrightarrow>R)\<close>]{\<open>\<turnstile> (P \<longrightarrow> Q), (Q \<longrightarrow> P)\<close>} |
|
1179 |
{\infer[\<open>(\<longrightarrow>R)\<close>]{\<open>P \<turnstile> Q, (Q \<longrightarrow> P)\<close>} |
|
1180 |
{\<open>P, Q \<turnstile> Q, P\<close>}}} |
|
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1181 |
\] |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1182 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1183 |
This example is typical of the sequent calculus: start with the |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1184 |
desired theorem and apply rules backwards in a fairly arbitrary |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1185 |
manner. This yields a surprisingly effective proof procedure. |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1186 |
Quantifiers add only few complications, since Isabelle handles |
58552 | 1187 |
parameters and schematic variables. See @{cite \<open>Chapter 10\<close> |
58618 | 1188 |
"paulson-ml2"} for further discussion.\<close> |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1189 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1190 |
|
58618 | 1191 |
subsubsection \<open>Simulating sequents by natural deduction\<close> |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1192 |
|
58618 | 1193 |
text \<open>Isabelle can represent sequents directly, as in the |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1194 |
object-logic LK. But natural deduction is easier to work with, and |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1195 |
most object-logics employ it. Fortunately, we can simulate the |
61493 | 1196 |
sequent \<open>P\<^sub>1, \<dots>, P\<^sub>m \<turnstile> Q\<^sub>1, \<dots>, Q\<^sub>n\<close> by the Isabelle formula |
1197 |
\<open>P\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> P\<^sub>m \<Longrightarrow> \<not> Q\<^sub>2 \<Longrightarrow> ... \<Longrightarrow> \<not> Q\<^sub>n \<Longrightarrow> Q\<^sub>1\<close> where the order of |
|
1198 |
the assumptions and the choice of \<open>Q\<^sub>1\<close> are arbitrary. |
|
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1199 |
Elim-resolution plays a key role in simulating sequent proofs. |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1200 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1201 |
We can easily handle reasoning on the left. Elim-resolution with |
61493 | 1202 |
the rules \<open>(\<or>E)\<close>, \<open>(\<bottom>E)\<close> and \<open>(\<exists>E)\<close> achieves |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1203 |
a similar effect as the corresponding sequent rules. For the other |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1204 |
connectives, we use sequent-style elimination rules instead of |
61493 | 1205 |
destruction rules such as \<open>(\<and>E1, 2)\<close> and \<open>(\<forall>E)\<close>. |
1206 |
But note that the rule \<open>(\<not>L)\<close> has no effect under our |
|
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1207 |
representation of sequents! |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1208 |
\[ |
61493 | 1209 |
\infer[\<open>(\<not>L)\<close>]{\<open>\<not> P, \<Gamma> \<turnstile> \<Delta>\<close>}{\<open>\<Gamma> \<turnstile> \<Delta>, P\<close>} |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1210 |
\] |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1211 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1212 |
What about reasoning on the right? Introduction rules can only |
61493 | 1213 |
affect the formula in the conclusion, namely \<open>Q\<^sub>1\<close>. The |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1214 |
other right-side formulae are represented as negated assumptions, |
61493 | 1215 |
\<open>\<not> Q\<^sub>2, \<dots>, \<not> Q\<^sub>n\<close>. In order to operate on one of these, it |
1216 |
must first be exchanged with \<open>Q\<^sub>1\<close>. Elim-resolution with the |
|
1217 |
\<open>swap\<close> rule has this effect: \<open>\<not> P \<Longrightarrow> (\<not> R \<Longrightarrow> P) \<Longrightarrow> R\<close> |
|
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1218 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1219 |
To ensure that swaps occur only when necessary, each introduction |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1220 |
rule is converted into a swapped form: it is resolved with the |
61493 | 1221 |
second premise of \<open>(swap)\<close>. The swapped form of \<open>(\<and>I)\<close>, which might be called \<open>(\<not>\<and>E)\<close>, is |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1222 |
@{text [display] "\<not> (P \<and> Q) \<Longrightarrow> (\<not> R \<Longrightarrow> P) \<Longrightarrow> (\<not> R \<Longrightarrow> Q) \<Longrightarrow> R"} |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1223 |
|
61493 | 1224 |
Similarly, the swapped form of \<open>(\<longrightarrow>I)\<close> is |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1225 |
@{text [display] "\<not> (P \<longrightarrow> Q) \<Longrightarrow> (\<not> R \<Longrightarrow> P \<Longrightarrow> Q) \<Longrightarrow> R"} |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1226 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1227 |
Swapped introduction rules are applied using elim-resolution, which |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1228 |
deletes the negated formula. Our representation of sequents also |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1229 |
requires the use of ordinary introduction rules. If we had no |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1230 |
regard for readability of intermediate goal states, we could treat |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1231 |
the right side more uniformly by representing sequents as @{text |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1232 |
[display] "P\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> P\<^sub>m \<Longrightarrow> \<not> Q\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> \<not> Q\<^sub>n \<Longrightarrow> \<bottom>"} |
58618 | 1233 |
\<close> |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1234 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1235 |
|
58618 | 1236 |
subsubsection \<open>Extra rules for the sequent calculus\<close> |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1237 |
|
61493 | 1238 |
text \<open>As mentioned, destruction rules such as \<open>(\<and>E1, 2)\<close> and |
1239 |
\<open>(\<forall>E)\<close> must be replaced by sequent-style elimination rules. |
|
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1240 |
In addition, we need rules to embody the classical equivalence |
61493 | 1241 |
between \<open>P \<longrightarrow> Q\<close> and \<open>\<not> P \<or> Q\<close>. The introduction |
1242 |
rules \<open>(\<or>I1, 2)\<close> are replaced by a rule that simulates |
|
1243 |
\<open>(\<or>R)\<close>: @{text [display] "(\<not> Q \<Longrightarrow> P) \<Longrightarrow> P \<or> Q"} |
|
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1244 |
|
61493 | 1245 |
The destruction rule \<open>(\<longrightarrow>E)\<close> is replaced by @{text [display] |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1246 |
"(P \<longrightarrow> Q) \<Longrightarrow> (\<not> P \<Longrightarrow> R) \<Longrightarrow> (Q \<Longrightarrow> R) \<Longrightarrow> R"} |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1247 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1248 |
Quantifier replication also requires special rules. In classical |
61493 | 1249 |
logic, \<open>\<exists>x. P x\<close> is equivalent to \<open>\<not> (\<forall>x. \<not> P x)\<close>; |
1250 |
the rules \<open>(\<exists>R)\<close> and \<open>(\<forall>L)\<close> are dual: |
|
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1251 |
\[ |
61493 | 1252 |
\infer[\<open>(\<exists>R)\<close>]{\<open>\<Gamma> \<turnstile> \<Delta>, \<exists>x. P x\<close>}{\<open>\<Gamma> \<turnstile> \<Delta>, \<exists>x. P x, P t\<close>} |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1253 |
\qquad |
61493 | 1254 |
\infer[\<open>(\<forall>L)\<close>]{\<open>\<forall>x. P x, \<Gamma> \<turnstile> \<Delta>\<close>}{\<open>P t, \<forall>x. P x, \<Gamma> \<turnstile> \<Delta>\<close>} |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1255 |
\] |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1256 |
Thus both kinds of quantifier may be replicated. Theorems requiring |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1257 |
multiple uses of a universal formula are easy to invent; consider |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1258 |
@{text [display] "(\<forall>x. P x \<longrightarrow> P (f x)) \<and> P a \<longrightarrow> P (f\<^sup>n a)"} for any |
61493 | 1259 |
\<open>n > 1\<close>. Natural examples of the multiple use of an |
1260 |
existential formula are rare; a standard one is \<open>\<exists>x. \<forall>y. P x |
|
1261 |
\<longrightarrow> P y\<close>. |
|
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1262 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1263 |
Forgoing quantifier replication loses completeness, but gains |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1264 |
decidability, since the search space becomes finite. Many useful |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1265 |
theorems can be proved without replication, and the search generally |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1266 |
delivers its verdict in a reasonable time. To adopt this approach, |
61493 | 1267 |
represent the sequent rules \<open>(\<exists>R)\<close>, \<open>(\<exists>L)\<close> and |
1268 |
\<open>(\<forall>R)\<close> by \<open>(\<exists>I)\<close>, \<open>(\<exists>E)\<close> and \<open>(\<forall>I)\<close>, |
|
1269 |
respectively, and put \<open>(\<forall>E)\<close> into elimination form: @{text |
|
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1270 |
[display] "\<forall>x. P x \<Longrightarrow> (P t \<Longrightarrow> Q) \<Longrightarrow> Q"} |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1271 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1272 |
Elim-resolution with this rule will delete the universal formula |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1273 |
after a single use. To replicate universal quantifiers, replace the |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1274 |
rule by @{text [display] "\<forall>x. P x \<Longrightarrow> (P t \<Longrightarrow> \<forall>x. P x \<Longrightarrow> Q) \<Longrightarrow> Q"} |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1275 |
|
61493 | 1276 |
To replicate existential quantifiers, replace \<open>(\<exists>I)\<close> by |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1277 |
@{text [display] "(\<not> (\<exists>x. P x) \<Longrightarrow> P t) \<Longrightarrow> \<exists>x. P x"} |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1278 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1279 |
All introduction rules mentioned above are also useful in swapped |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1280 |
form. |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1281 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1282 |
Replication makes the search space infinite; we must apply the rules |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1283 |
with care. The classical reasoner distinguishes between safe and |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1284 |
unsafe rules, applying the latter only when there is no alternative. |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1285 |
Depth-first search may well go down a blind alley; best-first search |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1286 |
is better behaved in an infinite search space. However, quantifier |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1287 |
replication is too expensive to prove any but the simplest theorems. |
58618 | 1288 |
\<close> |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1289 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
1290 |
|
58618 | 1291 |
subsection \<open>Rule declarations\<close> |
42928
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1292 |
|
58618 | 1293 |
text \<open>The proof tools of the Classical Reasoner depend on |
42928
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1294 |
collections of rules declared in the context, which are classified |
61477 | 1295 |
as introduction, elimination or destruction and as \<^emph>\<open>safe\<close> or |
1296 |
\<^emph>\<open>unsafe\<close>. In general, safe rules can be attempted blindly, |
|
42928
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1297 |
while unsafe rules must be used with care. A safe rule must never |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1298 |
reduce a provable goal to an unprovable set of subgoals. |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1299 |
|
61493 | 1300 |
The rule \<open>P \<Longrightarrow> P \<or> Q\<close> is unsafe because it reduces \<open>P |
1301 |
\<or> Q\<close> to \<open>P\<close>, which might turn out as premature choice of an |
|
42928
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1302 |
unprovable subgoal. Any rule is unsafe whose premises contain new |
61493 | 1303 |
unknowns. The elimination rule \<open>\<forall>x. P x \<Longrightarrow> (P t \<Longrightarrow> Q) \<Longrightarrow> Q\<close> is |
42928
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1304 |
unsafe, since it is applied via elim-resolution, which discards the |
61493 | 1305 |
assumption \<open>\<forall>x. P x\<close> and replaces it by the weaker |
1306 |
assumption \<open>P t\<close>. The rule \<open>P t \<Longrightarrow> \<exists>x. P x\<close> is |
|
1307 |
unsafe for similar reasons. The quantifier duplication rule \<open>\<forall>x. P x \<Longrightarrow> (P t \<Longrightarrow> \<forall>x. P x \<Longrightarrow> Q) \<Longrightarrow> Q\<close> is unsafe in a different sense: |
|
1308 |
since it keeps the assumption \<open>\<forall>x. P x\<close>, it is prone to |
|
42928
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1309 |
looping. In classical first-order logic, all rules are safe except |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1310 |
those mentioned above. |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1311 |
|
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1312 |
The safe~/ unsafe distinction is vague, and may be regarded merely |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1313 |
as a way of giving some rules priority over others. One could argue |
61493 | 1314 |
that \<open>(\<or>E)\<close> is unsafe, because repeated application of it |
42928
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1315 |
could generate exponentially many subgoals. Induction rules are |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1316 |
unsafe because inductive proofs are difficult to set up |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1317 |
automatically. Any inference is unsafe that instantiates an unknown |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1318 |
in the proof state --- thus matching must be used, rather than |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1319 |
unification. Even proof by assumption is unsafe if it instantiates |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1320 |
unknowns shared with other subgoals. |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1321 |
|
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1322 |
\begin{matharray}{rcl} |
61493 | 1323 |
@{command_def "print_claset"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\ |
1324 |
@{attribute_def intro} & : & \<open>attribute\<close> \\ |
|
1325 |
@{attribute_def elim} & : & \<open>attribute\<close> \\ |
|
1326 |
@{attribute_def dest} & : & \<open>attribute\<close> \\ |
|
1327 |
@{attribute_def rule} & : & \<open>attribute\<close> \\ |
|
1328 |
@{attribute_def iff} & : & \<open>attribute\<close> \\ |
|
1329 |
@{attribute_def swapped} & : & \<open>attribute\<close> \\ |
|
42928
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1330 |
\end{matharray} |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1331 |
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1332 |
@{rail \<open> |
42928
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1333 |
(@@{attribute intro} | @@{attribute elim} | @@{attribute dest}) ('!' | () | '?') @{syntax nat}? |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1334 |
; |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1335 |
@@{attribute rule} 'del' |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1336 |
; |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1337 |
@@{attribute iff} (((() | 'add') '?'?) | 'del') |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1338 |
\<close>} |
42928
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1339 |
|
61439 | 1340 |
\<^descr> @{command "print_claset"} prints the collection of rules |
42928
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1341 |
declared to the Classical Reasoner, i.e.\ the @{ML_type claset} |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1342 |
within the context. |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1343 |
|
61439 | 1344 |
\<^descr> @{attribute intro}, @{attribute elim}, and @{attribute dest} |
42928
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1345 |
declare introduction, elimination, and destruction rules, |
61477 | 1346 |
respectively. By default, rules are considered as \<^emph>\<open>unsafe\<close> |
61493 | 1347 |
(i.e.\ not applied blindly without backtracking), while ``\<open>!\<close>'' classifies as \<^emph>\<open>safe\<close>. Rule declarations marked by |
1348 |
``\<open>?\<close>'' coincide with those of Isabelle/Pure, cf.\ |
|
42928
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1349 |
\secref{sec:pure-meth-att} (i.e.\ are only applied in single steps |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1350 |
of the @{method rule} method). The optional natural number |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1351 |
specifies an explicit weight argument, which is ignored by the |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1352 |
automated reasoning tools, but determines the search order of single |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1353 |
rule steps. |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1354 |
|
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1355 |
Introduction rules are those that can be applied using ordinary |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1356 |
resolution. Their swapped forms are generated internally, which |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1357 |
will be applied using elim-resolution. Elimination rules are |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1358 |
applied using elim-resolution. Rules are sorted by the number of |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1359 |
new subgoals they will yield; rules that generate the fewest |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1360 |
subgoals will be tried first. Otherwise, later declarations take |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1361 |
precedence over earlier ones. |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1362 |
|
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1363 |
Rules already present in the context with the same classification |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1364 |
are ignored. A warning is printed if the rule has already been |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1365 |
added with some other classification, but the rule is added anyway |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1366 |
as requested. |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1367 |
|
61493 | 1368 |
\<^descr> @{attribute rule}~\<open>del\<close> deletes all occurrences of a |
42928
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1369 |
rule from the classical context, regardless of its classification as |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1370 |
introduction~/ elimination~/ destruction and safe~/ unsafe. |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1371 |
|
61439 | 1372 |
\<^descr> @{attribute iff} declares logical equivalences to the |
42928
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1373 |
Simplifier and the Classical reasoner at the same time. |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1374 |
Non-conditional rules result in a safe introduction and elimination |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1375 |
pair; conditional ones are considered unsafe. Rules with negative |
61493 | 1376 |
conclusion are automatically inverted (using \<open>\<not>\<close>-elimination |
42928
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1377 |
internally). |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1378 |
|
61493 | 1379 |
The ``\<open>?\<close>'' version of @{attribute iff} declares rules to |
42928
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1380 |
the Isabelle/Pure context only, and omits the Simplifier |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1381 |
declaration. |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1382 |
|
61439 | 1383 |
\<^descr> @{attribute swapped} turns an introduction rule into an |
61493 | 1384 |
elimination, by resolving with the classical swap principle \<open>\<not> P \<Longrightarrow> (\<not> R \<Longrightarrow> P) \<Longrightarrow> R\<close> in the second position. This is mainly for |
42928
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1385 |
illustrative purposes: the Classical Reasoner already swaps rules |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1386 |
internally as explained above. |
58618 | 1387 |
\<close> |
26782 | 1388 |
|
1389 |
||
58618 | 1390 |
subsection \<open>Structured methods\<close> |
43365 | 1391 |
|
58618 | 1392 |
text \<open> |
43365 | 1393 |
\begin{matharray}{rcl} |
61493 | 1394 |
@{method_def rule} & : & \<open>method\<close> \\ |
1395 |
@{method_def contradiction} & : & \<open>method\<close> \\ |
|
43365 | 1396 |
\end{matharray} |
1397 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1398 |
@{rail \<open> |
62969 | 1399 |
@@{method rule} @{syntax thms}? |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1400 |
\<close>} |
43365 | 1401 |
|
61439 | 1402 |
\<^descr> @{method rule} as offered by the Classical Reasoner is a |
43365 | 1403 |
refinement over the Pure one (see \secref{sec:pure-meth-att}). Both |
1404 |
versions work the same, but the classical version observes the |
|
1405 |
classical rule context in addition to that of Isabelle/Pure. |
|
1406 |
||
1407 |
Common object logics (HOL, ZF, etc.) declare a rich collection of |
|
1408 |
classical rules (even if these would qualify as intuitionistic |
|
1409 |
ones), but only few declarations to the rule context of |
|
1410 |
Isabelle/Pure (\secref{sec:pure-meth-att}). |
|
1411 |
||
61439 | 1412 |
\<^descr> @{method contradiction} solves some goal by contradiction, |
61493 | 1413 |
deriving any result from both \<open>\<not> A\<close> and \<open>A\<close>. Chained |
43365 | 1414 |
facts, which are guaranteed to participate, may appear in either |
1415 |
order. |
|
58618 | 1416 |
\<close> |
43365 | 1417 |
|
1418 |
||
58618 | 1419 |
subsection \<open>Fully automated methods\<close> |
26782 | 1420 |
|
58618 | 1421 |
text \<open> |
26782 | 1422 |
\begin{matharray}{rcl} |
61493 | 1423 |
@{method_def blast} & : & \<open>method\<close> \\ |
1424 |
@{method_def auto} & : & \<open>method\<close> \\ |
|
1425 |
@{method_def force} & : & \<open>method\<close> \\ |
|
1426 |
@{method_def fast} & : & \<open>method\<close> \\ |
|
1427 |
@{method_def slow} & : & \<open>method\<close> \\ |
|
1428 |
@{method_def best} & : & \<open>method\<close> \\ |
|
1429 |
@{method_def fastforce} & : & \<open>method\<close> \\ |
|
1430 |
@{method_def slowsimp} & : & \<open>method\<close> \\ |
|
1431 |
@{method_def bestsimp} & : & \<open>method\<close> \\ |
|
1432 |
@{method_def deepen} & : & \<open>method\<close> \\ |
|
26782 | 1433 |
\end{matharray} |
1434 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1435 |
@{rail \<open> |
42930 | 1436 |
@@{method blast} @{syntax nat}? (@{syntax clamod} * ) |
1437 |
; |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
1438 |
@@{method auto} (@{syntax nat} @{syntax nat})? (@{syntax clasimpmod} * ) |
26782 | 1439 |
; |
42930 | 1440 |
@@{method force} (@{syntax clasimpmod} * ) |
1441 |
; |
|
1442 |
(@@{method fast} | @@{method slow} | @@{method best}) (@{syntax clamod} * ) |
|
26782 | 1443 |
; |
44911 | 1444 |
(@@{method fastforce} | @@{method slowsimp} | @@{method bestsimp}) |
42930 | 1445 |
(@{syntax clasimpmod} * ) |
1446 |
; |
|
43367 | 1447 |
@@{method deepen} (@{syntax nat} ?) (@{syntax clamod} * ) |
1448 |
; |
|
42930 | 1449 |
@{syntax_def clamod}: |
62969 | 1450 |
(('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del') ':' @{syntax thms} |
42930 | 1451 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
1452 |
@{syntax_def clasimpmod}: ('simp' (() | 'add' | 'del' | 'only') | |
63650 | 1453 |
'cong' (() | 'add' | 'del') | |
1454 |
'split' (() | '!' | 'del') | |
|
26782 | 1455 |
'iff' (((() | 'add') '?'?) | 'del') | |
62969 | 1456 |
(('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del')) ':' @{syntax thms} |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1457 |
\<close>} |
26782 | 1458 |
|
61439 | 1459 |
\<^descr> @{method blast} is a separate classical tableau prover that |
42930 | 1460 |
uses the same classical rule declarations as explained before. |
1461 |
||
1462 |
Proof search is coded directly in ML using special data structures. |
|
1463 |
A successful proof is then reconstructed using regular Isabelle |
|
1464 |
inferences. It is faster and more powerful than the other classical |
|
1465 |
reasoning tools, but has major limitations too. |
|
1466 |
||
61459
5f2ddeb15c06
clarified nesting of paragraphs: indentation is taken into account more uniformly;
wenzelm
parents:
61458
diff
changeset
|
1467 |
\<^item> It does not use the classical wrapper tacticals, such as the |
61458 | 1468 |
integration with the Simplifier of @{method fastforce}. |
42930 | 1469 |
|
61459
5f2ddeb15c06
clarified nesting of paragraphs: indentation is taken into account more uniformly;
wenzelm
parents:
61458
diff
changeset
|
1470 |
\<^item> It does not perform higher-order unification, as needed by the |
61458 | 1471 |
rule @{thm [source=false] rangeI} in HOL. There are often |
1472 |
alternatives to such rules, for example @{thm [source=false] |
|
1473 |
range_eqI}. |
|
42930 | 1474 |
|
61459
5f2ddeb15c06
clarified nesting of paragraphs: indentation is taken into account more uniformly;
wenzelm
parents:
61458
diff
changeset
|
1475 |
\<^item> Function variables may only be applied to parameters of the |
61458 | 1476 |
subgoal. (This restriction arises because the prover does not use |
1477 |
higher-order unification.) If other function variables are present |
|
1478 |
then the prover will fail with the message |
|
1479 |
@{verbatim [display] \<open>Function unknown's argument not a bound variable\<close>} |
|
42930 | 1480 |
|
61459
5f2ddeb15c06
clarified nesting of paragraphs: indentation is taken into account more uniformly;
wenzelm
parents:
61458
diff
changeset
|
1481 |
\<^item> Its proof strategy is more general than @{method fast} but can |
61458 | 1482 |
be slower. If @{method blast} fails or seems to be running forever, |
1483 |
try @{method fast} and the other proof tools described below. |
|
42930 | 1484 |
|
1485 |
The optional integer argument specifies a bound for the number of |
|
1486 |
unsafe steps used in a proof. By default, @{method blast} starts |
|
1487 |
with a bound of 0 and increases it successively to 20. In contrast, |
|
61493 | 1488 |
\<open>(blast lim)\<close> tries to prove the goal using a search bound |
1489 |
of \<open>lim\<close>. Sometimes a slow proof using @{method blast} can |
|
42930 | 1490 |
be made much faster by supplying the successful search bound to this |
1491 |
proof method instead. |
|
1492 |
||
61439 | 1493 |
\<^descr> @{method auto} combines classical reasoning with |
42930 | 1494 |
simplification. It is intended for situations where there are a lot |
1495 |
of mostly trivial subgoals; it proves all the easy ones, leaving the |
|
1496 |
ones it cannot prove. Occasionally, attempting to prove the hard |
|
1497 |
ones may take a long time. |
|
1498 |
||
61493 | 1499 |
The optional depth arguments in \<open>(auto m n)\<close> refer to its |
1500 |
builtin classical reasoning procedures: \<open>m\<close> (default 4) is for |
|
1501 |
@{method blast}, which is tried first, and \<open>n\<close> (default 2) is |
|
43332 | 1502 |
for a slower but more general alternative that also takes wrappers |
1503 |
into account. |
|
42930 | 1504 |
|
61439 | 1505 |
\<^descr> @{method force} is intended to prove the first subgoal |
42930 | 1506 |
completely, using many fancy proof tools and performing a rather |
1507 |
exhaustive search. As a result, proof attempts may take rather long |
|
1508 |
or diverge easily. |
|
1509 |
||
61439 | 1510 |
\<^descr> @{method fast}, @{method best}, @{method slow} attempt to |
42930 | 1511 |
prove the first subgoal using sequent-style reasoning as explained |
1512 |
before. Unlike @{method blast}, they construct proofs directly in |
|
1513 |
Isabelle. |
|
26782 | 1514 |
|
42930 | 1515 |
There is a difference in search strategy and back-tracking: @{method |
1516 |
fast} uses depth-first search and @{method best} uses best-first |
|
1517 |
search (guided by a heuristic function: normally the total size of |
|
1518 |
the proof state). |
|
1519 |
||
1520 |
Method @{method slow} is like @{method fast}, but conducts a broader |
|
1521 |
search: it may, when backtracking from a failed proof attempt, undo |
|
1522 |
even the step of proving a subgoal by assumption. |
|
1523 |
||
61439 | 1524 |
\<^descr> @{method fastforce}, @{method slowsimp}, @{method bestsimp} |
47967
c422128d3889
discontinued obsolete method fastsimp / tactic fast_simp_tac;
wenzelm
parents:
47497
diff
changeset
|
1525 |
are like @{method fast}, @{method slow}, @{method best}, |
c422128d3889
discontinued obsolete method fastsimp / tactic fast_simp_tac;
wenzelm
parents:
47497
diff
changeset
|
1526 |
respectively, but use the Simplifier as additional wrapper. The name |
c422128d3889
discontinued obsolete method fastsimp / tactic fast_simp_tac;
wenzelm
parents:
47497
diff
changeset
|
1527 |
@{method fastforce}, reflects the behaviour of this popular method |
c422128d3889
discontinued obsolete method fastsimp / tactic fast_simp_tac;
wenzelm
parents:
47497
diff
changeset
|
1528 |
better without requiring an understanding of its implementation. |
42930 | 1529 |
|
61439 | 1530 |
\<^descr> @{method deepen} works by exhaustive search up to a certain |
43367 | 1531 |
depth. The start depth is 4 (unless specified explicitly), and the |
1532 |
depth is increased iteratively up to 10. Unsafe rules are modified |
|
1533 |
to preserve the formula they act on, so that it be used repeatedly. |
|
1534 |
This method can prove more goals than @{method fast}, but is much |
|
1535 |
slower, for example if the assumptions have many universal |
|
1536 |
quantifiers. |
|
1537 |
||
42930 | 1538 |
|
1539 |
Any of the above methods support additional modifiers of the context |
|
1540 |
of classical (and simplifier) rules, but the ones related to the |
|
61493 | 1541 |
Simplifier are explicitly prefixed by \<open>simp\<close> here. The |
42930 | 1542 |
semantics of these ad-hoc rule declarations is analogous to the |
1543 |
attributes given before. Facts provided by forward chaining are |
|
1544 |
inserted into the goal before commencing proof search. |
|
58618 | 1545 |
\<close> |
42930 | 1546 |
|
1547 |
||
63650 | 1548 |
subsection \<open>Partially automated methods\label{sec:classical:partial}\<close> |
42930 | 1549 |
|
58618 | 1550 |
text \<open>These proof methods may help in situations when the |
42930 | 1551 |
fully-automated tools fail. The result is a simpler subgoal that |
1552 |
can be tackled by other means, such as by manual instantiation of |
|
1553 |
quantifiers. |
|
1554 |
||
1555 |
\begin{matharray}{rcl} |
|
61493 | 1556 |
@{method_def safe} & : & \<open>method\<close> \\ |
1557 |
@{method_def clarify} & : & \<open>method\<close> \\ |
|
1558 |
@{method_def clarsimp} & : & \<open>method\<close> \\ |
|
42930 | 1559 |
\end{matharray} |
1560 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1561 |
@{rail \<open> |
42930 | 1562 |
(@@{method safe} | @@{method clarify}) (@{syntax clamod} * ) |
1563 |
; |
|
1564 |
@@{method clarsimp} (@{syntax clasimpmod} * ) |
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1565 |
\<close>} |
42930 | 1566 |
|
61439 | 1567 |
\<^descr> @{method safe} repeatedly performs safe steps on all subgoals. |
42930 | 1568 |
It is deterministic, with at most one outcome. |
1569 |
||
61439 | 1570 |
\<^descr> @{method clarify} performs a series of safe steps without |
50108 | 1571 |
splitting subgoals; see also @{method clarify_step}. |
42930 | 1572 |
|
61439 | 1573 |
\<^descr> @{method clarsimp} acts like @{method clarify}, but also does |
42930 | 1574 |
simplification. Note that if the Simplifier context includes a |
1575 |
splitter for the premises, the subgoal may still be split. |
|
58618 | 1576 |
\<close> |
26782 | 1577 |
|
1578 |
||
58618 | 1579 |
subsection \<open>Single-step tactics\<close> |
43366 | 1580 |
|
58618 | 1581 |
text \<open> |
50108 | 1582 |
\begin{matharray}{rcl} |
61493 | 1583 |
@{method_def safe_step} & : & \<open>method\<close> \\ |
1584 |
@{method_def inst_step} & : & \<open>method\<close> \\ |
|
1585 |
@{method_def step} & : & \<open>method\<close> \\ |
|
1586 |
@{method_def slow_step} & : & \<open>method\<close> \\ |
|
1587 |
@{method_def clarify_step} & : & \<open>method\<close> \\ |
|
50108 | 1588 |
\end{matharray} |
43366 | 1589 |
|
50070
e447ad4d6edd
avoid clash of terminology wrt. "semi-automated" in the sense of Isar (e.g. method "rule");
wenzelm
parents:
50068
diff
changeset
|
1590 |
These are the primitive tactics behind the automated proof methods |
e447ad4d6edd
avoid clash of terminology wrt. "semi-automated" in the sense of Isar (e.g. method "rule");
wenzelm
parents:
50068
diff
changeset
|
1591 |
of the Classical Reasoner. By calling them yourself, you can |
e447ad4d6edd
avoid clash of terminology wrt. "semi-automated" in the sense of Isar (e.g. method "rule");
wenzelm
parents:
50068
diff
changeset
|
1592 |
execute these procedures one step at a time. |
43366 | 1593 |
|
61439 | 1594 |
\<^descr> @{method safe_step} performs a safe step on the first subgoal. |
50108 | 1595 |
The safe wrapper tacticals are applied to a tactic that may include |
1596 |
proof by assumption or Modus Ponens (taking care not to instantiate |
|
1597 |
unknowns), or substitution. |
|
43366 | 1598 |
|
61439 | 1599 |
\<^descr> @{method inst_step} is like @{method safe_step}, but allows |
43366 | 1600 |
unknowns to be instantiated. |
1601 |
||
61439 | 1602 |
\<^descr> @{method step} is the basic step of the proof procedure, it |
50108 | 1603 |
operates on the first subgoal. The unsafe wrapper tacticals are |
1604 |
applied to a tactic that tries @{method safe}, @{method inst_step}, |
|
1605 |
or applies an unsafe rule from the context. |
|
43366 | 1606 |
|
61439 | 1607 |
\<^descr> @{method slow_step} resembles @{method step}, but allows |
50108 | 1608 |
backtracking between using safe rules with instantiation (@{method |
1609 |
inst_step}) and using unsafe rules. The resulting search space is |
|
1610 |
larger. |
|
43366 | 1611 |
|
61439 | 1612 |
\<^descr> @{method clarify_step} performs a safe step on the first |
50108 | 1613 |
subgoal; no splitting step is applied. For example, the subgoal |
61493 | 1614 |
\<open>A \<and> B\<close> is left as a conjunction. Proof by assumption, |
50108 | 1615 |
Modus Ponens, etc., may be performed provided they do not |
61493 | 1616 |
instantiate unknowns. Assumptions of the form \<open>x = t\<close> may |
50108 | 1617 |
be eliminated. The safe wrapper tactical is applied. |
58618 | 1618 |
\<close> |
43366 | 1619 |
|
1620 |
||
58618 | 1621 |
subsection \<open>Modifying the search step\<close> |
50071 | 1622 |
|
58618 | 1623 |
text \<open> |
50071 | 1624 |
\begin{mldecls} |
1625 |
@{index_ML_type wrapper: "(int -> tactic) -> (int -> tactic)"} \\[0.5ex] |
|
51703
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents:
50108
diff
changeset
|
1626 |
@{index_ML_op addSWrapper: "Proof.context * |
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents:
50108
diff
changeset
|
1627 |
(string * (Proof.context -> wrapper)) -> Proof.context"} \\ |
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents:
50108
diff
changeset
|
1628 |
@{index_ML_op addSbefore: "Proof.context * |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
1629 |
(string * (Proof.context -> int -> tactic)) -> Proof.context"} \\ |
51703
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents:
50108
diff
changeset
|
1630 |
@{index_ML_op addSafter: "Proof.context * |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
1631 |
(string * (Proof.context -> int -> tactic)) -> Proof.context"} \\ |
51703
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents:
50108
diff
changeset
|
1632 |
@{index_ML_op delSWrapper: "Proof.context * string -> Proof.context"} \\[0.5ex] |
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents:
50108
diff
changeset
|
1633 |
@{index_ML_op addWrapper: "Proof.context * |
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents:
50108
diff
changeset
|
1634 |
(string * (Proof.context -> wrapper)) -> Proof.context"} \\ |
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents:
50108
diff
changeset
|
1635 |
@{index_ML_op addbefore: "Proof.context * |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
1636 |
(string * (Proof.context -> int -> tactic)) -> Proof.context"} \\ |
51703
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents:
50108
diff
changeset
|
1637 |
@{index_ML_op addafter: "Proof.context * |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51703
diff
changeset
|
1638 |
(string * (Proof.context -> int -> tactic)) -> Proof.context"} \\ |
51703
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents:
50108
diff
changeset
|
1639 |
@{index_ML_op delWrapper: "Proof.context * string -> Proof.context"} \\[0.5ex] |
50071 | 1640 |
@{index_ML addSss: "Proof.context -> Proof.context"} \\ |
1641 |
@{index_ML addss: "Proof.context -> Proof.context"} \\ |
|
1642 |
\end{mldecls} |
|
1643 |
||
1644 |
The proof strategy of the Classical Reasoner is simple. Perform as |
|
1645 |
many safe inferences as possible; or else, apply certain safe rules, |
|
1646 |
allowing instantiation of unknowns; or else, apply an unsafe rule. |
|
61493 | 1647 |
The tactics also eliminate assumptions of the form \<open>x = t\<close> |
50071 | 1648 |
by substitution if they have been set up to do so. They may perform |
61493 | 1649 |
a form of Modus Ponens: if there are assumptions \<open>P \<longrightarrow> Q\<close> and |
1650 |
\<open>P\<close>, then replace \<open>P \<longrightarrow> Q\<close> by \<open>Q\<close>. |
|
50071 | 1651 |
|
1652 |
The classical reasoning tools --- except @{method blast} --- allow |
|
1653 |
to modify this basic proof strategy by applying two lists of |
|
61477 | 1654 |
arbitrary \<^emph>\<open>wrapper tacticals\<close> to it. The first wrapper list, |
50108 | 1655 |
which is considered to contain safe wrappers only, affects @{method |
1656 |
safe_step} and all the tactics that call it. The second one, which |
|
1657 |
may contain unsafe wrappers, affects the unsafe parts of @{method |
|
1658 |
step}, @{method slow_step}, and the tactics that call them. A |
|
50071 | 1659 |
wrapper transforms each step of the search, for example by |
1660 |
attempting other tactics before or after the original step tactic. |
|
1661 |
All members of a wrapper list are applied in turn to the respective |
|
1662 |
step tactic. |
|
1663 |
||
1664 |
Initially the two wrapper lists are empty, which means no |
|
1665 |
modification of the step tactics. Safe and unsafe wrappers are added |
|
59905 | 1666 |
to the context with the functions given below, supplying them with |
50071 | 1667 |
wrapper names. These names may be used to selectively delete |
1668 |
wrappers. |
|
1669 |
||
61493 | 1670 |
\<^descr> \<open>ctxt addSWrapper (name, wrapper)\<close> adds a new wrapper, |
50071 | 1671 |
which should yield a safe tactic, to modify the existing safe step |
1672 |
tactic. |
|
1673 |
||
61493 | 1674 |
\<^descr> \<open>ctxt addSbefore (name, tac)\<close> adds the given tactic as a |
61477 | 1675 |
safe wrapper, such that it is tried \<^emph>\<open>before\<close> each safe step of |
50071 | 1676 |
the search. |
1677 |
||
61493 | 1678 |
\<^descr> \<open>ctxt addSafter (name, tac)\<close> adds the given tactic as a |
50071 | 1679 |
safe wrapper, such that it is tried when a safe step of the search |
1680 |
would fail. |
|
1681 |
||
61493 | 1682 |
\<^descr> \<open>ctxt delSWrapper name\<close> deletes the safe wrapper with |
50071 | 1683 |
the given name. |
1684 |
||
61493 | 1685 |
\<^descr> \<open>ctxt addWrapper (name, wrapper)\<close> adds a new wrapper to |
50071 | 1686 |
modify the existing (unsafe) step tactic. |
1687 |
||
61493 | 1688 |
\<^descr> \<open>ctxt addbefore (name, tac)\<close> adds the given tactic as an |
50071 | 1689 |
unsafe wrapper, such that it its result is concatenated |
61477 | 1690 |
\<^emph>\<open>before\<close> the result of each unsafe step. |
50071 | 1691 |
|
61493 | 1692 |
\<^descr> \<open>ctxt addafter (name, tac)\<close> adds the given tactic as an |
61477 | 1693 |
unsafe wrapper, such that it its result is concatenated \<^emph>\<open>after\<close> |
50071 | 1694 |
the result of each unsafe step. |
1695 |
||
61493 | 1696 |
\<^descr> \<open>ctxt delWrapper name\<close> deletes the unsafe wrapper with |
50071 | 1697 |
the given name. |
1698 |
||
61493 | 1699 |
\<^descr> \<open>addSss\<close> adds the simpset of the context to its |
50071 | 1700 |
classical set. The assumptions and goal will be simplified, in a |
1701 |
rather safe way, after each safe step of the search. |
|
1702 |
||
61493 | 1703 |
\<^descr> \<open>addss\<close> adds the simpset of the context to its |
50071 | 1704 |
classical set. The assumptions and goal will be simplified, before |
1705 |
the each unsafe step of the search. |
|
58618 | 1706 |
\<close> |
50071 | 1707 |
|
1708 |
||
58618 | 1709 |
section \<open>Object-logic setup \label{sec:object-logic}\<close> |
26790 | 1710 |
|
58618 | 1711 |
text \<open> |
26790 | 1712 |
\begin{matharray}{rcl} |
61493 | 1713 |
@{command_def "judgment"} & : & \<open>theory \<rightarrow> theory\<close> \\ |
1714 |
@{method_def atomize} & : & \<open>method\<close> \\ |
|
1715 |
@{attribute_def atomize} & : & \<open>attribute\<close> \\ |
|
1716 |
@{attribute_def rule_format} & : & \<open>attribute\<close> \\ |
|
1717 |
@{attribute_def rulify} & : & \<open>attribute\<close> \\ |
|
26790 | 1718 |
\end{matharray} |
1719 |
||
1720 |
The very starting point for any Isabelle object-logic is a ``truth |
|
1721 |
judgment'' that links object-level statements to the meta-logic |
|
61493 | 1722 |
(with its minimal language of \<open>prop\<close> that covers universal |
1723 |
quantification \<open>\<And>\<close> and implication \<open>\<Longrightarrow>\<close>). |
|
26790 | 1724 |
|
1725 |
Common object-logics are sufficiently expressive to internalize rule |
|
61493 | 1726 |
statements over \<open>\<And>\<close> and \<open>\<Longrightarrow>\<close> within their own |
26790 | 1727 |
language. This is useful in certain situations where a rule needs |
1728 |
to be viewed as an atomic statement from the meta-level perspective, |
|
61493 | 1729 |
e.g.\ \<open>\<And>x. x \<in> A \<Longrightarrow> P x\<close> versus \<open>\<forall>x \<in> A. P x\<close>. |
26790 | 1730 |
|
1731 |
From the following language elements, only the @{method atomize} |
|
1732 |
method and @{attribute rule_format} attribute are occasionally |
|
1733 |
required by end-users, the rest is for those who need to setup their |
|
1734 |
own object-logic. In the latter case existing formulations of |
|
1735 |
Isabelle/FOL or Isabelle/HOL may be taken as realistic examples. |
|
1736 |
||
1737 |
Generic tools may refer to the information provided by object-logic |
|
1738 |
declarations internally. |
|
1739 |
||
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1740 |
@{rail \<open> |
46494
ea2ae63336f3
clarified outer syntax "constdecl", which is only local to some rail diagrams;
wenzelm
parents:
46277
diff
changeset
|
1741 |
@@{command judgment} @{syntax name} '::' @{syntax type} @{syntax mixfix}? |
26790 | 1742 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
1743 |
@@{attribute atomize} ('(' 'full' ')')? |
26790 | 1744 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
1745 |
@@{attribute rule_format} ('(' 'noasm' ')')? |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
1746 |
\<close>} |
26790 | 1747 |
|
61493 | 1748 |
\<^descr> @{command "judgment"}~\<open>c :: \<sigma> (mx)\<close> declares constant |
1749 |
\<open>c\<close> as the truth judgment of the current object-logic. Its |
|
1750 |
type \<open>\<sigma>\<close> should specify a coercion of the category of |
|
1751 |
object-level propositions to \<open>prop\<close> of the Pure meta-logic; |
|
1752 |
the mixfix annotation \<open>(mx)\<close> would typically just link the |
|
1753 |
object language (internally of syntactic category \<open>logic\<close>) |
|
1754 |
with that of \<open>prop\<close>. Only one @{command "judgment"} |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
1755 |
declaration may be given in any theory development. |
26790 | 1756 |
|
61439 | 1757 |
\<^descr> @{method atomize} (as a method) rewrites any non-atomic |
26790 | 1758 |
premises of a sub-goal, using the meta-level equations declared via |
1759 |
@{attribute atomize} (as an attribute) beforehand. As a result, |
|
1760 |
heavily nested goals become amenable to fundamental operations such |
|
61493 | 1761 |
as resolution (cf.\ the @{method (Pure) rule} method). Giving the ``\<open>(full)\<close>'' option here means to turn the whole subgoal into an |
26790 | 1762 |
object-statement (if possible), including the outermost parameters |
1763 |
and assumptions as well. |
|
1764 |
||
1765 |
A typical collection of @{attribute atomize} rules for a particular |
|
1766 |
object-logic would provide an internalization for each of the |
|
61493 | 1767 |
connectives of \<open>\<And>\<close>, \<open>\<Longrightarrow>\<close>, and \<open>\<equiv>\<close>. |
26790 | 1768 |
Meta-level conjunction should be covered as well (this is |
1769 |
particularly important for locales, see \secref{sec:locale}). |
|
1770 |
||
61439 | 1771 |
\<^descr> @{attribute rule_format} rewrites a theorem by the equalities |
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
1772 |
declared as @{attribute rulify} rules in the current object-logic. |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
1773 |
By default, the result is fully normalized, including assumptions |
61493 | 1774 |
and conclusions at any depth. The \<open>(no_asm)\<close> option |
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
1775 |
restricts the transformation to the conclusion of a rule. |
26790 | 1776 |
|
1777 |
In common object-logics (HOL, FOL, ZF), the effect of @{attribute |
|
1778 |
rule_format} is to replace (bounded) universal quantification |
|
61493 | 1779 |
(\<open>\<forall>\<close>) and implication (\<open>\<longrightarrow>\<close>) by the corresponding |
1780 |
rule statements over \<open>\<And>\<close> and \<open>\<Longrightarrow>\<close>. |
|
58618 | 1781 |
\<close> |
26790 | 1782 |
|
50083 | 1783 |
|
58618 | 1784 |
section \<open>Tracing higher-order unification\<close> |
50083 | 1785 |
|
58618 | 1786 |
text \<open> |
50083 | 1787 |
\begin{tabular}{rcll} |
61493 | 1788 |
@{attribute_def unify_trace_simp} & : & \<open>attribute\<close> & default \<open>false\<close> \\ |
1789 |
@{attribute_def unify_trace_types} & : & \<open>attribute\<close> & default \<open>false\<close> \\ |
|
1790 |
@{attribute_def unify_trace_bound} & : & \<open>attribute\<close> & default \<open>50\<close> \\ |
|
1791 |
@{attribute_def unify_search_bound} & : & \<open>attribute\<close> & default \<open>60\<close> \\ |
|
50083 | 1792 |
\end{tabular} |
61421 | 1793 |
\<^medskip> |
50083 | 1794 |
|
1795 |
Higher-order unification works well in most practical situations, |
|
1796 |
but sometimes needs extra care to identify problems. These tracing |
|
1797 |
options may help. |
|
1798 |
||
61439 | 1799 |
\<^descr> @{attribute unify_trace_simp} controls tracing of the |
50083 | 1800 |
simplification phase of higher-order unification. |
1801 |
||
61439 | 1802 |
\<^descr> @{attribute unify_trace_types} controls warnings of |
50083 | 1803 |
incompleteness, when unification is not considering all possible |
1804 |
instantiations of schematic type variables. |
|
1805 |
||
61439 | 1806 |
\<^descr> @{attribute unify_trace_bound} determines the depth where |
50083 | 1807 |
unification starts to print tracing information once it reaches |
1808 |
depth; 0 for full tracing. At the default value, tracing |
|
1809 |
information is almost never printed in practice. |
|
1810 |
||
61439 | 1811 |
\<^descr> @{attribute unify_search_bound} prevents unification from |
50083 | 1812 |
searching past the given depth. Because of this bound, higher-order |
1813 |
unification cannot return an infinite sequence, though it can return |
|
1814 |
an exponentially long one. The search rarely approaches the default |
|
1815 |
value in practice. If the search is cut off, unification prints a |
|
1816 |
warning ``Unification bound exceeded''. |
|
1817 |
||
1818 |
||
1819 |
\begin{warn} |
|
1820 |
Options for unification cannot be modified in a local context. Only |
|
1821 |
the global theory content is taken into account. |
|
1822 |
\end{warn} |
|
58618 | 1823 |
\<close> |
50083 | 1824 |
|
26782 | 1825 |
end |