| author | Christian Urban <christian dot urban at kcl dot ac dot uk> | 
| Sun, 15 Dec 2013 23:49:05 +0000 | |
| changeset 552 | 82c482467d75 | 
| parent 551 | be361e980acf | 
| child 554 | 638ed040e6f8 | 
| permissions | -rw-r--r-- | 
| 93 | 1 | theory Tactical | 
| 441 | 2 | imports Base First_Steps | 
| 93 | 3 | begin | 
| 4 | ||
| 5 | chapter {* Tactical Reasoning\label{chp:tactical} *}
 | |
| 6 | ||
| 7 | text {*
 | |
| 506 | 8 |    \begin{flushright}
 | 
| 9 |   {\em
 | |
| 10 | ``The first thing I would say is that when you write a program, think of\\ | |
| 11 | it primarily as a work of literature. You're trying to write something\\ | |
| 12 | that human beings are going to read. Don't think of it primarily as\\ | |
| 13 | something a computer is going to follow. The more effective you are\\ | |
| 14 | at making your program readable, the more effective it's going to\\ | |
| 15 | be: You'll understand it today, you'll understand it next week, and\\ | |
| 16 | your successors who are going to maintain and modify it will\\ | |
| 17 | understand it.''}\\[1ex] | |
| 18 | Donald E.~Knuth, from an interview in Dr.~Dobb's Journal, 1996. | |
| 19 |   \end{flushright}
 | |
| 20 | ||
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 21 | One of the main reason for descending to the ML-level of Isabelle is to be | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 22 | able to implement automatic proof procedures. Such proof procedures can | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 23 | considerably lessen the burden of manual reasoning. They are centred around | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 24 | the idea of refining a goal state using tactics. This is similar to the | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 25 |   \isacommand{apply}-style reasoning at the user-level, where goals are
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 26 | modified in a sequence of proof steps until all of them are discharged. | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 27 | In this chapter we will explain simple tactics and how to combine them using | 
| 363 | 28 | tactic combinators. We also describe the simplifier, simprocs and conversions. | 
| 93 | 29 | *} | 
| 30 | ||
| 500 | 31 | section {* Basics of Reasoning with Tactics\label{sec:basictactics}*}
 | 
