| author | Christian Urban <urbanc@in.tum.de> | 
| Tue, 03 Mar 2009 13:00:55 +0000 | |
| changeset 156 | e8f11280c762 | 
| parent 153 | c22b507e1407 | 
| child 157 | 76cdc8f562fc | 
| permissions | -rw-r--r-- | 
| 93 | 1 | theory Tactical | 
| 99 | 2 | imports Base FirstSteps | 
| 142 | 3 | uses "infix_conv.ML" | 
| 93 | 4 | begin | 
| 5 | ||
| 6 | chapter {* Tactical Reasoning\label{chp:tactical} *}
 | |
| 7 | ||
| 8 | text {*
 | |
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 9 | The main reason for descending to the ML-level of Isabelle is to be able to | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 10 | implement automatic proof procedures. Such proof procedures usually lessen | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 11 | considerably the burden of manual reasoning, for example, when introducing | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 12 | new definitions. These proof procedures are centred around refining a goal | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 13 |   state using tactics. This is similar to the \isacommand{apply}-style
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 14 | reasoning at the user level, where goals are modified in a sequence of proof | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 15 | steps until all of them are solved. However, there are also more structured | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 16 | operations available on the ML-level that help with the handling of | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 17 | variables and assumptions. | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 18 | |
| 93 | 19 | *} | 
| 20 | ||
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 21 | section {* Basics of Reasoning with Tactics*}
 | 
