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