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