author | wenzelm |
Sat, 03 Nov 2012 21:31:40 +0100 | |
changeset 50063 | 7e491da6bcbd |
parent 48985 | 5386df44a037 |
child 50064 | e08cc8b20564 |
permissions | -rw-r--r-- |
26782 | 1 |
theory Generic |
42651 | 2 |
imports Base Main |
26782 | 3 |
begin |
4 |
||
5 |
chapter {* Generic tools and packages \label{ch:gen-tools} *} |
|
6 |
||
42655 | 7 |
section {* Configuration options \label{sec:config} *} |
26782 | 8 |
|
40291 | 9 |
text {* Isabelle/Pure maintains a record of named configuration |
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: |
|
19 |
*} |
|
20 |
||
21 |
declare [[show_main_goal = false]] |
|
26782 | 22 |
|
42655 | 23 |
notepad |
24 |
begin |
|
25 |
note [[show_main_goal = true]] |
|
26 |
end |
|
27 |
||
28 |
text {* 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} |
|
28761
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
34 |
@{command_def "print_configs"} & : & @{text "context \<rightarrow>"} \\ |
26782 | 35 |
\end{matharray} |
36 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
37 |
@{rail " |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
38 |
@{syntax name} ('=' ('true' | 'false' | @{syntax int} | @{syntax float} | @{syntax name}))? |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
39 |
"} |
26782 | 40 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
41 |
\begin{description} |
26782 | 42 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
43 |
\item @{command "print_configs"} prints the available configuration |
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} |
26782 | 52 |
*} |
53 |
||
54 |
||
27040 | 55 |
section {* Basic proof tools *} |
26782 | 56 |
|
57 |
subsection {* Miscellaneous methods and attributes \label{sec:misc-meth-att} *} |
|
58 |
||
59 |
text {* |
|
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 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
73 |
@{rail " |
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}? |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
80 |
"} |
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, |
|
98 |
respectively \cite{isabelle-implementation}. The optional natural |
|
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 standard}@{text "\<^sup>*"} & : & @{text attribute} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
136 |
@{attribute_def no_vars}@{text "\<^sup>*"} & : & @{text attribute} \\ |
26782 | 137 |
\end{matharray} |
138 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
139 |
@{rail " |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
140 |
@@{attribute tagged} @{syntax name} @{syntax name} |
26782 | 141 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
142 |
@@{attribute untagged} @{syntax name} |
26782 | 143 |
; |
48205 | 144 |
@@{attribute THEN} ('[' @{syntax nat} ']')? @{syntax thmref} |
26782 | 145 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
146 |
(@@{attribute unfolded} | @@{attribute folded}) @{syntax thmrefs} |
26782 | 147 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
148 |
@@{attribute rotated} @{syntax int}? |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
149 |
"} |
26782 | 150 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
151 |
\begin{description} |
26782 | 152 |
|
28764 | 153 |
\item @{attribute tagged}~@{text "name value"} and @{attribute |
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
154 |
untagged}~@{text name} add and remove \emph{tags} of some theorem. |
26782 | 155 |
Tags may be any list of string pairs that serve as formal comment. |
28764 | 156 |
The first string is considered the tag name, the second its value. |
157 |
Note that @{attribute untagged} removes any tags of the same name. |
|
26782 | 158 |
|
48205 | 159 |
\item @{attribute THEN}~@{text a} composes rules by resolution; it |
160 |
resolves with the first premise of @{text a} (an alternative |
|
161 |
position may be also specified). See also @{ML_op "RS"} in |
|
162 |
\cite{isabelle-implementation}. |
|
26782 | 163 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
164 |
\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
|
165 |
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
|
166 |
definitions throughout a rule. |
26782 | 167 |
|
47497 | 168 |
\item @{attribute abs_def} turns an equation of the form @{prop "f x |
169 |
y \<equiv> t"} into @{prop "f \<equiv> \<lambda>x y. t"}, which ensures that @{method |
|
170 |
simp} or @{method unfold} steps always expand it. This also works |
|
171 |
for object-logic equality. |
|
172 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
173 |
\item @{attribute rotated}~@{text n} rotate the premises of a |
26782 | 174 |
theorem by @{text n} (default 1). |
175 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
176 |
\item @{attribute (Pure) elim_format} turns a destruction rule into |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
177 |
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
|
178 |
(PROP A \<Longrightarrow> PROP B) \<Longrightarrow> PROP B"}. |
26782 | 179 |
|
180 |
Note that the Classical Reasoner (\secref{sec:classical}) provides |
|
181 |
its own version of this operation. |
|
182 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
183 |
\item @{attribute standard} puts a theorem into the standard form of |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
184 |
object-rules at the outermost theory level. Note that this |
26782 | 185 |
operation violates the local proof context (including active |
186 |
locales). |
|
187 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
188 |
\item @{attribute no_vars} replaces schematic variables by free |
26782 | 189 |
ones; this is mainly for tuning output of pretty printed theorems. |
190 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
191 |
\end{description} |
26782 | 192 |
*} |
193 |
||
194 |
||
27044 | 195 |
subsection {* Low-level equational reasoning *} |
196 |
||
197 |
text {* |
|
198 |
\begin{matharray}{rcl} |
|
28761
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
199 |
@{method_def subst} & : & @{text method} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
200 |
@{method_def hypsubst} & : & @{text method} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
201 |
@{method_def split} & : & @{text method} \\ |
27044 | 202 |
\end{matharray} |
203 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
204 |
@{rail " |
42704 | 205 |
@@{method subst} ('(' 'asm' ')')? \\ ('(' (@{syntax nat}+) ')')? @{syntax thmref} |
27044 | 206 |
; |
44094
f7bbfdf4b4a7
updated documentation of method "split" according to e6a4bb832b46;
wenzelm
parents:
43367
diff
changeset
|
207 |
@@{method split} @{syntax thmrefs} |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
208 |
"} |
27044 | 209 |
|
210 |
These methods provide low-level facilities for equational reasoning |
|
211 |
that are intended for specialized applications only. Normally, |
|
212 |
single step calculations would be performed in a structured text |
|
213 |
(see also \secref{sec:calculation}), while the Simplifier methods |
|
214 |
provide the canonical way for automated normalization (see |
|
215 |
\secref{sec:simplifier}). |
|
216 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
217 |
\begin{description} |
27044 | 218 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
219 |
\item @{method subst}~@{text eq} performs a single substitution step |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
220 |
using rule @{text eq}, which may be either a meta or object |
27044 | 221 |
equality. |
222 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
223 |
\item @{method subst}~@{text "(asm) eq"} substitutes in an |
27044 | 224 |
assumption. |
225 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
226 |
\item @{method subst}~@{text "(i \<dots> j) eq"} performs several |
27044 | 227 |
substitutions in the conclusion. The numbers @{text i} to @{text j} |
228 |
indicate the positions to substitute at. Positions are ordered from |
|
229 |
the top of the term tree moving down from left to right. For |
|
230 |
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
|
231 |
where commutativity of @{text "+"} is applicable: 1 refers to @{text |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
232 |
"a + b"}, 2 to the whole term, and 3 to @{text "c + d"}. |
27044 | 233 |
|
234 |
If the positions in the list @{text "(i \<dots> j)"} are non-overlapping |
|
235 |
(e.g.\ @{text "(2 3)"} in @{text "(a + b) + (c + d)"}) you may |
|
236 |
assume all substitutions are performed simultaneously. Otherwise |
|
237 |
the behaviour of @{text subst} is not specified. |
|
238 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
239 |
\item @{method subst}~@{text "(asm) (i \<dots> j) eq"} performs the |
27071 | 240 |
substitutions in the assumptions. The positions refer to the |
241 |
assumptions in order from left to right. For example, given in a |
|
242 |
goal of the form @{text "P (a + b) \<Longrightarrow> P (c + d) \<Longrightarrow> \<dots>"}, position 1 of |
|
243 |
commutativity of @{text "+"} is the subterm @{text "a + b"} and |
|
244 |
position 2 is the subterm @{text "c + d"}. |
|
27044 | 245 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
246 |
\item @{method hypsubst} performs substitution using some |
27044 | 247 |
assumption; this only works for equations of the form @{text "x = |
248 |
t"} where @{text x} is a free or bound variable. |
|
249 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
250 |
\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
|
251 |
splitting using the given rules. Splitting is performed in the |
f7bbfdf4b4a7
updated documentation of method "split" according to e6a4bb832b46;
wenzelm
parents:
43367
diff
changeset
|
252 |
conclusion or some assumption of the subgoal, depending of the |
f7bbfdf4b4a7
updated documentation of method "split" according to e6a4bb832b46;
wenzelm
parents:
43367
diff
changeset
|
253 |
structure of the rule. |
27044 | 254 |
|
255 |
Note that the @{method simp} method already involves repeated |
|
44094
f7bbfdf4b4a7
updated documentation of method "split" according to e6a4bb832b46;
wenzelm
parents:
43367
diff
changeset
|
256 |
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
|
257 |
@{attribute split}, for example. |
27044 | 258 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
259 |
\end{description} |
27044 | 260 |
*} |
261 |
||
262 |
||
26782 | 263 |
subsection {* Further tactic emulations \label{sec:tactics} *} |
264 |
||
265 |
text {* |
|
266 |
The following improper proof methods emulate traditional tactics. |
|
267 |
These admit direct access to the goal state, which is normally |
|
268 |
considered harmful! In particular, this may involve both numbered |
|
269 |
goal addressing (default 1), and dynamic instantiation within the |
|
270 |
scope of some subgoal. |
|
271 |
||
272 |
\begin{warn} |
|
273 |
Dynamic instantiations refer to universally quantified parameters |
|
274 |
of a subgoal (the dynamic context) rather than fixed variables and |
|
275 |
term abbreviations of a (static) Isar context. |
|
276 |
\end{warn} |
|
277 |
||
278 |
Tactic emulation methods, unlike their ML counterparts, admit |
|
279 |
simultaneous instantiation from both dynamic and static contexts. |
|
280 |
If names occur in both contexts goal parameters hide locally fixed |
|
281 |
variables. Likewise, schematic variables refer to term |
|
282 |
abbreviations, if present in the static context. Otherwise the |
|
283 |
schematic variable is interpreted as a schematic variable and left |
|
284 |
to be solved by unification with certain parts of the subgoal. |
|
285 |
||
286 |
Note that the tactic emulation proof methods in Isabelle/Isar are |
|
287 |
consistently named @{text foo_tac}. Note also that variable names |
|
288 |
occurring on left hand sides of instantiations must be preceded by a |
|
289 |
question mark if they coincide with a keyword or contain dots. This |
|
290 |
is consistent with the attribute @{attribute "where"} (see |
|
291 |
\secref{sec:pure-meth-att}). |
|
292 |
||
293 |
\begin{matharray}{rcl} |
|
28761
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
294 |
@{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
|
295 |
@{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
|
296 |
@{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
|
297 |
@{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
|
298 |
@{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
|
299 |
@{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
|
300 |
@{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
|
301 |
@{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
|
302 |
@{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
|
303 |
@{method_def tactic}@{text "\<^sup>*"} & : & @{text method} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
304 |
@{method_def raw_tactic}@{text "\<^sup>*"} & : & @{text method} \\ |
26782 | 305 |
\end{matharray} |
306 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
307 |
@{rail " |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
308 |
(@@{method rule_tac} | @@{method erule_tac} | @@{method drule_tac} | |
42705 | 309 |
@@{method frule_tac} | @@{method cut_tac} | @@{method thin_tac}) @{syntax goal_spec}? \\ |
42617 | 310 |
( dynamic_insts @'in' @{syntax thmref} | @{syntax thmrefs} ) |
26782 | 311 |
; |
42705 | 312 |
@@{method subgoal_tac} @{syntax goal_spec}? (@{syntax prop} +) |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
313 |
; |
42705 | 314 |
@@{method rename_tac} @{syntax goal_spec}? (@{syntax name} +) |
26782 | 315 |
; |
42705 | 316 |
@@{method rotate_tac} @{syntax goal_spec}? @{syntax int}? |
26782 | 317 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
318 |
(@@{method tactic} | @@{method raw_tactic}) @{syntax text} |
26782 | 319 |
; |
320 |
||
42617 | 321 |
dynamic_insts: ((@{syntax name} '=' @{syntax term}) + @'and') |
322 |
"} |
|
26782 | 323 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
324 |
\begin{description} |
26782 | 325 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
326 |
\item @{method rule_tac} etc. do resolution of rules with explicit |
26782 | 327 |
instantiation. This works the same way as the ML tactics @{ML |
30397 | 328 |
res_inst_tac} etc. (see \cite{isabelle-implementation}) |
26782 | 329 |
|
330 |
Multiple rules may be only given if there is no instantiation; then |
|
331 |
@{method rule_tac} is the same as @{ML resolve_tac} in ML (see |
|
30397 | 332 |
\cite{isabelle-implementation}). |
26782 | 333 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
334 |
\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
|
335 |
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
|
336 |
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
|
337 |
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
|
338 |
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
|
339 |
closed rule statements. |
26782 | 340 |
|
46277 | 341 |
\item @{method thin_tac}~@{text \<phi>} deletes the specified premise |
342 |
from a subgoal. Note that @{text \<phi>} may contain schematic |
|
343 |
variables, to abbreviate the intended proposition; the first |
|
344 |
matching subgoal premise will be deleted. Removing useless premises |
|
345 |
from a subgoal increases its readability and can make search tactics |
|
346 |
run faster. |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
347 |
|
46271 | 348 |
\item @{method subgoal_tac}~@{text "\<phi>\<^sub>1 \<dots> \<phi>\<^sub>n"} adds the propositions |
349 |
@{text "\<phi>\<^sub>1 \<dots> \<phi>\<^sub>n"} as local premises to a subgoal, and poses the same |
|
350 |
as new subgoals (in the original context). |
|
26782 | 351 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
352 |
\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
|
353 |
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
|
354 |
\emph{suffix} of variables. |
26782 | 355 |
|
46274 | 356 |
\item @{method rotate_tac}~@{text n} rotates the premises of a |
357 |
subgoal by @{text n} positions: from right to left if @{text n} is |
|
26782 | 358 |
positive, and from left to right if @{text n} is negative; the |
46274 | 359 |
default value is 1. |
26782 | 360 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
361 |
\item @{method tactic}~@{text "text"} produces a proof method from |
26782 | 362 |
any ML text of type @{ML_type tactic}. Apart from the usual ML |
27223 | 363 |
environment and the current proof context, the ML code may refer to |
364 |
the locally bound values @{ML_text facts}, which indicates any |
|
365 |
current facts used for forward-chaining. |
|
26782 | 366 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
367 |
\item @{method raw_tactic} is similar to @{method tactic}, but |
27223 | 368 |
presents the goal state in its raw internal form, where simultaneous |
369 |
subgoals appear as conjunction of the logical framework instead of |
|
370 |
the usual split into several subgoals. While feature this is useful |
|
371 |
for debugging of complex method definitions, it should not never |
|
372 |
appear in production theories. |
|
26782 | 373 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
374 |
\end{description} |
26782 | 375 |
*} |
376 |
||
377 |
||
27040 | 378 |
section {* The Simplifier \label{sec:simplifier} *} |
26782 | 379 |
|
50063 | 380 |
text {* The Simplifier performs conditional and unconditional |
381 |
rewriting and uses contextual information: rule declarations in the |
|
382 |
background theory or local proof context are taken into account, as |
|
383 |
well as chained facts and subgoal premises (``local assumptions''). |
|
384 |
There are several general hooks that allow to modify the |
|
385 |
simplification strategy, or incorporate other proof tools that solve |
|
386 |
sub-problems, produce rewrite rules on demand etc. |
|
387 |
||
388 |
The default Simplifier setup of major object logics (HOL, HOLCF, |
|
389 |
FOL, ZF) makes the Simplifier ready for immediate use, without |
|
390 |
engaging into the internal structures. Thus it serves as |
|
391 |
general-purpose proof tool with the main focus on equational |
|
392 |
reasoning, and a bit more than that. *} |
|
393 |
||
394 |
||
395 |
subsection {* Simplification methods \label{sec:simp-meth} *} |
|
26782 | 396 |
|
397 |
text {* |
|
398 |
\begin{matharray}{rcl} |
|
28761
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
399 |
@{method_def simp} & : & @{text method} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
400 |
@{method_def simp_all} & : & @{text method} \\ |
26782 | 401 |
\end{matharray} |
402 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
403 |
@{rail " |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
404 |
(@@{method simp} | @@{method simp_all}) opt? (@{syntax simpmod} * ) |
26782 | 405 |
; |
406 |
||
40255
9ffbc25e1606
eliminated obsolete \_ escapes in rail environments;
wenzelm
parents:
35613
diff
changeset
|
407 |
opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use' | 'asm_lr' ) ')' |
26782 | 408 |
; |
50063 | 409 |
@{syntax_def simpmod}: ('add' | 'del' | 'only' | 'split' (() | 'add' | 'del') | |
410 |
'cong' (() | 'add' | 'del')) ':' @{syntax thmrefs} |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
411 |
"} |
26782 | 412 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
413 |
\begin{description} |
26782 | 414 |
|
50063 | 415 |
\item @{method simp} invokes the Simplifier on the first subgoal, |
416 |
after inserting chained facts as additional goal premises; further |
|
417 |
rule declarations may be included via @{text "(simp add: facts)"}. |
|
418 |
The proof method fails if the subgoal remains unchanged after |
|
419 |
simplification. |
|
26782 | 420 |
|
50063 | 421 |
Note that the original goal premises and chained facts are subject |
422 |
to simplification themselves, while declarations via @{text |
|
423 |
"add"}/@{text "del"} merely follow the policies of the object-logic |
|
424 |
to extract rewrite rules from theorems, without further |
|
425 |
simplification. This may lead to slightly different behavior in |
|
426 |
either case, which might be required precisely like that in some |
|
427 |
boundary situations to perform the intended simplification step! |
|
428 |
||
429 |
\medskip The @{text only} modifier first removes all other rewrite |
|
430 |
rules, looper tactics (including split rules), congruence rules, and |
|
431 |
then behaves like @{text add}. Implicit solvers remain, which means |
|
432 |
that trivial rules like reflexivity or introduction of @{text |
|
433 |
"True"} are available to solve the simplified subgoals, but also |
|
434 |
non-trivial tools like linear arithmetic in HOL. The latter may |
|
435 |
lead to some surprise of the meaning of ``only'' in Isabelle/HOL |
|
436 |
compared to English! |
|
26782 | 437 |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
438 |
\medskip The @{text split} modifiers add or delete rules for the |
26782 | 439 |
Splitter (see also \cite{isabelle-ref}), the default is to add. |
440 |
This works only if the Simplifier method has been properly setup to |
|
441 |
include the Splitter (all major object logics such HOL, HOLCF, FOL, |
|
442 |
ZF do this already). |
|
443 |
||
50063 | 444 |
\medskip The @{text cong} modifiers add or delete Simplifier |
445 |
congruence rules (see also \secref{sec:simp-rules}); the default is |
|
446 |
to add. |
|
447 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
448 |
\item @{method simp_all} is similar to @{method simp}, but acts on |
50063 | 449 |
all goals, working backwards from the last to the first one as usual |
450 |
in Isabelle.\footnote{The order is irrelevant for goals without |
|
451 |
schematic variables, so simplification might actually be performed |
|
452 |
in parallel here.} |
|
453 |
||
454 |
Chained facts are inserted into all subgoals, before the |
|
455 |
simplification process starts. Further rule declarations are the |
|
456 |
same as for @{method simp}. |
|
457 |
||
458 |
The proof method fails if all subgoals remain unchanged after |
|
459 |
simplification. |
|
26782 | 460 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
461 |
\end{description} |
26782 | 462 |
|
50063 | 463 |
By default the Simplifier methods above take local assumptions fully |
464 |
into account, using equational assumptions in the subsequent |
|
465 |
normalization process, or simplifying assumptions themselves. |
|
466 |
Further options allow to fine-tune the behavior of the Simplifier |
|
467 |
in this respect, corresponding to a variety of ML tactics as |
|
468 |
follows.\footnote{Unlike the corresponding Isar proof methods, the |
|
469 |
ML tactics do not insist in changing the goal state.} |
|
470 |
||
471 |
\begin{center} |
|
472 |
\small |
|
473 |
\begin{tabular}{|l|l|p{0.3\textwidth}|} |
|
474 |
\hline |
|
475 |
Isar method & ML tactic & behavior \\\hline |
|
476 |
||
477 |
@{text "(simp (no_asm))"} & @{ML simp_tac} & assumptions are ignored |
|
478 |
completely \\\hline |
|
26782 | 479 |
|
50063 | 480 |
@{text "(simp (no_asm_simp))"} & @{ML asm_simp_tac} & assumptions |
481 |
are used in the simplification of the conclusion but are not |
|
482 |
themselves simplified \\\hline |
|
483 |
||
484 |
@{text "(simp (no_asm_use))"} & @{ML full_simp_tac} & assumptions |
|
485 |
are simplified but are not used in the simplification of each other |
|
486 |
or the conclusion \\\hline |
|
26782 | 487 |
|
50063 | 488 |
@{text "(simp)"} & @{ML asm_full_simp_tac} & assumptions are used in |
489 |
the simplification of the conclusion and to simplify other |
|
490 |
assumptions \\\hline |
|
491 |
||
492 |
@{text "(simp (asm_lr))"} & @{ML asm_lr_simp_tac} & compatibility |
|
493 |
mode: an assumption is only used for simplifying assumptions which |
|
494 |
are to the right of it \\\hline |
|
495 |
||
496 |
\end{tabular} |
|
497 |
\end{center} |
|
26782 | 498 |
|
499 |
\medskip The Splitter package is usually configured to work as part |
|
500 |
of the Simplifier. The effect of repeatedly applying @{ML |
|
501 |
split_tac} can be simulated by ``@{text "(simp only: split: |
|
502 |
a\<^sub>1 \<dots> a\<^sub>n)"}''. There is also a separate @{text split} |
|
503 |
method available for single-step case splitting. |
|
504 |
*} |
|
505 |
||
506 |
||
50063 | 507 |
subsection {* Declaring rules \label{sec:simp-rules} *} |
26782 | 508 |
|
509 |
text {* |
|
510 |
\begin{matharray}{rcl} |
|
28761
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
511 |
@{command_def "print_simpset"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
512 |
@{attribute_def simp} & : & @{text attribute} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
513 |
@{attribute_def split} & : & @{text attribute} \\ |
50063 | 514 |
@{attribute_def cong} & : & @{text attribute} \\ |
26782 | 515 |
\end{matharray} |
516 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
517 |
@{rail " |
50063 | 518 |
(@@{attribute simp} | @@{attribute split} | @@{attribute cong}) |
519 |
(() | 'add' | 'del') |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
520 |
"} |
26782 | 521 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
522 |
\begin{description} |
26782 | 523 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
524 |
\item @{command "print_simpset"} prints the collection of rules |
26782 | 525 |
declared to the Simplifier, which is also known as ``simpset'' |
50063 | 526 |
internally. |
527 |
||
528 |
For historical reasons, simpsets may occur independently from the |
|
529 |
current context, but are conceptually dependent on it. When the |
|
530 |
Simplifier is invoked via one of its main entry points in the Isar |
|
531 |
source language (as proof method \secref{sec:simp-meth} or rule |
|
532 |
attribute \secref{sec:simp-meth}), its simpset is derived from the |
|
533 |
current proof context, and carries a back-reference to that for |
|
534 |
other tools that might get invoked internally (e.g.\ simplification |
|
535 |
procedures \secref{sec:simproc}). A mismatch of the context of the |
|
536 |
simpset and the context of the problem being simplified may lead to |
|
537 |
unexpected results. |
|
26782 | 538 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
539 |
\item @{attribute simp} declares simplification rules. |
26782 | 540 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
541 |
\item @{attribute split} declares case split rules. |
26782 | 542 |
|
45645 | 543 |
\item @{attribute cong} declares congruence rules to the Simplifier |
544 |
context. |
|
545 |
||
546 |
Congruence rules are equalities of the form @{text [display] |
|
547 |
"\<dots> \<Longrightarrow> f ?x\<^sub>1 \<dots> ?x\<^sub>n = f ?y\<^sub>1 \<dots> ?y\<^sub>n"} |
|
548 |
||
549 |
This controls the simplification of the arguments of @{text f}. For |
|
550 |
example, some arguments can be simplified under additional |
|
551 |
assumptions: @{text [display] "?P\<^sub>1 \<longleftrightarrow> ?Q\<^sub>1 \<Longrightarrow> (?Q\<^sub>1 \<Longrightarrow> ?P\<^sub>2 \<longleftrightarrow> ?Q\<^sub>2) \<Longrightarrow> |
|
552 |
(?P\<^sub>1 \<longrightarrow> ?P\<^sub>2) \<longleftrightarrow> (?Q\<^sub>1 \<longrightarrow> ?Q\<^sub>2)"} |
|
553 |
||
554 |
Given this rule, the simplifier assumes @{text "?Q\<^sub>1"} and extracts |
|
555 |
rewrite rules from it when simplifying @{text "?P\<^sub>2"}. Such local |
|
556 |
assumptions are effective for rewriting formulae such as @{text "x = |
|
557 |
0 \<longrightarrow> y + x = y"}. |
|
558 |
||
559 |
%FIXME |
|
560 |
%The local assumptions are also provided as theorems to the solver; |
|
561 |
%see \secref{sec:simp-solver} below. |
|
562 |
||
563 |
\medskip The following congruence rule for bounded quantifiers also |
|
564 |
supplies contextual information --- about the bound variable: |
|
565 |
@{text [display] "(?A = ?B) \<Longrightarrow> (\<And>x. x \<in> ?B \<Longrightarrow> ?P x \<longleftrightarrow> ?Q x) \<Longrightarrow> |
|
566 |
(\<forall>x \<in> ?A. ?P x) \<longleftrightarrow> (\<forall>x \<in> ?B. ?Q x)"} |
|
567 |
||
568 |
\medskip This congruence rule for conditional expressions can |
|
569 |
supply contextual information for simplifying the arms: |
|
570 |
@{text [display] "?p = ?q \<Longrightarrow> (?q \<Longrightarrow> ?a = ?c) \<Longrightarrow> (\<not> ?q \<Longrightarrow> ?b = ?d) \<Longrightarrow> |
|
571 |
(if ?p then ?a else ?b) = (if ?q then ?c else ?d)"} |
|
572 |
||
573 |
A congruence rule can also \emph{prevent} simplification of some |
|
574 |
arguments. Here is an alternative congruence rule for conditional |
|
575 |
expressions that conforms to non-strict functional evaluation: |
|
576 |
@{text [display] "?p = ?q \<Longrightarrow> (if ?p then ?a else ?b) = (if ?q then ?a else ?b)"} |
|
577 |
||
578 |
Only the first argument is simplified; the others remain unchanged. |
|
579 |
This can make simplification much faster, but may require an extra |
|
580 |
case split over the condition @{text "?q"} to prove the goal. |
|
50063 | 581 |
|
582 |
\end{description} |
|
45645 | 583 |
*} |
584 |
||
585 |
||
50063 | 586 |
subsection {* Configuration options \label{sec:simp-config} *} |
587 |
||
588 |
text {* |
|
589 |
\begin{tabular}{rcll} |
|
590 |
@{attribute_def simp_depth_limit} & : & @{text attribute} & default @{text 100} \\ |
|
591 |
@{attribute_def simp_trace} & : & @{text attribute} & default @{text false} \\ |
|
592 |
@{attribute_def simp_trace_depth_limit} & : & @{text attribute} & default @{text 1} \\ |
|
593 |
@{attribute_def simp_debug} & : & @{text attribute} & default @{text false} \\ |
|
594 |
\end{tabular} |
|
595 |
\medskip |
|
596 |
||
597 |
These configurations options control further aspects of the Simplifier. |
|
598 |
See also \secref{sec:config}. |
|
599 |
||
600 |
\begin{description} |
|
601 |
||
602 |
\item @{attribute simp_depth_limit} limits the number of recursive |
|
603 |
invocations of the Simplifier during conditional rewriting. |
|
604 |
||
605 |
\item @{attribute simp_trace} makes the Simplifier output internal |
|
606 |
operations. This includes rewrite steps, but also bookkeeping like |
|
607 |
modifications of the simpset. |
|
608 |
||
609 |
\item @{attribute simp_trace_depth_limit} limits the effect of |
|
610 |
@{attribute simp_trace} to the given depth of recursive Simplifier |
|
611 |
invocations (when solving conditions of rewrite rules). |
|
612 |
||
613 |
\item @{attribute simp_debug} makes the Simplifier output some extra |
|
614 |
information about internal operations. This includes any attempted |
|
615 |
invocation of simplification procedures. |
|
616 |
||
617 |
\end{description} |
|
618 |
*} |
|
619 |
||
620 |
||
621 |
subsection {* Simplification procedures \label{sec:simproc} *} |
|
26782 | 622 |
|
42925 | 623 |
text {* Simplification procedures are ML functions that produce proven |
624 |
rewrite rules on demand. They are associated with higher-order |
|
625 |
patterns that approximate the left-hand sides of equations. The |
|
626 |
Simplifier first matches the current redex against one of the LHS |
|
627 |
patterns; if this succeeds, the corresponding ML function is |
|
628 |
invoked, passing the Simplifier context and redex term. Thus rules |
|
629 |
may be specifically fashioned for particular situations, resulting |
|
630 |
in a more powerful mechanism than term rewriting by a fixed set of |
|
631 |
rules. |
|
632 |
||
633 |
Any successful result needs to be a (possibly conditional) rewrite |
|
634 |
rule @{text "t \<equiv> u"} that is applicable to the current redex. The |
|
635 |
rule will be applied just as any ordinary rewrite rule. It is |
|
636 |
expected to be already in \emph{internal form}, bypassing the |
|
637 |
automatic preprocessing of object-level equivalences. |
|
638 |
||
26782 | 639 |
\begin{matharray}{rcl} |
28761
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
640 |
@{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
|
641 |
simproc & : & @{text attribute} \\ |
26782 | 642 |
\end{matharray} |
643 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
644 |
@{rail " |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
645 |
@@{command simproc_setup} @{syntax name} '(' (@{syntax term} + '|') ')' '=' |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
646 |
@{syntax text} \\ (@'identifier' (@{syntax nameref}+))? |
26782 | 647 |
; |
648 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
649 |
@@{attribute simproc} (('add' ':')? | 'del' ':') (@{syntax name}+) |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
650 |
"} |
26782 | 651 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
652 |
\begin{description} |
26782 | 653 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
654 |
\item @{command "simproc_setup"} defines a named simplification |
26782 | 655 |
procedure that is invoked by the Simplifier whenever any of the |
656 |
given term patterns match the current redex. The implementation, |
|
657 |
which is provided as ML source text, needs to be of type @{ML_type |
|
658 |
"morphism -> simpset -> cterm -> thm option"}, where the @{ML_type |
|
659 |
cterm} represents the current redex @{text r} and the result is |
|
660 |
supposed to be some proven rewrite rule @{text "r \<equiv> r'"} (or a |
|
661 |
generalized version), or @{ML NONE} to indicate failure. The |
|
662 |
@{ML_type simpset} argument holds the full context of the current |
|
663 |
Simplifier invocation, including the actual Isar proof context. The |
|
664 |
@{ML_type morphism} informs about the difference of the original |
|
665 |
compilation context wrt.\ the one of the actual application later |
|
666 |
on. The optional @{keyword "identifier"} specifies theorems that |
|
667 |
represent the logical content of the abstract theory of this |
|
668 |
simproc. |
|
669 |
||
670 |
Morphisms and identifiers are only relevant for simprocs that are |
|
671 |
defined within a local target context, e.g.\ in a locale. |
|
672 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
673 |
\item @{text "simproc add: name"} and @{text "simproc del: name"} |
26782 | 674 |
add or delete named simprocs to the current Simplifier context. The |
675 |
default is to add a simproc. Note that @{command "simproc_setup"} |
|
676 |
already adds the new simproc to the subsequent context. |
|
677 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
678 |
\end{description} |
26782 | 679 |
*} |
680 |
||
681 |
||
42925 | 682 |
subsubsection {* Example *} |
683 |
||
684 |
text {* The following simplification procedure for @{thm |
|
685 |
[source=false, show_types] unit_eq} in HOL performs fine-grained |
|
686 |
control over rule application, beyond higher-order pattern matching. |
|
687 |
Declaring @{thm unit_eq} as @{attribute simp} directly would make |
|
688 |
the simplifier loop! Note that a version of this simplification |
|
689 |
procedure is already active in Isabelle/HOL. *} |
|
690 |
||
691 |
simproc_setup unit ("x::unit") = {* |
|
692 |
fn _ => fn _ => fn ct => |
|
693 |
if HOLogic.is_unit (term_of ct) then NONE |
|
694 |
else SOME (mk_meta_eq @{thm unit_eq}) |
|
695 |
*} |
|
696 |
||
697 |
text {* Since the Simplifier applies simplification procedures |
|
698 |
frequently, it is important to make the failure check in ML |
|
699 |
reasonably fast. *} |
|
700 |
||
701 |
||
50063 | 702 |
subsection {* Forward simplification \label{sec:simp-forward} *} |
26782 | 703 |
|
704 |
text {* |
|
705 |
\begin{matharray}{rcl} |
|
28761
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
706 |
@{attribute_def simplified} & : & @{text attribute} \\ |
26782 | 707 |
\end{matharray} |
708 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
709 |
@{rail " |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
710 |
@@{attribute simplified} opt? @{syntax thmrefs}? |
26782 | 711 |
; |
712 |
||
40255
9ffbc25e1606
eliminated obsolete \_ escapes in rail environments;
wenzelm
parents:
35613
diff
changeset
|
713 |
opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use') ')' |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
714 |
"} |
26782 | 715 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
716 |
\begin{description} |
26782 | 717 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
718 |
\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
|
719 |
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
|
720 |
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
|
721 |
The result is fully simplified by default, including assumptions and |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
722 |
conclusion; the options @{text no_asm} etc.\ tune the Simplifier in |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
723 |
the same way as the for the @{text simp} method. |
26782 | 724 |
|
725 |
Note that forward simplification restricts the simplifier to its |
|
726 |
most basic operation of term rewriting; solver and looper tactics |
|
727 |
\cite{isabelle-ref} are \emph{not} involved here. The @{text |
|
728 |
simplified} attribute should be only rarely required under normal |
|
729 |
circumstances. |
|
730 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
731 |
\end{description} |
26782 | 732 |
*} |
733 |
||
734 |
||
27040 | 735 |
section {* The Classical Reasoner \label{sec:classical} *} |
26782 | 736 |
|
42930 | 737 |
subsection {* Basic concepts *} |
42927
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
738 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
739 |
text {* Although Isabelle is generic, many users will be working in |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
740 |
some extension of classical first-order logic. Isabelle/ZF is built |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
741 |
upon theory FOL, while Isabelle/HOL conceptually contains |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
742 |
first-order logic as a fragment. Theorem-proving in predicate logic |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
743 |
is undecidable, but many automated strategies have been developed to |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
744 |
assist in this task. |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
745 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
746 |
Isabelle's classical reasoner is a generic package that accepts |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
747 |
certain information about a logic and delivers a suite of automatic |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
748 |
proof tools, based on rules that are classified and declared in the |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
749 |
context. These proof procedures are slow and simplistic compared |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
750 |
with high-end automated theorem provers, but they can save |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
751 |
considerable time and effort in practice. They can prove theorems |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
752 |
such as Pelletier's \cite{pelletier86} problems 40 and 41 in a few |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
753 |
milliseconds (including full proof reconstruction): *} |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
754 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
755 |
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
|
756 |
by blast |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
757 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
758 |
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
|
759 |
by blast |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
760 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
761 |
text {* The proof tools are generic. They are not restricted to |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
762 |
first-order logic, and have been heavily used in the development of |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
763 |
the Isabelle/HOL library and applications. The tactics can be |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
764 |
traced, and their components can be called directly; in this manner, |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
765 |
any proof can be viewed interactively. *} |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
766 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
767 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
768 |
subsubsection {* The sequent calculus *} |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
769 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
770 |
text {* Isabelle supports natural deduction, which is easy to use for |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
771 |
interactive proof. But natural deduction does not easily lend |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
772 |
itself to automation, and has a bias towards intuitionism. For |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
773 |
certain proofs in classical logic, it can not be called natural. |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
774 |
The \emph{sequent calculus}, a generalization of natural deduction, |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
775 |
is easier to automate. |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
776 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
777 |
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
|
778 |
and @{text "\<Delta>"} are sets of formulae.\footnote{For first-order |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
779 |
logic, sequents can equivalently be made from lists or multisets of |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
780 |
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
|
781 |
\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
|
782 |
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
|
783 |
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
|
784 |
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
|
785 |
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
|
786 |
valid. |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
787 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
788 |
Sequent rules are classified as \textbf{right} or \textbf{left}, |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
789 |
indicating which side of the @{text "\<turnstile>"} symbol they operate on. |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
790 |
Rules that operate on the right side are analogous to natural |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
791 |
deduction's introduction rules, and left rules are analogous to |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
792 |
elimination rules. The sequent calculus analogue of @{text "(\<longrightarrow>I)"} |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
793 |
is the rule |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
794 |
\[ |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
795 |
\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
|
796 |
\] |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
797 |
Applying the rule backwards, this breaks down some implication on |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
798 |
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
|
799 |
the sets of formulae that are unaffected by the inference. The |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
800 |
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
|
801 |
single rule |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
802 |
\[ |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
803 |
\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
|
804 |
\] |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
805 |
This breaks down some disjunction on the right side, replacing it by |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
806 |
both disjuncts. Thus, the sequent calculus is a kind of |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
807 |
multiple-conclusion logic. |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
808 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
809 |
To illustrate the use of multiple formulae on the right, let us |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
810 |
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
|
811 |
backwards, we reduce this formula to a basic sequent: |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
812 |
\[ |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
813 |
\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
|
814 |
{\infer[@{text "(\<longrightarrow>R)"}]{@{text "\<turnstile> (P \<longrightarrow> Q), (Q \<longrightarrow> P)"}} |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
815 |
{\infer[@{text "(\<longrightarrow>R)"}]{@{text "P \<turnstile> Q, (Q \<longrightarrow> P)"}} |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
816 |
{@{text "P, Q \<turnstile> Q, P"}}}} |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
817 |
\] |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
818 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
819 |
This example is typical of the sequent calculus: start with the |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
820 |
desired theorem and apply rules backwards in a fairly arbitrary |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
821 |
manner. This yields a surprisingly effective proof procedure. |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
822 |
Quantifiers add only few complications, since Isabelle handles |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
823 |
parameters and schematic variables. See \cite[Chapter |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
824 |
10]{paulson-ml2} for further discussion. *} |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
825 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
826 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
827 |
subsubsection {* Simulating sequents by natural deduction *} |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
828 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
829 |
text {* Isabelle can represent sequents directly, as in the |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
830 |
object-logic LK. But natural deduction is easier to work with, and |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
831 |
most object-logics employ it. Fortunately, we can simulate the |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
832 |
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
|
833 |
@{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
|
834 |
the assumptions and the choice of @{text "Q\<^sub>1"} are arbitrary. |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
835 |
Elim-resolution plays a key role in simulating sequent proofs. |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
836 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
837 |
We can easily handle reasoning on the left. Elim-resolution with |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
838 |
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
|
839 |
a similar effect as the corresponding sequent rules. For the other |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
840 |
connectives, we use sequent-style elimination rules instead of |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
841 |
destruction rules such as @{text "(\<and>E1, 2)"} and @{text "(\<forall>E)"}. |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
842 |
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
|
843 |
representation of sequents! |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
844 |
\[ |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
845 |
\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
|
846 |
\] |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
847 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
848 |
What about reasoning on the right? Introduction rules can only |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
849 |
affect the formula in the conclusion, namely @{text "Q\<^sub>1"}. The |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
850 |
other right-side formulae are represented as negated assumptions, |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
851 |
@{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
|
852 |
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
|
853 |
@{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
|
854 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
855 |
To ensure that swaps occur only when necessary, each introduction |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
856 |
rule is converted into a swapped form: it is resolved with the |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
857 |
second premise of @{text "(swap)"}. The swapped form of @{text |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
858 |
"(\<and>I)"}, which might be called @{text "(\<not>\<and>E)"}, is |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
859 |
@{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
|
860 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
861 |
Similarly, the swapped form of @{text "(\<longrightarrow>I)"} is |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
862 |
@{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
|
863 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
864 |
Swapped introduction rules are applied using elim-resolution, which |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
865 |
deletes the negated formula. Our representation of sequents also |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
866 |
requires the use of ordinary introduction rules. If we had no |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
867 |
regard for readability of intermediate goal states, we could treat |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
868 |
the right side more uniformly by representing sequents as @{text |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
869 |
[display] "P\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> P\<^sub>m \<Longrightarrow> \<not> Q\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> \<not> Q\<^sub>n \<Longrightarrow> \<bottom>"} |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
870 |
*} |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
871 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
872 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
873 |
subsubsection {* Extra rules for the sequent calculus *} |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
874 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
875 |
text {* As mentioned, destruction rules such as @{text "(\<and>E1, 2)"} and |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
876 |
@{text "(\<forall>E)"} must be replaced by sequent-style elimination rules. |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
877 |
In addition, we need rules to embody the classical equivalence |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
878 |
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
|
879 |
rules @{text "(\<or>I1, 2)"} are replaced by a rule that simulates |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
880 |
@{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
|
881 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
882 |
The destruction rule @{text "(\<longrightarrow>E)"} is replaced by @{text [display] |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
883 |
"(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
|
884 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
885 |
Quantifier replication also requires special rules. In classical |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
886 |
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
|
887 |
the rules @{text "(\<exists>R)"} and @{text "(\<forall>L)"} are dual: |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
888 |
\[ |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
889 |
\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
|
890 |
\qquad |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
891 |
\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
|
892 |
\] |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
893 |
Thus both kinds of quantifier may be replicated. Theorems requiring |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
894 |
multiple uses of a universal formula are easy to invent; consider |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
895 |
@{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
|
896 |
@{text "n > 1"}. Natural examples of the multiple use of an |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
897 |
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
|
898 |
\<longrightarrow> P y"}. |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
899 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
900 |
Forgoing quantifier replication loses completeness, but gains |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
901 |
decidability, since the search space becomes finite. Many useful |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
902 |
theorems can be proved without replication, and the search generally |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
903 |
delivers its verdict in a reasonable time. To adopt this approach, |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
904 |
represent the sequent rules @{text "(\<exists>R)"}, @{text "(\<exists>L)"} and |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
905 |
@{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
|
906 |
respectively, and put @{text "(\<forall>E)"} into elimination form: @{text |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
907 |
[display] "\<forall>x. P x \<Longrightarrow> (P t \<Longrightarrow> Q) \<Longrightarrow> Q"} |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
908 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
909 |
Elim-resolution with this rule will delete the universal formula |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
910 |
after a single use. To replicate universal quantifiers, replace the |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
911 |
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
|
912 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
913 |
To replicate existential quantifiers, replace @{text "(\<exists>I)"} by |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
914 |
@{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
|
915 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
916 |
All introduction rules mentioned above are also useful in swapped |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
917 |
form. |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
918 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
919 |
Replication makes the search space infinite; we must apply the rules |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
920 |
with care. The classical reasoner distinguishes between safe and |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
921 |
unsafe rules, applying the latter only when there is no alternative. |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
922 |
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
|
923 |
is better behaved in an infinite search space. However, quantifier |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
924 |
replication is too expensive to prove any but the simplest theorems. |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
925 |
*} |
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
926 |
|
c40adab7568e
moved/updated introduction to Classical Reasoner;
wenzelm
parents:
42925
diff
changeset
|
927 |
|
42928
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
928 |
subsection {* Rule declarations *} |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
929 |
|
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
930 |
text {* The proof tools of the Classical Reasoner depend on |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
931 |
collections of rules declared in the context, which are classified |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
932 |
as introduction, elimination or destruction and as \emph{safe} or |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
933 |
\emph{unsafe}. In general, safe rules can be attempted blindly, |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
934 |
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
|
935 |
reduce a provable goal to an unprovable set of subgoals. |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
936 |
|
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
937 |
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
|
938 |
\<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
|
939 |
unprovable subgoal. Any rule is unsafe whose premises contain new |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
940 |
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
|
941 |
unsafe, since it is applied via elim-resolution, which discards the |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
942 |
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
|
943 |
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
|
944 |
unsafe for similar reasons. The quantifier duplication rule @{text |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
945 |
"\<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
|
946 |
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
|
947 |
looping. In classical first-order logic, all rules are safe except |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
948 |
those mentioned above. |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
949 |
|
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
950 |
The safe~/ unsafe distinction is vague, and may be regarded merely |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
951 |
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
|
952 |
that @{text "(\<or>E)"} is unsafe, because repeated application of it |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
953 |
could generate exponentially many subgoals. Induction rules are |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
954 |
unsafe because inductive proofs are difficult to set up |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
955 |
automatically. Any inference is unsafe that instantiates an unknown |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
956 |
in the proof state --- thus matching must be used, rather than |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
957 |
unification. Even proof by assumption is unsafe if it instantiates |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
958 |
unknowns shared with other subgoals. |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
959 |
|
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
960 |
\begin{matharray}{rcl} |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
961 |
@{command_def "print_claset"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\ |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
962 |
@{attribute_def intro} & : & @{text attribute} \\ |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
963 |
@{attribute_def elim} & : & @{text attribute} \\ |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
964 |
@{attribute_def dest} & : & @{text attribute} \\ |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
965 |
@{attribute_def rule} & : & @{text attribute} \\ |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
966 |
@{attribute_def iff} & : & @{text attribute} \\ |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
967 |
@{attribute_def swapped} & : & @{text attribute} \\ |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
968 |
\end{matharray} |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
969 |
|
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
970 |
@{rail " |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
971 |
(@@{attribute intro} | @@{attribute elim} | @@{attribute dest}) ('!' | () | '?') @{syntax nat}? |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
972 |
; |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
973 |
@@{attribute rule} 'del' |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
974 |
; |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
975 |
@@{attribute iff} (((() | 'add') '?'?) | 'del') |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
976 |
"} |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
977 |
|
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
978 |
\begin{description} |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
979 |
|
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
980 |
\item @{command "print_claset"} prints the collection of rules |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
981 |
declared to the Classical Reasoner, i.e.\ the @{ML_type claset} |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
982 |
within the context. |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
983 |
|
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
984 |
\item @{attribute intro}, @{attribute elim}, and @{attribute dest} |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
985 |
declare introduction, elimination, and destruction rules, |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
986 |
respectively. By default, rules are considered as \emph{unsafe} |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
987 |
(i.e.\ not applied blindly without backtracking), while ``@{text |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
988 |
"!"}'' classifies as \emph{safe}. Rule declarations marked by |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
989 |
``@{text "?"}'' coincide with those of Isabelle/Pure, cf.\ |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
990 |
\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
|
991 |
of the @{method rule} method). The optional natural number |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
992 |
specifies an explicit weight argument, which is ignored by the |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
993 |
automated reasoning tools, but determines the search order of single |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
994 |
rule steps. |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
995 |
|
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
996 |
Introduction rules are those that can be applied using ordinary |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
997 |
resolution. Their swapped forms are generated internally, which |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
998 |
will be applied using elim-resolution. Elimination rules are |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
999 |
applied using elim-resolution. Rules are sorted by the number of |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1000 |
new subgoals they will yield; rules that generate the fewest |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1001 |
subgoals will be tried first. Otherwise, later declarations take |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1002 |
precedence over earlier ones. |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1003 |
|
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1004 |
Rules already present in the context with the same classification |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1005 |
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
|
1006 |
added with some other classification, but the rule is added anyway |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1007 |
as requested. |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1008 |
|
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1009 |
\item @{attribute rule}~@{text del} deletes all occurrences of a |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1010 |
rule from the classical context, regardless of its classification as |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1011 |
introduction~/ elimination~/ destruction and safe~/ unsafe. |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1012 |
|
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1013 |
\item @{attribute iff} declares logical equivalences to the |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1014 |
Simplifier and the Classical reasoner at the same time. |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1015 |
Non-conditional rules result in a safe introduction and elimination |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1016 |
pair; conditional ones are considered unsafe. Rules with negative |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1017 |
conclusion are automatically inverted (using @{text "\<not>"}-elimination |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1018 |
internally). |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1019 |
|
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1020 |
The ``@{text "?"}'' version of @{attribute iff} declares rules to |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1021 |
the Isabelle/Pure context only, and omits the Simplifier |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1022 |
declaration. |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1023 |
|
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1024 |
\item @{attribute swapped} turns an introduction rule into an |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1025 |
elimination, by resolving with the classical swap principle @{text |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1026 |
"\<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
|
1027 |
illustrative purposes: the Classical Reasoner already swaps rules |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1028 |
internally as explained above. |
9d946de41120
updated and re-unified classical rule declarations;
wenzelm
parents:
42927
diff
changeset
|
1029 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
1030 |
\end{description} |
26782 | 1031 |
*} |
1032 |
||
1033 |
||
43365 | 1034 |
subsection {* Structured methods *} |
1035 |
||
1036 |
text {* |
|
1037 |
\begin{matharray}{rcl} |
|
1038 |
@{method_def rule} & : & @{text method} \\ |
|
1039 |
@{method_def contradiction} & : & @{text method} \\ |
|
1040 |
\end{matharray} |
|
1041 |
||
1042 |
@{rail " |
|
1043 |
@@{method rule} @{syntax thmrefs}? |
|
1044 |
"} |
|
1045 |
||
1046 |
\begin{description} |
|
1047 |
||
1048 |
\item @{method rule} as offered by the Classical Reasoner is a |
|
1049 |
refinement over the Pure one (see \secref{sec:pure-meth-att}). Both |
|
1050 |
versions work the same, but the classical version observes the |
|
1051 |
classical rule context in addition to that of Isabelle/Pure. |
|
1052 |
||
1053 |
Common object logics (HOL, ZF, etc.) declare a rich collection of |
|
1054 |
classical rules (even if these would qualify as intuitionistic |
|
1055 |
ones), but only few declarations to the rule context of |
|
1056 |
Isabelle/Pure (\secref{sec:pure-meth-att}). |
|
1057 |
||
1058 |
\item @{method contradiction} solves some goal by contradiction, |
|
1059 |
deriving any result from both @{text "\<not> A"} and @{text A}. Chained |
|
1060 |
facts, which are guaranteed to participate, may appear in either |
|
1061 |
order. |
|
1062 |
||
1063 |
\end{description} |
|
1064 |
*} |
|
1065 |
||
1066 |
||
27040 | 1067 |
subsection {* Automated methods *} |
26782 | 1068 |
|
1069 |
text {* |
|
1070 |
\begin{matharray}{rcl} |
|
28761
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
1071 |
@{method_def blast} & : & @{text method} \\ |
42930 | 1072 |
@{method_def auto} & : & @{text method} \\ |
1073 |
@{method_def force} & : & @{text method} \\ |
|
28761
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
1074 |
@{method_def fast} & : & @{text method} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
1075 |
@{method_def slow} & : & @{text method} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
1076 |
@{method_def best} & : & @{text method} \\ |
44911 | 1077 |
@{method_def fastforce} & : & @{text method} \\ |
28761
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
1078 |
@{method_def slowsimp} & : & @{text method} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
1079 |
@{method_def bestsimp} & : & @{text method} \\ |
43367 | 1080 |
@{method_def deepen} & : & @{text method} \\ |
26782 | 1081 |
\end{matharray} |
1082 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
1083 |
@{rail " |
42930 | 1084 |
@@{method blast} @{syntax nat}? (@{syntax clamod} * ) |
1085 |
; |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
1086 |
@@{method auto} (@{syntax nat} @{syntax nat})? (@{syntax clasimpmod} * ) |
26782 | 1087 |
; |
42930 | 1088 |
@@{method force} (@{syntax clasimpmod} * ) |
1089 |
; |
|
1090 |
(@@{method fast} | @@{method slow} | @@{method best}) (@{syntax clamod} * ) |
|
26782 | 1091 |
; |
44911 | 1092 |
(@@{method fastforce} | @@{method slowsimp} | @@{method bestsimp}) |
42930 | 1093 |
(@{syntax clasimpmod} * ) |
1094 |
; |
|
43367 | 1095 |
@@{method deepen} (@{syntax nat} ?) (@{syntax clamod} * ) |
1096 |
; |
|
42930 | 1097 |
@{syntax_def clamod}: |
1098 |
(('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del') ':' @{syntax thmrefs} |
|
1099 |
; |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
1100 |
@{syntax_def clasimpmod}: ('simp' (() | 'add' | 'del' | 'only') | |
26782 | 1101 |
('cong' | 'split') (() | 'add' | 'del') | |
1102 |
'iff' (((() | 'add') '?'?) | 'del') | |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
1103 |
(('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del')) ':' @{syntax thmrefs} |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
1104 |
"} |
26782 | 1105 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
1106 |
\begin{description} |
26782 | 1107 |
|
42930 | 1108 |
\item @{method blast} is a separate classical tableau prover that |
1109 |
uses the same classical rule declarations as explained before. |
|
1110 |
||
1111 |
Proof search is coded directly in ML using special data structures. |
|
1112 |
A successful proof is then reconstructed using regular Isabelle |
|
1113 |
inferences. It is faster and more powerful than the other classical |
|
1114 |
reasoning tools, but has major limitations too. |
|
1115 |
||
1116 |
\begin{itemize} |
|
1117 |
||
1118 |
\item It does not use the classical wrapper tacticals, such as the |
|
44911 | 1119 |
integration with the Simplifier of @{method fastforce}. |
42930 | 1120 |
|
1121 |
\item It does not perform higher-order unification, as needed by the |
|
1122 |
rule @{thm [source=false] rangeI} in HOL. There are often |
|
1123 |
alternatives to such rules, for example @{thm [source=false] |
|
1124 |
range_eqI}. |
|
1125 |
||
1126 |
\item Function variables may only be applied to parameters of the |
|
1127 |
subgoal. (This restriction arises because the prover does not use |
|
1128 |
higher-order unification.) If other function variables are present |
|
1129 |
then the prover will fail with the message \texttt{Function Var's |
|
1130 |
argument not a bound variable}. |
|
1131 |
||
1132 |
\item Its proof strategy is more general than @{method fast} but can |
|
1133 |
be slower. If @{method blast} fails or seems to be running forever, |
|
1134 |
try @{method fast} and the other proof tools described below. |
|
1135 |
||
1136 |
\end{itemize} |
|
1137 |
||
1138 |
The optional integer argument specifies a bound for the number of |
|
1139 |
unsafe steps used in a proof. By default, @{method blast} starts |
|
1140 |
with a bound of 0 and increases it successively to 20. In contrast, |
|
1141 |
@{text "(blast lim)"} tries to prove the goal using a search bound |
|
1142 |
of @{text "lim"}. Sometimes a slow proof using @{method blast} can |
|
1143 |
be made much faster by supplying the successful search bound to this |
|
1144 |
proof method instead. |
|
1145 |
||
1146 |
\item @{method auto} combines classical reasoning with |
|
1147 |
simplification. It is intended for situations where there are a lot |
|
1148 |
of mostly trivial subgoals; it proves all the easy ones, leaving the |
|
1149 |
ones it cannot prove. Occasionally, attempting to prove the hard |
|
1150 |
ones may take a long time. |
|
1151 |
||
43332 | 1152 |
The optional depth arguments in @{text "(auto m n)"} refer to its |
1153 |
builtin classical reasoning procedures: @{text m} (default 4) is for |
|
1154 |
@{method blast}, which is tried first, and @{text n} (default 2) is |
|
1155 |
for a slower but more general alternative that also takes wrappers |
|
1156 |
into account. |
|
42930 | 1157 |
|
1158 |
\item @{method force} is intended to prove the first subgoal |
|
1159 |
completely, using many fancy proof tools and performing a rather |
|
1160 |
exhaustive search. As a result, proof attempts may take rather long |
|
1161 |
or diverge easily. |
|
1162 |
||
1163 |
\item @{method fast}, @{method best}, @{method slow} attempt to |
|
1164 |
prove the first subgoal using sequent-style reasoning as explained |
|
1165 |
before. Unlike @{method blast}, they construct proofs directly in |
|
1166 |
Isabelle. |
|
26782 | 1167 |
|
42930 | 1168 |
There is a difference in search strategy and back-tracking: @{method |
1169 |
fast} uses depth-first search and @{method best} uses best-first |
|
1170 |
search (guided by a heuristic function: normally the total size of |
|
1171 |
the proof state). |
|
1172 |
||
1173 |
Method @{method slow} is like @{method fast}, but conducts a broader |
|
1174 |
search: it may, when backtracking from a failed proof attempt, undo |
|
1175 |
even the step of proving a subgoal by assumption. |
|
1176 |
||
47967
c422128d3889
discontinued obsolete method fastsimp / tactic fast_simp_tac;
wenzelm
parents:
47497
diff
changeset
|
1177 |
\item @{method fastforce}, @{method slowsimp}, @{method bestsimp} |
c422128d3889
discontinued obsolete method fastsimp / tactic fast_simp_tac;
wenzelm
parents:
47497
diff
changeset
|
1178 |
are like @{method fast}, @{method slow}, @{method best}, |
c422128d3889
discontinued obsolete method fastsimp / tactic fast_simp_tac;
wenzelm
parents:
47497
diff
changeset
|
1179 |
respectively, but use the Simplifier as additional wrapper. The name |
c422128d3889
discontinued obsolete method fastsimp / tactic fast_simp_tac;
wenzelm
parents:
47497
diff
changeset
|
1180 |
@{method fastforce}, reflects the behaviour of this popular method |
c422128d3889
discontinued obsolete method fastsimp / tactic fast_simp_tac;
wenzelm
parents:
47497
diff
changeset
|
1181 |
better without requiring an understanding of its implementation. |
42930 | 1182 |
|
43367 | 1183 |
\item @{method deepen} works by exhaustive search up to a certain |
1184 |
depth. The start depth is 4 (unless specified explicitly), and the |
|
1185 |
depth is increased iteratively up to 10. Unsafe rules are modified |
|
1186 |
to preserve the formula they act on, so that it be used repeatedly. |
|
1187 |
This method can prove more goals than @{method fast}, but is much |
|
1188 |
slower, for example if the assumptions have many universal |
|
1189 |
quantifiers. |
|
1190 |
||
42930 | 1191 |
\end{description} |
1192 |
||
1193 |
Any of the above methods support additional modifiers of the context |
|
1194 |
of classical (and simplifier) rules, but the ones related to the |
|
1195 |
Simplifier are explicitly prefixed by @{text simp} here. The |
|
1196 |
semantics of these ad-hoc rule declarations is analogous to the |
|
1197 |
attributes given before. Facts provided by forward chaining are |
|
1198 |
inserted into the goal before commencing proof search. |
|
1199 |
*} |
|
1200 |
||
1201 |
||
1202 |
subsection {* Semi-automated methods *} |
|
1203 |
||
1204 |
text {* These proof methods may help in situations when the |
|
1205 |
fully-automated tools fail. The result is a simpler subgoal that |
|
1206 |
can be tackled by other means, such as by manual instantiation of |
|
1207 |
quantifiers. |
|
1208 |
||
1209 |
\begin{matharray}{rcl} |
|
1210 |
@{method_def safe} & : & @{text method} \\ |
|
1211 |
@{method_def clarify} & : & @{text method} \\ |
|
1212 |
@{method_def clarsimp} & : & @{text method} \\ |
|
1213 |
\end{matharray} |
|
1214 |
||
1215 |
@{rail " |
|
1216 |
(@@{method safe} | @@{method clarify}) (@{syntax clamod} * ) |
|
1217 |
; |
|
1218 |
@@{method clarsimp} (@{syntax clasimpmod} * ) |
|
1219 |
"} |
|
1220 |
||
1221 |
\begin{description} |
|
1222 |
||
1223 |
\item @{method safe} repeatedly performs safe steps on all subgoals. |
|
1224 |
It is deterministic, with at most one outcome. |
|
1225 |
||
43366 | 1226 |
\item @{method clarify} performs a series of safe steps without |
1227 |
splitting subgoals; see also @{ML clarify_step_tac}. |
|
42930 | 1228 |
|
1229 |
\item @{method clarsimp} acts like @{method clarify}, but also does |
|
1230 |
simplification. Note that if the Simplifier context includes a |
|
1231 |
splitter for the premises, the subgoal may still be split. |
|
26782 | 1232 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
1233 |
\end{description} |
26782 | 1234 |
*} |
1235 |
||
1236 |
||
43366 | 1237 |
subsection {* Single-step tactics *} |
1238 |
||
1239 |
text {* |
|
1240 |
\begin{matharray}{rcl} |
|
1241 |
@{index_ML safe_step_tac: "Proof.context -> int -> tactic"} \\ |
|
1242 |
@{index_ML inst_step_tac: "Proof.context -> int -> tactic"} \\ |
|
1243 |
@{index_ML step_tac: "Proof.context -> int -> tactic"} \\ |
|
1244 |
@{index_ML slow_step_tac: "Proof.context -> int -> tactic"} \\ |
|
1245 |
@{index_ML clarify_step_tac: "Proof.context -> int -> tactic"} \\ |
|
1246 |
\end{matharray} |
|
1247 |
||
1248 |
These are the primitive tactics behind the (semi)automated proof |
|
1249 |
methods of the Classical Reasoner. By calling them yourself, you |
|
1250 |
can execute these procedures one step at a time. |
|
1251 |
||
1252 |
\begin{description} |
|
1253 |
||
1254 |
\item @{ML safe_step_tac}~@{text "ctxt i"} performs a safe step on |
|
1255 |
subgoal @{text i}. The safe wrapper tacticals are applied to a |
|
1256 |
tactic that may include proof by assumption or Modus Ponens (taking |
|
1257 |
care not to instantiate unknowns), or substitution. |
|
1258 |
||
1259 |
\item @{ML inst_step_tac} is like @{ML safe_step_tac}, but allows |
|
1260 |
unknowns to be instantiated. |
|
1261 |
||
1262 |
\item @{ML step_tac}~@{text "ctxt i"} is the basic step of the proof |
|
1263 |
procedure. The unsafe wrapper tacticals are applied to a tactic |
|
1264 |
that tries @{ML safe_tac}, @{ML inst_step_tac}, or applies an unsafe |
|
1265 |
rule from the context. |
|
1266 |
||
1267 |
\item @{ML slow_step_tac} resembles @{ML step_tac}, but allows |
|
1268 |
backtracking between using safe rules with instantiation (@{ML |
|
1269 |
inst_step_tac}) and using unsafe rules. The resulting search space |
|
1270 |
is larger. |
|
1271 |
||
1272 |
\item @{ML clarify_step_tac}~@{text "ctxt i"} performs a safe step |
|
1273 |
on subgoal @{text i}. No splitting step is applied; for example, |
|
1274 |
the subgoal @{text "A \<and> B"} is left as a conjunction. Proof by |
|
1275 |
assumption, Modus Ponens, etc., may be performed provided they do |
|
1276 |
not instantiate unknowns. Assumptions of the form @{text "x = t"} |
|
1277 |
may be eliminated. The safe wrapper tactical is applied. |
|
1278 |
||
1279 |
\end{description} |
|
1280 |
*} |
|
1281 |
||
1282 |
||
27044 | 1283 |
section {* Object-logic setup \label{sec:object-logic} *} |
26790 | 1284 |
|
1285 |
text {* |
|
1286 |
\begin{matharray}{rcl} |
|
28761
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
1287 |
@{command_def "judgment"} & : & @{text "theory \<rightarrow> theory"} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
1288 |
@{method_def atomize} & : & @{text method} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
1289 |
@{attribute_def atomize} & : & @{text attribute} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
1290 |
@{attribute_def rule_format} & : & @{text attribute} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
1291 |
@{attribute_def rulify} & : & @{text attribute} \\ |
26790 | 1292 |
\end{matharray} |
1293 |
||
1294 |
The very starting point for any Isabelle object-logic is a ``truth |
|
1295 |
judgment'' that links object-level statements to the meta-logic |
|
1296 |
(with its minimal language of @{text prop} that covers universal |
|
1297 |
quantification @{text "\<And>"} and implication @{text "\<Longrightarrow>"}). |
|
1298 |
||
1299 |
Common object-logics are sufficiently expressive to internalize rule |
|
1300 |
statements over @{text "\<And>"} and @{text "\<Longrightarrow>"} within their own |
|
1301 |
language. This is useful in certain situations where a rule needs |
|
1302 |
to be viewed as an atomic statement from the meta-level perspective, |
|
1303 |
e.g.\ @{text "\<And>x. x \<in> A \<Longrightarrow> P x"} versus @{text "\<forall>x \<in> A. P x"}. |
|
1304 |
||
1305 |
From the following language elements, only the @{method atomize} |
|
1306 |
method and @{attribute rule_format} attribute are occasionally |
|
1307 |
required by end-users, the rest is for those who need to setup their |
|
1308 |
own object-logic. In the latter case existing formulations of |
|
1309 |
Isabelle/FOL or Isabelle/HOL may be taken as realistic examples. |
|
1310 |
||
1311 |
Generic tools may refer to the information provided by object-logic |
|
1312 |
declarations internally. |
|
1313 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
1314 |
@{rail " |
46494
ea2ae63336f3
clarified outer syntax "constdecl", which is only local to some rail diagrams;
wenzelm
parents:
46277
diff
changeset
|
1315 |
@@{command judgment} @{syntax name} '::' @{syntax type} @{syntax mixfix}? |
26790 | 1316 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
1317 |
@@{attribute atomize} ('(' 'full' ')')? |
26790 | 1318 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
1319 |
@@{attribute rule_format} ('(' 'noasm' ')')? |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40291
diff
changeset
|
1320 |
"} |
26790 | 1321 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
1322 |
\begin{description} |
26790 | 1323 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
1324 |
\item @{command "judgment"}~@{text "c :: \<sigma> (mx)"} declares constant |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
1325 |
@{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
|
1326 |
type @{text \<sigma>} should specify a coercion of the category of |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
1327 |
object-level propositions to @{text prop} of the Pure meta-logic; |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
1328 |
the mixfix annotation @{text "(mx)"} would typically just link the |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
1329 |
object language (internally of syntactic category @{text logic}) |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
1330 |
with that of @{text prop}. Only one @{command "judgment"} |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
1331 |
declaration may be given in any theory development. |
26790 | 1332 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
1333 |
\item @{method atomize} (as a method) rewrites any non-atomic |
26790 | 1334 |
premises of a sub-goal, using the meta-level equations declared via |
1335 |
@{attribute atomize} (as an attribute) beforehand. As a result, |
|
1336 |
heavily nested goals become amenable to fundamental operations such |
|
42626 | 1337 |
as resolution (cf.\ the @{method (Pure) rule} method). Giving the ``@{text |
26790 | 1338 |
"(full)"}'' option here means to turn the whole subgoal into an |
1339 |
object-statement (if possible), including the outermost parameters |
|
1340 |
and assumptions as well. |
|
1341 |
||
1342 |
A typical collection of @{attribute atomize} rules for a particular |
|
1343 |
object-logic would provide an internalization for each of the |
|
1344 |
connectives of @{text "\<And>"}, @{text "\<Longrightarrow>"}, and @{text "\<equiv>"}. |
|
1345 |
Meta-level conjunction should be covered as well (this is |
|
1346 |
particularly important for locales, see \secref{sec:locale}). |
|
1347 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
1348 |
\item @{attribute rule_format} rewrites a theorem by the equalities |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
1349 |
declared as @{attribute rulify} rules in the current object-logic. |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
1350 |
By default, the result is fully normalized, including assumptions |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
1351 |
and conclusions at any depth. The @{text "(no_asm)"} option |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
1352 |
restricts the transformation to the conclusion of a rule. |
26790 | 1353 |
|
1354 |
In common object-logics (HOL, FOL, ZF), the effect of @{attribute |
|
1355 |
rule_format} is to replace (bounded) universal quantification |
|
1356 |
(@{text "\<forall>"}) and implication (@{text "\<longrightarrow>"}) by the corresponding |
|
1357 |
rule statements over @{text "\<And>"} and @{text "\<Longrightarrow>"}. |
|
1358 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28754
diff
changeset
|
1359 |
\end{description} |
26790 | 1360 |
*} |
1361 |
||
26782 | 1362 |
end |