| author | nipkow | 
| Wed, 02 Sep 2009 12:20:17 +0200 | |
| changeset 32489 | 071e4ae83149 | 
| parent 30462 | 0b857a58b15e | 
| child 35613 | 9d3ff36ad4e1 | 
| permissions | -rw-r--r-- | 
| 26782 | 1 | theory Generic | 
| 26894 | 2 | imports Main | 
| 26782 | 3 | begin | 
| 4 | ||
| 5 | chapter {* Generic tools and packages \label{ch:gen-tools} *}
 | |
| 6 | ||
| 27040 | 7 | section {* Configuration options *}
 | 
| 26782 | 8 | |
| 9 | text {*
 | |
| 10 | Isabelle/Pure maintains a record of named configuration options | |
| 11 |   within the theory or proof context, with values of type @{ML_type
 | |
| 12 |   bool}, @{ML_type int}, or @{ML_type string}.  Tools may declare
 | |
| 13 | options in ML, and then refer to these values (relative to the | |
| 14 | context). Thus global reference variables are easily avoided. The | |
| 15 | user may change the value of a configuration option by means of an | |
| 16 | associated attribute of the same name. This form of context | |
| 17 |   declaration works particularly well with commands such as @{command
 | |
| 18 |   "declare"} or @{command "using"}.
 | |
| 19 | ||
| 20 | For historical reasons, some tools cannot take the full proof | |
| 21 | context into account and merely refer to the background theory. | |
| 22 | This is accommodated by configuration options being declared as | |
| 23 | ``global'', which may not be changed within a local context. | |
| 24 | ||
| 25 |   \begin{matharray}{rcll}
 | |
| 28761 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 26 |     @{command_def "print_configs"} & : & @{text "context \<rightarrow>"} \\
 | 
| 26782 | 27 |   \end{matharray}
 | 
| 28 | ||
| 29 |   \begin{rail}
 | |
| 30 |     name ('=' ('true' | 'false' | int | name))?
 | |
| 31 |   \end{rail}
 | |
| 32 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 33 |   \begin{description}
 | 
| 26782 | 34 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 35 |   \item @{command "print_configs"} prints the available configuration
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 36 | options, with names, types, and current values. | 
| 26782 | 37 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 38 |   \item @{text "name = value"} as an attribute expression modifies the
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 39 | 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 | 40 |   type.  For @{ML_type bool} the default value is @{text true}.  Any
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 41 | attempt to change a global option in a local context is ignored. | 
| 26782 | 42 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 43 |   \end{description}
 | 
| 26782 | 44 | *} | 
| 45 | ||
| 46 | ||
| 27040 | 47 | section {* Basic proof tools *}
 | 
| 26782 | 48 | |
| 49 | subsection {* Miscellaneous methods and attributes \label{sec:misc-meth-att} *}
 | |
