| author | Christian Urban <urbanc@in.tum.de> | 
| Sun, 08 Feb 2009 08:45:25 +0000 | |
| changeset 104 | 5dcad9348e4d | 
| parent 102 | 5e309df58557 | 
| child 105 | f49dc7e96235 | 
| 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 {*
 | 
|
8  | 
||
| 
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
 | 
9  | 
The main reason for descending to the ML-level of Isabelle is to be  | 
| 
 
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
 | 
10  | 
able to implement automatic proof procedures. Such proof procedures usually  | 
| 
 
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
 | 
11  | 
lessen considerably the burden of manual reasoning, for example, when  | 
| 
 
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
 | 
12  | 
introducing new definitions. These proof procedures are centred around  | 
| 
 
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
 | 
13  | 
  refining a goal state using tactics. This is similar to the @{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
 | 
14  | 
apply}-style reasoning at the user level, where goals are modified in a  | 
| 
 
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
 | 
15  | 
sequence of proof steps until all of them are solved. However, there are  | 
| 
 
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
 | 
16  | 
also more structured operations available on the ML-level that help with  | 
| 
 
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  | 
the handling of variables and assumptions.  | 
| 
 
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
 | 
18  | 
|
| 93 | 19  | 
*}  | 
20  | 
||
| 
102
 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 
Christian Urban <urbanc@in.tum.de> 
parents: 
99 
diff
changeset
 | 
21  | 
section {* Basics of Reasoning with Tactics*}
 | 
