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