| 50 | ||
| 51 | text {*
 | |
| 52 |   \begin{matharray}{rcl}
 | |
| 28761 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 53 |     @{method_def unfold} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 54 |     @{method_def fold} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 55 |     @{method_def insert} & : & @{text method} \\[0.5ex]
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 56 |     @{method_def erule}@{text "\<^sup>*"} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 57 |     @{method_def drule}@{text "\<^sup>*"} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 58 |     @{method_def frule}@{text "\<^sup>*"} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 59 |     @{method_def succeed} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 60 |     @{method_def fail} & : & @{text method} \\
 | 
| 26782 | 61 |   \end{matharray}
 | 
| 62 | ||
| 63 |   \begin{rail}
 | |
| 64 |     ('fold' | 'unfold' | 'insert') thmrefs
 | |
| 65 | ; | |
| 66 |     ('erule' | 'drule' | 'frule') ('('nat')')? thmrefs
 | |
| 67 | ; | |
| 68 |   \end{rail}
 | |
| 69 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 70 |   \begin{description}
 | 
| 26782 | 71 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 72 |   \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 | 73 | "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 | 74 | all goals; any chained facts provided are inserted into the goal and | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 75 | subject to rewriting as well. | 
| 26782 | 76 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 77 |   \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 | 78 | into all goals of the proof state. Note that current facts | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 79 | indicated for forward chaining are ignored. | 
| 26782 | 80 | |
| 30397 | 81 |   \item @{method erule}~@{text "a\<^sub>1 \<dots> a\<^sub>n"}, @{method
 | 
| 82 |   drule}~@{text "a\<^sub>1 \<dots> a\<^sub>n"}, and @{method frule}~@{text
 | |
| 83 |   "a\<^sub>1 \<dots> a\<^sub>n"} are similar to the basic @{method rule}
 | |
| 84 |   method (see \secref{sec:pure-meth-att}), but apply rules by
 | |
| 85 | elim-resolution, destruct-resolution, and forward-resolution, | |
| 86 |   respectively \cite{isabelle-implementation}.  The optional natural
 | |
| 87 | number argument (default 0) specifies additional assumption steps to | |
| 88 | be performed here. | |
| 26782 | 89 | |
| 90 | Note that these methods are improper ones, mainly serving for | |
| 91 | experimentation and tactic script emulation. Different modes of | |
| 92 | basic rule application are usually expressed in Isar at the proof | |
| 93 | language level, rather than via implicit proof state manipulations. | |
| 94 | For example, a proper single-step elimination would be done using | |
| 95 |   the plain @{method rule} method, with forward chaining of current
 | |
| 96 | facts. | |
| 97 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 98 |   \item @{method succeed} yields a single (unchanged) result; it is
 | 
| 26782 | 99 |   the identity of the ``@{text ","}'' method combinator (cf.\
 | 
| 28754 
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
 wenzelm parents: 
27248diff
changeset | 100 |   \secref{sec:proof-meth}).
 | 
| 26782 | 101 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 102 |   \item @{method fail} yields an empty result sequence; it is the
 | 
| 26782 | 103 |   identity of the ``@{text "|"}'' method combinator (cf.\
 | 
| 28754 
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
 wenzelm parents: 
27248diff
changeset | 104 |   \secref{sec:proof-meth}).
 | 
| 26782 | 105 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 106 |   \end{description}
 | 
| 26782 | 107 | |
| 108 |   \begin{matharray}{rcl}
 | |
| 28761 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 109 |     @{attribute_def tagged} & : & @{text attribute} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 110 |     @{attribute_def untagged} & : & @{text attribute} \\[0.5ex]
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 111 |     @{attribute_def THEN} & : & @{text attribute} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 112 |     @{attribute_def COMP} & : & @{text attribute} \\[0.5ex]
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 113 |     @{attribute_def unfolded} & : & @{text attribute} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 114 |     @{attribute_def folded} & : & @{text attribute} \\[0.5ex]
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 115 |     @{attribute_def rotated} & : & @{text attribute} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 116 |     @{attribute_def (Pure) elim_format} & : & @{text attribute} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 117 |     @{attribute_def standard}@{text "\<^sup>*"} & : & @{text attribute} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 118 |     @{attribute_def no_vars}@{text "\<^sup>*"} & : & @{text attribute} \\
 | 
| 26782 | 119 |   \end{matharray}
 | 
| 120 | ||
| 121 |   \begin{rail}
 | |
| 28764 | 122 | 'tagged' name name | 
| 26782 | 123 | ; | 
| 124 | 'untagged' name | |
| 125 | ; | |
| 126 |     ('THEN' | 'COMP') ('[' nat ']')? thmref
 | |
| 127 | ; | |
| 128 |     ('unfolded' | 'folded') thmrefs
 | |
| 129 | ; | |
| 130 | 'rotated' ( int )? | |
| 131 |   \end{rail}
 | |
| 132 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 133 |   \begin{description}
 | 
| 26782 | 134 | |
| 28764 | 135 |   \item @{attribute tagged}~@{text "name value"} and @{attribute
 | 
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 136 |   untagged}~@{text name} add and remove \emph{tags} of some theorem.
 | 
| 26782 | 137 | Tags may be any list of string pairs that serve as formal comment. | 
| 28764 | 138 | The first string is considered the tag name, the second its value. | 
| 139 |   Note that @{attribute untagged} removes any tags of the same name.
 | |
| 26782 | 140 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 141 |   \item @{attribute THEN}~@{text a} and @{attribute COMP}~@{text a}
 | 
| 26782 | 142 |   compose rules by resolution.  @{attribute THEN} resolves with the
 | 
| 143 |   first premise of @{text a} (an alternative position may be also
 | |
| 144 |   specified); the @{attribute COMP} version skips the automatic
 | |
| 30462 | 145 |   lifting process that is normally intended (cf.\ @{ML "op RS"} and
 | 
| 146 |   @{ML "op COMP"} in \cite{isabelle-implementation}).
 | |
| 26782 | 147 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 148 |   \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 | 149 |   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 | 150 | definitions throughout a rule. | 
| 26782 | 151 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 152 |   \item @{attribute rotated}~@{text n} rotate the premises of a
 | 
| 26782 | 153 |   theorem by @{text n} (default 1).
 | 
| 154 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 155 |   \item @{attribute (Pure) elim_format} turns a destruction rule into
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 156 |   elimination rule format, by resolving with the rule @{prop "PROP A \<Longrightarrow>
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 157 | (PROP A \<Longrightarrow> PROP B) \<Longrightarrow> PROP B"}. | 
| 26782 | 158 | |
| 159 |   Note that the Classical Reasoner (\secref{sec:classical}) provides
 | |
| 160 | its own version of this operation. | |
| 161 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 162 |   \item @{attribute standard} puts a theorem into the standard form of
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 163 | object-rules at the outermost theory level. Note that this | 
| 26782 | 164 | operation violates the local proof context (including active | 
| 165 | locales). | |
| 166 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 167 |   \item @{attribute no_vars} replaces schematic variables by free
 | 
| 26782 | 168 | ones; this is mainly for tuning output of pretty printed theorems. | 
| 169 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 170 |   \end{description}
 | 
| 26782 | 171 | *} | 
| 172 | ||
| 173 | ||
| 27044 | 174 | subsection {* Low-level equational reasoning *}
 | 
| 175 | ||
| 176 | text {*
 | |
| 177 |   \begin{matharray}{rcl}
 | |
| 28761 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 178 |     @{method_def subst} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 179 |     @{method_def hypsubst} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 180 |     @{method_def split} & : & @{text method} \\
 | 
| 27044 | 181 |   \end{matharray}
 | 
| 182 | ||
| 183 |   \begin{rail}
 | |
| 184 |     'subst' ('(' 'asm' ')')? ('(' (nat+) ')')? thmref
 | |
| 185 | ; | |
| 186 |     'split' ('(' 'asm' ')')? thmrefs
 | |
| 187 | ; | |
| 188 |   \end{rail}
 | |
| 189 | ||
| 190 | These methods provide low-level facilities for equational reasoning | |
| 191 | that are intended for specialized applications only. Normally, | |
| 192 | single step calculations would be performed in a structured text | |
| 193 |   (see also \secref{sec:calculation}), while the Simplifier methods
 | |
| 194 | provide the canonical way for automated normalization (see | |
| 195 |   \secref{sec:simplifier}).
 | |
| 196 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 197 |   \begin{description}
 | 
| 27044 | 198 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 199 |   \item @{method subst}~@{text eq} performs a single substitution step
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 200 |   using rule @{text eq}, which may be either a meta or object
 | 
| 27044 | 201 | equality. | 
| 202 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 203 |   \item @{method subst}~@{text "(asm) eq"} substitutes in an
 | 
| 27044 | 204 | assumption. | 
| 205 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 206 |   \item @{method subst}~@{text "(i \<dots> j) eq"} performs several
 | 
| 27044 | 207 |   substitutions in the conclusion. The numbers @{text i} to @{text j}
 | 
| 208 | indicate the positions to substitute at. Positions are ordered from | |
| 209 | the top of the term tree moving down from left to right. For | |
| 210 |   example, in @{text "(a + b) + (c + d)"} there are three positions
 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 211 |   where commutativity of @{text "+"} is applicable: 1 refers to @{text
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 212 |   "a + b"}, 2 to the whole term, and 3 to @{text "c + d"}.
 | 
| 27044 | 213 | |
| 214 |   If the positions in the list @{text "(i \<dots> j)"} are non-overlapping
 | |
| 215 |   (e.g.\ @{text "(2 3)"} in @{text "(a + b) + (c + d)"}) you may
 | |
| 216 | assume all substitutions are performed simultaneously. Otherwise | |
| 217 |   the behaviour of @{text subst} is not specified.
 | |
| 218 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 219 |   \item @{method subst}~@{text "(asm) (i \<dots> j) eq"} performs the
 | 
| 27071 | 220 | substitutions in the assumptions. The positions refer to the | 
| 221 | assumptions in order from left to right. For example, given in a | |
| 222 |   goal of the form @{text "P (a + b) \<Longrightarrow> P (c + d) \<Longrightarrow> \<dots>"}, position 1 of
 | |
| 223 |   commutativity of @{text "+"} is the subterm @{text "a + b"} and
 | |
| 224 |   position 2 is the subterm @{text "c + d"}.
 | |
| 27044 | 225 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 226 |   \item @{method hypsubst} performs substitution using some
 | 
| 27044 | 227 |   assumption; this only works for equations of the form @{text "x =
 | 
| 228 |   t"} where @{text x} is a free or bound variable.
 | |
| 229 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 230 |   \item @{method split}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} performs single-step case
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 231 | splitting using the given rules. By default, splitting is performed | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 232 |   in the conclusion of a goal; the @{text "(asm)"} option indicates to
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 233 | operate on assumptions instead. | 
| 27044 | 234 | |
| 235 |   Note that the @{method simp} method already involves repeated
 | |
| 236 | application of split rules as declared in the current context. | |
| 237 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 238 |   \end{description}
 | 
| 27044 | 239 | *} | 
| 240 | ||
| 241 | ||
| 26782 | 242 | subsection {* Further tactic emulations \label{sec:tactics} *}
 | 
| 243 | ||
| 244 | text {*
 | |
| 245 | The following improper proof methods emulate traditional tactics. | |
| 246 | These admit direct access to the goal state, which is normally | |
| 247 | considered harmful! In particular, this may involve both numbered | |
| 248 | goal addressing (default 1), and dynamic instantiation within the | |
| 249 | scope of some subgoal. | |
| 250 | ||
| 251 |   \begin{warn}
 | |
| 252 | Dynamic instantiations refer to universally quantified parameters | |
| 253 | of a subgoal (the dynamic context) rather than fixed variables and | |
| 254 | term abbreviations of a (static) Isar context. | |
| 255 |   \end{warn}
 | |
| 256 | ||
| 257 | Tactic emulation methods, unlike their ML counterparts, admit | |
| 258 | simultaneous instantiation from both dynamic and static contexts. | |
| 259 | If names occur in both contexts goal parameters hide locally fixed | |
| 260 | variables. Likewise, schematic variables refer to term | |
| 261 | abbreviations, if present in the static context. Otherwise the | |
| 262 | schematic variable is interpreted as a schematic variable and left | |
| 263 | to be solved by unification with certain parts of the subgoal. | |
| 264 | ||
| 265 | Note that the tactic emulation proof methods in Isabelle/Isar are | |
| 266 |   consistently named @{text foo_tac}.  Note also that variable names
 | |
| 267 | occurring on left hand sides of instantiations must be preceded by a | |
| 268 | question mark if they coincide with a keyword or contain dots. This | |
| 269 |   is consistent with the attribute @{attribute "where"} (see
 | |
| 270 |   \secref{sec:pure-meth-att}).
 | |
| 271 | ||
| 272 |   \begin{matharray}{rcl}
 | |
| 28761 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 273 |     @{method_def rule_tac}@{text "\<^sup>*"} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 274 |     @{method_def erule_tac}@{text "\<^sup>*"} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 275 |     @{method_def drule_tac}@{text "\<^sup>*"} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 276 |     @{method_def frule_tac}@{text "\<^sup>*"} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 277 |     @{method_def cut_tac}@{text "\<^sup>*"} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 278 |     @{method_def thin_tac}@{text "\<^sup>*"} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 279 |     @{method_def subgoal_tac}@{text "\<^sup>*"} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 280 |     @{method_def rename_tac}@{text "\<^sup>*"} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 281 |     @{method_def rotate_tac}@{text "\<^sup>*"} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 282 |     @{method_def tactic}@{text "\<^sup>*"} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 283 |     @{method_def raw_tactic}@{text "\<^sup>*"} & : & @{text method} \\
 | 
| 26782 | 284 |   \end{matharray}
 | 
| 285 | ||
| 286 |   \begin{rail}
 | |
| 287 | ( 'rule\_tac' | 'erule\_tac' | 'drule\_tac' | 'frule\_tac' | 'cut\_tac' | 'thin\_tac' ) goalspec? | |
| 288 | ( insts thmref | thmrefs ) | |
| 289 | ; | |
| 290 | 'subgoal\_tac' goalspec? (prop +) | |
| 291 | ; | |
| 292 | 'rename\_tac' goalspec? (name +) | |
| 293 | ; | |
| 294 | 'rotate\_tac' goalspec? int? | |
| 295 | ; | |
| 27223 | 296 |     ('tactic' | 'raw_tactic') text
 | 
| 26782 | 297 | ; | 
| 298 | ||
| 299 | insts: ((name '=' term) + 'and') 'in' | |
| 300 | ; | |
| 301 |   \end{rail}
 | |
| 302 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 303 | \begin{description}
 | 
| 26782 | 304 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 305 |   \item @{method rule_tac} etc. do resolution of rules with explicit
 | 
| 26782 | 306 |   instantiation.  This works the same way as the ML tactics @{ML
 | 
| 30397 | 307 |   res_inst_tac} etc. (see \cite{isabelle-implementation})
 | 
| 26782 | 308 | |
| 309 | Multiple rules may be only given if there is no instantiation; then | |
| 310 |   @{method rule_tac} is the same as @{ML resolve_tac} in ML (see
 | |
| 30397 | 311 |   \cite{isabelle-implementation}).
 | 
| 26782 | 312 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 313 |   \item @{method cut_tac} inserts facts into the proof state as
 | 
| 27209 | 314 |   assumption of a subgoal, see also @{ML Tactic.cut_facts_tac} in
 | 
| 30397 | 315 |   \cite{isabelle-implementation}.  Note that the scope of schematic
 | 
| 26782 | 316 | variables is spread over the main goal statement. Instantiations | 
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 317 |   may be given as well, see also ML tactic @{ML cut_inst_tac} in
 | 
| 30397 | 318 |   \cite{isabelle-implementation}.
 | 
| 26782 | 319 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 320 |   \item @{method thin_tac}~@{text \<phi>} deletes the specified assumption
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 321 |   from a subgoal; note that @{text \<phi>} may contain schematic variables.
 | 
| 30397 | 322 |   See also @{ML thin_tac} in \cite{isabelle-implementation}.
 | 
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 323 | |
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 324 |   \item @{method subgoal_tac}~@{text \<phi>} adds @{text \<phi>} as an
 | 
| 27239 | 325 |   assumption to a subgoal.  See also @{ML subgoal_tac} and @{ML
 | 
| 30397 | 326 |   subgoals_tac} in \cite{isabelle-implementation}.
 | 
| 26782 | 327 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 328 |   \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 | 329 |   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 | 330 |   \emph{suffix} of variables.
 | 
| 26782 | 331 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 332 |   \item @{method rotate_tac}~@{text n} rotates the assumptions of a
 | 
| 26782 | 333 |   goal by @{text n} positions: from right to left if @{text n} is
 | 
| 334 |   positive, and from left to right if @{text n} is negative; the
 | |
| 335 |   default value is 1.  See also @{ML rotate_tac} in
 | |
| 30397 | 336 |   \cite{isabelle-implementation}.
 | 
| 26782 | 337 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 338 |   \item @{method tactic}~@{text "text"} produces a proof method from
 | 
| 26782 | 339 |   any ML text of type @{ML_type tactic}.  Apart from the usual ML
 | 
| 27223 | 340 | environment and the current proof context, the ML code may refer to | 
| 341 |   the locally bound values @{ML_text facts}, which indicates any
 | |
| 342 | current facts used for forward-chaining. | |
| 26782 | 343 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 344 |   \item @{method raw_tactic} is similar to @{method tactic}, but
 | 
| 27223 | 345 | presents the goal state in its raw internal form, where simultaneous | 
| 346 | subgoals appear as conjunction of the logical framework instead of | |
| 347 | the usual split into several subgoals. While feature this is useful | |
| 348 | for debugging of complex method definitions, it should not never | |
| 349 | appear in production theories. | |
| 26782 | 350 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 351 |   \end{description}
 | 
| 26782 | 352 | *} | 
| 353 | ||
| 354 | ||
| 27040 | 355 | section {* The Simplifier \label{sec:simplifier} *}
 | 
| 26782 | 356 | |
| 27040 | 357 | subsection {* Simplification methods *}
 | 
| 26782 | 358 | |
| 359 | text {*
 | |
| 360 |   \begin{matharray}{rcl}
 | |
| 28761 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 361 |     @{method_def simp} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 362 |     @{method_def simp_all} & : & @{text method} \\
 | 
| 26782 | 363 |   \end{matharray}
 | 
| 364 | ||
| 365 |   \indexouternonterm{simpmod}
 | |
| 366 |   \begin{rail}
 | |
| 367 |     ('simp' | 'simp\_all') ('!' ?) opt? (simpmod *)
 | |
| 368 | ; | |
| 369 | ||
| 27092 | 370 |     opt: '(' ('no\_asm' | 'no\_asm\_simp' | 'no\_asm\_use' | 'asm\_lr' ) ')'
 | 
| 26782 | 371 | ; | 
| 372 |     simpmod: ('add' | 'del' | 'only' | 'cong' (() | 'add' | 'del') |
 | |
| 373 | 'split' (() | 'add' | 'del')) ':' thmrefs | |
| 374 | ; | |
| 375 |   \end{rail}
 | |
| 376 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 377 |   \begin{description}
 | 
| 26782 | 378 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 379 |   \item @{method simp} invokes the Simplifier, after declaring
 | 
| 26782 | 380 | additional rules according to the arguments given. Note that the | 
| 381 |   \railtterm{only} modifier first removes all other rewrite rules,
 | |
| 382 | congruences, and looper tactics (including splits), and then behaves | |
| 383 |   like \railtterm{add}.
 | |
| 384 | ||
| 385 |   \medskip The \railtterm{cong} modifiers add or delete Simplifier
 | |
| 386 |   congruence rules (see also \cite{isabelle-ref}), the default is to
 | |
| 387 | add. | |
| 388 | ||
| 389 |   \medskip The \railtterm{split} modifiers add or delete rules for the
 | |
| 390 |   Splitter (see also \cite{isabelle-ref}), the default is to add.
 | |
| 391 | This works only if the Simplifier method has been properly setup to | |
| 392 | include the Splitter (all major object logics such HOL, HOLCF, FOL, | |
| 393 | ZF do this already). | |
| 394 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 395 |   \item @{method simp_all} is similar to @{method simp}, but acts on
 | 
| 26782 | 396 | all goals (backwards from the last to the first one). | 
| 397 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 398 |   \end{description}
 | 
| 26782 | 399 | |
| 400 | By default the Simplifier methods take local assumptions fully into | |
| 401 | account, using equational assumptions in the subsequent | |
| 402 | normalization process, or simplifying assumptions themselves (cf.\ | |
| 30397 | 403 |   @{ML asm_full_simp_tac} in \cite{isabelle-ref}).  In structured
 | 
| 404 | proofs this is usually quite well behaved in practice: just the | |
| 405 | local premises of the actual goal are involved, additional facts may | |
| 406 |   be inserted via explicit forward-chaining (via @{command "then"},
 | |
| 407 |   @{command "from"}, @{command "using"} etc.).  The full context of
 | |
| 408 |   premises is only included if the ``@{text "!"}'' (bang) argument is
 | |
| 409 | given, which should be used with some care, though. | |
| 26782 | 410 | |
| 411 | Additional Simplifier options may be specified to tune the behavior | |
| 412 | further (mostly for unstructured scripts with many accidental local | |
| 413 |   facts): ``@{text "(no_asm)"}'' means assumptions are ignored
 | |
| 414 |   completely (cf.\ @{ML simp_tac}), ``@{text "(no_asm_simp)"}'' means
 | |
| 415 | assumptions are used in the simplification of the conclusion but are | |
| 416 |   not themselves simplified (cf.\ @{ML asm_simp_tac}), and ``@{text
 | |
| 417 | "(no_asm_use)"}'' means assumptions are simplified but are not used | |
| 418 |   in the simplification of each other or the conclusion (cf.\ @{ML
 | |
| 419 | full_simp_tac}). For compatibility reasons, there is also an option | |
| 420 |   ``@{text "(asm_lr)"}'', which means that an assumption is only used
 | |
| 421 |   for simplifying assumptions which are to the right of it (cf.\ @{ML
 | |
| 422 | asm_lr_simp_tac}). | |
| 423 | ||
| 27092 | 424 |   The configuration option @{text "depth_limit"} limits the number of
 | 
| 26782 | 425 | recursive invocations of the simplifier during conditional | 
| 426 | rewriting. | |
| 427 | ||
| 428 | \medskip The Splitter package is usually configured to work as part | |
| 429 |   of the Simplifier.  The effect of repeatedly applying @{ML
 | |
| 430 |   split_tac} can be simulated by ``@{text "(simp only: split:
 | |
| 431 |   a\<^sub>1 \<dots> a\<^sub>n)"}''.  There is also a separate @{text split}
 | |
| 432 | method available for single-step case splitting. | |
| 433 | *} | |
| 434 | ||
| 435 | ||
| 27040 | 436 | subsection {* Declaring rules *}
 | 
| 26782 | 437 | |
| 438 | text {*
 | |
| 439 |   \begin{matharray}{rcl}
 | |
| 28761 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 440 |     @{command_def "print_simpset"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 441 |     @{attribute_def simp} & : & @{text attribute} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 442 |     @{attribute_def cong} & : & @{text attribute} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 443 |     @{attribute_def split} & : & @{text attribute} \\
 | 
| 26782 | 444 |   \end{matharray}
 | 
| 445 | ||
| 446 |   \begin{rail}
 | |
| 447 |     ('simp' | 'cong' | 'split') (() | 'add' | 'del')
 | |
| 448 | ; | |
| 449 |   \end{rail}
 | |
| 450 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 451 |   \begin{description}
 | 
| 26782 | 452 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 453 |   \item @{command "print_simpset"} prints the collection of rules
 | 
| 26782 | 454 | declared to the Simplifier, which is also known as ``simpset'' | 
| 455 |   internally \cite{isabelle-ref}.
 | |
| 456 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 457 |   \item @{attribute simp} declares simplification rules.
 | 
| 26782 | 458 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 459 |   \item @{attribute cong} declares congruence rules.
 | 
| 26782 | 460 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 461 |   \item @{attribute split} declares case split rules.
 | 
| 26782 | 462 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 463 |   \end{description}
 | 
| 26782 | 464 | *} | 
| 465 | ||
| 466 | ||
| 27040 | 467 | subsection {* Simplification procedures *}
 | 
| 26782 | 468 | |
| 469 | text {*
 | |
| 470 |   \begin{matharray}{rcl}
 | |
| 28761 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 471 |     @{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 | 472 |     simproc & : & @{text attribute} \\
 | 
| 26782 | 473 |   \end{matharray}
 | 
| 474 | ||
| 475 |   \begin{rail}
 | |
| 476 |     'simproc\_setup' name '(' (term + '|') ')' '=' text \\ ('identifier' (nameref+))?
 | |
| 477 | ; | |
| 478 | ||
| 479 |     'simproc' (('add' ':')? | 'del' ':') (name+)
 | |
| 480 | ; | |
| 481 |   \end{rail}
 | |
| 482 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 483 |   \begin{description}
 | 
| 26782 | 484 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 485 |   \item @{command "simproc_setup"} defines a named simplification
 | 
| 26782 | 486 | procedure that is invoked by the Simplifier whenever any of the | 
| 487 | given term patterns match the current redex. The implementation, | |
| 488 |   which is provided as ML source text, needs to be of type @{ML_type
 | |
| 489 |   "morphism -> simpset -> cterm -> thm option"}, where the @{ML_type
 | |
| 490 |   cterm} represents the current redex @{text r} and the result is
 | |
| 491 |   supposed to be some proven rewrite rule @{text "r \<equiv> r'"} (or a
 | |
| 492 |   generalized version), or @{ML NONE} to indicate failure.  The
 | |
| 493 |   @{ML_type simpset} argument holds the full context of the current
 | |
| 494 | Simplifier invocation, including the actual Isar proof context. The | |
| 495 |   @{ML_type morphism} informs about the difference of the original
 | |
| 496 | compilation context wrt.\ the one of the actual application later | |
| 497 |   on.  The optional @{keyword "identifier"} specifies theorems that
 | |
| 498 | represent the logical content of the abstract theory of this | |
| 499 | simproc. | |
| 500 | ||
| 501 | Morphisms and identifiers are only relevant for simprocs that are | |
| 502 | defined within a local target context, e.g.\ in a locale. | |
| 503 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 504 |   \item @{text "simproc add: name"} and @{text "simproc del: name"}
 | 
| 26782 | 505 | add or delete named simprocs to the current Simplifier context. The | 
| 506 |   default is to add a simproc.  Note that @{command "simproc_setup"}
 | |
| 507 | already adds the new simproc to the subsequent context. | |
| 508 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 509 |   \end{description}
 | 
| 26782 | 510 | *} | 
| 511 | ||
| 512 | ||
| 27040 | 513 | subsection {* Forward simplification *}
 | 
| 26782 | 514 | |
| 515 | text {*
 | |
| 516 |   \begin{matharray}{rcl}
 | |
| 28761 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 517 |     @{attribute_def simplified} & : & @{text attribute} \\
 | 
| 26782 | 518 |   \end{matharray}
 | 
| 519 | ||
| 520 |   \begin{rail}
 | |
| 521 | 'simplified' opt? thmrefs? | |
| 522 | ; | |
| 523 | ||
| 26789 | 524 |     opt: '(' ('no\_asm' | 'no\_asm\_simp' | 'no\_asm\_use') ')'
 | 
| 26782 | 525 | ; | 
| 526 |   \end{rail}
 | |
| 527 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 528 |   \begin{description}
 | 
| 26782 | 529 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 530 |   \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 | 531 |   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 | 532 | 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 | 533 | The result is fully simplified by default, including assumptions and | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 534 |   conclusion; the options @{text no_asm} etc.\ tune the Simplifier in
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 535 |   the same way as the for the @{text simp} method.
 | 
| 26782 | 536 | |
| 537 | Note that forward simplification restricts the simplifier to its | |
| 538 | most basic operation of term rewriting; solver and looper tactics | |
| 539 |   \cite{isabelle-ref} are \emph{not} involved here.  The @{text
 | |
| 540 | simplified} attribute should be only rarely required under normal | |
| 541 | circumstances. | |
| 542 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 543 |   \end{description}
 | 
| 26782 | 544 | *} | 
| 545 | ||
| 546 | ||
| 27040 | 547 | section {* The Classical Reasoner \label{sec:classical} *}
 | 
| 26782 | 548 | |
| 27040 | 549 | subsection {* Basic methods *}
 | 
| 26782 | 550 | |
| 551 | text {*
 | |
| 552 |   \begin{matharray}{rcl}
 | |
| 28761 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 553 |     @{method_def rule} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 554 |     @{method_def contradiction} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 555 |     @{method_def intro} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 556 |     @{method_def elim} & : & @{text method} \\
 | 
| 26782 | 557 |   \end{matharray}
 | 
| 558 | ||
| 559 |   \begin{rail}
 | |
| 560 |     ('rule' | 'intro' | 'elim') thmrefs?
 | |
| 561 | ; | |
| 562 |   \end{rail}
 | |
| 563 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 564 |   \begin{description}
 | 
| 26782 | 565 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 566 |   \item @{method rule} as offered by the Classical Reasoner is a
 | 
| 26782 | 567 |   refinement over the primitive one (see \secref{sec:pure-meth-att}).
 | 
| 568 | Both versions essentially work the same, but the classical version | |
| 569 | observes the classical rule context in addition to that of | |
| 570 | Isabelle/Pure. | |
| 571 | ||
| 572 | Common object logics (HOL, ZF, etc.) declare a rich collection of | |
| 573 | classical rules (even if these would qualify as intuitionistic | |
| 574 | ones), but only few declarations to the rule context of | |
| 575 |   Isabelle/Pure (\secref{sec:pure-meth-att}).
 | |
| 576 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 577 |   \item @{method contradiction} solves some goal by contradiction,
 | 
| 26782 | 578 |   deriving any result from both @{text "\<not> A"} and @{text A}.  Chained
 | 
| 579 | facts, which are guaranteed to participate, may appear in either | |
| 580 | order. | |
| 581 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 582 |   \item @{method intro} and @{method elim} repeatedly refine some goal
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 583 | by intro- or elim-resolution, after having inserted any chained | 
| 26901 | 584 | facts. Exactly the rules given as arguments are taken into account; | 
| 585 | this allows fine-tuned decomposition of a proof problem, in contrast | |
| 586 | to common automated tools. | |
| 26782 | 587 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 588 |   \end{description}
 | 
| 26782 | 589 | *} | 
| 590 | ||
| 591 | ||
| 27040 | 592 | subsection {* Automated methods *}
 | 
| 26782 | 593 | |
| 594 | text {*
 | |
| 595 |   \begin{matharray}{rcl}
 | |
| 28761 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 596 |     @{method_def blast} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 597 |     @{method_def fast} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 598 |     @{method_def slow} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 599 |     @{method_def best} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 600 |     @{method_def safe} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 601 |     @{method_def clarify} & : & @{text method} \\
 | 
| 26782 | 602 |   \end{matharray}
 | 
| 603 | ||
| 604 |   \indexouternonterm{clamod}
 | |
| 605 |   \begin{rail}
 | |
| 606 |     'blast' ('!' ?) nat? (clamod *)
 | |
| 607 | ; | |
| 608 |     ('fast' | 'slow' | 'best' | 'safe' | 'clarify') ('!' ?) (clamod *)
 | |
| 609 | ; | |
| 610 | ||
| 611 |     clamod: (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del') ':' thmrefs
 | |
| 612 | ; | |
| 613 |   \end{rail}
 | |
| 614 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 615 |   \begin{description}
 | 
| 26782 | 616 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 617 |   \item @{method blast} refers to the classical tableau prover (see
 | 
| 30397 | 618 |   @{ML blast_tac} in \cite{isabelle-ref}).  The optional argument
 | 
| 619 | specifies a user-supplied search bound (default 20). | |
| 26782 | 620 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 621 |   \item @{method fast}, @{method slow}, @{method best}, @{method
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 622 |   safe}, and @{method clarify} refer to the generic classical
 | 
| 26782 | 623 |   reasoner.  See @{ML fast_tac}, @{ML slow_tac}, @{ML best_tac}, @{ML
 | 
| 30397 | 624 |   safe_tac}, and @{ML clarify_tac} in \cite{isabelle-ref} for more
 | 
| 625 | information. | |
| 26782 | 626 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 627 |   \end{description}
 | 
| 26782 | 628 | |
| 629 | Any of the above methods support additional modifiers of the context | |
| 630 | of classical rules. Their semantics is analogous to the attributes | |
| 631 | given before. Facts provided by forward chaining are inserted into | |
| 632 |   the goal before commencing proof search.  The ``@{text
 | |
| 633 | "!"}''~argument causes the full context of assumptions to be | |
| 634 | included as well. | |
| 635 | *} | |
| 636 | ||
| 637 | ||
| 27040 | 638 | subsection {* Combined automated methods \label{sec:clasimp} *}
 | 
| 26782 | 639 | |
| 640 | text {*
 | |
| 641 |   \begin{matharray}{rcl}
 | |
| 28761 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 642 |     @{method_def auto} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 643 |     @{method_def force} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 644 |     @{method_def clarsimp} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 645 |     @{method_def fastsimp} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 646 |     @{method_def slowsimp} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 647 |     @{method_def bestsimp} & : & @{text method} \\
 | 
| 26782 | 648 |   \end{matharray}
 | 
| 649 | ||
| 650 |   \indexouternonterm{clasimpmod}
 | |
| 651 |   \begin{rail}
 | |
| 652 | 'auto' '!'? (nat nat)? (clasimpmod *) | |
| 653 | ; | |
| 654 |     ('force' | 'clarsimp' | 'fastsimp' | 'slowsimp' | 'bestsimp') '!'? (clasimpmod *)
 | |
| 655 | ; | |
| 656 | ||
| 657 |     clasimpmod: ('simp' (() | 'add' | 'del' | 'only') |
 | |
| 658 |       ('cong' | 'split') (() | 'add' | 'del') |
 | |
| 659 | 'iff' (((() | 'add') '?'?) | 'del') | | |
| 660 |       (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del')) ':' thmrefs
 | |
| 661 |   \end{rail}
 | |
| 662 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 663 |   \begin{description}
 | 
| 26782 | 664 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 665 |   \item @{method auto}, @{method force}, @{method clarsimp}, @{method
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 666 |   fastsimp}, @{method slowsimp}, and @{method bestsimp} provide access
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 667 | to Isabelle's combined simplification and classical reasoning | 
| 26782 | 668 |   tactics.  These correspond to @{ML auto_tac}, @{ML force_tac}, @{ML
 | 
| 669 | clarsimp_tac}, and Classical Reasoner tactics with the Simplifier | |
| 30397 | 670 |   added as wrapper, see \cite{isabelle-ref} for more information.  The
 | 
| 671 | modifier arguments correspond to those given in | |
| 26782 | 672 |   \secref{sec:simplifier} and \secref{sec:classical}.  Just note that
 | 
| 673 |   the ones related to the Simplifier are prefixed by \railtterm{simp}
 | |
| 674 | here. | |
| 675 | ||
| 676 | Facts provided by forward chaining are inserted into the goal before | |
| 677 |   doing the search.  The ``@{text "!"}'' argument causes the full
 | |
| 678 | context of assumptions to be included as well. | |
| 679 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 680 |   \end{description}
 | 
| 26782 | 681 | *} | 
| 682 | ||
| 683 | ||
| 27040 | 684 | subsection {* Declaring rules *}
 | 
| 26782 | 685 | |
| 686 | text {*
 | |
| 687 |   \begin{matharray}{rcl}
 | |
| 28761 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 688 |     @{command_def "print_claset"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 689 |     @{attribute_def intro} & : & @{text attribute} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 690 |     @{attribute_def elim} & : & @{text attribute} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 691 |     @{attribute_def dest} & : & @{text attribute} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 692 |     @{attribute_def rule} & : & @{text attribute} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 693 |     @{attribute_def iff} & : & @{text attribute} \\
 | 
| 26782 | 694 |   \end{matharray}
 | 
| 695 | ||
| 696 |   \begin{rail}
 | |
| 697 |     ('intro' | 'elim' | 'dest') ('!' | () | '?') nat?
 | |
| 698 | ; | |
| 699 | 'rule' 'del' | |
| 700 | ; | |
| 701 | 'iff' (((() | 'add') '?'?) | 'del') | |
| 702 | ; | |
| 703 |   \end{rail}
 | |
| 704 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 705 |   \begin{description}
 | 
| 26782 | 706 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 707 |   \item @{command "print_claset"} prints the collection of rules
 | 
| 26782 | 708 | declared to the Classical Reasoner, which is also known as | 
| 709 |   ``claset'' internally \cite{isabelle-ref}.
 | |
| 710 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 711 |   \item @{attribute intro}, @{attribute elim}, and @{attribute dest}
 | 
| 26782 | 712 | declare introduction, elimination, and destruction rules, | 
| 713 |   respectively.  By default, rules are considered as \emph{unsafe}
 | |
| 714 |   (i.e.\ not applied blindly without backtracking), while ``@{text
 | |
| 715 |   "!"}'' classifies as \emph{safe}.  Rule declarations marked by
 | |
| 716 |   ``@{text "?"}'' coincide with those of Isabelle/Pure, cf.\
 | |
| 717 |   \secref{sec:pure-meth-att} (i.e.\ are only applied in single steps
 | |
| 718 |   of the @{method rule} method).  The optional natural number
 | |
| 719 | specifies an explicit weight argument, which is ignored by automated | |
| 720 | tools, but determines the search order of single rule steps. | |
| 721 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 722 |   \item @{attribute rule}~@{text del} deletes introduction,
 | 
| 26782 | 723 | elimination, or destruction rules from the context. | 
| 724 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 725 |   \item @{attribute iff} declares logical equivalences to the
 | 
| 26782 | 726 | Simplifier and the Classical reasoner at the same time. | 
| 727 | Non-conditional rules result in a ``safe'' introduction and | |
| 728 | elimination pair; conditional ones are considered ``unsafe''. Rules | |
| 729 |   with negative conclusion are automatically inverted (using @{text
 | |
| 26789 | 730 | "\<not>"}-elimination internally). | 
| 26782 | 731 | |
| 732 |   The ``@{text "?"}'' version of @{attribute iff} declares rules to
 | |
| 733 | the Isabelle/Pure context only, and omits the Simplifier | |
| 734 | declaration. | |
| 735 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 736 |   \end{description}
 | 
| 26782 | 737 | *} | 
| 738 | ||
| 739 | ||
| 27040 | 740 | subsection {* Classical operations *}
 | 
| 26782 | 741 | |
| 742 | text {*
 | |
| 743 |   \begin{matharray}{rcl}
 | |
| 28761 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 744 |     @{attribute_def swapped} & : & @{text attribute} \\
 | 
| 26782 | 745 |   \end{matharray}
 | 
| 746 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 747 |   \begin{description}
 | 
| 26782 | 748 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 749 |   \item @{attribute swapped} turns an introduction rule into an
 | 
| 26782 | 750 |   elimination, by resolving with the classical swap principle @{text
 | 
| 751 | "(\<not> B \<Longrightarrow> A) \<Longrightarrow> (\<not> A \<Longrightarrow> B)"}. | |
| 752 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 753 |   \end{description}
 | 
| 26782 | 754 | *} | 
| 755 | ||
| 756 | ||
| 27044 | 757 | section {* Object-logic setup \label{sec:object-logic} *}
 | 
| 26790 | 758 | |
| 759 | text {*
 | |
| 760 |   \begin{matharray}{rcl}
 | |
| 28761 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 761 |     @{command_def "judgment"} & : & @{text "theory \<rightarrow> theory"} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 762 |     @{method_def atomize} & : & @{text method} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 763 |     @{attribute_def atomize} & : & @{text attribute} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 764 |     @{attribute_def rule_format} & : & @{text attribute} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 765 |     @{attribute_def rulify} & : & @{text attribute} \\
 | 
| 26790 | 766 |   \end{matharray}
 | 
| 767 | ||
| 768 | The very starting point for any Isabelle object-logic is a ``truth | |
| 769 | judgment'' that links object-level statements to the meta-logic | |
| 770 |   (with its minimal language of @{text prop} that covers universal
 | |
| 771 |   quantification @{text "\<And>"} and implication @{text "\<Longrightarrow>"}).
 | |
| 772 | ||
| 773 | Common object-logics are sufficiently expressive to internalize rule | |
| 774 |   statements over @{text "\<And>"} and @{text "\<Longrightarrow>"} within their own
 | |
| 775 | language. This is useful in certain situations where a rule needs | |
| 776 | to be viewed as an atomic statement from the meta-level perspective, | |
| 777 |   e.g.\ @{text "\<And>x. x \<in> A \<Longrightarrow> P x"} versus @{text "\<forall>x \<in> A. P x"}.
 | |
| 778 | ||
| 779 |   From the following language elements, only the @{method atomize}
 | |
| 780 |   method and @{attribute rule_format} attribute are occasionally
 | |
| 781 | required by end-users, the rest is for those who need to setup their | |
| 782 | own object-logic. In the latter case existing formulations of | |
| 783 | Isabelle/FOL or Isabelle/HOL may be taken as realistic examples. | |
| 784 | ||
| 785 | Generic tools may refer to the information provided by object-logic | |
| 786 | declarations internally. | |
| 787 | ||
| 788 |   \begin{rail}
 | |
| 789 | 'judgment' constdecl | |
| 790 | ; | |
| 791 |     'atomize' ('(' 'full' ')')?
 | |
| 792 | ; | |
| 793 |     'rule\_format' ('(' 'noasm' ')')?
 | |
| 794 | ; | |
| 795 |   \end{rail}
 | |
| 796 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 797 |   \begin{description}
 | 
| 26790 | 798 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 799 |   \item @{command "judgment"}~@{text "c :: \<sigma> (mx)"} declares constant
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 800 |   @{text c} as the truth judgment of the current object-logic.  Its
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 801 |   type @{text \<sigma>} should specify a coercion of the category of
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 802 |   object-level propositions to @{text prop} of the Pure meta-logic;
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 803 |   the mixfix annotation @{text "(mx)"} would typically just link the
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 804 |   object language (internally of syntactic category @{text logic})
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 805 |   with that of @{text prop}.  Only one @{command "judgment"}
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 806 | declaration may be given in any theory development. | 
| 26790 | 807 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 808 |   \item @{method atomize} (as a method) rewrites any non-atomic
 | 
| 26790 | 809 | premises of a sub-goal, using the meta-level equations declared via | 
| 810 |   @{attribute atomize} (as an attribute) beforehand.  As a result,
 | |
| 811 | heavily nested goals become amenable to fundamental operations such | |
| 812 |   as resolution (cf.\ the @{method rule} method).  Giving the ``@{text
 | |
| 813 | "(full)"}'' option here means to turn the whole subgoal into an | |
| 814 | object-statement (if possible), including the outermost parameters | |
| 815 | and assumptions as well. | |
| 816 | ||
| 817 |   A typical collection of @{attribute atomize} rules for a particular
 | |
| 818 | object-logic would provide an internalization for each of the | |
| 819 |   connectives of @{text "\<And>"}, @{text "\<Longrightarrow>"}, and @{text "\<equiv>"}.
 | |
| 820 | Meta-level conjunction should be covered as well (this is | |
| 821 |   particularly important for locales, see \secref{sec:locale}).
 | |
| 822 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 823 |   \item @{attribute rule_format} rewrites a theorem by the equalities
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 824 |   declared as @{attribute rulify} rules in the current object-logic.
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 825 | By default, the result is fully normalized, including assumptions | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 826 |   and conclusions at any depth.  The @{text "(no_asm)"} option
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 827 | restricts the transformation to the conclusion of a rule. | 
| 26790 | 828 | |
| 829 |   In common object-logics (HOL, FOL, ZF), the effect of @{attribute
 | |
| 830 | rule_format} is to replace (bounded) universal quantification | |
| 831 |   (@{text "\<forall>"}) and implication (@{text "\<longrightarrow>"}) by the corresponding
 | |
| 832 |   rule statements over @{text "\<And>"} and @{text "\<Longrightarrow>"}.
 | |
| 833 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28754diff
changeset | 834 |   \end{description}
 | 
| 26790 | 835 | *} | 
| 836 | ||
| 26782 | 837 | end |