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