author | blanchet |
Mon, 06 Dec 2010 13:29:23 +0100 | |
changeset 40996 | 63112be4a469 |
parent 39507 | 839873937ddd |
child 41379 | b31d7a1cd08f |
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 |
|
32738 | 11 |
val trace_rules: bool Unsynchronized.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 |
32193
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents:
32091
diff
changeset
|
18 |
val apply: (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 |
36093
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
ballarin
parents:
33522
diff
changeset
|
49 |
val try_intros_tac: thm list -> thm list -> tactic |
17110 | 50 |
val rule: thm list -> method |
51 |
val erule: int -> thm list -> method |
|
52 |
val drule: int -> thm list -> method |
|
53 |
val frule: int -> thm list -> method |
|
27235 | 54 |
val set_tactic: (thm list -> tactic) -> Proof.context -> Proof.context |
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26291
diff
changeset
|
55 |
val tactic: string * Position.T -> Proof.context -> method |
27235 | 56 |
val raw_tactic: string * Position.T -> Proof.context -> method |
27729 | 57 |
type src = Args.src |
5824 | 58 |
datatype text = |
32193
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents:
32091
diff
changeset
|
59 |
Basic of Proof.context -> method | |
15703 | 60 |
Source of src | |
20030 | 61 |
Source_i of src | |
5824 | 62 |
Then of text list | |
63 |
Orelse of text list | |
|
64 |
Try of text | |
|
19186 | 65 |
Repeat1 of text | |
66 |
SelectGoals of int * text |
|
17857 | 67 |
val primitive_text: (thm -> thm) -> text |
68 |
val succeed_text: text |
|
17110 | 69 |
val default_text: text |
70 |
val this_text: text |
|
71 |
val done_text: text |
|
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
72 |
val sorry_text: bool -> text |
32193
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents:
32091
diff
changeset
|
73 |
val finish_text: text option * bool -> text |
27729 | 74 |
val print_methods: theory -> unit |
26892 | 75 |
val intern: theory -> xstring -> string |
76 |
val defined: theory -> string -> bool |
|
20289 | 77 |
val method: theory -> src -> Proof.context -> method |
78 |
val method_i: theory -> src -> Proof.context -> method |
|
30512 | 79 |
val syntax: 'a context_parser -> src -> Proof.context -> 'a * Proof.context |
80 |
val setup: binding -> (Proof.context -> method) context_parser -> string -> theory -> theory |
|
30575 | 81 |
val method_setup: bstring * Position.T -> Symbol_Pos.text * Position.T -> string -> |
82 |
theory -> theory |
|
30540 | 83 |
type modifier = (Proof.context -> Proof.context) * attribute |
84 |
val section: modifier parser list -> thm list context_parser |
|
85 |
val sections: modifier parser list -> thm list list context_parser |
|
30512 | 86 |
val parse: text parser |
5824 | 87 |
end; |
88 |
||
89 |
structure Method: METHOD = |
|
90 |
struct |
|
91 |
||
17110 | 92 |
(** generic tools **) |
93 |
||
94 |
(* goal addressing *) |
|
95 |
||
96 |
fun FINDGOAL tac st = |
|
97 |
let fun find i n = if i > n then Seq.fail else Seq.APPEND (tac i, find (i + 1) n) |
|
98 |
in find 1 (Thm.nprems_of st) st end; |
|
99 |
||
100 |
fun HEADGOAL tac = tac 1; |
|
101 |
||
102 |
||
5824 | 103 |
|
12324
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
wenzelm
parents:
12311
diff
changeset
|
104 |
(** proof methods **) |
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
wenzelm
parents:
12311
diff
changeset
|
105 |
|
17110 | 106 |
(* datatype method *) |
11731 | 107 |
|
18227 | 108 |
datatype method = Meth of thm list -> cases_tactic; |
11731 | 109 |
|
32193
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents:
32091
diff
changeset
|
110 |
fun apply meth ctxt = let val Meth m = meth ctxt in m end; |
11731 | 111 |
|
17756 | 112 |
val RAW_METHOD_CASES = Meth; |
11731 | 113 |
|
17110 | 114 |
fun RAW_METHOD tac = RAW_METHOD_CASES (NO_CASES o tac); |
12144 | 115 |
|
17110 | 116 |
fun METHOD_CASES tac = RAW_METHOD_CASES (fn facts => |
21687 | 117 |
Seq.THEN (ALLGOALS Goal.conjunction_tac, tac facts)); |
8372 | 118 |
|
21687 | 119 |
fun METHOD tac = RAW_METHOD (fn facts => ALLGOALS Goal.conjunction_tac THEN tac facts); |
5824 | 120 |
|
121 |
val fail = METHOD (K no_tac); |
|
122 |
val succeed = METHOD (K all_tac); |
|
123 |
||
124 |
||
17110 | 125 |
(* insert facts *) |
7419 | 126 |
|
127 |
local |
|
5824 | 128 |
|
21579 | 129 |
fun cut_rule_tac rule = |
130 |
Tactic.rtac (Drule.forall_intr_vars rule COMP_INCR revcut_rl); |
|
6981 | 131 |
|
7419 | 132 |
in |
5824 | 133 |
|
7419 | 134 |
fun insert_tac [] i = all_tac |
135 |
| insert_tac facts i = EVERY (map (fn th => cut_rule_tac th i) facts); |
|
6981 | 136 |
|
7555 | 137 |
val insert_facts = METHOD (ALLGOALS o insert_tac); |
7664 | 138 |
fun insert thms = METHOD (fn _ => ALLGOALS (insert_tac thms)); |
7419 | 139 |
|
9706 | 140 |
fun SIMPLE_METHOD tac = METHOD (fn facts => ALLGOALS (insert_tac facts) THEN tac); |
21592 | 141 |
fun SIMPLE_METHOD'' quant tac = METHOD (fn facts => quant (insert_tac facts THEN' tac)); |
142 |
val SIMPLE_METHOD' = SIMPLE_METHOD'' HEADGOAL; |
|
9706 | 143 |
|
12324
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
wenzelm
parents:
12311
diff
changeset
|
144 |
end; |
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
wenzelm
parents:
12311
diff
changeset
|
145 |
|
9706 | 146 |
|
17110 | 147 |
(* shuffle subgoals *) |
148 |
||
149 |
fun prefer i = METHOD (K (Tactic.defer_tac i THEN PRIMITIVE (Thm.permute_prems 0 ~1))); |
|
18939 | 150 |
fun defer opt_i = METHOD (K (Tactic.defer_tac (the_default 1 opt_i))); |
17110 | 151 |
|
152 |
||
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
153 |
(* cheating *) |
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
154 |
|
32969 | 155 |
fun cheating int ctxt = |
156 |
if int orelse ! quick_and_dirty then |
|
157 |
METHOD (K (Skip_Proof.cheat_tac (ProofContext.theory_of ctxt))) |
|
158 |
else error "Cheating requires quick_and_dirty mode!"; |
|
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
159 |
|
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
160 |
|
17110 | 161 |
(* unfold intro/elim rules *) |
162 |
||
21592 | 163 |
fun intro ths = SIMPLE_METHOD' (CHANGED_PROP o REPEAT_ALL_NEW (Tactic.match_tac ths)); |
164 |
fun elim ths = SIMPLE_METHOD' (CHANGED_PROP o REPEAT_ALL_NEW (Tactic.ematch_tac ths)); |
|
17110 | 165 |
|
166 |
||
12384 | 167 |
(* unfold/fold definitions *) |
168 |
||
35624 | 169 |
fun unfold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (Local_Defs.unfold_tac ctxt ths)); |
170 |
fun fold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (Local_Defs.fold_tac ctxt ths)); |
|
6532 | 171 |
|
12384 | 172 |
|
12829 | 173 |
(* atomize rule statements *) |
174 |
||
35625 | 175 |
fun atomize false = SIMPLE_METHOD' (CHANGED_PROP o Object_Logic.atomize_prems_tac) |
176 |
| atomize true = RAW_METHOD (K (HEADGOAL (CHANGED_PROP o Object_Logic.full_atomize_tac))); |
|
12829 | 177 |
|
178 |
||
18039 | 179 |
(* this -- resolve facts directly *) |
12384 | 180 |
|
17110 | 181 |
val this = METHOD (EVERY o map (HEADGOAL o Tactic.rtac)); |
9484 | 182 |
|
183 |
||
18039 | 184 |
(* fact -- composition by facts from context *) |
185 |
||
21592 | 186 |
fun fact [] ctxt = SIMPLE_METHOD' (ProofContext.some_fact_tac ctxt) |
187 |
| fact rules _ = SIMPLE_METHOD' (ProofContext.fact_tac rules); |
|
18039 | 188 |
|
189 |
||
17110 | 190 |
(* assumption *) |
7419 | 191 |
|
192 |
local |
|
193 |
||
19778 | 194 |
fun cond_rtac cond rule = SUBGOAL (fn (prop, i) => |
195 |
if cond (Logic.strip_assums_concl prop) |
|
196 |
then Tactic.rtac rule i else no_tac); |
|
7419 | 197 |
|
29857
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
kleing
parents:
29301
diff
changeset
|
198 |
in |
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
kleing
parents:
29301
diff
changeset
|
199 |
|
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
|
200 |
fun assm_tac ctxt = |
17110 | 201 |
assume_tac APPEND' |
23349 | 202 |
Goal.assume_rule_tac ctxt APPEND' |
19778 | 203 |
cond_rtac (can Logic.dest_equals) Drule.reflexive_thm APPEND' |
204 |
cond_rtac (can Logic.dest_term) Drule.termI; |
|
17110 | 205 |
|
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
|
206 |
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
|
207 |
|
23349 | 208 |
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
|
209 |
(fn [] => assm_tac ctxt |
23349 | 210 |
| [fact] => solve_tac [fact] |
211 |
| _ => K no_tac)); |
|
212 |
||
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
213 |
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
|
214 |
(FILTER Thm.no_prems ((if immed then all_assm_tac ctxt else all_tac) THEN flexflex_tac))); |
7419 | 215 |
|
216 |
end; |
|
217 |
||
218 |
||
17110 | 219 |
(* rule etc. -- single-step refinements *) |
12347 | 220 |
|
32738 | 221 |
val trace_rules = Unsynchronized.ref false; |
12347 | 222 |
|
17110 | 223 |
fun trace ctxt rules = |
21962 | 224 |
if ! trace_rules andalso not (null rules) then |
32091
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents:
31304
diff
changeset
|
225 |
Pretty.big_list "rules:" (map (Display.pretty_thm ctxt) rules) |
21962 | 226 |
|> Pretty.string_of |> tracing |
227 |
else (); |
|
12347 | 228 |
|
229 |
local |
|
230 |
||
18841 | 231 |
fun gen_rule_tac tac rules facts = |
232 |
(fn i => fn st => |
|
233 |
if null facts then tac rules i st |
|
234 |
else Seq.maps (fn rule => (tac o single) rule i st) (Drule.multi_resolves facts rules)) |
|
21687 | 235 |
THEN_ALL_NEW Goal.norm_hhf_tac; |
7130 | 236 |
|
10744 | 237 |
fun gen_arule_tac tac j rules facts = |
238 |
EVERY' (gen_rule_tac tac rules facts :: replicate j Tactic.assume_tac); |
|
239 |
||
11785 | 240 |
fun gen_some_rule_tac tac arg_rules ctxt facts = SUBGOAL (fn (goal, i) => |
241 |
let |
|
242 |
val rules = |
|
243 |
if not (null arg_rules) then arg_rules |
|
33369 | 244 |
else flat (Context_Rules.find_rules false facts goal ctxt) |
12055 | 245 |
in trace ctxt rules; tac rules facts i end); |
10309 | 246 |
|
10744 | 247 |
fun meth tac x = METHOD (HEADGOAL o tac x); |
248 |
fun meth' tac x y = METHOD (HEADGOAL o tac x y); |
|
8220 | 249 |
|
7419 | 250 |
in |
251 |
||
10744 | 252 |
val rule_tac = gen_rule_tac Tactic.resolve_tac; |
253 |
val rule = meth rule_tac; |
|
254 |
val some_rule_tac = gen_some_rule_tac rule_tac; |
|
255 |
val some_rule = meth' some_rule_tac; |
|
256 |
||
257 |
val erule = meth' (gen_arule_tac Tactic.eresolve_tac); |
|
258 |
val drule = meth' (gen_arule_tac Tactic.dresolve_tac); |
|
259 |
val frule = meth' (gen_arule_tac Tactic.forward_tac); |
|
5824 | 260 |
|
7419 | 261 |
end; |
262 |
||
263 |
||
25270 | 264 |
(* intros_tac -- pervasive search spanned by intro rules *) |
265 |
||
36093
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
ballarin
parents:
33522
diff
changeset
|
266 |
fun gen_intros_tac goals intros facts = |
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
ballarin
parents:
33522
diff
changeset
|
267 |
goals (insert_tac facts THEN' |
25270 | 268 |
REPEAT_ALL_NEW (resolve_tac intros)) |
269 |
THEN Tactic.distinct_subgoals_tac; |
|
270 |
||
36093
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
ballarin
parents:
33522
diff
changeset
|
271 |
val intros_tac = gen_intros_tac ALLGOALS; |
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
ballarin
parents:
33522
diff
changeset
|
272 |
val try_intros_tac = gen_intros_tac TRYALL; |
25270 | 273 |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37198
diff
changeset
|
274 |
|
8351 | 275 |
(* ML tactics *) |
276 |
||
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37198
diff
changeset
|
277 |
structure ML_Tactic = Proof_Data |
26472
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
278 |
( |
27235 | 279 |
type T = thm list -> tactic; |
26472
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
280 |
fun init _ = undefined; |
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
281 |
); |
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
282 |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37198
diff
changeset
|
283 |
val set_tactic = ML_Tactic.put; |
8351 | 284 |
|
26472
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
285 |
fun ml_tactic (txt, pos) ctxt = |
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
286 |
let |
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
287 |
val ctxt' = ctxt |> Context.proof_map |
27235 | 288 |
(ML_Context.expression pos |
289 |
"fun tactic (facts: thm list) : tactic" |
|
37198
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
290 |
"Context.map_proof (Method.set_tactic tactic)" (ML_Lex.read pos txt)); |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37198
diff
changeset
|
291 |
in Context.setmp_thread_data (SOME (Context.Proof ctxt)) (ML_Tactic.get ctxt') end; |
23425 | 292 |
|
293 |
fun tactic txt ctxt = METHOD (ml_tactic txt ctxt); |
|
294 |
fun raw_tactic txt ctxt = RAW_METHOD (ml_tactic txt ctxt); |
|
8351 | 295 |
|
296 |
||
5824 | 297 |
|
17110 | 298 |
(** method syntax **) |
299 |
||
300 |
(* method text *) |
|
301 |
||
302 |
type src = Args.src; |
|
5824 | 303 |
|
17110 | 304 |
datatype text = |
32193
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents:
32091
diff
changeset
|
305 |
Basic of Proof.context -> method | |
17110 | 306 |
Source of src | |
20030 | 307 |
Source_i of src | |
17110 | 308 |
Then of text list | |
309 |
Orelse of text list | |
|
310 |
Try of text | |
|
19186 | 311 |
Repeat1 of text | |
312 |
SelectGoals of int * text; |
|
17110 | 313 |
|
32193
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents:
32091
diff
changeset
|
314 |
fun primitive_text r = Basic (K (SIMPLE_METHOD (PRIMITIVE r))); |
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents:
32091
diff
changeset
|
315 |
val succeed_text = Basic (K succeed); |
17110 | 316 |
val default_text = Source (Args.src (("default", []), Position.none)); |
32193
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents:
32091
diff
changeset
|
317 |
val this_text = Basic (K this); |
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents:
32091
diff
changeset
|
318 |
val done_text = Basic (K (SIMPLE_METHOD all_tac)); |
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents:
32091
diff
changeset
|
319 |
fun sorry_text int = Basic (cheating int); |
17110 | 320 |
|
32193
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents:
32091
diff
changeset
|
321 |
fun finish_text (NONE, immed) = Basic (close immed) |
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents:
32091
diff
changeset
|
322 |
| finish_text (SOME txt, immed) = Then [txt, Basic (close immed)]; |
17110 | 323 |
|
324 |
||
325 |
(* method definitions *) |
|
5824 | 326 |
|
33522 | 327 |
structure Methods = Theory_Data |
22846 | 328 |
( |
33095
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
wenzelm
parents:
33092
diff
changeset
|
329 |
type T = ((src -> Proof.context -> method) * string) Name_Space.table; |
33159 | 330 |
val empty : T = Name_Space.empty_table "method"; |
16448 | 331 |
val extend = I; |
33522 | 332 |
fun merge data : T = Name_Space.merge_tables data; |
22846 | 333 |
); |
5824 | 334 |
|
22846 | 335 |
fun print_methods thy = |
336 |
let |
|
24116 | 337 |
val meths = Methods.get thy; |
33092
c859019d3ac5
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
wenzelm
parents:
32969
diff
changeset
|
338 |
fun prt_meth (name, (_, comment)) = Pretty.block |
22846 | 339 |
[Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment]; |
340 |
in |
|
33095
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
wenzelm
parents:
33092
diff
changeset
|
341 |
[Pretty.big_list "methods:" (map prt_meth (Name_Space.extern_table meths))] |
22846 | 342 |
|> Pretty.chunks |> Pretty.writeln |
343 |
end; |
|
7611 | 344 |
|
33092
c859019d3ac5
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
wenzelm
parents:
32969
diff
changeset
|
345 |
fun add_method name meth comment thy = thy |
33095
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
wenzelm
parents:
33092
diff
changeset
|
346 |
|> Methods.map (#2 o Name_Space.define true (Sign.naming_of thy) (name, (meth, comment))); |
31304 | 347 |
|
5824 | 348 |
|
349 |
(* get methods *) |
|
350 |
||
33095
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
wenzelm
parents:
33092
diff
changeset
|
351 |
val intern = Name_Space.intern o #1 o Methods.get; |
26892 | 352 |
val defined = Symtab.defined o #2 o Methods.get; |
353 |
||
20030 | 354 |
fun method_i thy = |
5824 | 355 |
let |
24116 | 356 |
val meths = #2 (Methods.get thy); |
5884 | 357 |
fun meth src = |
20030 | 358 |
let val ((name, _), pos) = Args.dest_src src in |
17412 | 359 |
(case Symtab.lookup meths name of |
15531 | 360 |
NONE => error ("Unknown proof method: " ^ quote name ^ Position.str_of pos) |
39507
839873937ddd
tuned signature of (Context_)Position.report variants;
wenzelm
parents:
37216
diff
changeset
|
361 |
| SOME (mth, _) => (Position.report pos (Markup.method name); mth src)) |
5824 | 362 |
end; |
363 |
in meth end; |
|
364 |
||
33095
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
wenzelm
parents:
33092
diff
changeset
|
365 |
fun method thy = method_i thy o Args.map_name (Name_Space.intern (#1 (Methods.get thy))); |
20030 | 366 |
|
5824 | 367 |
|
30512 | 368 |
(* method setup *) |
369 |
||
370 |
fun syntax scan = Args.context_syntax "method" scan; |
|
371 |
||
31304 | 372 |
fun setup name scan = |
373 |
add_method name |
|
374 |
(fn src => fn ctxt => let val (m, ctxt') = syntax scan src ctxt in m ctxt' end); |
|
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
375 |
|
26385
ae7564661e76
ML runtime compilation: pass position, tuned signature;
wenzelm
parents:
26291
diff
changeset
|
376 |
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
|
377 |
Context.theory_map (ML_Context.expression pos |
30544 | 378 |
"val (name, scan, comment): binding * (Proof.context -> Proof.method) context_parser * string" |
379 |
"Context.map_theory (Method.setup name scan comment)" |
|
37198
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
380 |
(ML_Lex.read Position.none ("(" ^ ML_Syntax.make_binding name ^ ", ") @ |
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
381 |
ML_Lex.read pos txt @ |
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
382 |
ML_Lex.read Position.none (", " ^ ML_Syntax.print_string cmt ^ ")"))); |
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
383 |
|
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
384 |
|
5884 | 385 |
|
17110 | 386 |
(** concrete syntax **) |
5824 | 387 |
|
5884 | 388 |
(* sections *) |
5824 | 389 |
|
20289 | 390 |
type modifier = (Proof.context -> Proof.context) * attribute; |
7268 | 391 |
|
392 |
local |
|
393 |
||
24010
2ef318813e1a
method section scanners: added [[declaration]] syntax, ignore sid-effects of thms;
wenzelm
parents:
23937
diff
changeset
|
394 |
fun thms ss = Scan.repeat (Scan.unless (Scan.lift (Scan.first ss)) Attrib.multi_thm) >> flat; |
30190 | 395 |
fun app (f, att) (context, ths) = Library.foldl_map att (Context.map_proof f context, ths); |
5824 | 396 |
|
30540 | 397 |
in |
398 |
||
24022 | 399 |
fun section ss = Scan.depend (fn context => (Scan.first ss -- Scan.pass context (thms ss)) :|-- |
400 |
(fn (m, ths) => Scan.succeed (app m (context, ths)))); |
|
5884 | 401 |
|
30540 | 402 |
fun sections ss = Scan.repeat (section ss); |
5824 | 403 |
|
7268 | 404 |
end; |
405 |
||
5824 | 406 |
|
30515 | 407 |
(* extra rule methods *) |
408 |
||
409 |
fun xrule_meth m = |
|
36950 | 410 |
Scan.lift (Scan.optional (Args.parens Parse.nat) 0) -- Attrib.thms >> |
30515 | 411 |
(fn (n, ths) => K (m n ths)); |
412 |
||
413 |
||
27813
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
414 |
(* outer parser *) |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
415 |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
416 |
fun is_symid_meth s = |
36959
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents:
36950
diff
changeset
|
417 |
s <> "|" andalso s <> "?" andalso s <> "+" andalso Token.ident_or_symbolic s; |
27813
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
418 |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
419 |
local |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
420 |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
421 |
fun meth4 x = |
36950 | 422 |
(Parse.position (Parse.xname >> rpair []) >> (Source o Args.src) || |
423 |
Parse.$$$ "(" |-- Parse.!!! (meth0 --| Parse.$$$ ")")) x |
|
27813
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
424 |
and meth3 x = |
36950 | 425 |
(meth4 --| Parse.$$$ "?" >> Try || |
426 |
meth4 --| Parse.$$$ "+" >> Repeat1 || |
|
427 |
meth4 -- (Parse.$$$ "[" |-- Scan.optional Parse.nat 1 --| Parse.$$$ "]") >> (SelectGoals o swap) || |
|
27813
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
428 |
meth4) x |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
429 |
and meth2 x = |
36950 | 430 |
(Parse.position (Parse.xname -- Args.parse1 is_symid_meth) >> (Source o Args.src) || |
27813
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
431 |
meth3) x |
36950 | 432 |
and meth1 x = (Parse.enum1 "," meth2 >> (fn [m] => m | ms => Then ms)) x |
433 |
and meth0 x = (Parse.enum1 "|" meth1 >> (fn [m] => m | ms => Orelse ms)) x; |
|
27813
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
434 |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
435 |
in val parse = meth3 end; |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
436 |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
437 |
|
18708 | 438 |
(* theory setup *) |
5824 | 439 |
|
26463 | 440 |
val _ = Context.>> (Context.map_theory |
30515 | 441 |
(setup (Binding.name "fail") (Scan.succeed (K fail)) "force failure" #> |
442 |
setup (Binding.name "succeed") (Scan.succeed (K succeed)) "succeed" #> |
|
443 |
setup (Binding.name "-") (Scan.succeed (K insert_facts)) |
|
444 |
"do nothing (insert current facts only)" #> |
|
445 |
setup (Binding.name "insert") (Attrib.thms >> (K o insert)) |
|
446 |
"insert theorems, ignoring facts (improper)" #> |
|
447 |
setup (Binding.name "intro") (Attrib.thms >> (K o intro)) |
|
448 |
"repeatedly apply introduction rules" #> |
|
449 |
setup (Binding.name "elim") (Attrib.thms >> (K o elim)) |
|
450 |
"repeatedly apply elimination rules" #> |
|
451 |
setup (Binding.name "unfold") (Attrib.thms >> unfold_meth) "unfold definitions" #> |
|
452 |
setup (Binding.name "fold") (Attrib.thms >> fold_meth) "fold definitions" #> |
|
453 |
setup (Binding.name "atomize") (Args.mode "full" >> (K o atomize)) |
|
454 |
"present local premises as object-level statements" #> |
|
455 |
setup (Binding.name "rule") (Attrib.thms >> some_rule) "apply some intro/elim rule" #> |
|
456 |
setup (Binding.name "erule") (xrule_meth erule) "apply rule in elimination manner (improper)" #> |
|
457 |
setup (Binding.name "drule") (xrule_meth drule) "apply rule in destruct manner (improper)" #> |
|
458 |
setup (Binding.name "frule") (xrule_meth frule) "apply rule in forward manner (improper)" #> |
|
459 |
setup (Binding.name "this") (Scan.succeed (K this)) "apply current facts as rules" #> |
|
460 |
setup (Binding.name "fact") (Attrib.thms >> fact) "composition by facts from context" #> |
|
461 |
setup (Binding.name "assumption") (Scan.succeed assumption) |
|
462 |
"proof by assumption, preferring facts" #> |
|
463 |
setup (Binding.name "rename_tac") (Args.goal_spec -- Scan.lift (Scan.repeat1 Args.name) >> |
|
464 |
(fn (quant, xs) => K (SIMPLE_METHOD'' quant (Tactic.rename_tac xs)))) |
|
465 |
"rename parameters of goal" #> |
|
36950 | 466 |
setup (Binding.name "rotate_tac") (Args.goal_spec -- Scan.lift (Scan.optional Parse.int 1) >> |
30515 | 467 |
(fn (quant, i) => K (SIMPLE_METHOD'' quant (Tactic.rotate_tac i)))) |
468 |
"rotate assumptions of goal" #> |
|
37198
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
469 |
setup (Binding.name "tactic") (Scan.lift Args.name_source_position >> tactic) |
30515 | 470 |
"ML tactic as proof method" #> |
37198
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents:
36959
diff
changeset
|
471 |
setup (Binding.name "raw_tactic") (Scan.lift Args.name_source_position >> raw_tactic) |
30515 | 472 |
"ML tactic as raw proof method")); |
5824 | 473 |
|
474 |
||
16145 | 475 |
(*final declarations of this structure!*) |
476 |
val unfold = unfold_meth; |
|
477 |
val fold = fold_meth; |
|
478 |
||
5824 | 479 |
end; |
480 |
||
32193
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents:
32091
diff
changeset
|
481 |
structure Basic_Method: BASIC_METHOD = Method; |
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents:
32091
diff
changeset
|
482 |
open Basic_Method; |
30508
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents:
30466
diff
changeset
|
483 |
|
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents:
30466
diff
changeset
|
484 |
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
|
485 |
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
|
486 |
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
|
487 |
val METHOD = Method.METHOD; |
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents:
30466
diff
changeset
|
488 |
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
|
489 |
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
|
490 |
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
|
491 |