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