| author | bulwahn | 
| Wed, 19 Oct 2011 09:11:20 +0200 | |
| changeset 45190 | 58e33a125f32 | 
| parent 43278 | 1fbdcebb364b | 
| child 46463 | 3d0629a9ffca | 
| permissions | -rw-r--r-- | 
| 32169 | 1  | 
(* Title: Pure/tactical.ML  | 
| 
2244
 
dacee519738a
Converted I/O operatios for Basis Library compatibility
 
paulson 
parents: 
2158 
diff
changeset
 | 
2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 0 | 3  | 
|
| 16179 | 4  | 
Tacticals.  | 
| 0 | 5  | 
*)  | 
6  | 
||
| 4602 | 7  | 
infix 1 THEN THEN' THEN_ALL_NEW;  | 
| 0 | 8  | 
infix 0 ORELSE APPEND INTLEAVE ORELSE' APPEND' INTLEAVE';  | 
| 671 | 9  | 
infix 0 THEN_ELSE;  | 
10  | 
||
| 0 | 11  | 
signature TACTICAL =  | 
| 11916 | 12  | 
sig  | 
| 23538 | 13  | 
type tactic = thm -> thm Seq.seq  | 
14  | 
val THEN: tactic * tactic -> tactic  | 
|
15  | 
val ORELSE: tactic * tactic -> tactic  | 
|
16  | 
val APPEND: tactic * tactic -> tactic  | 
|
17  | 
val INTLEAVE: tactic * tactic -> tactic  | 
|
18  | 
val THEN_ELSE: tactic * (tactic*tactic) -> tactic  | 
|
19  | 
  val THEN': ('a -> tactic) * ('a -> tactic) -> 'a -> tactic
 | 
|
20  | 
  val ORELSE': ('a -> tactic) * ('a -> tactic) -> 'a -> tactic
 | 
|
21  | 
  val APPEND': ('a -> tactic) * ('a -> tactic) -> 'a -> tactic
 | 
|
22  | 
  val INTLEAVE': ('a -> tactic) * ('a -> tactic) -> 'a -> tactic
 | 
|
23  | 
val all_tac: tactic  | 
|
24  | 
val no_tac: tactic  | 
|
25  | 
val DETERM: tactic -> tactic  | 
|
26  | 
val COND: (thm -> bool) -> tactic -> tactic -> tactic  | 
|
27  | 
val TRY: tactic -> tactic  | 
|
28  | 
val EVERY: tactic list -> tactic  | 
|
29  | 
  val EVERY': ('a -> tactic) list -> 'a -> tactic
 | 
|
30  | 
val EVERY1: (int -> tactic) list -> tactic  | 
|
31  | 
val FIRST: tactic list -> tactic  | 
|
32  | 
  val FIRST': ('a -> tactic) list -> 'a -> tactic
 | 
|
33  | 
val FIRST1: (int -> tactic) list -> tactic  | 
|
34  | 
val RANGE: (int -> tactic) list -> int -> tactic  | 
|
35  | 
val print_tac: string -> tactic  | 
|
36  | 
val pause_tac: tactic  | 
|
| 32738 | 37  | 
val trace_REPEAT: bool Unsynchronized.ref  | 
38  | 
val suppress_tracing: bool Unsynchronized.ref  | 
|
39  | 
val tracify: bool Unsynchronized.ref -> tactic -> tactic  | 
|
| 23538 | 40  | 
val traced_tac: (thm -> (thm * thm Seq.seq) option) -> tactic  | 
41  | 
val DETERM_UNTIL: (thm -> bool) -> tactic -> tactic  | 
|
42  | 
val REPEAT_DETERM_N: int -> tactic -> tactic  | 
|
43  | 
val REPEAT_DETERM: tactic -> tactic  | 
|
44  | 
val REPEAT: tactic -> tactic  | 
|
45  | 
val REPEAT_DETERM1: tactic -> tactic  | 
|
46  | 
val REPEAT1: tactic -> tactic  | 
|
47  | 
val FILTER: (thm -> bool) -> tactic -> tactic  | 
|
48  | 
val CHANGED: tactic -> tactic  | 
|
49  | 
val CHANGED_PROP: tactic -> tactic  | 
|
50  | 
val ALLGOALS: (int -> tactic) -> tactic  | 
|
51  | 
val SOMEGOAL: (int -> tactic) -> tactic  | 
|
52  | 
val FIRSTGOAL: (int -> tactic) -> tactic  | 
|
53  | 
val REPEAT_SOME: (int -> tactic) -> tactic  | 
|
54  | 
val REPEAT_DETERM_SOME: (int -> tactic) -> tactic  | 
|
55  | 
val REPEAT_FIRST: (int -> tactic) -> tactic  | 
|
| 
703
 
3a5cd2883581
Pure/tctical/REPEAT_DETERM_N,REPEAT_DETERM1,REPEAT_DETERM_FIRST,
 
lcp 
parents: 
671 
diff
changeset
 | 
56  | 
val REPEAT_DETERM_FIRST: (int -> tactic) -> tactic  | 
| 23538 | 57  | 
val TRYALL: (int -> tactic) -> tactic  | 
58  | 
val CSUBGOAL: ((cterm * int) -> tactic) -> int -> tactic  | 
|
59  | 
val SUBGOAL: ((term * int) -> tactic) -> int -> tactic  | 
|
60  | 
val CHANGED_GOAL: (int -> tactic) -> int -> tactic  | 
|
| 
34885
 
