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