author | wenzelm |
Mon, 10 Mar 2014 16:30:07 +0100 | |
changeset 56029 | 8bedca4bd5a3 |
parent 56026 | 893fe12639bc |
child 56032 | b034b9f0fa2a |
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 METHOD = |
|
8 |
sig |
|
30508
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents:
30466
diff
changeset
|
9 |
type method |
32193
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents:
32091
diff
changeset
|
10 |
val apply: (Proof.context -> method) -> Proof.context -> thm list -> cases_tactic |
18227 | 11 |
val RAW_METHOD_CASES: (thm list -> cases_tactic) -> method |
17110 | 12 |
val RAW_METHOD: (thm list -> tactic) -> method |
18227 | 13 |
val METHOD_CASES: (thm list -> cases_tactic) -> method |
17110 | 14 |
val METHOD: (thm list -> tactic) -> method |
15 |
val fail: method |
|
16 |
val succeed: method |
|
17 |
val insert_tac: thm list -> int -> tactic |
|
18 |
val insert: thm list -> method |
|
19 |
val insert_facts: method |
|
20 |
val SIMPLE_METHOD: tactic -> method |
|
21592 | 21 |
val SIMPLE_METHOD': (int -> tactic) -> method |
22 |
val SIMPLE_METHOD'': ((int -> tactic) -> tactic) -> (int -> tactic) -> method |
|
52059 | 23 |
val cheating: Proof.context -> bool -> method |
17110 | 24 |
val intro: thm list -> method |
25 |
val elim: thm list -> method |
|
20289 | 26 |
val unfold: thm list -> Proof.context -> method |
27 |
val fold: thm list -> Proof.context -> method |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
28 |
val atomize: bool -> Proof.context -> method |
17110 | 29 |
val this: method |
20289 | 30 |
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
|
31 |
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
|
32 |
val all_assm_tac: Proof.context -> tactic |
20289 | 33 |
val assumption: Proof.context -> method |
46466
61c7214b4885
tuned signature, according to actual usage of these operations;
wenzelm
parents:
45375
diff
changeset
|
34 |
val rule_trace: bool Config.T |
20289 | 35 |
val trace: Proof.context -> thm list -> unit |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
36 |
val rule_tac: Proof.context -> thm list -> thm list -> int -> tactic |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
37 |
val some_rule_tac: Proof.context -> thm list -> thm list -> int -> tactic |
25270 | 38 |
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
|
39 |
val try_intros_tac: thm list -> thm list -> tactic |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
40 |
val rule: Proof.context -> thm list -> method |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
41 |
val erule: Proof.context -> int -> thm list -> method |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
42 |
val drule: Proof.context -> int -> thm list -> method |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
43 |
val frule: Proof.context -> int -> thm list -> method |
27235 | 44 |
val set_tactic: (thm list -> tactic) -> Proof.context -> Proof.context |
55828
42ac3cfb89f6
clarified language markup: added "delimited" property;
wenzelm
parents:
55795
diff
changeset
|
45 |
val tactic: Symbol_Pos.source -> Proof.context -> method |
42ac3cfb89f6
clarified language markup: added "delimited" property;
wenzelm
parents:
55795
diff
changeset
|
46 |
val raw_tactic: Symbol_Pos.source -> Proof.context -> method |
27729 | 47 |
type src = Args.src |
55765 | 48 |
type combinator_info |
49 |
val no_combinator_info: combinator_info |
|
5824 | 50 |
datatype text = |
55742
a989bdaf8121
modernized Method.check_name/check_source (with reports) vs. strict Method.the_method (without interning nor reports), e.g. relevant for semantic completion;
wenzelm
parents:
55709
diff
changeset
|
51 |
Source of src | |
32193
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents:
32091
diff
changeset
|
52 |
Basic of Proof.context -> method | |
55765 | 53 |
Then of combinator_info * text list | |
54 |
Orelse of combinator_info * text list | |
|
55 |
Try of combinator_info * text | |
|
56 |
Repeat1 of combinator_info * text | |
|
57 |
Select_Goals of combinator_info * int * text |
|
54883
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
54742
diff
changeset
|
58 |
val primitive_text: (Proof.context -> thm -> thm) -> text |
17857 | 59 |
val succeed_text: text |
17110 | 60 |
val default_text: text |
61 |
val this_text: text |
|
62 |
val done_text: text |
|
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
63 |
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
|
64 |
val finish_text: text option * bool -> text |
55742
a989bdaf8121
modernized Method.check_name/check_source (with reports) vs. strict Method.the_method (without interning nor reports), e.g. relevant for semantic completion;
wenzelm
parents:
55709
diff
changeset
|
65 |
val print_methods: Proof.context -> unit |
a989bdaf8121
modernized Method.check_name/check_source (with reports) vs. strict Method.the_method (without interning nor reports), e.g. relevant for semantic completion;
wenzelm
parents:
55709
diff
changeset
|
66 |
val check_name: Proof.context -> xstring * Position.T -> string |
55997
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents:
55917
diff
changeset
|
67 |
val check_src: Proof.context -> src -> src |
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents:
55917
diff
changeset
|
68 |
val method: Proof.context -> src -> Proof.context -> method |
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents:
55917
diff
changeset
|
69 |
val method_cmd: Proof.context -> src -> Proof.context -> method |
30512 | 70 |
val syntax: 'a context_parser -> src -> Proof.context -> 'a * Proof.context |
71 |
val setup: binding -> (Proof.context -> method) context_parser -> string -> theory -> theory |
|
55828
42ac3cfb89f6
clarified language markup: added "delimited" property;
wenzelm
parents:
55795
diff
changeset
|
72 |
val method_setup: bstring * Position.T -> Symbol_Pos.source -> string -> theory -> theory |
30540 | 73 |
type modifier = (Proof.context -> Proof.context) * attribute |
74 |
val section: modifier parser list -> thm list context_parser |
|
75 |
val sections: modifier parser list -> thm list list context_parser |
|
49889
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
wenzelm
parents:
49866
diff
changeset
|
76 |
type text_range = text * Position.range |
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
wenzelm
parents:
49866
diff
changeset
|
77 |
val text: text_range option -> text option |
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
wenzelm
parents:
49866
diff
changeset
|
78 |
val position: text_range option -> Position.T |
55795 | 79 |
val reports_of: text_range -> Position.report list |
80 |
val report: text_range -> unit |
|
55761
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents:
55742
diff
changeset
|
81 |
val parse: text_range parser |
5824 | 82 |
end; |
83 |
||
84 |
structure Method: METHOD = |
|
85 |
struct |
|
86 |
||
12324
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
wenzelm
parents:
12311
diff
changeset
|
87 |
(** proof methods **) |
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
wenzelm
parents:
12311
diff
changeset
|
88 |
|
17110 | 89 |
(* datatype method *) |
11731 | 90 |
|
18227 | 91 |
datatype method = Meth of thm list -> cases_tactic; |
11731 | 92 |
|
32193
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents:
32091
diff
changeset
|
93 |
fun apply meth ctxt = let val Meth m = meth ctxt in m end; |
11731 | 94 |
|
17756 | 95 |
val RAW_METHOD_CASES = Meth; |
11731 | 96 |
|
17110 | 97 |
fun RAW_METHOD tac = RAW_METHOD_CASES (NO_CASES o tac); |
12144 | 98 |
|
17110 | 99 |
fun METHOD_CASES tac = RAW_METHOD_CASES (fn facts => |
21687 | 100 |
Seq.THEN (ALLGOALS Goal.conjunction_tac, tac facts)); |
8372 | 101 |
|
21687 | 102 |
fun METHOD tac = RAW_METHOD (fn facts => ALLGOALS Goal.conjunction_tac THEN tac facts); |
5824 | 103 |
|
104 |
val fail = METHOD (K no_tac); |
|
105 |
val succeed = METHOD (K all_tac); |
|
106 |
||
107 |
||
17110 | 108 |
(* insert facts *) |
7419 | 109 |
|
110 |
local |
|
5824 | 111 |
|
21579 | 112 |
fun cut_rule_tac rule = |
52732 | 113 |
rtac (Drule.forall_intr_vars rule COMP_INCR revcut_rl); |
6981 | 114 |
|
7419 | 115 |
in |
5824 | 116 |
|
51552
c713c9505f68
clarified Skip_Proof.cheat_tac: more standard tactic;
wenzelm
parents:
51383
diff
changeset
|
117 |
fun insert_tac [] _ = all_tac |
7419 | 118 |
| insert_tac facts i = EVERY (map (fn th => cut_rule_tac th i) facts); |
6981 | 119 |
|
7555 | 120 |
val insert_facts = METHOD (ALLGOALS o insert_tac); |
7664 | 121 |
fun insert thms = METHOD (fn _ => ALLGOALS (insert_tac thms)); |
7419 | 122 |
|
9706 | 123 |
fun SIMPLE_METHOD tac = METHOD (fn facts => ALLGOALS (insert_tac facts) THEN tac); |
21592 | 124 |
fun SIMPLE_METHOD'' quant tac = METHOD (fn facts => quant (insert_tac facts THEN' tac)); |
125 |
val SIMPLE_METHOD' = SIMPLE_METHOD'' HEADGOAL; |
|
9706 | 126 |
|
12324
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
wenzelm
parents:
12311
diff
changeset
|
127 |
end; |
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
wenzelm
parents:
12311
diff
changeset
|
128 |
|
9706 | 129 |
|
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
130 |
(* cheating *) |
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
131 |
|
52059 | 132 |
fun cheating ctxt int = METHOD (fn _ => fn st => |
133 |
if int orelse Config.get ctxt quick_and_dirty then |
|
51552
c713c9505f68
clarified Skip_Proof.cheat_tac: more standard tactic;
wenzelm
parents:
51383
diff
changeset
|
134 |
ALLGOALS Skip_Proof.cheat_tac st |
c713c9505f68
clarified Skip_Proof.cheat_tac: more standard tactic;
wenzelm
parents:
51383
diff
changeset
|
135 |
else error "Cheating requires quick_and_dirty mode!"); |
17356
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
136 |
|
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
137 |
|
17110 | 138 |
(* unfold intro/elim rules *) |
139 |
||
52732 | 140 |
fun intro ths = SIMPLE_METHOD' (CHANGED_PROP o REPEAT_ALL_NEW (match_tac ths)); |
141 |
fun elim ths = SIMPLE_METHOD' (CHANGED_PROP o REPEAT_ALL_NEW (ematch_tac ths)); |
|
17110 | 142 |
|
143 |
||
12384 | 144 |
(* unfold/fold definitions *) |
145 |
||
35624 | 146 |
fun unfold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (Local_Defs.unfold_tac ctxt ths)); |
147 |
fun fold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (Local_Defs.fold_tac ctxt ths)); |
|
6532 | 148 |
|
12384 | 149 |
|
12829 | 150 |
(* atomize rule statements *) |
151 |
||
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
152 |
fun atomize false ctxt = |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
153 |
SIMPLE_METHOD' (CHANGED_PROP o Object_Logic.atomize_prems_tac ctxt) |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
154 |
| atomize true ctxt = |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
155 |
RAW_METHOD (K (HEADGOAL (CHANGED_PROP o Object_Logic.full_atomize_tac ctxt))); |
12829 | 156 |
|
157 |
||
18039 | 158 |
(* this -- resolve facts directly *) |
12384 | 159 |
|
52732 | 160 |
val this = METHOD (EVERY o map (HEADGOAL o rtac)); |
9484 | 161 |
|
162 |
||
18039 | 163 |
(* fact -- composition by facts from context *) |
164 |
||
42360 | 165 |
fun fact [] ctxt = SIMPLE_METHOD' (Proof_Context.some_fact_tac ctxt) |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
166 |
| fact rules ctxt = SIMPLE_METHOD' (Proof_Context.fact_tac ctxt rules); |
18039 | 167 |
|
168 |
||
17110 | 169 |
(* assumption *) |
7419 | 170 |
|
171 |
local |
|
172 |
||
19778 | 173 |
fun cond_rtac cond rule = SUBGOAL (fn (prop, i) => |
174 |
if cond (Logic.strip_assums_concl prop) |
|
52732 | 175 |
then rtac rule i else no_tac); |
7419 | 176 |
|
29857
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
kleing
parents:
29301
diff
changeset
|
177 |
in |
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
kleing
parents:
29301
diff
changeset
|
178 |
|
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
|
179 |
fun assm_tac ctxt = |
17110 | 180 |
assume_tac APPEND' |
23349 | 181 |
Goal.assume_rule_tac ctxt APPEND' |
19778 | 182 |
cond_rtac (can Logic.dest_equals) Drule.reflexive_thm APPEND' |
183 |
cond_rtac (can Logic.dest_term) Drule.termI; |
|
17110 | 184 |
|
49846
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
wenzelm
parents:
48992
diff
changeset
|
185 |
fun all_assm_tac ctxt = |
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
wenzelm
parents:
48992
diff
changeset
|
186 |
let |
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
wenzelm
parents:
48992
diff
changeset
|
187 |
fun tac i st = |
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
wenzelm
parents:
48992
diff
changeset
|
188 |
if i > Thm.nprems_of st then all_tac st |
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
wenzelm
parents:
48992
diff
changeset
|
189 |
else ((assm_tac ctxt i THEN tac i) ORELSE tac (i + 1)) st; |
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
wenzelm
parents:
48992
diff
changeset
|
190 |
in tac 1 end; |
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
|
191 |
|
23349 | 192 |
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
|
193 |
(fn [] => assm_tac ctxt |
23349 | 194 |
| [fact] => solve_tac [fact] |
195 |
| _ => K no_tac)); |
|
196 |
||
49846
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
wenzelm
parents:
48992
diff
changeset
|
197 |
fun finish immed ctxt = |
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
wenzelm
parents:
48992
diff
changeset
|
198 |
METHOD (K ((if immed then all_assm_tac ctxt else all_tac) THEN flexflex_tac)); |
7419 | 199 |
|
200 |
end; |
|
201 |
||
202 |
||
17110 | 203 |
(* rule etc. -- single-step refinements *) |
12347 | 204 |
|
42616
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents:
42380
diff
changeset
|
205 |
val rule_trace = Attrib.setup_config_bool (Binding.name "rule_trace") (fn _ => false); |
12347 | 206 |
|
17110 | 207 |
fun trace ctxt rules = |
41379 | 208 |
if Config.get ctxt rule_trace andalso not (null rules) then |
51584 | 209 |
Pretty.big_list "rules:" (map (Display.pretty_thm_item ctxt) rules) |
21962 | 210 |
|> Pretty.string_of |> tracing |
211 |
else (); |
|
12347 | 212 |
|
213 |
local |
|
214 |
||
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
215 |
fun gen_rule_tac tac ctxt rules facts = |
18841 | 216 |
(fn i => fn st => |
217 |
if null facts then tac rules i st |
|
218 |
else Seq.maps (fn rule => (tac o single) rule i st) (Drule.multi_resolves facts rules)) |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
219 |
THEN_ALL_NEW Goal.norm_hhf_tac ctxt; |
7130 | 220 |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
221 |
fun gen_arule_tac tac ctxt j rules facts = |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
222 |
EVERY' (gen_rule_tac tac ctxt rules facts :: replicate j assume_tac); |
10744 | 223 |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
224 |
fun gen_some_rule_tac tac ctxt arg_rules facts = SUBGOAL (fn (goal, i) => |
11785 | 225 |
let |
226 |
val rules = |
|
227 |
if not (null arg_rules) then arg_rules |
|
33369 | 228 |
else flat (Context_Rules.find_rules false facts goal ctxt) |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
229 |
in trace ctxt rules; tac ctxt rules facts i end); |
10309 | 230 |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
231 |
fun meth tac x y = METHOD (HEADGOAL o tac x y); |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
232 |
fun meth' tac x y z = METHOD (HEADGOAL o tac x y z); |
8220 | 233 |
|
7419 | 234 |
in |
235 |
||
52732 | 236 |
val rule_tac = gen_rule_tac resolve_tac; |
10744 | 237 |
val rule = meth rule_tac; |
238 |
val some_rule_tac = gen_some_rule_tac rule_tac; |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
239 |
val some_rule = meth some_rule_tac; |
10744 | 240 |
|
52732 | 241 |
val erule = meth' (gen_arule_tac eresolve_tac); |
242 |
val drule = meth' (gen_arule_tac dresolve_tac); |
|
243 |
val frule = meth' (gen_arule_tac forward_tac); |
|
5824 | 244 |
|
7419 | 245 |
end; |
246 |
||
247 |
||
25270 | 248 |
(* intros_tac -- pervasive search spanned by intro rules *) |
249 |
||
36093
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
ballarin
parents:
33522
diff
changeset
|
250 |
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
|
251 |
goals (insert_tac facts THEN' |
25270 | 252 |
REPEAT_ALL_NEW (resolve_tac intros)) |
253 |
THEN Tactic.distinct_subgoals_tac; |
|
254 |
||
36093
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
ballarin
parents:
33522
diff
changeset
|
255 |
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
|
256 |
val try_intros_tac = gen_intros_tac TRYALL; |
25270 | 257 |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37198
diff
changeset
|
258 |
|
8351 | 259 |
(* ML tactics *) |
260 |
||
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37198
diff
changeset
|
261 |
structure ML_Tactic = Proof_Data |
26472
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
262 |
( |
27235 | 263 |
type T = thm list -> tactic; |
26472
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
264 |
fun init _ = undefined; |
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
265 |
); |
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
266 |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37198
diff
changeset
|
267 |
val set_tactic = ML_Tactic.put; |
8351 | 268 |
|
55828
42ac3cfb89f6
clarified language markup: added "delimited" property;
wenzelm
parents:
55795
diff
changeset
|
269 |
fun ml_tactic source ctxt = |
26472
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
270 |
let |
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents:
26463
diff
changeset
|
271 |
val ctxt' = ctxt |> Context.proof_map |
55828
42ac3cfb89f6
clarified language markup: added "delimited" property;
wenzelm
parents:
55795
diff
changeset
|
272 |
(ML_Context.expression (#pos source) |
27235 | 273 |
"fun tactic (facts: thm list) : tactic" |
55828
42ac3cfb89f6
clarified language markup: added "delimited" property;
wenzelm
parents:
55795
diff
changeset
|
274 |
"Context.map_proof (Method.set_tactic tactic)" (ML_Lex.read_source source)); |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37198
diff
changeset
|
275 |
in Context.setmp_thread_data (SOME (Context.Proof ctxt)) (ML_Tactic.get ctxt') end; |
23425 | 276 |
|
55828
42ac3cfb89f6
clarified language markup: added "delimited" property;
wenzelm
parents:
55795
diff
changeset
|
277 |
fun tactic source ctxt = METHOD (ml_tactic source ctxt); |
42ac3cfb89f6
clarified language markup: added "delimited" property;
wenzelm
parents:
55795
diff
changeset
|
278 |
fun raw_tactic source ctxt = RAW_METHOD (ml_tactic source ctxt); |
8351 | 279 |
|
280 |
||
5824 | 281 |
|
17110 | 282 |
(** method syntax **) |
283 |
||
284 |
(* method text *) |
|
285 |
||
286 |
type src = Args.src; |
|
5824 | 287 |
|
55765 | 288 |
datatype combinator_info = Combinator_Info of {keywords: Position.T list}; |
289 |
fun combinator_info keywords = Combinator_Info {keywords = keywords}; |
|
290 |
val no_combinator_info = combinator_info []; |
|
291 |
||
17110 | 292 |
datatype text = |
55742
a989bdaf8121
modernized Method.check_name/check_source (with reports) vs. strict Method.the_method (without interning nor reports), e.g. relevant for semantic completion;
wenzelm
parents:
55709
diff
changeset
|
293 |
Source of src | |
32193
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents:
32091
diff
changeset
|
294 |
Basic of Proof.context -> method | |
55765 | 295 |
Then of combinator_info * text list | |
296 |
Orelse of combinator_info * text list | |
|
297 |
Try of combinator_info * text | |
|
298 |
Repeat1 of combinator_info * text | |
|
299 |
Select_Goals of combinator_info * int * text; |
|
300 |
||
54883
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
54742
diff
changeset
|
301 |
fun primitive_text r = Basic (SIMPLE_METHOD o PRIMITIVE o r); |
32193
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents:
32091
diff
changeset
|
302 |
val succeed_text = Basic (K succeed); |
56029
8bedca4bd5a3
clarified Args.src: more abstract type, position refers to name only;
wenzelm
parents:
56026
diff
changeset
|
303 |
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
|
304 |
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
|
305 |
val done_text = Basic (K (SIMPLE_METHOD all_tac)); |
52059 | 306 |
fun sorry_text int = Basic (fn ctxt => cheating ctxt int); |
17110 | 307 |
|
49846
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
wenzelm
parents:
48992
diff
changeset
|
308 |
fun finish_text (NONE, immed) = Basic (finish immed) |
55765 | 309 |
| finish_text (SOME txt, immed) = Then (no_combinator_info, [txt, Basic (finish immed)]); |
17110 | 310 |
|
311 |
||
312 |
(* method definitions *) |
|
5824 | 313 |
|
33522 | 314 |
structure Methods = Theory_Data |
22846 | 315 |
( |
33095
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
wenzelm
parents:
33092
diff
changeset
|
316 |
type T = ((src -> Proof.context -> method) * string) Name_Space.table; |
33159 | 317 |
val empty : T = Name_Space.empty_table "method"; |
16448 | 318 |
val extend = I; |
33522 | 319 |
fun merge data : T = Name_Space.merge_tables data; |
22846 | 320 |
); |
5824 | 321 |
|
55742
a989bdaf8121
modernized Method.check_name/check_source (with reports) vs. strict Method.the_method (without interning nor reports), e.g. relevant for semantic completion;
wenzelm
parents:
55709
diff
changeset
|
322 |
val get_methods = Methods.get o Proof_Context.theory_of; |
a989bdaf8121
modernized Method.check_name/check_source (with reports) vs. strict Method.the_method (without interning nor reports), e.g. relevant for semantic completion;
wenzelm
parents:
55709
diff
changeset
|
323 |
|
a989bdaf8121
modernized Method.check_name/check_source (with reports) vs. strict Method.the_method (without interning nor reports), e.g. relevant for semantic completion;
wenzelm
parents:
55709
diff
changeset
|
324 |
fun print_methods ctxt = |
22846 | 325 |
let |
55742
a989bdaf8121
modernized Method.check_name/check_source (with reports) vs. strict Method.the_method (without interning nor reports), e.g. relevant for semantic completion;
wenzelm
parents:
55709
diff
changeset
|
326 |
val meths = get_methods ctxt; |
50301 | 327 |
fun prt_meth (name, (_, "")) = Pretty.mark_str name |
42813
6c841fa92fa2
optional description for 'attribute_setup' and 'method_setup';
wenzelm
parents:
42616
diff
changeset
|
328 |
| prt_meth (name, (_, comment)) = |
50301 | 329 |
Pretty.block |
330 |
(Pretty.mark_str name :: Pretty.str ":" :: Pretty.brk 2 :: Pretty.text comment); |
|
22846 | 331 |
in |
42358
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
wenzelm
parents:
41379
diff
changeset
|
332 |
[Pretty.big_list "methods:" (map prt_meth (Name_Space.extern_table ctxt meths))] |
22846 | 333 |
|> Pretty.chunks |> Pretty.writeln |
334 |
end; |
|
7611 | 335 |
|
33092
c859019d3ac5
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
wenzelm
parents:
32969
diff
changeset
|
336 |
fun add_method name meth comment thy = thy |
47005
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents:
46775
diff
changeset
|
337 |
|> Methods.map (Name_Space.define (Context.Theory thy) true (name, (meth, comment)) #> snd); |
31304 | 338 |
|
55997
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents:
55917
diff
changeset
|
339 |
|
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents:
55917
diff
changeset
|
340 |
(* check *) |
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents:
55917
diff
changeset
|
341 |
|
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents:
55917
diff
changeset
|
342 |
fun check_name ctxt = #1 o Name_Space.check (Context.Proof ctxt) (get_methods ctxt); |
56029
8bedca4bd5a3
clarified Args.src: more abstract type, position refers to name only;
wenzelm
parents:
56026
diff
changeset
|
343 |
fun check_src ctxt src = #1 (Args.check_src (Context.Proof ctxt) (get_methods ctxt) src); |
55997
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents:
55917
diff
changeset
|
344 |
|
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents:
55917
diff
changeset
|
345 |
|
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents:
55917
diff
changeset
|
346 |
(* get methods *) |
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents:
55917
diff
changeset
|
347 |
|
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents:
55917
diff
changeset
|
348 |
fun method ctxt = |
56029
8bedca4bd5a3
clarified Args.src: more abstract type, position refers to name only;
wenzelm
parents:
56026
diff
changeset
|
349 |
let val table = get_methods ctxt |
8bedca4bd5a3
clarified Args.src: more abstract type, position refers to name only;
wenzelm
parents:
56026
diff
changeset
|
350 |
in fn src => #1 (Name_Space.get table (#1 (Args.name_of_src src))) src end; |
55742
a989bdaf8121
modernized Method.check_name/check_source (with reports) vs. strict Method.the_method (without interning nor reports), e.g. relevant for semantic completion;
wenzelm
parents:
55709
diff
changeset
|
351 |
|
55997
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents:
55917
diff
changeset
|
352 |
fun method_cmd ctxt = method ctxt o check_src ctxt; |
20030 | 353 |
|
5824 | 354 |
|
30512 | 355 |
(* method setup *) |
356 |
||
357 |
fun syntax scan = Args.context_syntax "method" scan; |
|
358 |
||
31304 | 359 |
fun setup name scan = |
360 |
add_method name |
|
361 |
(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
|
362 |
|
55828
42ac3cfb89f6
clarified language markup: added "delimited" property;
wenzelm
parents:
55795
diff
changeset
|
363 |
fun method_setup name source cmt = |
42ac3cfb89f6
clarified language markup: added "delimited" property;
wenzelm
parents:
55795
diff
changeset
|
364 |
Context.theory_map (ML_Context.expression (#pos source) |
30544 | 365 |
"val (name, scan, comment): binding * (Proof.context -> Proof.method) context_parser * string" |
366 |
"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
|
367 |
(ML_Lex.read Position.none ("(" ^ ML_Syntax.make_binding name ^ ", ") @ |
55828
42ac3cfb89f6
clarified language markup: added "delimited" property;
wenzelm
parents:
55795
diff
changeset
|
368 |
ML_Lex.read_source source @ |
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
|
369 |
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
|
370 |
|
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents:
17314
diff
changeset
|
371 |
|
5884 | 372 |
|
17110 | 373 |
(** concrete syntax **) |
5824 | 374 |
|
5884 | 375 |
(* sections *) |
5824 | 376 |
|
20289 | 377 |
type modifier = (Proof.context -> Proof.context) * attribute; |
7268 | 378 |
|
379 |
local |
|
380 |
||
24010
2ef318813e1a
method section scanners: added [[declaration]] syntax, ignore sid-effects of thms;
wenzelm
parents:
23937
diff
changeset
|
381 |
fun thms ss = Scan.repeat (Scan.unless (Scan.lift (Scan.first ss)) Attrib.multi_thm) >> flat; |
46775
6287653e63ec
canonical argument order for attribute application;
wenzelm
parents:
46466
diff
changeset
|
382 |
fun app (f, att) ths context = fold_map (Thm.apply_attribute att) ths (Context.map_proof f context); |
5824 | 383 |
|
30540 | 384 |
in |
385 |
||
24022 | 386 |
fun section ss = Scan.depend (fn context => (Scan.first ss -- Scan.pass context (thms ss)) :|-- |
46775
6287653e63ec
canonical argument order for attribute application;
wenzelm
parents:
46466
diff
changeset
|
387 |
(fn (m, ths) => Scan.succeed (swap (app m ths context)))); |
5884 | 388 |
|
30540 | 389 |
fun sections ss = Scan.repeat (section ss); |
5824 | 390 |
|
7268 | 391 |
end; |
392 |
||
5824 | 393 |
|
30515 | 394 |
(* extra rule methods *) |
395 |
||
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
396 |
fun xrule_meth meth = |
36950 | 397 |
Scan.lift (Scan.optional (Args.parens Parse.nat) 0) -- Attrib.thms >> |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
398 |
(fn (n, ths) => fn ctxt => meth ctxt n ths); |
30515 | 399 |
|
400 |
||
55761
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents:
55742
diff
changeset
|
401 |
(* text range *) |
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents:
55742
diff
changeset
|
402 |
|
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents:
55742
diff
changeset
|
403 |
type text_range = text * Position.range; |
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents:
55742
diff
changeset
|
404 |
|
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents:
55742
diff
changeset
|
405 |
fun text NONE = NONE |
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents:
55742
diff
changeset
|
406 |
| text (SOME (txt, _)) = SOME txt; |
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents:
55742
diff
changeset
|
407 |
|
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents:
55742
diff
changeset
|
408 |
fun position NONE = Position.none |
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents:
55742
diff
changeset
|
409 |
| position (SOME (_, (pos, _))) = pos; |
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents:
55742
diff
changeset
|
410 |
|
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents:
55742
diff
changeset
|
411 |
|
55795 | 412 |
(* reports *) |
413 |
||
414 |
local |
|
415 |
||
416 |
fun keyword_positions (Source _) = [] |
|
417 |
| keyword_positions (Basic _) = [] |
|
418 |
| keyword_positions (Then (Combinator_Info {keywords}, texts)) = |
|
419 |
keywords @ maps keyword_positions texts |
|
420 |
| keyword_positions (Orelse (Combinator_Info {keywords}, texts)) = |
|
421 |
keywords @ maps keyword_positions texts |
|
422 |
| keyword_positions (Try (Combinator_Info {keywords}, text)) = |
|
423 |
keywords @ keyword_positions text |
|
424 |
| keyword_positions (Repeat1 (Combinator_Info {keywords}, text)) = |
|
425 |
keywords @ keyword_positions text |
|
426 |
| keyword_positions (Select_Goals (Combinator_Info {keywords}, _, text)) = |
|
427 |
keywords @ keyword_positions text; |
|
428 |
||
429 |
in |
|
430 |
||
431 |
fun reports_of ((text, (pos, _)): text_range) = |
|
432 |
(pos, Markup.language_method) :: |
|
55917
5438ed05e1c9
special treatment of method combinators like Args.$$$ keywords, although parsed via Parse.$$$;
wenzelm
parents:
55828
diff
changeset
|
433 |
maps (fn p => map (pair p) (Markup.keyword3 :: Completion.suppress_abbrevs "")) |
5438ed05e1c9
special treatment of method combinators like Args.$$$ keywords, although parsed via Parse.$$$;
wenzelm
parents:
55828
diff
changeset
|
434 |
(keyword_positions text); |
55795 | 435 |
|
436 |
val report = Position.reports o reports_of; |
|
437 |
||
438 |
end; |
|
439 |
||
440 |
||
27813
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
441 |
(* outer parser *) |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
442 |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
443 |
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
|
444 |
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
|
445 |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
446 |
local |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
447 |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
448 |
fun meth4 x = |
56029
8bedca4bd5a3
clarified Args.src: more abstract type, position refers to name only;
wenzelm
parents:
56026
diff
changeset
|
449 |
(Parse.position Parse.xname >> (fn name => Source (Args.src name [])) || |
55048
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
wenzelm
parents:
54883
diff
changeset
|
450 |
Scan.ahead Parse.cartouche |-- Parse.not_eof >> (fn tok => |
56029
8bedca4bd5a3
clarified Args.src: more abstract type, position refers to name only;
wenzelm
parents:
56026
diff
changeset
|
451 |
Source (Args.src ("cartouche", Token.pos_of tok) [tok])) || |
36950 | 452 |
Parse.$$$ "(" |-- Parse.!!! (meth0 --| Parse.$$$ ")")) x |
27813
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
453 |
and meth3 x = |
55765 | 454 |
(meth4 -- Parse.position (Parse.$$$ "?") |
455 |
>> (fn (m, (_, pos)) => Try (combinator_info [pos], m)) || |
|
456 |
meth4 -- Parse.position (Parse.$$$ "+") |
|
457 |
>> (fn (m, (_, pos)) => Repeat1 (combinator_info [pos], m)) || |
|
458 |
meth4 -- |
|
459 |
(Parse.position (Parse.$$$ "[") -- Scan.optional Parse.nat 1 -- Parse.position (Parse.$$$ "]")) |
|
460 |
>> (fn (m, (((_, pos1), n), (_, pos2))) => |
|
461 |
Select_Goals (combinator_info [pos1, pos2], n, m)) || |
|
27813
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
462 |
meth4) x |
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
463 |
and meth2 x = |
56029
8bedca4bd5a3
clarified Args.src: more abstract type, position refers to name only;
wenzelm
parents:
56026
diff
changeset
|
464 |
(Parse.position Parse.xname -- Args.parse1 is_symid_meth >> (Source o uncurry Args.src) || |
27813
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
465 |
meth3) x |
55765 | 466 |
and meth1 x = |
467 |
(Parse.enum1_positions "," meth2 |
|
468 |
>> (fn ([m], _) => m | (ms, ps) => Then (combinator_info ps, ms))) x |
|
469 |
and meth0 x = |
|
470 |
(Parse.enum1_positions "|" meth1 |
|
471 |
>> (fn ([m], _) => m | (ms, ps) => Orelse (combinator_info ps, ms))) x; |
|
27813
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
472 |
|
49866 | 473 |
in |
474 |
||
475 |
val parse = |
|
55709
4e5a83a46ded
clarified Token.range_of in accordance to Symbol_Pos.range;
wenzelm
parents:
55708
diff
changeset
|
476 |
Scan.trace meth3 >> (fn (m, toks) => (m, Token.range_of toks)); |
49866 | 477 |
|
55761
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents:
55742
diff
changeset
|
478 |
end; |
27813
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
479 |
|
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27751
diff
changeset
|
480 |
|
18708 | 481 |
(* theory setup *) |
5824 | 482 |
|
53171 | 483 |
val _ = Theory.setup |
42616
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents:
42380
diff
changeset
|
484 |
(setup (Binding.name "fail") (Scan.succeed (K fail)) "force failure" #> |
30515 | 485 |
setup (Binding.name "succeed") (Scan.succeed (K succeed)) "succeed" #> |
486 |
setup (Binding.name "-") (Scan.succeed (K insert_facts)) |
|
487 |
"do nothing (insert current facts only)" #> |
|
488 |
setup (Binding.name "insert") (Attrib.thms >> (K o insert)) |
|
489 |
"insert theorems, ignoring facts (improper)" #> |
|
490 |
setup (Binding.name "intro") (Attrib.thms >> (K o intro)) |
|
491 |
"repeatedly apply introduction rules" #> |
|
492 |
setup (Binding.name "elim") (Attrib.thms >> (K o elim)) |
|
493 |
"repeatedly apply elimination rules" #> |
|
494 |
setup (Binding.name "unfold") (Attrib.thms >> unfold_meth) "unfold definitions" #> |
|
495 |
setup (Binding.name "fold") (Attrib.thms >> fold_meth) "fold definitions" #> |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
496 |
setup (Binding.name "atomize") (Scan.lift (Args.mode "full") >> atomize) |
30515 | 497 |
"present local premises as object-level statements" #> |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
498 |
setup (Binding.name "rule") (Attrib.thms >> (fn ths => fn ctxt => some_rule ctxt ths)) |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53171
diff
changeset
|
499 |
"apply some intro/elim rule" #> |
30515 | 500 |
setup (Binding.name "erule") (xrule_meth erule) "apply rule in elimination manner (improper)" #> |
501 |
setup (Binding.name "drule") (xrule_meth drule) "apply rule in destruct manner (improper)" #> |
|
502 |
setup (Binding.name "frule") (xrule_meth frule) "apply rule in forward manner (improper)" #> |
|
503 |
setup (Binding.name "this") (Scan.succeed (K this)) "apply current facts as rules" #> |
|
504 |
setup (Binding.name "fact") (Attrib.thms >> fact) "composition by facts from context" #> |
|
505 |
setup (Binding.name "assumption") (Scan.succeed assumption) |
|
506 |
"proof by assumption, preferring facts" #> |
|
507 |
setup (Binding.name "rename_tac") (Args.goal_spec -- Scan.lift (Scan.repeat1 Args.name) >> |
|
52732 | 508 |
(fn (quant, xs) => K (SIMPLE_METHOD'' quant (rename_tac xs)))) |
30515 | 509 |
"rename parameters of goal" #> |
36950 | 510 |
setup (Binding.name "rotate_tac") (Args.goal_spec -- Scan.lift (Scan.optional Parse.int 1) >> |
52732 | 511 |
(fn (quant, i) => K (SIMPLE_METHOD'' quant (rotate_tac i)))) |
30515 | 512 |
"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
|
513 |
setup (Binding.name "tactic") (Scan.lift Args.name_source_position >> tactic) |
30515 | 514 |
"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
|
515 |
setup (Binding.name "raw_tactic") (Scan.lift Args.name_source_position >> raw_tactic) |
53171 | 516 |
"ML tactic as raw proof method"); |
5824 | 517 |
|
518 |
||
16145 | 519 |
(*final declarations of this structure!*) |
520 |
val unfold = unfold_meth; |
|
521 |
val fold = fold_meth; |
|
522 |
||
5824 | 523 |
end; |
524 |
||
30508
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents:
30466
diff
changeset
|
525 |
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
|
526 |
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
|
527 |
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
|
528 |
val METHOD = Method.METHOD; |
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents:
30466
diff
changeset
|
529 |
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
|
530 |
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
|
531 |
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
|
532 |