author | kleing |
Mon, 19 Sep 2005 14:20:45 +0200 | |
changeset 17483 | c6005bfc1630 |
parent 17412 | e26cb20ef0cc |
child 17496 | 26535df536ae |
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 |
val print_methods: theory -> unit |
17110 | 15 |
val Method: bstring -> (Args.src -> ProofContext.context -> method) -> string -> unit |
5824 | 16 |
end; |
17 |
||
18 |
signature METHOD = |
|
19 |
sig |
|
20 |
include BASIC_METHOD |
|
7419 | 21 |
val multi_resolve: thm list -> thm -> thm Seq.seq |
22 |
val multi_resolves: thm list -> thm list -> thm Seq.seq |
|
17110 | 23 |
val apply: method -> thm list -> RuleCases.tactic; |
24 |
val RAW_METHOD_CASES: (thm list -> RuleCases.tactic) -> method |
|
25 |
val RAW_METHOD: (thm list -> tactic) -> method |
|
26 |
val METHOD_CASES: (thm list -> RuleCases.tactic) -> method |
|
27 |
val METHOD: (thm list -> tactic) -> method |
|
28 |
val fail: method |
|
29 |
val succeed: method |
|
30 |
val insert_tac: thm list -> int -> tactic |
|
31 |
val insert: thm list -> method |
|
32 |
val insert_facts: method |
|
33 |
val SIMPLE_METHOD: tactic -> method |
|
34 |
val SIMPLE_METHOD': ((int -> tactic) -> tactic) -> (int -> tactic) -> method |
|
35 |
val defer: int option -> method |
|
36 |
val prefer: int -> method |
|
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
37 |
val cheating: bool -> ProofContext.context -> method |
17110 | 38 |
val intro: thm list -> method |
39 |
val elim: thm list -> method |
|
40 |
val unfold: thm list -> method |
|
41 |
val fold: thm list -> method |
|
42 |
val atomize: bool -> method |
|
43 |
val this: method |
|
44 |
val assumption: ProofContext.context -> method |
|
45 |
val close: bool -> ProofContext.context -> method |
|
46 |
val trace: ProofContext.context -> thm list -> unit |
|
6091 | 47 |
val rule_tac: thm list -> thm list -> int -> tactic |
17110 | 48 |
val some_rule_tac: thm list -> ProofContext.context -> thm list -> int -> tactic |
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 |
|
53 |
val rules_tac: ProofContext.context -> int option -> int -> tactic |
|
54 |
val bires_inst_tac: bool -> ProofContext.context -> (indexname * string) list -> |
|
55 |
thm -> int -> tactic |
|
56 |
val set_tactic: (ProofContext.context -> thm list -> tactic) -> unit |
|
57 |
val tactic: string -> ProofContext.context -> method |
|
58 |
type src |
|
5824 | 59 |
datatype text = |
17110 | 60 |
Basic of (ProofContext.context -> method) | |
15703 | 61 |
Source of src | |
5824 | 62 |
Then of text list | |
63 |
Orelse of text list | |
|
64 |
Try of text | |
|
65 |
Repeat1 of text |
|
17110 | 66 |
val default_text: text |
67 |
val this_text: text |
|
68 |
val done_text: text |
|
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
69 |
val sorry_text: bool -> text |
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
70 |
val finish_text: text option * bool -> text |
17110 | 71 |
exception METHOD_FAIL of (string * Position.T) * exn |
72 |
val method: theory -> src -> ProofContext.context -> method |
|
73 |
val add_methods: (bstring * (src -> ProofContext.context -> method) * string) list |
|
74 |
-> theory -> theory |
|
75 |
val add_method: bstring * (src -> ProofContext.context -> method) * string |
|
76 |
-> theory -> theory |
|
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
77 |
val method_setup: bstring * string * string -> theory -> theory |
17110 | 78 |
val syntax: (ProofContext.context * Args.T list -> 'a * (ProofContext.context * Args.T list)) |
79 |
-> src -> ProofContext.context -> ProofContext.context * 'a |
|
80 |
val simple_args: (Args.T list -> 'a * Args.T list) |
|
81 |
-> ('a -> ProofContext.context -> method) -> src -> ProofContext.context -> method |
|
82 |
val ctxt_args: (ProofContext.context -> method) -> src -> ProofContext.context -> method |
|
83 |
val no_args: method -> src -> ProofContext.context -> method |
|
84 |
type modifier |
|
85 |
val sectioned_args: (ProofContext.context * Args.T list -> |
|
86 |
'a * (ProofContext.context * Args.T list)) -> |
|
87 |
(Args.T list -> modifier * Args.T list) list -> |
|
88 |
('a -> ProofContext.context -> 'b) -> src -> ProofContext.context -> 'b |
|
89 |
val bang_sectioned_args: |
|
90 |
(Args.T list -> modifier * Args.T list) list -> |
|
91 |
(thm list -> ProofContext.context -> 'a) -> src -> ProofContext.context -> 'a |
|
92 |
val bang_sectioned_args': |
|
93 |
(Args.T list -> modifier * Args.T list) list -> |
|
94 |
(ProofContext.context * Args.T list -> 'a * (ProofContext.context * Args.T list)) -> |
|
95 |
('a -> thm list -> ProofContext.context -> 'b) -> src -> ProofContext.context -> 'b |
|
96 |
val only_sectioned_args: |
|
97 |
(Args.T list -> modifier * Args.T list) list -> |
|
98 |
(ProofContext.context -> 'a) -> src -> ProofContext.context -> 'a |
|
99 |
val thms_ctxt_args: (thm list -> ProofContext.context -> 'a) -> src -> |
|
100 |
ProofContext.context -> 'a |
|
101 |
val thms_args: (thm list -> 'a) -> src -> ProofContext.context -> 'a |
|
102 |
val thm_args: (thm -> 'a) -> src -> ProofContext.context -> 'a |
|
9539 | 103 |
val goal_args: (Args.T list -> 'a * Args.T list) -> ('a -> int -> tactic) |
17110 | 104 |
-> src -> ProofContext.context -> method |
105 |
val goal_args': (ProofContext.context * Args.T list -> |
|
106 |
'a * (ProofContext.context * Args.T list)) |
|
107 |
-> ('a -> int -> tactic) -> src -> ProofContext.context -> method |
|
108 |
val goal_args_ctxt: (Args.T list -> 'a * Args.T list) -> |
|
109 |
(ProofContext.context -> 'a -> int -> tactic) -> src -> ProofContext.context -> method |
|
110 |
val goal_args_ctxt': (ProofContext.context * Args.T list -> |
|
111 |
'a * (ProofContext.context * Args.T list)) -> |
|
112 |
(ProofContext.context -> 'a -> int -> tactic) -> src -> ProofContext.context -> method |
|
5824 | 113 |
end; |
114 |
||
115 |
structure Method: METHOD = |
|
116 |
struct |
|
117 |
||
17110 | 118 |
(** generic tools **) |
119 |
||
120 |
(* goal addressing *) |
|
121 |
||
122 |
fun FINDGOAL tac st = |
|
123 |
let fun find i n = if i > n then Seq.fail else Seq.APPEND (tac i, find (i + 1) n) |
|
124 |
in find 1 (Thm.nprems_of st) st end; |
|
125 |
||
126 |
fun HEADGOAL tac = tac 1; |
|
127 |
||
128 |
||
129 |
(* multi_resolve *) |
|
130 |
||
131 |
local |
|
132 |
||
133 |
fun res th i rule = |
|
134 |
Thm.biresolution false [(false, th)] i rule handle THM _ => Seq.empty; |
|
135 |
||
136 |
fun multi_res _ [] rule = Seq.single rule |
|
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
137 |
| multi_res i (th :: ths) rule = Seq.maps (res th i) (multi_res (i + 1) ths rule); |
17110 | 138 |
|
139 |
in |
|
140 |
||
141 |
val multi_resolve = multi_res 1; |
|
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
142 |
fun multi_resolves facts rules = Seq.maps (multi_resolve facts) (Seq.of_list rules); |
17110 | 143 |
|
144 |
end; |
|
145 |
||
15703 | 146 |
|
5824 | 147 |
|
12324
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
wenzelm
parents:
12311
diff
changeset
|
148 |
(** proof methods **) |
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
wenzelm
parents:
12311
diff
changeset
|
149 |
|
17110 | 150 |
(* datatype method *) |
11731 | 151 |
|
17110 | 152 |
datatype method = Method of thm list -> RuleCases.tactic; |
11731 | 153 |
|
17110 | 154 |
fun apply (Method m) = m; |
11731 | 155 |
|
17110 | 156 |
val RAW_METHOD_CASES = Method; |
11731 | 157 |
|
17110 | 158 |
fun RAW_METHOD tac = RAW_METHOD_CASES (NO_CASES o tac); |
12144 | 159 |
|
17110 | 160 |
fun METHOD_CASES tac = RAW_METHOD_CASES (fn facts => |
161 |
Seq.THEN (TRY Tactic.conjunction_tac, tac facts)); |
|
8372 | 162 |
|
17110 | 163 |
fun METHOD tac = RAW_METHOD (fn facts => |
164 |
TRY Tactic.conjunction_tac THEN tac facts); |
|
5824 | 165 |
|
166 |
val fail = METHOD (K no_tac); |
|
167 |
val succeed = METHOD (K all_tac); |
|
168 |
||
169 |
||
17110 | 170 |
(* insert facts *) |
7419 | 171 |
|
172 |
local |
|
5824 | 173 |
|
6981 | 174 |
fun cut_rule_tac raw_rule = |
175 |
let |
|
176 |
val rule = Drule.forall_intr_vars raw_rule; |
|
177 |
val revcut_rl = Drule.incr_indexes_wrt [] [] [] [rule] Drule.revcut_rl; |
|
7555 | 178 |
in Tactic.rtac (rule COMP revcut_rl) end; |
6981 | 179 |
|
7419 | 180 |
in |
5824 | 181 |
|
7419 | 182 |
fun insert_tac [] i = all_tac |
183 |
| insert_tac facts i = EVERY (map (fn th => cut_rule_tac th i) facts); |
|
6981 | 184 |
|
7555 | 185 |
val insert_facts = METHOD (ALLGOALS o insert_tac); |
7664 | 186 |
fun insert thms = METHOD (fn _ => ALLGOALS (insert_tac thms)); |
7419 | 187 |
|
9706 | 188 |
fun SIMPLE_METHOD tac = METHOD (fn facts => ALLGOALS (insert_tac facts) THEN tac); |
189 |
fun SIMPLE_METHOD' quant tac = METHOD (fn facts => quant (insert_tac facts THEN' tac)); |
|
190 |
||
12324
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
wenzelm
parents:
12311
diff
changeset
|
191 |
end; |
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
wenzelm
parents:
12311
diff
changeset
|
192 |
|
9706 | 193 |
|
17110 | 194 |
(* shuffle subgoals *) |
195 |
||
196 |
fun prefer i = METHOD (K (Tactic.defer_tac i THEN PRIMITIVE (Thm.permute_prems 0 ~1))); |
|
197 |
fun defer opt_i = METHOD (K (Tactic.defer_tac (if_none opt_i 1))); |
|
198 |
||
199 |
||
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
200 |
(* cheating *) |
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
201 |
|
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
202 |
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
|
203 |
(SkipProof.cheat_tac (ProofContext.theory_of ctxt)))); |
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
204 |
|
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
205 |
|
17110 | 206 |
(* unfold intro/elim rules *) |
207 |
||
208 |
fun intro ths = SIMPLE_METHOD' HEADGOAL (CHANGED_PROP o REPEAT_ALL_NEW (Tactic.match_tac ths)); |
|
209 |
fun elim ths = SIMPLE_METHOD' HEADGOAL (CHANGED_PROP o REPEAT_ALL_NEW (Tactic.ematch_tac ths)); |
|
210 |
||
211 |
||
12384 | 212 |
(* unfold/fold definitions *) |
213 |
||
16145 | 214 |
fun unfold_meth ths = SIMPLE_METHOD (CHANGED_PROP (rewrite_goals_tac ths)); |
215 |
fun fold_meth ths = SIMPLE_METHOD (CHANGED_PROP (fold_goals_tac ths)); |
|
6532 | 216 |
|
12384 | 217 |
|
12829 | 218 |
(* atomize rule statements *) |
219 |
||
220 |
fun atomize false = SIMPLE_METHOD' HEADGOAL (CHANGED_PROP o ObjectLogic.atomize_tac) |
|
221 |
| atomize true = RAW_METHOD (K (HEADGOAL (CHANGED_PROP o ObjectLogic.full_atomize_tac))); |
|
222 |
||
223 |
||
17110 | 224 |
(* this -- apply facts directly *) |
12384 | 225 |
|
17110 | 226 |
val this = METHOD (EVERY o map (HEADGOAL o Tactic.rtac)); |
9484 | 227 |
|
228 |
||
17110 | 229 |
(* assumption *) |
7419 | 230 |
|
231 |
local |
|
232 |
||
17110 | 233 |
fun asm_tac ths = |
234 |
foldr (op APPEND') (K no_tac) (map (fn th => Tactic.rtac th THEN_ALL_NEW assume_tac) ths); |
|
235 |
||
236 |
val refl_tac = SUBGOAL (fn (prop, i) => |
|
237 |
if can Logic.dest_equals (Logic.strip_assums_concl prop) |
|
238 |
then Tactic.rtac Drule.reflexive_thm i else no_tac); |
|
7419 | 239 |
|
17110 | 240 |
fun assm_tac ctxt = |
241 |
assume_tac APPEND' |
|
242 |
asm_tac (ProofContext.prems_of ctxt) APPEND' |
|
243 |
refl_tac; |
|
244 |
||
245 |
fun assumption_tac ctxt [] = assm_tac ctxt |
|
246 |
| assumption_tac _ [fact] = asm_tac [fact] |
|
247 |
| assumption_tac _ _ = K no_tac; |
|
7419 | 248 |
|
249 |
in |
|
250 |
||
17110 | 251 |
fun assumption ctxt = METHOD (HEADGOAL o assumption_tac ctxt); |
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
252 |
fun close immed ctxt = METHOD (K |
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
253 |
(FILTER Thm.no_prems ((if immed then ALLGOALS (assm_tac ctxt) else all_tac) THEN flexflex_tac))); |
7419 | 254 |
|
255 |
end; |
|
256 |
||
257 |
||
17110 | 258 |
(* rule etc. -- single-step refinements *) |
12347 | 259 |
|
17110 | 260 |
val trace_rules = ref false; |
12347 | 261 |
|
17110 | 262 |
fun trace ctxt rules = |
263 |
conditional (! trace_rules andalso not (null rules)) (fn () => |
|
264 |
Pretty.big_list "rules:" (map (ProofContext.pretty_thm ctxt) rules) |
|
265 |
|> Pretty.string_of |> tracing); |
|
12347 | 266 |
|
267 |
local |
|
268 |
||
10541
fdec07d4f047
resolveq(_cases)_tac moved to HOL/Tools/induct_method.ML;
wenzelm
parents:
10529
diff
changeset
|
269 |
fun gen_rule_tac tac rules [] i st = tac rules i st |
12324
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
wenzelm
parents:
12311
diff
changeset
|
270 |
| gen_rule_tac tac rules facts i st = |
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
271 |
Seq.maps (fn rule => (tac o single) rule i st) (multi_resolves facts rules); |
7130 | 272 |
|
10744 | 273 |
fun gen_arule_tac tac j rules facts = |
274 |
EVERY' (gen_rule_tac tac rules facts :: replicate j Tactic.assume_tac); |
|
275 |
||
11785 | 276 |
fun gen_some_rule_tac tac arg_rules ctxt facts = SUBGOAL (fn (goal, i) => |
277 |
let |
|
278 |
val rules = |
|
279 |
if not (null arg_rules) then arg_rules |
|
15570 | 280 |
else List.concat (ContextRules.find_rules false facts goal ctxt) |
12055 | 281 |
in trace ctxt rules; tac rules facts i end); |
10309 | 282 |
|
10744 | 283 |
fun meth tac x = METHOD (HEADGOAL o tac x); |
284 |
fun meth' tac x y = METHOD (HEADGOAL o tac x y); |
|
8220 | 285 |
|
7419 | 286 |
in |
287 |
||
10744 | 288 |
val rule_tac = gen_rule_tac Tactic.resolve_tac; |
289 |
val rule = meth rule_tac; |
|
290 |
val some_rule_tac = gen_some_rule_tac rule_tac; |
|
291 |
val some_rule = meth' some_rule_tac; |
|
292 |
||
293 |
val erule = meth' (gen_arule_tac Tactic.eresolve_tac); |
|
294 |
val drule = meth' (gen_arule_tac Tactic.dresolve_tac); |
|
295 |
val frule = meth' (gen_arule_tac Tactic.forward_tac); |
|
5824 | 296 |
|
7419 | 297 |
end; |
298 |
||
299 |
||
17110 | 300 |
(* rules -- intuitionistic proof search *) |
301 |
||
302 |
local |
|
303 |
||
304 |
val remdups_tac = SUBGOAL (fn (g, i) => |
|
305 |
let val prems = Logic.strip_assums_hyp g in |
|
306 |
REPEAT_DETERM_N (length prems - length (gen_distinct op aconv prems)) |
|
307 |
(Tactic.ematch_tac [Drule.remdups_rl] i THEN Tactic.eq_assume_tac i) |
|
308 |
end); |
|
309 |
||
310 |
fun REMDUPS tac = tac THEN_ALL_NEW remdups_tac; |
|
311 |
||
312 |
fun gen_eq_set e s1 s2 = |
|
313 |
length s1 = length s2 andalso |
|
314 |
gen_subset e (s1, s2) andalso gen_subset e (s2, s1); |
|
315 |
||
316 |
val bires_tac = Tactic.biresolution_from_nets_tac ContextRules.orderlist; |
|
317 |
||
318 |
fun safe_step_tac ctxt = |
|
319 |
ContextRules.Swrap ctxt |
|
320 |
(eq_assume_tac ORELSE' |
|
321 |
bires_tac true (ContextRules.netpair_bang ctxt)); |
|
8195 | 322 |
|
17110 | 323 |
fun unsafe_step_tac ctxt = |
324 |
ContextRules.wrap ctxt |
|
325 |
(assume_tac APPEND' |
|
326 |
bires_tac false (ContextRules.netpair_bang ctxt) APPEND' |
|
327 |
bires_tac false (ContextRules.netpair ctxt)); |
|
328 |
||
329 |
fun step_tac ctxt i = |
|
330 |
REPEAT_DETERM1 (REMDUPS (safe_step_tac ctxt) i) ORELSE |
|
331 |
REMDUPS (unsafe_step_tac ctxt) i; |
|
332 |
||
333 |
fun intprover_tac ctxt gs d lim = SUBGOAL (fn (g, i) => if d > lim then no_tac else |
|
334 |
let |
|
335 |
val ps = Logic.strip_assums_hyp g; |
|
336 |
val c = Logic.strip_assums_concl g; |
|
337 |
in |
|
338 |
if gen_mem (fn ((ps1, c1), (ps2, c2)) => |
|
339 |
c1 aconv c2 andalso gen_eq_set op aconv ps1 ps2) ((ps, c), gs) then no_tac |
|
340 |
else (step_tac ctxt THEN_ALL_NEW intprover_tac ctxt ((ps, c) :: gs) (d + 1) lim) i |
|
341 |
end); |
|
342 |
||
343 |
in |
|
344 |
||
345 |
fun rules_tac ctxt opt_lim = |
|
346 |
SELECT_GOAL (DEEPEN (2, if_none opt_lim 20) (intprover_tac ctxt [] 0) 4 1); |
|
347 |
||
348 |
end; |
|
8195 | 349 |
|
350 |
||
17110 | 351 |
(* rule_tac etc. -- refer to dynamic goal state!! *) |
12119 | 352 |
|
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
353 |
fun bires_inst_tac bires_flag ctxt insts thm = |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
354 |
let |
17110 | 355 |
val thy = ProofContext.theory_of ctxt; |
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
356 |
(* Separate type and term insts *) |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
357 |
fun has_type_var ((x, _), _) = (case Symbol.explode x of |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
358 |
"'"::cs => true | cs => false); |
15570 | 359 |
val Tinsts = List.filter has_type_var insts; |
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
360 |
val tinsts = filter_out has_type_var insts; |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
361 |
(* Tactic *) |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
362 |
fun tac i st = |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
363 |
let |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
364 |
(* Preprocess state: extract environment information: |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
365 |
- variables and their types |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
366 |
- type variables and their sorts |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
367 |
- parameters and their types *) |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
368 |
val (types, sorts) = types_sorts st; |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
369 |
(* Process type insts: Tinsts_env *) |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
370 |
fun absent xi = error |
14718 | 371 |
("No such variable in theorem: " ^ Syntax.string_of_vname xi); |
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
372 |
val (rtypes, rsorts) = types_sorts thm; |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
373 |
fun readT (xi, s) = |
15531 | 374 |
let val S = case rsorts xi of SOME S => S | NONE => absent xi; |
17110 | 375 |
val T = Sign.read_typ (thy, sorts) s; |
15798
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
berghofe
parents:
15713
diff
changeset
|
376 |
val U = TVar (xi, S); |
17110 | 377 |
in if Sign.typ_instance thy (T, U) then (U, T) |
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
378 |
else error |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
379 |
("Instantiation of " ^ Syntax.string_of_vname xi ^ " fails") |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
380 |
end; |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
381 |
val Tinsts_env = map readT Tinsts; |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
382 |
(* Preprocess rule: extract vars and their types, apply Tinsts *) |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
383 |
fun get_typ xi = |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
384 |
(case rtypes xi of |
15798
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
berghofe
parents:
15713
diff
changeset
|
385 |
SOME T => typ_subst_atomic Tinsts_env T |
15531 | 386 |
| NONE => absent xi); |
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
387 |
val (xis, ss) = Library.split_list tinsts; |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
388 |
val Ts = map get_typ xis; |
14718 | 389 |
val (_, _, Bi, _) = dest_state(st,i) |
390 |
val params = Logic.strip_params Bi |
|
391 |
(* params of subgoal i as string typ pairs *) |
|
392 |
val params = rev(Term.rename_wrt_term Bi params) |
|
393 |
(* as they are printed: bound variables with *) |
|
14508
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14215
diff
changeset
|
394 |
(* the same name are renamed during printing *) |
17314 | 395 |
fun types' (a, ~1) = (case AList.lookup (op =) params a of |
15531 | 396 |
NONE => types (a, ~1) |
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
397 |
| some => some) |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
398 |
| types' xi = types xi; |
15973 | 399 |
fun internal x = is_some (types' (x, ~1)); |
15703 | 400 |
val used = Drule.add_used thm (Drule.add_used st []); |
14718 | 401 |
val (ts, envT) = |
402 |
ProofContext.read_termTs_schematic ctxt internal types' sorts used (ss ~~ Ts); |
|
15798
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
berghofe
parents:
15713
diff
changeset
|
403 |
val envT' = map (fn (ixn, T) => |
15973 | 404 |
(TVar (ixn, the (rsorts ixn)), T)) envT @ Tinsts_env; |
14718 | 405 |
val cenv = |
406 |
map |
|
407 |
(fn (xi, t) => |
|
17110 | 408 |
pairself (Thm.cterm_of thy) (Var (xi, fastype_of t), t)) |
14718 | 409 |
(gen_distinct |
410 |
(fn ((x1, t1), (x2, t2)) => x1 = x2 andalso t1 aconv t2) |
|
411 |
(xis ~~ ts)); |
|
412 |
(* Lift and instantiate rule *) |
|
413 |
val {maxidx, ...} = rep_thm st; |
|
414 |
val paramTs = map #2 params |
|
415 |
and inc = maxidx+1 |
|
416 |
fun liftvar (Var ((a,j), T)) = |
|
16876 | 417 |
Var((a, j+inc), paramTs ---> Logic.incr_tvar inc T) |
14718 | 418 |
| liftvar t = raise TERM("Variable expected", [t]); |
419 |
fun liftterm t = list_abs_free |
|
420 |
(params, Logic.incr_indexes(paramTs,inc) t) |
|
421 |
fun liftpair (cv,ct) = |
|
422 |
(cterm_fun liftvar cv, cterm_fun liftterm ct) |
|
17110 | 423 |
val lifttvar = pairself (ctyp_of thy o Logic.incr_tvar inc); |
14718 | 424 |
val rule = Drule.instantiate |
15798
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
berghofe
parents:
15713
diff
changeset
|
425 |
(map lifttvar envT', map liftpair cenv) |
14718 | 426 |
(lift_rule (st, i) thm) |
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
427 |
in |
14718 | 428 |
if i > nprems_of st then no_tac st |
429 |
else st |> |
|
430 |
compose_tac (bires_flag, rule, nprems_of thm) i |
|
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
431 |
end |
14718 | 432 |
handle TERM (msg,_) => (warning msg; no_tac st) |
433 |
| THM (msg,_,_) => (warning msg; no_tac st); |
|
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
434 |
in |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
435 |
tac |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
436 |
end; |
8238 | 437 |
|
17110 | 438 |
local |
439 |
||
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
440 |
fun gen_inst _ tac _ (quant, ([], thms)) = |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
441 |
METHOD (fn facts => quant (insert_tac facts THEN' tac thms)) |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
442 |
| gen_inst inst_tac _ ctxt (quant, (insts, [thm])) = |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
443 |
METHOD (fn facts => |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
444 |
quant (insert_tac facts THEN' inst_tac ctxt insts thm)) |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
445 |
| gen_inst _ _ _ _ = error "Cannot have instantiations with multiple rules"; |
14718 | 446 |
|
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
447 |
(* Preserve Var indexes of rl; increment revcut_rl instead. |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
448 |
Copied from tactic.ML *) |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
449 |
fun make_elim_preserve rl = |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
450 |
let val {maxidx,...} = rep_thm rl |
17110 | 451 |
fun cvar xi = cterm_of ProtoPure.thy (Var(xi,propT)); |
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
452 |
val revcut_rl' = |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
453 |
instantiate ([], [(cvar("V",0), cvar("V",maxidx+1)), |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
454 |
(cvar("W",0), cvar("W",maxidx+1))]) revcut_rl |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
455 |
val arg = (false, rl, nprems_of rl) |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
456 |
val [th] = Seq.list_of (bicompose false arg 1 revcut_rl') |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
457 |
in th end |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
458 |
handle Bind => raise THM("make_elim_preserve", 1, [rl]); |
8238 | 459 |
|
17110 | 460 |
in |
461 |
||
462 |
val res_inst_meth = gen_inst (bires_inst_tac false) Tactic.resolve_tac; |
|
463 |
||
464 |
val eres_inst_meth = gen_inst (bires_inst_tac true) Tactic.eresolve_tac; |
|
465 |
||
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
466 |
val cut_inst_meth = |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
467 |
gen_inst |
17110 | 468 |
(fn ctxt => fn insts => bires_inst_tac false ctxt insts o make_elim_preserve) |
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
469 |
Tactic.cut_rules_tac; |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
470 |
|
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
471 |
val dres_inst_meth = |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
472 |
gen_inst |
17110 | 473 |
(fn ctxt => fn insts => bires_inst_tac true ctxt insts o make_elim_preserve) |
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
474 |
Tactic.dresolve_tac; |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
475 |
|
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
476 |
val forw_inst_meth = |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
477 |
gen_inst |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
478 |
(fn ctxt => fn insts => fn rule => |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
479 |
bires_inst_tac false ctxt insts (make_elim_preserve rule) THEN' |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
480 |
assume_tac) |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
481 |
Tactic.forward_tac; |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
482 |
|
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
483 |
fun subgoal_tac ctxt sprop = |
17110 | 484 |
DETERM o bires_inst_tac false ctxt [(("psi", 0), sprop)] cut_rl; |
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
485 |
|
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
486 |
fun subgoals_tac ctxt sprops = EVERY' (map (subgoal_tac ctxt) sprops); |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
487 |
|
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
488 |
fun thin_tac ctxt s = |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
489 |
bires_inst_tac true ctxt [(("V", 0), s)] thin_rl; |
8238 | 490 |
|
17110 | 491 |
end; |
8329 | 492 |
|
493 |
||
8351 | 494 |
(* ML tactics *) |
495 |
||
17110 | 496 |
val tactic_ref = ref ((fn _ => raise Match): ProofContext.context -> thm list -> tactic); |
8351 | 497 |
fun set_tactic f = tactic_ref := f; |
498 |
||
499 |
fun tactic txt ctxt = METHOD (fn facts => |
|
9631 | 500 |
(Context.use_mltext |
17110 | 501 |
("let fun tactic (ctxt: ProofContext.context) (facts: thm list) : tactic = \ |
16500
09d43301b195
refl_tac: avoid failure of unification, i.e. confusing trace msg;
wenzelm
parents:
16448
diff
changeset
|
502 |
\let val thm = ProofContext.get_thm_closure ctxt o PureThy.Name\n\ |
09d43301b195
refl_tac: avoid failure of unification, i.e. confusing trace msg;
wenzelm
parents:
16448
diff
changeset
|
503 |
\ and thms = ProofContext.get_thms_closure ctxt o PureThy.Name in\n" |
09d43301b195
refl_tac: avoid failure of unification, i.e. confusing trace msg;
wenzelm
parents:
16448
diff
changeset
|
504 |
^ txt ^ |
09d43301b195
refl_tac: avoid failure of unification, i.e. confusing trace msg;
wenzelm
parents:
16448
diff
changeset
|
505 |
"\nend in Method.set_tactic tactic end") |
15531 | 506 |
false NONE; |
507 |
Context.setmp (SOME (ProofContext.theory_of ctxt)) (! tactic_ref ctxt) facts)); |
|
8351 | 508 |
|
509 |
||
5824 | 510 |
|
17110 | 511 |
(** method syntax **) |
512 |
||
513 |
(* method text *) |
|
514 |
||
515 |
type src = Args.src; |
|
5824 | 516 |
|
17110 | 517 |
datatype text = |
518 |
Basic of (ProofContext.context -> method) | |
|
519 |
Source of src | |
|
520 |
Then of text list | |
|
521 |
Orelse of text list | |
|
522 |
Try of text | |
|
523 |
Repeat1 of text; |
|
524 |
||
525 |
val default_text = Source (Args.src (("default", []), Position.none)); |
|
526 |
val this_text = Basic (K this); |
|
527 |
val done_text = Basic (K (SIMPLE_METHOD all_tac)); |
|
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
528 |
val sorry_text = Basic o cheating; |
17110 | 529 |
|
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
530 |
fun finish_text (NONE, immed) = Basic (close immed) |
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
531 |
| finish_text (SOME txt, immed) = Then [txt, Basic (close immed)]; |
17110 | 532 |
|
533 |
||
534 |
(* method definitions *) |
|
5824 | 535 |
|
16448 | 536 |
structure MethodsData = TheoryDataFun |
537 |
(struct |
|
5824 | 538 |
val name = "Isar/methods"; |
17110 | 539 |
type T = (((src -> ProofContext.context -> method) * string) * stamp) NameSpace.table; |
5824 | 540 |
|
16347 | 541 |
val empty = NameSpace.empty_table; |
6546 | 542 |
val copy = I; |
16448 | 543 |
val extend = I; |
544 |
fun merge _ tables = NameSpace.merge_tables eq_snd tables handle Symtab.DUPS dups => |
|
16347 | 545 |
error ("Attempt to merge different versions of method(s) " ^ commas_quote dups); |
5824 | 546 |
|
16347 | 547 |
fun print _ meths = |
5824 | 548 |
let |
549 |
fun prt_meth (name, ((_, comment), _)) = Pretty.block |
|
6849 | 550 |
[Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment]; |
5824 | 551 |
in |
16347 | 552 |
[Pretty.big_list "methods:" (map prt_meth (NameSpace.extern_table meths))] |
9222 | 553 |
|> Pretty.chunks |> Pretty.writeln |
5824 | 554 |
end; |
16448 | 555 |
end); |
5824 | 556 |
|
15801 | 557 |
val _ = Context.add_setup [MethodsData.init]; |
5824 | 558 |
val print_methods = MethodsData.print; |
7611 | 559 |
|
5824 | 560 |
|
561 |
(* get methods *) |
|
562 |
||
5916 | 563 |
exception METHOD_FAIL of (string * Position.T) * exn; |
564 |
||
5824 | 565 |
fun method thy = |
566 |
let |
|
16347 | 567 |
val (space, meths) = MethodsData.get thy; |
5884 | 568 |
fun meth src = |
569 |
let |
|
570 |
val ((raw_name, _), pos) = Args.dest_src src; |
|
571 |
val name = NameSpace.intern space raw_name; |
|
572 |
in |
|
17412 | 573 |
(case Symtab.lookup meths name of |
15531 | 574 |
NONE => error ("Unknown proof method: " ^ quote name ^ Position.str_of pos) |
575 |
| SOME ((mth, _), _) => transform_failure (curry METHOD_FAIL (name, pos)) (mth src)) |
|
5824 | 576 |
end; |
577 |
in meth end; |
|
578 |
||
579 |
||
17110 | 580 |
(* add method *) |
5824 | 581 |
|
582 |
fun add_methods raw_meths thy = |
|
583 |
let |
|
16145 | 584 |
val new_meths = raw_meths |> map (fn (name, f, comment) => |
16347 | 585 |
(name, ((f, comment), stamp ()))); |
5824 | 586 |
|
17110 | 587 |
fun add meths = NameSpace.extend_table (Sign.naming_of thy) (meths, new_meths) |
16347 | 588 |
handle Symtab.DUPS dups => |
589 |
error ("Duplicate declaration of method(s) " ^ commas_quote dups); |
|
590 |
in MethodsData.map add thy end; |
|
5824 | 591 |
|
9194 | 592 |
val add_method = add_methods o Library.single; |
593 |
||
5824 | 594 |
fun Method name meth cmt = Context.>> (add_methods [(name, meth, cmt)]); |
595 |
||
596 |
||
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
597 |
(* method_setup *) |
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
598 |
|
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
599 |
fun method_setup (name, txt, cmt) = |
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
600 |
Context.use_let |
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
601 |
"val thm = PureThy.get_thm_closure (Context.the_context ()) o PureThy.Name;\n\ |
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
602 |
\val thms = PureThy.get_thms_closure (Context.the_context ()) o PureThy.Name;\n\ |
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
603 |
\val method: bstring * (Method.src -> ProofContext.context -> Proof.method) * string" |
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
604 |
"Method.add_method method" |
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
605 |
("(" ^ Library.quote name ^ ", " ^ txt ^ ", " ^ Library.quote cmt ^ ")"); |
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
606 |
|
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
607 |
|
5884 | 608 |
|
17110 | 609 |
(** concrete syntax **) |
5824 | 610 |
|
5884 | 611 |
(* basic *) |
612 |
||
17110 | 613 |
fun syntax (scan: |
614 |
(ProofContext.context * Args.T list -> 'a * (ProofContext.context * Args.T list))) = |
|
5884 | 615 |
Args.syntax "method" scan; |
5824 | 616 |
|
17110 | 617 |
fun simple_args scan f src ctxt : method = |
8351 | 618 |
#2 (syntax (Scan.lift (scan >> (fn x => f x ctxt))) src ctxt); |
619 |
||
17110 | 620 |
fun ctxt_args (f: ProofContext.context -> method) src ctxt = |
8282 | 621 |
#2 (syntax (Scan.succeed (f ctxt)) src ctxt); |
7555 | 622 |
|
623 |
fun no_args m = ctxt_args (K m); |
|
5884 | 624 |
|
625 |
||
626 |
(* sections *) |
|
5824 | 627 |
|
17110 | 628 |
type modifier = (ProofContext.context -> ProofContext.context) * ProofContext.context attribute; |
7268 | 629 |
|
630 |
local |
|
631 |
||
8381 | 632 |
fun sect ss = Scan.first (map Scan.lift ss); |
5884 | 633 |
fun thms ss = Scan.unless (sect ss) Attrib.local_thms; |
15570 | 634 |
fun thmss ss = Scan.repeat (thms ss) >> List.concat; |
5884 | 635 |
|
7268 | 636 |
fun apply (f, att) (ctxt, ths) = Thm.applys_attributes ((f ctxt, ths), [att]); |
5824 | 637 |
|
7268 | 638 |
fun section ss = (sect ss -- thmss ss) :-- (fn (m, ths) => Scan.depend (fn ctxt => |
639 |
Scan.succeed (apply m (ctxt, ths)))) >> #2; |
|
5884 | 640 |
|
7601 | 641 |
fun sectioned args ss = args -- Scan.repeat (section ss); |
5884 | 642 |
|
7268 | 643 |
in |
5824 | 644 |
|
5884 | 645 |
fun sectioned_args args ss f src ctxt = |
8282 | 646 |
let val (ctxt', (x, _)) = syntax (sectioned args ss) src ctxt |
5921 | 647 |
in f x ctxt' end; |
5884 | 648 |
|
7601 | 649 |
fun bang_sectioned_args ss f = sectioned_args Args.bang_facts ss f; |
9777 | 650 |
fun bang_sectioned_args' ss scan f = |
651 |
sectioned_args (Args.bang_facts -- scan >> swap) ss (uncurry f); |
|
7601 | 652 |
fun only_sectioned_args ss f = sectioned_args (Scan.succeed ()) ss (fn () => f); |
7268 | 653 |
|
8093 | 654 |
fun thms_ctxt_args f = sectioned_args (thmss []) [] f; |
655 |
fun thms_args f = thms_ctxt_args (K o f); |
|
9706 | 656 |
fun thm_args f = thms_args (fn [thm] => f thm | _ => error "Single theorem expected"); |
5824 | 657 |
|
7268 | 658 |
end; |
659 |
||
5824 | 660 |
|
12347 | 661 |
(* rules syntax *) |
662 |
||
663 |
local |
|
664 |
||
665 |
val introN = "intro"; |
|
666 |
val elimN = "elim"; |
|
667 |
val destN = "dest"; |
|
668 |
val ruleN = "rule"; |
|
669 |
||
670 |
fun modifier name kind kind' att = |
|
15531 | 671 |
Args.$$$ name |-- (kind >> K NONE || kind' |-- Args.nat --| Args.colon >> SOME) |
17110 | 672 |
>> (pair (I: ProofContext.context -> ProofContext.context) o att); |
12347 | 673 |
|
674 |
val rules_modifiers = |
|
12384 | 675 |
[modifier destN Args.bang_colon Args.bang ContextRules.dest_bang_local, |
12350 | 676 |
modifier destN Args.colon (Scan.succeed ()) ContextRules.dest_local, |
677 |
modifier elimN Args.bang_colon Args.bang ContextRules.elim_bang_local, |
|
678 |
modifier elimN Args.colon (Scan.succeed ()) ContextRules.elim_local, |
|
679 |
modifier introN Args.bang_colon Args.bang ContextRules.intro_bang_local, |
|
680 |
modifier introN Args.colon (Scan.succeed ()) ContextRules.intro_local, |
|
681 |
Args.del -- Args.colon >> K (I, ContextRules.rule_del_local)]; |
|
12347 | 682 |
|
683 |
in |
|
684 |
||
685 |
fun rules_args m = bang_sectioned_args' rules_modifiers (Scan.lift (Scan.option Args.nat)) m; |
|
686 |
||
687 |
fun rules_meth n prems ctxt = METHOD (fn facts => |
|
12350 | 688 |
HEADGOAL (insert_tac (prems @ facts) THEN' ObjectLogic.atomize_tac THEN' rules_tac ctxt n)); |
12347 | 689 |
|
690 |
end; |
|
691 |
||
692 |
||
9539 | 693 |
(* tactic syntax *) |
8238 | 694 |
|
10744 | 695 |
fun nat_thms_args f = uncurry f oo |
696 |
(#2 oo syntax (Scan.lift (Scan.optional (Args.parens Args.nat) 0) -- Attrib.local_thmss)); |
|
697 |
||
8238 | 698 |
val insts = |
9539 | 699 |
Scan.optional |
9565
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
wenzelm
parents:
9539
diff
changeset
|
700 |
(Args.enum1 "and" (Scan.lift (Args.name -- (Args.$$$ "=" |-- Args.!!! Args.name))) --| |
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
wenzelm
parents:
9539
diff
changeset
|
701 |
Scan.lift (Args.$$$ "in")) [] -- Attrib.local_thmss; |
8238 | 702 |
|
12119 | 703 |
fun inst_args f src ctxt = f ctxt (#2 (syntax (Args.goal_spec HEADGOAL -- insts) src ctxt)); |
8537 | 704 |
|
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
705 |
val insts_var = |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
706 |
Scan.optional |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
707 |
(Args.enum1 "and" (Scan.lift (Args.var -- (Args.$$$ "=" |-- Args.!!! Args.name))) --| |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
708 |
Scan.lift (Args.$$$ "in")) [] -- Attrib.local_thmss; |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
709 |
|
17110 | 710 |
fun inst_args_var f src ctxt = |
711 |
f ctxt (#2 (syntax (Args.goal_spec HEADGOAL -- insts_var) src ctxt)); |
|
8537 | 712 |
|
12119 | 713 |
fun goal_args' args tac src ctxt = #2 (syntax (Args.goal_spec HEADGOAL -- args >> |
14215
ebf291f3b449
Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents:
14174
diff
changeset
|
714 |
(fn (quant, s) => SIMPLE_METHOD' quant (tac s))) src ctxt); |
8537 | 715 |
|
9539 | 716 |
fun goal_args args tac = goal_args' (Scan.lift args) tac; |
8238 | 717 |
|
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
718 |
fun goal_args_ctxt' args tac src ctxt = |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
719 |
#2 (syntax (Args.goal_spec HEADGOAL -- args >> |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
720 |
(fn (quant, s) => SIMPLE_METHOD' quant (tac ctxt s))) src ctxt); |
8238 | 721 |
|
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
722 |
fun goal_args_ctxt args tac = goal_args_ctxt' (Scan.lift args) tac; |
5824 | 723 |
|
14718 | 724 |
|
9539 | 725 |
(* misc tactic emulations *) |
726 |
||
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
727 |
val subgoal_meth = goal_args_ctxt (Scan.repeat1 Args.name) subgoals_tac; |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
728 |
val thin_meth = goal_args_ctxt Args.name thin_tac; |
9539 | 729 |
val rename_meth = goal_args (Scan.repeat1 Args.name) Tactic.rename_params_tac; |
9631 | 730 |
val rotate_meth = goal_args (Scan.optional Args.int 1) Tactic.rotate_tac; |
9539 | 731 |
|
732 |
||
5824 | 733 |
(* pure_methods *) |
734 |
||
735 |
val pure_methods = |
|
736 |
[("fail", no_args fail, "force failure"), |
|
737 |
("succeed", no_args succeed, "succeed"), |
|
9587 | 738 |
("-", no_args insert_facts, "do nothing (insert current facts only)"), |
9539 | 739 |
("insert", thms_args insert, "insert theorems, ignoring facts (improper)"), |
16145 | 740 |
("unfold", thms_args unfold_meth, "unfold definitions"), |
12384 | 741 |
("intro", thms_args intro, "repeatedly apply introduction rules"), |
742 |
("elim", thms_args elim, "repeatedly apply elimination rules"), |
|
16145 | 743 |
("fold", thms_args fold_meth, "fold definitions"), |
12829 | 744 |
("atomize", (atomize o #2) oo syntax (Args.mode "full"), |
11962 | 745 |
"present local premises as object-level statements"), |
12347 | 746 |
("rules", rules_args rules_meth, "apply many rules, including proof search"), |
12384 | 747 |
("rule", thms_ctxt_args some_rule, "apply some intro/elim rule"), |
10744 | 748 |
("erule", nat_thms_args erule, "apply rule in elimination manner (improper)"), |
749 |
("drule", nat_thms_args drule, "apply rule in destruct manner (improper)"), |
|
750 |
("frule", nat_thms_args frule, "apply rule in forward manner (improper)"), |
|
8195 | 751 |
("this", no_args this, "apply current facts as rules"), |
8238 | 752 |
("assumption", ctxt_args assumption, "proof by assumption, preferring facts"), |
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
753 |
("rule_tac", inst_args_var res_inst_meth, "apply rule (dynamic instantiation)"), |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
754 |
("erule_tac", inst_args_var eres_inst_meth, "apply rule in elimination manner (dynamic instantiation)"), |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
755 |
("drule_tac", inst_args_var dres_inst_meth, "apply rule in destruct manner (dynamic instantiation)"), |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
756 |
("frule_tac", inst_args_var forw_inst_meth, "apply rule in forward manner (dynamic instantiation)"), |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
757 |
("cut_tac", inst_args_var cut_inst_meth, "cut rule (dynamic instantiation)"), |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
758 |
("subgoal_tac", subgoal_meth, "insert subgoal (dynamic instantiation)"), |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
759 |
("thin_tac", thin_meth, "remove premise (dynamic instantiation)"), |
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13650
diff
changeset
|
760 |
("rename_tac", rename_meth, "rename parameters of goal (dynamic instantiation)"), |
9631 | 761 |
("rotate_tac", rotate_meth, "rotate assumptions of goal"), |
8351 | 762 |
("tactic", simple_args Args.name tactic, "ML tactic as proof method")]; |
5824 | 763 |
|
15801 | 764 |
val _ = Context.add_setup [add_methods pure_methods]; |
5824 | 765 |
|
766 |
||
16145 | 767 |
(*final declarations of this structure!*) |
768 |
val unfold = unfold_meth; |
|
769 |
val fold = fold_meth; |
|
770 |
||
5824 | 771 |
end; |
772 |
||
773 |
structure BasicMethod: BASIC_METHOD = Method; |
|
774 |
open BasicMethod; |