| author | wenzelm |
| Mon, 23 Mar 2009 11:20:46 +0100 | |
| changeset 30666 | d6248d4508d5 |
| parent 30575 | 368e26dfba69 |
| child 31238 | 2291e4d628eb |
| permissions | -rw-r--r-- |
| 5824 | 1 |
(* Title: Pure/Isar/method.ML |
2 |
Author: Markus Wenzel, TU Muenchen |
|
3 |
||
| 17110 | 4 |
Isar proof methods. |
| 5824 | 5 |
*) |
6 |
||
7 |
signature BASIC_METHOD = |
|
8 |
sig |
|
| 17110 | 9 |
val FINDGOAL: (int -> thm -> 'a Seq.seq) -> thm -> 'a Seq.seq |
10 |
val HEADGOAL: (int -> thm -> 'a Seq.seq) -> thm -> 'a Seq.seq |
|
| 11731 | 11 |
val trace_rules: bool ref |
| 5824 | 12 |
end; |
13 |
||
14 |
signature METHOD = |
|
15 |
sig |
|
16 |
include BASIC_METHOD |
|
|
30508
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents:
30466
diff
changeset
|
17 |
type method |
|
25957
2cfb703fa8d8
improved apply: handle thread position, apply to context here;
wenzelm
parents:
25699
diff
changeset
|
18 |
val apply: Position.T -> (Proof.context -> method) -> Proof.context -> thm list -> cases_tactic |
| 18227 | 19 |
val RAW_METHOD_CASES: (thm list -> cases_tactic) -> method |
| 17110 | 20 |
val RAW_METHOD: (thm list -> tactic) -> method |
| 18227 | 21 |
val METHOD_CASES: (thm list -> cases_tactic) -> method |
| 17110 | 22 |
val METHOD: (thm list -> tactic) -> method |
23 |
val fail: method |
|
24 |
val succeed: method |
|
25 |
val insert_tac: thm list -> int -> tactic |
|
26 |
val insert: thm list -> method |
|
27 |
val insert_facts: method |
|
28 |
val SIMPLE_METHOD: tactic -> method |
|
| 21592 | 29 |
val SIMPLE_METHOD': (int -> tactic) -> method |
30 |
val SIMPLE_METHOD'': ((int -> tactic) -> tactic) -> (int -> tactic) -> method |
|
| 17110 | 31 |
val defer: int option -> method |
32 |
val prefer: int -> method |
|
| 20289 | 33 |
val cheating: bool -> Proof.context -> method |
| 17110 | 34 |
val intro: thm list -> method |
35 |
val elim: thm list -> method |
|
| 20289 | 36 |
val unfold: thm list -> Proof.context -> method |
37 |
val fold: thm list -> Proof.context -> method |
|
| 17110 | 38 |
val atomize: bool -> method |
39 |
val this: method |
|
| 20289 | 40 |
val fact: thm list -> Proof.context -> method |
|
30234
7dd251bce291
renamed Method.assumption_tac back to Method.assm_tac -- as assumption_tac it would have to be exactly the tactic behind the assumption method (with facts);
wenzelm
parents:
30190
diff
changeset
|
41 |
val assm_tac: Proof.context -> int -> tactic |
|
30567
cd8e20f86795
close/all_assm_tac: finish all subgoals from left to right (cf. Proof.goal_tac a28d83e903ce) -- NB: ALLGOALS/THEN_ALL_NEW operate from right to left;
wenzelm
parents:
30544
diff
changeset
|
42 |
val all_assm_tac: Proof.context -> tactic |
| 20289 | 43 |
val assumption: Proof.context -> method |
44 |
val close: bool -> Proof.context -> method |
|
45 |
val trace: Proof.context -> thm list -> unit |
|
| 6091 | 46 |
val rule_tac: thm list -> thm list -> int -> tactic |
| 20289 | 47 |
val some_rule_tac: thm list -> Proof.context -> thm list -> int -> tactic |
| 25270 | 48 |
val intros_tac: thm list -> thm list -> tactic |
| 17110 | 49 |
val rule: thm list -> method |
50 |
val erule: int -> thm list -> method |
|
51 |
val drule: int -> thm list -> method |
|
52 |
val frule: int -> thm list -> method |
|
| 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 |
| 30512 | 82 |
val syntax: 'a context_parser -> src -> Proof.context -> 'a * Proof.context |
83 |
val setup: binding -> (Proof.context -> method) context_parser -> string -> theory -> theory |
|
| 30575 | 84 |
val method_setup: bstring * Position.T -> Symbol_Pos.text * Position.T -> string -> |
85 |
theory -> theory |
|
| 30512 | 86 |
val simple_args: 'a parser -> ('a -> Proof.context -> method) -> src -> Proof.context -> method
|
| 20289 | 87 |
val ctxt_args: (Proof.context -> method) -> src -> Proof.context -> method |
88 |
val no_args: method -> src -> Proof.context -> method |
|
| 30540 | 89 |
type modifier = (Proof.context -> Proof.context) * attribute |
90 |
val section: modifier parser list -> thm list context_parser |
|
91 |
val sections: modifier parser list -> thm list list context_parser |
|
| 30512 | 92 |
val sectioned_args: 'a context_parser -> modifier parser list -> |
| 20289 | 93 |
('a -> Proof.context -> 'b) -> src -> Proof.context -> 'b
|
| 30512 | 94 |
val bang_sectioned_args: modifier parser list -> |
| 20289 | 95 |
(thm list -> Proof.context -> 'a) -> src -> Proof.context -> 'a |
| 30512 | 96 |
val bang_sectioned_args': modifier parser list -> 'a context_parser -> |
| 20289 | 97 |
('a -> thm list -> Proof.context -> 'b) -> src -> Proof.context -> 'b
|
| 30512 | 98 |
val only_sectioned_args: modifier parser list -> (Proof.context -> 'a) -> src -> |
| 20289 | 99 |
Proof.context -> 'a |
| 30512 | 100 |
val thms_ctxt_args: (thm list -> Proof.context -> 'a) -> src -> Proof.context -> 'a |
| 20289 | 101 |
val thms_args: (thm list -> 'a) -> src -> Proof.context -> 'a |
102 |
val thm_args: (thm -> 'a) -> src -> Proof.context -> 'a |
|
| 30512 | 103 |
val goal_args: 'a parser -> ('a -> int -> tactic) -> src -> Proof.context -> method
|
104 |
val goal_args': 'a context_parser -> ('a -> int -> tactic) -> src -> Proof.context -> method
|
|
105 |
val goal_args_ctxt: 'a parser -> (Proof.context -> 'a -> int -> tactic) -> src -> |
|
106 |
Proof.context -> method |
|
107 |
val goal_args_ctxt': 'a context_parser -> (Proof.context -> 'a -> int -> tactic) -> src -> |
|
108 |
Proof.context -> method |
|
109 |
val parse: text parser |
|
| 5824 | 110 |
end; |
111 |
||
112 |
structure Method: METHOD = |
|
113 |
struct |
|
114 |
||
| 17110 | 115 |
(** generic tools **) |
116 |
||
117 |
(* goal addressing *) |
|
118 |
||
119 |
fun FINDGOAL tac st = |
|
120 |
let fun find i n = if i > n then Seq.fail else Seq.APPEND (tac i, find (i + 1) n) |
|
121 |
in find 1 (Thm.nprems_of st) st end; |
|
122 |
||
123 |
fun HEADGOAL tac = tac 1; |
|
124 |
||
125 |
||
| 5824 | 126 |
|
|
12324
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
wenzelm
parents:
12311
diff
changeset
|
127 |
(** proof methods **) |
|
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
wenzelm
parents:
12311
diff
changeset
|
128 |
|
| 17110 | 129 |
(* datatype method *) |
| 11731 | 130 |
|
| 18227 | 131 |
datatype method = Meth of thm list -> cases_tactic; |
| 11731 | 132 |
|
|
25957
2cfb703fa8d8
improved apply: handle thread position, apply to context here;
wenzelm
parents:
25699
diff
changeset
|
133 |
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
|
134 |
(fn () => let val Meth meth = meth_fun ctxt in meth facts goal end) (); |
| 11731 | 135 |
|
| 17756 | 136 |
val RAW_METHOD_CASES = Meth; |
| 11731 | 137 |
|
| 17110 | 138 |
fun RAW_METHOD tac = RAW_METHOD_CASES (NO_CASES o tac); |
| 12144 | 139 |
|
| 17110 | 140 |
fun METHOD_CASES tac = RAW_METHOD_CASES (fn facts => |
| 21687 | 141 |
Seq.THEN (ALLGOALS Goal.conjunction_tac, tac facts)); |
| 8372 | 142 |
|
| 21687 | 143 |
fun METHOD tac = RAW_METHOD (fn facts => ALLGOALS Goal.conjunction_tac THEN tac facts); |
| 5824 | 144 |
|
145 |
val fail = METHOD (K no_tac); |
|
146 |
val succeed = METHOD (K all_tac); |
|
147 |
||
148 |
||
| 17110 | 149 |
(* insert facts *) |
| 7419 | 150 |
|
151 |
local |
|
| 5824 | 152 |
|
| 21579 | 153 |
fun cut_rule_tac rule = |
154 |
Tactic.rtac (Drule.forall_intr_vars rule COMP_INCR revcut_rl); |
|
| 6981 | 155 |
|
| 7419 | 156 |
in |
| 5824 | 157 |
|
| 7419 | 158 |
fun insert_tac [] i = all_tac |
159 |
| insert_tac facts i = EVERY (map (fn th => cut_rule_tac th i) facts); |
|
| 6981 | 160 |
|
| 7555 | 161 |
val insert_facts = METHOD (ALLGOALS o insert_tac); |
| 7664 | 162 |
fun insert thms = METHOD (fn _ => ALLGOALS (insert_tac thms)); |
| 7419 | 163 |
|
| 9706 | 164 |
fun SIMPLE_METHOD tac = METHOD (fn facts => ALLGOALS (insert_tac facts) THEN tac); |
| 21592 | 165 |
fun SIMPLE_METHOD'' quant tac = METHOD (fn facts => quant (insert_tac facts THEN' tac)); |
166 |
val SIMPLE_METHOD' = SIMPLE_METHOD'' HEADGOAL; |
|
| 9706 | 167 |
|
|
12324
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
wenzelm
parents:
12311
diff
changeset
|
168 |
end; |
|
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
wenzelm
parents:
12311
diff
changeset
|
169 |
|
| 9706 | 170 |
|
| 17110 | 171 |
(* shuffle subgoals *) |
172 |
||
173 |
fun prefer i = METHOD (K (Tactic.defer_tac i THEN PRIMITIVE (Thm.permute_prems 0 ~1))); |
|
| 18939 | 174 |
fun defer opt_i = METHOD (K (Tactic.defer_tac (the_default 1 opt_i))); |
| 17110 | 175 |
|
176 |
||
|
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
177 |
(* cheating *) |
|
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
178 |
|
|
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
179 |
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
|
180 |
(SkipProof.cheat_tac (ProofContext.theory_of ctxt)))); |
|
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
181 |
|
|
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
182 |
|
| 17110 | 183 |
(* unfold intro/elim rules *) |
184 |
||
| 21592 | 185 |
fun intro ths = SIMPLE_METHOD' (CHANGED_PROP o REPEAT_ALL_NEW (Tactic.match_tac ths)); |
186 |
fun elim ths = SIMPLE_METHOD' (CHANGED_PROP o REPEAT_ALL_NEW (Tactic.ematch_tac ths)); |
|
| 17110 | 187 |
|
188 |
||
| 12384 | 189 |
(* unfold/fold definitions *) |
190 |
||
| 18877 | 191 |
fun unfold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (LocalDefs.unfold_tac ctxt ths)); |
192 |
fun fold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (LocalDefs.fold_tac ctxt ths)); |
|
| 6532 | 193 |
|
| 12384 | 194 |
|
| 12829 | 195 |
(* atomize rule statements *) |
196 |
||
|
23590
ad95084a5c63
renamed ObjectLogic.atomize_tac to ObjectLogic.atomize_prems_tac;
wenzelm
parents:
23577
diff
changeset
|
197 |
fun atomize false = SIMPLE_METHOD' (CHANGED_PROP o ObjectLogic.atomize_prems_tac) |
| 12829 | 198 |
| atomize true = RAW_METHOD (K (HEADGOAL (CHANGED_PROP o ObjectLogic.full_atomize_tac))); |
199 |
||
200 |
||
| 18039 | 201 |
(* this -- resolve facts directly *) |
| 12384 | 202 |
|
| 17110 | 203 |
val this = METHOD (EVERY o map (HEADGOAL o Tactic.rtac)); |
| 9484 | 204 |
|
205 |
||
| 18039 | 206 |
(* fact -- composition by facts from context *) |
207 |
||
| 21592 | 208 |
fun fact [] ctxt = SIMPLE_METHOD' (ProofContext.some_fact_tac ctxt) |
209 |
| fact rules _ = SIMPLE_METHOD' (ProofContext.fact_tac rules); |
|
| 18039 | 210 |
|
211 |
||
| 17110 | 212 |
(* assumption *) |
| 7419 | 213 |
|
214 |
local |
|
215 |
||
| 19778 | 216 |
fun cond_rtac cond rule = SUBGOAL (fn (prop, i) => |
217 |
if cond (Logic.strip_assums_concl prop) |
|
218 |
then Tactic.rtac rule i else no_tac); |
|
| 7419 | 219 |
|
|
29857
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
kleing
parents:
29301
diff
changeset
|
220 |
in |
|
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
kleing
parents:
29301
diff
changeset
|
221 |
|
|
30234
7dd251bce291
renamed Method.assumption_tac back to Method.assm_tac -- as assumption_tac it would have to be exactly the tactic behind the assumption method (with facts);
wenzelm
parents:
30190
diff
changeset
|
222 |
fun assm_tac ctxt = |
| 17110 | 223 |
assume_tac APPEND' |
| 23349 | 224 |
Goal.assume_rule_tac ctxt APPEND' |
| 19778 | 225 |
cond_rtac (can Logic.dest_equals) Drule.reflexive_thm APPEND' |
226 |
cond_rtac (can Logic.dest_term) Drule.termI; |
|
| 17110 | 227 |
|
|
30567
cd8e20f86795
close/all_assm_tac: finish all subgoals from left to right (cf. Proof.goal_tac a28d83e903ce) -- NB: ALLGOALS/THEN_ALL_NEW operate from right to left;
wenzelm
parents:
30544
diff
changeset
|
228 |
fun all_assm_tac ctxt st = EVERY1 (replicate (Thm.nprems_of st) (assm_tac ctxt)) st; |
|
cd8e20f86795
close/all_assm_tac: finish all subgoals from left to right (cf. Proof.goal_tac a28d83e903ce) -- NB: ALLGOALS/THEN_ALL_NEW operate from right to left;
wenzelm
parents:
30544
diff
changeset
|
229 |
|
| 23349 | 230 |
fun assumption ctxt = METHOD (HEADGOAL o |
|
30234
7dd251bce291
renamed Method.assumption_tac back to Method.assm_tac -- as assumption_tac it would have to be exactly the tactic behind the assumption method (with facts);
wenzelm
parents:
30190
diff
changeset
|
231 |
(fn [] => assm_tac ctxt |
| 23349 | 232 |
| [fact] => solve_tac [fact] |
233 |
| _ => K no_tac)); |
|
234 |
||
|
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
235 |
fun close immed ctxt = METHOD (K |
|
30567
cd8e20f86795
close/all_assm_tac: finish all subgoals from left to right (cf. Proof.goal_tac a28d83e903ce) -- NB: ALLGOALS/THEN_ALL_NEW operate from right to left;
wenzelm
parents:
30544
diff
changeset
|
236 |
(FILTER Thm.no_prems ((if immed then all_assm_tac ctxt else all_tac) THEN flexflex_tac))); |
| 7419 | 237 |
|
238 |
end; |
|
239 |
||
240 |
||
| 17110 | 241 |
(* rule etc. -- single-step refinements *) |
| 12347 | 242 |
|
| 17110 | 243 |
val trace_rules = ref false; |
| 12347 | 244 |
|
| 17110 | 245 |
fun trace ctxt rules = |
| 21962 | 246 |
if ! trace_rules andalso not (null rules) then |
| 17110 | 247 |
Pretty.big_list "rules:" (map (ProofContext.pretty_thm ctxt) rules) |
| 21962 | 248 |
|> Pretty.string_of |> tracing |
249 |
else (); |
|
| 12347 | 250 |
|
251 |
local |
|
252 |
||
| 18841 | 253 |
fun gen_rule_tac tac rules facts = |
254 |
(fn i => fn st => |
|
255 |
if null facts then tac rules i st |
|
256 |
else Seq.maps (fn rule => (tac o single) rule i st) (Drule.multi_resolves facts rules)) |
|
| 21687 | 257 |
THEN_ALL_NEW Goal.norm_hhf_tac; |
| 7130 | 258 |
|
| 10744 | 259 |
fun gen_arule_tac tac j rules facts = |
260 |
EVERY' (gen_rule_tac tac rules facts :: replicate j Tactic.assume_tac); |
|
261 |
||
| 11785 | 262 |
fun gen_some_rule_tac tac arg_rules ctxt facts = SUBGOAL (fn (goal, i) => |
263 |
let |
|
264 |
val rules = |
|
265 |
if not (null arg_rules) then arg_rules |
|
|
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19307
diff
changeset
|
266 |
else flat (ContextRules.find_rules false facts goal ctxt) |
| 12055 | 267 |
in trace ctxt rules; tac rules facts i end); |
| 10309 | 268 |
|
| 10744 | 269 |
fun meth tac x = METHOD (HEADGOAL o tac x); |
270 |
fun meth' tac x y = METHOD (HEADGOAL o tac x y); |
|
| 8220 | 271 |
|
| 7419 | 272 |
in |
273 |
||
| 10744 | 274 |
val rule_tac = gen_rule_tac Tactic.resolve_tac; |
275 |
val rule = meth rule_tac; |
|
276 |
val some_rule_tac = gen_some_rule_tac rule_tac; |
|
277 |
val some_rule = meth' some_rule_tac; |
|
278 |
||
279 |
val erule = meth' (gen_arule_tac Tactic.eresolve_tac); |
|
280 |
val drule = meth' (gen_arule_tac Tactic.dresolve_tac); |
|
281 |
val frule = meth' (gen_arule_tac Tactic.forward_tac); |
|
| 5824 | 282 |
|
| 7419 | 283 |
end; |
284 |
||
285 |
||
| 25270 | 286 |
(* intros_tac -- pervasive search spanned by intro rules *) |
287 |
||
288 |
fun intros_tac intros facts = |
|
289 |
ALLGOALS (insert_tac facts THEN' |
|
290 |
REPEAT_ALL_NEW (resolve_tac intros)) |
|
291 |
THEN Tactic.distinct_subgoals_tac; |
|
292 |
||
293 |
||
| 8351 | 294 |
(* ML tactics *) |
295 |
||
|
26472
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
296 |
structure TacticData = ProofDataFun |
|
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
297 |
( |
| 27235 | 298 |
type T = thm list -> tactic; |
|
26472
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
299 |
fun init _ = undefined; |
|
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
300 |
); |
|
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
301 |
|
|
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
302 |
val set_tactic = TacticData.put; |
| 8351 | 303 |
|
|
26472
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
304 |
fun ml_tactic (txt, pos) ctxt = |
|
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
305 |
let |
|
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
306 |
val ctxt' = ctxt |> Context.proof_map |
| 27235 | 307 |
(ML_Context.expression pos |
308 |
"fun tactic (facts: thm list) : tactic" |
|
|
26472
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
309 |
"Context.map_proof (Method.set_tactic tactic)" txt); |
| 27235 | 310 |
in Context.setmp_thread_data (SOME (Context.Proof ctxt)) (TacticData.get ctxt') end; |
| 23425 | 311 |
|
312 |
fun tactic txt ctxt = METHOD (ml_tactic txt ctxt); |
|
313 |
fun raw_tactic txt ctxt = RAW_METHOD (ml_tactic txt ctxt); |
|
| 8351 | 314 |
|
315 |
||
| 5824 | 316 |
|
| 17110 | 317 |
(** method syntax **) |
318 |
||
319 |
(* method text *) |
|
320 |
||
321 |
type src = Args.src; |
|
| 5824 | 322 |
|
| 17110 | 323 |
datatype text = |
| 23349 | 324 |
Basic of (Proof.context -> method) * Position.T | |
| 17110 | 325 |
Source of src | |
| 20030 | 326 |
Source_i of src | |
| 17110 | 327 |
Then of text list | |
328 |
Orelse of text list | |
|
329 |
Try of text | |
|
| 19186 | 330 |
Repeat1 of text | |
331 |
SelectGoals of int * text; |
|
| 17110 | 332 |
|
| 23349 | 333 |
fun primitive_text r = Basic (K (SIMPLE_METHOD (PRIMITIVE r)), Position.none); |
334 |
val succeed_text = Basic (K succeed, Position.none); |
|
| 17110 | 335 |
val default_text = Source (Args.src (("default", []), Position.none));
|
| 23349 | 336 |
val this_text = Basic (K this, Position.none); |
337 |
val done_text = Basic (K (SIMPLE_METHOD all_tac), Position.none); |
|
338 |
fun sorry_text int = Basic (cheating int, Position.none); |
|
| 17110 | 339 |
|
| 23349 | 340 |
fun finish_text (NONE, immed) pos = Basic (close immed, pos) |
341 |
| finish_text (SOME txt, immed) pos = Then [txt, Basic (close immed, pos)]; |
|
| 17110 | 342 |
|
343 |
||
344 |
(* method definitions *) |
|
| 5824 | 345 |
|
| 24116 | 346 |
structure Methods = TheoryDataFun |
| 22846 | 347 |
( |
| 20289 | 348 |
type T = (((src -> Proof.context -> method) * string) * stamp) NameSpace.table; |
| 16347 | 349 |
val empty = NameSpace.empty_table; |
| 6546 | 350 |
val copy = I; |
| 16448 | 351 |
val extend = I; |
|
23655
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents:
23590
diff
changeset
|
352 |
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
|
353 |
error ("Attempt to merge different versions of method " ^ quote dup);
|
| 22846 | 354 |
); |
| 5824 | 355 |
|
| 22846 | 356 |
fun print_methods thy = |
357 |
let |
|
| 24116 | 358 |
val meths = Methods.get thy; |
| 22846 | 359 |
fun prt_meth (name, ((_, comment), _)) = Pretty.block |
360 |
[Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment]; |
|
361 |
in |
|
362 |
[Pretty.big_list "methods:" (map prt_meth (NameSpace.extern_table meths))] |
|
363 |
|> Pretty.chunks |> Pretty.writeln |
|
364 |
end; |
|
| 7611 | 365 |
|
| 5824 | 366 |
|
367 |
(* get methods *) |
|
368 |
||
| 26892 | 369 |
val intern = NameSpace.intern o #1 o Methods.get; |
370 |
val defined = Symtab.defined o #2 o Methods.get; |
|
371 |
||
| 20030 | 372 |
fun method_i thy = |
| 5824 | 373 |
let |
| 24116 | 374 |
val meths = #2 (Methods.get thy); |
| 5884 | 375 |
fun meth src = |
| 20030 | 376 |
let val ((name, _), pos) = Args.dest_src src in |
| 17412 | 377 |
(case Symtab.lookup meths name of |
| 15531 | 378 |
NONE => error ("Unknown proof method: " ^ quote name ^ Position.str_of pos)
|
| 27751 | 379 |
| SOME ((mth, _), _) => (Position.report (Markup.method name) pos; mth src)) |
| 5824 | 380 |
end; |
381 |
in meth end; |
|
382 |
||
| 24116 | 383 |
fun method thy = method_i thy o Args.map_name (NameSpace.intern (#1 (Methods.get thy))); |
| 20030 | 384 |
|
| 5824 | 385 |
|
| 17110 | 386 |
(* add method *) |
| 5824 | 387 |
|
388 |
fun add_methods raw_meths thy = |
|
389 |
let |
|
| 16145 | 390 |
val new_meths = raw_meths |> map (fn (name, f, comment) => |
| 29004 | 391 |
(Binding.name name, ((f, comment), stamp ()))); |
| 5824 | 392 |
|
| 30466 | 393 |
fun add meths = fold (snd oo NameSpace.define (Sign.naming_of thy)) new_meths meths |
|
23655
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents:
23590
diff
changeset
|
394 |
handle Symtab.DUP dup => error ("Duplicate declaration of method " ^ quote dup);
|
| 24116 | 395 |
in Methods.map add thy end; |
| 5824 | 396 |
|
| 9194 | 397 |
val add_method = add_methods o Library.single; |
398 |
||
| 5824 | 399 |
|
| 30512 | 400 |
(* method setup *) |
401 |
||
402 |
fun syntax scan = Args.context_syntax "method" scan; |
|
403 |
||
404 |
fun setup name scan comment = |
|
405 |
add_methods [(Binding.name_of name, |
|
406 |
fn src => fn ctxt => let val (m, ctxt') = syntax scan src ctxt in m ctxt' end, comment)]; |
|
|
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
407 |
|
|
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26291
diff
changeset
|
408 |
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
|
409 |
Context.theory_map (ML_Context.expression pos |
| 30544 | 410 |
"val (name, scan, comment): binding * (Proof.context -> Proof.method) context_parser * string" |
411 |
"Context.map_theory (Method.setup name scan comment)" |
|
412 |
("(" ^ ML_Syntax.make_binding name ^ ", " ^ txt ^ ", " ^ ML_Syntax.print_string cmt ^ ")"));
|
|
|
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
413 |
|
|
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
414 |
|
| 5884 | 415 |
|
| 17110 | 416 |
(** concrete syntax **) |
| 5824 | 417 |
|
|
27813
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
418 |
structure P = OuterParse; |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
419 |
|
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
420 |
|
| 5884 | 421 |
(* basic *) |
422 |
||
| 17110 | 423 |
fun simple_args scan f src ctxt : method = |
| 21879 | 424 |
fst (syntax (Scan.lift (scan >> (fn x => f x ctxt))) src ctxt); |
| 8351 | 425 |
|
| 20289 | 426 |
fun ctxt_args (f: Proof.context -> method) src ctxt = |
| 21879 | 427 |
fst (syntax (Scan.succeed (f ctxt)) src ctxt); |
| 7555 | 428 |
|
429 |
fun no_args m = ctxt_args (K m); |
|
| 5884 | 430 |
|
431 |
||
432 |
(* sections *) |
|
| 5824 | 433 |
|
| 20289 | 434 |
type modifier = (Proof.context -> Proof.context) * attribute; |
| 7268 | 435 |
|
436 |
local |
|
437 |
||
|
24010
2ef318813e1a
method section scanners: added [[declaration]] syntax, ignore sid-effects of thms;
wenzelm
parents:
23937
diff
changeset
|
438 |
fun thms ss = Scan.repeat (Scan.unless (Scan.lift (Scan.first ss)) Attrib.multi_thm) >> flat; |
| 30190 | 439 |
fun app (f, att) (context, ths) = Library.foldl_map att (Context.map_proof f context, ths); |
| 5824 | 440 |
|
| 30540 | 441 |
in |
442 |
||
| 24022 | 443 |
fun section ss = Scan.depend (fn context => (Scan.first ss -- Scan.pass context (thms ss)) :|-- |
444 |
(fn (m, ths) => Scan.succeed (app m (context, ths)))); |
|
| 5884 | 445 |
|
| 30540 | 446 |
fun sections ss = Scan.repeat (section ss); |
| 5824 | 447 |
|
| 5884 | 448 |
fun sectioned_args args ss f src ctxt = |
| 30540 | 449 |
let val ((x, _), ctxt') = syntax (args -- sections ss) src ctxt |
| 5921 | 450 |
in f x ctxt' end; |
| 5884 | 451 |
|
| 7601 | 452 |
fun bang_sectioned_args ss f = sectioned_args Args.bang_facts ss f; |
| 9777 | 453 |
fun bang_sectioned_args' ss scan f = |
454 |
sectioned_args (Args.bang_facts -- scan >> swap) ss (uncurry f); |
|
| 7601 | 455 |
fun only_sectioned_args ss f = sectioned_args (Scan.succeed ()) ss (fn () => f); |
| 7268 | 456 |
|
| 18999 | 457 |
fun thms_ctxt_args f = sectioned_args (thms []) [] f; |
| 8093 | 458 |
fun thms_args f = thms_ctxt_args (K o f); |
| 9706 | 459 |
fun thm_args f = thms_args (fn [thm] => f thm | _ => error "Single theorem expected"); |
| 5824 | 460 |
|
| 7268 | 461 |
end; |
462 |
||
| 5824 | 463 |
|
| 30515 | 464 |
(* extra rule methods *) |
465 |
||
466 |
fun xrule_meth m = |
|
467 |
Scan.lift (Scan.optional (Args.parens OuterParse.nat) 0) -- Attrib.thms >> |
|
468 |
(fn (n, ths) => K (m n ths)); |
|
469 |
||
470 |
||
| 9539 | 471 |
(* tactic syntax *) |
| 8238 | 472 |
|
| 30515 | 473 |
fun goal_args' args tac src ctxt = fst (syntax (Args.goal_spec -- args >> |
| 21592 | 474 |
(fn (quant, s) => SIMPLE_METHOD'' quant (tac s))) src ctxt); |
| 8537 | 475 |
|
| 9539 | 476 |
fun goal_args args tac = goal_args' (Scan.lift args) tac; |
| 8238 | 477 |
|
|
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
478 |
fun goal_args_ctxt' args tac src ctxt = |
| 30515 | 479 |
fst (syntax (Args.goal_spec -- args >> |
| 21592 | 480 |
(fn (quant, s) => SIMPLE_METHOD'' quant (tac ctxt s))) src ctxt); |
| 8238 | 481 |
|
|
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
482 |
fun goal_args_ctxt args tac = goal_args_ctxt' (Scan.lift args) tac; |
| 5824 | 483 |
|
| 14718 | 484 |
|
|
27813
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
485 |
(* outer parser *) |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
486 |
|
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
487 |
fun is_symid_meth s = |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
488 |
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
|
489 |
|
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
490 |
local |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
491 |
|
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
492 |
fun meth4 x = |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
493 |
(P.position (P.xname >> rpair []) >> (Source o Args.src) || |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
494 |
P.$$$ "(" |-- P.!!! (meth0 --| P.$$$ ")")) x
|
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
495 |
and meth3 x = |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
496 |
(meth4 --| P.$$$ "?" >> Try || |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
497 |
meth4 --| P.$$$ "+" >> Repeat1 || |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
498 |
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
|
499 |
meth4) x |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
500 |
and meth2 x = |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
501 |
(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
|
502 |
meth3) x |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
503 |
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
|
504 |
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
|
505 |
|
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
506 |
in val parse = meth3 end; |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
507 |
|
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
508 |
|
| 18708 | 509 |
(* theory setup *) |
| 5824 | 510 |
|
| 26463 | 511 |
val _ = Context.>> (Context.map_theory |
| 30515 | 512 |
(setup (Binding.name "fail") (Scan.succeed (K fail)) "force failure" #> |
513 |
setup (Binding.name "succeed") (Scan.succeed (K succeed)) "succeed" #> |
|
514 |
setup (Binding.name "-") (Scan.succeed (K insert_facts)) |
|
515 |
"do nothing (insert current facts only)" #> |
|
516 |
setup (Binding.name "insert") (Attrib.thms >> (K o insert)) |
|
517 |
"insert theorems, ignoring facts (improper)" #> |
|
518 |
setup (Binding.name "intro") (Attrib.thms >> (K o intro)) |
|
519 |
"repeatedly apply introduction rules" #> |
|
520 |
setup (Binding.name "elim") (Attrib.thms >> (K o elim)) |
|
521 |
"repeatedly apply elimination rules" #> |
|
522 |
setup (Binding.name "unfold") (Attrib.thms >> unfold_meth) "unfold definitions" #> |
|
523 |
setup (Binding.name "fold") (Attrib.thms >> fold_meth) "fold definitions" #> |
|
524 |
setup (Binding.name "atomize") (Args.mode "full" >> (K o atomize)) |
|
525 |
"present local premises as object-level statements" #> |
|
526 |
setup (Binding.name "rule") (Attrib.thms >> some_rule) "apply some intro/elim rule" #> |
|
527 |
setup (Binding.name "erule") (xrule_meth erule) "apply rule in elimination manner (improper)" #> |
|
528 |
setup (Binding.name "drule") (xrule_meth drule) "apply rule in destruct manner (improper)" #> |
|
529 |
setup (Binding.name "frule") (xrule_meth frule) "apply rule in forward manner (improper)" #> |
|
530 |
setup (Binding.name "this") (Scan.succeed (K this)) "apply current facts as rules" #> |
|
531 |
setup (Binding.name "fact") (Attrib.thms >> fact) "composition by facts from context" #> |
|
532 |
setup (Binding.name "assumption") (Scan.succeed assumption) |
|
533 |
"proof by assumption, preferring facts" #> |
|
534 |
setup (Binding.name "rename_tac") (Args.goal_spec -- Scan.lift (Scan.repeat1 Args.name) >> |
|
535 |
(fn (quant, xs) => K (SIMPLE_METHOD'' quant (Tactic.rename_tac xs)))) |
|
536 |
"rename parameters of goal" #> |
|
537 |
setup (Binding.name "rotate_tac") (Args.goal_spec -- Scan.lift (Scan.optional P.int 1) >> |
|
538 |
(fn (quant, i) => K (SIMPLE_METHOD'' quant (Tactic.rotate_tac i)))) |
|
539 |
"rotate assumptions of goal" #> |
|
540 |
setup (Binding.name "tactic") (Scan.lift (P.position Args.name) >> tactic) |
|
541 |
"ML tactic as proof method" #> |
|
542 |
setup (Binding.name "raw_tactic") (Scan.lift (P.position Args.name) >> raw_tactic) |
|
543 |
"ML tactic as raw proof method")); |
|
| 5824 | 544 |
|
545 |
||
| 16145 | 546 |
(*final declarations of this structure!*) |
547 |
val unfold = unfold_meth; |
|
548 |
val fold = fold_meth; |
|
549 |
||
| 5824 | 550 |
end; |
551 |
||
552 |
structure BasicMethod: BASIC_METHOD = Method; |
|
553 |
open BasicMethod; |
|
|
30508
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents:
30466
diff
changeset
|
554 |
|
|
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents:
30466
diff
changeset
|
555 |
val RAW_METHOD_CASES = Method.RAW_METHOD_CASES; |
|
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents:
30466
diff
changeset
|
556 |
val RAW_METHOD = Method.RAW_METHOD; |
|
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents:
30466
diff
changeset
|
557 |
val METHOD_CASES = Method.METHOD_CASES; |
|
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents:
30466
diff
changeset
|
558 |
val METHOD = Method.METHOD; |
|
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents:
30466
diff
changeset
|
559 |
val SIMPLE_METHOD = Method.SIMPLE_METHOD; |
|
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents:
30466
diff
changeset
|
560 |
val SIMPLE_METHOD' = Method.SIMPLE_METHOD'; |
|
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents:
30466
diff
changeset
|
561 |
val SIMPLE_METHOD'' = Method.SIMPLE_METHOD''; |
|
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents:
30466
diff
changeset
|
562 |