| 93 | 32 | |
| 33 | 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 | 34 |   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 | 35 | into ML. Suppose the following proof. | 
| 93 | 36 | *} | 
| 37 | ||
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 38 | lemma disj_swap: | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 39 | shows "P \<or> Q \<Longrightarrow> Q \<or> P" | 
| 93 | 40 | apply(erule disjE) | 
| 41 | apply(rule disjI2) | |
| 42 | apply(assumption) | |
| 43 | apply(rule disjI1) | |
| 44 | apply(assumption) | |
| 45 | done | |
| 46 | ||
| 451 
fc074e669f9f
disabled foobar_prove; updated to new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
449diff
changeset | 47 | |
| 
fc074e669f9f
disabled foobar_prove; updated to new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
449diff
changeset | 48 | |
| 93 | 49 | text {*
 | 
| 50 | This proof translates to the following ML-code. | |
| 51 | ||
| 52 | @{ML_response_fake [display,gray]
 | |
| 53 | "let | |
| 54 |   val ctxt = @{context}
 | |
| 55 |   val goal = @{prop \"P \<or> Q \<Longrightarrow> Q \<or> P\"}
 | |
| 56 | in | |
| 57 | Goal.prove ctxt [\"P\", \"Q\"] [] goal | |
| 363 | 58 |    (fn _ =>  etac @{thm disjE} 1
 | 
| 59 |              THEN rtac @{thm disjI2} 1
 | |
| 60 | THEN atac 1 | |
| 61 |              THEN rtac @{thm disjI1} 1
 | |
| 62 | THEN atac 1) | |
| 93 | 63 | end" "?P \<or> ?Q \<Longrightarrow> ?Q \<or> ?P"} | 
| 64 | ||
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 65 |   To start the proof, the function @{ML_ind prove in Goal} sets up a goal
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 66 |   state for proving the goal @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"}. We can give this
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 67 | function some assumptions in the third argument (there are no assumption in | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 68 | the proof at hand). The second argument stands for a list of variables | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 69 | (given as strings). This list contains the variables that will be turned | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 70 |   into schematic variables once the goal is proved (in our case @{text P} and
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 71 |   @{text Q}). The last argument is the tactic that proves the goal. This
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 72 | tactic can make use of the local assumptions (there are none in this | 
| 363 | 73 |   example). The tactics @{ML_ind etac in Tactic}, @{ML_ind rtac in Tactic} and
 | 
| 74 |   @{ML_ind atac in Tactic} in the code above correspond roughly to @{text
 | |
| 75 |   erule}, @{text rule} and @{text assumption}, respectively. The combinator
 | |
| 76 |   @{ML THEN} strings the tactics together.
 | |
| 93 | 77 | |
| 437 | 78 |   TBD: Write something about @{ML_ind prove_multi in Goal}.
 | 
| 79 | ||
| 93 | 80 |   \begin{readmore}
 | 
| 318 
efb5fff99c96
split up the first-steps section into two chapters
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 81 |   To learn more about the function @{ML_ind prove in Goal} see
 | 
| 
efb5fff99c96
split up the first-steps section into two chapters
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 82 |   \isccite{sec:results} and the file @{ML_file "Pure/goal.ML"}.  See @{ML_file
 | 
| 289 
08ffafe2585d
adapted to changes in Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
288diff
changeset | 83 |   "Pure/tactic.ML"} and @{ML_file "Pure/tactical.ML"} for the code of basic
 | 
| 99 | 84 | tactics and tactic combinators; see also Chapters 3 and 4 in the old | 
| 318 
efb5fff99c96
split up the first-steps section into two chapters
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 85 | Isabelle Reference Manual, and Chapter 3 in the Isabelle/Isar Implementation | 
| 
efb5fff99c96
split up the first-steps section into two chapters
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 86 | Manual. | 
| 93 | 87 |   \end{readmore}
 | 
| 88 | ||
| 89 | ||
| 436 
373f99b1221a
added something about raw_tactic
 Christian Urban <urbanc@in.tum.de> parents: 
424diff
changeset | 90 |   \index{tactic@ {\tt\slshape{}tactic}}
 | 
| 
373f99b1221a
added something about raw_tactic
 Christian Urban <urbanc@in.tum.de> parents: 
424diff
changeset | 91 |   \index{raw\_tactic@ {\tt\slshape{}raw\_tactic}}
 | 
| 318 
efb5fff99c96
split up the first-steps section into two chapters
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 92 | During the development of automatic proof procedures, you will often find it | 
| 
efb5fff99c96
split up the first-steps section into two chapters
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 93 | necessary to test a tactic on examples. This can be conveniently done with | 
| 
efb5fff99c96
split up the first-steps section into two chapters
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 94 |   the command \isacommand{apply}@{text "(tactic \<verbopen> \<dots> \<verbclose>)"}. 
 | 
| 93 | 95 | Consider the following sequence of tactics | 
| 318 
efb5fff99c96
split up the first-steps section into two chapters
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 96 | |
| 93 | 97 | *} | 
| 98 | ||
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 99 | ML %grayML{*val foo_tac = 
 | 
| 93 | 100 |       (etac @{thm disjE} 1
 | 
| 101 |        THEN rtac @{thm disjI2} 1
 | |
| 102 | THEN atac 1 | |
| 103 |        THEN rtac @{thm disjI1} 1
 | |
| 104 | THEN atac 1)*} | |
| 105 | ||
| 106 | text {* and the Isabelle proof: *}
 | |
| 107 | ||
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 108 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 109 | shows "P \<or> Q \<Longrightarrow> Q \<or> P" | 
| 93 | 110 | apply(tactic {* foo_tac *})
 | 
| 111 | done | |
| 112 | ||
| 113 | text {*
 | |
| 104 | 114 |   By using @{text "tactic \<verbopen> \<dots> \<verbclose>"} you can call from the 
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 115 |   user-level of Isabelle the tactic @{ML foo_tac} or 
 | 
| 436 
373f99b1221a
added something about raw_tactic
 Christian Urban <urbanc@in.tum.de> parents: 
424diff
changeset | 116 | any other function that returns a tactic. There are some goal transformation | 
| 
373f99b1221a
added something about raw_tactic
 Christian Urban <urbanc@in.tum.de> parents: 
424diff
changeset | 117 |   that are performed by @{text "tactic"}. They can be avoided by using 
 | 
| 
373f99b1221a
added something about raw_tactic
 Christian Urban <urbanc@in.tum.de> parents: 
424diff
changeset | 118 |   @{text "raw_tactic"} instead.
 | 
| 93 | 119 | |
| 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 | 120 |   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 | 121 |   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 | 122 | 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 | 123 |   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 | 124 | of goals is sometimes wanted, but usually it is not. To avoid the explicit numbering, | 
| 238 
29787dcf7b2e
added something about TRY and TRYALL
 Christian Urban <urbanc@in.tum.de> parents: 
232diff
changeset | 125 | you can write | 
| 93 | 126 | *} | 
| 127 | ||
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 128 | ML %grayML{*val foo_tac' = 
 | 
| 93 | 129 |       (etac @{thm disjE} 
 | 
| 130 |        THEN' rtac @{thm disjI2} 
 | |
| 131 | THEN' atac | |
| 132 |        THEN' rtac @{thm disjI1} 
 | |
| 238 
29787dcf7b2e
added something about TRY and TRYALL
 Christian Urban <urbanc@in.tum.de> parents: 
232diff
changeset | 133 |        THEN' atac)*}text_raw{*\label{tac:footacprime}*}
 | 
| 93 | 134 | |
| 135 | text {* 
 | |
| 363 | 136 |   where @{ML_ind THEN' in Tactical} is used instead of @{ML THEN in
 | 
| 137 |   Tactical}. (For most combinators that combine tactics---@{ML THEN} is only
 | |
| 138 |   one such combinator---a ``primed'' version exists.)  With @{ML foo_tac'} you
 | |
| 139 | can give the number for the subgoal explicitly when the tactic is called. So | |
| 140 | in the next proof you can first discharge the second subgoal, and | |
| 141 | subsequently the first. | |
| 93 | 142 | *} | 
| 143 | ||
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 144 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 145 | shows "P1 \<or> Q1 \<Longrightarrow> Q1 \<or> P1" | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 146 | and "P2 \<or> Q2 \<Longrightarrow> Q2 \<or> P2" | 
| 93 | 147 | apply(tactic {* foo_tac' 2 *})
 | 
| 148 | apply(tactic {* foo_tac' 1 *})
 | |
| 149 | done | |
| 150 | ||
| 151 | 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 | 152 | This kind of addressing is more difficult to achieve when the goal is | 
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 153 | hard-coded inside the tactic. | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 154 | |
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 155 |   The tactics @{ML foo_tac} and @{ML foo_tac'} are very specific for analysing
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 156 |   goals being only of the form @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"}. If the goal is not of
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 157 |   this form, then these tactics return the error message:\footnote{To be
 | 
| 363 | 158 | precise, tactics do not produce this error message; the message originates from the | 
| 159 |   \isacommand{apply} wrapper that calls the tactic.}
 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 160 | |
| 99 | 161 | |
| 162 |   \begin{isabelle}
 | |
| 163 |   @{text "*** empty result sequence -- proof command failed"}\\
 | |
| 164 |   @{text "*** At command \"apply\"."}
 | |
| 165 |   \end{isabelle}
 | |
| 166 | ||
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 167 | This means they failed. The reason for this error message is that tactics | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 168 | are functions mapping a goal state to a (lazy) sequence of successor | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 169 | states. Hence the type of a tactic is: | 
| 109 | 170 | *} | 
| 93 | 171 | |
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 172 | ML %grayML{*type tactic = thm -> thm Seq.seq*}
 | 
| 93 | 173 | |
| 109 | 174 | text {*
 | 
| 175 | By convention, if a tactic fails, then it should return the empty sequence. | |
| 176 | Therefore, if you write your own tactics, they should not raise exceptions | |
| 177 | willy-nilly; only in very grave failure situations should a tactic raise the | |
| 178 |   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 | 179 | |
| 363 | 180 |   The simplest tactics are @{ML_ind no_tac in Tactical} and 
 | 
| 181 |   @{ML_ind all_tac in Tactical}. The first returns the empty sequence and 
 | |
| 182 | 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 | 183 | *} | 
| 
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 | 184 | |
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 185 | ML %grayML{*fun no_tac thm = Seq.empty*}
 | 
| 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 | 186 | |
| 
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 | 187 | 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 | 188 |   which means @{ML no_tac} always fails. The second returns the given theorem wrapped 
 | 
| 173 
d820cb5873ea
used latex package boxedminipage
 Christian Urban <urbanc@in.tum.de> parents: 
172diff
changeset | 189 | 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 | 190 | *} | 
| 
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 | 191 | |
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 192 | ML %grayML{*fun all_tac thm = Seq.single 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 | 193 | |
| 
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 | 194 | text {*
 | 
| 109 | 195 |   which means @{ML all_tac} always succeeds, but also does not make any progress 
 | 
| 196 | with the proof. | |
| 93 | 197 | |
| 109 | 198 | The lazy list of possible successor goal states shows through at the user-level | 
| 99 | 199 |   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 | 200 | 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 | 201 |   @{ML foo_tac'}: either using the first assumption or the second.
 | 
| 93 | 202 | *} | 
| 203 | ||
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 204 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 205 | shows "\<lbrakk>P \<or> Q; P \<or> Q\<rbrakk> \<Longrightarrow> Q \<or> P" | 
| 93 | 206 | apply(tactic {* foo_tac' 1 *})
 | 
| 207 | back | |
| 208 | done | |
| 209 | ||
| 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 | 210 | |
| 93 | 211 | text {*
 | 
| 99 | 212 |   By using \isacommand{back}, we construct the proof that uses the
 | 
| 109 | 213 | second assumption. While in the proof above, it does not really matter which | 
| 214 | assumption is used, in more interesting cases provability might depend | |
| 215 | on exploring different possibilities. | |
| 99 | 216 | |
| 93 | 217 |   \begin{readmore}
 | 
| 218 |   See @{ML_file "Pure/General/seq.ML"} for the implementation of lazy
 | |
| 109 | 219 | sequences. In day-to-day Isabelle programming, however, one rarely | 
| 220 | constructs sequences explicitly, but uses the predefined tactics and | |
| 221 | tactic combinators instead. | |
| 93 | 222 |   \end{readmore}
 | 
| 223 | ||
| 104 | 224 | It might be surprising that tactics, which transform | 
| 109 | 225 | 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 | 226 | (sequences). The surprise resolves by knowing that every | 
| 104 | 227 | 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 | 228 |   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 | 229 | 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 | 230 | *} | 
| 
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 | 231 | |
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 232 | ML %grayML{*fun my_print_tac ctxt thm =
 | 
| 132 | 233 | let | 
| 440 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
437diff
changeset | 234 | val _ = tracing (Pretty.string_of (pretty_thm_no_vars ctxt thm)) | 
| 132 | 235 | in | 
| 236 | Seq.single thm | |
| 237 | 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 | 238 | |
| 363 | 239 | text {*
 | 
| 240 | which prints out the given theorem (using the string-function defined in | |
| 241 |   Section~\ref{sec:printing}) and then behaves like @{ML all_tac}. With this
 | |
| 242 | tactic we are in the position to inspect every goal state in a proof. For | |
| 243 |   this consider the proof in Figure~\ref{fig:goalstates}: as can be seen,
 | |
| 244 | internally every goal state is an implication of the form | |
| 245 | ||
| 551 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 246 |   @{text[display] "A\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^sub>n \<Longrightarrow> #C"}
 | 
| 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 247 | |
| 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 248 |   where @{term C} is the goal to be proved and the @{term "A\<^sub>i"} are
 | 
| 363 | 249 | the subgoals. So after setting up the lemma, the goal state is always of the | 
| 250 |   form @{text "C \<Longrightarrow> #C"}; when the proof is finished we are left with @{text
 | |
| 251 |   "#C"}. Since the goal @{term C} can potentially be an implication, there is a
 | |
| 252 | ``protector'' wrapped around it (the wrapper is the outermost constant | |
| 253 |   @{text "Const (\"prop\", bool \<Rightarrow> bool)"}; in the figure we make it visible
 | |
| 254 |   as a @{text #}). This wrapper prevents that premises of @{text C} are
 | |
| 255 | misinterpreted as open subgoals. While tactics can operate on the subgoals | |
| 551 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 256 |   (the @{text "A\<^sub>i"} above), they are expected to leave the conclusion
 | 
| 363 | 257 |   @{term C} intact, with the exception of possibly instantiating schematic
 | 
| 258 | variables. This instantiations of schematic variables can be observed | |
| 259 | on the user level. Have a look at the following definition and proof. | |
| 260 | *} | |
| 261 | ||
| 109 | 262 | text_raw {*
 | 
| 263 |   \begin{figure}[p]
 | |
| 173 
d820cb5873ea
used latex package boxedminipage
 Christian Urban <urbanc@in.tum.de> parents: 
172diff
changeset | 264 |   \begin{boxedminipage}{\textwidth}
 | 
| 109 | 265 |   \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 | 266 | *} | 
| 303 
05e6a33edef6
added an antiquotation for printing the raw proof state; polished the example about proof state
 Christian Urban <urbanc@in.tum.de> parents: 
302diff
changeset | 267 | notation (output) "prop"  ("#_" [1000] 1000)
 | 
| 
05e6a33edef6
added an antiquotation for printing the raw proof state; polished the example about proof state
 Christian Urban <urbanc@in.tum.de> parents: 
302diff
changeset | 268 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 269 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 270 | 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 | 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:\\ | |
| 303 
05e6a33edef6
added an antiquotation for printing the raw proof state; polished the example about proof state
 Christian Urban <urbanc@in.tum.de> parents: 
302diff
changeset | 281 |       @{raw_goal_state}
 | 
| 109 | 282 |       \end{tabular}}
 | 
| 283 |       \end{minipage}\medskip
 | |
| 93 | 284 | *} | 
| 285 | ||
| 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 | 286 | 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 | 287 | 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 | 288 | |
| 109 | 289 | 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 | 290 |       @{subgoals [display]} 
 | 
| 109 | 291 |       \end{minipage}\medskip
 | 
| 292 | ||
| 293 |       \begin{minipage}{\textwidth}
 | |
| 294 |       \small\colorbox{gray!20}{
 | |
| 295 |       \begin{tabular}{@ {}l@ {}}
 | |
| 296 | internal goal state:\\ | |
| 303 
05e6a33edef6
added an antiquotation for printing the raw proof state; polished the example about proof state
 Christian Urban <urbanc@in.tum.de> parents: 
302diff
changeset | 297 |       @{raw_goal_state}
 | 
| 109 | 298 |       \end{tabular}}
 | 
| 299 |       \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 | 300 | *} | 
| 
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 | 301 | |
| 
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 | 302 | 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 | 303 | 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 | 304 | |
| 109 | 305 | 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 | 306 |       @{subgoals [display]} 
 | 
| 109 | 307 |       \end{minipage}\medskip
 | 
| 308 | ||
| 309 |       \begin{minipage}{\textwidth}
 | |
| 310 |       \small\colorbox{gray!20}{
 | |
| 311 |       \begin{tabular}{@ {}l@ {}}
 | |
| 312 | internal goal state:\\ | |
| 303 
05e6a33edef6
added an antiquotation for printing the raw proof state; polished the example about proof state
 Christian Urban <urbanc@in.tum.de> parents: 
302diff
changeset | 313 |       @{raw_goal_state}
 | 
| 109 | 314 |       \end{tabular}}
 | 
| 315 |       \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 | 316 | *} | 
| 
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 | 317 | |
| 
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 | 318 | 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 | 319 | 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 | 320 | |
| 109 | 321 | 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 | 322 |       @{subgoals [display]} 
 | 
| 109 | 323 |       \end{minipage}\medskip 
 | 
| 324 | ||
| 325 |       \begin{minipage}{\textwidth}
 | |
| 326 |       \small\colorbox{gray!20}{
 | |
| 327 |       \begin{tabular}{@ {}l@ {}}
 | |
| 328 | internal goal state:\\ | |
| 303 
05e6a33edef6
added an antiquotation for printing the raw proof state; polished the example about proof state
 Christian Urban <urbanc@in.tum.de> parents: 
302diff
changeset | 329 |       @{raw_goal_state}
 | 
| 109 | 330 |       \end{tabular}}
 | 
| 331 |       \end{minipage}\medskip   
 | |
| 332 | *} | |
| 303 
05e6a33edef6
added an antiquotation for printing the raw proof state; polished the example about proof state
 Christian Urban <urbanc@in.tum.de> parents: 
302diff
changeset | 333 | (*<*)oops(*>*) | 
| 109 | 334 | text_raw {*  
 | 
| 335 |   \end{isabelle}
 | |
| 173 
d820cb5873ea
used latex package boxedminipage
 Christian Urban <urbanc@in.tum.de> parents: 
172diff
changeset | 336 |   \end{boxedminipage}
 | 
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 337 |   \caption{The figure shows an Isabelle snippet of a proof where each
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 338 |   intermediate goal state is printed by the Isabelle system and by @{ML
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 339 | my_print_tac}. The latter shows the goal state as represented internally | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 340 | (highlighted boxes). This tactic shows that every goal state in Isabelle is | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 341 |   represented by a theorem: when you start the proof of \mbox{@{text "\<lbrakk>A; B\<rbrakk> \<Longrightarrow>
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 342 |   A \<and> B"}} the theorem is @{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B) \<Longrightarrow> #(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}; when
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 343 |   you finish the proof the theorem is @{text "#(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and>
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 344 |   B)"}.\label{fig:goalstates}}
 | 
| 109 | 345 |   \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 | 346 | *} | 
| 
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 | 347 | |
| 358 | 348 | definition | 
| 349 | EQ_TRUE | |
| 350 | where | |
| 351 | "EQ_TRUE X \<equiv> (X = True)" | |
| 352 | ||
| 422 | 353 | schematic_lemma test: | 
| 358 | 354 | shows "EQ_TRUE ?X" | 
| 363 | 355 | unfolding EQ_TRUE_def | 
| 356 | by (rule refl) | |
| 358 | 357 | |
| 358 | text {*
 | |
| 422 | 359 |   By using \isacommand{schematic\_lemma} it is possible to prove lemmas including
 | 
| 360 |   meta-variables on the user level. However, the proved theorem is not @{text "EQ_TRUE ?X"}, 
 | |
| 361 |   as one might expect, but @{thm test}. We can test this with:
 | |
| 358 | 362 | |
| 363 |   \begin{isabelle}
 | |
| 364 |   \isacommand{thm}~@{thm [source] test}\\
 | |
| 365 |   @{text ">"}~@{thm test}
 | |
| 366 |   \end{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 | 367 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 368 |   The reason for this result is that the schematic variable @{text "?X"} 
 | 
| 363 | 369 |   is instantiated inside the proof to be @{term True} and then the 
 | 
| 370 | instantiation propagated to the ``outside''. | |
| 359 | 371 | |
| 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 | 372 |   \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 | 373 |   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 | 374 |   \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 | 375 | |
| 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 | 376 | *} | 
| 
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 | 377 | |
| 422 | 378 | |
| 194 | 379 | section {* Simple Tactics\label{sec:simpletacs} *}
 | 
| 93 | 380 | |
| 99 | 381 | text {*
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 382 | In this section we will introduce more of the simple tactics in Isabelle. The | 
| 363 | 383 |   first one is @{ML_ind print_tac in Tactical}, which is quite useful 
 | 
| 173 
d820cb5873ea
used latex package boxedminipage
 Christian Urban <urbanc@in.tum.de> parents: 
172diff
changeset | 384 | for low-level debugging of tactics. It just prints out a message and the current | 
| 
d820cb5873ea
used latex package boxedminipage
 Christian Urban <urbanc@in.tum.de> parents: 
172diff
changeset | 385 |   goal state. Unlike @{ML my_print_tac} shown earlier, it prints the goal state 
 | 
| 
d820cb5873ea
used latex package boxedminipage
 Christian Urban <urbanc@in.tum.de> parents: 
172diff
changeset | 386 | as the user would see it. For example, processing the proof | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 387 | *} | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 388 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 389 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 390 | shows "False \<Longrightarrow> True" | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 391 | apply(tactic {* print_tac "foo message" *})
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 392 | txt{*gives:
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 393 |      \begin{isabelle}
 | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 394 |      @{text "foo message"}\\[3mm] 
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 395 |      @{prop "False \<Longrightarrow> True"}\\
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 396 |      @{text " 1. False \<Longrightarrow> True"}\\
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 397 |      \end{isabelle}
 | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 398 | *} | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 399 | (*<*)oops(*>*) | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 400 | |
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 401 | text {*
 | 
| 363 | 402 | A simple tactic for easy discharge of any proof obligations, even difficult ones, is | 
| 403 |   @{ML_ind cheat_tac in Skip_Proof} in the structure @{ML_struct Skip_Proof}. 
 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 404 |   This tactic corresponds to the Isabelle command \isacommand{sorry} and is 
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 405 | sometimes useful during the development of tactics. | 
| 192 | 406 | *} | 
| 407 | ||
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 408 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 409 | shows "False" and "Goldbach_conjecture" | 
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 410 | apply(tactic {* Skip_Proof.cheat_tac 1 *})
 | 
| 192 | 411 | txt{*\begin{minipage}{\textwidth}
 | 
| 412 |      @{subgoals [display]}
 | |
| 413 |      \end{minipage}*}
 | |
| 414 | (*<*)oops(*>*) | |
| 415 | ||
| 416 | text {*
 | |
| 363 | 417 |   Another simple tactic is the function @{ML_ind atac in Tactic}, which, as shown 
 | 
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 418 | earlier, corresponds to the assumption method. | 
| 99 | 419 | *} | 
| 420 | ||
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 421 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 422 | shows "P \<Longrightarrow> P" | 
| 93 | 423 | apply(tactic {* atac 1 *})
 | 
| 109 | 424 | txt{*\begin{minipage}{\textwidth}
 | 
| 425 |      @{subgoals [display]}
 | |
| 426 |      \end{minipage}*}
 | |
| 427 | (*<*)oops(*>*) | |
| 93 | 428 | |
| 99 | 429 | text {*
 | 
| 363 | 430 |   Similarly, @{ML_ind rtac in Tactic}, @{ML_ind dtac in Tactic}, @{ML_ind etac
 | 
| 431 |   in Tactic} and @{ML_ind ftac in Tactic} correspond (roughly) to @{text
 | |
| 432 |   rule}, @{text drule}, @{text erule} and @{text frule}, respectively. Each of
 | |
| 433 | them takes a theorem as argument and attempts to apply it to a goal. Below | |
| 434 | are three self-explanatory examples. | |
| 99 | 435 | *} | 
| 436 | ||
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 437 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 438 | shows "P \<and> Q" | 
| 93 | 439 | apply(tactic {* rtac @{thm conjI} 1 *})
 | 
| 104 | 440 | txt{*\begin{minipage}{\textwidth}
 | 
| 441 |      @{subgoals [display]}
 | |
| 442 |      \end{minipage}*}
 | |
| 93 | 443 | (*<*)oops(*>*) | 
| 444 | ||
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 445 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 446 | shows "P \<and> Q \<Longrightarrow> False" | 
| 93 | 447 | apply(tactic {* etac @{thm conjE} 1 *})
 | 
| 104 | 448 | txt{*\begin{minipage}{\textwidth}
 | 
| 449 |      @{subgoals [display]}
 | |
| 450 |      \end{minipage}*}
 | |
| 93 | 451 | (*<*)oops(*>*) | 
| 452 | ||
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 453 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 454 | shows "False \<and> True \<Longrightarrow> False" | 
| 93 | 455 | apply(tactic {* dtac @{thm conjunct2} 1 *})
 | 
| 104 | 456 | txt{*\begin{minipage}{\textwidth}
 | 
| 457 |      @{subgoals [display]}
 | |
| 458 |      \end{minipage}*}
 | |
| 93 | 459 | (*<*)oops(*>*) | 
| 460 | ||
| 461 | text {*
 | |
| 369 | 462 |   The function @{ML_ind resolve_tac in Tactic} is similar to @{ML rtac}, except that it
 | 
| 363 | 463 | expects a list of theorems as argument. From this list it will apply the | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 464 | 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 | 465 | explored via the lazy sequences mechanism). Given the code | 
| 93 | 466 | *} | 
| 467 | ||
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 468 | ML %grayML{*val resolve_xmp_tac = resolve_tac [@{thm impI}, @{thm conjI}]*}
 | 
| 99 | 469 | |
| 470 | text {*
 | |
| 471 |   an example for @{ML resolve_tac} is the following proof where first an outermost 
 | |
| 472 | implication is analysed and then an outermost conjunction. | |
| 473 | *} | |
| 474 | ||
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 475 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 476 | shows "C \<longrightarrow> (A \<and> B)" | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 477 | and "(A \<longrightarrow> B) \<and> C" | 
| 238 
29787dcf7b2e
added something about TRY and TRYALL
 Christian Urban <urbanc@in.tum.de> parents: 
232diff
changeset | 478 | apply(tactic {* resolve_xmp_tac 1 *})
 | 
| 
29787dcf7b2e
added something about TRY and TRYALL
 Christian Urban <urbanc@in.tum.de> parents: 
232diff
changeset | 479 | apply(tactic {* resolve_xmp_tac 2 *})
 | 
| 104 | 480 | txt{*\begin{minipage}{\textwidth}
 | 
| 481 |      @{subgoals [display]} 
 | |
| 482 |      \end{minipage}*}
 | |
| 99 | 483 | (*<*)oops(*>*) | 
| 484 | ||
| 485 | text {* 
 | |
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
184diff
changeset | 486 | Similar versions taking a list of theorems exist for the tactics | 
| 363 | 487 |   @{ML dtac} (@{ML_ind dresolve_tac in Tactic}), @{ML etac} 
 | 
| 488 |   (@{ML_ind eresolve_tac in Tactic}) and so on.
 | |
| 489 | ||
| 490 |   Another simple tactic is @{ML_ind cut_facts_tac in Tactic}. It inserts a
 | |
| 491 | list of theorems into the assumptions of the current goal state. Below we | |
| 492 |   will insert the definitions for the constants @{term True} and @{term
 | |
| 493 | False}. So | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 494 | *} | 
| 99 | 495 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 496 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 497 | shows "True \<noteq> False" | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 498 | apply(tactic {* cut_facts_tac [@{thm True_def}, @{thm False_def}] 1 *})
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 499 | txt{*produces the goal state
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 500 |      \begin{isabelle}
 | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 501 |      @{subgoals [display]} 
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 502 |      \end{isabelle}*}
 | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 503 | (*<*)oops(*>*) | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 504 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 505 | text {*
 | 
| 109 | 506 | Often proofs on the ML-level involve elaborate operations on assumptions and | 
| 507 |   @{text "\<And>"}-quantified variables. To do such operations using the basic tactics 
 | |
| 128 | 508 | shown so far is very unwieldy and brittle. Some convenience and | 
| 363 | 509 |   safety is provided by the functions @{ML_ind FOCUS in Subgoal} and 
 | 
| 510 |   @{ML_ind SUBPROOF in Subgoal}. These tactics fix the parameters 
 | |
| 298 | 511 | and bind the various components of a goal state to a record. | 
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 512 |   To see what happens, suppose the function defined in Figure~\ref{fig:sptac}, which
 | 
| 363 | 513 | takes a record and just prints out the contents of this record. Then consider | 
| 514 | the proof: | |
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 515 | *} | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 516 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 517 | |
| 99 | 518 | text_raw{*
 | 
| 173 
d820cb5873ea
used latex package boxedminipage
 Christian Urban <urbanc@in.tum.de> parents: 
172diff
changeset | 519 | \begin{figure}[t]
 | 
| 177 | 520 | \begin{minipage}{\textwidth}
 | 
| 99 | 521 | \begin{isabelle}
 | 
| 522 | *} | |
| 294 
ee9d53fbb56b
made changes for SUBPROOF and sat_tac
 Christian Urban <urbanc@in.tum.de> parents: 
292diff
changeset | 523 | |
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 524 | ML %grayML{*fun foc_tac {prems, params, asms, concl, context, schematics} = 
 | 
| 132 | 525 | let | 
| 440 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
437diff
changeset | 526 | fun assgn1 f1 f2 xs = | 
| 541 
96d10631eec2
added readme and fixed output in Subgoal.FOCUS section
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
517diff
changeset | 527 | let | 
| 
96d10631eec2
added readme and fixed output in Subgoal.FOCUS section
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
517diff
changeset | 528 | val out = map (fn (x, y) => Pretty.enum ":=" "" "" [f1 x, f2 y]) xs | 
| 
96d10631eec2
added readme and fixed output in Subgoal.FOCUS section
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
517diff
changeset | 529 | in | 
| 
96d10631eec2
added readme and fixed output in Subgoal.FOCUS section
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
517diff
changeset | 530 | Pretty.list "" "" out | 
| 
96d10631eec2
added readme and fixed output in Subgoal.FOCUS section
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
517diff
changeset | 531 | end | 
| 440 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
437diff
changeset | 532 | |
| 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
437diff
changeset | 533 | fun assgn2 f xs = assgn1 f f xs | 
| 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
437diff
changeset | 534 | |
| 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
437diff
changeset | 535 | val pps = map (fn (s, pp) => Pretty.block [Pretty.str s, pp]) | 
| 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
437diff
changeset | 536 |     [("params: ", assgn1 Pretty.str (pretty_cterm context) params),
 | 
| 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
437diff
changeset | 537 |      ("assumptions: ", pretty_cterms context asms),
 | 
| 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
437diff
changeset | 538 |      ("conclusion: ", pretty_cterm context concl),
 | 
| 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
437diff
changeset | 539 |      ("premises: ", pretty_thms_no_vars context prems),   
 | 
| 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
437diff
changeset | 540 |      ("schematics: ", assgn2 (pretty_cterm context) (snd schematics))]
 | 
| 132 | 541 | in | 
| 440 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
437diff
changeset | 542 | tracing (Pretty.string_of (Pretty.chunks pps)); all_tac | 
| 132 | 543 | end*} | 
| 299 
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
298diff
changeset | 544 | |
| 99 | 545 | text_raw{*
 | 
| 546 | \end{isabelle}
 | |
| 177 | 547 | \end{minipage}
 | 
| 298 | 548 | \caption{A function that prints out the various parameters provided by 
 | 
| 299 
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
298diff
changeset | 549 |  @{ML FOCUS in Subgoal} and @{ML SUBPROOF}. It uses the functions defined 
 | 
| 
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
298diff
changeset | 550 |   in Section~\ref{sec:printing} for extracting strings from @{ML_type cterm}s 
 | 
| 
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
298diff
changeset | 551 |   and @{ML_type thm}s.\label{fig:sptac}}
 | 
| 99 | 552 | \end{figure}
 | 
| 553 | *} | |
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 554 | |
| 422 | 555 | schematic_lemma | 
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 556 | shows "\<And>x y. A x y \<Longrightarrow> B y x \<longrightarrow> C (?z y) x" | 
| 299 
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
298diff
changeset | 557 | apply(tactic {* Subgoal.FOCUS foc_tac @{context} 1 *})
 | 
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 558 | |
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 559 | txt {*
 | 
| 109 | 560 | The tactic produces the following printout: | 
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 561 | |
| 99 | 562 |   \begin{quote}\small
 | 
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 563 |   \begin{tabular}{ll}
 | 
| 440 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
437diff
changeset | 564 |   params:      & @{text "x:= x"}, @{text "y:= y"}\\
 | 
| 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 | 565 |   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 | 566 |   conclusion:  & @{term "B y x \<longrightarrow> C (z y) x"}\\
 | 
| 541 
96d10631eec2
added readme and fixed output in Subgoal.FOCUS section
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
517diff
changeset | 567 |   premises:    & @{term "A x y"}\\
 | 
| 
96d10631eec2
added readme and fixed output in Subgoal.FOCUS section
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
517diff
changeset | 568 |   schematics:  & @{text "?z:=z"}
 | 
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 569 |   \end{tabular}
 | 
| 99 | 570 |   \end{quote}
 | 
| 571 | ||
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 572 |   The @{text params} and @{text schematics} stand for list of pairs where the
 | 
| 363 | 573 |   left-hand side of @{text ":="} is replaced by the right-hand side inside the
 | 
| 574 |   tactic.  Notice that in the actual output the variables @{term x} and @{term
 | |
| 575 | y} have a brown colour. Although they are parameters in the original goal, | |
| 576 | they are fixed inside the tactic. By convention these fixed variables are | |
| 577 |   printed in brown colour.  Similarly the schematic variable @{text ?z}. The
 | |
| 578 |   assumption, or premise, @{prop "A x y"} is bound as @{ML_type cterm} to the
 | |
| 579 |   record-variable @{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 | 580 | |
| 99 | 581 |   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 | 582 | *} | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 583 | |
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 584 | apply(rule impI) | 
| 299 
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
298diff
changeset | 585 | apply(tactic {* Subgoal.FOCUS foc_tac @{context} 1 *})
 | 
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 586 | |
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 587 | txt {*
 | 
| 118 
5f003fdf2653
polished and added more material to the package chapter
 Christian Urban <urbanc@in.tum.de> parents: 
114diff
changeset | 588 | then the tactic prints out: | 
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 589 | |
| 99 | 590 |   \begin{quote}\small
 | 
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 591 |   \begin{tabular}{ll}
 | 
| 440 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
437diff
changeset | 592 |   params:      & @{text "x:= x"}, @{text "y:= y"}\\
 | 
| 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 | 593 |   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 | 594 |   conclusion:  & @{term "C (z y) x"}\\
 | 
| 541 
96d10631eec2
added readme and fixed output in Subgoal.FOCUS section
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
517diff
changeset | 595 |   premises:    & @{term "A x y"}, @{term "B y x"}\\
 | 
| 
96d10631eec2
added readme and fixed output in Subgoal.FOCUS section
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
517diff
changeset | 596 |   schematics:  & @{text "?z:=z"}
 | 
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 597 |   \end{tabular}
 | 
| 99 | 598 |   \end{quote}
 | 
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 599 | *} | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 600 | (*<*)oops(*>*) | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 601 | |
| 99 | 602 | text {*
 | 
| 109 | 603 |   Now also @{term "B y x"} is an assumption bound to @{text asms} and @{text prems}.
 | 
| 99 | 604 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 605 |   One difference between the tactics @{ML SUBPROOF} and @{ML FOCUS in Subgoal}
 | 
| 301 
2728e8daebc0
replaced "writeln" with "tracing"
 Christian Urban <urbanc@in.tum.de> parents: 
299diff
changeset | 606 | is that the former expects that the goal is solved completely, which the | 
| 363 | 607 |   latter does not. Another is that @{ML SUBPROOF} cannot instantiate any schematic
 | 
| 301 
2728e8daebc0
replaced "writeln" with "tracing"
 Christian Urban <urbanc@in.tum.de> parents: 
299diff
changeset | 608 | variables. | 
| 
2728e8daebc0
replaced "writeln" with "tracing"
 Christian Urban <urbanc@in.tum.de> parents: 
299diff
changeset | 609 | |
| 411 
24fc00319c4a
polised the section about Subgoal.FOCUS
 Christian Urban <urbanc@in.tum.de> parents: 
410diff
changeset | 610 |   Observe that inside @{ML FOCUS in Subgoal} and @{ML SUBPROOF}, we are in a goal
 | 
| 
24fc00319c4a
polised the section about Subgoal.FOCUS
 Christian Urban <urbanc@in.tum.de> parents: 
410diff
changeset | 611 |   state where there is only a conclusion. This means the tactics @{ML dtac} and 
 | 
| 
24fc00319c4a
polised the section about Subgoal.FOCUS
 Christian Urban <urbanc@in.tum.de> parents: 
410diff
changeset | 612 | the like are of no use for manipulating the goal state. The assumptions inside | 
| 
24fc00319c4a
polised the section about Subgoal.FOCUS
 Christian Urban <urbanc@in.tum.de> parents: 
410diff
changeset | 613 |   @{ML FOCUS in Subgoal} and @{ML SUBPROOF} are given as cterms and theorems in 
 | 
| 
24fc00319c4a
polised the section about Subgoal.FOCUS
 Christian Urban <urbanc@in.tum.de> parents: 
410diff
changeset | 614 |   the arguments @{text "asms"} and @{text "prems"}, respectively. This 
 | 
| 
24fc00319c4a
polised the section about Subgoal.FOCUS
 Christian Urban <urbanc@in.tum.de> parents: 
410diff
changeset | 615 | means we can apply them using the usual assumption tactics. With this you can | 
| 
24fc00319c4a
polised the section about Subgoal.FOCUS
 Christian Urban <urbanc@in.tum.de> parents: 
410diff
changeset | 616 |   for example easily implement a tactic that behaves almost like @{ML atac}:
 | 
| 99 | 617 | *} | 
| 618 | ||
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 619 | ML %grayML{*val atac' = Subgoal.FOCUS (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 | 620 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 621 | text {*
 | 
| 109 | 622 |   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 | 623 | *} | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 624 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 625 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 626 | shows "\<lbrakk>B x y; A x y; C x y\<rbrakk> \<Longrightarrow> A x y" | 
| 104 | 627 | apply(tactic {* atac' @{context} 1 *})
 | 
| 109 | 628 | txt{* it will produce
 | 
| 99 | 629 |       @{subgoals [display]} *}
 | 
| 630 | (*<*)oops(*>*) | |
| 631 | ||
| 104 | 632 | text {*
 | 
| 301 
2728e8daebc0
replaced "writeln" with "tracing"
 Christian Urban <urbanc@in.tum.de> parents: 
299diff
changeset | 633 |   Notice that @{ML atac'} inside @{ML FOCUS in Subgoal} calls @{ML
 | 
| 
2728e8daebc0
replaced "writeln" with "tracing"
 Christian Urban <urbanc@in.tum.de> parents: 
299diff
changeset | 634 |   resolve_tac} with the subgoal number @{text "1"} and also the outer call to
 | 
| 
2728e8daebc0
replaced "writeln" with "tracing"
 Christian Urban <urbanc@in.tum.de> parents: 
299diff
changeset | 635 |   @{ML FOCUS in Subgoal} in the \isacommand{apply}-step uses @{text "1"}. This
 | 
| 
2728e8daebc0
replaced "writeln" with "tracing"
 Christian Urban <urbanc@in.tum.de> parents: 
299diff
changeset | 636 |   is another advantage of @{ML FOCUS in Subgoal} and @{ML SUBPROOF}: the
 | 
| 
2728e8daebc0
replaced "writeln" with "tracing"
 Christian Urban <urbanc@in.tum.de> parents: 
299diff
changeset | 637 | addressing inside it is completely local to the tactic inside the | 
| 
2728e8daebc0
replaced "writeln" with "tracing"
 Christian Urban <urbanc@in.tum.de> parents: 
299diff
changeset | 638 |   subproof. It is therefore possible to also apply @{ML atac'} to the second
 | 
| 
2728e8daebc0
replaced "writeln" with "tracing"
 Christian Urban <urbanc@in.tum.de> parents: 
299diff
changeset | 639 | goal by just writing: | 
| 104 | 640 | |
| 641 | *} | |
| 642 | ||
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 643 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 644 | shows "True" | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 645 | and "\<lbrakk>B x y; A x y; C x y\<rbrakk> \<Longrightarrow> A x y" | 
| 104 | 646 | apply(tactic {* atac' @{context} 2 *})
 | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 647 | apply(rule TrueI) | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 648 | done | 
| 104 | 649 | |
| 93 | 650 | text {*
 | 
| 451 
fc074e669f9f
disabled foobar_prove; updated to new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
449diff
changeset | 651 |   To sum up, both @{ML FOCUS in Subgoal} and @{ML SUBPROOF} are rather
 | 
| 411 
24fc00319c4a
polised the section about Subgoal.FOCUS
 Christian Urban <urbanc@in.tum.de> parents: 
410diff
changeset | 652 | convenient, but can impose a considerable run-time penalty in automatic | 
| 
24fc00319c4a
polised the section about Subgoal.FOCUS
 Christian Urban <urbanc@in.tum.de> parents: 
410diff
changeset | 653 | proofs. If speed is of the essence, then maybe the two lower level combinators | 
| 
24fc00319c4a
polised the section about Subgoal.FOCUS
 Christian Urban <urbanc@in.tum.de> parents: 
410diff
changeset | 654 | described next are more appropriate. | 
| 
24fc00319c4a
polised the section about Subgoal.FOCUS
 Christian Urban <urbanc@in.tum.de> parents: 
410diff
changeset | 655 | |
| 
24fc00319c4a
polised the section about Subgoal.FOCUS
 Christian Urban <urbanc@in.tum.de> parents: 
410diff
changeset | 656 | |
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 657 |   \begin{readmore}
 | 
| 299 
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
298diff
changeset | 658 |   The functions @{ML FOCUS in Subgoal} and @{ML SUBPROOF} are defined in 
 | 
| 298 | 659 |   @{ML_file "Pure/subgoal.ML"} and also described in 
 | 
| 660 |   \isccite{sec:results}. 
 | |
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 661 |   \end{readmore}
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 662 | |
| 411 
24fc00319c4a
polised the section about Subgoal.FOCUS
 Christian Urban <urbanc@in.tum.de> parents: 
410diff
changeset | 663 |   Similar but less powerful functions than @{ML FOCUS in Subgoal},
 | 
| 
24fc00319c4a
polised the section about Subgoal.FOCUS
 Christian Urban <urbanc@in.tum.de> parents: 
410diff
changeset | 664 |   respectively @{ML SUBPROOF}, are @{ML_ind SUBGOAL in Tactical} and @{ML_ind
 | 
| 
24fc00319c4a
polised the section about Subgoal.FOCUS
 Christian Urban <urbanc@in.tum.de> parents: 
410diff
changeset | 665 | CSUBGOAL in Tactical}. They allow you to inspect a given subgoal (the former | 
| 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 |   presents the subgoal as a @{ML_type term}, while the latter as a @{ML_type
 | 
| 363 | 667 | cterm}). With them you can implement a tactic that applies a rule according | 
| 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 | 668 | to the topmost logic connective in the subgoal (to illustrate this we only | 
| 411 
24fc00319c4a
polised the section about Subgoal.FOCUS
 Christian Urban <urbanc@in.tum.de> parents: 
410diff
changeset | 669 | analyse a few connectives). The code of the tactic is as follows. | 
| 93 | 670 | *} | 
| 671 | ||
| 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 | 672 | ML %linenosgray{*fun select_tac (t, i) =
 | 
| 99 | 673 | 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 | 674 |      @{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 | 675 |    | @{term "op \<Longrightarrow>"} $ _ $ t' => select_tac (t', i)
 | 
| 99 | 676 |    | @{term "op \<and>"} $ _ $ _ => rtac @{thm conjI} i
 | 
| 677 |    | @{term "op \<longrightarrow>"} $ _ $ _ => rtac @{thm impI} i
 | |
| 678 |    | @{term "Not"} $ _ => rtac @{thm notI} i
 | |
| 679 |    | Const (@{const_name "All"}, _) $ _ => rtac @{thm allI} i
 | |
| 238 
29787dcf7b2e
added something about TRY and TRYALL
 Christian Urban <urbanc@in.tum.de> parents: 
232diff
changeset | 680 |    | _ => all_tac*}text_raw{*\label{tac:selecttac}*}
 | 
| 99 | 681 | |
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 682 | text {*
 | 
| 109 | 683 | The input of the function is a term representing the subgoal and a number | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
184diff
changeset | 684 | specifying the subgoal of interest. In Line 3 you need to descend under the | 
| 109 | 685 |   outermost @{term "Trueprop"} in order to get to the connective you like to
 | 
| 686 |   analyse. Otherwise goals like @{prop "A \<and> B"} are not properly
 | |
| 687 | analysed. Similarly with meta-implications in the next line. While for the | |
| 688 |   first five patterns we can use the @{text "@term"}-antiquotation to
 | |
| 689 | construct the patterns, the pattern in Line 8 cannot be constructed in this | |
| 690 | way. The reason is that an antiquotation would fix the type of the | |
| 363 | 691 | quantified variable. So you really have to construct this pattern using the | 
| 692 | basic term-constructors. This is not necessary in the other cases, because their | |
| 156 | 693 |   type is always fixed to function types involving only the type @{typ
 | 
| 298 | 694 |   bool}. (See Section \ref{sec:terms_types_manually} about constructing terms
 | 
| 156 | 695 |   manually.) For the catch-all pattern, we chose to just return @{ML all_tac}. 
 | 
| 696 |   Consequently, @{ML select_tac} never fails.
 | |
| 697 | ||
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 698 | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 699 | 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 | 700 | *} | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 701 | |
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 702 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 703 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 704 | shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F" | 
| 104 | 705 | apply(tactic {* SUBGOAL select_tac 4 *})
 | 
| 706 | apply(tactic {* SUBGOAL select_tac 3 *})
 | |
| 707 | apply(tactic {* SUBGOAL select_tac 2 *})
 | |
| 99 | 708 | 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 | 709 | txt{* \begin{minipage}{\textwidth}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 710 |       @{subgoals [display]}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 711 |       \end{minipage} *}
 | 
| 99 | 712 | (*<*)oops(*>*) | 
| 713 | ||
| 714 | text {*
 | |
| 363 | 715 | where in all but the last the tactic applies an introduction rule. | 
| 109 | 716 | Note that we applied the tactic to the goals in ``reverse'' order. | 
| 717 | This is a trick in order to be independent from the subgoals that are | |
| 718 | 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 | 719 | *} | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 720 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 721 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 722 | shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F" | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 723 | apply(tactic {* SUBGOAL select_tac 1 *})
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 724 | apply(tactic {* SUBGOAL select_tac 3 *})
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 725 | apply(tactic {* SUBGOAL select_tac 4 *})
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 726 | apply(tactic {* SUBGOAL select_tac 5 *})
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 727 | (*<*)oops(*>*) | 
| 99 | 728 | |
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 729 | text {*
 | 
| 109 | 730 | then we have to be careful to not apply the tactic to the two subgoals produced by | 
| 731 | 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 | 732 | the ``reverse application'' is easy to implement. | 
| 104 | 733 | |
| 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 | 734 | Of course, this example is | 
| 149 | 735 | contrived: there are much simpler methods available in Isabelle for | 
| 243 | 736 | implementing a tactic analysing a goal according to its topmost | 
| 149 | 737 | connective. These simpler methods use tactic combinators, which we will | 
| 363 | 738 | explain in the next section. But before that we will show how | 
| 739 | tactic application can be constrained. | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 740 | |
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 741 |   \begin{readmore}
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 742 |   The functions @{ML SUBGOAL} and @{ML CSUBGOAL} are defined in @{ML_file "Pure/tactical.ML"}.
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 743 |   \end{readmore}
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 744 | |
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 745 | |
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 746 |   Since @{ML_ind rtac in Tactic} and the like use higher-order unification, an
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 747 | automatic proof procedure based on them might become too fragile, if it just | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 748 | applies theorems as shown above. The reason is that a number of theorems | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 749 | introduce schematic variables into the goal state. Consider for example the | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 750 | proof | 
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 751 | *} | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 752 | |
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 753 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 754 | shows "\<forall>x \<in> A. P x \<Longrightarrow> Q x" | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 755 | apply(tactic {* dtac @{thm bspec} 1 *})
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 756 | txt{*\begin{minipage}{\textwidth}
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 757 |      @{subgoals [display]} 
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 758 |      \end{minipage}*}
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 759 | (*<*)oops(*>*) | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 760 | |
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 761 | text {*
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 762 |   where the application of theorem @{text bspec} generates two subgoals involving the
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 763 |   new schematic variable @{text "?x"}. Now, if you are not careful, tactics 
 | 
| 363 | 764 | applied to the first subgoal might instantiate this schematic variable in such a | 
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 765 |   way that the second subgoal becomes unprovable. If it is clear what the @{text "?x"}
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 766 | should be, then this situation can be avoided by introducing a more | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 767 |   constrained version of the @{text bspec}-theorem. One way to give such 
 | 
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 768 | constraints is by pre-instantiating theorems with other theorems. | 
| 363 | 769 |   The function @{ML_ind RS in Drule}, for example, does this.
 | 
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 770 | |
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 771 |   @{ML_response_fake [display,gray]
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 772 |   "@{thm disjI1} RS @{thm conjI}" "\<lbrakk>?P1; ?Q\<rbrakk> \<Longrightarrow> (?P1 \<or> ?Q1) \<and> ?Q"}
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 773 | |
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 774 |   In this example it instantiates the first premise of the @{text conjI}-theorem 
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 775 |   with the theorem @{text disjI1}. If the instantiation is impossible, as in the 
 | 
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 776 | case of | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 777 | |
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 778 |   @{ML_response_fake_both [display,gray]
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 779 |   "@{thm conjI} RS @{thm mp}" 
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 780 | "*** Exception- THM (\"RSN: no unifiers\", 1, | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 781 | [\"\<lbrakk>?P; ?Q\<rbrakk> \<Longrightarrow> ?P \<and> ?Q\", \"\<lbrakk>?P \<longrightarrow> ?Q; ?P\<rbrakk> \<Longrightarrow> ?Q\"]) raised"} | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 782 | |
| 363 | 783 |   then the function raises an exception. The function @{ML_ind  RSN in Drule} 
 | 
| 784 |   is similar to @{ML RS}, but takes an additional number as argument. This 
 | |
| 785 | number makes explicit which premise should be instantiated. | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 786 | |
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 787 | If you want to instantiate more than one premise of a theorem, you can use | 
| 363 | 788 |   the function @{ML_ind MRS in Drule}:
 | 
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 789 | |
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 790 |   @{ML_response_fake [display,gray]
 | 
| 466 | 791 |   "[@{thm disjI1}, @{thm disjI2}] MRS @{thm conjI}" 
 | 
| 792 | "\<lbrakk>?P2; ?Q1\<rbrakk> \<Longrightarrow> (?P2 \<or> ?Q2) \<and> (?P1 \<or> ?Q1)"} | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 793 | |
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 794 | If you need to instantiate lists of theorems, you can use the | 
| 510 
500d1abbc98c
updated to Isabelle 15 Feb
 Christian Urban <urbanc@in.tum.de> parents: 
506diff
changeset | 795 |   functions @{ML_ind RL in Drule} and @{ML_ind OF in Drule}. For 
 | 
| 363 | 796 | example in the code below, every theorem in the second list is | 
| 797 | instantiated with every theorem in the first. | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 798 | |
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 799 |   @{ML_response_fake [display,gray]
 | 
| 363 | 800 | "let | 
| 801 |   val list1 = [@{thm impI}, @{thm disjI2}]
 | |
| 802 |   val list2 = [@{thm conjI}, @{thm disjI1}]
 | |
| 803 | in | |
| 466 | 804 | list1 RL list2 | 
| 363 | 805 | end" | 
| 466 | 806 | "[\<lbrakk>?P1 \<Longrightarrow> ?Q1; ?Q\<rbrakk> \<Longrightarrow> (?P1 \<longrightarrow> ?Q1) \<and> ?Q, | 
| 807 | \<lbrakk>?Q1; ?Q\<rbrakk> \<Longrightarrow> (?P1 \<or> ?Q1) \<and> ?Q, | |
| 808 | (?P1 \<Longrightarrow> ?Q1) \<Longrightarrow> (?P1 \<longrightarrow> ?Q1) \<or> ?Q, | |
| 809 | ?Q1 \<Longrightarrow> (?P1 \<or> ?Q1) \<or> ?Q]"} | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 810 | |
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 811 |   \begin{readmore}
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 812 | The combinators for instantiating theorems with other theorems are | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 813 |   defined in @{ML_file "Pure/drule.ML"}.
 | 
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 814 |   \end{readmore}
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 815 | |
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 816 | Higher-order unification is also often in the way when applying certain | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 817 | congruence theorems. For example we would expect that the theorem | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 818 |   @{thm [source] cong}
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 819 | |
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 820 |   \begin{isabelle}
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 821 |   \isacommand{thm}~@{thm [source] cong}\\
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 822 |   @{text "> "}~@{thm cong}
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 823 |   \end{isabelle}
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 824 | |
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 825 | is applicable in the following proof producing the subgoals | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 826 |   @{term "t x = s u"} and @{term "y = w"}. 
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 827 | *} | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 828 | |
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 829 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 830 | fixes x y u w::"'a" | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 831 | shows "t x y = s u w" | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 832 | apply(rule cong) | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 833 | txt{*\begin{minipage}{\textwidth}
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 834 |      @{subgoals [display]}
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 835 |      \end{minipage}*}
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 836 | (*<*)oops(*>*) | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 837 | |
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 838 | text {*
 | 
| 363 | 839 |   As you can see this is unfortunately \emph{not} the case if we apply @{thm [source] 
 | 
| 840 |   cong} with the method @{text rule}. The problem is
 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 841 | that higher-order unification produces an instantiation that is not the | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 842 |   intended one. While we can use \isacommand{back} to interactively find the
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 843 | intended instantiation, this is not an option for an automatic proof | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 844 |   procedure. Fortunately, the tactic @{ML_ind cong_tac in Cong_Tac} helps 
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 845 | with applying congruence theorems and finding the intended instantiation. | 
| 363 | 846 | For example | 
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 847 | *} | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 848 | |
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 849 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 850 | fixes x y u w::"'a" | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 851 | shows "t x y = s u w" | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 852 | apply(tactic {* Cong_Tac.cong_tac @{thm cong} 1 *})
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 853 | txt{*\begin{minipage}{\textwidth}
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 854 |      @{subgoals [display]}
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 855 |      \end{minipage}*}
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 856 | (*<*)oops(*>*) | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 857 | |
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 858 | text {*
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 859 | However, frequently it is necessary to explicitly match a theorem against a goal | 
| 363 | 860 | state and in doing so construct manually an appropriate instantiation. Imagine | 
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 861 | you have the theorem | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 862 | *} | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 863 | |
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 864 | lemma rule: | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 865 | shows "\<lbrakk>f = g; x = y\<rbrakk> \<Longrightarrow> R (f x) (g y)" | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 866 | sorry | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 867 | |
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 868 | text {* 
 | 
| 551 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 869 |   and you want to apply it to the goal @{term "(t\<^sub>1 t\<^sub>2) \<le>
 | 
| 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 870 | (s\<^sub>1 s\<^sub>2)"}. Since in the theorem all variables are | 
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 871 |   schematic, we have a nasty higher-order unification problem and @{text rtac}
 | 
| 363 | 872 | will not be able to apply this rule in the way we want. For the goal at hand | 
| 873 |   we want to use it so that @{term R} is instantiated to the constant @{text \<le>} and
 | |
| 874 | similarly ``obvious'' instantiations for the other variables. To achieve this | |
| 875 | we need to match the conclusion of | |
| 876 |   @{thm [source] rule} against the goal reading off an instantiation for
 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 877 |   @{thm [source] rule}. For this the function @{ML_ind first_order_match in Thm} 
 | 
| 363 | 878 |   matches two @{ML_type cterm}s and produces, in the successful case, a matcher 
 | 
| 879 | that can be used to instantiate the theorem. The instantiation | |
| 465 | 880 |   can be done with the function @{ML_ind instantiate_normalize in Drule}. To automate 
 | 
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 881 | this we implement the following function. | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 882 | *} | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 883 | |
| 363 | 884 | ML %linenosgray{*fun fo_rtac thm = Subgoal.FOCUS (fn {concl, ...} =>
 | 
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 885 | let | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 886 | val concl_pat = Drule.strip_imp_concl (cprop_of thm) | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 887 | val insts = Thm.first_order_match (concl_pat, concl) | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 888 | in | 
| 465 | 889 | rtac (Drule.instantiate_normalize insts thm) 1 | 
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 890 | end)*} | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 891 | |
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 892 | text {*
 | 
| 363 | 893 |   Note that we use @{ML FOCUS in Subgoal} because it gives us directly access
 | 
| 894 | to the conclusion of the goal state, but also because this function | |
| 895 | takes care of correctly handling parameters that might be present | |
| 896 |   in the goal state. In Line 3 we use the function @{ML_ind strip_imp_concl in Drule}
 | |
| 897 |   for calculating the conclusion of a theorem (produced as @{ML_type cterm}).
 | |
| 898 |   An example of @{ML fo_rtac} is as follows.
 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 899 | *} | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 900 | |
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 901 | lemma | 
| 551 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 902 | shows "\<And>t\<^sub>1 s\<^sub>1 (t\<^sub>2::'a) (s\<^sub>2::'a). (t\<^sub>1 t\<^sub>2) \<le> (s\<^sub>1 s\<^sub>2)" | 
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 903 | apply(tactic {* fo_rtac @{thm rule} @{context} 1 *})
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 904 | txt{*\begin{minipage}{\textwidth}
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 905 |      @{subgoals [display]}
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 906 |      \end{minipage}*}
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 907 | (*<*)oops(*>*) | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 908 | |
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 909 | text {*
 | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 910 | We obtain the intended subgoals and also the parameters are correctly | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 911 | introduced in both of them. Such manual instantiations are quite frequently | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 912 | necessary in order to appropriately constrain the application of theorems. | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 913 | Otherwise one can end up with ``wild'' higher-order unification problems | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 914 | that make automatic proofs fail. | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 915 | |
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 916 |   \begin{readmore}
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 917 |   Functions for matching @{ML_type cterm}s are defined in @{ML_file "Pure/thm.ML"}.
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 918 | Functions for instantiating schematic variables in theorems are defined | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 919 |   in @{ML_file "Pure/drule.ML"}.
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 920 |   \end{readmore}
 | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 921 | *} | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 922 | |
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 923 | section {* Tactic Combinators *}
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 924 | |
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 925 | text {* 
 | 
| 109 | 926 | The purpose of tactic combinators is to build compound tactics out of | 
| 363 | 927 |   smaller tactics. In the previous section we already used @{ML_ind THEN in Tactical}, 
 | 
| 928 | which just strings together two tactics in a sequence. For example: | |
| 93 | 929 | *} | 
| 930 | ||
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 931 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 932 | shows "(Foo \<and> Bar) \<and> False" | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 933 | 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 | 934 | txt {* \begin{minipage}{\textwidth}
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 935 |        @{subgoals [display]} 
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 936 |        \end{minipage} *}
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 937 | (*<*)oops(*>*) | 
| 99 | 938 | |
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 939 | text {*
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 940 | If you want to avoid the hard-coded subgoal addressing in each component, | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 941 |   then, as seen earlier, you can use the ``primed'' version of @{ML THEN}. 
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 942 | For example: | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 943 | *} | 
| 93 | 944 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 945 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 946 | shows "(Foo \<and> Bar) \<and> False" | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 947 | 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 | 948 | txt {* \begin{minipage}{\textwidth}
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 949 |        @{subgoals [display]} 
 | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 950 |        \end{minipage} *}
 | 
| 93 | 951 | (*<*)oops(*>*) | 
| 952 | ||
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 953 | text {* 
 | 
| 404 | 954 | Here you have to specify the subgoal of interest only once and it is | 
| 955 | consistently applied to the component. For most tactic combinators such a | |
| 956 | ``primed'' version exists and in what follows we will usually prefer it over | |
| 957 | the ``unprimed'' one. | |
| 958 | ||
| 959 |   The tactic combinator @{ML_ind RANGE in Tactical} takes a list of @{text n}
 | |
| 960 |   tactics, say, and applies them to each of the first @{text n} subgoals. 
 | |
| 961 | For example below we first apply the introduction rule for conjunctions | |
| 962 | and then apply a tactic to each of the two subgoals. | |
| 963 | *} | |
| 964 | ||
| 965 | lemma | |
| 966 | shows "A \<Longrightarrow> True \<and> A" | |
| 967 | apply(tactic {* (rtac @{thm conjI} 
 | |
| 968 |                  THEN' RANGE [rtac @{thm TrueI}, atac]) 1 *})
 | |
| 969 | txt {* \begin{minipage}{\textwidth}
 | |
| 970 |        @{subgoals [display]} 
 | |
| 971 |        \end{minipage} *}
 | |
| 972 | (*<*)oops(*>*) | |
| 973 | ||
| 974 | text {*
 | |
| 975 | If there is a list of tactics that should all be tried out in sequence on | |
| 976 |   one specified subgoal, you can use the combinator @{ML_ind EVERY' in
 | |
| 977 |   Tactical}. For example the function @{ML foo_tac'} from
 | |
| 978 |   page~\pageref{tac:footacprime} can also be written as:
 | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 979 | *} | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 980 | |
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 981 | ML %grayML{*val foo_tac'' = EVERY' [etac @{thm disjE}, rtac @{thm disjI2}, 
 | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 982 |                         atac, rtac @{thm disjI1}, atac]*}
 | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 983 | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 984 | text {*
 | 
| 109 | 985 | There is even another way of implementing this tactic: in automatic proof | 
| 986 | procedures (in contrast to tactics that might be called by the user) there | |
| 987 | are often long lists of tactics that are applied to the first | |
| 988 |   subgoal. Instead of writing the code above and then calling @{ML "foo_tac'' 1"}, 
 | |
| 989 | you can also just write | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 990 | *} | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 991 | |
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 992 | ML %grayML{*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 | 993 |                        atac, rtac @{thm disjI1}, atac]*}
 | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 994 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 995 | text {*
 | 
| 118 
5f003fdf2653
polished and added more material to the package chapter
 Christian Urban <urbanc@in.tum.de> parents: 
114diff
changeset | 996 |   and call @{ML foo_tac1}.  
 | 
| 109 | 997 | |
| 363 | 998 |   With the combinators @{ML THEN'}, @{ML EVERY'} and @{ML_ind  EVERY1 in Tactical} it must be
 | 
| 109 | 999 | guaranteed that all component tactics successfully apply; otherwise the | 
| 1000 | whole tactic will fail. If you rather want to try out a number of tactics, | |
| 363 | 1001 |   then you can use the combinator @{ML_ind  ORELSE' in Tactical} for two tactics, and @{ML_ind
 | 
| 1002 |    FIRST' in Tactical} (or @{ML_ind  FIRST1 in Tactical}) for a list of tactics. For example, the tactic
 | |
| 109 | 1003 | |
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 1004 | *} | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 1005 | |
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 1006 | ML %grayML{*val orelse_xmp_tac = rtac @{thm disjI1} ORELSE' rtac @{thm conjI}*}
 | 
| 99 | 1007 | |
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 1008 | text {*
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1009 |   will first try out whether theorem @{text disjI} applies and in case of failure 
 | 
| 243 | 1010 |   will try @{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 | 1011 | *} | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 1012 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1013 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1014 | shows "True \<and> False" and "Foo \<or> Bar" | 
| 238 
29787dcf7b2e
added something about TRY and TRYALL
 Christian Urban <urbanc@in.tum.de> parents: 
232diff
changeset | 1015 | apply(tactic {* orelse_xmp_tac 2 *})
 | 
| 
29787dcf7b2e
added something about TRY and TRYALL
 Christian Urban <urbanc@in.tum.de> parents: 
232diff
changeset | 1016 | apply(tactic {* orelse_xmp_tac 1 *})
 | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1017 | txt {* which results in the goal state
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1018 |        \begin{isabelle}
 | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1019 |        @{subgoals [display]} 
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1020 |        \end{isabelle}
 | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1021 | *} | 
| 93 | 1022 | (*<*)oops(*>*) | 
| 1023 | ||
| 1024 | ||
| 1025 | text {*
 | |
| 109 | 1026 |   Using @{ML FIRST'} we can simplify our @{ML select_tac} from Page~\pageref{tac:selecttac} 
 | 
| 1027 | as follows: | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1028 | *} | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1029 | |
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 1030 | ML %grayML{*val select_tac' = FIRST' [rtac @{thm conjI}, rtac @{thm impI}, 
 | 
| 238 
29787dcf7b2e
added something about TRY and TRYALL
 Christian Urban <urbanc@in.tum.de> parents: 
232diff
changeset | 1031 |                           rtac @{thm notI}, rtac @{thm allI}, K all_tac]*}text_raw{*\label{tac:selectprime}*}
 | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1032 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1033 | 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 | 1034 |   Since we like to mimic the behaviour of @{ML select_tac} as closely as possible, 
 | 
| 109 | 1035 |   we must include @{ML all_tac} at the end of the list, otherwise the tactic will
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1036 |   fail if no theorem applies (we also have to wrap @{ML all_tac} using the 
 | 
| 109 | 1037 |   @{ML K}-combinator, because it does not take a subgoal number as argument). You can 
 | 
| 1038 | 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 | 1039 | *} | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1040 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1041 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1042 | shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F" | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1043 | apply(tactic {* select_tac' 4 *})
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1044 | apply(tactic {* select_tac' 3 *})
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1045 | apply(tactic {* select_tac' 2 *})
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1046 | apply(tactic {* select_tac' 1 *})
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1047 | txt{* \begin{minipage}{\textwidth}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1048 |       @{subgoals [display]}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1049 |       \end{minipage} *}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1050 | (*<*)oops(*>*) | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1051 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1052 | text {* 
 | 
| 109 | 1053 | Since such repeated applications of a tactic to the reverse order of | 
| 1054 |   \emph{all} subgoals is quite common, there is the tactic combinator 
 | |
| 363 | 1055 |   @{ML_ind  ALLGOALS in Tactical} 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 | 1056 | write: *} | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1057 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1058 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1059 | shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F" | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1060 | apply(tactic {* ALLGOALS select_tac' *})
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1061 | txt{* \begin{minipage}{\textwidth}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1062 |       @{subgoals [display]}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1063 |       \end{minipage} *}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1064 | (*<*)oops(*>*) | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1065 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1066 | text {*
 | 
| 109 | 1067 |   Remember that we chose to implement @{ML select_tac'} so that it 
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1068 |   always  succeeds by fact of having @{ML all_tac} at the end of the tactic
 | 
| 363 | 1069 |   list. The same can be achieved with the tactic combinator @{ML_ind  TRY in Tactical}.
 | 
| 243 | 1070 | For example: | 
| 1071 | *} | |
| 1072 | ||
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 1073 | ML %grayML{*val select_tac'' = TRY o FIRST' [rtac @{thm conjI}, rtac @{thm impI}, 
 | 
| 298 | 1074 |                                  rtac @{thm notI}, rtac @{thm allI}]*}
 | 
| 243 | 1075 | text_raw{*\label{tac:selectprimeprime}*}
 | 
| 1076 | ||
| 1077 | text {*
 | |
| 1078 |   This tactic behaves in the same way as @{ML select_tac'}: it tries out
 | |
| 1079 | one of the given tactics and if none applies leaves the goal state | |
| 1080 | unchanged. This, however, can be potentially very confusing when visible to | |
| 1081 | the user, for example, in cases where the goal is the form | |
| 1082 | ||
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1083 | *} | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1084 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1085 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1086 | shows "E \<Longrightarrow> F" | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1087 | apply(tactic {* select_tac' 1 *})
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1088 | txt{* \begin{minipage}{\textwidth}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1089 |       @{subgoals [display]}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1090 |       \end{minipage} *}
 | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1091 | (*<*)oops(*>*) | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1092 | |
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1093 | text {*
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1094 |   In this case no theorem applies. But because we wrapped the tactic in a @{ML
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1095 | TRY}, it does not fail anymore. The problem is that for the user there is | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1096 | little chance to see whether progress in the proof has been made, or not. By | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1097 | convention therefore, tactics visible to the user should either change | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1098 | something or fail. | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1099 | |
| 109 | 1100 |   To comply with this convention, we could simply delete the @{ML "K all_tac"}
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1101 |   in @{ML select_tac'} or delete @{ML TRY} from @{ML select_tac''}. But for
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1102 |   the sake of argument, let us suppose that this deletion is \emph{not} an
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1103 |   option. In such cases, you can use the combinator @{ML_ind CHANGED in
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1104 | Tactical} to make sure the subgoal has been changed by the tactic. Because | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1105 | now | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1106 | *} | 
| 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1107 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1108 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1109 | shows "E \<Longrightarrow> F" | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1110 | apply(tactic {* CHANGED (select_tac' 1) *})(*<*)?(*>*)
 | 
| 109 | 1111 | 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 | 1112 |   \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 | 1113 |   @{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 | 1114 |   @{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 | 1115 |   \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 | 1116 | *}(*<*)oops(*>*) | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 1117 | |
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 1118 | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1119 | text {*
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1120 |   We can further extend the @{ML select_tac}s so that they not just apply to the topmost
 | 
| 109 | 1121 | connective, but also to the ones immediately ``underneath'', i.e.~analyse the goal | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1122 |   completely. For this you can use the tactic combinator @{ML_ind REPEAT in Tactical}. As an example 
 | 
| 109 | 1123 | 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 | 1124 | *} | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1125 | |
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 1126 | ML %grayML{*val repeat_xmp_tac = REPEAT (CHANGED (select_tac' 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 | 1127 | |
| 109 | 1128 | 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 | 1129 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1130 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1131 | shows "((\<not>A) \<and> (\<forall>x. B x)) \<and> (C \<longrightarrow> D)" | 
| 238 
29787dcf7b2e
added something about TRY and TRYALL
 Christian Urban <urbanc@in.tum.de> parents: 
232diff
changeset | 1132 | apply(tactic {* repeat_xmp_tac *})
 | 
| 109 | 1133 | txt{* produces
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1134 |       \begin{isabelle}
 | 
| 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 | 1135 |       @{subgoals [display]}
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1136 |       \end{isabelle} *}
 | 
| 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 | 1137 | (*<*)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 | 1138 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1139 | text {*
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1140 |   Here it is crucial that @{ML select_tac'} is prefixed with @{ML CHANGED},
 | 
| 109 | 1141 |   because otherwise @{ML REPEAT} runs into an infinite loop (it applies the
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1142 |   tactic as long as it succeeds). The function @{ML_ind REPEAT1 in Tactical}
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1143 | is similar, but runs the tactic at least once (failing if this is not | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1144 | possible). | 
| 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 | 1145 | |
| 238 
29787dcf7b2e
added something about TRY and TRYALL
 Christian Urban <urbanc@in.tum.de> parents: 
232diff
changeset | 1146 |   If you are after the ``primed'' version of @{ML repeat_xmp_tac}, then you 
 | 
| 243 | 1147 | can 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 | 1148 | *} | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1149 | |
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 1150 | ML %grayML{*val repeat_xmp_tac' = REPEAT o CHANGED o select_tac'*}
 | 
| 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 | 1151 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1152 | 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 | 1153 |   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 | 1154 | |
| 243 | 1155 |   If you look closely at the goal state above, then you see the tactics @{ML repeat_xmp_tac}
 | 
| 238 
29787dcf7b2e
added something about TRY and TRYALL
 Christian Urban <urbanc@in.tum.de> parents: 
232diff
changeset | 1156 |   and @{ML repeat_xmp_tac'} are not yet quite what we are after: the problem is
 | 
| 109 | 1157 | that goals 2 and 3 are not analysed. This is because the tactic | 
| 1158 | is applied repeatedly only to the first subgoal. To analyse also all | |
| 363 | 1159 |   resulting subgoals, you can use the tactic combinator @{ML_ind  REPEAT_ALL_NEW in Tactical}. 
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1160 | Supposing the 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 | 1161 | *} | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1162 | |
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 1163 | ML %grayML{*val repeat_all_new_xmp_tac = REPEAT_ALL_NEW (CHANGED o select_tac')*}
 | 
| 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 | 1164 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1165 | text {* 
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1166 | you can 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 | 1167 | *} | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1168 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1169 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1170 | shows "((\<not>A) \<and> (\<forall>x. B x)) \<and> (C \<longrightarrow> D)" | 
| 238 
29787dcf7b2e
added something about TRY and TRYALL
 Christian Urban <urbanc@in.tum.de> parents: 
232diff
changeset | 1171 | apply(tactic {* repeat_all_new_xmp_tac 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 | 1172 | 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 | 1173 |       @{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 | 1174 |       \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 | 1175 | (*<*)oops(*>*) | 
| 93 | 1176 | |
| 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 | 1177 | text {*
 | 
| 109 | 1178 | 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 | 1179 |   include in @{ML select_tac'}. 
 | 
| 109 | 1180 | |
| 1181 | 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 | 1182 |   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 | 1183 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1184 | *} | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1185 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1186 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1187 | shows "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R" | 
| 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 | 1188 | apply(tactic {* etac @{thm disjE} 1 *})
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1189 | txt{* applies the rule to the first assumption yielding the goal state:
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1190 |       \begin{isabelle}
 | 
| 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 | 1191 |       @{subgoals [display]}
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1192 |       \end{isabelle}
 | 
| 109 | 1193 | |
| 1194 | After typing | |
| 1195 | *} | |
| 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 | 1196 | (*<*) | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1197 | oops | 
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1198 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1199 | shows "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R" | 
| 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 | 1200 | 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 | 1201 | (*>*) | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1202 | back | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1203 | txt{* the rule now applies to the second assumption.
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1204 |       \begin{isabelle}
 | 
| 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 | 1205 |       @{subgoals [display]}
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1206 |       \end{isabelle} *}
 | 
| 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 | 1207 | (*<*)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 | 1208 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1209 | 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 | 1210 | Sometimes this leads to confusing behaviour of tactics and also has | 
| 109 | 1211 | the potential to explode the search space for tactics. | 
| 1212 | These problems can be avoided by prefixing the tactic with the tactic | |
| 363 | 1213 |   combinator @{ML_ind  DETERM in Tactical}.
 | 
| 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 | 1214 | *} | 
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1215 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1216 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1217 | shows "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R" | 
| 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 | 1218 | 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 | 1219 | 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 | 1220 |       @{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 | 1221 |       \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 | 1222 | (*<*)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 | 1223 | text {*
 | 
| 118 
5f003fdf2653
polished and added more material to the package chapter
 Christian Urban <urbanc@in.tum.de> parents: 
114diff
changeset | 1224 | 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 | 1225 |   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 | 1226 | 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 | 1227 | |
| 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 1228 |   \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 | 1229 |   @{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 | 1230 |   @{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 | 1231 |   \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 | 1232 | |
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1233 | |
| 369 | 1234 |   \footnote{\bf FIXME: say something about @{ML_ind COND in Tactical} and COND'}
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1235 |   \footnote{\bf FIXME: PARALLEL-CHOICE PARALLEL-GOALS}
 | 
| 238 
29787dcf7b2e
added something about TRY and TRYALL
 Christian Urban <urbanc@in.tum.de> parents: 
232diff
changeset | 1236 | |
| 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 | 1237 |   \begin{readmore}
 | 
| 289 
08ffafe2585d
adapted to changes in Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
288diff
changeset | 1238 |   Most tactic combinators described in this section are defined in @{ML_file "Pure/tactical.ML"}.
 | 
| 238 
29787dcf7b2e
added something about TRY and TRYALL
 Christian Urban <urbanc@in.tum.de> parents: 
232diff
changeset | 1239 |   Some combinators for the purpose of proof search are implemented in @{ML_file "Pure/search.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 | 1240 |   \end{readmore}
 | 
| 314 | 1241 | *} | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 1242 | |
| 314 | 1243 | text {*
 | 
| 313 
1ca2f41770cc
polished the exercises about constructing terms
 Christian Urban <urbanc@in.tum.de> parents: 
308diff
changeset | 1244 |   \begin{exercise}\label{ex:dyckhoff}
 | 
| 370 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1245 |   Dyckhoff presents in \cite{Dyckhoff92} inference rules of a sequent
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1246 | calculus, called G4ip, for intuitionistic propositional logic. The | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1247 | interesting feature of this calculus is that no contraction rule is needed | 
| 370 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1248 | in order to be complete. As a result the rules can be applied exhaustively, which | 
| 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1249 | in turn leads to simple decision procedure for propositional intuitionistic logic. | 
| 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1250 | The task is to implement this decision procedure as a tactic. His rules are | 
| 314 | 1251 | |
| 1252 |   \begin{center}
 | |
| 1253 |   \def\arraystretch{2.3}
 | |
| 1254 |   \begin{tabular}{cc}
 | |
| 1255 |   \infer[Ax]{A,\varGamma \Rightarrow A}{} &
 | |
| 1256 |   \infer[False]{False,\varGamma \Rightarrow G}{}\\
 | |
| 1257 | ||
| 1258 |   \infer[\wedge_L]{A \wedge B, \varGamma \Rightarrow G}{A, B, \varGamma \Rightarrow G} &
 | |
| 1259 | \infer[\wedge_R] | |
| 1260 |   {\varGamma \Rightarrow A\wedge B}{\varGamma \Rightarrow A & \varGamma \Rightarrow B}\\
 | |
| 313 
1ca2f41770cc
polished the exercises about constructing terms
 Christian Urban <urbanc@in.tum.de> parents: 
308diff
changeset | 1261 | |
| 314 | 1262 | \infer[\vee_L] | 
| 1263 |   {A\vee B, \varGamma \Rightarrow G}{A,\varGamma \Rightarrow G & B,\varGamma \Rightarrow G} &
 | |
| 1264 |   \infer[\vee_{R_1}]
 | |
| 1265 |   {\varGamma \Rightarrow A \vee B}{\varGamma \Rightarrow A} \hspace{3mm}
 | |
| 1266 |   \infer[\vee_{R_2}]
 | |
| 1267 |   {\varGamma \Rightarrow A \vee B}{\varGamma \Rightarrow B}\\
 | |
| 1268 | ||
| 1269 |   \infer[\longrightarrow_{L_1}]
 | |
| 1270 |   {A\longrightarrow B, A, \varGamma \Rightarrow G}{B, A, \varGamma \Rightarrow G} &
 | |
| 1271 | \infer[\longrightarrow_R] | |
| 1272 |   {\varGamma \Rightarrow A\longrightarrow B}{A,\varGamma \Rightarrow B}\\
 | |
| 1273 | ||
| 1274 |   \infer[\longrightarrow_{L_2}]
 | |
| 1275 |   {(C \wedge D)\longrightarrow B, \varGamma \Rightarrow G}
 | |
| 1276 |   {C\longrightarrow (D \longrightarrow B), \varGamma \Rightarrow G} &
 | |
| 1277 | ||
| 1278 |   \infer[\longrightarrow_{L_3}]
 | |
| 1279 |   {(C \vee D)\longrightarrow B, \varGamma \Rightarrow G}
 | |
| 1280 |   {C\longrightarrow B, D\longrightarrow B, \varGamma \Rightarrow G}\\
 | |
| 1281 | ||
| 1282 |   \multicolumn{2}{c}{
 | |
| 1283 |   \infer[\longrightarrow_{L_4}]
 | |
| 1284 |   {(C \longrightarrow D)\longrightarrow B, \varGamma \Rightarrow G}
 | |
| 1285 |   {D\longrightarrow B, \varGamma \Rightarrow C \longrightarrow D &
 | |
| 1286 | B, \varGamma \Rightarrow G}}\\ | |
| 1287 |   \end{tabular}
 | |
| 1288 |   \end{center}
 | |
| 1289 | ||
| 370 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1290 | Note that in Isabelle right rules need to be implemented as | 
| 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1291 | introduction rule, the left rules as elimination rules. You have to to | 
| 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1292 |   prove separate theorems corresponding to $\longrightarrow_{L_{1..4}}$. The
 | 
| 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1293 | tactic should explore all possibilites of applying these rules to a | 
| 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1294 | propositional formula until a goal state is reached in which all subgoals | 
| 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1295 |   are discharged. For this you can use the tactic combinator @{ML_ind
 | 
| 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1296 |   DEPTH_SOLVE in Search} in the structure @{ML_struct Search}.
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1297 |   \end{exercise}
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1298 | |
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1299 |   \begin{exercise}
 | 
| 370 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1300 | Add to the sequent calculus from the previous exercise also rules for | 
| 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1301 | equality and run your tactic on the de Bruijn formulae discussed | 
| 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1302 |   in Exercise~\ref{ex:debruijn}.
 | 
| 313 
1ca2f41770cc
polished the exercises about constructing terms
 Christian Urban <urbanc@in.tum.de> parents: 
308diff
changeset | 1303 |   \end{exercise}
 | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 1304 | *} | 
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 1305 | |
| 388 
0b337dedc306
added Skip_Proof.mk_thm and some pointers about concurrency
 Christian Urban <urbanc@in.tum.de> parents: 
384diff
changeset | 1306 | section {* Simplifier Tactics\label{sec:simplifier} *}
 | 
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 1307 | |
| 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 1308 | text {*
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1309 | A lot of convenience in reasoning with Isabelle derives from its | 
| 370 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1310 | powerful simplifier. We will describe it in this section. However, | 
| 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1311 | due to its complexity, we can mostly only scratch the surface. | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1312 | |
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1313 | The power of the simplifier is a strength and a weakness at the same time, | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1314 | because you can easily make the simplifier run into a loop and in general | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1315 | its behaviour can be difficult to predict. There is also a multitude of | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1316 | options that you can configure to change the behaviour of the | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1317 | simplifier. There are the following five main tactics behind the simplifier | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1318 | (in parentheses is their user-level counterpart): | 
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1319 | |
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1320 |   \begin{isabelle}
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1321 |   \begin{center}
 | 
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1322 |   \begin{tabular}{l@ {\hspace{2cm}}l}
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1323 |    @{ML_ind simp_tac in Simplifier}            & @{text "(simp (no_asm))"} \\
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1324 |    @{ML_ind asm_simp_tac in Simplifier}        & @{text "(simp (no_asm_simp))"} \\
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1325 |    @{ML_ind full_simp_tac in Simplifier}       & @{text "(simp (no_asm_use))"} \\
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1326 |    @{ML_ind asm_lr_simp_tac in Simplifier}     & @{text "(simp (asm_lr))"} \\
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1327 |    @{ML_ind asm_full_simp_tac in Simplifier}   & @{text "(simp)"}
 | 
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1328 |   \end{tabular}
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1329 |   \end{center}
 | 
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1330 |   \end{isabelle}
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1331 | |
| 370 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1332 | All these tactics take a simpset and an integer as argument (the latter as usual | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1333 | to specify the goal to be analysed). So the proof | 
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1334 | *} | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1335 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1336 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1337 | shows "Suc (1 + 2) < 3 + 2" | 
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1338 | apply(simp) | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1339 | done | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1340 | |
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1341 | text {*
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1342 | corresponds on the ML-level to the tactic | 
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1343 | *} | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1344 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1345 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1346 | shows "Suc (1 + 2) < 3 + 2" | 
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1347 | apply(tactic {* asm_full_simp_tac @{context} 1 *})
 | 
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1348 | done | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1349 | |
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1350 | text {*
 | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1351 | If the simplifier cannot make any progress, then it leaves the goal unchanged, | 
| 209 | 1352 | i.e., does not raise any error message. That means if you use it to unfold a | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1353 | definition for a constant and this constant is not present in the goal state, | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1354 | you can still safely apply the simplifier. | 
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1355 | |
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1356 |   \footnote{\bf FIXME: show rewriting of cterms}
 | 
| 308 
c90f4ec30d43
improvements from the workshop
 Christian Urban <urbanc@in.tum.de> parents: 
307diff
changeset | 1357 | |
| 412 | 1358 | There is one restriction you have to keep in mind when using the simplifier: | 
| 413 | 1359 | it can only deal with rewriting rules whose left-hand sides are higher order | 
| 1360 |   pattern (see Section \ref{sec:univ} on unification). Whether a term is a pattern
 | |
| 1361 |   or not can be tested with the function @{ML_ind pattern in Pattern} from
 | |
| 1362 |   the structure @{ML_struct Pattern}. If a rule is not a pattern and you want
 | |
| 1363 | to use it for rewriting, then you have to use simprocs or conversions, both | |
| 1364 | of which we shall describe in the next section. | |
| 412 | 1365 | |
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1366 | When using the simplifier, the crucial information you have to provide is | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1367 | the simpset. If this information is not handled with care, then, as | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1368 | mentioned above, the simplifier can easily run into a loop. Therefore a good | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1369 | rule of thumb is to use simpsets that are as minimal as possible. It might | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1370 | be surprising that a simpset is more complex than just a simple collection | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1371 | of theorems. One reason for the complexity is that the simplifier must be | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1372 | able to rewrite inside terms and should also be able to rewrite according to | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1373 | theorems that have premises. | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1374 | |
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1375 | The rewriting inside terms requires congruence theorems, which | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1376 | are typically meta-equalities of the form | 
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1377 | |
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1378 |   \begin{isabelle}
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1379 |   \begin{center}
 | 
| 551 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 1380 |   \mbox{\inferrule{@{text "t\<^sub>1 \<equiv> s\<^sub>1 \<dots> t\<^sub>n \<equiv> s\<^sub>n"}}
 | 
| 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 1381 |                   {@{text "constr t\<^sub>1\<dots>t\<^sub>n \<equiv> constr s\<^sub>1\<dots>s\<^sub>n"}}}
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1382 |   \end{center}
 | 
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1383 |   \end{isabelle}
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1384 | |
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1385 |   with @{text "constr"} being a constant, like @{const "If"}, @{const "Let"}
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1386 | and so on. Every simpset contains only one congruence rule for each | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1387 | term-constructor, which however can be overwritten. The user can declare | 
| 370 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1388 |   lemmas to be congruence rules using the attribute @{text "[cong]"}. Note that 
 | 
| 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1389 | in HOL these congruence theorems are usually stated as equations, which are | 
| 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1390 | then internally transformed into meta-equations. | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1391 | |
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1392 | The rewriting with theorems involving premises requires what is in Isabelle | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1393 | called a subgoaler, a solver and a looper. These can be arbitrary tactics | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1394 | that can be installed in a simpset and which are executed at various stages | 
| 370 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1395 | during simplification. | 
| 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1396 | |
| 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1397 | Simpsets can also include simprocs, which produce rewrite rules on | 
| 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1398 | demand according to a pattern (see next section for a detailed description | 
| 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1399 | of simpsets). Another component are split-rules, which can simplify for | 
| 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1400 | example the ``then'' and ``else'' branches of if-statements under the | 
| 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1401 | corresponding preconditions. | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1402 | |
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1403 |   \begin{readmore}
 | 
| 458 
242e81f4d461
updated to post-2011 Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
456diff
changeset | 1404 |   For more information about the simplifier see @{ML_file "Pure/raw_simplifier.ML"}
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1405 |   and @{ML_file "Pure/simplifier.ML"}. The generic splitter is implemented in 
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1406 |   @{ML_file "Provers/splitter.ML"}.
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1407 |   \end{readmore}
 | 
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1408 | |
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1409 | |
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1410 |   \footnote{\bf FIXME: Find the right place to mention this: Discrimination 
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1411 |   nets are implemented in @{ML_file "Pure/net.ML"}.}
 | 
| 160 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 1412 | |
| 370 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1413 | The most common combinators for modifying simpsets are: | 
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1414 | |
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1415 |   \begin{isabelle}
 | 
| 370 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1416 |   \begin{tabular}{l@ {\hspace{10mm}}l}
 | 
| 458 
242e81f4d461
updated to post-2011 Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
456diff
changeset | 1417 |   @{ML_ind addsimps in Raw_Simplifier}    & @{ML_ind delsimps in Raw_Simplifier}\\
 | 
| 
242e81f4d461
updated to post-2011 Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
456diff
changeset | 1418 |   @{ML_ind addsimprocs in Raw_Simplifier} & @{ML_ind delsimprocs in Raw_Simplifier}\\
 | 
| 503 
3778bddfb824
updated to Isabelle 24 November
 Christian Urban <urbanc@in.tum.de> parents: 
489diff
changeset | 1419 |   @{ML_ind add_cong in Raw_Simplifier}    & @{ML_ind del_cong in Raw_Simplifier}\\
 | 
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1420 |   \end{tabular}
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1421 |   \end{isabelle}
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1422 | |
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1423 | *} | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1424 | |
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1425 | text_raw {*
 | 
| 173 
d820cb5873ea
used latex package boxedminipage
 Christian Urban <urbanc@in.tum.de> parents: 
172diff
changeset | 1426 | \begin{figure}[t]
 | 
| 177 | 1427 | \begin{minipage}{\textwidth}
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1428 | \begin{isabelle}*}
 | 
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 1429 | ML %grayML{*fun print_ss ctxt ss =
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1430 | let | 
| 458 
242e81f4d461
updated to post-2011 Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
456diff
changeset | 1431 |   val {simps, congs, procs, ...} = Raw_Simplifier.dest_ss ss
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1432 | |
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1433 | fun name_sthm (nm, thm) = | 
| 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1434 | Pretty.enclose (nm ^ ": ") "" [pretty_thm_no_vars ctxt thm] | 
| 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1435 | fun name_cthm ((_, nm), thm) = | 
| 440 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
437diff
changeset | 1436 | Pretty.enclose (nm ^ ": ") "" [pretty_thm_no_vars ctxt thm] | 
| 163 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1437 | fun name_ctrm (nm, ctrm) = | 
| 440 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
437diff
changeset | 1438 | Pretty.enclose (nm ^ ": ") "" [pretty_cterms ctxt ctrm] | 
| 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
437diff
changeset | 1439 | |
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1440 | val pps = [Pretty.big_list "Simplification rules:" (map name_sthm simps), | 
| 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1441 | Pretty.big_list "Congruences rules:" (map name_cthm congs), | 
| 440 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
437diff
changeset | 1442 | Pretty.big_list "Simproc patterns:" (map name_ctrm procs)] | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1443 | in | 
| 440 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
437diff
changeset | 1444 | pps |> Pretty.chunks | 
| 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
437diff
changeset | 1445 | |> pwriteln | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1446 | end*} | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1447 | text_raw {* 
 | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1448 | \end{isabelle}
 | 
| 177 | 1449 | \end{minipage}
 | 
| 458 
242e81f4d461
updated to post-2011 Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
456diff
changeset | 1450 | \caption{The function @{ML_ind dest_ss in Raw_Simplifier} returns a record containing
 | 
| 163 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1451 | all printable information stored in a simpset. We are here only interested in the | 
| 231 | 1452 |   simplification rules, congruence rules and simprocs.\label{fig:printss}}
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1453 | \end{figure} *}
 | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1454 | |
| 318 
efb5fff99c96
split up the first-steps section into two chapters
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 1455 | |
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1456 | text {* 
 | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
184diff
changeset | 1457 |   To see how they work, consider the function in Figure~\ref{fig:printss}, which
 | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
184diff
changeset | 1458 | prints out some parts of a simpset. If you use it to print out the components of the | 
| 458 
242e81f4d461
updated to post-2011 Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
456diff
changeset | 1459 |   empty simpset, i.e., @{ML_ind empty_ss in Raw_Simplifier}
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1460 | |
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1461 |   @{ML_response_fake [display,gray]
 | 
| 163 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1462 |   "print_ss @{context} empty_ss"
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1463 | "Simplification rules: | 
| 163 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1464 | Congruences rules: | 
| 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1465 | Simproc patterns:"} | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1466 | |
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1467 | you can see it contains nothing. This simpset is usually not useful, except as a | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1468 |   building block to build bigger simpsets. For example you can add to @{ML empty_ss} 
 | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1469 |   the simplification rule @{thm [source] Diff_Int} as follows:
 | 
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1470 | *} | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1471 | |
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1472 | ML %grayML{*val ss1 = put_simpset empty_ss @{context} addsimps [@{thm Diff_Int} RS @{thm eq_reflection}] *}
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1473 | |
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1474 | text {*
 | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1475 | Printing then out the components of the simpset gives: | 
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1476 | |
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1477 |   @{ML_response_fake [display,gray]
 | 
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1478 |   "print_ss @{context} (Raw_Simplifier.simpset_of ss1)"
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1479 | "Simplification rules: | 
| 158 
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
157diff
changeset | 1480 | ??.unknown: A - B \<inter> C \<equiv> A - B \<union> (A - C) | 
| 163 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1481 | Congruences rules: | 
| 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1482 | Simproc patterns:"} | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1483 | |
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1484 |   \footnote{\bf FIXME: Why does it print out ??.unknown}
 | 
| 158 
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
157diff
changeset | 1485 | |
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1486 |   Adding also the congruence rule @{thm [source] UN_cong} 
 | 
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1487 | *} | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1488 | |
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1489 | ML %grayML{*val ss2 = ss1 |> Simplifier.add_cong (@{thm UN_cong} RS @{thm eq_reflection}) *}
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1490 | |
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1491 | text {*
 | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1492 | gives | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1493 | |
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1494 |   @{ML_response_fake [display,gray]
 | 
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1495 |   "print_ss @{context} (Raw_Simplifier.simpset_of ss2)"
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1496 | "Simplification rules: | 
| 158 
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
157diff
changeset | 1497 | ??.unknown: A - B \<inter> C \<equiv> A - B \<union> (A - C) | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1498 | Congruences rules: | 
| 163 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1499 | UNION: \<lbrakk>A = B; \<And>x. x \<in> B \<Longrightarrow> C x = D x\<rbrakk> \<Longrightarrow> \<Union>x\<in>A. C x \<equiv> \<Union>x\<in>B. D x | 
| 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1500 | Simproc patterns:"} | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1501 | |
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1502 |   Notice that we had to add these lemmas as meta-equations. The @{ML empty_ss} 
 | 
| 370 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1503 | expects this form of the simplification and congruence rules. This is | 
| 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1504 |   different, if we use for example the simpset @{ML HOL_basic_ss} (see below), 
 | 
| 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1505 | where rules are usually added as equation. However, even | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1506 |   when adding these lemmas to @{ML empty_ss} we do not end up with anything useful yet.
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1507 |   In the context of HOL, the first really useful simpset is @{ML_ind
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1508 | HOL_basic_ss in Simpdata}. While printing out the components of this simpset | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1509 | |
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1510 |   @{ML_response_fake [display,gray]
 | 
| 163 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1511 |   "print_ss @{context} HOL_basic_ss"
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1512 | "Simplification rules: | 
| 163 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1513 | Congruences rules: | 
| 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1514 | Simproc patterns:"} | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1515 | |
| 370 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1516 | also produces ``nothing'', the printout is somewhat misleading. In fact | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1517 |   the @{ML HOL_basic_ss} is setup so that it can solve goals of the
 | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
184diff
changeset | 1518 | form | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
184diff
changeset | 1519 | |
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
184diff
changeset | 1520 |   \begin{isabelle}
 | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
184diff
changeset | 1521 |   @{thm TrueI}, @{thm refl[no_vars]}, @{term "t \<equiv> t"} and @{thm FalseE[no_vars]}; 
 | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
184diff
changeset | 1522 |   \end{isabelle}
 | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
184diff
changeset | 1523 | |
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1524 | and also resolve with assumptions. For example: | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1525 | *} | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1526 | |
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1527 | lemma | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1528 | shows "True" | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1529 | and "t = t" | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1530 | and "t \<equiv> t" | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1531 | and "False \<Longrightarrow> Foo" | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1532 | and "\<lbrakk>A; B; C\<rbrakk> \<Longrightarrow> A" | 
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1533 | apply(tactic {* ALLGOALS (simp_tac (put_simpset HOL_basic_ss @{context})) *})
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1534 | done | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1535 | |
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1536 | text {*
 | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1537 | This behaviour is not because of simplification rules, but how the subgoaler, solver | 
| 369 | 1538 |   and looper are set up in @{ML HOL_basic_ss}.
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1539 | |
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1540 |   The simpset @{ML_ind HOL_ss} is an extension of @{ML HOL_basic_ss} containing 
 | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1541 | already many useful simplification and congruence rules for the logical | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1542 | connectives in HOL. | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1543 | |
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1544 |   @{ML_response_fake [display,gray]
 | 
| 163 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1545 |   "print_ss @{context} HOL_ss"
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1546 | "Simplification rules: | 
| 158 
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
157diff
changeset | 1547 | Pure.triv_forall_equality: (\<And>x. PROP V) \<equiv> PROP V | 
| 
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
157diff
changeset | 1548 | HOL.the_eq_trivial: THE x. x = y \<equiv> y | 
| 
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
157diff
changeset | 1549 | HOL.the_sym_eq_trivial: THE ya. y = ya \<equiv> y | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1550 | \<dots> | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1551 | Congruences rules: | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1552 | HOL.simp_implies: \<dots> | 
| 158 
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
157diff
changeset | 1553 | \<Longrightarrow> (PROP P =simp=> PROP Q) \<equiv> (PROP P' =simp=> PROP Q') | 
| 163 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1554 | op -->: \<lbrakk>P \<equiv> P'; P' \<Longrightarrow> Q \<equiv> Q'\<rbrakk> \<Longrightarrow> P \<longrightarrow> Q \<equiv> P' \<longrightarrow> Q' | 
| 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1555 | Simproc patterns: | 
| 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1556 | \<dots>"} | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1557 | |
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1558 |   \begin{readmore}
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1559 |   The simplifier for HOL is set up in @{ML_file "HOL/Tools/simpdata.ML"}.
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1560 |   The simpset @{ML HOL_ss} is implemented in @{ML_file "HOL/HOL.thy"}.
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1561 |   \end{readmore}
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1562 | |
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1563 | The simplifier is often used to unfold definitions in a proof. For this the | 
| 458 
242e81f4d461
updated to post-2011 Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
456diff
changeset | 1564 |   simplifier implements the tactic @{ML_ind rewrite_goal_tac in Raw_Simplifier}.\footnote{\bf FIXME: 
 | 
| 243 | 1565 | see LocalDefs infrastructure.} Suppose for example the | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1566 | definition | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1567 | *} | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1568 | |
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1569 | definition "MyTrue \<equiv> True" | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1570 | |
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
184diff
changeset | 1571 | text {*
 | 
| 370 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1572 | then we can use this tactic to unfold the definition of this constant. | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
184diff
changeset | 1573 | *} | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
184diff
changeset | 1574 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1575 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1576 | shows "MyTrue \<Longrightarrow> True" | 
| 552 
82c482467d75
updated to latest isabelle
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
551diff
changeset | 1577 | apply(tactic {* rewrite_goal_tac @{context} @{thms MyTrue_def} 1 *})
 | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
184diff
changeset | 1578 | txt{* producing the goal state
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1579 |       \begin{isabelle}
 | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1580 |       @{subgoals [display]}
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1581 |       \end{isabelle} *}
 | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1582 | (*<*)oops(*>*) | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1583 | |
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1584 | text {*
 | 
| 370 
2494b5b7a85d
added something about show_types references
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 1585 |   If you want to unfold definitions in \emph{all} subgoals, not just one, 
 | 
| 458 
242e81f4d461
updated to post-2011 Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
456diff
changeset | 1586 |   then use the the tactic @{ML_ind rewrite_goals_tac in Raw_Simplifier}.
 | 
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1587 | *} | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1588 | |
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1589 | |
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1590 | text_raw {*
 | 
| 173 
d820cb5873ea
used latex package boxedminipage
 Christian Urban <urbanc@in.tum.de> parents: 
172diff
changeset | 1591 | \begin{figure}[p]
 | 
| 
d820cb5873ea
used latex package boxedminipage
 Christian Urban <urbanc@in.tum.de> parents: 
172diff
changeset | 1592 | \begin{boxedminipage}{\textwidth}
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1593 | \begin{isabelle} *}
 | 
| 475 
25371f74c768
updated to post-2011-1 Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
466diff
changeset | 1594 | type_synonym prm = "(nat \<times> nat) list" | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1595 | consts perm :: "prm \<Rightarrow> 'a \<Rightarrow> 'a"  ("_ \<bullet> _" [80,80] 80)
 | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1596 | |
| 229 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1597 | overloading | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1598 | perm_nat \<equiv> "perm :: prm \<Rightarrow> nat \<Rightarrow> nat" | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1599 |   perm_prod \<equiv> "perm :: prm \<Rightarrow> ('a\<times>'b) \<Rightarrow> ('a\<times>'b)"
 | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1600 | perm_list \<equiv> "perm :: prm \<Rightarrow> 'a list \<Rightarrow> 'a list" | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1601 | begin | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1602 | |
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1603 | fun swap::"nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1604 | where | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1605 | "swap a b c = (if c=a then b else (if c=b then a else c))" | 
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1606 | |
| 229 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1607 | primrec perm_nat | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1608 | where | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1609 | "perm_nat [] c = c" | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1610 | | "perm_nat (ab#pi) c = swap (fst ab) (snd ab) (perm_nat pi c)" | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1611 | |
| 229 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1612 | fun perm_prod | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1613 | where | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1614 | "perm_prod pi (x, y) = (pi\<bullet>x, pi\<bullet>y)" | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1615 | |
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1616 | primrec perm_list | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1617 | where | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1618 | "perm_list pi [] = []" | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1619 | | "perm_list pi (x#xs) = (pi\<bullet>x)#(perm_list pi xs)" | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1620 | |
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1621 | end | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1622 | |
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1623 | lemma perm_append[simp]: | 
| 551 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 1624 | fixes c::"nat" and pi\<^sub>1 pi\<^sub>2::"prm" | 
| 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 1625 | shows "((pi\<^sub>1@pi\<^sub>2)\<bullet>c) = (pi\<^sub>1\<bullet>(pi\<^sub>2\<bullet>c))" | 
| 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 1626 | by (induct pi\<^sub>1) (auto) | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1627 | |
| 229 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1628 | lemma perm_bij[simp]: | 
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1629 | fixes c d::"nat" and pi::"prm" | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1630 | shows "(pi\<bullet>c = pi\<bullet>d) = (c = d)" | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1631 | by (induct pi) (auto) | 
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1632 | |
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1633 | lemma perm_rev[simp]: | 
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1634 | fixes c::"nat" and pi::"prm" | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1635 | shows "pi\<bullet>((rev pi)\<bullet>c) = c" | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1636 | by (induct pi arbitrary: c) (auto) | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1637 | |
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1638 | lemma perm_compose: | 
| 551 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 1639 | fixes c::"nat" and pi\<^sub>1 pi\<^sub>2::"prm" | 
| 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 1640 | shows "pi\<^sub>1\<bullet>(pi\<^sub>2\<bullet>c) = (pi\<^sub>1\<bullet>pi\<^sub>2)\<bullet>(pi\<^sub>1\<bullet>c)" | 
| 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 1641 | by (induct pi\<^sub>2) (auto) | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1642 | text_raw {*
 | 
| 173 
d820cb5873ea
used latex package boxedminipage
 Christian Urban <urbanc@in.tum.de> parents: 
172diff
changeset | 1643 | \end{isabelle}
 | 
| 
d820cb5873ea
used latex package boxedminipage
 Christian Urban <urbanc@in.tum.de> parents: 
172diff
changeset | 1644 | \end{boxedminipage}
 | 
| 229 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1645 | \caption{A simple theory about permutations over @{typ nat}s. The point is that the
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1646 |   lemma @{thm [source] perm_compose} cannot be directly added to the simplifier, as
 | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1647 | it would cause the simplifier to loop. It can still be used as a simplification | 
| 229 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1648 | rule if the permutation in the right-hand side is sufficiently protected. | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1649 |   \label{fig:perms}}
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1650 | \end{figure} *}
 | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1651 | |
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1652 | |
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1653 | text {*
 | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1654 | The simplifier is often used in order to bring terms into a normal form. | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1655 | Unfortunately, often the situation arises that the corresponding | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1656 | simplification rules will cause the simplifier to run into an infinite | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1657 | loop. Consider for example the simple theory about permutations over natural | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1658 |   numbers shown in Figure~\ref{fig:perms}. The purpose of the lemmas is to
 | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1659 | push permutations as far inside as possible, where they might disappear by | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1660 |   Lemma @{thm [source] perm_rev}. However, to fully normalise all instances,
 | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1661 |   it would be desirable to add also the lemma @{thm [source] perm_compose} to
 | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1662 | the simplifier for pushing permutations over other permutations. Unfortunately, | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1663 | the right-hand side of this lemma is again an instance of the left-hand side | 
| 209 | 1664 | and so causes an infinite loop. There seems to be no easy way to reformulate | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1665 | this rule and so one ends up with clunky proofs like: | 
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1666 | *} | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1667 | |
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1668 | lemma | 
| 551 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 1669 | fixes c d::"nat" and pi\<^sub>1 pi\<^sub>2::"prm" | 
| 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 1670 | shows "pi\<^sub>1\<bullet>(c, pi\<^sub>2\<bullet>((rev pi\<^sub>1)\<bullet>d)) = (pi\<^sub>1\<bullet>c, (pi\<^sub>1\<bullet>pi\<^sub>2)\<bullet>d)" | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1671 | apply(simp) | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1672 | apply(rule trans) | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1673 | apply(rule perm_compose) | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1674 | apply(simp) | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1675 | done | 
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1676 | |
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1677 | text {*
 | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1678 | It is however possible to create a single simplifier tactic that solves | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1679 | such proofs. The trick is to introduce an auxiliary constant for permutations | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1680 | and split the simplification into two phases (below actually three). Let | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1681 | assume the auxiliary constant is | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1682 | *} | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1683 | |
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1684 | definition | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1685 |   perm_aux :: "prm \<Rightarrow> 'a \<Rightarrow> 'a" ("_ \<bullet>aux _" [80,80] 80)
 | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1686 | where | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1687 | "pi \<bullet>aux c \<equiv> pi \<bullet> c" | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1688 | |
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1689 | text {* Now the two lemmas *}
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1690 | |
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1691 | lemma perm_aux_expand: | 
| 551 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 1692 | fixes c::"nat" and pi\<^sub>1 pi\<^sub>2::"prm" | 
| 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 1693 | shows "pi\<^sub>1\<bullet>(pi\<^sub>2\<bullet>c) = pi\<^sub>1 \<bullet>aux (pi\<^sub>2\<bullet>c)" | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1694 | unfolding perm_aux_def by (rule refl) | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1695 | |
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1696 | lemma perm_compose_aux: | 
| 551 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 1697 | fixes c::"nat" and pi\<^sub>1 pi\<^sub>2::"prm" | 
| 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 1698 | shows "pi\<^sub>1\<bullet>(pi\<^sub>2\<bullet>aux c) = (pi\<^sub>1\<bullet>pi\<^sub>2) \<bullet>aux (pi\<^sub>1\<bullet>c)" | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1699 | unfolding perm_aux_def by (rule perm_compose) | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1700 | |
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1701 | text {* 
 | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1702 |   are simple consequence of the definition and @{thm [source] perm_compose}. 
 | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1703 |   More importantly, the lemma @{thm [source] perm_compose_aux} can be safely 
 | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1704 | added to the simplifier, because now the right-hand side is not anymore an instance | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1705 | of the left-hand side. In a sense it freezes all redexes of permutation compositions | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1706 | after one step. In this way, we can split simplification of permutations | 
| 213 | 1707 | into three phases without the user noticing anything about the auxiliary | 
| 231 | 1708 | constant. We first freeze any instance of permutation compositions in the term using | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1709 |   lemma @{thm [source] "perm_aux_expand"} (Line 9);
 | 
| 231 | 1710 | then simplify all other permutations including pushing permutations over | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1711 |   other permutations by rule @{thm [source] perm_compose_aux} (Line 10); and
 | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1712 | finally ``unfreeze'' all instances of permutation compositions by unfolding | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1713 | the definition of the auxiliary constant. | 
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1714 | *} | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1715 | |
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1716 | ML %linenosgray{*fun perm_simp_tac ctxt =
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1717 | let | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1718 |   val thms1 = [@{thm perm_aux_expand}]
 | 
| 229 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
219diff
changeset | 1719 |   val thms2 = [@{thm perm_append}, @{thm perm_bij}, @{thm perm_rev}, 
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1720 |                @{thm perm_compose_aux}] @ @{thms perm_prod.simps} @ 
 | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1721 |                @{thms perm_list.simps} @ @{thms perm_nat.simps}
 | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1722 |   val thms3 = [@{thm perm_aux_def}]
 | 
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1723 | val ss = put_simpset HOL_basic_ss ctxt | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1724 | in | 
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1725 | simp_tac (ss addsimps thms1) | 
| 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1726 | THEN' simp_tac (ss addsimps thms2) | 
| 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1727 | THEN' simp_tac (ss addsimps thms3) | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1728 | end*} | 
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
152diff
changeset | 1729 | |
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1730 | text {*
 | 
| 209 | 1731 | For all three phases we have to build simpsets adding specific lemmas. As is sufficient | 
| 232 | 1732 |   for our purposes here, we can add these lemmas to @{ML HOL_basic_ss} in order to obtain
 | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1733 | the desired results. Now we can solve the following lemma | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1734 | *} | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1735 | |
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1736 | lemma | 
| 551 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 1737 | fixes c d::"nat" and pi\<^sub>1 pi\<^sub>2::"prm" | 
| 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 1738 | shows "pi\<^sub>1\<bullet>(c, pi\<^sub>2\<bullet>((rev pi\<^sub>1)\<bullet>d)) = (pi\<^sub>1\<bullet>c, (pi\<^sub>1\<bullet>pi\<^sub>2)\<bullet>d)" | 
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1739 | apply(tactic {* perm_simp_tac @{context} 1 *})
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1740 | done | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1741 | |
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1742 | |
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1743 | text {*
 | 
| 209 | 1744 | in one step. This tactic can deal with most instances of normalising permutations. | 
| 1745 | In order to solve all cases we have to deal with corner-cases such as the | |
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1746 | lemma being an exact instance of the permutation composition lemma. This can | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1747 | often be done easier by implementing a simproc or a conversion. Both will be | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1748 | explained in the next two chapters. | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1749 | |
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1750 |   (FIXME: Is it interesting to say something about @{term "op =simp=>"}?)
 | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1751 | |
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1752 | (FIXME: What are the second components of the congruence rules---something to | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1753 | do with weak congruence constants?) | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1754 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
239diff
changeset | 1755 |   (FIXME: what are @{ML mksimps_pairs}; used in Nominal.thy)
 | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
239diff
changeset | 1756 | |
| 250 
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
 Christian Urban <urbanc@in.tum.de> parents: 
243diff
changeset | 1757 |   (FIXME: explain @{ML simplify} and @{ML "Simplifier.rewrite_rule"} etc.)
 | 
| 
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
 Christian Urban <urbanc@in.tum.de> parents: 
243diff
changeset | 1758 | |
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1759 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1760 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1761 | section {* Simprocs *}
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1762 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1763 | text {*
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1764 | In Isabelle you can also implement custom simplification procedures, called | 
| 149 | 1765 |   \emph{simprocs}. Simprocs can be triggered by the simplifier on a specified
 | 
| 1766 | term-pattern and rewrite a term according to a theorem. They are useful in | |
| 1767 | cases where a rewriting rule must be produced on ``demand'' or when | |
| 1768 | 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 | 1769 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1770 | To see how simprocs work, let us first write a simproc that just prints out | 
| 132 | 1771 | 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 | 1772 | you can use the function: | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1773 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1774 | |
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1775 | ML %linenosgray{*fun fail_simproc ctxt redex = 
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1776 | let | 
| 543 
cd28458c2add
tuned
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
541diff
changeset | 1777 | val _ = pwriteln (Pretty.block | 
| 
cd28458c2add
tuned
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
541diff
changeset | 1778 | [Pretty.str "The redex: ", pretty_cterm ctxt redex]) | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1779 | in | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1780 | NONE | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1781 | end*} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1782 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1783 | text {*
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1784 |   This function takes a simpset and a redex (a @{ML_type cterm}) as
 | 
| 132 | 1785 | 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 | 1786 | 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 | 1787 |   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 | 1788 |   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 | 1789 |   that the simproc is triggered by the pattern @{term "Suc n"}. This can be
 | 
| 149 | 1790 | 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 | 1791 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1792 | |
| 243 | 1793 | simproc_setup %gray fail ("Suc n") = {* K fail_simproc *}
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1794 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1795 | text {*
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1796 | where the second argument specifies the pattern and the right-hand side | 
| 232 | 1797 |   contains the code of the simproc (we have to use @{ML K} since we are ignoring
 | 
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 1798 | an argument about morphisms. | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1799 | After this, the simplifier is aware of the simproc and you can test whether | 
| 131 | 1800 | it fires on the lemma: | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1801 | *} | 
| 120 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: 
118diff
changeset | 1802 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1803 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1804 | shows "Suc 0 = 1" | 
| 178 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1805 | apply(simp) | 
| 378 
8d160d79b48c
section about matching and unification of types
 Christian Urban <urbanc@in.tum.de> parents: 
370diff
changeset | 1806 | txt{*
 | 
| 
8d160d79b48c
section about matching and unification of types
 Christian Urban <urbanc@in.tum.de> parents: 
370diff
changeset | 1807 |   \begin{minipage}{\textwidth}
 | 
| 
8d160d79b48c
section about matching and unification of types
 Christian Urban <urbanc@in.tum.de> parents: 
370diff
changeset | 1808 |   \small@{text "> The redex: Suc 0"}\\
 | 
| 213 | 1809 |   @{text "> The redex: Suc 0"}\\
 | 
| 378 
8d160d79b48c
section about matching and unification of types
 Christian Urban <urbanc@in.tum.de> parents: 
370diff
changeset | 1810 |   \end{minipage}*}(*<*)oops(*>*)
 | 
| 
8d160d79b48c
section about matching and unification of types
 Christian Urban <urbanc@in.tum.de> parents: 
370diff
changeset | 1811 | |
| 
8d160d79b48c
section about matching and unification of types
 Christian Urban <urbanc@in.tum.de> parents: 
370diff
changeset | 1812 | text {* 
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1813 | 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 | 1814 | 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 | 1815 |   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 | 1816 | 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 | 1817 | |
| 131 | 1818 | We can add or delete the simproc from the current simpset by the usual | 
| 132 | 1819 |   \isacommand{declare}-statement. For example the simproc will be deleted
 | 
| 1820 | with the declaration | |
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1821 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1822 | |
| 243 | 1823 | declare [[simproc del: fail]] | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1824 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1825 | text {*
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1826 |   If you want to see what happens with just \emph{this} simproc, without any 
 | 
| 243 | 1827 |   interference from other rewrite rules, you can call @{text fail} 
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1828 | as follows: | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1829 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1830 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1831 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1832 | shows "Suc 0 = 1" | 
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1833 | apply(tactic {* simp_tac (put_simpset 
 | 
| 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1834 |   HOL_basic_ss @{context} addsimprocs [@{simproc fail}]) 1*})
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1835 | (*<*)oops(*>*) | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1836 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1837 | text {*
 | 
| 131 | 1838 |   Now the message shows up only once since the term @{term "1::nat"} is 
 | 
| 1839 | left unchanged. | |
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1840 | |
| 178 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1841 |   Setting up a simproc using the command \isacommand{simproc\_setup} will
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1842 | 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 | 1843 | want this, then you have to use a slightly different method for setting | 
| 243 | 1844 |   up the simproc. First the function @{ML fail_simproc} needs to be modified
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1845 | to | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1846 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1847 | |
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1848 | ML %grayML{*fun fail_simproc' ctxt redex = 
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1849 | let | 
| 543 
cd28458c2add
tuned
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
541diff
changeset | 1850 | val _ = pwriteln (Pretty.block | 
| 
cd28458c2add
tuned
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
541diff
changeset | 1851 | [Pretty.str "The redex:", pretty_term ctxt redex]) | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1852 | in | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1853 | NONE | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1854 | end*} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1855 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1856 | text {*
 | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1857 |   Here the redex is given as a @{ML_type term}, instead of a @{ML_type cterm}
 | 
| 441 | 1858 |   (therefore we printing it out using the function @{ML pretty_term in Pretty}).
 | 
| 149 | 1859 | We can turn this function into a proper simproc using the function | 
| 449 | 1860 |   @{ML Simplifier.simproc_global_i}:
 | 
| 93 | 1861 | *} | 
| 1862 | ||
| 105 
f49dc7e96235
added more to the Tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 1863 | |
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1864 | ML %grayML{*fun fail' ctxt = 
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1865 | let | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1866 |   val thy = @{theory}
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1867 |   val pat = [@{term "Suc n"}]
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1868 | in | 
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1869 | Simplifier.simproc_global_i thy "fail_simproc'" pat (K fail_simproc' ctxt) | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1870 | end*} | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1871 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1872 | text {*
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1873 |   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 | 1874 | The function also takes a list of patterns that can trigger the simproc. | 
| 132 | 1875 | Now the simproc is set up and can be explicitly added using | 
| 458 
242e81f4d461
updated to post-2011 Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
456diff
changeset | 1876 |   @{ML_ind addsimprocs in Raw_Simplifier} to a simpset whenever
 | 
| 132 | 1877 | needed. | 
| 1878 | ||
| 1879 | Simprocs are applied from inside to outside and from left to right. You can | |
| 1880 | see this in the proof | |
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1881 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1882 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1883 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1884 | shows "Suc (Suc 0) = (Suc 1)" | 
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1885 | apply(tactic {* simp_tac ((put_simpset HOL_basic_ss @{context}) addsimprocs [fail' @{context}]) 1*})
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1886 | (*<*)oops(*>*) | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1887 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1888 | text {*
 | 
| 243 | 1889 |   The simproc @{ML fail'} prints out the sequence 
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1890 | |
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1891 | @{text [display]
 | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1892 | "> Suc 0 | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1893 | > Suc (Suc 0) | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1894 | > Suc 1"} | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1895 | |
| 131 | 1896 | 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 | 1897 | rewrites terms according to the equation: | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1898 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1899 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1900 | lemma plus_one: | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1901 | shows "Suc n \<equiv> n + 1" by simp | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1902 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1903 | text {*
 | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1904 | Simprocs expect that the given equation is a meta-equation, however the | 
| 131 | 1905 | equation can contain preconditions (the simproc then will only fire if the | 
| 132 | 1906 | preconditions can be solved). To see that one has relatively precise control over | 
| 131 | 1907 | the rewriting with simprocs, let us further assume we want that the simproc | 
| 1908 |   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 | 1909 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1910 | |
| 131 | 1911 | |
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1912 | ML %grayML{*fun plus_one_simproc ctxt redex =
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1913 | case redex of | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1914 |     @{term "Suc 0"} => NONE
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1915 |   | _ => SOME @{thm plus_one}*}
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1916 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1917 | text {*
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1918 | and set up the simproc as follows. | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1919 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1920 | |
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1921 | ML %grayML{*fun plus_one ctxt =
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1922 | let | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1923 |   val thy = @{theory}
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1924 |   val pat = [@{term "Suc n"}] 
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1925 | in | 
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1926 | Simplifier.simproc_global_i thy "sproc +1" pat (K plus_one_simproc ctxt) | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 1927 | end*} | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1928 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1929 | text {*
 | 
| 132 | 1930 | 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 | 1931 |   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 | 1932 |   a theorem if the term is not @{term "Suc 0"}. The result you can see
 | 
| 131 | 1933 | in the following proof | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1934 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1935 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1936 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 1937 | shows "P (Suc (Suc (Suc 0))) (Suc 0)" | 
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1938 | apply(tactic {* simp_tac (put_simpset HOL_basic_ss @{context} addsimprocs [plus_one @{context}]) 1*})
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1939 | txt{*
 | 
| 131 | 1940 | where the simproc produces the goal state | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1941 |   \begin{isabelle}
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1942 |   @{subgoals[display]}
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 1943 |   \end{isabelle}
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1944 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1945 | (*<*)oops(*>*) | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1946 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1947 | text {*
 | 
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1948 | As usual with rewriting you have to worry about looping: you already have | 
| 243 | 1949 |   a loop with @{ML plus_one}, if you apply it with the default simpset (because
 | 
| 1950 |   the default simpset contains a rule which just does the opposite of @{ML plus_one},
 | |
| 132 | 1951 |   namely rewriting @{text [quotes] "+ 1"} to a successor). So you have to be careful 
 | 
| 1952 | 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 | 1953 | |
| 132 | 1954 |   Next let us implement a simproc that replaces terms of the form @{term "Suc n"}
 | 
| 232 | 1955 |   with the number @{text n} increased by one. First we implement a function that
 | 
| 132 | 1956 | 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 | 1957 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1958 | |
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 1959 | ML %grayML{*fun dest_suc_trm ((Const (@{const_name "Suc"}, _)) $ t) = 1 + dest_suc_trm t
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1960 | | dest_suc_trm t = snd (HOLogic.dest_number t)*} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 1961 | |
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1962 | text {* 
 | 
| 316 
74f0a06f751f
further polishing of index generation
 Christian Urban <urbanc@in.tum.de> parents: 
315diff
changeset | 1963 |   It uses the library function @{ML_ind  dest_number in HOLogic} that transforms
 | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1964 |   (Isabelle) terms, like @{term "0::nat"}, @{term "1::nat"}, @{term "2::nat"} and so
 | 
| 131 | 1965 |   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 | 1966 | the term is not a number. The next function expects a pair consisting of a term | 
| 131 | 1967 |   @{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 | 1968 | *} | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1969 | |
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1970 | 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 | 1971 | let | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1972 |   val num = HOLogic.mk_number @{typ "nat"} n
 | 
| 132 | 1973 | 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 | 1974 | in | 
| 214 
7e04dc2368b0
updated to latest Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
213diff
changeset | 1975 | Goal.prove ctxt [] [] goal (K (Arith_Data.arith_tac ctxt 1)) | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1976 | end*} | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1977 | |
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1978 | text {*
 | 
| 132 | 1979 | From the integer value it generates the corresponding number term, called | 
| 1980 |   @{text num} (Line 3), and then generates the meta-equation @{text "t \<equiv> num"} 
 | |
| 1981 | (Line 4), which it proves by the arithmetic tactic in Line 6. | |
| 1982 | ||
| 219 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
217diff
changeset | 1983 |   For our purpose at the moment, proving the meta-equation using @{ML
 | 
| 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
217diff
changeset | 1984 | arith_tac in Arith_Data} is fine, but there is also an alternative employing | 
| 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
217diff
changeset | 1985 | the simplifier with a special simpset. For the kind of lemmas we | 
| 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
217diff
changeset | 1986 |   want to prove here, the simpset @{text "num_ss"} should suffice.
 | 
| 132 | 1987 | *} | 
| 131 | 1988 | |
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 1989 | ML %grayML{*fun get_thm_alt ctxt (t, n) =
 | 
| 132 | 1990 | let | 
| 1991 |   val num = HOLogic.mk_number @{typ "nat"} n
 | |
| 1992 | val goal = Logic.mk_equals (t, num) | |
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 1993 |   val num_ss = put_simpset HOL_ss ctxt addsimps @{thms semiring_norm}
 | 
| 132 | 1994 | in | 
| 1995 | Goal.prove ctxt [] [] goal (K (simp_tac num_ss 1)) | |
| 1996 | end*} | |
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 1997 | |
| 132 | 1998 | text {*
 | 
| 1999 |   The advantage of @{ML get_thm_alt} is that it leaves very little room for 
 | |
| 2000 | something to go wrong; in contrast it is much more difficult to predict | |
| 219 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
217diff
changeset | 2001 |   what happens with @{ML arith_tac in Arith_Data}, especially in more complicated 
 | 
| 231 | 2002 |   circumstances. The disadvantage of @{ML get_thm_alt} is to find a simpset
 | 
| 132 | 2003 | that is sufficiently powerful to solve every instance of the lemmas | 
| 2004 | we like to prove. This requires careful tuning, but is often necessary in | |
| 2005 |   ``production code''.\footnote{It would be of great help if there is another
 | |
| 2006 | way than tracing the simplifier to obtain the lemmas that are successfully | |
| 2007 | applied during simplification. Alas, there is none.} | |
| 2008 | ||
| 2009 | Anyway, either version can be used in the function that produces the actual | |
| 2010 | theorem for the simproc. | |
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 2011 | *} | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 2012 | |
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 2013 | ML %grayML{*fun nat_number_simproc ctxt t =
 | 
| 515 
4b105b97069c
updated to forthcoming Isabelle 2012
 Christian Urban <urbanc@in.tum.de> parents: 
513diff
changeset | 2014 | SOME (get_thm_alt ctxt (t, dest_suc_trm t)) | 
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 2015 | handle TERM _ => NONE *} | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 2016 | |
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 2017 | text {*
 | 
| 243 | 2018 |   This function uses the fact that @{ML dest_suc_trm} might raise an exception 
 | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 2019 |   @{ML TERM}. In this case there is nothing that can be rewritten and therefore no
 | 
| 131 | 2020 |   theorem is produced (i.e.~the function returns @{ML NONE}). To try out the simproc 
 | 
| 2021 | 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 | 2022 | *} | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 2023 | |
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 2024 | ML %grayML{*fun nat_number ctxt =
 | 
| 132 | 2025 | let | 
| 2026 |   val thy = @{theory}
 | |
| 2027 |   val pat = [@{term "Suc n"}]
 | |
| 2028 | in | |
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 2029 | Simplifier.simproc_global_i thy "nat_number" pat (K nat_number_simproc ctxt) | 
| 132 | 2030 | end*} | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 2031 | |
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 2032 | text {* 
 | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 2033 | Now in the lemma | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 2034 | *} | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 2035 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 2036 | lemma | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 2037 | shows "P (Suc (Suc 2)) (Suc 99) (0::nat) (Suc 4 + Suc 0) (Suc (0 + 0))" | 
| 544 
501491d56798
updated to simplifier change
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
543diff
changeset | 2038 | apply(tactic {* simp_tac (put_simpset HOL_ss @{context} addsimprocs [nat_number @{context}]) 1*})
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 2039 | txt {* 
 | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 2040 | you obtain the more legible goal state | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 2041 |   \begin{isabelle}
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 2042 |   @{subgoals [display]}
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 2043 |   \end{isabelle}*}
 | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 2044 | (*<*)oops(*>*) | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 2045 | |
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 2046 | text {*
 | 
| 132 | 2047 |   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 | 2048 |   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 | 2049 | 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 | 2050 | |
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 2051 |   \begin{exercise}\label{ex:addsimproc}
 | 
| 551 
be361e980acf
updated subscripts
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
544diff
changeset | 2052 |   Write a simproc that replaces terms of the form @{term "t\<^sub>1 + t\<^sub>2"} by their 
 | 
| 130 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 2053 | 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 | 2054 |   @{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 | 2055 |   \end{exercise}
 | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 2056 | |
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 2057 | (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 | 2058 | one can say about them?) | 
| 
a21d7b300616
polished the section about simprocs and added an exercise
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 2059 | *} | 
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 2060 | |
| 137 | 2061 | section {* Conversions\label{sec:conversion} *}
 | 
| 132 | 2062 | |
| 135 | 2063 | text {*
 | 
| 406 
f399b6855546
implemented suggestion from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
405diff
changeset | 2064 | Conversions are a thin layer on top of Isabelle's inference kernel, and can | 
| 
f399b6855546
implemented suggestion from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
405diff
changeset | 2065 | be viewed as a controllable, bare-bone version of Isabelle's simplifier. | 
| 412 | 2066 |   The purpose of conversions is to manipulate @{ML_type cterm}s. However,
 | 
| 406 
f399b6855546
implemented suggestion from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
405diff
changeset | 2067 | we will also show in this section how conversions can be applied to theorems | 
| 412 | 2068 | and to goal states. The type of conversions is | 
| 135 | 2069 | *} | 
| 2070 | ||
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 2071 | ML %grayML{*type conv = cterm -> thm*}
 | 
| 135 | 2072 | |
| 2073 | text {*
 | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2074 | whereby the produced theorem is always a meta-equality. A simple conversion | 
| 316 
74f0a06f751f
further polishing of index generation
 Christian Urban <urbanc@in.tum.de> parents: 
315diff
changeset | 2075 |   is the function @{ML_ind  all_conv in Conv}, which maps a @{ML_type cterm} to an
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2076 | instance of the (meta)reflexivity theorem. For example: | 
| 135 | 2077 | |
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 2078 |   @{ML_response_fake [display,gray]
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2079 |   "Conv.all_conv @{cterm \"Foo \<or> Bar\"}"
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2080 | "Foo \<or> Bar \<equiv> Foo \<or> Bar"} | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2081 | |
| 316 
74f0a06f751f
further polishing of index generation
 Christian Urban <urbanc@in.tum.de> parents: 
315diff
changeset | 2082 |   Another simple conversion is @{ML_ind  no_conv in Conv} which always raises the 
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2083 |   exception @{ML CTERM}.
 | 
| 135 | 2084 | |
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 2085 |   @{ML_response_fake [display,gray]
 | 
| 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 2086 |   "Conv.no_conv @{cterm True}" 
 | 
| 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 2087 | "*** Exception- CTERM (\"no conversion\", []) raised"} | 
| 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 2088 | |
| 316 
74f0a06f751f
further polishing of index generation
 Christian Urban <urbanc@in.tum.de> parents: 
315diff
changeset | 2089 |   A more interesting conversion is the function @{ML_ind  beta_conversion in Thm}: it 
 | 
| 160 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 2090 | produces a meta-equation between a term and its beta-normal form. For example | 
| 142 | 2091 | |
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 2092 |   @{ML_response_fake [display,gray]
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2093 | "let | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2094 |   val add = @{cterm \"\<lambda>x y. x + (y::nat)\"}
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2095 |   val two = @{cterm \"2::nat\"}
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2096 |   val ten = @{cterm \"10::nat\"}
 | 
| 513 | 2097 | val ctrm = Thm.apply (Thm.apply add two) ten | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2098 | in | 
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2099 | Thm.beta_conversion true ctrm | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2100 | end" | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2101 | "((\<lambda>x y. x + y) 2) 10 \<equiv> 2 + 10"} | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2102 | |
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2103 | If you run this example, you will notice that the actual response is the | 
| 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2104 |   seemingly nonsensical @{term
 | 
| 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2105 | "2 + 10 \<equiv> 2 + (10::nat)"}. The reason is that the pretty-printer for | 
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2106 |   @{ML_type cterm}s eta-normalises (sic) terms and therefore produces this output.
 | 
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2107 | If we get hold of the ``raw'' representation of the produced theorem, | 
| 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2108 | we obtain the expected result. | 
| 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2109 | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2110 |   @{ML_response [display,gray]
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2111 | "let | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2112 |   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 | 2113 |   val two = @{cterm \"2::nat\"}
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2114 |   val ten = @{cterm \"10::nat\"}
 | 
| 513 | 2115 | val ctrm = Thm.apply (Thm.apply add two) ten | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2116 | in | 
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2117 | Thm.prop_of (Thm.beta_conversion true ctrm) | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2118 | end" | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2119 | "Const (\"==\",\<dots>) $ | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2120 | (Abs (\"x\",\<dots>,Abs (\"y\",\<dots>,\<dots>)) $\<dots>$\<dots>) $ | 
| 418 | 2121 | (Const (\"Groups.plus_class.plus\",\<dots>) $ \<dots> $ \<dots>)"} | 
| 142 | 2122 | |
| 512 | 2123 | or in the pretty-printed form | 
| 509 
dcefee89bf62
updated part about beta in Tactiacal section
 Christian Urban <urbanc@in.tum.de> parents: 
506diff
changeset | 2124 | |
| 
dcefee89bf62
updated part about beta in Tactiacal section
 Christian Urban <urbanc@in.tum.de> parents: 
506diff
changeset | 2125 |   @{ML_response_fake [display,gray]
 | 
| 
dcefee89bf62
updated part about beta in Tactiacal section
 Christian Urban <urbanc@in.tum.de> parents: 
506diff
changeset | 2126 | "let | 
| 512 | 2127 |    val add = @{cterm \"\<lambda>x y. x + (y::nat)\"}
 | 
| 2128 |    val two = @{cterm \"2::nat\"}
 | |
| 2129 |    val ten = @{cterm \"10::nat\"}
 | |
| 513 | 2130 | val ctrm = Thm.apply (Thm.apply add two) ten | 
| 512 | 2131 |    val ctxt = @{context}
 | 
| 2132 | |> Config.put eta_contract false | |
| 2133 | |> Config.put show_abbrevs false | |
| 509 
dcefee89bf62
updated part about beta in Tactiacal section
 Christian Urban <urbanc@in.tum.de> parents: 
506diff
changeset | 2134 | in | 
| 
dcefee89bf62
updated part about beta in Tactiacal section
 Christian Urban <urbanc@in.tum.de> parents: 
506diff
changeset | 2135 | Thm.prop_of (Thm.beta_conversion true ctrm) | 
| 
dcefee89bf62
updated part about beta in Tactiacal section
 Christian Urban <urbanc@in.tum.de> parents: 
506diff
changeset | 2136 | |> pretty_term ctxt | 
| 
dcefee89bf62
updated part about beta in Tactiacal section
 Christian Urban <urbanc@in.tum.de> parents: 
506diff
changeset | 2137 | |> pwriteln | 
| 
dcefee89bf62
updated part about beta in Tactiacal section
 Christian Urban <urbanc@in.tum.de> parents: 
506diff
changeset | 2138 | end" | 
| 
dcefee89bf62
updated part about beta in Tactiacal section
 Christian Urban <urbanc@in.tum.de> parents: 
506diff
changeset | 2139 | "(\<lambda>x y. x + y) 2 10 \<equiv> 2 + 10"} | 
| 
dcefee89bf62
updated part about beta in Tactiacal section
 Christian Urban <urbanc@in.tum.de> parents: 
506diff
changeset | 2140 | |
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2141 |   The argument @{ML true} in @{ML beta_conversion in Thm} indicates that 
 | 
| 243 | 2142 | the right-hand side should be fully beta-normalised. If instead | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2143 |   @{ML false} is given, then only a single beta-reduction is performed 
 | 
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2144 | on the outer-most level. | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2145 | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2146 | The main point of conversions is that they can be used for rewriting | 
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2147 |   @{ML_type cterm}s. One example is the function 
 | 
| 316 
74f0a06f751f
further polishing of index generation
 Christian Urban <urbanc@in.tum.de> parents: 
315diff
changeset | 2148 |   @{ML_ind  rewr_conv in Conv}, which expects a meta-equation as an 
 | 
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2149 | argument. Suppose the following meta-equation. | 
| 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2150 | |
| 135 | 2151 | *} | 
| 2152 | ||
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 2153 | lemma true_conj1: | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 2154 | shows "True \<and> P \<equiv> P" by simp | 
| 135 | 2155 | |
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2156 | text {* 
 | 
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2157 |   It can be used for example to rewrite @{term "True \<and> (Foo \<longrightarrow> Bar)"} 
 | 
| 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2158 |   to @{term "Foo \<longrightarrow> Bar"}. The code is as follows.
 | 
| 139 | 2159 | |
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 2160 |   @{ML_response_fake [display,gray]
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2161 | "let | 
| 149 | 2162 |   val ctrm = @{cterm \"True \<and> (Foo \<longrightarrow> Bar)\"}
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2163 | in | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2164 |   Conv.rewr_conv @{thm true_conj1} ctrm
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2165 | end" | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2166 | "True \<and> (Foo \<longrightarrow> Bar) \<equiv> Foo \<longrightarrow> Bar"} | 
| 139 | 2167 | |
| 316 
74f0a06f751f
further polishing of index generation
 Christian Urban <urbanc@in.tum.de> parents: 
315diff
changeset | 2168 |   Note, however, that the function @{ML_ind  rewr_conv in Conv} only rewrites the 
 | 
| 160 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 2169 |   outer-most level of the @{ML_type cterm}. If the given @{ML_type cterm} does not match 
 | 
| 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 2170 | exactly the | 
| 316 
74f0a06f751f
further polishing of index generation
 Christian Urban <urbanc@in.tum.de> parents: 
315diff
changeset | 2171 |   left-hand side of the theorem, then  @{ML_ind  rewr_conv in Conv} fails, raising 
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2172 |   the exception @{ML CTERM}.
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2173 | |
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2174 | 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 | 2175 | combining several conversions into one. For this you can use conversion | 
| 369 | 2176 |   combinators. The simplest conversion combinator is @{ML_ind then_conv in Conv}, 
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2177 | which applies one conversion after another. For example | 
| 139 | 2178 | |
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 2179 |   @{ML_response_fake [display,gray]
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2180 | "let | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2181 | val conv1 = Thm.beta_conversion false | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2182 |   val conv2 = Conv.rewr_conv @{thm true_conj1}
 | 
| 513 | 2183 |   val ctrm = Thm.apply @{cterm \"\<lambda>x. x \<and> False\"} @{cterm \"True\"}
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2184 | in | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2185 | (conv1 then_conv conv2) ctrm | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2186 | end" | 
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 2187 | "(\<lambda>x. x \<and> False) True \<equiv> False"} | 
| 139 | 2188 | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2189 | where we first beta-reduce the term and then rewrite according to | 
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2190 |   @{thm [source] true_conj1}. (When running this example recall the 
 | 
| 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2191 | problem with the pretty-printer normalising all terms.) | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2192 | |
| 369 | 2193 |   The conversion combinator @{ML_ind else_conv in Conv} tries out the 
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2194 | 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 | 2195 | |
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 2196 |   @{ML_response_fake [display,gray]
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2197 | "let | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2198 |   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 | 2199 |   val ctrm1 = @{cterm \"True \<and> Q\"}
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2200 |   val ctrm2 = @{cterm \"P \<or> (True \<and> Q)\"}
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2201 | in | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2202 | (conv ctrm1, conv ctrm2) | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2203 | end" | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2204 | "(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 | 2205 | |
| 406 
f399b6855546
implemented suggestion from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
405diff
changeset | 2206 |   Here the conversion @{thm [source] true_conj1} only applies
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2207 | in the first case, but fails in the second. The whole conversion | 
| 256 
1fb8d62c88a0
added some first index-information
 Christian Urban <urbanc@in.tum.de> parents: 
255diff
changeset | 2208 |   does not fail, however, because the combinator @{ML else_conv in Conv} will then 
 | 
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2209 |   try out @{ML all_conv in Conv}, which always succeeds. The same
 | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2210 |   behaviour can also be achieved with conversion combinator @{ML_ind try_conv in Conv}.
 | 
| 174 | 2211 | For example | 
| 2212 | ||
| 2213 |   @{ML_response_fake [display,gray]
 | |
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2214 | "let | 
| 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2215 |   val conv = Conv.try_conv (Conv.rewr_conv @{thm true_conj1})
 | 
| 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2216 |   val ctrm = @{cterm \"True \<or> P\"}
 | 
| 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2217 | in | 
| 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2218 | conv ctrm | 
| 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2219 | end" | 
| 174 | 2220 | "True \<or> P \<equiv> True \<or> P"} | 
| 2221 | ||
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2222 | Rewriting with more than one theorem can be done using the function | 
| 424 | 2223 |   @{ML_ind rewrs_conv in Conv} from the structure @{ML_struct Conv}.
 | 
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2224 | |
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2225 | |
| 149 | 2226 |   Apart from the function @{ML beta_conversion in Thm}, which is able to fully
 | 
| 2227 | 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 | 2228 |   only apply to the outer-most level of a @{ML_type cterm}. In what follows we
 | 
| 369 | 2229 |   will lift this restriction. The combinators @{ML_ind fun_conv in Conv} 
 | 
| 2230 |   and @{ML_ind arg_conv in Conv} will apply
 | |
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2231 | a conversion to the first, respectively second, argument of an application. | 
| 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2232 | For example | 
| 139 | 2233 | |
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 2234 |   @{ML_response_fake [display,gray]
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2235 | "let | 
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2236 |   val conv = Conv.arg_conv (Conv.rewr_conv @{thm true_conj1})
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2237 |   val ctrm = @{cterm \"P \<or> (True \<and> Q)\"}
 | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2238 | in | 
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2239 | conv ctrm | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2240 | end" | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2241 | "P \<or> (True \<and> Q) \<equiv> P \<or> Q"} | 
| 139 | 2242 | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2243 |   The reason for this behaviour is that @{text "(op \<or>)"} expects two
 | 
| 160 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 2244 |   arguments.  Therefore the term must be of the form @{text "(Const \<dots> $ t1) $ t2"}. The
 | 
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2245 |   conversion is then applied to @{text "t2"}, which in the example above
 | 
| 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2246 |   stands for @{term "True \<and> Q"}. The function @{ML fun_conv in Conv} would apply
 | 
| 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2247 |   the conversion to the term @{text "(Const \<dots> $ t1)"}.
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2248 | |
| 316 
74f0a06f751f
further polishing of index generation
 Christian Urban <urbanc@in.tum.de> parents: 
315diff
changeset | 2249 |   The function @{ML_ind  abs_conv in Conv} applies a conversion under an 
 | 
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2250 | abstraction. For example: | 
| 139 | 2251 | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2252 |   @{ML_response_fake [display,gray]
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2253 | "let | 
| 243 | 2254 |   val conv = Conv.rewr_conv @{thm true_conj1} 
 | 
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2255 |   val ctrm = @{cterm \"\<lambda>P. True \<and> (P \<and> Foo)\"}
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2256 | in | 
| 243 | 2257 |   Conv.abs_conv (K conv) @{context} ctrm
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2258 | end" | 
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2259 | "\<lambda>P. True \<and> (P \<and> Foo) \<equiv> \<lambda>P. P \<and> Foo"} | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2260 | |
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2261 | Note that this conversion needs a context as an argument. We also give the | 
| 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2262 |   conversion as @{text "(K conv)"}, which is a function that ignores its
 | 
| 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2263 | argument (the argument being a sufficiently freshened version of the | 
| 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2264 | variable that is abstracted and a context). The conversion that goes under | 
| 316 
74f0a06f751f
further polishing of index generation
 Christian Urban <urbanc@in.tum.de> parents: 
315diff
changeset | 2265 |   an application is @{ML_ind  combination_conv in Conv}. It expects two
 | 
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2266 | conversions as arguments, each of which is applied to the corresponding | 
| 292 
41a802bbb7df
added more to the ML-antiquotation section
 Christian Urban <urbanc@in.tum.de> parents: 
291diff
changeset | 2267 | ``branch'' of the application. An abbreviation for this conversion is the | 
| 316 
74f0a06f751f
further polishing of index generation
 Christian Urban <urbanc@in.tum.de> parents: 
315diff
changeset | 2268 |   function @{ML_ind  comb_conv in Conv}, which applies the same conversion
 | 
| 292 
41a802bbb7df
added more to the ML-antiquotation section
 Christian Urban <urbanc@in.tum.de> parents: 
291diff
changeset | 2269 | to both branches. | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2270 | |
| 160 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 2271 | We can now apply all these functions in a conversion that recursively | 
| 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 2272 |   descends a term and applies a ``@{thm [source] true_conj1}''-conversion 
 | 
| 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 2273 | in all possible positions. | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2274 | *} | 
| 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2275 | |
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2276 | ML %linenosgray{*fun true_conj1_conv ctxt ctrm =
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2277 | case (Thm.term_of ctrm) of | 
| 142 | 2278 |     @{term "op \<and>"} $ @{term True} $ _ => 
 | 
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2279 | (Conv.arg_conv (true_conj1_conv ctxt) then_conv | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2280 |          Conv.rewr_conv @{thm true_conj1}) ctrm
 | 
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2281 | | _ $ _ => Conv.comb_conv (true_conj1_conv ctxt) ctrm | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2282 | | Abs _ => Conv.abs_conv (fn (_, ctxt) => true_conj1_conv ctxt) ctxt ctrm | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2283 | | _ => Conv.all_conv ctrm*} | 
| 139 | 2284 | |
| 2285 | text {*
 | |
| 406 
f399b6855546
implemented suggestion from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
405diff
changeset | 2286 |   This function ``fires'' if the term is of the form @{text "(True \<and> \<dots>)"}. 
 | 
| 
f399b6855546
implemented suggestion from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
405diff
changeset | 2287 | It descends under applications (Line 6) and abstractions | 
| 
f399b6855546
implemented suggestion from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
405diff
changeset | 2288 | (Line 7); otherwise it leaves the term unchanged (Line 8). In Line 2 | 
| 149 | 2289 |   we need to transform the @{ML_type cterm} into a @{ML_type term} in order
 | 
| 2290 | to be able to pattern-match the term. To see this | |
| 160 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 2291 | conversion in action, consider the following example: | 
| 139 | 2292 | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2293 | @{ML_response_fake [display,gray]
 | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2294 | "let | 
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2295 |   val conv = true_conj1_conv @{context}
 | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2296 |   val ctrm = @{cterm \"distinct [1::nat, x] \<longrightarrow> True \<and> 1 \<noteq> x\"}
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2297 | in | 
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2298 | conv ctrm | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2299 | end" | 
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 2300 | "distinct [1, x] \<longrightarrow> True \<and> 1 \<noteq> x \<equiv> distinct [1, x] \<longrightarrow> 1 \<noteq> x"} | 
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2301 | *} | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2302 | |
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2303 | text {*
 | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2304 |   Conversions that traverse terms, like @{ML true_conj1_conv} above, can be
 | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2305 |   implemented more succinctly with the combinators @{ML_ind bottom_conv in
 | 
| 424 | 2306 |   Conv} and @{ML_ind top_conv in Conv}. For example:
 | 
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2307 | *} | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2308 | |
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 2309 | ML %grayML{*fun true_conj1_conv ctxt =
 | 
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2310 | let | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2311 |   val conv = Conv.try_conv (Conv.rewr_conv @{thm true_conj1})
 | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2312 | in | 
| 424 | 2313 | Conv.bottom_conv (K conv) ctxt | 
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2314 | end*} | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2315 | |
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2316 | text {*
 | 
| 408 | 2317 | If we regard terms as trees with variables and constants on the top, then | 
| 424 | 2318 |   @{ML bottom_conv in Conv} traverses first the the term and
 | 
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2319 |   on the ``way down'' applies the conversion, whereas @{ML top_conv in
 | 
| 424 | 2320 | Conv} applies the conversion on the ``way up''. To see the difference, | 
| 412 | 2321 | assume the following two meta-equations | 
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2322 | *} | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2323 | |
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2324 | lemma conj_assoc: | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2325 | fixes A B C::bool | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2326 | shows "A \<and> (B \<and> C) \<equiv> (A \<and> B) \<and> C" | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2327 | and "(A \<and> B) \<and> C \<equiv> A \<and> (B \<and> C)" | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2328 | by simp_all | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2329 | |
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2330 | text {*
 | 
| 412 | 2331 |   and the @{ML_type cterm} @{text "(a \<and> (b \<and> c)) \<and> d"}. Here you should
 | 
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2332 | pause for a moment to be convinced that rewriting top-down and bottom-up | 
| 412 | 2333 | according to the two meta-equations produces two results. Below these | 
| 2334 | two results are calculated. | |
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2335 | |
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2336 |   @{ML_response_fake [display, gray]
 | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2337 | "let | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2338 |   val ctxt = @{context}
 | 
| 424 | 2339 |   val conv = Conv.try_conv (Conv.rewrs_conv @{thms conj_assoc})
 | 
| 2340 | val conv_top = Conv.top_conv (K conv) ctxt | |
| 2341 | val conv_bot = Conv.bottom_conv (K conv) ctxt | |
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2342 |   val ctrm = @{cterm \"(a \<and> (b \<and> c)) \<and> d\"}
 | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2343 | in | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2344 | (conv_top ctrm, conv_bot ctrm) | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2345 | end" | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2346 | "(\"(a \<and> (b \<and> c)) \<and> d \<equiv> a \<and> (b \<and> (c \<and> d))\", | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2347 | \"(a \<and> (b \<and> c)) \<and> d \<equiv> (a \<and> b) \<and> (c \<and> d)\")"} | 
| 139 | 2348 | |
| 412 | 2349 | To see how much control you have over rewriting with conversions, let us | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2350 | make the task a bit more complicated by rewriting according to the rule | 
| 149 | 2351 |   @{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 | 2352 | the conversion should be as follows. | 
| 135 | 2353 | *} | 
| 2354 | ||
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 2355 | ML %grayML{*fun if_true1_simple_conv ctxt ctrm =
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2356 | case Thm.term_of ctrm of | 
| 142 | 2357 |     Const (@{const_name If}, _) $ _ =>
 | 
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2358 | Conv.arg_conv (true_conj1_conv ctxt) ctrm | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2359 | | _ => Conv.no_conv ctrm | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2360 | |
| 424 | 2361 | val if_true1_conv = Conv.top_sweep_conv if_true1_simple_conv*} | 
| 135 | 2362 | |
| 139 | 2363 | text {*
 | 
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2364 | In the first function we only treat the top-most redex and also only the | 
| 412 | 2365 |   success case. As default we return @{ML no_conv in Conv}.  To apply this
 | 
| 2366 |   ``simple'' conversion inside a term, we use in the last line the combinator @{ML_ind
 | |
| 424 | 2367 | top_sweep_conv in Conv}, which traverses the term starting from the | 
| 406 
f399b6855546
implemented suggestion from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
405diff
changeset | 2368 | root and applies it to all the redexes on the way, but stops in each branch as | 
| 
f399b6855546
implemented suggestion from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
405diff
changeset | 2369 | soon as it found one redex. Here is an example for this conversion: | 
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2370 | |
| 139 | 2371 | |
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 2372 |   @{ML_response_fake [display,gray]
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2373 | "let | 
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2374 |   val ctrm = @{cterm \"if P (True \<and> 1 \<noteq> (2::nat)) 
 | 
| 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2375 | then True \<and> True else True \<and> False\"} | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2376 | in | 
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2377 |   if_true1_conv @{context} ctrm
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2378 | end" | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2379 | "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 | 2380 | \<equiv> if P (1 \<noteq> 2) then True \<and> True else True \<and> False"} | 
| 135 | 2381 | *} | 
| 2382 | ||
| 2383 | text {*
 | |
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2384 |   So far we only applied conversions to @{ML_type cterm}s. Conversions can, however, 
 | 
| 316 
74f0a06f751f
further polishing of index generation
 Christian Urban <urbanc@in.tum.de> parents: 
315diff
changeset | 2385 |   also work on theorems using the function @{ML_ind  fconv_rule in Conv}. As an example, 
 | 
| 412 | 2386 |   consider again the conversion @{ML true_conj1_conv} and the lemma:
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2387 | *} | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2388 | |
| 362 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 2389 | lemma foo_test: | 
| 
a5e7ab090abf
added something about manual instantiations of theorems
 Christian Urban <urbanc@in.tum.de> parents: 
361diff
changeset | 2390 | shows "P \<or> (True \<and> \<not>P)" by simp | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2391 | |
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2392 | text {*
 | 
| 412 | 2393 | Using the conversion you can transform this theorem into a | 
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2394 | new theorem as follows | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2395 | |
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2396 |   @{ML_response_fake [display,gray]
 | 
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2397 | "let | 
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2398 |   val conv = Conv.fconv_rule (true_conj1_conv @{context}) 
 | 
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2399 |   val thm = @{thm foo_test}
 | 
| 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2400 | in | 
| 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2401 | conv thm | 
| 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2402 | end" | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2403 | "?P \<or> \<not> ?P"} | 
| 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2404 | |
| 412 | 2405 |   Finally, Isabelle provides function @{ML_ind CONVERSION in Tactical} 
 | 
| 2406 | for turning conversions into tactics. This needs some predefined conversion | |
| 2407 | combinators that traverse a goal | |
| 410 | 2408 | state and can selectively apply the conversion. The combinators for | 
| 2409 | the goal state are: | |
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2410 | |
| 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2411 |   \begin{itemize}
 | 
| 369 | 2412 |   \item @{ML_ind params_conv in Conv} for converting under parameters
 | 
| 412 | 2413 |   (i.e.~where a goal state is of the form @{text "\<And>x. P x \<Longrightarrow> Q x"})
 | 
| 2414 | ||
| 2415 |   \item @{ML_ind prems_conv in Conv} for applying a conversion to 
 | |
| 2416 | premises of a goal state, and | |
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2417 | |
| 369 | 2418 |   \item @{ML_ind concl_conv in Conv} for applying a conversion to the
 | 
| 412 | 2419 | conclusion of a goal state. | 
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2420 |   \end{itemize}
 | 
| 139 | 2421 | |
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2422 |   Assume we want to apply @{ML true_conj1_conv} only in the conclusion
 | 
| 160 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 2423 |   of the goal, and @{ML if_true1_conv} should only apply to the premises.
 | 
| 145 
f1ba430a5e7d
some very slight polishing on the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
142diff
changeset | 2424 | Here is a tactic doing exactly that: | 
| 135 | 2425 | *} | 
| 2426 | ||
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
515diff
changeset | 2427 | ML %grayML{*fun true1_tac ctxt =
 | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
184diff
changeset | 2428 | CONVERSION | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
184diff
changeset | 2429 | (Conv.params_conv ~1 (fn ctxt => | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
184diff
changeset | 2430 | (Conv.prems_conv ~1 (if_true1_conv ctxt) then_conv | 
| 405 
f8d020bbc2c0
improved section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
404diff
changeset | 2431 | Conv.concl_conv ~1 (true_conj1_conv ctxt))) ctxt)*} | 
| 142 | 2432 | |
| 2433 | text {* 
 | |
| 406 
f399b6855546
implemented suggestion from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
405diff
changeset | 2434 |   We call the conversions with the argument @{ML "~1"}. By this we 
 | 
| 
f399b6855546
implemented suggestion from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
405diff
changeset | 2435 | analyse all parameters, all premises and the conclusion of a goal | 
| 
f399b6855546
implemented suggestion from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
405diff
changeset | 2436 |   state. If we call @{ML concl_conv in Conv} with 
 | 
| 
f399b6855546
implemented suggestion from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
405diff
changeset | 2437 |   a non-negative number, say @{text n}, then this conversions will 
 | 
| 
f399b6855546
implemented suggestion from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
405diff
changeset | 2438 |   skip the first @{text n} premises and applies the conversion to the 
 | 
| 
f399b6855546
implemented suggestion from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
405diff
changeset | 2439 | ``rest'' (similar for parameters and conclusions). To test the | 
| 
f399b6855546
implemented suggestion from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
405diff
changeset | 2440 | tactic, consider the proof | 
| 142 | 2441 | *} | 
| 139 | 2442 | |
| 142 | 2443 | lemma | 
| 2444 | "if True \<and> P then P else True \<and> False \<Longrightarrow> | |
| 148 | 2445 | (if True \<and> Q then True \<and> Q else P) \<longrightarrow> True \<and> (True \<and> Q)" | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
184diff
changeset | 2446 | apply(tactic {* true1_tac @{context} 1 *})
 | 
| 147 
6dafb0815ae6
more polishing of the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
146diff
changeset | 2447 | txt {* where the tactic yields the goal state
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 2448 |   \begin{isabelle}
 | 
| 177 | 2449 |   @{subgoals [display]}
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 2450 |   \end{isabelle}*}
 | 
| 142 | 2451 | (*<*)oops(*>*) | 
| 135 | 2452 | |
| 2453 | text {*
 | |
| 148 | 2454 |   As you can see, the premises are rewritten according to @{ML if_true1_conv}, while
 | 
| 410 | 2455 |   the conclusion according to @{ML true_conj1_conv}. If we only have one
 | 
| 2456 | conversion that should be uniformly applied to the whole goal state, we | |
| 424 | 2457 |   can simplify @{ML true1_tac} using @{ML_ind top_conv in Conv}.
 | 
| 412 | 2458 | |
| 2459 | Conversions are also be helpful for constructing meta-equality theorems. | |
| 332 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2460 | Such theorems are often definitions. As an example consider the following | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2461 | two ways of defining the identity function in Isabelle. | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2462 | *} | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2463 | |
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2464 | definition id1::"'a \<Rightarrow> 'a" | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2465 | where "id1 x \<equiv> x" | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2466 | |
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2467 | definition id2::"'a \<Rightarrow> 'a" | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2468 | where "id2 \<equiv> \<lambda>x. x" | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2469 | |
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2470 | text {*
 | 
| 335 
163ac0662211
reorganised the certified terms section; tuned
 Christian Urban <urbanc@in.tum.de> parents: 
334diff
changeset | 2471 |   Although both definitions define the same function, the theorems @{thm
 | 
| 412 | 2472 |   [source] id1_def} and @{thm [source] id2_def} are different meta-equations. However it is
 | 
| 2473 |   easy to transform one into the other. The function @{ML_ind abs_def
 | |
| 2474 |   in Drule} from the structure @{ML_struct Drule} can automatically transform 
 | |
| 2475 |   theorem @{thm [source] id1_def}
 | |
| 334 | 2476 |   into @{thm [source] id2_def} by abstracting all variables on the 
 | 
| 2477 | left-hand side in the right-hand side. | |
| 332 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2478 | |
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2479 |   @{ML_response_fake [display,gray]
 | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2480 |   "Drule.abs_def @{thm id1_def}"
 | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2481 | "id1 \<equiv> \<lambda>x. x"} | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2482 | |
| 406 
f399b6855546
implemented suggestion from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
405diff
changeset | 2483 | Unfortunately, Isabelle has no built-in function that transforms the | 
| 
f399b6855546
implemented suggestion from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
405diff
changeset | 2484 | theorems in the other direction. We can implement one using | 
| 334 | 2485 | conversions. The feature of conversions we are using is that if we apply a | 
| 2486 |   @{ML_type cterm} to a conversion we obtain a meta-equality theorem (recall
 | |
| 2487 | that the type of conversions is an abbreviation for | |
| 2488 |   @{ML_type "cterm -> thm"}). The code of the transformation is below.
 | |
| 332 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2489 | *} | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2490 | |
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2491 | ML %linenosgray{*fun unabs_def ctxt def =
 | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2492 | let | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2493 | val (lhs, rhs) = Thm.dest_equals (cprop_of def) | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2494 | val xs = strip_abs_vars (term_of rhs) | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2495 | val (_, ctxt') = Variable.add_fixes (map fst xs) ctxt | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2496 | |
| 475 
25371f74c768
updated to post-2011-1 Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
466diff
changeset | 2497 | val thy = Proof_Context.theory_of ctxt' | 
| 332 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2498 | val cxs = map (cterm_of thy o Free) xs | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2499 | val new_lhs = Drule.list_comb (lhs, cxs) | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2500 | |
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2501 | fun get_conv [] = Conv.rewr_conv def | 
| 334 | 2502 | | get_conv (_::xs) = Conv.fun_conv (get_conv xs) | 
| 332 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2503 | in | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2504 | get_conv xs new_lhs |> | 
| 475 
25371f74c768
updated to post-2011-1 Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
466diff
changeset | 2505 | singleton (Proof_Context.export ctxt' ctxt) | 
| 332 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2506 | end*} | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2507 | |
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2508 | text {*
 | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2509 |   In Line 3 we destruct the meta-equality into the @{ML_type cterm}s
 | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2510 | corresponding to the left-hand and right-hand side of the meta-equality. The | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2511 |   assumption in @{ML unabs_def} is that the right-hand side is an
 | 
| 334 | 2512 | abstraction. We compute the possibly empty list of abstracted variables in | 
| 369 | 2513 |   Line 4 using the function @{ML_ind strip_abs_vars in Term}. For this we have to
 | 
| 412 | 2514 |   first transform the @{ML_type cterm} into a @{ML_type term}. In Line 5 we
 | 
| 406 
f399b6855546
implemented suggestion from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
405diff
changeset | 2515 |   import these variables into the context @{text "ctxt'"}, in order to
 | 
| 334 | 2516 | export them again in Line 15. In Line 8 we certify the list of variables, | 
| 2517 |   which in turn we apply to the @{ML_text lhs} of the definition using the
 | |
| 2518 |   function @{ML_ind list_comb in Drule}. In Line 11 and 12 we generate the
 | |
| 2519 | conversion according to the length of the list of (abstracted) variables. If | |
| 2520 | there are none, then we just return the conversion corresponding to the | |
| 2521 | original definition. If there are variables, then we have to prefix this | |
| 2522 |   conversion with @{ML_ind fun_conv in Conv}. Now in Line 14 we only have to
 | |
| 2523 | apply the new left-hand side to the generated conversion and obtain the the | |
| 2524 | theorem we want to construct. As mentioned above, in Line 15 we only have to | |
| 406 
f399b6855546
implemented suggestion from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
405diff
changeset | 2525 |   export the context @{text "ctxt'"} back to @{text "ctxt"} in order to 
 | 
| 
f399b6855546
implemented suggestion from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
405diff
changeset | 2526 | produce meta-variables for the theorem. An example is as follows. | 
| 332 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2527 | |
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2528 |   @{ML_response_fake [display, gray]
 | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2529 |   "unabs_def @{context} @{thm id2_def}"
 | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2530 | "id2 ?x1 \<equiv> ?x1"} | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2531 | *} | 
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2532 | |
| 
6fba3a3e80a3
added the function unabs_def in the conversions section
 Christian Urban <urbanc@in.tum.de> parents: 
331diff
changeset | 2533 | text {*
 | 
| 243 | 2534 | To sum up this section, conversions are more general than the simplifier | 
| 2535 | or simprocs, but you have to do more work yourself. Also conversions are | |
| 2536 | often much less efficient than the simplifier. The advantage of conversions, | |
| 406 
f399b6855546
implemented suggestion from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
405diff
changeset | 2537 | however, is that they provide much less room for non-termination. | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2538 | |
| 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 | 2539 |   \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 | 2540 | Write a tactic that does the same as the simproc in exercise | 
| 291 
077c764c8d8b
polished the section on conversions
 Christian Urban <urbanc@in.tum.de> parents: 
289diff
changeset | 2541 |   \ref{ex:addsimproc}, but is based on conversions. You can make
 | 
| 166 
00d153e32a53
improvments to the solutions suggested by Sacha B?hme
 Christian Urban <urbanc@in.tum.de> parents: 
163diff
changeset | 2542 |   the same assumptions as in \ref{ex:addsimproc}. 
 | 
| 152 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 2543 |   \end{exercise}
 | 
| 
8084c353d196
added material to the endless story of the simplifier
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 2544 | |
| 172 
ec47352e99c2
improved the solution for the simproc/conversion exercise
 Christian Urban <urbanc@in.tum.de> parents: 
170diff
changeset | 2545 |   \begin{exercise}\label{ex:compare}
 | 
| 174 | 2546 |   Compare your solutions of Exercises~\ref{ex:addsimproc} and \ref{ex:addconversion},
 | 
| 172 
ec47352e99c2
improved the solution for the simproc/conversion exercise
 Christian Urban <urbanc@in.tum.de> parents: 
170diff
changeset | 2547 | and try to determine which way of rewriting such terms is faster. For this you might | 
| 
ec47352e99c2
improved the solution for the simproc/conversion exercise
 Christian Urban <urbanc@in.tum.de> parents: 
170diff
changeset | 2548 |   have to construct quite large terms. Also see Recipe \ref{rec:timing} for information 
 | 
| 
ec47352e99c2
improved the solution for the simproc/conversion exercise
 Christian Urban <urbanc@in.tum.de> parents: 
170diff
changeset | 2549 | about timing. | 
| 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 | 2550 |   \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 | 2551 | |
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2552 |   \begin{readmore}
 | 
| 424 | 2553 |   See @{ML_file "Pure/conv.ML"}
 | 
| 384 | 2554 | for more information about conversion combinators. | 
| 243 | 2555 |   Some basic conversions are defined in  @{ML_file "Pure/thm.ML"},
 | 
| 458 
242e81f4d461
updated to post-2011 Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
456diff
changeset | 2556 |   @{ML_file "Pure/drule.ML"} and @{ML_file "Pure/raw_simplifier.ML"}.
 | 
| 146 
4aa8a80e37ff
some polishing about conversions
 Christian Urban <urbanc@in.tum.de> parents: 
145diff
changeset | 2557 |   \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 | 2558 | |
| 135 | 2559 | *} | 
| 2560 | ||
| 184 | 2561 | text {*
 | 
| 2562 |   (FIXME: check whether @{ML Pattern.match_rew} and @{ML Pattern.rewrite_term}
 | |
| 2563 | are of any use/efficient) | |
| 2564 | *} | |
| 135 | 2565 | |
| 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 | 2566 | |
| 216 
fcedd5bd6a35
added a declaration section (for Amine)
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 2567 | section {* Declarations (TBD) *}
 | 
| 
fcedd5bd6a35
added a declaration section (for Amine)
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 2568 | |
| 500 | 2569 | section {* Structured Proofs\label{sec:structured} (TBD) *}
 | 
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2570 | |
| 129 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 2571 | text {* TBD *}
 | 
| 
e0d368a45537
started a section about simprocs
 Christian Urban <urbanc@in.tum.de> parents: 
128diff
changeset | 2572 | |
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2573 | lemma True | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2574 | proof | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2575 | |
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2576 |   {
 | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2577 | fix A B C | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2578 | assume r: "A & B \<Longrightarrow> C" | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2579 | assume A B | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2580 | then have "A & B" .. | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2581 | then have C by (rule r) | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2582 | } | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2583 | |
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2584 |   {
 | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2585 | fix A B C | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2586 | assume r: "A & B \<Longrightarrow> C" | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2587 | assume A B | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2588 | note conjI [OF this] | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2589 | note r [OF this] | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2590 | } | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2591 | oops | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2592 | |
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2593 | ML {* 
 | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2594 |   val ctxt0 = @{context};
 | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2595 | val ctxt = ctxt0; | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2596 | val (_, ctxt) = Variable.add_fixes ["A", "B", "C"] ctxt; | 
| 217 | 2597 |   val ([r], ctxt) = Assumption.add_assumes [@{cprop "A & B \<Longrightarrow> C"}] ctxt
 | 
| 2598 |   val (this, ctxt) = Assumption.add_assumes [@{cprop "A"}, @{cprop "B"}] ctxt;
 | |
| 95 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2599 |   val this = [@{thm conjI} OF this]; 
 | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2600 | val this = r OF this; | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2601 | val this = Assumption.export false ctxt ctxt0 this | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2602 | val this = Variable.export ctxt ctxt0 [this] | 
| 
7235374f34c8
added some preliminary notes about SUBPROOF
 Christian Urban <urbanc@in.tum.de> parents: 
93diff
changeset | 2603 | *} | 
| 93 | 2604 | |
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 2605 | section {* Summary *}
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 2606 | |
| 363 | 2607 | text {*
 | 
| 368 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 2608 | |
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 2609 |   \begin{conventions}
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 2610 |   \begin{itemize}
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 2611 |   \item Reference theorems with the antiquotation @{text "@{thm \<dots>}"}.
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 2612 | \item If a tactic is supposed to fail, return the empty sequence. | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 2613 | \item If you apply a tactic to a sequence of subgoals, apply it | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 2614 | in reverse order (i.e.~start with the last subgoal). | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 2615 | \item Use simpsets that are as small as possible. | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 2616 |   \end{itemize}
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 2617 |   \end{conventions}
 | 
| 
b1a458a03a8e
new parts in the tactical section
 Christian Urban <urbanc@in.tum.de> parents: 
363diff
changeset | 2618 | |
| 363 | 2619 | *} | 
| 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 | 2620 | |
| 139 | 2621 | end |