| author | wenzelm | 
| Thu, 02 Mar 2023 16:09:22 +0100 | |
| changeset 77483 | 291f5848bf55 | 
| parent 76987 | 4c275405faae | 
| child 79501 | bce98b5dfec6 | 
| permissions | -rw-r--r-- | 
| 61656 | 1 | (*:maxLineLen=78:*) | 
| 2 | ||
| 29755 | 3 | theory Isar | 
| 4 | imports Base | |
| 5 | begin | |
| 20472 | 6 | |
| 58618 | 7 | chapter \<open>Isar language elements\<close> | 
| 29759 | 8 | |
| 61854 | 9 | text \<open> | 
| 76987 | 10 | The Isar proof language (see also \<^cite>\<open>\<open>\S2\<close> in "isabelle-isar-ref"\<close>) | 
| 61854 | 11 | consists of three main categories of language elements: | 
| 29759 | 12 | |
| 61854 | 13 | \<^enum> Proof \<^emph>\<open>commands\<close> define the primary language of transactions of the | 
| 14 |   underlying Isar/VM interpreter. Typical examples are @{command "fix"},
 | |
| 15 |   @{command "assume"}, @{command "show"}, @{command "proof"}, and @{command
 | |
| 16 | "qed"}. | |
| 39842 | 17 | |
| 61854 | 18 | Composing proof commands according to the rules of the Isar/VM leads to | 
| 19 | expressions of structured proof text, such that both the machine and the | |
| 20 | human reader can give it a meaning as formal reasoning. | |
| 20472 | 21 | |
| 61854 | 22 | \<^enum> Proof \<^emph>\<open>methods\<close> define a secondary language of mixed forward-backward | 
| 23 |   refinement steps involving facts and goals. Typical examples are @{method
 | |
| 24 |   rule}, @{method unfold}, and @{method simp}.
 | |
| 29759 | 25 | |
| 61854 | 26 | Methods can occur in certain well-defined parts of the Isar proof language, | 
| 27 |   say as arguments to @{command "proof"}, @{command "qed"}, or @{command
 | |
| 28 | "by"}. | |
| 39842 | 29 | |
| 61854 | 30 | \<^enum> \<^emph>\<open>Attributes\<close> define a tertiary language of small annotations to theorems | 
| 31 | being defined or referenced. Attributes can modify both the context and the | |
| 32 | theorem. | |
| 39842 | 33 | |
| 61854 | 34 |   Typical examples are @{attribute intro} (which affects the context), and
 | 
| 35 |   @{attribute symmetric} (which affects the theorem).
 | |
| 58618 | 36 | \<close> | 
| 29759 | 37 | |
| 38 | ||
| 58618 | 39 | section \<open>Proof commands\<close> | 
| 20520 | 40 | |
| 61854 | 41 | text \<open> | 
| 42 | A \<^emph>\<open>proof command\<close> is state transition of the Isar/VM proof interpreter. | |
| 39849 | 43 | |
| 61854 | 44 | In principle, Isar proof commands could be defined in user-space as well. | 
| 45 | The system is built like that in the first place: one part of the commands | |
| 46 | are primitive, the other part is defined as derived elements. Adding to the | |
| 47 | genuine structured proof language requires profound understanding of the | |
| 48 | Isar/VM machinery, though, so this is beyond the scope of this manual. | |
| 39842 | 49 | |
| 61854 | 50 | What can be done realistically is to define some diagnostic commands that | 
| 51 | inspect the general state of the Isar/VM, and report some feedback to the | |
| 52 | user. Typically this involves checking of the linguistic \<^emph>\<open>mode\<close> of a proof | |
| 53 | state, or peeking at the pending goals (if available). | |
| 39845 | 54 | |
| 61854 | 55 | Another common application is to define a toplevel command that poses a | 
| 56 | problem to the user as Isar proof state and processes the final result | |
| 57 | relatively to the context. Thus a proof can be incorporated into the context | |
| 58 | of some user-space tool, without modifying the Isar proof language itself. | |
| 59 | \<close> | |
| 39842 | 60 | |
| 58618 | 61 | text %mlref \<open> | 
| 39842 | 62 |   \begin{mldecls}
 | 
| 73764 | 63 |   @{define_ML_type Proof.state} \\
 | 
| 64 |   @{define_ML Proof.assert_forward: "Proof.state -> Proof.state"} \\
 | |
| 65 |   @{define_ML Proof.assert_chain: "Proof.state -> Proof.state"} \\
 | |
| 66 |   @{define_ML Proof.assert_backward: "Proof.state -> Proof.state"} \\
 | |
| 67 |   @{define_ML Proof.simple_goal: "Proof.state -> {context: Proof.context, goal: thm}"} \\
 | |