6587c24ef6d8
added SOLVED' -- a more direct version of THEN_ALL_NEW (K no_tac) -- strictly speaking it does not even depend on subgoal addressing, but it would be too confusing without it;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
61  | 
val SOLVED': (int -> tactic) -> int -> tactic  | 
| 23538 | 62  | 
val THEN_ALL_NEW: (int -> tactic) * (int -> tactic) -> int -> tactic  | 
63  | 
val REPEAT_ALL_NEW: (int -> tactic) -> int -> tactic  | 
|
64  | 
val PRIMSEQ: (thm -> thm Seq.seq) -> tactic  | 
|
65  | 
val PRIMITIVE: (thm -> thm) -> tactic  | 
|
66  | 
val SINGLE: tactic -> thm -> thm option  | 
|
67  | 
val CONVERSION: conv -> int -> tactic  | 
|
| 11916 | 68  | 
end;  | 
| 0 | 69  | 
|
| 13108 | 70  | 
structure Tactical : TACTICAL =  | 
| 0 | 71  | 
struct  | 
72  | 
||
73  | 
(**** Tactics ****)  | 
|
74  | 
||
75  | 
(*A tactic maps a proof tree to a sequence of proof trees:  | 
|
76  | 
if length of sequence = 0 then the tactic does not apply;  | 
|
77  | 
if length > 1 then backtracking on the alternatives can occur.*)  | 
|
78  | 
||
| 4270 | 79  | 
type tactic = thm -> thm Seq.seq;  | 
| 0 | 80  | 
|
81  | 
||
82  | 
(*** LCF-style tacticals ***)  | 
|
83  | 
||
84  | 
(*the tactical THEN performs one tactic followed by another*)  | 
|
| 17344 | 85  | 
fun (tac1 THEN tac2) st = Seq.maps tac2 (tac1 st);  | 
| 0 | 86  | 
|
87  | 
||
88  | 
(*The tactical ORELSE uses the first tactic that returns a nonempty sequence.  | 
|
89  | 
Like in LCF, ORELSE commits to either tac1 or tac2 immediately.  | 
|
90  | 
Does not backtrack to tac2 if tac1 was initially chosen. *)  | 
|
| 1502 | 91  | 
fun (tac1 ORELSE tac2) st =  | 
| 4270 | 92  | 
case Seq.pull(tac1 st) of  | 
| 15531 | 93  | 
NONE => tac2 st  | 
| 4270 | 94  | 
| sequencecell => Seq.make(fn()=> sequencecell);  | 
| 0 | 95  | 
|
96  | 
||
97  | 
(*The tactical APPEND combines the results of two tactics.  | 
|
98  | 
Like ORELSE, but allows backtracking on both tac1 and tac2.  | 
|
99  | 
The tactic tac2 is not applied until needed.*)  | 
|
| 13108 | 100  | 
fun (tac1 APPEND tac2) st =  | 
| 19861 | 101  | 
Seq.append (tac1 st) (Seq.make(fn()=> Seq.pull (tac2 st)));  | 
| 0 | 102  | 
|
103  | 
(*Like APPEND, but interleaves results of tac1 and tac2.*)  | 
|
| 13108 | 104  | 
fun (tac1 INTLEAVE tac2) st =  | 
| 4270 | 105  | 
Seq.interleave(tac1 st,  | 
106  | 
Seq.make(fn()=> Seq.pull (tac2 st)));  | 
|
| 0 | 107  | 
|
| 671 | 108  | 
(*Conditional tactic.  | 
| 
2244
 
dacee519738a
Converted I/O operatios for Basis Library compatibility
 
paulson 
parents: 
2158 
diff
changeset
 | 
109  | 
tac1 ORELSE tac2 = tac1 THEN_ELSE (all_tac, tac2)  | 
| 
 
dacee519738a
Converted I/O operatios for Basis Library compatibility
 
paulson 
parents: 
2158 
diff
changeset
 | 
110  | 
tac1 THEN tac2 = tac1 THEN_ELSE (tac2, no_tac)  | 
| 671 | 111  | 
*)  | 
| 13108 | 112  | 
fun (tac THEN_ELSE (tac1, tac2)) st =  | 
| 4270 | 113  | 
case Seq.pull(tac st) of  | 
| 17344 | 114  | 
NONE => tac2 st (*failed; try tactic 2*)  | 
115  | 
| seqcell => Seq.maps tac1 (Seq.make(fn()=> seqcell)); (*succeeded; use tactic 1*)  | 
|
| 671 | 116  | 
|
117  | 
||
| 0 | 118  | 
(*Versions for combining tactic-valued functions, as in  | 
119  | 
SOMEGOAL (resolve_tac rls THEN' assume_tac) *)  | 
|
| 1502 | 120  | 
fun (tac1 THEN' tac2) x = tac1 x THEN tac2 x;  | 
121  | 
fun (tac1 ORELSE' tac2) x = tac1 x ORELSE tac2 x;  | 
|
122  | 
fun (tac1 APPEND' tac2) x = tac1 x APPEND tac2 x;  | 
|
123  | 
fun (tac1 INTLEAVE' tac2) x = tac1 x INTLEAVE tac2 x;  | 
|
| 0 | 124  | 
|
125  | 
(*passes all proofs through unchanged; identity of THEN*)  | 
|
| 4270 | 126  | 
fun all_tac st = Seq.single st;  | 
| 0 | 127  | 
|
128  | 
(*passes no proofs through; identity of ORELSE and APPEND*)  | 
|
| 4270 | 129  | 
fun no_tac st = Seq.empty;  | 
| 0 | 130  | 
|
131  | 
||
132  | 
(*Make a tactic deterministic by chopping the tail of the proof sequence*)  | 
|
| 12851 | 133  | 
fun DETERM tac = Seq.DETERM tac;  | 
| 0 | 134  | 
|
135  | 
(*Conditional tactical: testfun controls which tactic to use next.  | 
|
136  | 
Beware: due to eager evaluation, both thentac and elsetac are evaluated.*)  | 
|
| 1502 | 137  | 
fun COND testfun thenf elsef = (fn prf =>  | 
| 0 | 138  | 
if testfun prf then thenf prf else elsef prf);  | 
139  | 
||
140  | 
(*Do the tactic or else do nothing*)  | 
|
141  | 
fun TRY tac = tac ORELSE all_tac;  | 
|
142  | 
||
| 
2672
 
85d7e800d754
Replaced "flat" by the Basis Library function List.concat
 
paulson 
parents: 
2627 
diff
changeset
 | 
143  | 
(*** List-oriented tactics ***)  | 
| 
 
85d7e800d754
Replaced "flat" by the Basis Library function List.concat
 
paulson 
parents: 
2627 
diff
changeset
 | 
144  | 
|
| 
 
85d7e800d754
Replaced "flat" by the Basis Library function List.concat
 
paulson 
parents: 
2627 
diff
changeset
 | 
145  | 
local  | 
| 
 
85d7e800d754
Replaced "flat" by the Basis Library function List.concat
 
paulson 
parents: 
2627 
diff
changeset
 | 
146  | 
(*This version of EVERY avoids backtracking over repeated states*)  | 
| 
 
85d7e800d754
Replaced "flat" by the Basis Library function List.concat
 
paulson 
parents: 
2627 
diff
changeset
 | 
147  | 
|
| 13108 | 148  | 
fun EVY (trail, []) st =  | 
| 15531 | 149  | 
Seq.make (fn()=> SOME(st,  | 
| 13108 | 150  | 
Seq.make (fn()=> Seq.pull (evyBack trail))))  | 
151  | 
| EVY (trail, tac::tacs) st =  | 
|
152  | 
case Seq.pull(tac st) of  | 
|
| 15531 | 153  | 
NONE => evyBack trail (*failed: backtrack*)  | 
154  | 
| SOME(st',q) => EVY ((st',q,tacs)::trail, tacs) st'  | 
|
| 4270 | 155  | 
and evyBack [] = Seq.empty (*no alternatives*)  | 
| 
2672
 
85d7e800d754
Replaced "flat" by the Basis Library function List.concat
 
paulson 
parents: 
2627 
diff
changeset
 | 
156  | 
| evyBack ((st',q,tacs)::trail) =  | 
| 13108 | 157  | 
case Seq.pull q of  | 
| 15531 | 158  | 
NONE => evyBack trail  | 
| 
22360
 
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
 
wenzelm 
parents: 
20664 
diff
changeset
 | 
159  | 
| SOME(st,q') => if Thm.eq_thm (st',st)  | 
| 13108 | 160  | 
then evyBack ((st',q',tacs)::trail)  | 
161  | 
else EVY ((st,q',tacs)::trail, tacs) st  | 
|
| 
2672
 
85d7e800d754
Replaced "flat" by the Basis Library function List.concat
 
paulson 
parents: 
2627 
diff
changeset
 | 
162  | 
in  | 
| 
 
85d7e800d754
Replaced "flat" by the Basis Library function List.concat
 
paulson 
parents: 
2627 
diff
changeset
 | 
163  | 
|
| 
 
85d7e800d754
Replaced "flat" by the Basis Library function List.concat
 
paulson 
parents: 
2627 
diff
changeset
 | 
164  | 
(* EVERY [tac1,...,tacn] equals tac1 THEN ... THEN tacn *)  | 
| 
 
85d7e800d754
Replaced "flat" by the Basis Library function List.concat
 
paulson 
parents: 
2627 
diff
changeset
 | 
165  | 
fun EVERY tacs = EVY ([], tacs);  | 
| 
 
85d7e800d754
Replaced "flat" by the Basis Library function List.concat
 
paulson 
parents: 
2627 
diff
changeset
 | 
166  | 
end;  | 
| 2627 | 167  | 
|
| 0 | 168  | 
|
| 1502 | 169  | 
(* EVERY' [tac1,...,tacn] i equals tac1 i THEN ... THEN tacn i *)  | 
| 
2672
 
85d7e800d754
Replaced "flat" by the Basis Library function List.concat
 
paulson 
parents: 
2627 
diff
changeset
 | 
170  | 
fun EVERY' tacs i = EVERY (map (fn f => f i) tacs);  | 
| 0 | 171  | 
|
172  | 
(*Apply every tactic to 1*)  | 
|
| 1502 | 173  | 
fun EVERY1 tacs = EVERY' tacs 1;  | 
| 0 | 174  | 
|
175  | 
(* FIRST [tac1,...,tacn] equals tac1 ORELSE ... ORELSE tacn *)  | 
|
| 23178 | 176  | 
fun FIRST tacs = fold_rev (curry op ORELSE) tacs no_tac;  | 
| 0 | 177  | 
|
| 1502 | 178  | 
(* FIRST' [tac1,...,tacn] i equals tac1 i ORELSE ... ORELSE tacn i *)  | 
| 23178 | 179  | 
fun FIRST' tacs = fold_rev (curry op ORELSE') tacs (K no_tac);  | 
| 0 | 180  | 
|
181  | 
(*Apply first tactic to 1*)  | 
|
| 1502 | 182  | 
fun FIRST1 tacs = FIRST' tacs 1;  | 
| 0 | 183  | 
|
| 11916 | 184  | 
(*Apply tactics on consecutive subgoals*)  | 
185  | 
fun RANGE [] _ = all_tac  | 
|
186  | 
| RANGE (tac :: tacs) i = RANGE tacs (i + 1) THEN tac i;  | 
|
187  | 
||
| 0 | 188  | 
|
189  | 
(*** Tracing tactics ***)  | 
|
190  | 
||
191  | 
(*Print the current proof state and pass it on.*)  | 
|
| 
32145
 
220c9e439d39
clarified pretty_goals, pretty_thm_aux: plain context;
 
wenzelm 
parents: 
32089 
diff
changeset
 | 
192  | 
fun print_tac msg st =  | 
| 
 
220c9e439d39
clarified pretty_goals, pretty_thm_aux: plain context;
 
wenzelm 
parents: 
32089 
diff
changeset
 | 
193  | 
(tracing (msg ^ "\n" ^  | 
| 
39125
 
f45d332a90e3
pretty_goals: turned some global references and function arguments into configuration options (goals_limit = 10, goals_total = true, show_main_goal = false) depending on the context;
 
wenzelm 
parents: 
34885 
diff
changeset
 | 
194  | 
Pretty.string_of (Pretty.chunks (Goal_Display.pretty_goals_without_context st)));  | 
| 
32145
 
220c9e439d39
clarified pretty_goals, pretty_thm_aux: plain context;
 
wenzelm 
parents: 
32089 
diff
changeset
 | 
195  | 
Seq.single st);  | 
| 0 | 196  | 
|
197  | 
(*Pause until a line is typed -- if non-empty then fail. *)  | 
|
| 13108 | 198  | 
fun pause_tac st =  | 
| 12262 | 199  | 
(tracing "** Press RETURN to continue:";  | 
| 
24359
 
44556727197a
TextIO.inputLine: non-critical (assume exclusive ownership);
 
wenzelm 
parents: 
23922 
diff
changeset
 | 
200  | 
if TextIO.inputLine TextIO.stdIn = SOME "\n" then Seq.single st  | 
| 12262 | 201  | 
else (tracing "Goodbye"; Seq.empty));  | 
| 0 | 202  | 
|
203  | 
exception TRACE_EXIT of thm  | 
|
204  | 
and TRACE_QUIT;  | 
|
205  | 
||
| 
631
 
8bc44f7bbab8
Pure/tctical/suppress_tracing: new; can now switch tracing off until the
 
lcp 
parents: 
230 
diff
changeset
 | 
206  | 
(*Tracing flags*)  | 
| 32738 | 207  | 
val trace_REPEAT= Unsynchronized.ref false  | 
208  | 
and suppress_tracing = Unsynchronized.ref false;  | 
|
| 
631
 
8bc44f7bbab8
Pure/tctical/suppress_tracing: new; can now switch tracing off until the
 
lcp 
parents: 
230 
diff
changeset
 | 
209  | 
|
| 0 | 210  | 
(*Handle all tracing commands for current state and tactic *)  | 
| 13108 | 211  | 
fun exec_trace_command flag (tac, st) =  | 
| 
24359
 
44556727197a
TextIO.inputLine: non-critical (assume exclusive ownership);
 
wenzelm 
parents: 
23922 
diff
changeset
 | 
212  | 
case TextIO.inputLine TextIO.stdIn of  | 
| 
23139
 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 
wenzelm 
parents: 
22596 
diff
changeset
 | 
213  | 
SOME "\n" => tac st  | 
| 
 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 
wenzelm 
parents: 
22596 
diff
changeset
 | 
214  | 
| SOME "f\n" => Seq.empty  | 
| 32738 | 215  | 
| SOME "o\n" => (flag := false; tac st)  | 
216  | 
| SOME "s\n" => (suppress_tracing := true; tac st)  | 
|
| 
23139
 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 
wenzelm 
parents: 
22596 
diff
changeset
 | 
217  | 
| SOME "x\n" => (tracing "Exiting now"; raise (TRACE_EXIT st))  | 
| 
 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 
wenzelm 
parents: 
22596 
diff
changeset
 | 
218  | 
| SOME "quit\n" => raise TRACE_QUIT  | 
| 12262 | 219  | 
| _ => (tracing  | 
| 0 | 220  | 
"Type RETURN to continue or...\n\  | 
221  | 
\ f - to fail here\n\  | 
|
222  | 
\ o - to switch tracing off\n\  | 
|
| 
631
 
8bc44f7bbab8
Pure/tctical/suppress_tracing: new; can now switch tracing off until the
 
lcp 
parents: 
230 
diff
changeset
 | 
223  | 
\ s - to suppress tracing until next entry to a tactical\n\  | 
| 0 | 224  | 
\ x - to exit at this point\n\  | 
225  | 
\ quit - to abort this tracing run\n\  | 
|
| 1502 | 226  | 
\** Well? " ; exec_trace_command flag (tac, st));  | 
| 0 | 227  | 
|
228  | 
||
229  | 
(*Extract from a tactic, a thm->thm seq function that handles tracing*)  | 
|
| 1502 | 230  | 
fun tracify flag tac st =  | 
| 
32145
 
220c9e439d39
clarified pretty_goals, pretty_thm_aux: plain context;
 
wenzelm 
parents: 
32089 
diff
changeset
 | 
231  | 
if !flag andalso not (!suppress_tracing) then  | 
| 
32168
 
116461b8fc01
eliminated print_goals_without_context -- proper pretty printing;
 
wenzelm 
parents: 
32145 
diff
changeset
 | 
232  | 
(tracing (Pretty.string_of (Pretty.chunks  | 
| 
39125
 
f45d332a90e3
pretty_goals: turned some global references and function arguments into configuration options (goals_limit = 10, goals_total = true, show_main_goal = false) depending on the context;
 
wenzelm 
parents: 
34885 
diff
changeset
 | 
233  | 
(Goal_Display.pretty_goals_without_context st @  | 
| 
32168
 
116461b8fc01
eliminated print_goals_without_context -- proper pretty printing;
 
wenzelm 
parents: 
32145 
diff
changeset
 | 
234  | 
[Pretty.str "** Press RETURN to continue:"])));  | 
| 
 
116461b8fc01
eliminated print_goals_without_context -- proper pretty printing;
 
wenzelm 
parents: 
32145 
diff
changeset
 | 
235  | 
exec_trace_command flag (tac, st))  | 
| 1502 | 236  | 
else tac st;  | 
| 0 | 237  | 
|
238  | 
(*Create a tactic whose outcome is given by seqf, handling TRACE_EXIT*)  | 
|
| 13108 | 239  | 
fun traced_tac seqf st =  | 
| 
631
 
8bc44f7bbab8
Pure/tctical/suppress_tracing: new; can now switch tracing off until the
 
lcp 
parents: 
230 
diff
changeset
 | 
240  | 
(suppress_tracing := false;  | 
| 4270 | 241  | 
Seq.make (fn()=> seqf st  | 
| 15531 | 242  | 
handle TRACE_EXIT st' => SOME(st', Seq.empty)));  | 
| 0 | 243  | 
|
244  | 
||
| 
8149
 
941afb897532
added tacticals DETERM_UNTIL and DETERM_UNTIL_SOLVED
 
oheimb 
parents: 
7686 
diff
changeset
 | 
245  | 
(*Deterministic DO..UNTIL: only retains the first outcome; tail recursive.  | 
| 
 
941afb897532
added tacticals DETERM_UNTIL and DETERM_UNTIL_SOLVED
 
oheimb 
parents: 
7686 
diff
changeset
 | 
246  | 
Forces repitition until predicate on state is fulfilled.*)  | 
| 13108 | 247  | 
fun DETERM_UNTIL p tac =  | 
| 
8149
 
941afb897532
added tacticals DETERM_UNTIL and DETERM_UNTIL_SOLVED
 
oheimb 
parents: 
7686 
diff
changeset
 | 
248  | 
let val tac = tracify trace_REPEAT tac  | 
| 15531 | 249  | 
fun drep st = if p st then SOME (st, Seq.empty)  | 
| 
8149
 
941afb897532
added tacticals DETERM_UNTIL and DETERM_UNTIL_SOLVED
 
oheimb 
parents: 
7686 
diff
changeset
 | 
250  | 
else (case Seq.pull(tac st) of  | 
| 15531 | 251  | 
NONE => NONE  | 
252  | 
| SOME(st',_) => drep st')  | 
|
| 
8149
 
941afb897532
added tacticals DETERM_UNTIL and DETERM_UNTIL_SOLVED
 
oheimb 
parents: 
7686 
diff
changeset
 | 
253  | 
in traced_tac drep end;  | 
| 
 
941afb897532
added tacticals DETERM_UNTIL and DETERM_UNTIL_SOLVED
 
oheimb 
parents: 
7686 
diff
changeset
 | 
254  | 
|
| 13108 | 255  | 
(*Deterministic REPEAT: only retains the first outcome;  | 
| 
703
 
3a5cd2883581
Pure/tctical/REPEAT_DETERM_N,REPEAT_DETERM1,REPEAT_DETERM_FIRST,
 
lcp 
parents: 
671 
diff
changeset
 | 
256  | 
uses less space than REPEAT; tail recursive.  | 
| 
 
3a5cd2883581
Pure/tctical/REPEAT_DETERM_N,REPEAT_DETERM1,REPEAT_DETERM_FIRST,
 
lcp 
parents: 
671 
diff
changeset
 | 
257  | 
If non-negative, n bounds the number of repetitions.*)  | 
| 13108 | 258  | 
fun REPEAT_DETERM_N n tac =  | 
| 1502 | 259  | 
let val tac = tracify trace_REPEAT tac  | 
| 15531 | 260  | 
fun drep 0 st = SOME(st, Seq.empty)  | 
| 
2244
 
dacee519738a
Converted I/O operatios for Basis Library compatibility
 
paulson 
parents: 
2158 
diff
changeset
 | 
261  | 
| drep n st =  | 
| 4270 | 262  | 
(case Seq.pull(tac st) of  | 
| 15531 | 263  | 
NONE => SOME(st, Seq.empty)  | 
264  | 
| SOME(st',_) => drep (n-1) st')  | 
|
| 
703
 
3a5cd2883581
Pure/tctical/REPEAT_DETERM_N,REPEAT_DETERM1,REPEAT_DETERM_FIRST,
 
lcp 
parents: 
671 
diff
changeset
 | 
265  | 
in traced_tac (drep n) end;  | 
| 
 
3a5cd2883581
Pure/tctical/REPEAT_DETERM_N,REPEAT_DETERM1,REPEAT_DETERM_FIRST,
 
lcp 
parents: 
671 
diff
changeset
 | 
266  | 
|
| 
 
3a5cd2883581
Pure/tctical/REPEAT_DETERM_N,REPEAT_DETERM1,REPEAT_DETERM_FIRST,
 
lcp 
parents: 
671 
diff
changeset
 | 
267  | 
(*Allows any number of repetitions*)  | 
| 
 
3a5cd2883581
Pure/tctical/REPEAT_DETERM_N,REPEAT_DETERM1,REPEAT_DETERM_FIRST,
 
lcp 
parents: 
671 
diff
changeset
 | 
268  | 
val REPEAT_DETERM = REPEAT_DETERM_N ~1;  | 
| 0 | 269  | 
|
270  | 
(*General REPEAT: maintains a stack of alternatives; tail recursive*)  | 
|
| 13108 | 271  | 
fun REPEAT tac =  | 
| 1502 | 272  | 
let val tac = tracify trace_REPEAT tac  | 
| 13108 | 273  | 
fun rep qs st =  | 
| 4270 | 274  | 
case Seq.pull(tac st) of  | 
| 15531 | 275  | 
NONE => SOME(st, Seq.make(fn()=> repq qs))  | 
276  | 
| SOME(st',q) => rep (q::qs) st'  | 
|
277  | 
and repq [] = NONE  | 
|
| 4270 | 278  | 
| repq(q::qs) = case Seq.pull q of  | 
| 15531 | 279  | 
NONE => repq qs  | 
280  | 
| SOME(st,q) => rep (q::qs) st  | 
|
| 0 | 281  | 
in traced_tac (rep []) end;  | 
282  | 
||
283  | 
(*Repeat 1 or more times*)  | 
|
| 
703
 
3a5cd2883581
Pure/tctical/REPEAT_DETERM_N,REPEAT_DETERM1,REPEAT_DETERM_FIRST,
 
lcp 
parents: 
671 
diff
changeset
 | 
284  | 
fun REPEAT_DETERM1 tac = DETERM tac THEN REPEAT_DETERM tac;  | 
| 0 | 285  | 
fun REPEAT1 tac = tac THEN REPEAT tac;  | 
286  | 
||
287  | 
||
288  | 
(** Filtering tacticals **)  | 
|
289  | 
||
| 4270 | 290  | 
fun FILTER pred tac st = Seq.filter pred (tac st);  | 
| 0 | 291  | 
|
| 
13650
 
31bd2a8cdbe2
fixing the cut_tac method to work when there are no instantiations and the
 
paulson 
parents: 
13108 
diff
changeset
 | 
292  | 
(*Accept only next states that change the theorem somehow*)  | 
| 13108 | 293  | 
fun CHANGED tac st =  | 
294  | 
let fun diff st' = not (Thm.eq_thm (st, st'));  | 
|
295  | 
in Seq.filter diff (tac st) end;  | 
|
| 0 | 296  | 
|
| 
13650
 
31bd2a8cdbe2
fixing the cut_tac method to work when there are no instantiations and the
 
paulson 
parents: 
13108 
diff
changeset
 | 
297  | 
(*Accept only next states that change the theorem's prop field  | 
| 
 
31bd2a8cdbe2
fixing the cut_tac method to work when there are no instantiations and the
 
paulson 
parents: 
13108 
diff
changeset
 | 
298  | 
(changes to signature, hyps, etc. don't count)*)  | 
| 13108 | 299  | 
fun CHANGED_PROP tac st =  | 
| 
22360
 
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
 
wenzelm 
parents: 
20664 
diff
changeset
 | 
300  | 
let fun diff st' = not (Thm.eq_thm_prop (st, st'));  | 
| 13108 | 301  | 
in Seq.filter diff (tac st) end;  | 
| 10821 | 302  | 
|
| 0 | 303  | 
|
304  | 
(*** Tacticals based on subgoal numbering ***)  | 
|
305  | 
||
| 13108 | 306  | 
(*For n subgoals, performs tac(n) THEN ... THEN tac(1)  | 
| 1502 | 307  | 
Essential to work backwards since tac(i) may add/delete subgoals at i. *)  | 
| 13108 | 308  | 
fun ALLGOALS tac st =  | 
| 1502 | 309  | 
let fun doall 0 = all_tac  | 
| 
2244
 
dacee519738a
Converted I/O operatios for Basis Library compatibility
 
paulson 
parents: 
2158 
diff
changeset
 | 
310  | 
| doall n = tac(n) THEN doall(n-1)  | 
| 1502 | 311  | 
in doall(nprems_of st)st end;  | 
| 0 | 312  | 
|
| 1502 | 313  | 
(*For n subgoals, performs tac(n) ORELSE ... ORELSE tac(1) *)  | 
| 13108 | 314  | 
fun SOMEGOAL tac st =  | 
| 1502 | 315  | 
let fun find 0 = no_tac  | 
| 
2244
 
dacee519738a
Converted I/O operatios for Basis Library compatibility
 
paulson 
parents: 
2158 
diff
changeset
 | 
316  | 
| find n = tac(n) ORELSE find(n-1)  | 
| 1502 | 317  | 
in find(nprems_of st)st end;  | 
| 0 | 318  | 
|
| 1502 | 319  | 
(*For n subgoals, performs tac(1) ORELSE ... ORELSE tac(n).  | 
| 0 | 320  | 
More appropriate than SOMEGOAL in some cases.*)  | 
| 13108 | 321  | 
fun FIRSTGOAL tac st =  | 
| 1502 | 322  | 
let fun find (i,n) = if i>n then no_tac else tac(i) ORELSE find (i+1,n)  | 
323  | 
in find(1, nprems_of st)st end;  | 
|
| 0 | 324  | 
|
| 1502 | 325  | 
(*Repeatedly solve some using tac. *)  | 
326  | 
fun REPEAT_SOME tac = REPEAT1 (SOMEGOAL (REPEAT1 o tac));  | 
|
327  | 
fun REPEAT_DETERM_SOME tac = REPEAT_DETERM1 (SOMEGOAL (REPEAT_DETERM1 o tac));  | 
|
| 0 | 328  | 
|
| 1502 | 329  | 
(*Repeatedly solve the first possible subgoal using tac. *)  | 
330  | 
fun REPEAT_FIRST tac = REPEAT1 (FIRSTGOAL (REPEAT1 o tac));  | 
|
331  | 
fun REPEAT_DETERM_FIRST tac = REPEAT_DETERM1 (FIRSTGOAL (REPEAT_DETERM1 o tac));  | 
|
| 0 | 332  | 
|
| 1502 | 333  | 
(*For n subgoals, tries to apply tac to n,...1 *)  | 
334  | 
fun TRYALL tac = ALLGOALS (TRY o tac);  | 
|
| 0 | 335  | 
|
336  | 
||
337  | 
(*Make a tactic for subgoal i, if there is one. *)  | 
|
| 23224 | 338  | 
fun CSUBGOAL goalfun i st =  | 
339  | 
(case SOME (Thm.cprem_of st i) handle THM _ => NONE of  | 
|
| 
16510
 
606d919ad3c3
tuned SUBGOAL: Logic.nth_prem instead of List.nth o prems_of;
 
wenzelm 
parents: 
16179 
diff
changeset
 | 
340  | 
SOME goal => goalfun (goal, i) st  | 
| 
 
606d919ad3c3
tuned SUBGOAL: Logic.nth_prem instead of List.nth o prems_of;
 
wenzelm 
parents: 
16179 
diff
changeset
 | 
341  | 
| NONE => Seq.empty);  | 
| 0 | 342  | 
|
| 23224 | 343  | 
fun SUBGOAL goalfun =  | 
344  | 
CSUBGOAL (fn (goal, i) => goalfun (Thm.term_of goal, i));  | 
|
345  | 
||
| 5141 | 346  | 
(*Returns all states that have changed in subgoal i, counted from the LAST  | 
347  | 
subgoal. For stac, for example.*)  | 
|
| 13108 | 348  | 
fun CHANGED_GOAL tac i st =  | 
| 
30145
 
09817540ccae
tuned CHANGED_GOAL: use Thm.cprem_of instead of selecting from Thm.prems_of;
 
wenzelm 
parents: 
29272 
diff
changeset
 | 
349  | 
let val np = Thm.nprems_of st  | 
| 7686 | 350  | 
val d = np-i (*distance from END*)  | 
| 
30145
 
09817540ccae
tuned CHANGED_GOAL: use Thm.cprem_of instead of selecting from Thm.prems_of;
 
wenzelm 
parents: 
29272 
diff
changeset
 | 
351  | 
val t = Thm.term_of (Thm.cprem_of st i)  | 
| 13108 | 352  | 
fun diff st' =  | 
| 
30145
 
09817540ccae
tuned CHANGED_GOAL: use Thm.cprem_of instead of selecting from Thm.prems_of;
 
wenzelm 
parents: 
29272 
diff
changeset
 | 
353  | 
Thm.nprems_of st' - d <= 0 (*the subgoal no longer exists*)  | 
| 13108 | 354  | 
orelse  | 
| 
30145
 
09817540ccae
tuned CHANGED_GOAL: use Thm.cprem_of instead of selecting from Thm.prems_of;
 
wenzelm 
parents: 
29272 
diff
changeset
 | 
355  | 
not (Pattern.aeconv (t, Thm.term_of (Thm.cprem_of st' (Thm.nprems_of st' - d))))  | 
| 5141 | 356  | 
in Seq.filter diff (tac i st) end  | 
| 43278 | 357  | 
handle General.Subscript => Seq.empty (*no subgoal i*);  | 
| 5141 | 358  | 
|
| 
34885
 
6587c24ef6d8
added SOLVED' -- a more direct version of THEN_ALL_NEW (K no_tac) -- strictly speaking it does not even depend on subgoal addressing, but it would be too confusing without it;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
359  | 
(*Returns all states where some subgoals have been solved. For  | 
| 
 
6587c24ef6d8
added SOLVED' -- a more direct version of THEN_ALL_NEW (K no_tac) -- strictly speaking it does not even depend on subgoal addressing, but it would be too confusing without it;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
360  | 
subgoal-based tactics this means subgoal i has been solved  | 
| 
 
6587c24ef6d8
added SOLVED' -- a more direct version of THEN_ALL_NEW (K no_tac) -- strictly speaking it does not even depend on subgoal addressing, but it would be too confusing without it;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
361  | 
altogether -- no new subgoals have emerged.*)  | 
| 
 
6587c24ef6d8
added SOLVED' -- a more direct version of THEN_ALL_NEW (K no_tac) -- strictly speaking it does not even depend on subgoal addressing, but it would be too confusing without it;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
362  | 
fun SOLVED' tac i st =  | 
| 
 
6587c24ef6d8
added SOLVED' -- a more direct version of THEN_ALL_NEW (K no_tac) -- strictly speaking it does not even depend on subgoal addressing, but it would be too confusing without it;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
363  | 
tac i st |> Seq.filter (fn st' => nprems_of st' < nprems_of st);  | 
| 
 
6587c24ef6d8
added SOLVED' -- a more direct version of THEN_ALL_NEW (K no_tac) -- strictly speaking it does not even depend on subgoal addressing, but it would be too confusing without it;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
364  | 
|
| 
 
6587c24ef6d8
added SOLVED' -- a more direct version of THEN_ALL_NEW (K no_tac) -- strictly speaking it does not even depend on subgoal addressing, but it would be too confusing without it;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
365  | 
(*Apply second tactic to all subgoals emerging from the first --  | 
| 
 
6587c24ef6d8
added SOLVED' -- a more direct version of THEN_ALL_NEW (K no_tac) -- strictly speaking it does not even depend on subgoal addressing, but it would be too confusing without it;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
366  | 
following usual convention for subgoal-based tactics.*)  | 
| 4602 | 367  | 
fun (tac1 THEN_ALL_NEW tac2) i st =  | 
| 8535 | 368  | 
st |> (tac1 i THEN (fn st' => Seq.INTERVAL tac2 i (i + nprems_of st' - nprems_of st) st'));  | 
| 4602 | 369  | 
|
| 
34885
 
6587c24ef6d8
added SOLVED' -- a more direct version of THEN_ALL_NEW (K no_tac) -- strictly speaking it does not even depend on subgoal addressing, but it would be too confusing without it;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
370  | 
(*Repeatedly dig into any emerging subgoals.*)  | 
| 8341 | 371  | 
fun REPEAT_ALL_NEW tac =  | 
372  | 
tac THEN_ALL_NEW (TRY o (fn i => REPEAT_ALL_NEW tac i));  | 
|
373  | 
||
| 
15006
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
13664 
diff
changeset
 | 
374  | 
(*Makes a tactic whose effect on a state is given by thmfun: thm->thm seq.*)  | 
| 
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
13664 
diff
changeset
 | 
375  | 
fun PRIMSEQ thmfun st = thmfun st handle THM _ => Seq.empty;  | 
| 
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
13664 
diff
changeset
 | 
376  | 
|
| 
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
13664 
diff
changeset
 | 
377  | 
(*Makes a tactic whose effect on a state is given by thmfun: thm->thm.*)  | 
| 
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
13664 
diff
changeset
 | 
378  | 
fun PRIMITIVE thmfun = PRIMSEQ (Seq.single o thmfun);  | 
| 
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
13664 
diff
changeset
 | 
379  | 
|
| 23538 | 380  | 
(*Inverse (more or less) of PRIMITIVE*)  | 
| 15570 | 381  | 
fun SINGLE tacf = Option.map fst o Seq.pull o tacf  | 
| 19455 | 382  | 
|
| 23538 | 383  | 
(*Conversions as tactics*)  | 
| 23584 | 384  | 
fun CONVERSION cv i st = Seq.single (Conv.gconv_rule cv i st)  | 
| 23561 | 385  | 
handle THM _ => Seq.empty  | 
386  | 
| CTERM _ => Seq.empty  | 
|
387  | 
| TERM _ => Seq.empty  | 
|
388  | 
| TYPE _ => Seq.empty;  | 
|
| 23538 | 389  | 
|
| 0 | 390  | 
end;  | 
| 1502 | 391  | 
|
392  | 
open Tactical;  |