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