| 68 |   @{define_ML Proof.goal: "Proof.state ->
 | |
| 39842 | 69 |   {context: Proof.context, facts: thm list, goal: thm}"} \\
 | 
| 73764 | 70 |   @{define_ML Proof.raw_goal: "Proof.state ->
 | 
| 39842 | 71 |   {context: Proof.context, facts: thm list, goal: thm}"} \\
 | 
| 73764 | 72 |   @{define_ML Proof.theorem: "Method.text option ->
 | 
| 39845 | 73 | (thm list list -> Proof.context -> Proof.context) -> | 
| 74 | (term * term list) list list -> Proof.context -> Proof.state"} \\ | |
| 39842 | 75 |   \end{mldecls}
 | 
| 76 | ||
| 69597 | 77 | \<^descr> Type \<^ML_type>\<open>Proof.state\<close> represents Isar proof states. This is a | 
| 61854 | 78 | block-structured configuration with proof context, linguistic mode, and | 
| 79 | optional goal. The latter consists of goal context, goal facts | |
| 80 |   (``\<open>using\<close>''), and tactical goal state (see \secref{sec:tactical-goals}).
 | |
| 39842 | 81 | |
| 61854 | 82 | The general idea is that the facts shall contribute to the refinement of | 
| 83 | some parts of the tactical goal --- how exactly is defined by the proof | |
| 84 | method that is applied in that situation. | |
| 39842 | 85 | |
| 69597 | 86 | \<^descr> \<^ML>\<open>Proof.assert_forward\<close>, \<^ML>\<open>Proof.assert_chain\<close>, \<^ML>\<open>Proof.assert_backward\<close> are partial identity functions that fail unless a | 
| 61854 | 87 | certain linguistic mode is active, namely ``\<open>proof(state)\<close>'', | 
| 88 | ``\<open>proof(chain)\<close>'', ``\<open>proof(prove)\<close>'', respectively (using the terminology | |
| 76987 | 89 | of \<^cite>\<open>"isabelle-isar-ref"\<close>). | 
| 39842 | 90 | |
| 61854 | 91 | It is advisable study the implementations of existing proof commands for | 
| 92 | suitable modes to be asserted. | |
| 39842 | 93 | |
| 69597 | 94 | \<^descr> \<^ML>\<open>Proof.simple_goal\<close>~\<open>state\<close> returns the structured Isar goal (if | 
| 61854 | 95 |   available) in the form seen by ``simple'' methods (like @{method simp} or
 | 
| 96 |   @{method blast}). The Isar goal facts are already inserted as premises into
 | |
| 69597 | 97 | the subgoals, which are presented individually as in \<^ML>\<open>Proof.goal\<close>. | 
| 39842 | 98 | |
| 69597 | 99 | \<^descr> \<^ML>\<open>Proof.goal\<close>~\<open>state\<close> returns the structured Isar goal (if available) | 
| 61854 | 100 |   in the form seen by regular methods (like @{method rule}). The auxiliary
 | 
| 101 | internal encoding of Pure conjunctions is split into individual subgoals as | |
| 102 | usual. | |
| 39842 | 103 | |
| 69597 | 104 | \<^descr> \<^ML>\<open>Proof.raw_goal\<close>~\<open>state\<close> returns the structured Isar goal (if | 
| 61854 | 105 |   available) in the raw internal form seen by ``raw'' methods (like @{method
 | 
| 69597 | 106 | induct}). This form is rarely appropriate for diagnostic tools; \<^ML>\<open>Proof.simple_goal\<close> or \<^ML>\<open>Proof.goal\<close> should be used in most situations. | 
| 39842 | 107 | |
| 69597 | 108 | \<^descr> \<^ML>\<open>Proof.theorem\<close>~\<open>before_qed after_qed statement ctxt\<close> initializes a | 
| 61854 | 109 | toplevel Isar proof state within a given context. | 
| 39845 | 110 | |
| 61854 | 111 | The optional \<open>before_qed\<close> method is applied at the end of the proof, just | 
| 112 | before extracting the result (this feature is rarely used). | |
| 39845 | 113 | |
| 61854 | 114 | The \<open>after_qed\<close> continuation receives the extracted result in order to apply | 
| 115 | it to the final context in a suitable way (e.g.\ storing named facts). Note | |
| 69597 | 116 | that at this generic level the target context is specified as \<^ML_type>\<open>Proof.context\<close>, but the usual wrapping of toplevel proofs into command | 
| 117 | transactions will provide a \<^ML_type>\<open>local_theory\<close> here | |
| 61854 | 118 |   (\chref{ch:local-theory}). This affects the way how results are stored.
 | 
| 39845 | 119 | |
| 61854 | 120 | The \<open>statement\<close> is given as a nested list of terms, each associated with | 
| 121 |   optional @{keyword "is"} patterns as usual in the Isar source language. The
 | |
| 122 | original nested list structure over terms is turned into one over theorems | |
| 123 | when \<open>after_qed\<close> is invoked. | |
| 58618 | 124 | \<close> | 
| 39842 | 125 | |
| 20520 | 126 | |
| 58618 | 127 | text %mlantiq \<open> | 
| 39843 | 128 |   \begin{matharray}{rcl}
 | 
| 61493 | 129 |   @{ML_antiquotation_def "Isar.goal"} & : & \<open>ML_antiquotation\<close> \\
 | 
| 39843 | 130 |   \end{matharray}
 | 
| 131 | ||
| 61854 | 132 |   \<^descr> \<open>@{Isar.goal}\<close> refers to the regular goal state (if available) of the
 | 
| 133 | current proof state managed by the Isar toplevel --- as abstract value. | |
| 39843 | 134 | |
| 61854 | 135 |   This only works for diagnostic ML commands, such as @{command ML_val} or
 | 
| 136 |   @{command ML_command}.
 | |
| 58618 | 137 | \<close> | 
| 39843 | 138 | |
| 61854 | 139 | text %mlex \<open> | 
| 140 | The following example peeks at a certain goal configuration. | |
| 141 | \<close> | |
| 39843 | 142 | |
| 40964 | 143 | notepad | 
| 144 | begin | |
| 39846 | 145 | have A and B and C | 
| 58728 | 146 | ML_val | 
| 147 |      \<open>val n = Thm.nprems_of (#goal @{Isar.goal});
 | |
| 69597 | 148 | \<^assert> (n = 3);\<close> | 
| 58801 | 149 | sorry | 
| 150 | end | |
| 39843 | 151 | |
| 61854 | 152 | text \<open> | 
| 153 | Here we see 3 individual subgoals in the same way as regular proof methods | |
| 154 | would do. | |
| 155 | \<close> | |
| 39843 | 156 | |
| 20520 | 157 | |
| 58618 | 158 | section \<open>Proof methods\<close> | 
| 20472 | 159 | |
| 61841 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
61656diff
changeset | 160 | text \<open> | 
| 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
61656diff
changeset | 161 | A \<open>method\<close> is a function \<open>thm\<^sup>* \<rightarrow> context * goal \<rightarrow> (context \<times> goal)\<^sup>*\<^sup>*\<close> | 
| 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
61656diff
changeset | 162 | that operates on the full Isar goal configuration with context, goal facts, | 
| 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
61656diff
changeset | 163 | and tactical goal state and enumerates possible follow-up goal states. Under | 
| 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
61656diff
changeset | 164 | normal circumstances, the goal context remains unchanged, but it is also | 
| 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
61656diff
changeset | 165 | possible to declare named extensions of the proof context (\<^emph>\<open>cases\<close>). | 
| 39847 | 166 | |
| 61854 | 167 | This means a proof method is like a structurally enhanced tactic (cf.\ | 
| 168 |   \secref{sec:tactics}). The well-formedness conditions for tactics need to
 | |
| 169 | hold for methods accordingly, with the following additions. | |
| 39847 | 170 | |
| 61854 | 171 | \<^item> Goal addressing is further limited either to operate uniformly on \<^emph>\<open>all\<close> | 
| 172 | subgoals, or specifically on the \<^emph>\<open>first\<close> subgoal. | |
| 39847 | 173 | |
| 61854 | 174 | Exception: old-style tactic emulations that are embedded into the method | 
| 175 |   space, e.g.\ @{method rule_tac}.
 | |
| 39847 | 176 | |
| 61854 | 177 | \<^item> A non-trivial method always needs to make progress: an identical follow-up | 
| 178 | goal state has to be avoided.\<^footnote>\<open>This enables the user to write method | |
| 179 | expressions like \<open>meth\<^sup>+\<close> without looping, while the trivial do-nothing case | |
| 180 | can be recovered via \<open>meth\<^sup>?\<close>.\<close> | |
| 39847 | 181 | |
| 61854 | 182 |   Exception: trivial stuttering steps, such as ``@{method -}'' or @{method
 | 
| 183 | succeed}. | |
| 39847 | 184 | |
| 61854 | 185 | \<^item> Goal facts passed to the method must not be ignored. If there is no | 
| 186 | sensible use of facts outside the goal state, facts should be inserted into | |
| 187 | the subgoals that are addressed by the method. | |
| 39847 | 188 | |
| 189 | ||
| 61416 | 190 | \<^medskip> | 
| 61854 | 191 | Syntactically, the language of proof methods appears as arguments to Isar | 
| 192 |   commands like @{command "by"} or @{command apply}. User-space additions are
 | |
| 193 | reasonably easy by plugging suitable method-valued parser functions into the | |
| 194 |   framework, using the @{command method_setup} command, for example.
 | |
| 39843 | 195 | |
| 39844 | 196 | To get a better idea about the range of possibilities, consider the | 
| 61854 | 197 | following Isar proof schemes. This is the general form of structured proof | 
| 198 | text: | |
| 39843 | 199 | |
| 61416 | 200 | \<^medskip> | 
| 39843 | 201 |   \begin{tabular}{l}
 | 
| 61493 | 202 |   @{command from}~\<open>facts\<^sub>1\<close>~@{command have}~\<open>props\<close>~@{command using}~\<open>facts\<^sub>2\<close> \\
 | 
| 203 |   @{command proof}~\<open>(initial_method)\<close> \\
 | |
| 204 | \quad\<open>body\<close> \\ | |
| 205 |   @{command qed}~\<open>(terminal_method)\<close> \\
 | |
| 39843 | 206 |   \end{tabular}
 | 
| 61416 | 207 | \<^medskip> | 
| 39843 | 208 | |
| 61854 | 209 | The goal configuration consists of \<open>facts\<^sub>1\<close> and \<open>facts\<^sub>2\<close> appended in that | 
| 210 | order, and various \<open>props\<close> being claimed. The \<open>initial_method\<close> is invoked | |
| 211 | with facts and goals together and refines the problem to something that is | |
| 212 | handled recursively in the proof \<open>body\<close>. The \<open>terminal_method\<close> has another | |
| 213 | chance to finish any remaining subgoals, but it does not see the facts of | |
| 214 | the initial step. | |
| 39843 | 215 | |
| 61416 | 216 | \<^medskip> | 
| 217 | This pattern illustrates unstructured proof scripts: | |
| 39843 | 218 | |
| 61416 | 219 | \<^medskip> | 
| 39843 | 220 |   \begin{tabular}{l}
 | 
| 61493 | 221 |   @{command have}~\<open>props\<close> \\
 | 
| 222 |   \quad@{command using}~\<open>facts\<^sub>1\<close>~@{command apply}~\<open>method\<^sub>1\<close> \\
 | |
| 223 |   \quad@{command apply}~\<open>method\<^sub>2\<close> \\
 | |
| 224 |   \quad@{command using}~\<open>facts\<^sub>3\<close>~@{command apply}~\<open>method\<^sub>3\<close> \\
 | |
| 39843 | 225 |   \quad@{command done} \\
 | 
| 226 |   \end{tabular}
 | |
| 61416 | 227 | \<^medskip> | 
| 39843 | 228 | |
| 61854 | 229 | The \<open>method\<^sub>1\<close> operates on the original claim while using \<open>facts\<^sub>1\<close>. Since | 
| 230 |   the @{command apply} command structurally resets the facts, the \<open>method\<^sub>2\<close>
 | |
| 231 | will operate on the remaining goal state without facts. The \<open>method\<^sub>3\<close> will | |
| 232 | see again a collection of \<open>facts\<^sub>3\<close> that has been inserted into the script | |
| 233 | explicitly. | |
| 39843 | 234 | |
| 61416 | 235 | \<^medskip> | 
| 61854 | 236 | Empirically, any Isar proof method can be categorized as follows. | 
| 39843 | 237 | |
| 61854 | 238 | \<^enum> \<^emph>\<open>Special method with cases\<close> with named context additions associated | 
| 239 | with the follow-up goal state. | |
| 39843 | 240 | |
| 61854 | 241 |     Example: @{method "induct"}, which is also a ``raw'' method since it
 | 
| 242 | operates on the internal representation of simultaneous claims as Pure | |
| 243 |     conjunction (\secref{sec:logic-aux}), instead of separate subgoals
 | |
| 244 |     (\secref{sec:tactical-goals}).
 | |
| 39843 | 245 | |
| 61854 | 246 | \<^enum> \<^emph>\<open>Structured method\<close> with strong emphasis on facts outside the goal | 
| 247 | state. | |
| 39847 | 248 | |
| 61854 | 249 |     Example: @{method "rule"}, which captures the key ideas behind structured
 | 
| 250 | reasoning in Isar in its purest form. | |
| 39843 | 251 | |
| 61854 | 252 | \<^enum> \<^emph>\<open>Simple method\<close> with weaker emphasis on facts, which are inserted into | 
| 253 | subgoals to emulate old-style tactical ``premises''. | |
| 39843 | 254 | |
| 61854 | 255 |     Examples: @{method "simp"}, @{method "blast"}, @{method "auto"}.
 | 
| 39843 | 256 | |
| 61854 | 257 | \<^enum> \<^emph>\<open>Old-style tactic emulation\<close> with detailed numeric goal addressing and | 
| 258 | explicit references to entities of the internal goal state (which are | |
| 259 | otherwise invisible from proper Isar proof text). The naming convention | |
| 260 | \<open>foo_tac\<close> makes this special non-standard status clear. | |
| 39843 | 261 | |
| 61854 | 262 |     Example: @{method "rule_tac"}.
 | 
| 39843 | 263 | |
| 39847 | 264 | When implementing proof methods, it is advisable to study existing | 
| 61854 | 265 | implementations carefully and imitate the typical ``boiler plate'' for | 
| 266 | context-sensitive parsing and further combinators to wrap-up tactic | |
| 267 | expressions as methods.\<^footnote>\<open>Aliases or abbreviations of the standard method | |
| 268 | combinators should be avoided. Note that from Isabelle99 until Isabelle2009 | |
| 269 | the system did provide various odd combinations of method syntax wrappers | |
| 270 | that made applications more complicated than necessary.\<close> | |
| 58618 | 271 | \<close> | 
| 39847 | 272 | |
| 58618 | 273 | text %mlref \<open> | 
| 39847 | 274 |   \begin{mldecls}
 | 
| 73764 | 275 |   @{define_ML_type Proof.method} \\
 | 
| 276 |   @{define_ML CONTEXT_METHOD: "(thm list -> context_tactic) -> Proof.method"} \\
 | |
| 277 |   @{define_ML METHOD: "(thm list -> tactic) -> Proof.method"} \\
 | |
| 278 |   @{define_ML SIMPLE_METHOD: "tactic -> Proof.method"} \\
 | |
| 279 |   @{define_ML SIMPLE_METHOD': "(int -> tactic) -> Proof.method"} \\
 | |
| 280 |   @{define_ML Method.insert_tac: "Proof.context -> thm list -> int -> tactic"} \\
 | |
| 281 |   @{define_ML Method.setup: "binding -> (Proof.context -> Proof.method) context_parser ->
 | |
| 39847 | 282 | string -> theory -> theory"} \\ | 
| 283 |   \end{mldecls}
 | |
| 284 | ||
| 69597 | 285 | \<^descr> Type \<^ML_type>\<open>Proof.method\<close> represents proof methods as abstract type. | 
| 39847 | 286 | |
| 69597 | 287 | \<^descr> \<^ML>\<open>CONTEXT_METHOD\<close>~\<open>(fn facts => context_tactic)\<close> wraps \<open>context_tactic\<close> | 
| 61841 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
61656diff
changeset | 288 | depending on goal facts as a general proof method that may change the proof | 
| 73764 | 289 |   context dynamically. A typical operation is \<^ML>\<open>Proof_Context.update_cases\<close>, which is wrapped up as combinator @{define_ML
 | 
| 61841 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
61656diff
changeset | 290 | CONTEXT_CASES} for convenience. | 
| 39847 | 291 | |
| 69597 | 292 | \<^descr> \<^ML>\<open>METHOD\<close>~\<open>(fn facts => tactic)\<close> wraps \<open>tactic\<close> depending on goal facts | 
| 61854 | 293 | as regular proof method; the goal context is passed via method syntax. | 
| 39847 | 294 | |
| 69597 | 295 | \<^descr> \<^ML>\<open>SIMPLE_METHOD\<close>~\<open>tactic\<close> wraps a tactic that addresses all subgoals | 
| 61854 | 296 | uniformly as simple proof method. Goal facts are already inserted into all | 
| 297 | subgoals before \<open>tactic\<close> is applied. | |
| 39847 | 298 | |
| 69597 | 299 | \<^descr> \<^ML>\<open>SIMPLE_METHOD'\<close>~\<open>tactic\<close> wraps a tactic that addresses a specific | 
| 61854 | 300 | subgoal as simple proof method that operates on subgoal 1. Goal facts are | 
| 301 | inserted into the subgoal then the \<open>tactic\<close> is applied. | |
| 39847 | 302 | |
| 69597 | 303 | \<^descr> \<^ML>\<open>Method.insert_tac\<close>~\<open>ctxt facts i\<close> inserts \<open>facts\<close> into subgoal \<open>i\<close>. | 
| 304 | This is convenient to reproduce part of the \<^ML>\<open>SIMPLE_METHOD\<close> or \<^ML>\<open>SIMPLE_METHOD'\<close> wrapping within regular \<^ML>\<open>METHOD\<close>, for example. | |
| 39847 | 305 | |
| 69597 | 306 | \<^descr> \<^ML>\<open>Method.setup\<close>~\<open>name parser description\<close> provides the functionality of | 
| 61854 | 307 |   the Isar command @{command method_setup} as ML function.
 | 
| 58618 | 308 | \<close> | 
| 39847 | 309 | |
| 61854 | 310 | text %mlex \<open> | 
| 76987 | 311 |   See also @{command method_setup} in \<^cite>\<open>"isabelle-isar-ref"\<close> which
 | 
| 61854 | 312 | includes some abstract examples. | 
| 39851 | 313 | |
| 61416 | 314 | \<^medskip> | 
| 61854 | 315 | The following toy examples illustrate how the goal facts and state are | 
| 316 |   passed to proof methods. The predefined proof method called ``@{method
 | |
| 69597 | 317 | tactic}'' wraps ML source of type \<^ML_type>\<open>tactic\<close> (abstracted over | 
| 318 | \<^ML_text>\<open>facts\<close>). This allows immediate experimentation without parsing of | |
| 61854 | 319 | concrete syntax. | 
| 320 | \<close> | |
| 20472 | 321 | |
| 40964 | 322 | notepad | 
| 323 | begin | |
| 57340 | 324 | fix A B :: bool | 
| 39847 | 325 | assume a: A and b: B | 
| 326 | ||
| 327 | have "A \<and> B" | |
| 69597 | 328 |     apply (tactic \<open>resolve_tac \<^context> @{thms conjI} 1\<close>)
 | 
| 329 | using a apply (tactic \<open>resolve_tac \<^context> facts 1\<close>) | |
| 330 | using b apply (tactic \<open>resolve_tac \<^context> facts 1\<close>) | |
| 39847 | 331 | done | 
| 332 | ||
| 333 | have "A \<and> B" | |
| 334 | using a and b | |
| 57340 | 335 |     ML_val \<open>@{Isar.goal}\<close>
 | 
| 69597 | 336 | apply (tactic \<open>Method.insert_tac \<^context> facts 1\<close>) | 
| 337 |     apply (tactic \<open>(resolve_tac \<^context> @{thms conjI} THEN_ALL_NEW assume_tac \<^context>) 1\<close>)
 | |
| 39847 | 338 | done | 
| 40964 | 339 | end | 
| 39847 | 340 | |
| 61416 | 341 | text \<open> | 
| 342 | \<^medskip> | |
| 61854 | 343 | The next example implements a method that simplifies the first subgoal by | 
| 344 | rewrite rules that are given as arguments. | |
| 345 | \<close> | |
| 39848 | 346 | |
| 57340 | 347 | method_setup my_simp = | 
| 348 | \<open>Attrib.thms >> (fn thms => fn ctxt => | |
| 39848 | 349 | SIMPLE_METHOD' (fn i => | 
| 350 | CHANGED (asm_full_simp_tac | |
| 57340 | 351 | (put_simpset HOL_basic_ss ctxt addsimps thms) i)))\<close> | 
| 352 | "rewrite subgoal by given rules" | |
| 39848 | 353 | |
| 61854 | 354 | text \<open> | 
| 355 |   The concrete syntax wrapping of @{command method_setup} always
 | |
| 356 | passes-through the proof context at the end of parsing, but it is not used | |
| 357 | in this example. | |
| 39848 | 358 | |
| 69597 | 359 | The \<^ML>\<open>Attrib.thms\<close> parser produces a list of theorems from the usual Isar | 
| 61854 | 360 |   syntax involving attribute expressions etc.\ (syntax category @{syntax
 | 
| 76987 | 361 | thms}) \<^cite>\<open>"isabelle-isar-ref"\<close>. The resulting \<^ML_text>\<open>thms\<close> are | 
| 69597 | 362 | added to \<^ML>\<open>HOL_basic_ss\<close> which already contains the basic Simplifier | 
| 61854 | 363 | setup for HOL. | 
| 39848 | 364 | |
| 69597 | 365 | The tactic \<^ML>\<open>asm_full_simp_tac\<close> is the one that is also used in method | 
| 366 |   @{method simp} by default. The extra wrapping by the \<^ML>\<open>CHANGED\<close> tactical
 | |
| 61854 | 367 | ensures progress of simplification: identical goal states are filtered out | 
| 368 | explicitly to make the raw tactic conform to standard Isar method behaviour. | |
| 39848 | 369 | |
| 61416 | 370 | \<^medskip> | 
| 371 |   Method @{method my_simp} can be used in Isar proofs like this:
 | |
| 58618 | 372 | \<close> | 
| 39847 | 373 | |
| 40964 | 374 | notepad | 
| 375 | begin | |
| 57340 | 376 | fix a b c :: 'a | 
| 39851 | 377 | assume a: "a = b" | 
| 378 | assume b: "b = c" | |
| 379 | have "a = c" by (my_simp a b) | |
| 40964 | 380 | end | 
| 39851 | 381 | |
| 61854 | 382 | text \<open> | 
| 383 | Here is a similar method that operates on all subgoals, instead of just the | |
| 384 | first one.\<close> | |
| 39851 | 385 | |
| 57340 | 386 | method_setup my_simp_all = | 
| 387 | \<open>Attrib.thms >> (fn thms => fn ctxt => | |
| 39851 | 388 | SIMPLE_METHOD | 
| 389 | (CHANGED | |
| 390 | (ALLGOALS (asm_full_simp_tac | |
| 57340 | 391 | (put_simpset HOL_basic_ss ctxt addsimps thms)))))\<close> | 
| 392 | "rewrite all subgoals by given rules" | |
| 39851 | 393 | |
| 40964 | 394 | notepad | 
| 395 | begin | |
| 57340 | 396 | fix a b c :: 'a | 
| 39851 | 397 | assume a: "a = b" | 
| 398 | assume b: "b = c" | |
| 399 | have "a = c" and "c = b" by (my_simp_all a b) | |
| 40964 | 400 | end | 
| 39848 | 401 | |
| 61416 | 402 | text \<open> | 
| 403 | \<^medskip> | |
| 61854 | 404 | Apart from explicit arguments, common proof methods typically work with a | 
| 405 | default configuration provided by the context. As a shortcut to rule | |
| 406 |   management we use a cheap solution via the @{command named_theorems} command
 | |
| 407 | to declare a dynamic fact in the context. | |
| 408 | \<close> | |
| 39848 | 409 | |
| 57946 
6a26aa5fa65e
updated documentation concerning 'named_theorems';
 wenzelm parents: 
57342diff
changeset | 410 | named_theorems my_simp | 
| 39847 | 411 | |
| 61854 | 412 | text \<open> | 
| 413 | The proof method is now defined as before, but we append the explicit | |
| 414 | arguments and the rules from the context. | |
| 415 | \<close> | |
| 39847 | 416 | |
| 57340 | 417 | method_setup my_simp' = | 
| 418 | \<open>Attrib.thms >> (fn thms => fn ctxt => | |
| 57946 
6a26aa5fa65e
updated documentation concerning 'named_theorems';
 wenzelm parents: 
57342diff
changeset | 419 | let | 
| 69597 | 420 | val my_simps = Named_Theorems.get ctxt \<^named_theorems>\<open>my_simp\<close> | 
| 57946 
6a26aa5fa65e
updated documentation concerning 'named_theorems';
 wenzelm parents: 
57342diff
changeset | 421 | in | 
| 
6a26aa5fa65e
updated documentation concerning 'named_theorems';
 wenzelm parents: 
57342diff
changeset | 422 | SIMPLE_METHOD' (fn i => | 
| 
6a26aa5fa65e
updated documentation concerning 'named_theorems';
 wenzelm parents: 
57342diff
changeset | 423 | CHANGED (asm_full_simp_tac | 
| 
6a26aa5fa65e
updated documentation concerning 'named_theorems';
 wenzelm parents: 
57342diff
changeset | 424 | (put_simpset HOL_basic_ss ctxt | 
| 
6a26aa5fa65e
updated documentation concerning 'named_theorems';
 wenzelm parents: 
57342diff
changeset | 425 | addsimps (thms @ my_simps)) i)) | 
| 
6a26aa5fa65e
updated documentation concerning 'named_theorems';
 wenzelm parents: 
57342diff
changeset | 426 | end)\<close> | 
| 57340 | 427 | "rewrite subgoal by given rules and my_simp rules from the context" | 
| 39847 | 428 | |
| 58618 | 429 | text \<open> | 
| 61416 | 430 | \<^medskip> | 
| 61854 | 431 |   Method @{method my_simp'} can be used in Isar proofs like this:
 | 
| 58618 | 432 | \<close> | 
| 39847 | 433 | |
| 40964 | 434 | notepad | 
| 435 | begin | |
| 57342 | 436 | fix a b c :: 'a | 
| 39847 | 437 | assume [my_simp]: "a \<equiv> b" | 
| 438 | assume [my_simp]: "b \<equiv> c" | |
| 39848 | 439 | have "a \<equiv> c" by my_simp' | 
| 40964 | 440 | end | 
| 39847 | 441 | |
| 61416 | 442 | text \<open> | 
| 443 | \<^medskip> | |
| 61854 | 444 |   The @{method my_simp} variants defined above are ``simple'' methods, i.e.\
 | 
| 69597 | 445 | the goal facts are merely inserted as goal premises by the \<^ML>\<open>SIMPLE_METHOD'\<close> or \<^ML>\<open>SIMPLE_METHOD\<close> wrapper. For proof methods that are | 
| 61854 | 446 |   similar to the standard collection of @{method simp}, @{method blast},
 | 
| 447 |   @{method fast}, @{method auto} there is little more that can be done.
 | |
| 39847 | 448 | |
| 61854 | 449 | Note that using the primary goal facts in the same manner as the method | 
| 450 | arguments obtained via concrete syntax or the context does not meet the | |
| 451 | requirement of ``strong emphasis on facts'' of regular proof methods, | |
| 452 | because rewrite rules as used above can be easily ignored. A proof text | |
| 453 |   ``@{command using}~\<open>foo\<close>~@{command "by"}~\<open>my_simp\<close>'' where \<open>foo\<close> is not used
 | |
| 454 | would deceive the reader. | |
| 39847 | 455 | |
| 61416 | 456 | \<^medskip> | 
| 61854 | 457 | The technical treatment of rules from the context requires further | 
| 69597 | 458 | attention. Above we rebuild a fresh \<^ML_type>\<open>simpset\<close> from the arguments | 
| 61854 | 459 | and \<^emph>\<open>all\<close> rules retrieved from the context on every invocation of the | 
| 460 | method. This does not scale to really large collections of rules, which | |
| 461 | easily emerges in the context of a big theory library, for example. | |
| 39847 | 462 | |
| 39848 | 463 | This is an inherent limitation of the simplistic rule management via | 
| 61854 | 464 |   @{command named_theorems}, because it lacks tool-specific storage and
 | 
| 465 | retrieval. More realistic applications require efficient index-structures | |
| 466 | that organize theorems in a customized manner, such as a discrimination net | |
| 467 | that is indexed by the left-hand sides of rewrite rules. For variations on | |
| 69597 | 468 | the Simplifier, re-use of the existing type \<^ML_type>\<open>simpset\<close> is adequate, | 
| 61854 | 469 | but scalability would require it be maintained statically within the context | 
| 470 | data, not dynamically on each tool invocation. | |
| 471 | \<close> | |
| 39847 | 472 | |
| 29759 | 473 | |
| 58618 | 474 | section \<open>Attributes \label{sec:attributes}\<close>
 | 
| 20472 | 475 | |
| 61854 | 476 | text \<open> | 
| 477 | An \<^emph>\<open>attribute\<close> is a function \<open>context \<times> thm \<rightarrow> context \<times> thm\<close>, which means | |
| 478 | both a (generic) context and a theorem can be modified simultaneously. In | |
| 479 | practice this mixed form is very rare, instead attributes are presented | |
| 480 | either as \<^emph>\<open>declaration attribute:\<close> \<open>thm \<rightarrow> context \<rightarrow> context\<close> or \<^emph>\<open>rule | |
| 61493 | 481 | attribute:\<close> \<open>context \<rightarrow> thm \<rightarrow> thm\<close>. | 
| 39849 | 482 | |
| 61854 | 483 | Attributes can have additional arguments via concrete syntax. There is a | 
| 484 | collection of context-sensitive parsers for various logical entities (types, | |
| 485 | terms, theorems). These already take care of applying morphisms to the | |
| 486 | arguments when attribute expressions are moved into a different context (see | |
| 487 |   also \secref{sec:morphisms}).
 | |
| 39849 | 488 | |
| 61854 | 489 | When implementing declaration attributes, it is important to operate exactly | 
| 490 | on the variant of the generic context that is provided by the system, which | |
| 491 | is either global theory context or local proof context. In particular, the | |
| 492 | background theory of a local context must not be modified in this | |
| 493 | situation! | |
| 494 | \<close> | |
| 39849 | 495 | |
| 58618 | 496 | text %mlref \<open> | 
| 39849 | 497 |   \begin{mldecls}
 | 
| 73764 | 498 |   @{define_ML_type attribute} \\
 | 
| 499 |   @{define_ML Thm.rule_attribute: "thm list ->
 | |
| 61853 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 wenzelm parents: 
61841diff
changeset | 500 | (Context.generic -> thm -> thm) -> attribute"} \\ | 
| 73764 | 501 |   @{define_ML Thm.declaration_attribute: "
 | 
| 39849 | 502 | (thm -> Context.generic -> Context.generic) -> attribute"} \\ | 
| 73764 | 503 |   @{define_ML Attrib.setup: "binding -> attribute context_parser ->
 | 
| 39849 | 504 | string -> theory -> theory"} \\ | 
| 505 |   \end{mldecls}
 | |
| 506 | ||
| 69597 | 507 | \<^descr> Type \<^ML_type>\<open>attribute\<close> represents attributes as concrete type alias. | 
| 39849 | 508 | |
| 69597 | 509 | \<^descr> \<^ML>\<open>Thm.rule_attribute\<close>~\<open>thms (fn context => rule)\<close> wraps a | 
| 510 | context-dependent rule (mapping on \<^ML_type>\<open>thm\<close>) as attribute. | |
| 61853 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 wenzelm parents: 
61841diff
changeset | 511 | |
| 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 wenzelm parents: 
61841diff
changeset | 512 | The \<open>thms\<close> are additional parameters: when forming an abstract closure, the | 
| 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 wenzelm parents: 
61841diff
changeset | 513 | system may provide dummy facts that are propagated according to strict | 
| 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 wenzelm parents: 
61841diff
changeset | 514 | evaluation discipline. In that case, \<open>rule\<close> is bypassed. | 
| 39849 | 515 | |
| 69597 | 516 | \<^descr> \<^ML>\<open>Thm.declaration_attribute\<close>~\<open>(fn thm => decl)\<close> wraps a | 
| 517 | theorem-dependent declaration (mapping on \<^ML_type>\<open>Context.generic\<close>) as | |
| 61853 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 wenzelm parents: 
61841diff
changeset | 518 | attribute. | 
| 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 wenzelm parents: 
61841diff
changeset | 519 | |
| 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 wenzelm parents: 
61841diff
changeset | 520 | When forming an abstract closure, the system may provide a dummy fact as | 
| 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 wenzelm parents: 
61841diff
changeset | 521 | \<open>thm\<close>. In that case, \<open>decl\<close> is bypassed. | 
| 39849 | 522 | |
| 69597 | 523 | \<^descr> \<^ML>\<open>Attrib.setup\<close>~\<open>name parser description\<close> provides the functionality of | 
| 61854 | 524 |   the Isar command @{command attribute_setup} as ML function.
 | 
| 58618 | 525 | \<close> | 
| 39849 | 526 | |
| 58618 | 527 | text %mlantiq \<open> | 
| 45592 | 528 |   \begin{matharray}{rcl}
 | 
| 61493 | 529 |   @{ML_antiquotation_def attributes} & : & \<open>ML_antiquotation\<close> \\
 | 
| 45592 | 530 |   \end{matharray}
 | 
| 531 | ||
| 69597 | 532 | \<^rail>\<open> | 
| 45592 | 533 |   @@{ML_antiquotation attributes} attributes
 | 
| 69597 | 534 | \<close> | 
| 45592 | 535 | |
| 61854 | 536 |   \<^descr> \<open>@{attributes [\<dots>]}\<close> embeds attribute source representation into the ML
 | 
| 69597 | 537 | text, which is particularly useful with declarations like \<^ML>\<open>Local_Theory.note\<close>. Attribute names are internalized at compile time, but | 
| 61854 | 538 | the source is unevaluated. This means attributes with formal arguments | 
| 539 | (types, terms, theorems) may be subject to odd effects of dynamic scoping! | |
| 58618 | 540 | \<close> | 
| 45592 | 541 | |
| 61854 | 542 | text %mlex \<open> | 
| 76987 | 543 |   See also @{command attribute_setup} in \<^cite>\<open>"isabelle-isar-ref"\<close> which
 | 
| 61854 | 544 | includes some abstract examples. | 
| 545 | \<close> | |
| 30272 | 546 | |
| 20472 | 547 | end |