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