| 93 | 22 | |
| 23 | text {*
 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 24 |   To see how tactics work, let us first transcribe a simple \isacommand{apply}-style proof 
 | 
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 25 | into ML. Suppose the following proof. | 
| 93 | 26 | *} | 
| 27 | ||
| 28 | lemma disj_swap: "P \<or> Q \<Longrightarrow> Q \<or> P" | |
| 29 | apply(erule disjE) | |
| 30 | apply(rule disjI2) | |
| 31 | apply(assumption) | |
| 32 | apply(rule disjI1) | |
| 33 | apply(assumption) | |
| 34 | done | |
| 35 | ||
| 36 | text {*
 | |
| 37 | This proof translates to the following ML-code. | |
| 38 | ||
| 39 | @{ML_response_fake [display,gray]
 | |
| 40 | "let | |
| 41 |   val ctxt = @{context}
 | |
| 42 |   val goal = @{prop \"P \<or> Q \<Longrightarrow> Q \<or> P\"}
 | |
| 43 | in | |
| 44 | Goal.prove ctxt [\"P\", \"Q\"] [] goal | |
| 45 | (fn _ => | |
| 46 |       etac @{thm disjE} 1
 | |
| 47 |       THEN rtac @{thm disjI2} 1
 | |
| 48 | THEN atac 1 | |
| 49 |       THEN rtac @{thm disjI1} 1
 | |
| 50 | THEN atac 1) | |
| 51 | end" "?P \<or> ?Q \<Longrightarrow> ?Q \<or> ?P"} | |
| 52 | ||
| 53 |   To start the proof, the function @{ML "Goal.prove"}~@{text "ctxt xs As C
 | |
| 99 | 54 |   tac"} sets up a goal state for proving the goal @{text C} 
 | 
| 55 |   (that is @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"} in the proof at hand) under the
 | |
| 56 |   assumptions @{text As} (happens to be empty) with the variables
 | |
| 93 | 57 |   @{text xs} that will be generalised once the goal is proved (in our case
 | 
| 58 |   @{text P} and @{text Q}). The @{text "tac"} is the tactic that proves the goal;
 | |
| 59 | it can make use of the local assumptions (there are none in this example). | |
| 120 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: 
118diff
changeset | 60 |   The functions @{ML etac}, @{ML rtac} and @{ML atac} in the code above correspond to 
 | 
| 93 | 61 |   @{text erule}, @{text rule} and @{text assumption}, respectively. 
 | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 62 |   The operator @{ML THEN} strings the tactics together. 
 | 
| 93 | 63 | |
| 64 |   \begin{readmore}
 | |
| 99 | 65 |   To learn more about the function @{ML Goal.prove} see \isccite{sec:results}
 | 
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 66 |   and the file @{ML_file "Pure/goal.ML"}.  See @{ML_file
 | 
| 99 | 67 |   "Pure/tactic.ML"} and @{ML_file "Pure/tctical.ML"} for the code of basic
 | 
| 68 | tactics and tactic combinators; see also Chapters 3 and 4 in the old | |
| 69 | Isabelle Reference Manual. | |
| 93 | 70 |   \end{readmore}
 | 
| 71 | ||
| 120 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: 
118diff
changeset | 72 | Note that in the code above we use antiquotations for referencing the theorems. Many theorems | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 73 | also have ML-bindings with the same name. Therefore, we could also just have | 
| 118 
5f003fdf2653
polished and added more material to the package chapter
 Christian Urban <urbanc@in.tum.de> parents: 
114diff
changeset | 74 |   written @{ML "etac disjE 1"}, or in case where there are no ML-binding obtain
 | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 75 |   the theorem dynamically using the function @{ML thm}; for example 
 | 
| 109 | 76 |   \mbox{@{ML  "etac (thm \"disjE\") 1"}}. Both ways however are considered bad style! 
 | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 77 | The reason | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 78 |   is that the binding for @{ML disjE} can be re-assigned by the user and thus
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 79 | one does not have complete control over which theorem is actually | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 80 | applied. This problem is nicely prevented by using antiquotations, because | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 81 | then the theorems are fixed statically at compile-time. | 
| 93 | 82 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 83 | During the development of automatic proof procedures, you will often find it | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 84 | necessary to test a tactic on examples. This can be conveniently | 
| 93 | 85 |   done with the command \isacommand{apply}@{text "(tactic \<verbopen> \<dots> \<verbclose>)"}. 
 | 
| 86 | Consider the following sequence of tactics | |
| 87 | *} | |
| 88 | ||
| 89 | ML{*val foo_tac = 
 | |
| 90 |       (etac @{thm disjE} 1
 | |
| 91 |        THEN rtac @{thm disjI2} 1
 | |
| 92 | THEN atac 1 | |
| 93 |        THEN rtac @{thm disjI1} 1
 | |
| 94 | THEN atac 1)*} | |
| 95 | ||
| 96 | text {* and the Isabelle proof: *}
 | |
| 97 | ||
| 98 | lemma "P \<or> Q \<Longrightarrow> Q \<or> P" | |
| 99 | apply(tactic {* foo_tac *})
 | |
| 100 | done | |
| 101 | ||
| 102 | text {*
 | |
| 104 | 103 |   By using @{text "tactic \<verbopen> \<dots> \<verbclose>"} you can call from the 
 | 
| 104 |   user level of Isabelle the tactic @{ML foo_tac} or 
 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 105 | any other function that returns a tactic. | 
| 93 | 106 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 107 |   The tactic @{ML foo_tac} is just a sequence of simple tactics stringed
 | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 108 |   together by @{ML THEN}. As can be seen, each simple tactic in @{ML foo_tac}
 | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 109 | has a hard-coded number that stands for the subgoal analysed by the | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 110 |   tactic (@{text "1"} stands for the first, or top-most, subgoal). This hard-coding
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 111 | of goals is sometimes wanted, but usually it is not. To avoid the explicit numbering, | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 112 |   you can write\label{tac:footacprime}
 | 
| 93 | 113 | *} | 
| 114 | ||
| 115 | ML{*val foo_tac' = 
 | |
| 116 |       (etac @{thm disjE} 
 | |
| 117 |        THEN' rtac @{thm disjI2} 
 | |
| 118 | THEN' atac | |
| 119 |        THEN' rtac @{thm disjI1} 
 | |
| 120 | THEN' atac)*} | |
| 121 | ||
| 122 | text {* 
 | |
| 123 | and then give the number for the subgoal explicitly when the tactic is | |
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 124 | called. So in the next proof you can first discharge the second subgoal, and | 
| 109 | 125 | subsequently the first. | 
| 93 | 126 | *} | 
| 127 | ||
| 128 | lemma "P1 \<or> Q1 \<Longrightarrow> Q1 \<or> P1" | |
| 129 | and "P2 \<or> Q2 \<Longrightarrow> Q2 \<or> P2" | |
| 130 | apply(tactic {* foo_tac' 2 *})
 | |
| 131 | apply(tactic {* foo_tac' 1 *})
 | |
| 132 | done | |
| 133 | ||
| 134 | text {*
 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 135 | This kind of addressing is more difficult to achieve when the goal is | 
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 136 | hard-coded inside the tactic. For most operators that combine tactics | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 137 |   (@{ML THEN} is only one such operator) a ``primed'' version exists.
 | 
| 99 | 138 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 139 |   The tactics @{ML foo_tac} and @{ML foo_tac'} are very specific for
 | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 140 |   analysing goals being only of the form @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"}. If the goal is not
 | 
| 109 | 141 | of this form, then they return the error message: | 
| 99 | 142 | |
| 143 |   \begin{isabelle}
 | |
| 144 |   @{text "*** empty result sequence -- proof command failed"}\\
 | |
| 145 |   @{text "*** At command \"apply\"."}
 | |
| 146 |   \end{isabelle}
 | |
| 147 | ||
| 109 | 148 | This means the tactics failed. The reason for this error message is that tactics | 
| 149 | are functions mapping a goal state to a (lazy) sequence of successor states. | |
| 150 | Hence the type of a tactic is: | |
| 151 | *} | |
| 93 | 152 | |
| 109 | 153 | ML{*type tactic = thm -> thm Seq.seq*}
 | 
| 93 | 154 | |
| 109 | 155 | text {*
 | 
| 156 | By convention, if a tactic fails, then it should return the empty sequence. | |
| 157 | Therefore, if you write your own tactics, they should not raise exceptions | |
| 158 | willy-nilly; only in very grave failure situations should a tactic raise the | |
| 159 |   exception @{text THM}.
 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 160 | |
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 161 |   The simplest tactics are @{ML no_tac} and @{ML all_tac}. The first returns
 | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 162 | the empty sequence and is defined as | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 163 | *} | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 164 | |
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 165 | ML{*fun no_tac thm = Seq.empty*}
 | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 166 | |
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 167 | text {*
 | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 168 |   which means @{ML no_tac} always fails. The second returns the given theorem wrapped 
 | 
| 109 | 169 | up in a single member sequence; it is defined as | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 170 | *} | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 171 | |
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 172 | ML{*fun all_tac thm = Seq.single thm*}
 | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 173 | |
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 174 | text {*
 | 
| 109 | 175 |   which means @{ML all_tac} always succeeds, but also does not make any progress 
 | 
| 176 | with the proof. | |
| 93 | 177 | |
| 109 | 178 | The lazy list of possible successor goal states shows through at the user-level | 
| 99 | 179 |   of Isabelle when using the command \isacommand{back}. For instance in the 
 | 
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 180 | following proof there are two possibilities for how to apply | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 181 |   @{ML foo_tac'}: either using the first assumption or the second.
 | 
| 93 | 182 | *} | 
| 183 | ||
| 184 | lemma "\<lbrakk>P \<or> Q; P \<or> Q\<rbrakk> \<Longrightarrow> Q \<or> P" | |
| 185 | apply(tactic {* foo_tac' 1 *})
 | |
| 186 | back | |
| 187 | done | |
| 188 | ||
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 189 | |
| 93 | 190 | text {*
 | 
| 99 | 191 |   By using \isacommand{back}, we construct the proof that uses the
 | 
| 109 | 192 | second assumption. While in the proof above, it does not really matter which | 
| 193 | assumption is used, in more interesting cases provability might depend | |
| 194 | on exploring different possibilities. | |
| 99 | 195 | |
| 93 | 196 |   \begin{readmore}
 | 
| 197 |   See @{ML_file "Pure/General/seq.ML"} for the implementation of lazy
 | |
| 109 | 198 | sequences. In day-to-day Isabelle programming, however, one rarely | 
| 199 | constructs sequences explicitly, but uses the predefined tactics and | |
| 200 | tactic combinators instead. | |
| 93 | 201 |   \end{readmore}
 | 
| 202 | ||
| 104 | 203 | It might be surprising that tactics, which transform | 
| 109 | 204 | one goal state to the next, are functions from theorems to theorem | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 205 | (sequences). The surprise resolves by knowing that every | 
| 104 | 206 | goal state is indeed a theorem. To shed more light on this, | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 207 |   let us modify the code of @{ML all_tac} to obtain the following
 | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 208 | tactic | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 209 | *} | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 210 | |
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 211 | ML{*fun my_print_tac ctxt thm =
 | 
| 132 | 212 | let | 
| 213 | val _ = warning (str_of_thm ctxt thm) | |
| 214 | in | |
| 215 | Seq.single thm | |
| 216 | end*} | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 217 | |
| 109 | 218 | text_raw {*
 | 
| 219 |   \begin{figure}[p]
 | |
| 220 |   \begin{isabelle}
 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 221 | *} | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 222 | lemma shows "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B" | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 223 | apply(tactic {* my_print_tac @{context} *})
 | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 224 | |
| 109 | 225 | txt{* \begin{minipage}{\textwidth}
 | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 226 |       @{subgoals [display]} 
 | 
| 109 | 227 |       \end{minipage}\medskip      
 | 
| 228 | ||
| 229 |       \begin{minipage}{\textwidth}
 | |
| 230 |       \small\colorbox{gray!20}{
 | |
| 231 |       \begin{tabular}{@ {}l@ {}}
 | |
| 232 | internal goal state:\\ | |
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 233 |       @{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}
 | 
| 109 | 234 |       \end{tabular}}
 | 
| 235 |       \end{minipage}\medskip
 | |
| 93 | 236 | *} | 
| 237 | ||
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 238 | apply(rule conjI) | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 239 | apply(tactic {* my_print_tac @{context} *})
 | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 240 | |
| 109 | 241 | txt{* \begin{minipage}{\textwidth}
 | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 242 |       @{subgoals [display]} 
 | 
| 109 | 243 |       \end{minipage}\medskip
 | 
| 244 | ||
| 245 |       \begin{minipage}{\textwidth}
 | |
| 246 |       \small\colorbox{gray!20}{
 | |
| 247 |       \begin{tabular}{@ {}l@ {}}
 | |
| 248 | internal goal state:\\ | |
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 249 |       @{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}
 | 
| 109 | 250 |       \end{tabular}}
 | 
| 251 |       \end{minipage}\medskip
 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 252 | *} | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 253 | |
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 254 | apply(assumption) | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 255 | apply(tactic {* my_print_tac @{context} *})
 | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 256 | |
| 109 | 257 | txt{* \begin{minipage}{\textwidth}
 | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 258 |       @{subgoals [display]} 
 | 
| 109 | 259 |       \end{minipage}\medskip
 | 
| 260 | ||
| 261 |       \begin{minipage}{\textwidth}
 | |
| 262 |       \small\colorbox{gray!20}{
 | |
| 263 |       \begin{tabular}{@ {}l@ {}}
 | |
| 264 | internal goal state:\\ | |
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 265 |       @{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}
 | 
| 109 | 266 |       \end{tabular}}
 | 
| 267 |       \end{minipage}\medskip
 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 268 | *} | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 269 | |
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 270 | apply(assumption) | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 271 | apply(tactic {* my_print_tac @{context} *})
 | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 272 | |
| 109 | 273 | txt{* \begin{minipage}{\textwidth}
 | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 274 |       @{subgoals [display]} 
 | 
| 109 | 275 |       \end{minipage}\medskip 
 | 
| 276 | ||
| 277 |       \begin{minipage}{\textwidth}
 | |
| 278 |       \small\colorbox{gray!20}{
 | |
| 279 |       \begin{tabular}{@ {}l@ {}}
 | |
| 280 | internal goal state:\\ | |
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 281 |       @{text "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B"}
 | 
| 109 | 282 |       \end{tabular}}
 | 
| 283 |       \end{minipage}\medskip   
 | |
| 284 | *} | |
| 285 | done | |
| 286 | text_raw {*  
 | |
| 287 |   \end{isabelle}
 | |
| 118 
5f003fdf2653
polished and added more material to the package chapter
 Christian Urban <urbanc@in.tum.de> parents: 
114diff
changeset | 288 |   \caption{The figure shows a proof where each intermediate goal state is
 | 
| 
5f003fdf2653
polished and added more material to the package chapter
 Christian Urban <urbanc@in.tum.de> parents: 
114diff
changeset | 289 |   printed by the Isabelle system and by @{ML my_print_tac}. The latter shows
 | 
| 
5f003fdf2653
polished and added more material to the package chapter
 Christian Urban <urbanc@in.tum.de> parents: 
114diff
changeset | 290 | the goal state as represented internally (highlighted boxes). This | 
| 
5f003fdf2653
polished and added more material to the package chapter
 Christian Urban <urbanc@in.tum.de> parents: 
114diff
changeset | 291 | illustrates that every goal state in Isabelle is represented by a theorem: | 
| 156 | 292 |   when you start the proof of \mbox{@{text "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B"}} the theorem is
 | 
| 293 |   @{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}; when you finish the proof the
 | |
| 118 
5f003fdf2653
polished and added more material to the package chapter
 Christian Urban <urbanc@in.tum.de> parents: 
114diff
changeset | 294 |   theorem is @{text "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B"}.\label{fig:goalstates}}
 | 
| 109 | 295 |   \end{figure}
 | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 296 | *} | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 297 | |
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 298 | |
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 299 | text {*
 | 
| 109 | 300 | which prints out the given theorem (using the string-function defined in | 
| 301 |   Section~\ref{sec:printing}) and then behaves like @{ML all_tac}. With this
 | |
| 302 | tactic we are in the position to inspect every goal state in a | |
| 303 |   proof. Consider now the proof in Figure~\ref{fig:goalstates}: as can be seen, 
 | |
| 304 | internally every goal state is an implication of the form | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 305 | |
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 306 |   @{text[display] "A\<^isub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^isub>n \<Longrightarrow> (C)"}
 | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 307 | |
| 109 | 308 |   where @{term C} is the goal to be proved and the @{term "A\<^isub>i"} are
 | 
| 309 | the subgoals. So after setting up the lemma, the goal state is always of the | |
| 310 |   form @{text "C \<Longrightarrow> (C)"}; when the proof is finished we are left with @{text
 | |
| 311 |   "(C)"}. Since the goal @{term C} can potentially be an implication, there is
 | |
| 312 |   a ``protector'' wrapped around it (in from of an outermost constant @{text
 | |
| 120 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: 
118diff
changeset | 313 | "Const (\"prop\", bool \<Rightarrow> bool)"}; however this constant | 
| 149 | 314 |   is invisible in the figure). This wrapper prevents that premises of @{text C} are
 | 
| 109 | 315 | mis-interpreted as open subgoals. While tactics can operate on the subgoals | 
| 316 |   (the @{text "A\<^isub>i"} above), they are expected to leave the conclusion
 | |
| 317 |   @{term C} intact, with the exception of possibly instantiating schematic
 | |
| 318 | variables. If you use the predefined tactics, which we describe in the next | |
| 319 | section, this will always be the case. | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 320 | |
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 321 |   \begin{readmore}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 322 |   For more information about the internals of goals see \isccite{sec:tactical-goals}.
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 323 |   \end{readmore}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 324 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 325 | *} | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 326 | |
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 327 | section {* Simple Tactics *}
 | 
| 93 | 328 | |
| 99 | 329 | text {*
 | 
| 109 | 330 |   Let us start with the tactic @{ML print_tac}, which is quite useful for low-level 
 | 
| 156 | 331 | debugging of tactics. It just prints out a message and the current goal state | 
| 332 |   (unlike @{ML my_print_tac} it prints the goal state as the user would see it). 
 | |
| 333 | For example, processing the proof | |
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 334 | *} | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 335 | |
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 336 | lemma shows "False \<Longrightarrow> True" | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 337 | apply(tactic {* print_tac "foo message" *})
 | 
| 109 | 338 | txt{*gives:\medskip
 | 
| 339 | ||
| 340 |      \begin{minipage}{\textwidth}\small
 | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 341 |      @{text "foo message"}\\[3mm] 
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 342 |      @{prop "False \<Longrightarrow> True"}\\
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 343 |      @{text " 1. False \<Longrightarrow> True"}\\
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 344 |      \end{minipage}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 345 | *} | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 346 | (*<*)oops(*>*) | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 347 | |
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 348 | text {*
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 349 |   Another simple tactic is the function @{ML atac}, which, as shown in the previous
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 350 | section, corresponds to the assumption command. | 
| 99 | 351 | *} | 
| 352 | ||
| 353 | lemma shows "P \<Longrightarrow> P" | |
| 93 | 354 | apply(tactic {* atac 1 *})
 | 
| 109 | 355 | txt{*\begin{minipage}{\textwidth}
 | 
| 356 |      @{subgoals [display]}
 | |
| 357 |      \end{minipage}*}
 | |
| 358 | (*<*)oops(*>*) | |
| 93 | 359 | |
| 99 | 360 | text {*
 | 
| 361 |   Similarly, @{ML rtac}, @{ML dtac}, @{ML etac} and @{ML ftac} correspond
 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 362 |   to @{text rule}, @{text drule}, @{text erule} and @{text frule}, 
 | 
| 109 | 363 | respectively. Each of them takes a theorem as argument and attempts to | 
| 364 | apply it to a goal. Below are three self-explanatory examples. | |
| 99 | 365 | *} | 
| 366 | ||
| 367 | lemma shows "P \<and> Q" | |
| 93 | 368 | apply(tactic {* rtac @{thm conjI} 1 *})
 | 
| 104 | 369 | txt{*\begin{minipage}{\textwidth}
 | 
| 370 |      @{subgoals [display]}
 | |
| 371 |      \end{minipage}*}
 | |
| 93 | 372 | (*<*)oops(*>*) | 
| 373 | ||
| 99 | 374 | lemma shows "P \<and> Q \<Longrightarrow> False" | 
| 93 | 375 | apply(tactic {* etac @{thm conjE} 1 *})
 | 
| 104 | 376 | txt{*\begin{minipage}{\textwidth}
 | 
| 377 |      @{subgoals [display]}
 | |
| 378 |      \end{minipage}*}
 | |
| 93 | 379 | (*<*)oops(*>*) | 
| 380 | ||
| 381 | lemma shows "False \<and> True \<Longrightarrow> False" | |
| 382 | apply(tactic {* dtac @{thm conjunct2} 1 *})
 | |
| 104 | 383 | txt{*\begin{minipage}{\textwidth}
 | 
| 384 |      @{subgoals [display]}
 | |
| 385 |      \end{minipage}*}
 | |
| 93 | 386 | (*<*)oops(*>*) | 
| 387 | ||
| 388 | text {*
 | |
| 109 | 389 | Note the number in each tactic call. Also as mentioned in the previous section, most | 
| 156 | 390 | basic tactics take such a number as argument; it addresses the subgoal they are analysing. | 
| 109 | 391 | In the proof below, we first split up the conjunction in the second subgoal by | 
| 392 | focusing on this subgoal first. | |
| 99 | 393 | *} | 
| 394 | ||
| 395 | lemma shows "Foo" and "P \<and> Q" | |
| 396 | apply(tactic {* rtac @{thm conjI} 2 *})
 | |
| 104 | 397 | txt {*\begin{minipage}{\textwidth}
 | 
| 398 |       @{subgoals [display]}
 | |
| 399 |       \end{minipage}*}
 | |
| 99 | 400 | (*<*)oops(*>*) | 
| 401 | ||
| 109 | 402 | text {*
 | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 403 |   The function @{ML resolve_tac} is similar to @{ML rtac}, except that it
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 404 | expects a list of theorems as arguments. From this list it will apply the | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 405 | first applicable theorem (later theorems that are also applicable can be | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 406 | explored via the lazy sequences mechanism). Given the code | 
| 93 | 407 | *} | 
| 408 | ||
| 99 | 409 | ML{*val resolve_tac_xmp = resolve_tac [@{thm impI}, @{thm conjI}]*}
 | 
| 410 | ||
| 411 | text {*
 | |
| 412 |   an example for @{ML resolve_tac} is the following proof where first an outermost 
 | |
| 413 | implication is analysed and then an outermost conjunction. | |
| 414 | *} | |
| 415 | ||
| 416 | lemma shows "C \<longrightarrow> (A \<and> B)" and "(A \<longrightarrow> B) \<and> C" | |
| 417 | apply(tactic {* resolve_tac_xmp 1 *})
 | |
| 418 | apply(tactic {* resolve_tac_xmp 2 *})
 | |
| 104 | 419 | txt{*\begin{minipage}{\textwidth}
 | 
| 420 |      @{subgoals [display]} 
 | |
| 421 |      \end{minipage}*}
 | |
| 99 | 422 | (*<*)oops(*>*) | 
| 423 | ||
| 424 | text {* 
 | |
| 109 | 425 | Similarly versions taking a list of theorems exist for the tactics | 
| 426 |   @{ML dtac} (@{ML dresolve_tac}), @{ML etac} (@{ML eresolve_tac}) and so on.
 | |
| 427 | ||
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 428 | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 429 |   Another simple tactic is @{ML cut_facts_tac}. It inserts a list of theorems
 | 
| 109 | 430 | into the assumptions of the current goal state. For example | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 431 | *} | 
| 99 | 432 | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 433 | lemma shows "True \<noteq> False" | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 434 | apply(tactic {* cut_facts_tac [@{thm True_def}, @{thm False_def}] 1 *})
 | 
| 109 | 435 | txt{*produces the goal state\medskip
 | 
| 436 | ||
| 437 |      \begin{minipage}{\textwidth}
 | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 438 |      @{subgoals [display]} 
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 439 |      \end{minipage}*}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 440 | (*<*)oops(*>*) | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 441 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 442 | text {*
 | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 443 | Since rules are applied using higher-order unification, an automatic proof | 
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 444 | procedure might become too fragile, if it just applies inference rules as | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 445 | shown above. The reason is that a number of rules introduce meta-variables | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 446 | into the goal state. Consider for example the proof | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 447 | *} | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 448 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 449 | lemma shows "\<forall>x\<in>A. P x \<Longrightarrow> Q x" | 
| 118 
5f003fdf2653
polished and added more material to the package chapter
 Christian Urban <urbanc@in.tum.de> parents: 
114diff
changeset | 450 | apply(tactic {* dtac @{thm bspec} 1 *})
 | 
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 451 | txt{*\begin{minipage}{\textwidth}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 452 |      @{subgoals [display]} 
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 453 |      \end{minipage}*}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 454 | (*<*)oops(*>*) | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 455 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 456 | text {*
 | 
| 149 | 457 |   where the application of rule @{text bspec} generates two subgoals involving the
 | 
| 109 | 458 |   meta-variable @{text "?x"}. Now, if you are not careful, tactics 
 | 
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 459 | applied to the first subgoal might instantiate this meta-variable in such a | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 460 |   way that the second subgoal becomes unprovable. If it is clear what the @{text "?x"}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 461 | should be, then this situation can be avoided by introducing a more | 
| 109 | 462 |   constraint version of the @{text bspec}-rule. Such constraints can be given by
 | 
| 463 | pre-instantiating theorems with other theorems. One function to do this is | |
| 464 |   @{ML RS}
 | |
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 465 | |
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 466 |   @{ML_response_fake [display,gray]
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 467 |   "@{thm disjI1} RS @{thm conjI}" "\<lbrakk>?P1; ?Q\<rbrakk> \<Longrightarrow> (?P1 \<or> ?Q1) \<and> ?Q"}
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 468 | |
| 109 | 469 |   which in the example instantiates the first premise of the @{text conjI}-rule 
 | 
| 470 |   with the rule @{text disjI1}. If the instantiation is impossible, as in the 
 | |
| 471 | case of | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 472 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 473 |   @{ML_response_fake_both [display,gray]
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 474 |   "@{thm conjI} RS @{thm mp}" 
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 475 | "*** Exception- THM (\"RSN: no unifiers\", 1, | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 476 | [\"\<lbrakk>?P; ?Q\<rbrakk> \<Longrightarrow> ?P \<and> ?Q\", \"\<lbrakk>?P \<longrightarrow> ?Q; ?P\<rbrakk> \<Longrightarrow> ?Q\"]) raised"} | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 477 | |
| 109 | 478 |   then the function raises an exception. The function @{ML RSN} is similar to @{ML RS}, but 
 | 
| 479 | takes an additional number as argument that makes explicit which premise | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 480 | should be instantiated. | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 481 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 482 | To improve readability of the theorems we produce below, we shall use | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 483 | the following function | 
| 99 | 484 | *} | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 485 | |
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 486 | ML{*fun no_vars ctxt thm =
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 487 | let | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 488 | val ((_, [thm']), _) = Variable.import_thms true [thm] ctxt | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 489 | in | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 490 | thm' | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 491 | end*} | 
| 93 | 492 | |
| 99 | 493 | text {*
 | 
| 109 | 494 | that transform the schematic variables of a theorem into free variables. | 
| 495 |   Using this function for the first @{ML RS}-expression above would produce
 | |
| 496 | the more readable result: | |
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 497 | |
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 498 |   @{ML_response_fake [display,gray]
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 499 |   "no_vars @{context} (@{thm disjI1} RS @{thm conjI})" "\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> (P \<or> Qa) \<and> Q"}
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 500 | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 501 | If you want to instantiate more than one premise of a theorem, you can use | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 502 |   the function @{ML MRS}:
 | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 503 | |
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 504 |   @{ML_response_fake [display,gray]
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 505 |   "no_vars @{context} ([@{thm disjI1}, @{thm disjI2}] MRS @{thm conjI})" 
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 506 | "\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> (P \<or> Qa) \<and> (Pa \<or> Q)"} | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 507 | |
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 508 | If you need to instantiate lists of theorems, you can use the | 
| 109 | 509 |   functions @{ML RL} and @{ML MRL}. For example in the code below,
 | 
| 510 | every theorem in the second list is instantiated with every | |
| 511 | theorem in the first. | |
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 512 | |
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 513 |   @{ML_response_fake [display,gray]
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 514 |    "[@{thm impI}, @{thm disjI2}] RL [@{thm conjI}, @{thm disjI1}]" 
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 515 | "[\<lbrakk>P \<Longrightarrow> Q; Qa\<rbrakk> \<Longrightarrow> (P \<longrightarrow> Q) \<and> Qa, | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 516 | \<lbrakk>Q; Qa\<rbrakk> \<Longrightarrow> (P \<or> Q) \<and> Qa, | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 517 | (P \<Longrightarrow> Q) \<Longrightarrow> (P \<longrightarrow> Q) \<or> Qa, | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 518 | Q \<Longrightarrow> (P \<or> Q) \<or> Qa]"} | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 519 | |
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 520 |   \begin{readmore}
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 521 |   The combinators for instantiating theorems are defined in @{ML_file "Pure/drule.ML"}.
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 522 |   \end{readmore}
 | 
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 523 | |
| 109 | 524 | Often proofs on the ML-level involve elaborate operations on assumptions and | 
| 525 |   @{text "\<And>"}-quantified variables. To do such operations using the basic tactics 
 | |
| 128 | 526 | shown so far is very unwieldy and brittle. Some convenience and | 
| 99 | 527 |   safety is provided by the tactic @{ML SUBPROOF}. This tactic fixes the parameters 
 | 
| 109 | 528 | and binds the various components of a goal state to a record. | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 529 |   To see what happens, assume the function defined in Figure~\ref{fig:sptac}, which
 | 
| 109 | 530 | takes a record and just prints out the content of this record (using the | 
| 531 |   string transformation functions from in Section~\ref{sec:printing}). Consider
 | |
| 532 | now the proof: | |
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 533 | *} | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 534 | |
| 99 | 535 | text_raw{*
 | 
| 536 | \begin{figure}
 | |
| 537 | \begin{isabelle}
 | |
| 538 | *} | |
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 539 | ML{*fun sp_tac {prems, params, asms, concl, context, schematics} = 
 | 
| 132 | 540 | let | 
| 541 | val str_of_params = str_of_cterms context params | |
| 542 | val str_of_asms = str_of_cterms context asms | |
| 543 | val str_of_concl = str_of_cterm context concl | |
| 544 | val str_of_prems = str_of_thms context prems | |
| 545 | val str_of_schms = str_of_cterms context (snd schematics) | |
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 546 | |
| 132 | 547 |   val _ = (warning ("params: " ^ str_of_params);
 | 
| 548 |            warning ("schematics: " ^ str_of_schms);
 | |
| 549 |            warning ("assumptions: " ^ str_of_asms);
 | |
| 550 |            warning ("conclusion: " ^ str_of_concl);
 | |
| 551 |            warning ("premises: " ^ str_of_prems))
 | |
| 552 | in | |
| 553 | no_tac | |
| 554 | end*} | |
| 99 | 555 | text_raw{*
 | 
| 556 | \end{isabelle}
 | |
| 557 | \caption{A function that prints out the various parameters provided by the tactic
 | |
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 558 |  @{ML SUBPROOF}. It uses the functions defined in Section~\ref{sec:printing} for
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 559 |   extracting strings from @{ML_type cterm}s and @{ML_type thm}s.\label{fig:sptac}}
 | 
| 99 | 560 | \end{figure}
 | 
| 561 | *} | |
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 562 | |
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 563 | |
| 99 | 564 | lemma shows "\<And>x y. A x y \<Longrightarrow> B y x \<longrightarrow> C (?z y) x" | 
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 565 | apply(tactic {* SUBPROOF sp_tac @{context} 1 *})?
 | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 566 | |
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 567 | txt {*
 | 
| 109 | 568 | The tactic produces the following printout: | 
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 569 | |
| 99 | 570 |   \begin{quote}\small
 | 
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 571 |   \begin{tabular}{ll}
 | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 572 |   params:      & @{term x}, @{term y}\\
 | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 573 |   schematics:  & @{term z}\\
 | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 574 |   assumptions: & @{term "A x y"}\\
 | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 575 |   conclusion:  & @{term "B y x \<longrightarrow> C (z y) x"}\\
 | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 576 |   premises:    & @{term "A x y"}
 | 
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 577 |   \end{tabular}
 | 
| 99 | 578 |   \end{quote}
 | 
| 579 | ||
| 149 | 580 |   Notice in the actual output the brown colour of the variables @{term x} and 
 | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 581 |   @{term y}. Although they are parameters in the original goal, they are fixed inside
 | 
| 109 | 582 | the subproof. By convention these fixed variables are printed in brown colour. | 
| 583 |   Similarly the schematic variable @{term z}. The assumption, or premise, 
 | |
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 584 |   @{prop "A x y"} is bound as @{ML_type cterm} to the record-variable 
 | 
| 109 | 585 |   @{text asms}, but also as @{ML_type thm} to @{text prems}.
 | 
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 586 | |
| 109 | 587 |   Notice also that we had to append @{text [quotes] "?"} to the
 | 
| 588 |   \isacommand{apply}-command. The reason is that @{ML SUBPROOF} normally
 | |
| 589 |   expects that the subgoal is solved completely.  Since in the function @{ML
 | |
| 590 |   sp_tac} we returned the tactic @{ML no_tac}, the subproof obviously
 | |
| 591 | fails. The question-mark allows us to recover from this failure in a | |
| 592 | graceful manner so that the warning messages are not overwritten by an | |
| 593 | ``empty sequence'' error message. | |
| 99 | 594 | |
| 595 |   If we continue the proof script by applying the @{text impI}-rule
 | |
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 596 | *} | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 597 | |
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 598 | apply(rule impI) | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 599 | apply(tactic {* SUBPROOF sp_tac @{context} 1 *})?
 | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 600 | |
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 601 | txt {*
 | 
| 118 
5f003fdf2653
polished and added more material to the package chapter
 Christian Urban <urbanc@in.tum.de> parents: 
114diff
changeset | 602 | then the tactic prints out: | 
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 603 | |
| 99 | 604 |   \begin{quote}\small
 | 
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 605 |   \begin{tabular}{ll}
 | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 606 |   params:      & @{term x}, @{term y}\\
 | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 607 |   schematics:  & @{term z}\\
 | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 608 |   assumptions: & @{term "A x y"}, @{term "B y x"}\\
 | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 609 |   conclusion:  & @{term "C (z y) x"}\\
 | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 610 |   premises:    & @{term "A x y"}, @{term "B y x"}
 | 
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 611 |   \end{tabular}
 | 
| 99 | 612 |   \end{quote}
 | 
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 613 | *} | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 614 | (*<*)oops(*>*) | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 615 | |
| 99 | 616 | text {*
 | 
| 109 | 617 |   Now also @{term "B y x"} is an assumption bound to @{text asms} and @{text prems}.
 | 
| 99 | 618 | |
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 619 |   One convenience of @{ML SUBPROOF} is that we can apply the assumptions
 | 
| 99 | 620 |   using the usual tactics, because the parameter @{text prems} 
 | 
| 109 | 621 | contains them as theorems. With this you can easily | 
| 622 |   implement a tactic that behaves almost like @{ML atac}:
 | |
| 99 | 623 | *} | 
| 624 | ||
| 104 | 625 | ML{*val atac' = SUBPROOF (fn {prems, ...} => resolve_tac prems 1)*}
 | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 626 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 627 | text {*
 | 
| 109 | 628 |   If you apply @{ML atac'} to the next lemma
 | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 629 | *} | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 630 | |
| 109 | 631 | lemma shows "\<lbrakk>B x y; A x y; C x y\<rbrakk> \<Longrightarrow> A x y" | 
| 104 | 632 | apply(tactic {* atac' @{context} 1 *})
 | 
| 109 | 633 | txt{* it will produce
 | 
| 99 | 634 |       @{subgoals [display]} *}
 | 
| 635 | (*<*)oops(*>*) | |
| 636 | ||
| 104 | 637 | text {*
 | 
| 109 | 638 |   The restriction in this tactic which is not present in @{ML atac} is 
 | 
| 639 | that it cannot instantiate any | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 640 | schematic variable. This might be seen as a defect, but it is actually | 
| 104 | 641 |   an advantage in the situations for which @{ML SUBPROOF} was designed: 
 | 
| 109 | 642 | the reason is that, as mentioned before, instantiation of schematic variables can affect | 
| 104 | 643 |   several goals and can render them unprovable. @{ML SUBPROOF} is meant 
 | 
| 644 | to avoid this. | |
| 645 | ||
| 109 | 646 |   Notice that @{ML atac'} inside @{ML SUBPROOF} calls @{ML resolve_tac} with 
 | 
| 647 |   the subgoal number @{text "1"} and also the outer call to @{ML SUBPROOF} in 
 | |
| 648 |   the \isacommand{apply}-step uses @{text "1"}. This is another advantage 
 | |
| 649 |   of @{ML SUBPROOF}: the addressing  inside it is completely 
 | |
| 650 | local to the tactic inside the subproof. It is therefore possible to also apply | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 651 |   @{ML atac'} to the second goal by just writing:
 | 
| 104 | 652 | *} | 
| 653 | ||
| 109 | 654 | lemma shows "True" and "\<lbrakk>B x y; A x y; C x y\<rbrakk> \<Longrightarrow> A x y" | 
| 104 | 655 | apply(tactic {* atac' @{context} 2 *})
 | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 656 | apply(rule TrueI) | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 657 | done | 
| 104 | 658 | |
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 659 | |
| 93 | 660 | text {*
 | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 661 |   \begin{readmore}
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 662 |   The function @{ML SUBPROOF} is defined in @{ML_file "Pure/subgoal.ML"} and
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 663 |   also described in \isccite{sec:results}. 
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 664 |   \end{readmore}
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 665 | |
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
150diff
changeset | 666 |   Similar but less powerful functions than @{ML SUBPROOF} are @{ML SUBGOAL}
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
150diff
changeset | 667 |   and @{ML CSUBGOAL}. They allow you to inspect a given subgoal (the former
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
150diff
changeset | 668 |   presents the subgoal as a @{ML_type term}, while the latter as a @{ML_type
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
150diff
changeset | 669 | cterm}). With this you can implement a tactic that applies a rule according | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
150diff
changeset | 670 | to the topmost logic connective in the subgoal (to illustrate this we only | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
150diff
changeset | 671 | analyse a few connectives). The code of the tactic is as | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
150diff
changeset | 672 |   follows.\label{tac:selecttac}
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
150diff
changeset | 673 | |
| 93 | 674 | *} | 
| 675 | ||
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
150diff
changeset | 676 | ML %linenosgray{*fun select_tac (t, i) =
 | 
| 99 | 677 | case t of | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
150diff
changeset | 678 |      @{term "Trueprop"} $ t' => select_tac (t', i)
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
150diff
changeset | 679 |    | @{term "op \<Longrightarrow>"} $ _ $ t' => select_tac (t', i)
 | 
| 99 | 680 |    | @{term "op \<and>"} $ _ $ _ => rtac @{thm conjI} i
 | 
| 681 |    | @{term "op \<longrightarrow>"} $ _ $ _ => rtac @{thm impI} i
 | |
| 682 |    | @{term "Not"} $ _ => rtac @{thm notI} i
 | |
| 683 |    | Const (@{const_name "All"}, _) $ _ => rtac @{thm allI} i
 | |
| 104 | 684 | | _ => all_tac*} | 
| 99 | 685 | |
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 686 | text {*
 | 
| 109 | 687 | The input of the function is a term representing the subgoal and a number | 
| 688 | specifying the subgoal of interest. In line 3 you need to descend under the | |
| 689 |   outermost @{term "Trueprop"} in order to get to the connective you like to
 | |
| 690 |   analyse. Otherwise goals like @{prop "A \<and> B"} are not properly
 | |
| 691 | analysed. Similarly with meta-implications in the next line. While for the | |
| 692 |   first five patterns we can use the @{text "@term"}-antiquotation to
 | |
| 693 | construct the patterns, the pattern in Line 8 cannot be constructed in this | |
| 694 | way. The reason is that an antiquotation would fix the type of the | |
| 695 | quantified variable. So you really have to construct the pattern using the | |
| 156 | 696 | basic term-constructors. This is not necessary in other cases, because their | 
| 697 |   type is always fixed to function types involving only the type @{typ
 | |
| 698 |   bool}. (See Section \ref{sec:terms_types_manually} about constructing terms
 | |
| 699 |   manually.) For the catch-all pattern, we chose to just return @{ML all_tac}. 
 | |
| 700 |   Consequently, @{ML select_tac} never fails.
 | |
| 701 | ||
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 702 | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 703 | Let us now see how to apply this tactic. Consider the four goals: | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 704 | *} | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 705 | |
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 706 | |
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 707 | lemma shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F" | 
| 104 | 708 | apply(tactic {* SUBGOAL select_tac 4 *})
 | 
| 709 | apply(tactic {* SUBGOAL select_tac 3 *})
 | |
| 710 | apply(tactic {* SUBGOAL select_tac 2 *})
 | |
| 99 | 711 | apply(tactic {* SUBGOAL select_tac 1 *})
 | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 712 | txt{* \begin{minipage}{\textwidth}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 713 |       @{subgoals [display]}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 714 |       \end{minipage} *}
 | 
| 99 | 715 | (*<*)oops(*>*) | 
| 716 | ||
| 717 | text {*
 | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 718 | where in all but the last the tactic applied an introduction rule. | 
| 109 | 719 | Note that we applied the tactic to the goals in ``reverse'' order. | 
| 720 | This is a trick in order to be independent from the subgoals that are | |
| 721 | produced by the rule. If we had applied it in the other order | |
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 722 | *} | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 723 | |
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 724 | lemma shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F" | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 725 | apply(tactic {* SUBGOAL select_tac 1 *})
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 726 | apply(tactic {* SUBGOAL select_tac 3 *})
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 727 | apply(tactic {* SUBGOAL select_tac 4 *})
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 728 | apply(tactic {* SUBGOAL select_tac 5 *})
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 729 | (*<*)oops(*>*) | 
| 99 | 730 | |
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 731 | text {*
 | 
| 109 | 732 | then we have to be careful to not apply the tactic to the two subgoals produced by | 
| 733 | the first goal. To do this can result in quite messy code. In contrast, | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 734 | the ``reverse application'' is easy to implement. | 
| 104 | 735 | |
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
150diff
changeset | 736 | Of course, this example is | 
| 149 | 737 | contrived: there are much simpler methods available in Isabelle for | 
| 738 | implementing a proof procedure analysing a goal according to its topmost | |
| 739 | connective. These simpler methods use tactic combinators, which we will | |
| 740 | explain in the next section. | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 741 | |
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 742 | *} | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 743 | |
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 744 | section {* Tactic Combinators *}
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 745 | |
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 746 | text {* 
 | 
| 109 | 747 | The purpose of tactic combinators is to build compound tactics out of | 
| 748 |   smaller tactics. In the previous section we already used @{ML THEN}, which
 | |
| 749 | just strings together two tactics in a sequence. For example: | |
| 93 | 750 | *} | 
| 751 | ||
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 752 | lemma shows "(Foo \<and> Bar) \<and> False" | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 753 | apply(tactic {* rtac @{thm conjI} 1 THEN rtac @{thm conjI} 1 *})
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 754 | txt {* \begin{minipage}{\textwidth}
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 755 |        @{subgoals [display]} 
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 756 |        \end{minipage} *}
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 757 | (*<*)oops(*>*) | 
| 99 | 758 | |
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 759 | text {*
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 760 | If you want to avoid the hard-coded subgoal addressing, then you can use | 
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 761 |   the ``primed'' version of @{ML THEN}. For example:
 | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 762 | *} | 
| 93 | 763 | |
| 99 | 764 | lemma shows "(Foo \<and> Bar) \<and> False" | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 765 | apply(tactic {* (rtac @{thm conjI} THEN' rtac @{thm conjI}) 1 *})
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 766 | txt {* \begin{minipage}{\textwidth}
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 767 |        @{subgoals [display]} 
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 768 |        \end{minipage} *}
 | 
| 93 | 769 | (*<*)oops(*>*) | 
| 770 | ||
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 771 | text {* 
 | 
| 109 | 772 | Here you only have to specify the subgoal of interest only once and | 
| 773 | it is consistently applied to the component tactics. | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 774 | For most tactic combinators such a ``primed'' version exists and | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 775 | in what follows we will usually prefer it over the ``unprimed'' one. | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 776 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 777 | If there is a list of tactics that should all be tried out in | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 778 |   sequence, you can use the combinator @{ML EVERY'}. For example
 | 
| 109 | 779 |   the function @{ML foo_tac'} from page~\pageref{tac:footacprime} can also 
 | 
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 780 | be written as: | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 781 | *} | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 782 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 783 | ML{*val foo_tac'' = EVERY' [etac @{thm disjE}, rtac @{thm disjI2}, 
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 784 |                         atac, rtac @{thm disjI1}, atac]*}
 | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 785 | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 786 | text {*
 | 
| 109 | 787 | There is even another way of implementing this tactic: in automatic proof | 
| 788 | procedures (in contrast to tactics that might be called by the user) there | |
| 789 | are often long lists of tactics that are applied to the first | |
| 790 |   subgoal. Instead of writing the code above and then calling @{ML "foo_tac'' 1"}, 
 | |
| 791 | you can also just write | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 792 | *} | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 793 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 794 | ML{*val foo_tac1 = EVERY1 [etac @{thm disjE}, rtac @{thm disjI2}, 
 | 
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 795 |                        atac, rtac @{thm disjI1}, atac]*}
 | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 796 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 797 | text {*
 | 
| 118 
5f003fdf2653
polished and added more material to the package chapter
 Christian Urban <urbanc@in.tum.de> parents: 
114diff
changeset | 798 |   and call @{ML foo_tac1}.  
 | 
| 109 | 799 | |
| 800 |   With the combinators @{ML THEN'}, @{ML EVERY'} and @{ML EVERY1} it must be
 | |
| 801 | guaranteed that all component tactics successfully apply; otherwise the | |
| 802 | whole tactic will fail. If you rather want to try out a number of tactics, | |
| 803 |   then you can use the combinator @{ML ORELSE'} for two tactics, and @{ML
 | |
| 804 |   FIRST'} (or @{ML FIRST1}) for a list of tactics. For example, the tactic
 | |
| 805 | ||
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 806 | *} | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 807 | |
| 131 | 808 | ML{*val orelse_xmp = rtac @{thm disjI1} ORELSE' rtac @{thm conjI}*}
 | 
| 99 | 809 | |
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 810 | text {*
 | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 811 |   will first try out whether rule @{text disjI} applies and after that 
 | 
| 109 | 812 |   @{text conjI}. To see this consider the proof
 | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 813 | *} | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 814 | |
| 99 | 815 | lemma shows "True \<and> False" and "Foo \<or> Bar" | 
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 816 | apply(tactic {* orelse_xmp 2 *})
 | 
| 99 | 817 | apply(tactic {* orelse_xmp 1 *})
 | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 818 | txt {* which results in the goal state
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 819 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 820 |        \begin{minipage}{\textwidth}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 821 |        @{subgoals [display]} 
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 822 |        \end{minipage}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 823 | *} | 
| 93 | 824 | (*<*)oops(*>*) | 
| 825 | ||
| 826 | ||
| 827 | text {*
 | |
| 109 | 828 |   Using @{ML FIRST'} we can simplify our @{ML select_tac} from Page~\pageref{tac:selecttac} 
 | 
| 829 | as follows: | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 830 | *} | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 831 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 832 | ML{*val select_tac' = FIRST' [rtac @{thm conjI}, rtac @{thm impI}, 
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 833 |                           rtac @{thm notI}, rtac @{thm allI}, K all_tac]*}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 834 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 835 | text {*
 | 
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 836 |   Since we like to mimic the behaviour of @{ML select_tac} as closely as possible, 
 | 
| 109 | 837 |   we must include @{ML all_tac} at the end of the list, otherwise the tactic will
 | 
| 118 
5f003fdf2653
polished and added more material to the package chapter
 Christian Urban <urbanc@in.tum.de> parents: 
114diff
changeset | 838 |   fail if no rule applies (we also have to wrap @{ML all_tac} using the 
 | 
| 109 | 839 |   @{ML K}-combinator, because it does not take a subgoal number as argument). You can 
 | 
| 840 | test the tactic on the same goals: | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 841 | *} | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 842 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 843 | lemma shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F" | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 844 | apply(tactic {* select_tac' 4 *})
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 845 | apply(tactic {* select_tac' 3 *})
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 846 | apply(tactic {* select_tac' 2 *})
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 847 | apply(tactic {* select_tac' 1 *})
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 848 | txt{* \begin{minipage}{\textwidth}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 849 |       @{subgoals [display]}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 850 |       \end{minipage} *}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 851 | (*<*)oops(*>*) | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 852 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 853 | text {* 
 | 
| 109 | 854 | Since such repeated applications of a tactic to the reverse order of | 
| 855 |   \emph{all} subgoals is quite common, there is the tactic combinator 
 | |
| 856 |   @{ML ALLGOALS} that simplifies this. Using this combinator you can simply 
 | |
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 857 | write: *} | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 858 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 859 | lemma shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F" | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 860 | apply(tactic {* ALLGOALS select_tac' *})
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 861 | txt{* \begin{minipage}{\textwidth}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 862 |       @{subgoals [display]}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 863 |       \end{minipage} *}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 864 | (*<*)oops(*>*) | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 865 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 866 | text {*
 | 
| 109 | 867 |   Remember that we chose to implement @{ML select_tac'} so that it 
 | 
| 868 | always succeeds. This can be potentially very confusing for the user, | |
| 869 | for example, in cases where the goal is the form | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 870 | *} | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 871 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 872 | lemma shows "E \<Longrightarrow> F" | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 873 | apply(tactic {* select_tac' 1 *})
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 874 | txt{* \begin{minipage}{\textwidth}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 875 |       @{subgoals [display]}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 876 |       \end{minipage} *}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 877 | (*<*)oops(*>*) | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 878 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 879 | text {*
 | 
| 109 | 880 | In this case no rule applies. The problem for the user is that there is little | 
| 881 | chance to see whether or not progress in the proof has been made. By convention | |
| 882 | therefore, tactics visible to the user should either change something or fail. | |
| 883 | ||
| 884 |   To comply with this convention, we could simply delete the @{ML "K all_tac"}
 | |
| 885 |   from the end of the theorem list. As a result @{ML select_tac'} would only
 | |
| 886 | succeed on goals where it can make progress. But for the sake of argument, | |
| 887 |   let us suppose that this deletion is \emph{not} an option. In such cases, you can
 | |
| 888 |   use the combinator @{ML CHANGED} to make sure the subgoal has been changed
 | |
| 889 | by the tactic. Because now | |
| 890 | ||
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 891 | *} | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 892 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 893 | lemma shows "E \<Longrightarrow> F" | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 894 | apply(tactic {* CHANGED (select_tac' 1) *})(*<*)?(*>*)
 | 
| 109 | 895 | txt{* gives the error message:
 | 
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 896 |   \begin{isabelle}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 897 |   @{text "*** empty result sequence -- proof command failed"}\\
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 898 |   @{text "*** At command \"apply\"."}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 899 |   \end{isabelle} 
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 900 | *}(*<*)oops(*>*) | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 901 | |
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 902 | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 903 | text {*
 | 
| 109 | 904 |   We can further extend @{ML select_tac'} so that it not just applies to the topmost
 | 
| 905 | connective, but also to the ones immediately ``underneath'', i.e.~analyse the goal | |
| 906 |   completely. For this you can use the tactic combinator @{ML REPEAT}. As an example 
 | |
| 907 | suppose the following tactic | |
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 908 | *} | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 909 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 910 | ML{*val repeat_xmp = REPEAT (CHANGED (select_tac' 1)) *}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 911 | |
| 109 | 912 | text {* which applied to the proof *}
 | 
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 913 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 914 | lemma shows "((\<not>A) \<and> (\<forall>x. B x)) \<and> (C \<longrightarrow> D)" | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 915 | apply(tactic {* repeat_xmp *})
 | 
| 109 | 916 | txt{* produces
 | 
| 917 | ||
| 918 |       \begin{minipage}{\textwidth}
 | |
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 919 |       @{subgoals [display]}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 920 |       \end{minipage} *}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 921 | (*<*)oops(*>*) | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 922 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 923 | text {*
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 924 |   Here it is crucial that @{ML select_tac'} is prefixed with @{ML CHANGED}, 
 | 
| 109 | 925 |   because otherwise @{ML REPEAT} runs into an infinite loop (it applies the
 | 
| 926 | tactic as long as it succeeds). The function | |
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 927 |   @{ML REPEAT1} is similar, but runs the tactic at least once (failing if 
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 928 | this is not possible). | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 929 | |
| 156 | 930 |   If you are after the ``primed'' version of @{ML repeat_xmp}, then you 
 | 
| 109 | 931 | need to implement it as | 
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 932 | *} | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 933 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 934 | ML{*val repeat_xmp' = REPEAT o CHANGED o select_tac'*}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 935 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 936 | text {*
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 937 |   since there are no ``primed'' versions of @{ML REPEAT} and @{ML CHANGED}.
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 938 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 939 |   If you look closely at the goal state above, the tactics @{ML repeat_xmp}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 940 |   and @{ML repeat_xmp'} are not yet quite what we are after: the problem is
 | 
| 109 | 941 | that goals 2 and 3 are not analysed. This is because the tactic | 
| 942 | is applied repeatedly only to the first subgoal. To analyse also all | |
| 943 |   resulting subgoals, you can use the tactic combinator @{ML REPEAT_ALL_NEW}. 
 | |
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 944 | Suppose the tactic | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 945 | *} | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 946 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 947 | ML{*val repeat_all_new_xmp = REPEAT_ALL_NEW (CHANGED o select_tac')*}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 948 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 949 | text {* 
 | 
| 109 | 950 | you see that the following goal | 
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 951 | *} | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 952 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 953 | lemma shows "((\<not>A) \<and> (\<forall>x. B x)) \<and> (C \<longrightarrow> D)" | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 954 | apply(tactic {* repeat_all_new_xmp 1 *})
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 955 | txt{* \begin{minipage}{\textwidth}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 956 |       @{subgoals [display]}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 957 |       \end{minipage} *}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 958 | (*<*)oops(*>*) | 
| 93 | 959 | |
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 960 | text {*
 | 
| 109 | 961 | is completely analysed according to the theorems we chose to | 
| 120 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: 
118diff
changeset | 962 |   include in @{ML select_tac'}. 
 | 
| 109 | 963 | |
| 964 | Recall that tactics produce a lazy sequence of successor goal states. These | |
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 965 |   states can be explored using the command \isacommand{back}. For example
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 966 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 967 | *} | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 968 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 969 | lemma "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R" | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 970 | apply(tactic {* etac @{thm disjE} 1 *})
 | 
| 109 | 971 | txt{* applies the rule to the first assumption yielding the goal state:\smallskip
 | 
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 972 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 973 |       \begin{minipage}{\textwidth}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 974 |       @{subgoals [display]}
 | 
| 109 | 975 |       \end{minipage}\smallskip 
 | 
| 976 | ||
| 977 | After typing | |
| 978 | *} | |
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 979 | (*<*) | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 980 | oops | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 981 | lemma "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R" | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 982 | apply(tactic {* etac @{thm disjE} 1 *})
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 983 | (*>*) | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 984 | back | 
| 109 | 985 | txt{* the rule now applies to the second assumption.\smallskip
 | 
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 986 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 987 |       \begin{minipage}{\textwidth}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 988 |       @{subgoals [display]}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 989 |       \end{minipage} *}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 990 | (*<*)oops(*>*) | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 991 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 992 | text {*
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 993 | Sometimes this leads to confusing behaviour of tactics and also has | 
| 109 | 994 | the potential to explode the search space for tactics. | 
| 995 | These problems can be avoided by prefixing the tactic with the tactic | |
| 996 |   combinator @{ML DETERM}.
 | |
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 997 | *} | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 998 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 999 | lemma "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R" | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1000 | apply(tactic {* DETERM (etac @{thm disjE} 1) *})
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1001 | txt {*\begin{minipage}{\textwidth}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1002 |       @{subgoals [display]}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1003 |       \end{minipage} *}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1004 | (*<*)oops(*>*) | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1005 | text {*
 | 
| 118 
5f003fdf2653
polished and added more material to the package chapter
 Christian Urban <urbanc@in.tum.de> parents: 
114diff
changeset | 1006 | This combinator will prune the search space to just the first successful application. | 
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1007 |   Attempting to apply \isacommand{back} in this goal states gives the
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1008 | error message: | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1009 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1010 |   \begin{isabelle}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1011 |   @{text "*** back: no alternatives"}\\
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1012 |   @{text "*** At command \"back\"."}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1013 |   \end{isabelle}
 | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1014 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1015 |   \begin{readmore}
 | 
| 109 | 1016 |   Most tactic combinators described in this section are defined in @{ML_file "Pure/tctical.ML"}.
 | 
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1017 |   \end{readmore}
 | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1018 | |
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 1019 | *} | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 1020 | |
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1021 | section {* Simplifier Tactics (TBD) *}
 | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 1022 | |
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 1023 | text {*
 | 
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1024 | A lot of convenience in the reasoning with Isabelle derives from its | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1025 | powerful simplifier. There are the following five main tactics behind | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1026 | the simplifier (in parentheses is their userlevel counterpart) | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1027 | |
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1028 |   \begin{isabelle}
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1029 |   \begin{tabular}{l@ {\hspace{2cm}}l}
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1030 |    @{ML simp_tac}            & @{text "(simp (no_asm))"} \\
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1031 |    @{ML asm_simp_tac}        & @{text "(simp (no_asm_simp))"} \\
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1032 |    @{ML full_simp_tac}       & @{text "(simp (no_asm_use))"} \\
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1033 |    @{ML asm_lr_simp_tac}     & @{text "(simp (asm_lr))"} \\
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1034 |    @{ML asm_full_simp_tac}   & @{text "(simp)"}
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1035 |   \end{tabular}
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1036 |   \end{isabelle}
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1037 | |
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1038 | All of them take a simpset and an interger as argument (the latter to specify the goal | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1039 | to be analysed). So the proof | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1040 | *} | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1041 | |
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1042 | lemma "Suc (1 + 2) < 3 + 2" | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1043 | apply(simp) | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1044 | done | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1045 | |
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1046 | text {*
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1047 | corrsponds on the ML-level to the tactic | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1048 | *} | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1049 | |
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1050 | lemma "Suc (1 + 2) < 3 + 2" | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1051 | apply(tactic {* asm_full_simp_tac @{simpset} 1 *})
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1052 | done | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1053 | |
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1054 | thm imp_cong simp_implies_cong | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1055 | |
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1056 | text {*
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1057 | The crucial information the developer has to control is the simpset | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1058 | to be used by the simplifier. If not handled with care, then the | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1059 | simplifier can easily run forever. | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1060 | |
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1061 | It might be surprising that a simpset is quite more complex than just | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1062 | a simple list of theorems. One reason for the complexity is that the | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1063 | simplifier must be able to rewrite inside terms and should also be | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1064 | able to rewrite according to rules that have precoditions. The rewriting | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1065 | inside terms requires congruence rules which are meta-equalities | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1066 | typical of the form | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1067 | |
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1068 |   \begin{isabelle}
 | 
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1069 |   \mbox{\inferrule{@{text "t\<^isub>i \<equiv> s\<^isub>i"}}
 | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1070 |                   {@{text "contr t\<^isub>1\<dots>t\<^isub>n \<equiv> contr s\<^isub>1\<dots>s\<^isub>n"}}}
 | 
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1071 |   \end{isabelle}
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1072 | |
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1073 |   with @{text "constr"} being a term-constructor. However there are also more complicated 
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1074 | congruence rules. The user can declare lemmas to be congruence rules using the | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1075 |   attribute @{text "[cong]"} (theses lemmas are usually equations that are internally 
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1076 | transformed into meta-equations (FIXME: check that). | 
| 120 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: 
118diff
changeset | 1077 | |
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1078 | The rewriting with rules involving preconditions requires solvers. However | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1079 | there are also simprocs, which can produce rewrite rules on demand (see | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1080 | next section). Another component are split rules, which can simplify for example | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1081 | the then and else branches of if-statements under the corresponding | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1082 | precoditions. (FIXME: loopers and subgoalers?) | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1083 | |
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1084 | The most common combinators to modify simpsets are | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1085 | |
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1086 |   \begin{isabelle}
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1087 |   \begin{tabular}{ll}
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1088 |   @{ML addsimps} & @{ML delsimps}\\
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1089 |   @{ML addcongs} & @{ML delcongs}\\
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1090 |   \end{tabular}
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1091 |   \end{isabelle}
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1092 | |
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1093 |   The most frequently used simpsets in HOL are @{ML HOL_basic_ss} and @{ML HOL_ss}.
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1094 | *} | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1095 | |
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1096 | ML {* warning (Pretty.string_of (MetaSimplifier.pretty_ss HOL_ss)) *}
 | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1097 | |
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1098 | ML {*
 | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1099 | fun get_parts ss = | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1100 | let | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1101 |   val ({rules, ...}, {congs, procs, loop_tacs, solvers, ...}) = MetaSimplifier.rep_ss ss
 | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1102 | in | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1103 | (rules, congs, procs, loop_tacs, solvers) | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1104 | end | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1105 | *} | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1106 | |
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1107 | ML {*
 | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1108 | Pretty.big_list | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1109 | *} | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1110 | |
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1111 | |
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1112 | ML {*
 | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1113 | fun print_ss ss = | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1114 | let | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1115 | val pretty_thms = map (enclose " " "\n" o Display.string_of_thm) | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1116 | |
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1117 |     fun pretty_cong (name, {thm, lhs}) =
 | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1118 | name ^ ":" ^ (Display.string_of_thm thm); | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1119 | |
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1120 | val (rules, congs, procs, loop_tacs, solvers) = get_parts ss; | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1121 | val smps = map #thm (Net.entries rules); | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1122 | |
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1123 | in | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1124 | "simplification rules:\n" ^ | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1125 | (implode (pretty_thms smps)) ^ | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1126 | "congruences:" ^ (commas (map pretty_cong (fst congs))) ^ "\n" ^ | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1127 | "loopers:" ^ (commas (map (quote o fst) loop_tacs)) | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1128 | end; | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1129 | *} | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1130 | |
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1131 | thm FalseE | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1132 | |
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1133 | thm simp_impliesI | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1134 | lemma "P (Suc 0) \<Longrightarrow> P ?x" | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1135 | apply(tactic {* CHANGED (simp_tac HOL_basic_ss 1) *})
 | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1136 | oops | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1137 | |
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1138 | ML {* warning (print_ss HOL_ss) *}
 | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1139 | |
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1140 | text {*
 | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1141 |   @{ML HOL_basic_ss} can deal essentially only with goals of the form: 
 | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1142 |   @{thm TrueI}, @{thm refl[no_vars]} @{term "x \<equiv> x"} and @{thm FalseE}, 
 | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1143 | and resolve with assumptions. | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1144 | *} | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1145 | |
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1146 | ML {* setsubgoaler *}
 | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1147 | |
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1148 | text {*
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1149 | (FIXME: what do the simplifier tactics do when nothing can be rewritten) | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1150 | |
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1151 | |
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1152 |   (FIXME: @{ML rewrite_goals_tac}, @{ML ObjectLogic.full_atomize_tac}, 
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1153 |   @{ML ObjectLogic.rulify_tac})
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1154 | |
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1155 | |
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1156 |   \begin{readmore}
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1157 |   For more information about the simplifier see @{ML_file "Pure/meta_simplifier.ML"}
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1158 |   and @{ML_file "Pure/simplifier.ML"}.
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1159 |   \end{readmore}
 | 
| 120 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: 
118diff
changeset | 1160 | |
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1161 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1162 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1163 | section {* Simprocs *}
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1164 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1165 | text {*
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1166 | In Isabelle you can also implement custom simplification procedures, called | 
| 149 | 1167 |   \emph{simprocs}. Simprocs can be triggered by the simplifier on a specified
 | 
| 1168 | term-pattern and rewrite a term according to a theorem. They are useful in | |
| 1169 | cases where a rewriting rule must be produced on ``demand'' or when | |
| 1170 | rewriting by simplification is too unpredictable and potentially loops. | |
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1171 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1172 | To see how simprocs work, let us first write a simproc that just prints out | 
| 132 | 1173 | the pattern which triggers it and otherwise does nothing. For this | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1174 | you can use the function: | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1175 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1176 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1177 | ML %linenosgray{*fun fail_sp_aux simpset redex = 
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1178 | let | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1179 | val ctxt = Simplifier.the_context simpset | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1180 |   val _ = warning ("The redex: " ^ (str_of_cterm ctxt redex))
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1181 | in | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1182 | NONE | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1183 | end*} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1184 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1185 | text {*
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1186 |   This function takes a simpset and a redex (a @{ML_type cterm}) as
 | 
| 132 | 1187 | arguments. In Lines 3 and~4, we first extract the context from the given | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1188 | simpset and then print out a message containing the redex. The function | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1189 |   returns @{ML NONE} (standing for an optional @{ML_type thm}) since at the
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1190 |   moment we are \emph{not} interested in actually rewriting anything. We want
 | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1191 |   that the simproc is triggered by the pattern @{term "Suc n"}. This can be
 | 
| 149 | 1192 | done by adding the simproc to the current simpset as follows | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1193 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1194 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1195 | simproc_setup fail_sp ("Suc n") = {* K fail_sp_aux *}
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1196 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1197 | text {*
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1198 | where the second argument specifies the pattern and the right-hand side | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1199 |   contains the code of the simproc (we have to use @{ML K} since we ignoring
 | 
| 131 | 1200 |   an argument about morphisms\footnote{FIXME: what does the morphism do?}). 
 | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1201 | After this, the simplifier is aware of the simproc and you can test whether | 
| 131 | 1202 | it fires on the lemma: | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1203 | *} | 
| 120 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: 
118diff
changeset | 1204 | |
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1205 | lemma shows "Suc 0 = 1" | 
| 131 | 1206 | apply(simp) | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1207 | (*<*)oops(*>*) | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1208 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1209 | text {*
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1210 | This will print out the message twice: once for the left-hand side and | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1211 | once for the right-hand side. The reason is that during simplification the | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1212 |   simplifier will at some point reduce the term @{term "1::nat"} to @{term "Suc
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1213 | 0"}, and then the simproc ``fires'' also on that term. | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1214 | |
| 131 | 1215 | We can add or delete the simproc from the current simpset by the usual | 
| 132 | 1216 |   \isacommand{declare}-statement. For example the simproc will be deleted
 | 
| 1217 | with the declaration | |
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1218 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1219 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1220 | declare [[simproc del: fail_sp]] | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1221 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1222 | text {*
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1223 |   If you want to see what happens with just \emph{this} simproc, without any 
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1224 |   interference from other rewrite rules, you can call @{text fail_sp} 
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1225 | as follows: | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1226 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1227 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1228 | lemma shows "Suc 0 = 1" | 
| 132 | 1229 |   apply(tactic {* simp_tac (HOL_basic_ss addsimprocs [@{simproc fail_sp}]) 1*})
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1230 | (*<*)oops(*>*) | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1231 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1232 | text {*
 | 
| 131 | 1233 |   Now the message shows up only once since the term @{term "1::nat"} is 
 | 
| 1234 | left unchanged. | |
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1235 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1236 |   Setting up a simproc using the command \isacommand{setup\_simproc} will
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1237 | always add automatically the simproc to the current simpset. If you do not | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1238 | want this, then you have to use a slightly different method for setting | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1239 |   up the simproc. First the function @{ML fail_sp_aux} needs to be modified
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1240 | to | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1241 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1242 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1243 | ML{*fun fail_sp_aux' simpset redex = 
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1244 | let | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1245 | val ctxt = Simplifier.the_context simpset | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1246 |   val _ = warning ("The redex: " ^ (Syntax.string_of_term ctxt redex))
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1247 | in | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1248 | NONE | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1249 | end*} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1250 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1251 | text {*
 | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1252 |   Here the redex is given as a @{ML_type term}, instead of a @{ML_type cterm}
 | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1253 |   (therefore we printing it out using the function @{ML string_of_term in Syntax}).
 | 
| 149 | 1254 | We can turn this function into a proper simproc using the function | 
| 1255 |   @{ML Simplifier.simproc_i}:
 | |
| 93 | 1256 | *} | 
| 1257 | ||
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 1258 | |
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1259 | ML{*val fail_sp' = 
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1260 | let | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1261 |   val thy = @{theory}
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1262 |   val pat = [@{term "Suc n"}]
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1263 | in | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1264 | Simplifier.simproc_i thy "fail_sp'" pat (K fail_sp_aux') | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1265 | end*} | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1266 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1267 | text {*
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1268 |   Here the pattern is given as @{ML_type term} (instead of @{ML_type cterm}).
 | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1269 | The function also takes a list of patterns that can trigger the simproc. | 
| 132 | 1270 | Now the simproc is set up and can be explicitly added using | 
| 149 | 1271 |   @{ML addsimprocs} to a simpset whenerver
 | 
| 132 | 1272 | needed. | 
| 1273 | ||
| 1274 | Simprocs are applied from inside to outside and from left to right. You can | |
| 1275 | see this in the proof | |
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1276 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1277 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1278 | lemma shows "Suc (Suc 0) = (Suc 1)" | 
| 149 | 1279 |   apply(tactic {* simp_tac (HOL_basic_ss addsimprocs [fail_sp']) 1*})
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1280 | (*<*)oops(*>*) | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1281 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1282 | text {*
 | 
| 132 | 1283 |   The simproc @{ML fail_sp'} prints out the sequence 
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1284 | |
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1285 | @{text [display]
 | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1286 | "> Suc 0 | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1287 | > Suc (Suc 0) | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1288 | > Suc 1"} | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1289 | |
| 131 | 1290 | To see how a simproc applies a theorem, let us implement a simproc that | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1291 | rewrites terms according to the equation: | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1292 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1293 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1294 | lemma plus_one: | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1295 | shows "Suc n \<equiv> n + 1" by simp | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1296 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1297 | text {*
 | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1298 | Simprocs expect that the given equation is a meta-equation, however the | 
| 131 | 1299 | equation can contain preconditions (the simproc then will only fire if the | 
| 132 | 1300 | preconditions can be solved). To see that one has relatively precise control over | 
| 131 | 1301 | the rewriting with simprocs, let us further assume we want that the simproc | 
| 1302 |   only rewrites terms ``greater'' than @{term "Suc 0"}. For this we can write 
 | |
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1303 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1304 | |
| 131 | 1305 | |
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1306 | ML{*fun plus_one_sp_aux ss redex =
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1307 | case redex of | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1308 |     @{term "Suc 0"} => NONE
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1309 |   | _ => SOME @{thm plus_one}*}
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1310 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1311 | text {*
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1312 | and set up the simproc as follows. | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1313 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1314 | |
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1315 | ML{*val plus_one_sp =
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1316 | let | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1317 |   val thy = @{theory}
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1318 |   val pat = [@{term "Suc n"}] 
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1319 | in | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1320 | Simplifier.simproc_i thy "sproc +1" pat (K plus_one_sp_aux) | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1321 | end*} | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1322 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1323 | text {*
 | 
| 132 | 1324 | Now the simproc is set up so that it is triggered by terms | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1325 |   of the form @{term "Suc n"}, but inside the simproc we only produce
 | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1326 |   a theorem if the term is not @{term "Suc 0"}. The result you can see
 | 
| 131 | 1327 | in the following proof | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1328 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1329 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1330 | lemma shows "P (Suc (Suc (Suc 0))) (Suc 0)" | 
| 149 | 1331 |   apply(tactic {* simp_tac (HOL_basic_ss addsimprocs [plus_one_sp]) 1*})
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1332 | txt{*
 | 
| 131 | 1333 | where the simproc produces the goal state | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1334 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1335 |   @{subgoals[display]}
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1336 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1337 | (*<*)oops(*>*) | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1338 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1339 | text {*
 | 
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1340 | As usual with rewriting you have to worry about looping: you already have | 
| 132 | 1341 |   a loop with @{ML plus_one_sp}, if you apply it with the default simpset (because
 | 
| 1342 |   the default simpset contains a rule which just does the opposite of @{ML plus_one_sp},
 | |
| 1343 |   namely rewriting @{text [quotes] "+ 1"} to a successor). So you have to be careful 
 | |
| 1344 | in choosing the right simpset to which you add a simproc. | |
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1345 | |
| 132 | 1346 |   Next let us implement a simproc that replaces terms of the form @{term "Suc n"}
 | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1347 |   with the number @{text n} increase by one. First we implement a function that
 | 
| 132 | 1348 | takes a term and produces the corresponding integer value. | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1349 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1350 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1351 | ML{*fun dest_suc_trm ((Const (@{const_name "Suc"}, _)) $ t) = 1 + dest_suc_trm t
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1352 | | dest_suc_trm t = snd (HOLogic.dest_number t)*} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1353 | |
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1354 | text {* 
 | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1355 |   It uses the library function @{ML dest_number in HOLogic} that transforms
 | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1356 |   (Isabelle) terms, like @{term "0::nat"}, @{term "1::nat"}, @{term "2::nat"} and so
 | 
| 131 | 1357 |   on, into integer values. This function raises the exception @{ML TERM}, if
 | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1358 | the term is not a number. The next function expects a pair consisting of a term | 
| 131 | 1359 |   @{text t} (containing @{term Suc}s) and the corresponding integer value @{text n}. 
 | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1360 | *} | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1361 | |
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1362 | ML %linenosgray{*fun get_thm ctxt (t, n) =
 | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1363 | let | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1364 |   val num = HOLogic.mk_number @{typ "nat"} n
 | 
| 132 | 1365 | val goal = Logic.mk_equals (t, num) | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1366 | in | 
| 132 | 1367 | Goal.prove ctxt [] [] goal (K (arith_tac ctxt 1)) | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1368 | end*} | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1369 | |
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1370 | text {*
 | 
| 132 | 1371 | From the integer value it generates the corresponding number term, called | 
| 1372 |   @{text num} (Line 3), and then generates the meta-equation @{text "t \<equiv> num"} 
 | |
| 1373 | (Line 4), which it proves by the arithmetic tactic in Line 6. | |
| 1374 | ||
| 134 | 1375 |   For our purpose at the moment, proving the meta-equation using @{ML arith_tac} is
 | 
| 132 | 1376 | fine, but there is also an alternative employing the simplifier with a very | 
| 1377 | restricted simpset. For the kind of lemmas we want to prove, the simpset | |
| 1378 |   @{text "num_ss"} in the code will suffice.
 | |
| 1379 | *} | |
| 131 | 1380 | |
| 132 | 1381 | ML{*fun get_thm_alt ctxt (t, n) =
 | 
| 1382 | let | |
| 1383 |   val num = HOLogic.mk_number @{typ "nat"} n
 | |
| 1384 | val goal = Logic.mk_equals (t, num) | |
| 1385 |   val num_ss = HOL_ss addsimps [@{thm One_nat_def}, @{thm Let_def}] @ 
 | |
| 1386 |           @{thms nat_number} @ @{thms neg_simps} @ @{thms plus_nat.simps}
 | |
| 1387 | in | |
| 1388 | Goal.prove ctxt [] [] goal (K (simp_tac num_ss 1)) | |
| 1389 | end*} | |
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1390 | |
| 132 | 1391 | text {*
 | 
| 1392 |   The advantage of @{ML get_thm_alt} is that it leaves very little room for 
 | |
| 1393 | something to go wrong; in contrast it is much more difficult to predict | |
| 1394 |   what happens with @{ML arith_tac}, especially in more complicated 
 | |
| 1395 |   circumstances. The disatvantage of @{ML get_thm_alt} is to find a simpset
 | |
| 1396 | that is sufficiently powerful to solve every instance of the lemmas | |
| 1397 | we like to prove. This requires careful tuning, but is often necessary in | |
| 1398 |   ``production code''.\footnote{It would be of great help if there is another
 | |
| 1399 | way than tracing the simplifier to obtain the lemmas that are successfully | |
| 1400 | applied during simplification. Alas, there is none.} | |
| 1401 | ||
| 1402 | Anyway, either version can be used in the function that produces the actual | |
| 1403 | theorem for the simproc. | |
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1404 | *} | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1405 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1406 | ML{*fun nat_number_sp_aux ss t =
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1407 | let | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1408 | val ctxt = Simplifier.the_context ss | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1409 | in | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1410 | SOME (get_thm ctxt (t, dest_suc_trm t)) | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1411 | handle TERM _ => NONE | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1412 | end*} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1413 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1414 | text {*
 | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1415 |   This function uses the fact that @{ML dest_suc_trm} might throw an exception 
 | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1416 |   @{ML TERM}. In this case there is nothing that can be rewritten and therefore no
 | 
| 131 | 1417 |   theorem is produced (i.e.~the function returns @{ML NONE}). To try out the simproc 
 | 
| 1418 | on an example, you can set it up as follows: | |
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1419 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1420 | |
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1421 | ML{*val nat_number_sp =
 | 
| 132 | 1422 | let | 
| 1423 |   val thy = @{theory}
 | |
| 1424 |   val pat = [@{term "Suc n"}]
 | |
| 1425 | in | |
| 1426 | Simplifier.simproc_i thy "nat_number" pat (K nat_number_sp_aux) | |
| 1427 | end*} | |
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1428 | |
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1429 | text {* 
 | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1430 | Now in the lemma | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1431 | *} | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1432 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1433 | lemma "P (Suc (Suc 2)) (Suc 99) (0::nat) (Suc 4 + Suc 0) (Suc (0 + 0))" | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1434 |   apply(tactic {* simp_tac (HOL_ss addsimprocs [nat_number_sp]) 1*})
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1435 | txt {* 
 | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1436 | you obtain the more legible goal state | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1437 | |
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1438 |   @{subgoals [display]}
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1439 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1440 | (*<*)oops(*>*) | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1441 | |
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1442 | text {*
 | 
| 132 | 1443 |   where the simproc rewrites all @{term "Suc"}s except in the last argument. There it cannot 
 | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1444 |   rewrite anything, because it does not know how to transform the term @{term "Suc (0 + 0)"}
 | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1445 | into a number. To solve this problem have a look at the next exercise. | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1446 | |
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1447 |   \begin{exercise}\label{ex:addsimproc}
 | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1448 |   Write a simproc that replaces terms of the form @{term "t\<^isub>1 + t\<^isub>2"} by their 
 | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1449 | result. You can assume the terms are ``proper'' numbers, that is of the form | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1450 |   @{term "0::nat"}, @{term "1::nat"}, @{term "2::nat"} and so on.
 | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1451 |   \end{exercise}
 | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1452 | |
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1453 | (FIXME: We did not do anything with morphisms. Anything interesting | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1454 | one can say about them?) | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1455 | *} | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1456 | |
| 137 | 1457 | section {* Conversions\label{sec:conversion} *}
 | 
| 132 | 1458 | |
| 135 | 1459 | text {*
 | 
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 1460 | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1461 | Conversions are a thin layer on top of Isabelle's inference kernel, and | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1462 | can be viewed be as a controllable, bare-bone version of Isabelle's simplifier. | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1463 | One difference between conversions and the simplifier is that the former | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1464 |   act on @{ML_type cterm}s while the latter acts on @{ML_type thm}s. 
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1465 | However, we will also show in this section how conversions can be applied | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1466 | to theorems via tactics. The type for conversions is | 
| 135 | 1467 | *} | 
| 1468 | ||
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 1469 | ML{*type conv = Thm.cterm -> Thm.thm*}
 | 
| 135 | 1470 | |
| 1471 | text {*
 | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1472 | whereby the produced theorem is always a meta-equality. A simple conversion | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1473 |   is the function @{ML "Conv.all_conv"}, which maps a @{ML_type cterm} to an
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1474 | instance of the (meta)reflexivity theorem. For example: | 
| 135 | 1475 | |
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 1476 |   @{ML_response_fake [display,gray]
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1477 |   "Conv.all_conv @{cterm \"Foo \<or> Bar\"}"
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1478 | "Foo \<or> Bar \<equiv> Foo \<or> Bar"} | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1479 | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1480 |   Another simple conversion is @{ML Conv.no_conv} which always raises the 
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1481 |   exception @{ML CTERM}.
 | 
| 135 | 1482 | |
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 1483 |   @{ML_response_fake [display,gray]
 | 
| 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 1484 |   "Conv.no_conv @{cterm True}" 
 | 
| 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 1485 | "*** Exception- CTERM (\"no conversion\", []) raised"} | 
| 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 1486 | |
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1487 |   A more interesting conversion is the function @{ML "Thm.beta_conversion"}: it 
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1488 | produces an equation between a term and its beta-normal form. For example | 
| 142 | 1489 | |
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 1490 |   @{ML_response_fake [display,gray]
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1491 | "let | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1492 |   val add = @{cterm \"\<lambda>x y. x + (y::nat)\"}
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1493 |   val two = @{cterm \"2::nat\"}
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1494 |   val ten = @{cterm \"10::nat\"}
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1495 | in | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1496 | Thm.beta_conversion true (Thm.capply (Thm.capply add two) ten) | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1497 | end" | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1498 | "((\<lambda>x y. x + y) 2) 10 \<equiv> 2 + 10"} | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1499 | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1500 |   Note that the actual response in this example is @{term "2 + 10 \<equiv> 2 + (10::nat)"}, 
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1501 |   since the pretty printer for @{ML_type cterm}s already beta-normalises terms.
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1502 | But by the way how we constructed the term (using the function | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1503 |   @{ML Thm.capply}, which is the application @{ML $} for @{ML_type cterm}s),
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1504 | we can be sure the left-hand side must contain beta-redexes. Indeed | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1505 | if we obtain the ``raw'' representation of the produced theorem, we | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1506 | can see the difference: | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1507 | |
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1508 |   @{ML_response [display,gray]
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1509 | "let | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1510 |   val add = @{cterm \"\<lambda>x y. x + (y::nat)\"}
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1511 |   val two = @{cterm \"2::nat\"}
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1512 |   val ten = @{cterm \"10::nat\"}
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1513 | in | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1514 | #prop (rep_thm | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1515 | (Thm.beta_conversion true (Thm.capply (Thm.capply add two) ten))) | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1516 | end" | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1517 | "Const (\"==\",\<dots>) $ | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1518 | (Abs (\"x\",\<dots>,Abs (\"y\",\<dots>,\<dots>)) $\<dots>$\<dots>) $ | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1519 | (Const (\"HOL.plus_class.plus\",\<dots>) $ \<dots> $ \<dots>)"} | 
| 142 | 1520 | |
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1521 |   The argument @{ML true} in @{ML Thm.beta_conversion} indicates that 
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1522 | the right-hand side will be fully beta-normalised. If instead | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1523 |   @{ML false} is given, then only a single beta-reduction is performed 
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1524 | on the outer-most level. For example | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1525 | |
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1526 |   @{ML_response_fake [display,gray]
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1527 | "let | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1528 |   val add = @{cterm \"\<lambda>x y. x + (y::nat)\"}
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1529 |   val two = @{cterm \"2::nat\"}
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1530 | in | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1531 | Thm.beta_conversion false (Thm.capply add two) | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1532 | end" | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1533 | "((\<lambda>x y. x + y) 2) \<equiv> \<lambda>y. 2 + y"} | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1534 | |
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1535 | Again, we actually see as output only the fully normalised term | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1536 |   @{text "\<lambda>y. 2 + y"}.
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1537 | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1538 | The main point of conversions is that they can be used for rewriting | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1539 |   @{ML_type cterm}s. To do this you can use the function @{ML
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1540 | "Conv.rewr_conv"} which expects a meta-equation as an argument. Suppose we | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1541 |   want to rewrite a @{ML_type cterm} according to the (meta)equation:
 | 
| 135 | 1542 | *} | 
| 1543 | ||
| 139 | 1544 | lemma true_conj1: "True \<and> P \<equiv> P" by simp | 
| 135 | 1545 | |
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1546 | text {* 
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1547 | You can see how this function works with the example | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1548 |   @{term "True \<and> (Foo \<longrightarrow> Bar)"}:
 | 
| 139 | 1549 | |
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 1550 |   @{ML_response_fake [display,gray]
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1551 | "let | 
| 149 | 1552 |   val ctrm = @{cterm \"True \<and> (Foo \<longrightarrow> Bar)\"}
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1553 | in | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1554 |   Conv.rewr_conv @{thm true_conj1} ctrm
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1555 | end" | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1556 | "True \<and> (Foo \<longrightarrow> Bar) \<equiv> Foo \<longrightarrow> Bar"} | 
| 139 | 1557 | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1558 |   Note, however, that the function @{ML Conv.rewr_conv} only rewrites the 
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1559 |   outer-most level of the @{ML_type cterm}. If the given @{ML_type cterm} does not match the 
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1560 |   left-hand side of the theorem, then  @{ML Conv.rewr_conv} raises 
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1561 |   the exception @{ML CTERM}.
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1562 | |
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1563 | This very primitive way of rewriting can be made more powerful by | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1564 | combining several conversions into one. For this you can use conversion | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1565 |   combinators. The simplest conversion combinator is @{ML then_conv}, 
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1566 | which applies one conversion after another. For example | 
| 139 | 1567 | |
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 1568 |   @{ML_response_fake [display,gray]
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1569 | "let | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1570 | val conv1 = Thm.beta_conversion false | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1571 |   val conv2 = Conv.rewr_conv @{thm true_conj1}
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1572 |   val ctrm = Thm.capply @{cterm \"\<lambda>x. x \<and> False\"} @{cterm \"True\"}
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1573 | in | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1574 | (conv1 then_conv conv2) ctrm | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1575 | end" | 
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 1576 | "(\<lambda>x. x \<and> False) True \<equiv> False"} | 
| 139 | 1577 | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1578 | where we first beta-reduce the term and then rewrite according to | 
| 149 | 1579 |   @{thm [source] true_conj1}. (Recall the problem with the pretty-printer
 | 
| 1580 | normalising all terms.) | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1581 | |
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1582 |   The conversion combinator @{ML else_conv} tries out the 
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1583 | first one, and if it does not apply, tries the second. For example | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1584 | |
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 1585 |   @{ML_response_fake [display,gray]
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1586 | "let | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1587 |   val conv = Conv.rewr_conv @{thm true_conj1} else_conv Conv.all_conv
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1588 |   val ctrm1 = @{cterm \"True \<and> Q\"}
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1589 |   val ctrm2 = @{cterm \"P \<or> (True \<and> Q)\"}
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1590 | in | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1591 | (conv ctrm1, conv ctrm2) | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1592 | end" | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1593 | "(True \<and> Q \<equiv> Q, P \<or> True \<and> Q \<equiv> P \<or> True \<and> Q)"} | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1594 | |
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1595 |   Here the conversion of @{thm [source] true_conj1} only applies
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1596 | in the first case, but fails in the second. The whole conversion | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
150diff
changeset | 1597 |   does not fail, however, because the combinator @{ML Conv.else_conv} will then 
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1598 |   try out @{ML Conv.all_conv}, which always succeeds.
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1599 | |
| 149 | 1600 |   Apart from the function @{ML beta_conversion in Thm}, which is able to fully
 | 
| 1601 | beta-normalise a term, the conversions so far are restricted in that they | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1602 |   only apply to the outer-most level of a @{ML_type cterm}. In what follows we
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1603 |   will lift this restriction. The combinator @{ML Conv.arg_conv} will apply
 | 
| 149 | 1604 | the conversion to the first argument of an application, that is the term | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1605 |   must be of the form @{ML "t1 $ t2" for t1 t2} and the conversion is applied
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1606 |   to @{text t2}.  For example
 | 
| 139 | 1607 | |
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 1608 |   @{ML_response_fake [display,gray]
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1609 | "let | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1610 |   val conv = Conv.rewr_conv @{thm true_conj1}
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1611 |   val ctrm = @{cterm \"P \<or> (True \<and> Q)\"}
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1612 | in | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1613 | Conv.arg_conv conv ctrm | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1614 | end" | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1615 | "P \<or> (True \<and> Q) \<equiv> P \<or> Q"} | 
| 139 | 1616 | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1617 |   The reason for this behaviour is that @{text "(op \<or>)"} expects two
 | 
| 149 | 1618 |   arguments.  So the term must be of the form @{text "(Const \<dots> $ t1) $ t2"}. The
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1619 |   conversion is then applied to @{text "t2"} which in the example above
 | 
| 150 
cb39c41548bd
added a comment about Conv.fun_conv
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1620 |   stands for @{term "True \<and> Q"}. The function @{ML Conv.fun_conv} applies
 | 
| 
cb39c41548bd
added a comment about Conv.fun_conv
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1621 | the conversion to the first argument of an application. | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1622 | |
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1623 |   The function @{ML Conv.abs_conv} applies a conversion under an abstractions.
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1624 | For example: | 
| 139 | 1625 | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1626 |   @{ML_response_fake [display,gray]
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1627 | "let | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1628 |   val conv = K (Conv.rewr_conv @{thm true_conj1}) 
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1629 |   val ctrm = @{cterm \"\<lambda>P. True \<and> P \<and> Foo\"}
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1630 | in | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1631 |   Conv.abs_conv conv @{context} ctrm
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1632 | end" | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1633 | "\<lambda>P. True \<and> P \<and> Foo \<equiv> \<lambda>P. P \<and> Foo"} | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1634 | |
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1635 | The conversion that goes under an application is | 
| 148 | 1636 |   @{ML Conv.combination_conv}. It expects two conversions as arguments, 
 | 
| 1637 | each of which is applied to the corresponding ``branch'' of the application. | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1638 | |
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1639 | We can now apply all these functions in a | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1640 | conversion that recursively descends a term and applies a conversion in all | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1641 | possible positions. | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1642 | *} | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1643 | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1644 | ML %linenosgray{*fun all_true1_conv ctxt ctrm =
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1645 | case (Thm.term_of ctrm) of | 
| 142 | 1646 |     @{term "op \<and>"} $ @{term True} $ _ => 
 | 
| 1647 | (Conv.arg_conv (all_true1_conv ctxt) then_conv | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1648 |          Conv.rewr_conv @{thm true_conj1}) ctrm
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1649 | | _ $ _ => Conv.combination_conv | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1650 | (all_true1_conv ctxt) (all_true1_conv ctxt) ctrm | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1651 | | Abs _ => Conv.abs_conv (fn (_, ctxt) => all_true1_conv ctxt) ctxt ctrm | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1652 | | _ => Conv.all_conv ctrm*} | 
| 139 | 1653 | |
| 1654 | text {*
 | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1655 | This functions descends under applications (Line 6 and 7) and abstractions | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1656 |   (Line 8); and ``fires'' if the outer-most constant is an @{text "True \<and> \<dots>"}
 | 
| 149 | 1657 | (Lines 3-5); otherwise it leaves the term unchanged (Line 9). In Line 2 | 
| 1658 |   we need to transform the @{ML_type cterm} into a @{ML_type term} in order
 | |
| 1659 | to be able to pattern-match the term. To see this | |
| 1660 | conversion in action, consider the following example | |
| 139 | 1661 | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1662 | @{ML_response_fake [display,gray]
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1663 | "let | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1664 |   val ctxt = @{context}
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1665 |   val ctrm = @{cterm \"distinct [1, x] \<longrightarrow> True \<and> 1 \<noteq> x\"}
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1666 | in | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1667 | all_true1_conv ctxt ctrm | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1668 | end" | 
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 1669 | "distinct [1, x] \<longrightarrow> True \<and> 1 \<noteq> x \<equiv> distinct [1, x] \<longrightarrow> 1 \<noteq> x"} | 
| 139 | 1670 | |
| 149 | 1671 | where the conversion is applied ``deep'' inside the term. | 
| 1672 | ||
| 1673 | To see how much control you have about rewriting by using conversions, let us | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1674 | make the task a bit more complicated by rewriting according to the rule | 
| 149 | 1675 |   @{text true_conj1}, but only in the first arguments of @{term If}s. Then 
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1676 | the conversion should be as follows. | 
| 135 | 1677 | *} | 
| 1678 | ||
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1679 | ML{*fun if_true1_conv ctxt ctrm =
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1680 | case Thm.term_of ctrm of | 
| 142 | 1681 |     Const (@{const_name If}, _) $ _ =>
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1682 | Conv.arg_conv (all_true1_conv ctxt) ctrm | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1683 | | _ $ _ => Conv.combination_conv | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1684 | (if_true1_conv ctxt) (if_true1_conv ctxt) ctrm | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1685 | | Abs _ => Conv.abs_conv (fn (_, ctxt) => if_true1_conv ctxt) ctxt ctrm | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1686 | | _ => Conv.all_conv ctrm *} | 
| 135 | 1687 | |
| 139 | 1688 | text {*
 | 
| 149 | 1689 | Here is an example for this conversion: | 
| 139 | 1690 | |
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 1691 |   @{ML_response_fake [display,gray]
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1692 | "let | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1693 |   val ctxt = @{context}
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1694 | val ctrm = | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1695 |      @{cterm \"if P (True \<and> 1 \<noteq> 2) then True \<and> True else True \<and> False\"}
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1696 | in | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1697 | if_true1_conv ctxt ctrm | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1698 | end" | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1699 | "if P (True \<and> 1 \<noteq> 2) then True \<and> True else True \<and> False | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1700 | \<equiv> if P (1 \<noteq> 2) then True \<and> True else True \<and> False"} | 
| 135 | 1701 | *} | 
| 1702 | ||
| 1703 | text {*
 | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1704 |   So far we only applied conversions to @{ML_type cterm}s. Conversions can, however, 
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1705 |   also work on theorems using the function @{ML "Conv.fconv_rule"}. As an example, 
 | 
| 149 | 1706 |   consider the conversion @{ML all_true1_conv} and the lemma:
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1707 | *} | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1708 | |
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1709 | lemma foo_test: "P \<or> (True \<and> \<not>P)" by simp | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1710 | |
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1711 | text {*
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1712 | Using the conversion you can transform this theorem into a new theorem | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1713 | as follows | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1714 | |
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1715 |   @{ML_response_fake [display,gray]
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1716 |   "Conv.fconv_rule (all_true1_conv @{context}) @{thm foo_test}" 
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1717 | "?P \<or> \<not> ?P"} | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1718 | |
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1719 | Finally, conversions can also be turned into tactics and then applied to | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1720 |   goal states. This can be done with the help of the function @{ML CONVERSION},
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1721 | and also some predefined conversion combinator which traverse a goal | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1722 |   state. The combinators for the goal state are: @{ML Conv.params_conv} for
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1723 |   going under parameters (i.e.~where goals are of the form @{text "\<And>x. P \<Longrightarrow>
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1724 |   Q"}); the function @{ML Conv.prems_conv} for applying the conversion to all
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1725 |   premises of a goal, and @{ML Conv.concl_conv} for applying the conversion to
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1726 | the conclusion of a goal. | 
| 139 | 1727 | |
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 1728 |   Assume we want to apply @{ML all_true1_conv} only in the conclusion
 | 
| 149 | 1729 |   of the goal, and @{ML if_true1_conv} should only be applied to the premises.
 | 
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 1730 | Here is a tactic doing exactly that: | 
| 135 | 1731 | *} | 
| 1732 | ||
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 1733 | ML{*val true1_tac = CSUBGOAL (fn (goal, i) =>
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1734 | let | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1735 | val ctxt = ProofContext.init (Thm.theory_of_cterm goal) | 
| 139 | 1736 | in | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1737 | CONVERSION | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1738 | (Conv.params_conv ~1 (fn ctxt => | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
150diff
changeset | 1739 | (Conv.prems_conv ~1 (if_true1_conv ctxt) then_conv | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
150diff
changeset | 1740 | Conv.concl_conv ~1 (all_true1_conv ctxt))) ctxt) i | 
| 142 | 1741 | end)*} | 
| 1742 | ||
| 1743 | text {* 
 | |
| 148 | 1744 |   We call the conversions with the argument @{ML "~1"}. This is to 
 | 
| 1745 | analyse all parameters, premises and conclusions. If we call them with | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1746 |   a non-negative number, say @{text n}, then these conversions will 
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1747 |   only be called on @{text n} premises (similar for parameters and
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1748 | conclusions). To test the tactic, consider the proof | 
| 142 | 1749 | *} | 
| 139 | 1750 | |
| 142 | 1751 | lemma | 
| 1752 | "if True \<and> P then P else True \<and> False \<Longrightarrow> | |
| 148 | 1753 | (if True \<and> Q then True \<and> Q else P) \<longrightarrow> True \<and> (True \<and> Q)" | 
| 142 | 1754 | apply(tactic {* true1_tac 1 *})
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1755 | txt {* where the tactic yields the goal state
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1756 | |
| 149 | 1757 |        @{subgoals [display]}*}
 | 
| 142 | 1758 | (*<*)oops(*>*) | 
| 135 | 1759 | |
| 1760 | text {*
 | |
| 148 | 1761 |   As you can see, the premises are rewritten according to @{ML if_true1_conv}, while
 | 
| 1762 |   the conclusion according to @{ML all_true1_conv}.
 | |
| 1763 | ||
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1764 | To sum up this section, conversions are not as powerful as the simplifier | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1765 | and simprocs; the advantage of conversions, however, is that you never have | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 1766 | to worry about non-termination. | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1767 | |
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
150diff
changeset | 1768 |   \begin{exercise}\label{ex:addconversion}
 | 
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1769 | Write a tactic that does the same as the simproc in exercise | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1770 |   \ref{ex:addsimproc}, but is based in conversions. That means replace terms
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1771 |   of the form @{term "t\<^isub>1 + t\<^isub>2"} by their result. You can make
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1772 |   the same assumptions as in \ref{ex:addsimproc}. FIXME: the current solution
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1773 | requires some additional functions. | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1774 |   \end{exercise}
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1775 | |
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1776 |   \begin{exercise}
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1777 | Compare which way of rewriting such terms is more efficient. For this | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1778 | you might have to construct quite large terms. | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
150diff
changeset | 1779 |   \end{exercise}
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
150diff
changeset | 1780 | |
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1781 |   \begin{readmore}
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1782 |   See @{ML_file "Pure/conv.ML"} for more information about conversion combinators. 
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1783 |   Further conversions are defined in  @{ML_file "Pure/thm.ML"},
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1784 |   @{ML_file "Pure/drule.ML"} and @{ML_file "Pure/meta_simplifier.ML"}.
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1785 |   \end{readmore}
 | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
150diff
changeset | 1786 | |
| 135 | 1787 | *} | 
| 1788 | ||
| 1789 | ||
| 1790 | ||
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
150diff
changeset | 1791 | |
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1792 | section {* Structured Proofs (TBD) *}
 | 
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1793 | |
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1794 | text {* TBD *}
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1795 | |
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1796 | lemma True | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1797 | proof | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1798 | |
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1799 |   {
 | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1800 | fix A B C | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1801 | assume r: "A & B \<Longrightarrow> C" | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1802 | assume A B | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1803 | then have "A & B" .. | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1804 | then have C by (rule r) | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1805 | } | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1806 | |
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1807 |   {
 | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1808 | fix A B C | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1809 | assume r: "A & B \<Longrightarrow> C" | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1810 | assume A B | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1811 | note conjI [OF this] | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1812 | note r [OF this] | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1813 | } | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1814 | oops | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1815 | |
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1816 | ML {* fun prop ctxt s =
 | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1817 | Thm.cterm_of (ProofContext.theory_of ctxt) (Syntax.read_prop ctxt s) *} | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1818 | |
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1819 | ML {* 
 | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1820 |   val ctxt0 = @{context};
 | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1821 | val ctxt = ctxt0; | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1822 | val (_, ctxt) = Variable.add_fixes ["A", "B", "C"] ctxt; | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1823 | val ([r], ctxt) = Assumption.add_assumes [prop ctxt "A & B \<Longrightarrow> C"] ctxt; | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1824 | val (this, ctxt) = Assumption.add_assumes [prop ctxt "A", prop ctxt "B"] ctxt; | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1825 |   val this = [@{thm conjI} OF this]; 
 | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1826 | val this = r OF this; | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1827 | val this = Assumption.export false ctxt ctxt0 this | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1828 | val this = Variable.export ctxt ctxt0 [this] | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 1829 | *} | 
| 93 | 1830 | |
| 1831 | ||
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
99diff
changeset | 1832 | |
| 139 | 1833 | end |