| 93 | 22  | 
|
23  | 
text {*
 | 
|
| 
102
 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 
Christian Urban <urbanc@in.tum.de> 
parents: 
99 
diff
changeset
 | 
24  | 
  To see how tactics work, let us first transcribe a simple \isacommand{apply}-style proof 
 | 
| 93 | 25  | 
into ML. Consider the following proof.  | 
26  | 
*}  | 
|
27  | 
||
28  | 
lemma disj_swap: "P \<or> Q \<Longrightarrow> Q \<or> P"  | 
|
29  | 
apply(erule disjE)  | 
|
30  | 
apply(rule disjI2)  | 
|
31  | 
apply(assumption)  | 
|
32  | 
apply(rule disjI1)  | 
|
33  | 
apply(assumption)  | 
|
34  | 
done  | 
|
35  | 
||
36  | 
text {*
 | 
|
37  | 
This proof translates to the following ML-code.  | 
|
38  | 
||
39  | 
@{ML_response_fake [display,gray]
 | 
|
40  | 
"let  | 
|
41  | 
  val ctxt = @{context}
 | 
|
42  | 
  val goal = @{prop \"P \<or> Q \<Longrightarrow> Q \<or> P\"}
 | 
|
43  | 
in  | 
|
44  | 
Goal.prove ctxt [\"P\", \"Q\"] [] goal  | 
|
45  | 
(fn _ =>  | 
|
46  | 
      etac @{thm disjE} 1
 | 
|
47  | 
      THEN rtac @{thm disjI2} 1
 | 
|
48  | 
THEN atac 1  | 
|
49  | 
      THEN rtac @{thm disjI1} 1
 | 
|
50  | 
THEN atac 1)  | 
|
51  | 
end" "?P \<or> ?Q \<Longrightarrow> ?Q \<or> ?P"}  | 
|
52  | 
||
53  | 
  To start the proof, the function @{ML "Goal.prove"}~@{text "ctxt xs As C
 | 
|
| 99 | 54  | 
  tac"} sets up a goal state for proving the goal @{text C} 
 | 
55  | 
  (that is @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"} in the proof at hand) under the
 | 
|
56  | 
  assumptions @{text As} (happens to be empty) with the variables
 | 
|
| 93 | 57  | 
  @{text xs} that will be generalised once the goal is proved (in our case
 | 
58  | 
  @{text P} and @{text Q}). The @{text "tac"} is the tactic that proves the goal;
 | 
|
59  | 
it can make use of the local assumptions (there are none in this example).  | 
|
60  | 
  The functions @{ML etac}, @{ML rtac} and @{ML atac} correspond to 
 | 
|
61  | 
  @{text erule}, @{text rule} and @{text assumption}, respectively. 
 | 
|
| 
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
 | 
62  | 
  The operator @{ML THEN} strings the tactics together. A difference
 | 
| 
 
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
 | 
63  | 
  between the \isacommand{apply}-script and the ML-code is that the
 | 
| 
 
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
 | 
64  | 
  former causes the lemma to be stored under the name @{text "disj_swap"}, 
 | 
| 
 
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
 | 
65  | 
whereas the latter does not include any code for this.  | 
| 93 | 66  | 
|
67  | 
  \begin{readmore}
 | 
|
| 99 | 68  | 
  To learn more about the function @{ML Goal.prove} see \isccite{sec:results}
 | 
69  | 
  and the file @{ML_file "Pure/goal.ML"}. For more information about the
 | 
|
70  | 
  internals of goals see \isccite{sec:tactical-goals}.  See @{ML_file
 | 
|
71  | 
  "Pure/tactic.ML"} and @{ML_file "Pure/tctical.ML"} for the code of basic
 | 
|
72  | 
tactics and tactic combinators; see also Chapters 3 and 4 in the old  | 
|
73  | 
Isabelle Reference Manual.  | 
|
| 93 | 74  | 
  \end{readmore}
 | 
75  | 
||
76  | 
Note that we used antiquotations for referencing the theorems. We could also  | 
|
77  | 
  just have written @{ML "etac disjE 1"} and so on, but this is considered bad
 | 
|
78  | 
  style. The reason is that the binding for @{ML disjE} can be re-assigned by 
 | 
|
79  | 
the user and thus one does not have complete control over which theorem is  | 
|
80  | 
actually applied. This problem is nicely prevented by using antiquotations,  | 
|
81  | 
because then the theorems are fixed statically at compile-time.  | 
|
82  | 
||
| 
102
 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 
Christian Urban <urbanc@in.tum.de> 
parents: 
99 
diff
changeset
 | 
83  | 
During the development of automatic proof procedures, you will often find it  | 
| 
 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 
Christian Urban <urbanc@in.tum.de> 
parents: 
99 
diff
changeset
 | 
84  | 
necessary to test a tactic on examples. This can be conveniently  | 
| 93 | 85  | 
  done with the command \isacommand{apply}@{text "(tactic \<verbopen> \<dots> \<verbclose>)"}. 
 | 
86  | 
Consider the following sequence of tactics  | 
|
87  | 
*}  | 
|
88  | 
||
89  | 
ML{*val foo_tac = 
 | 
|
90  | 
      (etac @{thm disjE} 1
 | 
|
91  | 
       THEN rtac @{thm disjI2} 1
 | 
|
92  | 
THEN atac 1  | 
|
93  | 
       THEN rtac @{thm disjI1} 1
 | 
|
94  | 
THEN atac 1)*}  | 
|
95  | 
||
96  | 
text {* and the Isabelle proof: *}
 | 
|
97  | 
||
98  | 
lemma "P \<or> Q \<Longrightarrow> Q \<or> P"  | 
|
99  | 
apply(tactic {* foo_tac *})
 | 
|
100  | 
done  | 
|
101  | 
||
102  | 
text {*
 | 
|
| 104 | 103  | 
  By using @{text "tactic \<verbopen> \<dots> \<verbclose>"} you can call from the 
 | 
104  | 
  user level of Isabelle the tactic @{ML foo_tac} or 
 | 
|
| 
102
 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 
Christian Urban <urbanc@in.tum.de> 
parents: 
99 
diff
changeset
 | 
105  | 
any other function that returns a tactic.  | 
| 93 | 106  | 
|
| 
102
 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 
Christian Urban <urbanc@in.tum.de> 
parents: 
99 
diff
changeset
 | 
107  | 
  The tactic @{ML foo_tac} is just a sequence of simple tactics stringed
 | 
| 
 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 
Christian Urban <urbanc@in.tum.de> 
parents: 
99 
diff
changeset
 | 
108  | 
  together by @{ML THEN}. As can be seen, each simple tactic in @{ML foo_tac}
 | 
| 
 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 
Christian Urban <urbanc@in.tum.de> 
parents: 
99 
diff
changeset
 | 
109  | 
has a hard-coded number that stands for the subgoal analysed by the  | 
| 
 
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
 | 
110  | 
  tactic (@{text "1"} stands for the first, or top-most, subgoal). This is 
 | 
| 
 
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
 | 
111  | 
sometimes wanted, but usually not. To avoid the explicit numbering in  | 
| 
 
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
 | 
112  | 
the tactic, you can write  | 
| 93 | 113  | 
*}  | 
114  | 
||
115  | 
ML{*val foo_tac' = 
 | 
|
116  | 
      (etac @{thm disjE} 
 | 
|
117  | 
       THEN' rtac @{thm disjI2} 
 | 
|
118  | 
THEN' atac  | 
|
119  | 
       THEN' rtac @{thm disjI1} 
 | 
|
120  | 
THEN' atac)*}  | 
|
121  | 
||
122  | 
text {* 
 | 
|
123  | 
and then give the number for the subgoal explicitly when the tactic is  | 
|
| 
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
 | 
124  | 
  called. For every operator that combines tactics (@{ML THEN} is only one 
 | 
| 
 
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
 | 
125  | 
such operator), a primed version exists. So in the next proof you  | 
| 
 
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
 | 
126  | 
can first discharge the second subgoal, and after that the first.  | 
| 93 | 127  | 
*}  | 
128  | 
||
129  | 
lemma "P1 \<or> Q1 \<Longrightarrow> Q1 \<or> P1"  | 
|
130  | 
and "P2 \<or> Q2 \<Longrightarrow> Q2 \<or> P2"  | 
|
131  | 
apply(tactic {* foo_tac' 2 *})
 | 
|
132  | 
apply(tactic {* foo_tac' 1 *})
 | 
|
133  | 
done  | 
|
134  | 
||
135  | 
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
 | 
136  | 
This kind of addressing is more difficult to achieve when the goal is  | 
| 
 
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
 | 
137  | 
hard-coded inside the tactic.  | 
| 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
 | 
| 99 | 141  | 
  of this form, then @{ML foo_tac} throws the error message:
 | 
142  | 
||
143  | 
  \begin{isabelle}
 | 
|
144  | 
  @{text "*** empty result sequence -- proof command failed"}\\
 | 
|
145  | 
  @{text "*** At command \"apply\"."}
 | 
|
146  | 
  \end{isabelle}
 | 
|
147  | 
||
148  | 
Meaning the tactic failed. The reason for this error message is that tactics  | 
|
149  | 
are functions that map a goal state to a (lazy) sequence of successor states,  | 
|
| 
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
 | 
150  | 
hence the type of a tactic is:  | 
| 93 | 151  | 
|
152  | 
  @{text [display, gray] "type tactic = thm -> thm Seq.seq"}
 | 
|
153  | 
||
| 99 | 154  | 
It is custom that if a tactic fails, it should return the empty sequence:  | 
| 
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
 | 
155  | 
your own tactics should not raise exceptions willy-nilly.  | 
| 
 
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
 | 
156  | 
|
| 
 
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
 | 
157  | 
  The simplest tactics are @{ML no_tac} and @{ML all_tac}. The first returns
 | 
| 
 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 
Christian Urban <urbanc@in.tum.de> 
parents: 
99 
diff
changeset
 | 
158  | 
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
 | 
159  | 
*}  | 
| 
 
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
 | 
160  | 
|
| 
 
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
 | 
161  | 
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
 | 
162  | 
|
| 
 
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
 | 
163  | 
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
 | 
164  | 
  which means @{ML no_tac} always fails. The second returns the given theorem wrapped 
 | 
| 
 
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  | 
as a single member sequence. It 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
 | 
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  | 
|
| 
 
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  | 
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
 | 
169  | 
|
| 
 
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  | 
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
 | 
171  | 
  which means @{ML all_tac} always succeeds (but also does not make any progress 
 | 
| 
 
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  | 
with the proof).  | 
| 93 | 173  | 
|
| 99 | 174  | 
The lazy list of possible successor states shows through to the user-level  | 
175  | 
  of Isabelle when using the command \isacommand{back}. For instance in the 
 | 
|
176  | 
following proof, there are two possibilities for how to apply  | 
|
177  | 
  @{ML foo_tac'}.
 | 
|
| 93 | 178  | 
*}  | 
179  | 
||
180  | 
lemma "\<lbrakk>P \<or> Q; P \<or> Q\<rbrakk> \<Longrightarrow> Q \<or> P"  | 
|
181  | 
apply(tactic {* foo_tac' 1 *})
 | 
|
182  | 
back  | 
|
183  | 
done  | 
|
184  | 
||
| 
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
 | 
185  | 
|
| 93 | 186  | 
text {*
 | 
| 99 | 187  | 
  By using \isacommand{back}, we construct the proof that uses the
 | 
| 
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
 | 
188  | 
second assumption. In more interesting situations, different possibilities  | 
| 
 
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
 | 
189  | 
can lead to different proofs and even often need to be explored when  | 
| 
 
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
 | 
190  | 
a first proof attempt is unsuccessful.  | 
| 99 | 191  | 
|
| 93 | 192  | 
  \begin{readmore}
 | 
193  | 
  See @{ML_file "Pure/General/seq.ML"} for the implementation of lazy
 | 
|
194  | 
sequences. However in day-to-day Isabelle programming, one rarely  | 
|
195  | 
constructs sequences explicitly, but uses the predefined functions  | 
|
| 99 | 196  | 
instead.  | 
| 93 | 197  | 
  \end{readmore}
 | 
198  | 
||
| 104 | 199  | 
It might be surprising that tactics, which transform  | 
| 
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
 | 
200  | 
one proof state to the next, are functions from theorems to theorem  | 
| 
 
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
 | 
201  | 
(sequences). The surprise resolves by knowing that every  | 
| 104 | 202  | 
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
 | 
203  | 
  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
 | 
204  | 
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
 | 
205  | 
*}  | 
| 
 
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
 | 
206  | 
|
| 
 
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  | 
ML{*fun my_print_tac ctxt 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
 | 
208  | 
let  | 
| 
 
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  | 
val _ = warning (str_of_thm ctxt 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
 | 
210  | 
in  | 
| 
 
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  | 
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
 | 
212  | 
end*}  | 
| 
 
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  | 
|
| 
 
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
 | 
214  | 
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
 | 
215  | 
which prints out the given theorem (using the string-function defined  | 
| 
 
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
 | 
216  | 
  in Section~\ref{sec:printing}) and then behaves like @{ML all_tac}. We
 | 
| 
 
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
 | 
217  | 
now can inspect every proof state in the follwing proof. On the left-hand  | 
| 
 
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
 | 
218  | 
side we show the goal state as shown by the system; on the right-hand  | 
| 104 | 219  | 
  side the print out from @{ML my_print_tac}.
 | 
| 
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
 | 
220  | 
*}  | 
| 
 
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
 | 
221  | 
|
| 
 
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
 | 
222  | 
lemma shows "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B"  | 
| 
 
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
 | 
223  | 
apply(tactic {* my_print_tac @{context} *})
 | 
| 
 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 
Christian Urban <urbanc@in.tum.de> 
parents: 
99 
diff
changeset
 | 
224  | 
|
| 
 
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
 | 
225  | 
txt{* \small 
 | 
| 
 
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  | 
      \begin{tabular}{@ {}l@ {}p{0.7\textwidth}@ {}}
 | 
| 
 
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  | 
      \begin{minipage}[t]{0.3\textwidth}
 | 
| 
 
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
 | 
228  | 
      @{subgoals [display]} 
 | 
| 
 
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  | 
      \end{minipage} &
 | 
| 
 
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
 | 
230  | 
      \hfill@{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}
 | 
| 
 
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
 | 
231  | 
      \end{tabular}
 | 
| 93 | 232  | 
*}  | 
233  | 
||
| 
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
 | 
234  | 
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
 | 
235  | 
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
 | 
236  | 
|
| 
 
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
 | 
237  | 
txt{* \small 
 | 
| 
 
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
 | 
238  | 
      \begin{tabular}{@ {}l@ {}p{0.76\textwidth}@ {}}
 | 
| 
 
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
 | 
239  | 
      \begin{minipage}[t]{0.26\textwidth}
 | 
| 
 
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
 | 
240  | 
      @{subgoals [display]} 
 | 
| 
 
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  | 
      \end{minipage} &
 | 
| 
 
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  | 
      \hfill@{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}
 | 
| 
 
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  | 
      \end{tabular}
 | 
| 
 
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
 | 
244  | 
*}  | 
| 
 
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  | 
|
| 
 
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
 | 
246  | 
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
 | 
247  | 
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
 | 
248  | 
|
| 
 
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
 | 
249  | 
txt{* \small 
 | 
| 
 
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
 | 
250  | 
      \begin{tabular}{@ {}l@ {}p{0.7\textwidth}@ {}}
 | 
| 
 
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
 | 
251  | 
      \begin{minipage}[t]{0.3\textwidth}
 | 
| 
 
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
 | 
252  | 
      @{subgoals [display]} 
 | 
| 
 
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
 | 
253  | 
      \end{minipage} &
 | 
| 
 
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
 | 
254  | 
      \hfill@{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}
 | 
| 
 
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  | 
      \end{tabular}
 | 
| 
 
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  | 
|
| 
 
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(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
 | 
259  | 
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
 | 
260  | 
|
| 
 
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  | 
txt{* \small 
 | 
| 
 
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
 | 
262  | 
      \begin{tabular}{@ {}l@ {}p{0.7\textwidth}@ {}}
 | 
| 
 
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
 | 
263  | 
      \begin{minipage}[t]{0.3\textwidth}
 | 
| 
 
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
 | 
264  | 
      @{subgoals [display]} 
 | 
| 
 
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
 | 
265  | 
      \end{minipage} &
 | 
| 
 
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
 | 
266  | 
      \hfill@{text "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B"}
 | 
| 
 
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
 | 
267  | 
      \end{tabular}
 | 
| 
 
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
 | 
268  | 
*}  | 
| 
 
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
 | 
269  | 
|
| 
 
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
 | 
270  | 
done  | 
| 
 
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  | 
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
 | 
273  | 
As can be seen, internally every goal state is an implication of the form  | 
| 
 
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  | 
|
| 
 
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  | 
  @{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
 | 
276  | 
|
| 
 
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  | 
  where @{term C} is the goal to be proved and the @{term "A\<^isub>i"} are the  
 | 
| 104 | 278  | 
subgoals. So in the first step the goal state is always of the form  | 
279  | 
  @{text "C \<Longrightarrow> (C)"}. Since the goal @{term C} can potentially be an implication, 
 | 
|
280  | 
there is a ``protector'' wrapped around it (in from of an outermost constant  | 
|
281  | 
  @{text "Const (\"prop\", bool \<Rightarrow> bool)"} applied to each goal;
 | 
|
282  | 
however this constant is invisible in the print out above). This  | 
|
283  | 
prevents that premises are misinterpreted as open subgoals.  | 
|
| 
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
 | 
284  | 
  While tactics can operate on the subgoals (the @{text "A\<^isub>i"} above), they 
 | 
| 
 
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
 | 
285  | 
  are expected to leave the conclusion @{term C} intact, with the 
 | 
| 
 
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
 | 
286  | 
exception of possibly instantiating schematic variables.  | 
| 
 
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
 | 
287  | 
|
| 
 
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
 | 
288  | 
*}  | 
| 
 
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
 | 
289  | 
|
| 
 
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
 | 
290  | 
section {* Simple Tactics *}
 | 
| 93 | 291  | 
|
| 99 | 292  | 
text {*
 | 
293  | 
  As seen above, the function @{ML atac} corresponds to the assumption tactic.
 | 
|
294  | 
*}  | 
|
295  | 
||
296  | 
lemma shows "P \<Longrightarrow> P"  | 
|
| 93 | 297  | 
apply(tactic {* atac 1 *})
 | 
298  | 
done  | 
|
299  | 
||
| 99 | 300  | 
text {*
 | 
301  | 
  Similarly, @{ML rtac}, @{ML dtac}, @{ML etac} and @{ML ftac} correspond
 | 
|
| 
102
 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 
Christian Urban <urbanc@in.tum.de> 
parents: 
99 
diff
changeset
 | 
302  | 
  to @{text rule}, @{text drule}, @{text erule} and @{text frule}, 
 | 
| 104 | 303  | 
respectively. Each of them takes a theorem as argument. Below are three  | 
304  | 
examples with the resulting goal state. How  | 
|
305  | 
they work should be self-explanatory.  | 
|
| 99 | 306  | 
*}  | 
307  | 
||
308  | 
lemma shows "P \<and> Q"  | 
|
| 93 | 309  | 
apply(tactic {* rtac @{thm conjI} 1 *})
 | 
| 104 | 310  | 
txt{*\begin{minipage}{\textwidth}
 | 
311  | 
     @{subgoals [display]}
 | 
|
312  | 
     \end{minipage}*}
 | 
|
| 93 | 313  | 
(*<*)oops(*>*)  | 
314  | 
||
| 99 | 315  | 
lemma shows "P \<and> Q \<Longrightarrow> False"  | 
| 93 | 316  | 
apply(tactic {* etac @{thm conjE} 1 *})
 | 
| 104 | 317  | 
txt{*\begin{minipage}{\textwidth}
 | 
318  | 
     @{subgoals [display]}
 | 
|
319  | 
     \end{minipage}*}
 | 
|
| 93 | 320  | 
(*<*)oops(*>*)  | 
321  | 
||
322  | 
lemma shows "False \<and> True \<Longrightarrow> False"  | 
|
323  | 
apply(tactic {* dtac @{thm conjunct2} 1 *})
 | 
|
| 104 | 324  | 
txt{*\begin{minipage}{\textwidth}
 | 
325  | 
     @{subgoals [display]}
 | 
|
326  | 
     \end{minipage}*}
 | 
|
| 93 | 327  | 
(*<*)oops(*>*)  | 
328  | 
||
329  | 
text {*
 | 
|
| 99 | 330  | 
As mentioned above, most basic tactics take a number as argument, which  | 
331  | 
addresses to subgoal they are analysing.  | 
|
332  | 
*}  | 
|
333  | 
||
334  | 
lemma shows "Foo" and "P \<and> Q"  | 
|
335  | 
apply(tactic {* rtac @{thm conjI} 2 *})
 | 
|
| 104 | 336  | 
txt {*\begin{minipage}{\textwidth}
 | 
337  | 
      @{subgoals [display]}
 | 
|
338  | 
      \end{minipage}*}
 | 
|
| 99 | 339  | 
(*<*)oops(*>*)  | 
340  | 
||
341  | 
text {* 
 | 
|
342  | 
  Corresponding to @{ML rtac}, there is also the tactic @{ML resolve_tac}, which
 | 
|
343  | 
however expects a list of theorems as arguments. From this list it will apply with  | 
|
344  | 
the first applicable theorem (later theorems that are also applicable can be  | 
|
345  | 
explored via the lazy sequences mechanism). Given the abbreviation  | 
|
| 93 | 346  | 
*}  | 
347  | 
||
| 99 | 348  | 
ML{*val resolve_tac_xmp = resolve_tac [@{thm impI}, @{thm conjI}]*}
 | 
349  | 
||
350  | 
text {*
 | 
|
351  | 
  an example for @{ML resolve_tac} is the following proof where first an outermost 
 | 
|
352  | 
implication is analysed and then an outermost conjunction.  | 
|
353  | 
*}  | 
|
354  | 
||
355  | 
lemma shows "C \<longrightarrow> (A \<and> B)" and "(A \<longrightarrow> B) \<and> C"  | 
|
356  | 
apply(tactic {* resolve_tac_xmp 1 *})
 | 
|
357  | 
apply(tactic {* resolve_tac_xmp 2 *})
 | 
|
| 104 | 358  | 
txt{*\begin{minipage}{\textwidth}
 | 
359  | 
     @{subgoals [display]} 
 | 
|
360  | 
     \end{minipage}*}
 | 
|
| 99 | 361  | 
(*<*)oops(*>*)  | 
362  | 
||
363  | 
text {* 
 | 
|
364  | 
  Similarly versions exists for @{ML atac} (@{ML assume_tac}), @{ML etac} 
 | 
|
365  | 
  (@{ML eresolve_tac}) and so on.
 | 
|
366  | 
||
367  | 
  The tactic @{ML print_tac} is useful for low-level debugging of tactics: it
 | 
|
368  | 
prints out a message and the current goal state.  | 
|
369  | 
*}  | 
|
370  | 
||
371  | 
lemma shows "False \<Longrightarrow> True"  | 
|
| 93 | 372  | 
apply(tactic {* print_tac "foo message" *})
 | 
373  | 
(*<*)oops(*>*)  | 
|
374  | 
||
| 99 | 375  | 
text {*
 | 
| 104 | 376  | 
|
| 99 | 377  | 
(FIXME explain RS MRS)  | 
| 
95
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
378  | 
|
| 104 | 379  | 
Often proofs involve elaborate operations on assumptions and  | 
380  | 
  @{text "\<And>"}-quantified variables. To do such operations on the ML-level 
 | 
|
381  | 
using the basic tactics is very unwieldy and brittle. Some convenience and  | 
|
| 99 | 382  | 
  safety is provided by the tactic @{ML SUBPROOF}. This tactic fixes the parameters 
 | 
383  | 
and binds the various components of a proof state into a record.  | 
|
| 
95
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
384  | 
*}  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
385  | 
|
| 99 | 386  | 
text_raw{*
 | 
387  | 
\begin{figure}
 | 
|
388  | 
\begin{isabelle}
 | 
|
389  | 
*}  | 
|
| 
95
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
390  | 
ML{*fun sp_tac {prems, params, asms, concl, context, schematics} = 
 | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
391  | 
let  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
392  | 
val str_of_params = str_of_cterms context params  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
393  | 
val str_of_asms = str_of_cterms context asms  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
394  | 
val str_of_concl = str_of_cterm context concl  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
395  | 
val str_of_prems = str_of_thms context prems  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
396  | 
val str_of_schms = str_of_cterms context (snd schematics)  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
397  | 
|
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
398  | 
    val _ = (warning ("params: " ^ str_of_params);
 | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
399  | 
             warning ("schematics: " ^ str_of_schms);
 | 
| 
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
 | 
400  | 
             warning ("assumptions: " ^ str_of_asms);
 | 
| 
 
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
 | 
401  | 
             warning ("conclusion: " ^ str_of_concl);
 | 
| 
 
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
 | 
402  | 
             warning ("premises: " ^ str_of_prems))
 | 
| 
95
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
403  | 
in  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
404  | 
no_tac  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
405  | 
end*}  | 
| 99 | 406  | 
text_raw{*
 | 
407  | 
\end{isabelle}
 | 
|
408  | 
\caption{A function that prints out the various parameters provided by the tactic
 | 
|
409  | 
 @{ML SUBPROOF}. It uses the functions extracting strings from @{ML_type cterm}s 
 | 
|
410  | 
  and @{ML_type thm}s, which are defined in Section~\ref{sec:printing}.\label{fig:sptac}}
 | 
|
411  | 
\end{figure}
 | 
|
412  | 
*}  | 
|
| 
95
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
413  | 
|
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
414  | 
text {*
 | 
| 99 | 415  | 
  To see what happens, assume the function defined in Figure~\ref{fig:sptac}, which
 | 
416  | 
takes a record as argument and just prints out the content of this record (using the  | 
|
417  | 
  string transformation functions defined in Section~\ref{sec:printing}). Consider
 | 
|
418  | 
now the proof  | 
|
| 
95
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
419  | 
*}  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
420  | 
|
| 99 | 421  | 
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
 | 
422  | 
apply(tactic {* SUBPROOF sp_tac @{context} 1 *})?
 | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
423  | 
|
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
424  | 
txt {*
 | 
| 99 | 425  | 
which yields the printout:  | 
| 
95
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
426  | 
|
| 99 | 427  | 
  \begin{quote}\small
 | 
| 
95
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
428  | 
  \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
 | 
429  | 
  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
 | 
430  | 
  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
 | 
431  | 
  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
 | 
432  | 
  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
 | 
433  | 
  premises:    & @{term "A x y"}
 | 
| 
95
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
434  | 
  \end{tabular}
 | 
| 99 | 435  | 
  \end{quote}
 | 
436  | 
||
437  | 
  Note in the actual output the brown colour of the variables @{term x} and 
 | 
|
438  | 
  @{term y}. Although parameters in the original goal, they are fixed inside
 | 
|
439  | 
  the subproof. Similarly the schematic variable @{term z}. The assumption 
 | 
|
| 104 | 440  | 
  @{prop "A x y"} is bound once as @{ML_type cterm} to the record-variable 
 | 
441  | 
  @{text asms} and another time 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
 | 
442  | 
|
| 99 | 443  | 
  Notice also that we had to append @{text "?"} to \isacommand{apply}. The 
 | 
444  | 
  reason is that @{ML SUBPROOF} normally expects that the subgoal is solved completely.
 | 
|
445  | 
  Since in the function @{ML sp_tac} we returned the tactic @{ML no_tac}, the subproof
 | 
|
446  | 
obviously fails. The question-mark allows us to recover from this failure  | 
|
447  | 
in a graceful manner so that the warning messages are not overwritten  | 
|
| 104 | 448  | 
by an error message.  | 
| 99 | 449  | 
|
450  | 
  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
 | 
451  | 
*}  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
452  | 
|
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
453  | 
apply(rule impI)  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
454  | 
apply(tactic {* SUBPROOF sp_tac @{context} 1 *})?
 | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
455  | 
|
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
456  | 
txt {*
 | 
| 99 | 457  | 
  then @{ML SUBPROOF} prints out 
 | 
| 
95
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
458  | 
|
| 99 | 459  | 
  \begin{quote}\small
 | 
| 
95
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
460  | 
  \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
 | 
461  | 
  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
 | 
462  | 
  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
 | 
463  | 
  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
 | 
464  | 
  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
 | 
465  | 
  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
 | 
466  | 
  \end{tabular}
 | 
| 99 | 467  | 
  \end{quote}
 | 
| 
95
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
468  | 
*}  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
469  | 
(*<*)oops(*>*)  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
470  | 
|
| 99 | 471  | 
text {*
 | 
| 104 | 472  | 
  where we now also have @{term "B y x"} as an assumption.
 | 
| 99 | 473  | 
|
474  | 
  One convenience of @{ML SUBPROOF} is that we can apply assumption
 | 
|
475  | 
  using the usual tactics, because the parameter @{text prems} 
 | 
|
476  | 
contains the assumptions as theorems. With this we can easily  | 
|
477  | 
  implement a tactic that almost behaves like @{ML atac}:
 | 
|
478  | 
*}  | 
|
479  | 
||
| 104 | 480  | 
ML{*val atac' = SUBPROOF (fn {prems, ...} => resolve_tac prems 1)*}
 | 
| 99 | 481  | 
lemma shows "\<And>x y. \<lbrakk>B x y; A x y; C x y\<rbrakk> \<Longrightarrow> A x y"  | 
| 104 | 482  | 
apply(tactic {* atac' @{context} 1 *})
 | 
| 99 | 483  | 
txt{* yields
 | 
484  | 
      @{subgoals [display]} *}
 | 
|
485  | 
(*<*)oops(*>*)  | 
|
486  | 
||
| 104 | 487  | 
text {*
 | 
488  | 
The restriction in this tactic is that it cannot instantiate any  | 
|
489  | 
schematic variables. This might be seen as a defect, but is actually  | 
|
490  | 
  an advantage in the situations for which @{ML SUBPROOF} was designed: 
 | 
|
491  | 
the reason is that instantiation of schematic variables can affect  | 
|
492  | 
  several goals and can render them unprovable. @{ML SUBPROOF} is meant 
 | 
|
493  | 
to avoid this.  | 
|
494  | 
||
495  | 
  Notice that @{ML atac'} calls @{ML resolve_tac} with the subgoal
 | 
|
496  | 
  number @{text "1"} and also the ``outer'' call to @{ML SUBPROOF} in 
 | 
|
497  | 
  the \isacommand{apply}-step uses @{text "1"}. Another advantage 
 | 
|
498  | 
  of @{ML SUBGOAL} is that the addressing  inside it is completely 
 | 
|
499  | 
local to the proof inside. It is therefore possible to also apply  | 
|
500  | 
  @{ML atac'} to the second goal:
 | 
|
501  | 
*}  | 
|
502  | 
||
503  | 
lemma shows "True" and "\<And>x y. \<lbrakk>B x y; A x y; C x y\<rbrakk> \<Longrightarrow> A x y"  | 
|
504  | 
apply(tactic {* atac' @{context} 2 *})
 | 
|
505  | 
txt{* This gives:
 | 
|
506  | 
      @{subgoals [display]} *}
 | 
|
507  | 
(*<*)oops(*>*)  | 
|
508  | 
||
| 
95
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
509  | 
|
| 93 | 510  | 
text {*
 | 
| 104 | 511  | 
  A similar but less powerful function than @{ML SUBPROOF} is @{ML SUBGOAL}. 
 | 
512  | 
It allows you to inspect a subgoal specified by a number. With this we can  | 
|
513  | 
implement a little tactic that applies a rule corresponding to its  | 
|
514  | 
topmost connective. The tactic should only apply ``safe'' rules (that is  | 
|
515  | 
which do not render the goal unprovable). For this we can write:  | 
|
| 93 | 516  | 
*}  | 
517  | 
||
| 99 | 518  | 
ML %linenumbers{*fun select_tac (t,i) =
 | 
519  | 
case t of  | 
|
520  | 
     @{term "Trueprop"} $ t' => select_tac (t',i)
 | 
|
521  | 
   | @{term "op \<and>"} $ _ $ _ => rtac @{thm conjI} i
 | 
|
522  | 
   | @{term "op \<longrightarrow>"} $ _ $ _ => rtac @{thm impI} i
 | 
|
523  | 
   | @{term "Not"} $ _ => rtac @{thm notI} i
 | 
|
524  | 
   | Const (@{const_name "All"}, _) $ _ => rtac @{thm allI} i
 | 
|
| 104 | 525  | 
| _ => all_tac*}  | 
| 99 | 526  | 
|
| 104 | 527  | 
lemma shows "A \<and> B" "A \<longrightarrow> B" "\<forall>x. C x" "D \<Longrightarrow> E"  | 
528  | 
apply(tactic {* SUBGOAL select_tac 4 *})
 | 
|
529  | 
apply(tactic {* SUBGOAL select_tac 3 *})
 | 
|
530  | 
apply(tactic {* SUBGOAL select_tac 2 *})
 | 
|
| 99 | 531  | 
apply(tactic {* SUBGOAL select_tac 1 *})
 | 
532  | 
txt{* @{subgoals [display]} *}
 | 
|
533  | 
(*<*)oops(*>*)  | 
|
534  | 
||
535  | 
text {*
 | 
|
536  | 
However, this example is contrived, as there are much simpler ways  | 
|
537  | 
to implement a proof procedure like the one above. They will be explained  | 
|
538  | 
in the next section.  | 
|
539  | 
||
| 104 | 540  | 
(Notice that we applied the goals in reverse order)  | 
541  | 
||
| 99 | 542  | 
  A variant of @{ML SUBGOAL} is @{ML CSUBGOAL} which allows access to the goal
 | 
543  | 
  as @{ML_type cterm} instead of a @{ML_type term}.
 | 
|
| 93 | 544  | 
*}  | 
545  | 
||
| 99 | 546  | 
|
| 93 | 547  | 
section {* Operations on Tactics *}
 | 
548  | 
||
| 99 | 549  | 
text {* @{ML THEN} *}
 | 
| 93 | 550  | 
|
| 99 | 551  | 
lemma shows "(Foo \<and> Bar) \<and> False"  | 
552  | 
apply(tactic {* (rtac @{thm conjI} 1) 
 | 
|
553  | 
                 THEN (rtac @{thm conjI} 1) *})
 | 
|
| 93 | 554  | 
txt {* @{subgoals [display]} *}
 | 
555  | 
(*<*)oops(*>*)  | 
|
556  | 
||
| 99 | 557  | 
ML{*val orelse_xmp = (rtac @{thm disjI1} ORELSE' rtac @{thm conjI})*}
 | 
558  | 
||
559  | 
lemma shows "True \<and> False" and "Foo \<or> Bar"  | 
|
560  | 
apply(tactic {* orelse_xmp 1 *})
 | 
|
561  | 
apply(tactic {* orelse_xmp 3 *})
 | 
|
| 93 | 562  | 
txt {* @{subgoals [display]} *}
 | 
563  | 
(*<*)oops(*>*)  | 
|
564  | 
||
565  | 
||
566  | 
text {*
 | 
|
| 99 | 567  | 
  @{ML EVERY} @{ML REPEAT} @{ML DETERM}
 | 
| 93 | 568  | 
|
569  | 
  @{ML rewrite_goals_tac}
 | 
|
570  | 
  @{ML cut_facts_tac}
 | 
|
571  | 
  @{ML ObjectLogic.full_atomize_tac}
 | 
|
572  | 
  @{ML ObjectLogic.rulify_tac}
 | 
|
573  | 
  @{ML resolve_tac}
 | 
|
574  | 
*}  | 
|
575  | 
||
| 
95
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
576  | 
section {* Structured Proofs *}
 | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
577  | 
|
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
578  | 
lemma True  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
579  | 
proof  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
580  | 
|
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
581  | 
  {
 | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
582  | 
fix A B C  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
583  | 
assume r: "A & B \<Longrightarrow> C"  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
584  | 
assume A B  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
585  | 
then have "A & B" ..  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
586  | 
then have C by (rule r)  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
587  | 
}  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
588  | 
|
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
589  | 
  {
 | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
590  | 
fix A B C  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
591  | 
assume r: "A & B \<Longrightarrow> C"  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
592  | 
assume A B  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
593  | 
note conjI [OF this]  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
594  | 
note r [OF this]  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
595  | 
}  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
596  | 
oops  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
597  | 
|
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
598  | 
ML {* fun prop ctxt s =
 | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
599  | 
Thm.cterm_of (ProofContext.theory_of ctxt) (Syntax.read_prop ctxt s) *}  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
600  | 
|
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
601  | 
ML {* 
 | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
602  | 
  val ctxt0 = @{context};
 | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
603  | 
val ctxt = ctxt0;  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
604  | 
val (_, ctxt) = Variable.add_fixes ["A", "B", "C"] ctxt;  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
605  | 
val ([r], ctxt) = Assumption.add_assumes [prop ctxt "A & B \<Longrightarrow> C"] ctxt;  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
606  | 
val (this, ctxt) = Assumption.add_assumes [prop ctxt "A", prop ctxt "B"] ctxt;  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
607  | 
  val this = [@{thm conjI} OF this]; 
 | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
608  | 
val this = r OF this;  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
609  | 
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
 | 
610  | 
val this = Variable.export ctxt ctxt0 [this]  | 
| 
 
7235374f34c8
added some preliminary notes about SUBPROOF
 
Christian Urban <urbanc@in.tum.de> 
parents: 
93 
diff
changeset
 | 
611  | 
*}  | 
| 93 | 612  | 
|
613  | 
||
| 
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
 | 
614  | 
|
| 93 | 615  | 
end  |