author | wenzelm |
Thu, 28 Aug 2008 00:33:04 +0200 | |
changeset 28030 | 8b197e2bc66a |
parent 27813 | 96fbe385a0d0 |
child 29004 | a5a91f387791 |
permissions | -rw-r--r-- |
5824 | 1 |
(* Title: Pure/Isar/method.ML |
2 |
ID: $Id$ |
|
3 |
Author: Markus Wenzel, TU Muenchen |
|
4 |
||
17110 | 5 |
Isar proof methods. |
5824 | 6 |
*) |
7 |
||
8 |
signature BASIC_METHOD = |
|
9 |
sig |
|
17110 | 10 |
val FINDGOAL: (int -> thm -> 'a Seq.seq) -> thm -> 'a Seq.seq |
11 |
val HEADGOAL: (int -> thm -> 'a Seq.seq) -> thm -> 'a Seq.seq |
|
12 |
type method |
|
11731 | 13 |
val trace_rules: bool ref |
5824 | 14 |
end; |
15 |
||
16 |
signature METHOD = |
|
17 |
sig |
|
18 |
include BASIC_METHOD |
|
25957
2cfb703fa8d8
improved apply: handle thread position, apply to context here;
wenzelm
parents:
25699
diff
changeset
|
19 |
val apply: Position.T -> (Proof.context -> method) -> Proof.context -> thm list -> cases_tactic |
18227 | 20 |
val RAW_METHOD_CASES: (thm list -> cases_tactic) -> method |
17110 | 21 |
val RAW_METHOD: (thm list -> tactic) -> method |
18227 | 22 |
val METHOD_CASES: (thm list -> cases_tactic) -> method |
17110 | 23 |
val METHOD: (thm list -> tactic) -> method |
24 |
val fail: method |
|
25 |
val succeed: method |
|
26 |
val insert_tac: thm list -> int -> tactic |
|
27 |
val insert: thm list -> method |
|
28 |
val insert_facts: method |
|
29 |
val SIMPLE_METHOD: tactic -> method |
|
21592 | 30 |
val SIMPLE_METHOD': (int -> tactic) -> method |
31 |
val SIMPLE_METHOD'': ((int -> tactic) -> tactic) -> (int -> tactic) -> method |
|
17110 | 32 |
val defer: int option -> method |
33 |
val prefer: int -> method |
|
20289 | 34 |
val cheating: bool -> Proof.context -> method |
17110 | 35 |
val intro: thm list -> method |
36 |
val elim: thm list -> method |
|
20289 | 37 |
val unfold: thm list -> Proof.context -> method |
38 |
val fold: thm list -> Proof.context -> method |
|
17110 | 39 |
val atomize: bool -> method |
40 |
val this: method |
|
20289 | 41 |
val fact: thm list -> Proof.context -> method |
42 |
val assumption: Proof.context -> method |
|
43 |
val close: bool -> Proof.context -> method |
|
44 |
val trace: Proof.context -> thm list -> unit |
|
6091 | 45 |
val rule_tac: thm list -> thm list -> int -> tactic |
20289 | 46 |
val some_rule_tac: thm list -> Proof.context -> thm list -> int -> tactic |
25270 | 47 |
val intros_tac: thm list -> thm list -> tactic |
17110 | 48 |
val rule: thm list -> method |
49 |
val erule: int -> thm list -> method |
|
50 |
val drule: int -> thm list -> method |
|
51 |
val frule: int -> thm list -> method |
|
20289 | 52 |
val iprover_tac: Proof.context -> int option -> int -> tactic |
27235 | 53 |
val set_tactic: (thm list -> tactic) -> Proof.context -> Proof.context |
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26291
diff
changeset
|
54 |
val tactic: string * Position.T -> Proof.context -> method |
27235 | 55 |
val raw_tactic: string * Position.T -> Proof.context -> method |
27729 | 56 |
type src = Args.src |
5824 | 57 |
datatype text = |
23349 | 58 |
Basic of (Proof.context -> method) * Position.T | |
15703 | 59 |
Source of src | |
20030 | 60 |
Source_i of src | |
5824 | 61 |
Then of text list | |
62 |
Orelse of text list | |
|
63 |
Try of text | |
|
19186 | 64 |
Repeat1 of text | |
65 |
SelectGoals of int * text |
|
17857 | 66 |
val primitive_text: (thm -> thm) -> text |
67 |
val succeed_text: text |
|
17110 | 68 |
val default_text: text |
69 |
val this_text: text |
|
70 |
val done_text: text |
|
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
71 |
val sorry_text: bool -> text |
23349 | 72 |
val finish_text: text option * bool -> Position.T -> text |
27729 | 73 |
val print_methods: theory -> unit |
26892 | 74 |
val intern: theory -> xstring -> string |
75 |
val defined: theory -> string -> bool |
|
20289 | 76 |
val method: theory -> src -> Proof.context -> method |
77 |
val method_i: theory -> src -> Proof.context -> method |
|
78 |
val add_methods: (bstring * (src -> Proof.context -> method) * string) list |
|
17110 | 79 |
-> theory -> theory |
20289 | 80 |
val add_method: bstring * (src -> Proof.context -> method) * string |
17110 | 81 |
-> theory -> theory |
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26291
diff
changeset
|
82 |
val method_setup: bstring -> string * Position.T -> string -> theory -> theory |
18999 | 83 |
val syntax: (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list)) |
21879 | 84 |
-> src -> Proof.context -> 'a * Proof.context |
17110 | 85 |
val simple_args: (Args.T list -> 'a * Args.T list) |
20289 | 86 |
-> ('a -> Proof.context -> method) -> src -> Proof.context -> method |
87 |
val ctxt_args: (Proof.context -> method) -> src -> Proof.context -> method |
|
88 |
val no_args: method -> src -> Proof.context -> method |
|
17110 | 89 |
type modifier |
18999 | 90 |
val sectioned_args: (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list)) -> |
17110 | 91 |
(Args.T list -> modifier * Args.T list) list -> |
20289 | 92 |
('a -> Proof.context -> 'b) -> src -> Proof.context -> 'b |
17110 | 93 |
val bang_sectioned_args: |
94 |
(Args.T list -> modifier * Args.T list) list -> |
|
20289 | 95 |
(thm list -> Proof.context -> 'a) -> src -> Proof.context -> 'a |
17110 | 96 |
val bang_sectioned_args': |
97 |
(Args.T list -> modifier * Args.T list) list -> |
|
18999 | 98 |
(Context.generic * Args.T list -> 'a * (Context.generic * Args.T list)) -> |
20289 | 99 |
('a -> thm list -> Proof.context -> 'b) -> src -> Proof.context -> 'b |
17110 | 100 |
val only_sectioned_args: |
101 |
(Args.T list -> modifier * Args.T list) list -> |
|
20289 | 102 |
(Proof.context -> 'a) -> src -> Proof.context -> 'a |
103 |
val thms_ctxt_args: (thm list -> Proof.context -> 'a) -> src -> |
|
104 |
Proof.context -> 'a |
|
105 |
val thms_args: (thm list -> 'a) -> src -> Proof.context -> 'a |
|
106 |
val thm_args: (thm -> 'a) -> src -> Proof.context -> 'a |
|
9539 | 107 |
val goal_args: (Args.T list -> 'a * Args.T list) -> ('a -> int -> tactic) |
20289 | 108 |
-> src -> Proof.context -> method |
18999 | 109 |
val goal_args': (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list)) |
20289 | 110 |
-> ('a -> int -> tactic) -> src -> Proof.context -> method |
17110 | 111 |
val goal_args_ctxt: (Args.T list -> 'a * Args.T list) -> |
20289 | 112 |
(Proof.context -> 'a -> int -> tactic) -> src -> Proof.context -> method |
18999 | 113 |
val goal_args_ctxt': (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list)) -> |
20289 | 114 |
(Proof.context -> 'a -> int -> tactic) -> src -> Proof.context -> method |
22118 | 115 |
val parse: OuterLex.token list -> text * OuterLex.token list |
5824 | 116 |
end; |
117 |
||
118 |
structure Method: METHOD = |
|
119 |
struct |
|
120 |
||
17110 | 121 |
(** generic tools **) |
122 |
||
123 |
(* goal addressing *) |
|
124 |
||
125 |
fun FINDGOAL tac st = |
|
126 |
let fun find i n = if i > n then Seq.fail else Seq.APPEND (tac i, find (i + 1) n) |
|
127 |
in find 1 (Thm.nprems_of st) st end; |
|
128 |
||
129 |
fun HEADGOAL tac = tac 1; |
|
130 |
||
131 |
||
5824 | 132 |
|
12324
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
wenzelm
parents:
12311
diff
changeset
|
133 |
(** proof methods **) |
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
wenzelm
parents:
12311
diff
changeset
|
134 |
|
17110 | 135 |
(* datatype method *) |
11731 | 136 |
|
18227 | 137 |
datatype method = Meth of thm list -> cases_tactic; |
11731 | 138 |
|
25957
2cfb703fa8d8
improved apply: handle thread position, apply to context here;
wenzelm
parents:
25699
diff
changeset
|
139 |
fun apply pos meth_fun ctxt facts goal = Position.setmp_thread_data_seq pos |
2cfb703fa8d8
improved apply: handle thread position, apply to context here;
wenzelm
parents:
25699
diff
changeset
|
140 |
(fn () => let val Meth meth = meth_fun ctxt in meth facts goal end) (); |
11731 | 141 |
|
17756 | 142 |
val RAW_METHOD_CASES = Meth; |
11731 | 143 |
|
17110 | 144 |
fun RAW_METHOD tac = RAW_METHOD_CASES (NO_CASES o tac); |
12144 | 145 |
|
17110 | 146 |
fun METHOD_CASES tac = RAW_METHOD_CASES (fn facts => |
21687 | 147 |
Seq.THEN (ALLGOALS Goal.conjunction_tac, tac facts)); |
8372 | 148 |
|
21687 | 149 |
fun METHOD tac = RAW_METHOD (fn facts => ALLGOALS Goal.conjunction_tac THEN tac facts); |
5824 | 150 |
|
151 |
val fail = METHOD (K no_tac); |
|
152 |
val succeed = METHOD (K all_tac); |
|
153 |
||
154 |
||
17110 | 155 |
(* insert facts *) |
7419 | 156 |
|
157 |
local |
|
5824 | 158 |
|
21579 | 159 |
fun cut_rule_tac rule = |
160 |
Tactic.rtac (Drule.forall_intr_vars rule COMP_INCR revcut_rl); |
|
6981 | 161 |
|
7419 | 162 |
in |
5824 | 163 |
|
7419 | 164 |
fun insert_tac [] i = all_tac |
165 |
| insert_tac facts i = EVERY (map (fn th => cut_rule_tac th i) facts); |
|
6981 | 166 |
|
7555 | 167 |
val insert_facts = METHOD (ALLGOALS o insert_tac); |
7664 | 168 |
fun insert thms = METHOD (fn _ => ALLGOALS (insert_tac thms)); |
7419 | 169 |
|
9706 | 170 |
fun SIMPLE_METHOD tac = METHOD (fn facts => ALLGOALS (insert_tac facts) THEN tac); |
21592 | 171 |
fun SIMPLE_METHOD'' quant tac = METHOD (fn facts => quant (insert_tac facts THEN' tac)); |
172 |
val SIMPLE_METHOD' = SIMPLE_METHOD'' HEADGOAL; |
|
9706 | 173 |
|
12324
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
wenzelm
parents:
12311
diff
changeset
|
174 |
end; |
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
wenzelm
parents:
12311
diff
changeset
|
175 |
|
9706 | 176 |
|
17110 | 177 |
(* shuffle subgoals *) |
178 |
||
179 |
fun prefer i = METHOD (K (Tactic.defer_tac i THEN PRIMITIVE (Thm.permute_prems 0 ~1))); |
|
18939 | 180 |
fun defer opt_i = METHOD (K (Tactic.defer_tac (the_default 1 opt_i))); |
17110 | 181 |
|
182 |
||
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
183 |
(* cheating *) |
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
184 |
|
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
185 |
fun cheating int ctxt = METHOD (K (setmp quick_and_dirty (int orelse ! quick_and_dirty) |
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
186 |
(SkipProof.cheat_tac (ProofContext.theory_of ctxt)))); |
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
187 |
|
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
188 |
|
17110 | 189 |
(* unfold intro/elim rules *) |
190 |
||
21592 | 191 |
fun intro ths = SIMPLE_METHOD' (CHANGED_PROP o REPEAT_ALL_NEW (Tactic.match_tac ths)); |
192 |
fun elim ths = SIMPLE_METHOD' (CHANGED_PROP o REPEAT_ALL_NEW (Tactic.ematch_tac ths)); |
|
17110 | 193 |
|
194 |
||
12384 | 195 |
(* unfold/fold definitions *) |
196 |
||
18877 | 197 |
fun unfold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (LocalDefs.unfold_tac ctxt ths)); |
198 |
fun fold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (LocalDefs.fold_tac ctxt ths)); |
|
6532 | 199 |
|
12384 | 200 |
|
12829 | 201 |
(* atomize rule statements *) |
202 |
||
23590
ad95084a5c63
renamed ObjectLogic.atomize_tac to ObjectLogic.atomize_prems_tac;
wenzelm
parents:
23577
diff
changeset
|
203 |
fun atomize false = SIMPLE_METHOD' (CHANGED_PROP o ObjectLogic.atomize_prems_tac) |
12829 | 204 |
| atomize true = RAW_METHOD (K (HEADGOAL (CHANGED_PROP o ObjectLogic.full_atomize_tac))); |
205 |
||
206 |
||
18039 | 207 |
(* this -- resolve facts directly *) |
12384 | 208 |
|
17110 | 209 |
val this = METHOD (EVERY o map (HEADGOAL o Tactic.rtac)); |
9484 | 210 |
|
211 |
||
18039 | 212 |
(* fact -- composition by facts from context *) |
213 |
||
21592 | 214 |
fun fact [] ctxt = SIMPLE_METHOD' (ProofContext.some_fact_tac ctxt) |
215 |
| fact rules _ = SIMPLE_METHOD' (ProofContext.fact_tac rules); |
|
18039 | 216 |
|
217 |
||
17110 | 218 |
(* assumption *) |
7419 | 219 |
|
220 |
local |
|
221 |
||
19778 | 222 |
fun cond_rtac cond rule = SUBGOAL (fn (prop, i) => |
223 |
if cond (Logic.strip_assums_concl prop) |
|
224 |
then Tactic.rtac rule i else no_tac); |
|
7419 | 225 |
|
25957
2cfb703fa8d8
improved apply: handle thread position, apply to context here;
wenzelm
parents:
25699
diff
changeset
|
226 |
fun legacy_tac st = |
2cfb703fa8d8
improved apply: handle thread position, apply to context here;
wenzelm
parents:
25699
diff
changeset
|
227 |
(legacy_feature |
26291 | 228 |
("Implicit use of prems in assumption proof" ^ Position.str_of (Position.thread_data ())); |
23395
15fb6637690e
method assumption: uniform treatment of prems as legacy feature;
wenzelm
parents:
23349
diff
changeset
|
229 |
all_tac st); |
15fb6637690e
method assumption: uniform treatment of prems as legacy feature;
wenzelm
parents:
23349
diff
changeset
|
230 |
|
15fb6637690e
method assumption: uniform treatment of prems as legacy feature;
wenzelm
parents:
23349
diff
changeset
|
231 |
fun assm_tac ctxt = |
17110 | 232 |
assume_tac APPEND' |
23349 | 233 |
Goal.assume_rule_tac ctxt APPEND' |
25957
2cfb703fa8d8
improved apply: handle thread position, apply to context here;
wenzelm
parents:
25699
diff
changeset
|
234 |
(CHANGED o solve_tac (Assumption.prems_of ctxt) THEN' K legacy_tac) APPEND' |
19778 | 235 |
cond_rtac (can Logic.dest_equals) Drule.reflexive_thm APPEND' |
236 |
cond_rtac (can Logic.dest_term) Drule.termI; |
|
17110 | 237 |
|
7419 | 238 |
in |
239 |
||
23349 | 240 |
fun assumption ctxt = METHOD (HEADGOAL o |
23395
15fb6637690e
method assumption: uniform treatment of prems as legacy feature;
wenzelm
parents:
23349
diff
changeset
|
241 |
(fn [] => assm_tac ctxt |
23349 | 242 |
| [fact] => solve_tac [fact] |
243 |
| _ => K no_tac)); |
|
244 |
||
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
245 |
fun close immed ctxt = METHOD (K |
23349 | 246 |
(FILTER Thm.no_prems |
23395
15fb6637690e
method assumption: uniform treatment of prems as legacy feature;
wenzelm
parents:
23349
diff
changeset
|
247 |
((if immed then ALLGOALS (assm_tac ctxt) else all_tac) THEN flexflex_tac))); |
7419 | 248 |
|
249 |
end; |
|
250 |
||
251 |
||
17110 | 252 |
(* rule etc. -- single-step refinements *) |
12347 | 253 |
|
17110 | 254 |
val trace_rules = ref false; |
12347 | 255 |
|
17110 | 256 |
fun trace ctxt rules = |
21962 | 257 |
if ! trace_rules andalso not (null rules) then |
17110 | 258 |
Pretty.big_list "rules:" (map (ProofContext.pretty_thm ctxt) rules) |
21962 | 259 |
|> Pretty.string_of |> tracing |
260 |
else (); |
|
12347 | 261 |
|
262 |
local |
|
263 |
||
18841 | 264 |
fun gen_rule_tac tac rules facts = |
265 |
(fn i => fn st => |
|
266 |
if null facts then tac rules i st |
|
267 |
else Seq.maps (fn rule => (tac o single) rule i st) (Drule.multi_resolves facts rules)) |
|
21687 | 268 |
THEN_ALL_NEW Goal.norm_hhf_tac; |
7130 | 269 |
|
10744 | 270 |
fun gen_arule_tac tac j rules facts = |
271 |
EVERY' (gen_rule_tac tac rules facts :: replicate j Tactic.assume_tac); |
|
272 |
||
11785 | 273 |
fun gen_some_rule_tac tac arg_rules ctxt facts = SUBGOAL (fn (goal, i) => |
274 |
let |
|
275 |
val rules = |
|
276 |
if not (null arg_rules) then arg_rules |
|
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19307
diff
changeset
|
277 |
else flat (ContextRules.find_rules false facts goal ctxt) |
12055 | 278 |
in trace ctxt rules; tac rules facts i end); |
10309 | 279 |
|
10744 | 280 |
fun meth tac x = METHOD (HEADGOAL o tac x); |
281 |
fun meth' tac x y = METHOD (HEADGOAL o tac x y); |
|
8220 | 282 |
|
7419 | 283 |
in |
284 |
||
10744 | 285 |
val rule_tac = gen_rule_tac Tactic.resolve_tac; |
286 |
val rule = meth rule_tac; |
|
287 |
val some_rule_tac = gen_some_rule_tac rule_tac; |
|
288 |
val some_rule = meth' some_rule_tac; |
|
289 |
||
290 |
val erule = meth' (gen_arule_tac Tactic.eresolve_tac); |
|
291 |
val drule = meth' (gen_arule_tac Tactic.dresolve_tac); |
|
292 |
val frule = meth' (gen_arule_tac Tactic.forward_tac); |
|
5824 | 293 |
|
7419 | 294 |
end; |
295 |
||
296 |
||
25270 | 297 |
(* intros_tac -- pervasive search spanned by intro rules *) |
298 |
||
299 |
fun intros_tac intros facts = |
|
300 |
ALLGOALS (insert_tac facts THEN' |
|
301 |
REPEAT_ALL_NEW (resolve_tac intros)) |
|
302 |
THEN Tactic.distinct_subgoals_tac; |
|
303 |
||
304 |
||
17587 | 305 |
(* iprover -- intuitionistic proof search *) |
17110 | 306 |
|
307 |
local |
|
308 |
||
309 |
val remdups_tac = SUBGOAL (fn (g, i) => |
|
310 |
let val prems = Logic.strip_assums_hyp g in |
|
19046
bc5c6c9b114e
removed distinct, renamed gen_distinct to distinct;
wenzelm
parents:
18999
diff
changeset
|
311 |
REPEAT_DETERM_N (length prems - length (distinct op aconv prems)) |
17110 | 312 |
(Tactic.ematch_tac [Drule.remdups_rl] i THEN Tactic.eq_assume_tac i) |
313 |
end); |
|
314 |
||
315 |
fun REMDUPS tac = tac THEN_ALL_NEW remdups_tac; |
|
316 |
||
317 |
val bires_tac = Tactic.biresolution_from_nets_tac ContextRules.orderlist; |
|
318 |
||
319 |
fun safe_step_tac ctxt = |
|
320 |
ContextRules.Swrap ctxt |
|
321 |
(eq_assume_tac ORELSE' |
|
322 |
bires_tac true (ContextRules.netpair_bang ctxt)); |
|
8195 | 323 |
|
17110 | 324 |
fun unsafe_step_tac ctxt = |
325 |
ContextRules.wrap ctxt |
|
326 |
(assume_tac APPEND' |
|
327 |
bires_tac false (ContextRules.netpair_bang ctxt) APPEND' |
|
328 |
bires_tac false (ContextRules.netpair ctxt)); |
|
329 |
||
330 |
fun step_tac ctxt i = |
|
331 |
REPEAT_DETERM1 (REMDUPS (safe_step_tac ctxt) i) ORELSE |
|
332 |
REMDUPS (unsafe_step_tac ctxt) i; |
|
333 |
||
334 |
fun intprover_tac ctxt gs d lim = SUBGOAL (fn (g, i) => if d > lim then no_tac else |
|
335 |
let |
|
336 |
val ps = Logic.strip_assums_hyp g; |
|
337 |
val c = Logic.strip_assums_concl g; |
|
338 |
in |
|
18921 | 339 |
if member (fn ((ps1, c1), (ps2, c2)) => |
19307 | 340 |
c1 aconv c2 andalso |
341 |
length ps1 = length ps2 andalso |
|
342 |
gen_eq_set (op aconv) (ps1, ps2)) gs (ps, c) then no_tac |
|
17110 | 343 |
else (step_tac ctxt THEN_ALL_NEW intprover_tac ctxt ((ps, c) :: gs) (d + 1) lim) i |
344 |
end); |
|
345 |
||
346 |
in |
|
347 |
||
17587 | 348 |
fun iprover_tac ctxt opt_lim = |
18939 | 349 |
SELECT_GOAL (DEEPEN (2, the_default 20 opt_lim) (intprover_tac ctxt [] 0) 4 1); |
17110 | 350 |
|
351 |
end; |
|
8195 | 352 |
|
353 |
||
8351 | 354 |
(* ML tactics *) |
355 |
||
26472
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
356 |
structure TacticData = ProofDataFun |
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
357 |
( |
27235 | 358 |
type T = thm list -> tactic; |
26472
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
359 |
fun init _ = undefined; |
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
360 |
); |
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
361 |
|
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
362 |
val set_tactic = TacticData.put; |
8351 | 363 |
|
26472
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
364 |
fun ml_tactic (txt, pos) ctxt = |
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
365 |
let |
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
366 |
val ctxt' = ctxt |> Context.proof_map |
27235 | 367 |
(ML_Context.expression pos |
368 |
"fun tactic (facts: thm list) : tactic" |
|
26472
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
369 |
"Context.map_proof (Method.set_tactic tactic)" txt); |
27235 | 370 |
in Context.setmp_thread_data (SOME (Context.Proof ctxt)) (TacticData.get ctxt') end; |
23425 | 371 |
|
372 |
fun tactic txt ctxt = METHOD (ml_tactic txt ctxt); |
|
373 |
fun raw_tactic txt ctxt = RAW_METHOD (ml_tactic txt ctxt); |
|
8351 | 374 |
|
375 |
||
5824 | 376 |
|
17110 | 377 |
(** method syntax **) |
378 |
||
379 |
(* method text *) |
|
380 |
||
381 |
type src = Args.src; |
|
5824 | 382 |
|
17110 | 383 |
datatype text = |
23349 | 384 |
Basic of (Proof.context -> method) * Position.T | |
17110 | 385 |
Source of src | |
20030 | 386 |
Source_i of src | |
17110 | 387 |
Then of text list | |
388 |
Orelse of text list | |
|
389 |
Try of text | |
|
19186 | 390 |
Repeat1 of text | |
391 |
SelectGoals of int * text; |
|
17110 | 392 |
|
23349 | 393 |
fun primitive_text r = Basic (K (SIMPLE_METHOD (PRIMITIVE r)), Position.none); |
394 |
val succeed_text = Basic (K succeed, Position.none); |
|
17110 | 395 |
val default_text = Source (Args.src (("default", []), Position.none)); |
23349 | 396 |
val this_text = Basic (K this, Position.none); |
397 |
val done_text = Basic (K (SIMPLE_METHOD all_tac), Position.none); |
|
398 |
fun sorry_text int = Basic (cheating int, Position.none); |
|
17110 | 399 |
|
23349 | 400 |
fun finish_text (NONE, immed) pos = Basic (close immed, pos) |
401 |
| finish_text (SOME txt, immed) pos = Then [txt, Basic (close immed, pos)]; |
|
17110 | 402 |
|
403 |
||
404 |
(* method definitions *) |
|
5824 | 405 |
|
24116 | 406 |
structure Methods = TheoryDataFun |
22846 | 407 |
( |
20289 | 408 |
type T = (((src -> Proof.context -> method) * string) * stamp) NameSpace.table; |
16347 | 409 |
val empty = NameSpace.empty_table; |
6546 | 410 |
val copy = I; |
16448 | 411 |
val extend = I; |
23655
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents:
23590
diff
changeset
|
412 |
fun merge _ tables : T = NameSpace.merge_tables (eq_snd (op =)) tables handle Symtab.DUP dup => |
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents:
23590
diff
changeset
|
413 |
error ("Attempt to merge different versions of method " ^ quote dup); |
22846 | 414 |
); |
5824 | 415 |
|
22846 | 416 |
fun print_methods thy = |
417 |
let |
|
24116 | 418 |
val meths = Methods.get thy; |
22846 | 419 |
fun prt_meth (name, ((_, comment), _)) = Pretty.block |
420 |
[Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment]; |
|
421 |
in |
|
422 |
[Pretty.big_list "methods:" (map prt_meth (NameSpace.extern_table meths))] |
|
423 |
|> Pretty.chunks |> Pretty.writeln |
|
424 |
end; |
|
7611 | 425 |
|
5824 | 426 |
|
427 |
(* get methods *) |
|
428 |
||
26892 | 429 |
val intern = NameSpace.intern o #1 o Methods.get; |
430 |
val defined = Symtab.defined o #2 o Methods.get; |
|
431 |
||
20030 | 432 |
fun method_i thy = |
5824 | 433 |
let |
24116 | 434 |
val meths = #2 (Methods.get thy); |
5884 | 435 |
fun meth src = |
20030 | 436 |
let val ((name, _), pos) = Args.dest_src src in |
17412 | 437 |
(case Symtab.lookup meths name of |
15531 | 438 |
NONE => error ("Unknown proof method: " ^ quote name ^ Position.str_of pos) |
27751 | 439 |
| SOME ((mth, _), _) => (Position.report (Markup.method name) pos; mth src)) |
5824 | 440 |
end; |
441 |
in meth end; |
|
442 |
||
24116 | 443 |
fun method thy = method_i thy o Args.map_name (NameSpace.intern (#1 (Methods.get thy))); |
20030 | 444 |
|
5824 | 445 |
|
17110 | 446 |
(* add method *) |
5824 | 447 |
|
448 |
fun add_methods raw_meths thy = |
|
449 |
let |
|
16145 | 450 |
val new_meths = raw_meths |> map (fn (name, f, comment) => |
16347 | 451 |
(name, ((f, comment), stamp ()))); |
5824 | 452 |
|
23086 | 453 |
fun add meths = NameSpace.extend_table (Sign.naming_of thy) new_meths meths |
23655
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents:
23590
diff
changeset
|
454 |
handle Symtab.DUP dup => error ("Duplicate declaration of method " ^ quote dup); |
24116 | 455 |
in Methods.map add thy end; |
5824 | 456 |
|
9194 | 457 |
val add_method = add_methods o Library.single; |
458 |
||
5824 | 459 |
|
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
460 |
(* method_setup *) |
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
461 |
|
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26291
diff
changeset
|
462 |
fun method_setup name (txt, pos) cmt = |
26472
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
463 |
Context.theory_map (ML_Context.expression pos |
22118 | 464 |
"val method: bstring * (Method.src -> Proof.context -> Proof.method) * string" |
22086 | 465 |
"Context.map_theory (Method.add_method method)" |
26472
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
466 |
("(" ^ quote name ^ ", " ^ txt ^ ", " ^ quote cmt ^ ")")); |
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
467 |
|
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
468 |
|
5884 | 469 |
|
17110 | 470 |
(** concrete syntax **) |
5824 | 471 |
|
27813
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
472 |
structure P = OuterParse; |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
473 |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
474 |
|
5884 | 475 |
(* basic *) |
476 |
||
18999 | 477 |
fun syntax scan = Args.context_syntax "method" scan; |
5824 | 478 |
|
17110 | 479 |
fun simple_args scan f src ctxt : method = |
21879 | 480 |
fst (syntax (Scan.lift (scan >> (fn x => f x ctxt))) src ctxt); |
8351 | 481 |
|
20289 | 482 |
fun ctxt_args (f: Proof.context -> method) src ctxt = |
21879 | 483 |
fst (syntax (Scan.succeed (f ctxt)) src ctxt); |
7555 | 484 |
|
485 |
fun no_args m = ctxt_args (K m); |
|
5884 | 486 |
|
487 |
||
488 |
(* sections *) |
|
5824 | 489 |
|
20289 | 490 |
type modifier = (Proof.context -> Proof.context) * attribute; |
7268 | 491 |
|
492 |
local |
|
493 |
||
24010
2ef318813e1a
method section scanners: added [[declaration]] syntax, ignore sid-effects of thms;
wenzelm
parents:
23937
diff
changeset
|
494 |
fun thms ss = Scan.repeat (Scan.unless (Scan.lift (Scan.first ss)) Attrib.multi_thm) >> flat; |
19186 | 495 |
fun app (f, att) (context, ths) = foldl_map att (Context.map_proof f context, ths); |
5824 | 496 |
|
24022 | 497 |
fun section ss = Scan.depend (fn context => (Scan.first ss -- Scan.pass context (thms ss)) :|-- |
498 |
(fn (m, ths) => Scan.succeed (app m (context, ths)))); |
|
5884 | 499 |
|
7601 | 500 |
fun sectioned args ss = args -- Scan.repeat (section ss); |
5884 | 501 |
|
7268 | 502 |
in |
5824 | 503 |
|
5884 | 504 |
fun sectioned_args args ss f src ctxt = |
21879 | 505 |
let val ((x, _), ctxt') = syntax (sectioned args ss) src ctxt |
5921 | 506 |
in f x ctxt' end; |
5884 | 507 |
|
7601 | 508 |
fun bang_sectioned_args ss f = sectioned_args Args.bang_facts ss f; |
9777 | 509 |
fun bang_sectioned_args' ss scan f = |
510 |
sectioned_args (Args.bang_facts -- scan >> swap) ss (uncurry f); |
|
7601 | 511 |
fun only_sectioned_args ss f = sectioned_args (Scan.succeed ()) ss (fn () => f); |
7268 | 512 |
|
18999 | 513 |
fun thms_ctxt_args f = sectioned_args (thms []) [] f; |
8093 | 514 |
fun thms_args f = thms_ctxt_args (K o f); |
9706 | 515 |
fun thm_args f = thms_args (fn [thm] => f thm | _ => error "Single theorem expected"); |
5824 | 516 |
|
7268 | 517 |
end; |
518 |
||
5824 | 519 |
|
17587 | 520 |
(* iprover syntax *) |
12347 | 521 |
|
522 |
local |
|
523 |
||
524 |
val introN = "intro"; |
|
525 |
val elimN = "elim"; |
|
526 |
val destN = "dest"; |
|
527 |
val ruleN = "rule"; |
|
528 |
||
529 |
fun modifier name kind kind' att = |
|
27813
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
530 |
Args.$$$ name |-- (kind >> K NONE || kind' |-- P.nat --| Args.colon >> SOME) |
20289 | 531 |
>> (pair (I: Proof.context -> Proof.context) o att); |
12347 | 532 |
|
17587 | 533 |
val iprover_modifiers = |
18728 | 534 |
[modifier destN Args.bang_colon Args.bang ContextRules.dest_bang, |
535 |
modifier destN Args.colon (Scan.succeed ()) ContextRules.dest, |
|
536 |
modifier elimN Args.bang_colon Args.bang ContextRules.elim_bang, |
|
537 |
modifier elimN Args.colon (Scan.succeed ()) ContextRules.elim, |
|
538 |
modifier introN Args.bang_colon Args.bang ContextRules.intro_bang, |
|
539 |
modifier introN Args.colon (Scan.succeed ()) ContextRules.intro, |
|
540 |
Args.del -- Args.colon >> K (I, ContextRules.rule_del)]; |
|
12347 | 541 |
|
542 |
in |
|
543 |
||
18640 | 544 |
val iprover_meth = |
27813
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
545 |
bang_sectioned_args' iprover_modifiers (Scan.lift (Scan.option P.nat)) |
18640 | 546 |
(fn n => fn prems => fn ctxt => METHOD (fn facts => |
547 |
HEADGOAL (insert_tac (prems @ facts) THEN' |
|
23590
ad95084a5c63
renamed ObjectLogic.atomize_tac to ObjectLogic.atomize_prems_tac;
wenzelm
parents:
23577
diff
changeset
|
548 |
ObjectLogic.atomize_prems_tac THEN' iprover_tac ctxt n))); |
12347 | 549 |
|
550 |
end; |
|
551 |
||
552 |
||
9539 | 553 |
(* tactic syntax *) |
8238 | 554 |
|
10744 | 555 |
fun nat_thms_args f = uncurry f oo |
27813
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
556 |
(fst oo syntax (Scan.lift (Scan.optional (Args.parens P.nat) 0) -- Attrib.thms)); |
10744 | 557 |
|
21879 | 558 |
fun goal_args' args tac src ctxt = fst (syntax (Args.goal_spec HEADGOAL -- args >> |
21592 | 559 |
(fn (quant, s) => SIMPLE_METHOD'' quant (tac s))) src ctxt); |
8537 | 560 |
|
9539 | 561 |
fun goal_args args tac = goal_args' (Scan.lift args) tac; |
8238 | 562 |
|
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
563 |
fun goal_args_ctxt' args tac src ctxt = |
21879 | 564 |
fst (syntax (Args.goal_spec HEADGOAL -- args >> |
21592 | 565 |
(fn (quant, s) => SIMPLE_METHOD'' quant (tac ctxt s))) src ctxt); |
8238 | 566 |
|
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
567 |
fun goal_args_ctxt args tac = goal_args_ctxt' (Scan.lift args) tac; |
5824 | 568 |
|
14718 | 569 |
|
27813
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
570 |
(* outer parser *) |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
571 |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
572 |
fun is_symid_meth s = |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
573 |
s <> "|" andalso s <> "?" andalso s <> "+" andalso OuterLex.ident_or_symbolic s; |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
574 |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
575 |
local |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
576 |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
577 |
fun meth4 x = |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
578 |
(P.position (P.xname >> rpair []) >> (Source o Args.src) || |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
579 |
P.$$$ "(" |-- P.!!! (meth0 --| P.$$$ ")")) x |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
580 |
and meth3 x = |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
581 |
(meth4 --| P.$$$ "?" >> Try || |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
582 |
meth4 --| P.$$$ "+" >> Repeat1 || |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
583 |
meth4 -- (P.$$$ "[" |-- Scan.optional P.nat 1 --| P.$$$ "]") >> (SelectGoals o swap) || |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
584 |
meth4) x |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
585 |
and meth2 x = |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
586 |
(P.position (P.xname -- Args.parse1 is_symid_meth) >> (Source o Args.src) || |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
587 |
meth3) x |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
588 |
and meth1 x = (P.enum1 "," meth2 >> (fn [m] => m | ms => Then ms)) x |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
589 |
and meth0 x = (P.enum1 "|" meth1 >> (fn [m] => m | ms => Orelse ms)) x; |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
590 |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
591 |
in val parse = meth3 end; |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
592 |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
593 |
|
18708 | 594 |
(* theory setup *) |
5824 | 595 |
|
26463 | 596 |
val _ = Context.>> (Context.map_theory |
597 |
(add_methods |
|
598 |
[("fail", no_args fail, "force failure"), |
|
599 |
("succeed", no_args succeed, "succeed"), |
|
600 |
("-", no_args insert_facts, "do nothing (insert current facts only)"), |
|
601 |
("insert", thms_args insert, "insert theorems, ignoring facts (improper)"), |
|
602 |
("intro", thms_args intro, "repeatedly apply introduction rules"), |
|
603 |
("elim", thms_args elim, "repeatedly apply elimination rules"), |
|
604 |
("unfold", thms_ctxt_args unfold_meth, "unfold definitions"), |
|
605 |
("fold", thms_ctxt_args fold_meth, "fold definitions"), |
|
606 |
("atomize", (atomize o fst) oo syntax (Args.mode "full"), |
|
607 |
"present local premises as object-level statements"), |
|
608 |
("iprover", iprover_meth, "intuitionistic proof search"), |
|
609 |
("rule", thms_ctxt_args some_rule, "apply some intro/elim rule"), |
|
610 |
("erule", nat_thms_args erule, "apply rule in elimination manner (improper)"), |
|
611 |
("drule", nat_thms_args drule, "apply rule in destruct manner (improper)"), |
|
612 |
("frule", nat_thms_args frule, "apply rule in forward manner (improper)"), |
|
613 |
("this", no_args this, "apply current facts as rules"), |
|
614 |
("fact", thms_ctxt_args fact, "composition by facts from context"), |
|
615 |
("assumption", ctxt_args assumption, "proof by assumption, preferring facts"), |
|
27244 | 616 |
("rename_tac", goal_args (Scan.repeat1 Args.name) Tactic.rename_tac, |
26463 | 617 |
"rename parameters of goal"), |
27813
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
618 |
("rotate_tac", goal_args (Scan.optional P.int 1) Tactic.rotate_tac, |
26463 | 619 |
"rotate assumptions of goal"), |
27813
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
620 |
("tactic", simple_args (P.position Args.name) tactic, "ML tactic as proof method"), |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
621 |
("raw_tactic", simple_args (P.position Args.name) raw_tactic, |
26463 | 622 |
"ML tactic as raw proof method")])); |
5824 | 623 |
|
624 |
||
16145 | 625 |
(*final declarations of this structure!*) |
626 |
val unfold = unfold_meth; |
|
627 |
val fold = fold_meth; |
|
628 |
||
5824 | 629 |
end; |
630 |
||
631 |
structure BasicMethod: BASIC_METHOD = Method; |
|
632 |
open BasicMethod; |