| author | paulson | 
| Thu, 07 Jan 2016 17:42:01 +0000 | |
| changeset 62088 | 8463e386eaec | 
| parent 61917 | 35ec3757d3c1 | 
| child 62795 | 063d2f23cdf6 | 
| 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  | 
||
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
7  | 
infix 1 CONTEXT_THEN_ALL_NEW;  | 
| 61834 | 8  | 
|
9  | 
signature BASIC_METHOD =  | 
|
10  | 
sig  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
11  | 
type context_state = Proof.context * thm  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
12  | 
type context_tactic = context_state -> context_state Seq.result Seq.seq  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
13  | 
val CONTEXT_CASES: Rule_Cases.cases -> tactic -> context_tactic  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
14  | 
val CONTEXT_SUBGOAL: (term * int -> context_tactic) -> int -> context_tactic  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
15  | 
val CONTEXT_THEN_ALL_NEW: (int -> context_tactic) * (int -> tactic) -> int -> context_tactic  | 
| 61834 | 16  | 
end;  | 
17  | 
||
| 5824 | 18  | 
signature METHOD =  | 
19  | 
sig  | 
|
| 61834 | 20  | 
include BASIC_METHOD  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
21  | 
type method = thm list -> context_tactic  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
22  | 
val CONTEXT: Proof.context -> thm Seq.seq -> context_state Seq.result Seq.seq  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
23  | 
val CONTEXT_TACTIC: tactic -> context_tactic  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
24  | 
val NO_CONTEXT_TACTIC: Proof.context -> context_tactic -> tactic  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
25  | 
val CONTEXT_METHOD: (thm list -> context_tactic) -> method  | 
| 17110 | 26  | 
val METHOD: (thm list -> tactic) -> method  | 
27  | 
val fail: method  | 
|
28  | 
val succeed: method  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
29  | 
val insert_tac: Proof.context -> thm list -> int -> tactic  | 
| 17110 | 30  | 
val insert: thm list -> method  | 
31  | 
val SIMPLE_METHOD: tactic -> method  | 
|
| 21592 | 32  | 
val SIMPLE_METHOD': (int -> tactic) -> method  | 
33  | 
val SIMPLE_METHOD'': ((int -> tactic) -> tactic) -> (int -> tactic) -> method  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
34  | 
val goal_cases_tac: string list -> context_tactic  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
35  | 
val cheating: bool -> method  | 
| 58957 | 36  | 
val intro: Proof.context -> thm list -> method  | 
37  | 
val elim: Proof.context -> thm list -> method  | 
|
| 20289 | 38  | 
val unfold: thm list -> Proof.context -> method  | 
39  | 
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
 | 
40  | 
val atomize: bool -> Proof.context -> method  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59067 
diff
changeset
 | 
41  | 
val this: Proof.context -> method  | 
| 20289 | 42  | 
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
 | 
43  | 
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
 | 
44  | 
val all_assm_tac: Proof.context -> tactic  | 
| 20289 | 45  | 
val assumption: Proof.context -> method  | 
| 
46466
 
61c7214b4885
tuned signature, according to actual usage of these operations;
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
46  | 
val rule_trace: bool Config.T  | 
| 20289 | 47  | 
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
 | 
48  | 
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
 | 
49  | 
val some_rule_tac: Proof.context -> thm list -> thm list -> int -> tactic  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59067 
diff
changeset
 | 
50  | 
val intros_tac: Proof.context -> thm list -> thm list -> tactic  | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59067 
diff
changeset
 | 
51  | 
val try_intros_tac: Proof.context -> 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
 | 
52  | 
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
 | 
53  | 
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
 | 
54  | 
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
 | 
55  | 
val frule: Proof.context -> int -> thm list -> method  | 
| 
58018
 
beb4b7c0bb30
proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
 
wenzelm 
parents: 
58016 
diff
changeset
 | 
56  | 
val set_tactic: (morphism -> thm list -> tactic) -> Context.generic -> Context.generic  | 
| 55765 | 57  | 
type combinator_info  | 
58  | 
val no_combinator_info: combinator_info  | 
|
| 
59660
 
49e498cedd02
support structural composition (THEN_ALL_NEW) for proof methods;
 
wenzelm 
parents: 
59498 
diff
changeset
 | 
59  | 
datatype combinator = Then | Then_All_New | Orelse | Try | Repeat1 | Select_Goals of int  | 
| 5824 | 60  | 
datatype text =  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
58007 
diff
changeset
 | 
61  | 
Source of Token.src |  | 
| 
32193
 
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
 
wenzelm 
parents: 
32091 
diff
changeset
 | 
62  | 
Basic of Proof.context -> method |  | 
| 58005 | 63  | 
Combinator of combinator_info * combinator * text list  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
58007 
diff
changeset
 | 
64  | 
val map_source: (Token.src -> Token.src) -> text -> text  | 
| 
54883
 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 
wenzelm 
parents: 
54742 
diff
changeset
 | 
65  | 
val primitive_text: (Proof.context -> thm -> thm) -> text  | 
| 17857 | 66  | 
val succeed_text: text  | 
| 
60618
 
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
 
wenzelm 
parents: 
60609 
diff
changeset
 | 
67  | 
val standard_text: text  | 
| 17110 | 68  | 
val this_text: text  | 
69  | 
val done_text: text  | 
|
| 
17356
 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 
wenzelm 
parents: 
17314 
diff
changeset
 | 
70  | 
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
 | 
71  | 
val finish_text: text option * bool -> text  | 
| 
59917
 
9830c944670f
more uniform "verbose" option to print name space;
 
wenzelm 
parents: 
59914 
diff
changeset
 | 
72  | 
val print_methods: bool -> Proof.context -> unit  | 
| 
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
 | 
73  | 
val check_name: Proof.context -> xstring * Position.T -> string  | 
| 59907 | 74  | 
val check_src: Proof.context -> Token.src -> Token.src  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61812 
diff
changeset
 | 
75  | 
val check_text: Proof.context -> text -> text  | 
| 
61917
 
35ec3757d3c1
check and report source at most once, notably in body of "match" method;
 
wenzelm 
parents: 
61843 
diff
changeset
 | 
76  | 
val checked_text: text -> bool  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
58007 
diff
changeset
 | 
77  | 
val method_syntax: (Proof.context -> method) context_parser ->  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
58007 
diff
changeset
 | 
78  | 
Token.src -> Proof.context -> method  | 
| 30512 | 79  | 
val setup: binding -> (Proof.context -> method) context_parser -> string -> theory -> theory  | 
| 
57935
 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 
wenzelm 
parents: 
57863 
diff
changeset
 | 
80  | 
val local_setup: binding -> (Proof.context -> method) context_parser -> string ->  | 
| 
 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 
wenzelm 
parents: 
57863 
diff
changeset
 | 
81  | 
local_theory -> local_theory  | 
| 59064 | 82  | 
val method_setup: bstring * Position.T -> Input.source -> string -> local_theory -> local_theory  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
58007 
diff
changeset
 | 
83  | 
val method: Proof.context -> Token.src -> Proof.context -> method  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
58007 
diff
changeset
 | 
84  | 
val method_closure: Proof.context -> Token.src -> Token.src  | 
| 59909 | 85  | 
val closure: bool Config.T  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
58007 
diff
changeset
 | 
86  | 
val method_cmd: Proof.context -> Token.src -> Proof.context -> method  | 
| 60553 | 87  | 
val detect_closure_state: thm -> bool  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
88  | 
val STATIC: (unit -> unit) -> context_tactic  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
89  | 
val RUNTIME: context_tactic -> context_tactic  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
90  | 
val sleep: Time.time -> context_tactic  | 
| 
58007
 
671c607fb4af
just one context for Method.evaluate (in contrast to a989bdaf8121, but in accordance to old global situation);
 
wenzelm 
parents: 
58006 
diff
changeset
 | 
91  | 
val evaluate: text -> Proof.context -> method  | 
| 
58048
 
aa6296d09e0e
more explicit Method.modifier with reported position;
 
wenzelm 
parents: 
58034 
diff
changeset
 | 
92  | 
  type modifier = {init: Proof.context -> Proof.context, attribute: attribute, pos: Position.T}
 | 
| 
 
aa6296d09e0e
more explicit Method.modifier with reported position;
 
wenzelm 
parents: 
58034 
diff
changeset
 | 
93  | 
val modifier: attribute -> Position.T -> modifier  | 
| 
59982
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
94  | 
val old_section_parser: bool Config.T  | 
| 
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
95  | 
val sections: modifier parser list -> unit context_parser  | 
| 
49889
 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 
wenzelm 
parents: 
49866 
diff
changeset
 | 
96  | 
type text_range = text * Position.range  | 
| 
 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 
wenzelm 
parents: 
49866 
diff
changeset
 | 
97  | 
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
 | 
98  | 
val position: text_range option -> Position.T  | 
| 55795 | 99  | 
val reports_of: text_range -> Position.report list  | 
100  | 
val report: text_range -> unit  | 
|
| 59981 | 101  | 
val parser: int -> text_range parser  | 
| 
55761
 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 
wenzelm 
parents: 
55742 
diff
changeset
 | 
102  | 
val parse: text_range parser  | 
| 5824 | 103  | 
end;  | 
104  | 
||
105  | 
structure Method: METHOD =  | 
|
106  | 
struct  | 
|
107  | 
||
| 
12324
 
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
 
wenzelm 
parents: 
12311 
diff
changeset
 | 
108  | 
(** proof methods **)  | 
| 
 
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
 
wenzelm 
parents: 
12311 
diff
changeset
 | 
109  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
110  | 
(* tactics with proof context / cases *)  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
111  | 
|
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
112  | 
type context_state = Proof.context * thm;  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
113  | 
type context_tactic = context_state -> context_state Seq.result Seq.seq;  | 
| 61834 | 114  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
115  | 
fun CONTEXT ctxt : thm Seq.seq -> context_state Seq.result Seq.seq =  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
116  | 
Seq.map (Seq.Result o pair ctxt);  | 
| 61834 | 117  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
118  | 
fun CONTEXT_TACTIC tac : context_tactic =  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
119  | 
fn (ctxt, st) => CONTEXT ctxt (tac st);  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
120  | 
|
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
121  | 
fun NO_CONTEXT_TACTIC ctxt (tac: context_tactic) st =  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
122  | 
tac (ctxt, st) |> Seq.filter_results |> Seq.map snd;  | 
| 61834 | 123  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
124  | 
fun CONTEXT_CASES cases tac : context_tactic =  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
125  | 
fn (ctxt, st) => CONTEXT (Proof_Context.update_cases cases ctxt) (tac st);  | 
| 61834 | 126  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
127  | 
fun CONTEXT_SUBGOAL tac i : context_tactic =  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
128  | 
fn (ctxt, st) =>  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
129  | 
(case try Logic.nth_prem (i, Thm.prop_of st) of  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
130  | 
SOME goal => tac (goal, i) (ctxt, st)  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
131  | 
| NONE => Seq.empty);  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
132  | 
|
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
133  | 
fun (tac1 CONTEXT_THEN_ALL_NEW tac2) i : context_tactic =  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
134  | 
fn (ctxt, st) =>  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
135  | 
(ctxt, st) |> tac1 i |> Seq.maps_results (fn (ctxt', st') =>  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
136  | 
CONTEXT ctxt' ((Seq.INTERVAL tac2 i (i + Thm.nprems_of st' - Thm.nprems_of st)) st'));  | 
| 61834 | 137  | 
|
138  | 
||
139  | 
(* type method *)  | 
|
| 11731 | 140  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
141  | 
type method = thm list -> context_tactic;  | 
| 11731 | 142  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
143  | 
fun CONTEXT_METHOD tac : method =  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
144  | 
fn facts => CONTEXT_TACTIC (ALLGOALS Goal.conjunction_tac) #> Seq.maps_results (tac facts);  | 
| 59953 | 145  | 
|
146  | 
fun METHOD tac : method =  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
147  | 
fn facts => CONTEXT_TACTIC (ALLGOALS Goal.conjunction_tac THEN tac facts);  | 
| 5824 | 148  | 
|
149  | 
val fail = METHOD (K no_tac);  | 
|
150  | 
val succeed = METHOD (K all_tac);  | 
|
151  | 
||
152  | 
||
| 17110 | 153  | 
(* insert facts *)  | 
| 7419 | 154  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
155  | 
fun insert_tac _ [] _ = all_tac  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
156  | 
| insert_tac ctxt facts i =  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
157  | 
EVERY (map (fn r => resolve_tac ctxt [Drule.forall_intr_vars r COMP_INCR revcut_rl] i) facts);  | 
| 5824 | 158  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
159  | 
fun insert thms =  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
160  | 
CONTEXT_METHOD (fn _ => fn (ctxt, st) => st |> ALLGOALS (insert_tac ctxt thms) |> CONTEXT ctxt);  | 
| 6981 | 161  | 
|
| 5824 | 162  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
163  | 
fun SIMPLE_METHOD tac =  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
164  | 
CONTEXT_METHOD (fn facts => fn (ctxt, st) =>  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
165  | 
st |> (ALLGOALS (insert_tac ctxt facts) THEN tac) |> CONTEXT ctxt);  | 
| 7419 | 166  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
167  | 
fun SIMPLE_METHOD'' quant tac =  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
168  | 
CONTEXT_METHOD (fn facts => fn (ctxt, st) =>  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
169  | 
st |> quant (insert_tac ctxt facts THEN' tac) |> CONTEXT ctxt);  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
170  | 
|
| 21592 | 171  | 
val SIMPLE_METHOD' = SIMPLE_METHOD'' HEADGOAL;  | 
| 9706 | 172  | 
|
173  | 
||
| 60578 | 174  | 
(* goals as cases *)  | 
175  | 
||
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
176  | 
fun goal_cases_tac case_names : context_tactic =  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
177  | 
fn (ctxt, st) =>  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
178  | 
let  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
179  | 
val cases =  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
180  | 
(if null case_names then map string_of_int (1 upto Thm.nprems_of st) else case_names)  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
181  | 
|> map (rpair [] o rpair [])  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
182  | 
|> Rule_Cases.make_common ctxt (Thm.prop_of (Rule_Cases.internalize_params st));  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
183  | 
in CONTEXT_CASES cases all_tac (ctxt, st) end;  | 
| 60578 | 184  | 
|
185  | 
||
| 
17356
 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 
wenzelm 
parents: 
17314 
diff
changeset
 | 
186  | 
(* cheating *)  | 
| 
 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 
wenzelm 
parents: 
17314 
diff
changeset
 | 
187  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
188  | 
fun cheating int = CONTEXT_METHOD (fn _ => fn (ctxt, st) =>  | 
| 52059 | 189  | 
if int orelse Config.get ctxt quick_and_dirty then  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
190  | 
CONTEXT ctxt (ALLGOALS (Skip_Proof.cheat_tac ctxt) st)  | 
| 
51552
 
c713c9505f68
clarified Skip_Proof.cheat_tac: more standard tactic;
 
wenzelm 
parents: 
51383 
diff
changeset
 | 
191  | 
else error "Cheating requires quick_and_dirty mode!");  | 
| 
17356
 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 
wenzelm 
parents: 
17314 
diff
changeset
 | 
192  | 
|
| 
 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 
wenzelm 
parents: 
17314 
diff
changeset
 | 
193  | 
|
| 17110 | 194  | 
(* unfold intro/elim rules *)  | 
195  | 
||
| 58957 | 196  | 
fun intro ctxt ths = SIMPLE_METHOD' (CHANGED_PROP o REPEAT_ALL_NEW (match_tac ctxt ths));  | 
197  | 
fun elim ctxt ths = SIMPLE_METHOD' (CHANGED_PROP o REPEAT_ALL_NEW (ematch_tac ctxt ths));  | 
|
| 17110 | 198  | 
|
199  | 
||
| 12384 | 200  | 
(* unfold/fold definitions *)  | 
201  | 
||
| 35624 | 202  | 
fun unfold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (Local_Defs.unfold_tac ctxt ths));  | 
203  | 
fun fold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (Local_Defs.fold_tac ctxt ths));  | 
|
| 6532 | 204  | 
|
| 12384 | 205  | 
|
| 12829 | 206  | 
(* atomize rule statements *)  | 
207  | 
||
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53171 
diff
changeset
 | 
208  | 
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
 | 
209  | 
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
 | 
210  | 
| atomize true ctxt =  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
211  | 
CONTEXT_TACTIC o K (HEADGOAL (CHANGED_PROP o Object_Logic.full_atomize_tac ctxt));  | 
| 12829 | 212  | 
|
213  | 
||
| 18039 | 214  | 
(* this -- resolve facts directly *)  | 
| 12384 | 215  | 
|
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59067 
diff
changeset
 | 
216  | 
fun this ctxt = METHOD (EVERY o map (HEADGOAL o resolve_tac ctxt o single));  | 
| 9484 | 217  | 
|
218  | 
||
| 18039 | 219  | 
(* fact -- composition by facts from context *)  | 
220  | 
||
| 42360 | 221  | 
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
 | 
222  | 
| fact rules ctxt = SIMPLE_METHOD' (Proof_Context.fact_tac ctxt rules);  | 
| 18039 | 223  | 
|
224  | 
||
| 17110 | 225  | 
(* assumption *)  | 
| 7419 | 226  | 
|
227  | 
local  | 
|
228  | 
||
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59067 
diff
changeset
 | 
229  | 
fun cond_rtac ctxt cond rule = SUBGOAL (fn (prop, i) =>  | 
| 19778 | 230  | 
if cond (Logic.strip_assums_concl prop)  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59067 
diff
changeset
 | 
231  | 
then resolve_tac ctxt [rule] i else no_tac);  | 
| 7419 | 232  | 
|
| 
29857
 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 
kleing 
parents: 
29301 
diff
changeset
 | 
233  | 
in  | 
| 
 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 
kleing 
parents: 
29301 
diff
changeset
 | 
234  | 
|
| 
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
 | 
235  | 
fun assm_tac ctxt =  | 
| 
58963
 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 
wenzelm 
parents: 
58957 
diff
changeset
 | 
236  | 
assume_tac ctxt APPEND'  | 
| 23349 | 237  | 
Goal.assume_rule_tac ctxt APPEND'  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59067 
diff
changeset
 | 
238  | 
cond_rtac ctxt (can Logic.dest_equals) Drule.reflexive_thm APPEND'  | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59067 
diff
changeset
 | 
239  | 
cond_rtac ctxt (can Logic.dest_term) Drule.termI;  | 
| 17110 | 240  | 
|
| 
49846
 
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
241  | 
fun all_assm_tac ctxt =  | 
| 
 
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
242  | 
let  | 
| 
 
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
243  | 
fun tac i st =  | 
| 
 
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
244  | 
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
 | 
245  | 
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
 | 
246  | 
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
 | 
247  | 
|
| 23349 | 248  | 
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
 | 
249  | 
(fn [] => assm_tac ctxt  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59067 
diff
changeset
 | 
250  | 
| [fact] => solve_tac ctxt [fact]  | 
| 23349 | 251  | 
| _ => K no_tac));  | 
252  | 
||
| 
49846
 
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
253  | 
fun finish immed ctxt =  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58837 
diff
changeset
 | 
254  | 
METHOD (K ((if immed then all_assm_tac ctxt else all_tac) THEN flexflex_tac ctxt));  | 
| 7419 | 255  | 
|
256  | 
end;  | 
|
257  | 
||
258  | 
||
| 17110 | 259  | 
(* rule etc. -- single-step refinements *)  | 
| 12347 | 260  | 
|
| 56204 | 261  | 
val rule_trace = Attrib.setup_config_bool @{binding rule_trace} (fn _ => false);
 | 
| 12347 | 262  | 
|
| 17110 | 263  | 
fun trace ctxt rules =  | 
| 41379 | 264  | 
if Config.get ctxt rule_trace andalso not (null rules) then  | 
| 61268 | 265  | 
Pretty.big_list "rules:" (map (Thm.pretty_thm_item ctxt) rules)  | 
| 21962 | 266  | 
|> Pretty.string_of |> tracing  | 
267  | 
else ();  | 
|
| 12347 | 268  | 
|
269  | 
local  | 
|
270  | 
||
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53171 
diff
changeset
 | 
271  | 
fun gen_rule_tac tac ctxt rules facts =  | 
| 18841 | 272  | 
(fn i => fn st =>  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59067 
diff
changeset
 | 
273  | 
if null facts then tac ctxt rules i st  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58837 
diff
changeset
 | 
274  | 
else  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59067 
diff
changeset
 | 
275  | 
Seq.maps (fn rule => (tac ctxt o single) rule i st)  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58837 
diff
changeset
 | 
276  | 
(Drule.multi_resolves (SOME ctxt) facts rules))  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53171 
diff
changeset
 | 
277  | 
THEN_ALL_NEW Goal.norm_hhf_tac ctxt;  | 
| 7130 | 278  | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53171 
diff
changeset
 | 
279  | 
fun gen_arule_tac tac ctxt j rules facts =  | 
| 
58963
 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 
wenzelm 
parents: 
58957 
diff
changeset
 | 
280  | 
EVERY' (gen_rule_tac tac ctxt rules facts :: replicate j (assume_tac ctxt));  | 
| 10744 | 281  | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53171 
diff
changeset
 | 
282  | 
fun gen_some_rule_tac tac ctxt arg_rules facts = SUBGOAL (fn (goal, i) =>  | 
| 11785 | 283  | 
let  | 
284  | 
val rules =  | 
|
285  | 
if not (null arg_rules) then arg_rules  | 
|
| 61049 | 286  | 
else flat (Context_Rules.find_rules ctxt false facts goal);  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53171 
diff
changeset
 | 
287  | 
in trace ctxt rules; tac ctxt rules facts i end);  | 
| 10309 | 288  | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53171 
diff
changeset
 | 
289  | 
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
 | 
290  | 
fun meth' tac x y z = METHOD (HEADGOAL o tac x y z);  | 
| 8220 | 291  | 
|
| 7419 | 292  | 
in  | 
293  | 
||
| 52732 | 294  | 
val rule_tac = gen_rule_tac resolve_tac;  | 
| 10744 | 295  | 
val rule = meth rule_tac;  | 
296  | 
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
 | 
297  | 
val some_rule = meth some_rule_tac;  | 
| 10744 | 298  | 
|
| 52732 | 299  | 
val erule = meth' (gen_arule_tac eresolve_tac);  | 
300  | 
val drule = meth' (gen_arule_tac dresolve_tac);  | 
|
301  | 
val frule = meth' (gen_arule_tac forward_tac);  | 
|
| 5824 | 302  | 
|
| 7419 | 303  | 
end;  | 
304  | 
||
305  | 
||
| 25270 | 306  | 
(* intros_tac -- pervasive search spanned by intro rules *)  | 
307  | 
||
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59067 
diff
changeset
 | 
308  | 
fun gen_intros_tac goals ctxt intros facts =  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
309  | 
goals (insert_tac ctxt facts THEN'  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59067 
diff
changeset
 | 
310  | 
REPEAT_ALL_NEW (resolve_tac ctxt intros))  | 
| 25270 | 311  | 
THEN Tactic.distinct_subgoals_tac;  | 
312  | 
||
| 
36093
 
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
 
ballarin 
parents: 
33522 
diff
changeset
 | 
313  | 
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
 | 
314  | 
val try_intros_tac = gen_intros_tac TRYALL;  | 
| 25270 | 315  | 
|
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
316  | 
|
| 58016 | 317  | 
|
318  | 
(** method syntax **)  | 
|
319  | 
||
320  | 
(* context data *)  | 
|
| 8351 | 321  | 
|
| 58016 | 322  | 
structure Data = Generic_Data  | 
| 
26472
 
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
323  | 
(  | 
| 58016 | 324  | 
type T =  | 
325  | 
((Token.src -> Proof.context -> method) * string) Name_Space.table * (*methods*)  | 
|
| 
58018
 
beb4b7c0bb30
proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
 
wenzelm 
parents: 
58016 
diff
changeset
 | 
326  | 
(morphism -> thm list -> tactic) option; (*ML tactic*)  | 
| 58016 | 327  | 
val empty : T = (Name_Space.empty_table "method", NONE);  | 
328  | 
val extend = I;  | 
|
329  | 
fun merge ((tab, tac), (tab', tac')) : T =  | 
|
330  | 
(Name_Space.merge_tables (tab, tab'), merge_options (tac, tac'));  | 
|
| 
26472
 
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
331  | 
);  | 
| 
 
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
332  | 
|
| 58016 | 333  | 
val get_methods = fst o Data.get;  | 
334  | 
val map_methods = Data.map o apfst;  | 
|
335  | 
||
336  | 
||
337  | 
(* ML tactic *)  | 
|
338  | 
||
339  | 
val set_tactic = Data.map o apsnd o K o SOME;  | 
|
340  | 
||
341  | 
fun the_tactic context =  | 
|
342  | 
(case snd (Data.get context) of  | 
|
343  | 
SOME tac => tac  | 
|
344  | 
| NONE => raise Fail "Undefined ML tactic");  | 
|
| 8351 | 345  | 
|
| 
58018
 
beb4b7c0bb30
proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
 
wenzelm 
parents: 
58016 
diff
changeset
 | 
346  | 
val parse_tactic =  | 
| 
 
beb4b7c0bb30
proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
 
wenzelm 
parents: 
58016 
diff
changeset
 | 
347  | 
Scan.state :|-- (fn context =>  | 
| 
 
beb4b7c0bb30
proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
 
wenzelm 
parents: 
58016 
diff
changeset
 | 
348  | 
Scan.lift (Args.text_declaration (fn source =>  | 
| 
 
beb4b7c0bb30
proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
 
wenzelm 
parents: 
58016 
diff
changeset
 | 
349  | 
let  | 
| 
 
beb4b7c0bb30
proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
 
wenzelm 
parents: 
58016 
diff
changeset
 | 
350  | 
val context' = context |>  | 
| 59064 | 351  | 
ML_Context.expression (Input.range_of source)  | 
| 
58991
 
92b6f4e68c5a
more careful ML source positions, for improved PIDE markup;
 
wenzelm 
parents: 
58979 
diff
changeset
 | 
352  | 
"tactic" "Morphism.morphism -> thm list -> tactic"  | 
| 
 
92b6f4e68c5a
more careful ML source positions, for improved PIDE markup;
 
wenzelm 
parents: 
58979 
diff
changeset
 | 
353  | 
"Method.set_tactic tactic"  | 
| 59067 | 354  | 
(ML_Lex.read "fn morphism: Morphism.morphism => fn facts: thm list =>" @  | 
| 
58991
 
92b6f4e68c5a
more careful ML source positions, for improved PIDE markup;
 
wenzelm 
parents: 
58979 
diff
changeset
 | 
355  | 
ML_Lex.read_source false source);  | 
| 
58018
 
beb4b7c0bb30
proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
 
wenzelm 
parents: 
58016 
diff
changeset
 | 
356  | 
val tac = the_tactic context';  | 
| 
 
beb4b7c0bb30
proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
 
wenzelm 
parents: 
58016 
diff
changeset
 | 
357  | 
in  | 
| 
 
beb4b7c0bb30
proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
 
wenzelm 
parents: 
58016 
diff
changeset
 | 
358  | 
fn phi =>  | 
| 
 
beb4b7c0bb30
proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
 
wenzelm 
parents: 
58016 
diff
changeset
 | 
359  | 
set_tactic (fn _ => Context.setmp_thread_data (SOME context) (tac phi))  | 
| 
 
beb4b7c0bb30
proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
 
wenzelm 
parents: 
58016 
diff
changeset
 | 
360  | 
end)) >> (fn decl => Morphism.form (the_tactic (Morphism.form decl context))));  | 
| 8351 | 361  | 
|
362  | 
||
| 17110 | 363  | 
(* method text *)  | 
364  | 
||
| 55765 | 365  | 
datatype combinator_info = Combinator_Info of {keywords: Position.T list};
 | 
366  | 
fun combinator_info keywords = Combinator_Info {keywords = keywords};
 | 
|
367  | 
val no_combinator_info = combinator_info [];  | 
|
368  | 
||
| 
59660
 
49e498cedd02
support structural composition (THEN_ALL_NEW) for proof methods;
 
wenzelm 
parents: 
59498 
diff
changeset
 | 
369  | 
datatype combinator = Then | Then_All_New | Orelse | Try | Repeat1 | Select_Goals of int;  | 
| 58005 | 370  | 
|
| 17110 | 371  | 
datatype text =  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
58007 
diff
changeset
 | 
372  | 
Source of Token.src |  | 
| 
32193
 
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
 
wenzelm 
parents: 
32091 
diff
changeset
 | 
373  | 
Basic of Proof.context -> method |  | 
| 58005 | 374  | 
Combinator of combinator_info * combinator * text list;  | 
| 55765 | 375  | 
|
| 58006 | 376  | 
fun map_source f (Source src) = Source (f src)  | 
377  | 
| map_source _ (Basic meth) = Basic meth  | 
|
378  | 
| map_source f (Combinator (info, comb, txts)) = Combinator (info, comb, map (map_source f) txts);  | 
|
379  | 
||
| 
54883
 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 
wenzelm 
parents: 
54742 
diff
changeset
 | 
380  | 
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
 | 
381  | 
val succeed_text = Basic (K succeed);  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61812 
diff
changeset
 | 
382  | 
val standard_text = Source (Token.make_src ("standard", Position.none) []);
 | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59067 
diff
changeset
 | 
383  | 
val this_text = Basic this;  | 
| 
32193
 
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
 
wenzelm 
parents: 
32091 
diff
changeset
 | 
384  | 
val done_text = Basic (K (SIMPLE_METHOD all_tac));  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
385  | 
fun sorry_text int = Basic (fn _ => cheating int);  | 
| 17110 | 386  | 
|
| 
49846
 
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
387  | 
fun finish_text (NONE, immed) = Basic (finish immed)  | 
| 58005 | 388  | 
| finish_text (SOME txt, immed) =  | 
389  | 
Combinator (no_combinator_info, Then, [txt, Basic (finish immed)]);  | 
|
| 17110 | 390  | 
|
391  | 
||
392  | 
(* method definitions *)  | 
|
| 5824 | 393  | 
|
| 57937 | 394  | 
fun transfer_methods ctxt =  | 
| 
57935
 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 
wenzelm 
parents: 
57863 
diff
changeset
 | 
395  | 
let  | 
| 58016 | 396  | 
val meths0 = get_methods (Context.Theory (Proof_Context.theory_of ctxt));  | 
397  | 
val meths' = Name_Space.merge_tables (meths0, get_methods (Context.Proof ctxt));  | 
|
398  | 
in Context.proof_map (map_methods (K meths')) ctxt 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
 | 
399  | 
|
| 
59917
 
9830c944670f
more uniform "verbose" option to print name space;
 
wenzelm 
parents: 
59914 
diff
changeset
 | 
400  | 
fun print_methods verbose ctxt =  | 
| 22846 | 401  | 
let  | 
| 58016 | 402  | 
val meths = get_methods (Context.Proof ctxt);  | 
| 50301 | 403  | 
fun prt_meth (name, (_, "")) = Pretty.mark_str name  | 
| 
42813
 
6c841fa92fa2
optional description for 'attribute_setup' and 'method_setup';
 
wenzelm 
parents: 
42616 
diff
changeset
 | 
404  | 
| prt_meth (name, (_, comment)) =  | 
| 50301 | 405  | 
Pretty.block  | 
406  | 
(Pretty.mark_str name :: Pretty.str ":" :: Pretty.brk 2 :: Pretty.text comment);  | 
|
| 22846 | 407  | 
in  | 
| 
59917
 
9830c944670f
more uniform "verbose" option to print name space;
 
wenzelm 
parents: 
59914 
diff
changeset
 | 
408  | 
[Pretty.big_list "methods:" (map prt_meth (Name_Space.markup_table verbose ctxt meths))]  | 
| 
56334
 
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
 
wenzelm 
parents: 
56278 
diff
changeset
 | 
409  | 
|> Pretty.writeln_chunks  | 
| 22846 | 410  | 
end;  | 
| 7611 | 411  | 
|
| 
57935
 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 
wenzelm 
parents: 
57863 
diff
changeset
 | 
412  | 
|
| 
 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 
wenzelm 
parents: 
57863 
diff
changeset
 | 
413  | 
(* define *)  | 
| 
 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 
wenzelm 
parents: 
57863 
diff
changeset
 | 
414  | 
|
| 
 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 
wenzelm 
parents: 
57863 
diff
changeset
 | 
415  | 
fun define_global binding meth comment thy =  | 
| 
 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 
wenzelm 
parents: 
57863 
diff
changeset
 | 
416  | 
let  | 
| 
 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 
wenzelm 
parents: 
57863 
diff
changeset
 | 
417  | 
val context = Context.Theory thy;  | 
| 
 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 
wenzelm 
parents: 
57863 
diff
changeset
 | 
418  | 
val (name, meths') =  | 
| 58016 | 419  | 
Name_Space.define context true (binding, (meth, comment)) (get_methods context);  | 
420  | 
in (name, Context.the_theory (map_methods (K meths') context)) end;  | 
|
| 
57935
 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 
wenzelm 
parents: 
57863 
diff
changeset
 | 
421  | 
|
| 
57941
 
57200bdc2aa7
localized command 'method_setup' and 'attribute_setup';
 
wenzelm 
parents: 
57938 
diff
changeset
 | 
422  | 
fun define binding meth comment =  | 
| 
 
57200bdc2aa7
localized command 'method_setup' and 'attribute_setup';
 
wenzelm 
parents: 
57938 
diff
changeset
 | 
423  | 
Local_Theory.background_theory_result (define_global binding meth comment)  | 
| 
 
57200bdc2aa7
localized command 'method_setup' and 'attribute_setup';
 
wenzelm 
parents: 
57938 
diff
changeset
 | 
424  | 
#-> (fn name =>  | 
| 
 
57200bdc2aa7
localized command 'method_setup' and 'attribute_setup';
 
wenzelm 
parents: 
57938 
diff
changeset
 | 
425  | 
Local_Theory.map_contexts (K transfer_methods)  | 
| 58016 | 426  | 
#> Local_Theory.generic_alias map_methods binding name  | 
| 
57941
 
57200bdc2aa7
localized command 'method_setup' and 'attribute_setup';
 
wenzelm 
parents: 
57938 
diff
changeset
 | 
427  | 
#> pair name);  | 
| 31304 | 428  | 
|
| 
55997
 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 
wenzelm 
parents: 
55917 
diff
changeset
 | 
429  | 
|
| 
 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 
wenzelm 
parents: 
55917 
diff
changeset
 | 
430  | 
(* check *)  | 
| 
 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 
wenzelm 
parents: 
55917 
diff
changeset
 | 
431  | 
|
| 58016 | 432  | 
fun check_name ctxt =  | 
433  | 
let val context = Context.Proof ctxt  | 
|
434  | 
in #1 o Name_Space.check context (get_methods context) end;  | 
|
435  | 
||
| 59907 | 436  | 
fun check_src ctxt =  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61812 
diff
changeset
 | 
437  | 
#1 o Token.check_src ctxt (get_methods o Context.Proof);  | 
| 
55997
 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 
wenzelm 
parents: 
55917 
diff
changeset
 | 
438  | 
|
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61812 
diff
changeset
 | 
439  | 
fun check_text ctxt (Source src) = Source (check_src ctxt src)  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61812 
diff
changeset
 | 
440  | 
| check_text _ (Basic m) = Basic m  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61812 
diff
changeset
 | 
441  | 
| check_text ctxt (Combinator (x, y, body)) = Combinator (x, y, map (check_text ctxt) body);  | 
| 
59914
 
d1ddcd8df4e4
proper treatment of internal method name as already checked Token.src;
 
wenzelm 
parents: 
59909 
diff
changeset
 | 
442  | 
|
| 
61917
 
35ec3757d3c1
check and report source at most once, notably in body of "match" method;
 
wenzelm 
parents: 
61843 
diff
changeset
 | 
443  | 
fun checked_text (Source src) = Token.checked_src src  | 
| 
 
35ec3757d3c1
check and report source at most once, notably in body of "match" method;
 
wenzelm 
parents: 
61843 
diff
changeset
 | 
444  | 
| checked_text (Basic _) = true  | 
| 
 
35ec3757d3c1
check and report source at most once, notably in body of "match" method;
 
wenzelm 
parents: 
61843 
diff
changeset
 | 
445  | 
| checked_text (Combinator (_, _, body)) = forall checked_text body;  | 
| 
 
35ec3757d3c1
check and report source at most once, notably in body of "match" method;
 
wenzelm 
parents: 
61843 
diff
changeset
 | 
446  | 
|
| 
55997
 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 
wenzelm 
parents: 
55917 
diff
changeset
 | 
447  | 
|
| 30512 | 448  | 
(* method setup *)  | 
449  | 
||
| 
57935
 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 
wenzelm 
parents: 
57863 
diff
changeset
 | 
450  | 
fun method_syntax scan src ctxt : method =  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
58007 
diff
changeset
 | 
451  | 
let val (m, ctxt') = Token.syntax scan src ctxt in m ctxt' end;  | 
| 
57935
 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 
wenzelm 
parents: 
57863 
diff
changeset
 | 
452  | 
|
| 
 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 
wenzelm 
parents: 
57863 
diff
changeset
 | 
453  | 
fun setup binding scan comment = define_global binding (method_syntax scan) comment #> snd;  | 
| 
 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 
wenzelm 
parents: 
57863 
diff
changeset
 | 
454  | 
fun local_setup binding scan comment = define binding (method_syntax scan) comment #> snd;  | 
| 
17356
 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 
wenzelm 
parents: 
17314 
diff
changeset
 | 
455  | 
|
| 
58979
 
162a4c2e97bc
more careful ML source positions, for improved PIDE markup;
 
wenzelm 
parents: 
58978 
diff
changeset
 | 
456  | 
fun method_setup name source comment =  | 
| 
 
162a4c2e97bc
more careful ML source positions, for improved PIDE markup;
 
wenzelm 
parents: 
58978 
diff
changeset
 | 
457  | 
ML_Lex.read_source false source  | 
| 59064 | 458  | 
|> ML_Context.expression (Input.range_of source) "parser"  | 
| 
58991
 
92b6f4e68c5a
more careful ML source positions, for improved PIDE markup;
 
wenzelm 
parents: 
58979 
diff
changeset
 | 
459  | 
"(Proof.context -> Proof.method) context_parser"  | 
| 
58979
 
162a4c2e97bc
more careful ML source positions, for improved PIDE markup;
 
wenzelm 
parents: 
58978 
diff
changeset
 | 
460  | 
    ("Context.map_proof (Method.local_setup " ^ ML_Syntax.atomic (ML_Syntax.make_binding name) ^
 | 
| 
 
162a4c2e97bc
more careful ML source positions, for improved PIDE markup;
 
wenzelm 
parents: 
58978 
diff
changeset
 | 
461  | 
" parser " ^ ML_Syntax.print_string comment ^ ")")  | 
| 
57941
 
57200bdc2aa7
localized command 'method_setup' and 'attribute_setup';
 
wenzelm 
parents: 
57938 
diff
changeset
 | 
462  | 
|> Context.proof_map;  | 
| 
17356
 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 
wenzelm 
parents: 
17314 
diff
changeset
 | 
463  | 
|
| 
 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 
wenzelm 
parents: 
17314 
diff
changeset
 | 
464  | 
|
| 58006 | 465  | 
(* prepare methods *)  | 
466  | 
||
467  | 
fun method ctxt =  | 
|
| 58016 | 468  | 
let val table = get_methods (Context.Proof ctxt)  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
58007 
diff
changeset
 | 
469  | 
in fn src => #1 (Name_Space.get table (#1 (Token.name_of_src src))) src end;  | 
| 58006 | 470  | 
|
| 59909 | 471  | 
fun method_closure ctxt src =  | 
| 58006 | 472  | 
let  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61812 
diff
changeset
 | 
473  | 
val src' = map Token.init_assignable src;  | 
| 59909 | 474  | 
val ctxt' = Context_Position.not_really ctxt;  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
475  | 
val _ = Seq.pull (method ctxt' src' ctxt' [] (ctxt', Goal.protect 0 Drule.dummy_thm));  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61812 
diff
changeset
 | 
476  | 
in map Token.closure src' end;  | 
| 58006 | 477  | 
|
| 59909 | 478  | 
val closure = Config.bool (Config.declare ("Method.closure", @{here}) (K (Config.Bool true)));
 | 
479  | 
||
480  | 
fun method_cmd ctxt =  | 
|
481  | 
check_src ctxt #>  | 
|
482  | 
Config.get ctxt closure ? method_closure ctxt #>  | 
|
483  | 
method ctxt;  | 
|
| 58006 | 484  | 
|
485  | 
||
| 60609 | 486  | 
(* static vs. runtime state *)  | 
| 60553 | 487  | 
|
488  | 
fun detect_closure_state st =  | 
|
489  | 
(case try Logic.dest_term (Thm.concl_of (perhaps (try Goal.conclude) st)) of  | 
|
490  | 
NONE => false  | 
|
491  | 
| SOME t => Term.is_dummy_pattern t);  | 
|
492  | 
||
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
493  | 
fun STATIC test : context_tactic =  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
494  | 
fn (ctxt, st) =>  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
495  | 
if detect_closure_state st then (test (); Seq.single (Seq.Result (ctxt, st))) else Seq.empty;  | 
| 60609 | 496  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
497  | 
fun RUNTIME (tac: context_tactic) (ctxt, st) =  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
498  | 
if detect_closure_state st then Seq.empty else tac (ctxt, st);  | 
| 60553 | 499  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
500  | 
fun sleep t = RUNTIME (fn ctxt_st => (OS.Process.sleep t; Seq.single (Seq.Result ctxt_st)));  | 
| 60554 | 501  | 
|
| 60553 | 502  | 
|
| 
58003
 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 
wenzelm 
parents: 
58002 
diff
changeset
 | 
503  | 
(* evaluate method text *)  | 
| 
 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 
wenzelm 
parents: 
58002 
diff
changeset
 | 
504  | 
|
| 
 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 
wenzelm 
parents: 
58002 
diff
changeset
 | 
505  | 
local  | 
| 
 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 
wenzelm 
parents: 
58002 
diff
changeset
 | 
506  | 
|
| 
 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 
wenzelm 
parents: 
58002 
diff
changeset
 | 
507  | 
val op THEN = Seq.THEN;  | 
| 
 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 
wenzelm 
parents: 
58002 
diff
changeset
 | 
508  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
509  | 
fun BYPASS_CONTEXT (tac: tactic) =  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
510  | 
fn result =>  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
511  | 
(case result of  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
512  | 
Seq.Error _ => Seq.single result  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
513  | 
| Seq.Result (ctxt, st) => tac st |> CONTEXT ctxt);  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
514  | 
|
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
515  | 
val preparation = BYPASS_CONTEXT (ALLGOALS Goal.conjunction_tac);  | 
| 
59660
 
49e498cedd02
support structural composition (THEN_ALL_NEW) for proof methods;
 
wenzelm 
parents: 
59498 
diff
changeset
 | 
516  | 
|
| 
 
49e498cedd02
support structural composition (THEN_ALL_NEW) for proof methods;
 
wenzelm 
parents: 
59498 
diff
changeset
 | 
517  | 
fun RESTRICT_GOAL i n method =  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
518  | 
BYPASS_CONTEXT (PRIMITIVE (Goal.restrict i n)) THEN  | 
| 
59660
 
49e498cedd02
support structural composition (THEN_ALL_NEW) for proof methods;
 
wenzelm 
parents: 
59498 
diff
changeset
 | 
519  | 
method THEN  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
520  | 
BYPASS_CONTEXT (PRIMITIVE (Goal.unrestrict i));  | 
| 
59660
 
49e498cedd02
support structural composition (THEN_ALL_NEW) for proof methods;
 
wenzelm 
parents: 
59498 
diff
changeset
 | 
521  | 
|
| 
 
49e498cedd02
support structural composition (THEN_ALL_NEW) for proof methods;
 
wenzelm 
parents: 
59498 
diff
changeset
 | 
522  | 
fun SELECT_GOAL method i = RESTRICT_GOAL i 1 method;  | 
| 
 
49e498cedd02
support structural composition (THEN_ALL_NEW) for proof methods;
 
wenzelm 
parents: 
59498 
diff
changeset
 | 
523  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
524  | 
fun (method1 THEN_ALL_NEW method2) i (result : context_state Seq.result) =  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
525  | 
(case result of  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
526  | 
Seq.Error _ => Seq.single result  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
527  | 
| Seq.Result (_, st) =>  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
528  | 
result |> method1 i  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
529  | 
|> Seq.maps (fn result' =>  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
530  | 
(case result' of  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
531  | 
Seq.Error _ => Seq.single result'  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
532  | 
| Seq.Result (_, st') =>  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
533  | 
result' |> Seq.INTERVAL method2 i (i + Thm.nprems_of st' - Thm.nprems_of st))))  | 
| 
58003
 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 
wenzelm 
parents: 
58002 
diff
changeset
 | 
534  | 
|
| 58005 | 535  | 
fun COMBINATOR1 comb [meth] = comb meth  | 
536  | 
| COMBINATOR1 _ _ = raise Fail "Method combinator requires exactly one argument";  | 
|
537  | 
||
538  | 
fun combinator Then = Seq.EVERY  | 
|
| 
59660
 
49e498cedd02
support structural composition (THEN_ALL_NEW) for proof methods;
 
wenzelm 
parents: 
59498 
diff
changeset
 | 
539  | 
| combinator Then_All_New =  | 
| 
 
49e498cedd02
support structural composition (THEN_ALL_NEW) for proof methods;
 
wenzelm 
parents: 
59498 
diff
changeset
 | 
540  | 
(fn [] => Seq.single  | 
| 
 
49e498cedd02
support structural composition (THEN_ALL_NEW) for proof methods;
 
wenzelm 
parents: 
59498 
diff
changeset
 | 
541  | 
| methods =>  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
542  | 
preparation THEN (foldl1 (op THEN_ALL_NEW) (map SELECT_GOAL methods) 1))  | 
| 58005 | 543  | 
| combinator Orelse = Seq.FIRST  | 
544  | 
| combinator Try = COMBINATOR1 Seq.TRY  | 
|
545  | 
| combinator Repeat1 = COMBINATOR1 Seq.REPEAT1  | 
|
| 
59660
 
49e498cedd02
support structural composition (THEN_ALL_NEW) for proof methods;
 
wenzelm 
parents: 
59498 
diff
changeset
 | 
546  | 
| combinator (Select_Goals n) =  | 
| 
 
49e498cedd02
support structural composition (THEN_ALL_NEW) for proof methods;
 
wenzelm 
parents: 
59498 
diff
changeset
 | 
547  | 
COMBINATOR1 (fn method => preparation THEN RESTRICT_GOAL 1 n method);  | 
| 
58003
 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 
wenzelm 
parents: 
58002 
diff
changeset
 | 
548  | 
|
| 
 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 
wenzelm 
parents: 
58002 
diff
changeset
 | 
549  | 
in  | 
| 
 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 
wenzelm 
parents: 
58002 
diff
changeset
 | 
550  | 
|
| 
58007
 
671c607fb4af
just one context for Method.evaluate (in contrast to a989bdaf8121, but in accordance to old global situation);
 
wenzelm 
parents: 
58006 
diff
changeset
 | 
551  | 
fun evaluate text ctxt =  | 
| 
58003
 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 
wenzelm 
parents: 
58002 
diff
changeset
 | 
552  | 
let  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
553  | 
fun eval0 m facts = Seq.single #> Seq.maps_results (m facts);  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
554  | 
fun eval (Basic m) = eval0 (m ctxt)  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
555  | 
| eval (Source src) = eval0 (method_cmd ctxt src ctxt)  | 
| 58005 | 556  | 
| eval (Combinator (_, c, txts)) =  | 
557  | 
let  | 
|
558  | 
val comb = combinator c;  | 
|
| 
58007
 
671c607fb4af
just one context for Method.evaluate (in contrast to a989bdaf8121, but in accordance to old global situation);
 
wenzelm 
parents: 
58006 
diff
changeset
 | 
559  | 
val meths = map eval txts;  | 
| 
 
671c607fb4af
just one context for Method.evaluate (in contrast to a989bdaf8121, but in accordance to old global situation);
 
wenzelm 
parents: 
58006 
diff
changeset
 | 
560  | 
in fn facts => comb (map (fn meth => meth facts) meths) end;  | 
| 
58003
 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 
wenzelm 
parents: 
58002 
diff
changeset
 | 
561  | 
val meth = eval text;  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
562  | 
in fn facts => fn ctxt_st => meth facts (Seq.Result ctxt_st) end;  | 
| 
58003
 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 
wenzelm 
parents: 
58002 
diff
changeset
 | 
563  | 
|
| 
 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 
wenzelm 
parents: 
58002 
diff
changeset
 | 
564  | 
end;  | 
| 
 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 
wenzelm 
parents: 
58002 
diff
changeset
 | 
565  | 
|
| 
 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 
wenzelm 
parents: 
58002 
diff
changeset
 | 
566  | 
|
| 5884 | 567  | 
|
| 17110 | 568  | 
(** concrete syntax **)  | 
| 5824 | 569  | 
|
| 
58048
 
aa6296d09e0e
more explicit Method.modifier with reported position;
 
wenzelm 
parents: 
58034 
diff
changeset
 | 
570  | 
(* type modifier *)  | 
| 
 
aa6296d09e0e
more explicit Method.modifier with reported position;
 
wenzelm 
parents: 
58034 
diff
changeset
 | 
571  | 
|
| 
 
aa6296d09e0e
more explicit Method.modifier with reported position;
 
wenzelm 
parents: 
58034 
diff
changeset
 | 
572  | 
type modifier =  | 
| 
 
aa6296d09e0e
more explicit Method.modifier with reported position;
 
wenzelm 
parents: 
58034 
diff
changeset
 | 
573  | 
  {init: Proof.context -> Proof.context, attribute: attribute, pos: Position.T};
 | 
| 5824 | 574  | 
|
| 
58048
 
aa6296d09e0e
more explicit Method.modifier with reported position;
 
wenzelm 
parents: 
58034 
diff
changeset
 | 
575  | 
fun modifier attribute pos : modifier = {init = I, attribute = attribute, pos = pos};
 | 
| 
 
aa6296d09e0e
more explicit Method.modifier with reported position;
 
wenzelm 
parents: 
58034 
diff
changeset
 | 
576  | 
|
| 
 
aa6296d09e0e
more explicit Method.modifier with reported position;
 
wenzelm 
parents: 
58034 
diff
changeset
 | 
577  | 
|
| 
 
aa6296d09e0e
more explicit Method.modifier with reported position;
 
wenzelm 
parents: 
58034 
diff
changeset
 | 
578  | 
(* sections *)  | 
| 7268 | 579  | 
|
| 
59982
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
580  | 
val old_section_parser =  | 
| 
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
581  | 
  Config.bool (Config.declare ("Method.old_section_parser", @{here}) (K (Config.Bool false)));
 | 
| 
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
582  | 
|
| 
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
583  | 
local  | 
| 
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
584  | 
|
| 
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
585  | 
fun thms ss =  | 
| 61476 | 586  | 
Scan.repeats (Scan.unless (Scan.lift (Scan.first ss)) Attrib.multi_thm);  | 
| 
59982
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
587  | 
|
| 
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
588  | 
fun app {init, attribute, pos = _} ths context =
 | 
| 
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
589  | 
fold_map (Thm.apply_attribute attribute) ths (Context.map_proof init context);  | 
| 
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
590  | 
|
| 
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
591  | 
fun section ss = Scan.depend (fn context => (Scan.first ss -- Scan.pass context (thms ss)) :|--  | 
| 
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
592  | 
(fn (m, ths) => Scan.succeed (swap (app m ths context))));  | 
| 
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
593  | 
|
| 
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
594  | 
in  | 
| 
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
595  | 
|
| 
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
596  | 
fun old_sections ss = Scan.repeat (section ss) >> K ();  | 
| 
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
597  | 
|
| 
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
598  | 
end;  | 
| 
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
599  | 
|
| 7268 | 600  | 
local  | 
601  | 
||
| 
58048
 
aa6296d09e0e
more explicit Method.modifier with reported position;
 
wenzelm 
parents: 
58034 
diff
changeset
 | 
602  | 
fun sect (modifier : modifier parser) = Scan.depend (fn context =>  | 
| 60211 | 603  | 
Scan.ahead Parse.not_eof -- Scan.trace modifier -- Scan.repeat (Scan.unless modifier Parse.xthm)  | 
604  | 
    >> (fn ((tok0, ({init, attribute, pos}, modifier_toks)), xthms) =>
 | 
|
| 
58029
 
2137e60b6f6d
clarified Method.section: explicit declaration with static closure;
 
wenzelm 
parents: 
58026 
diff
changeset
 | 
605  | 
let  | 
| 58068 | 606  | 
val decl =  | 
| 60211 | 607  | 
(case Token.get_value tok0 of  | 
| 58068 | 608  | 
SOME (Token.Declaration decl) => decl  | 
609  | 
| _ =>  | 
|
610  | 
let  | 
|
611  | 
val ctxt = Context.proof_of context;  | 
|
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61812 
diff
changeset
 | 
612  | 
val prep_att = Attrib.check_src ctxt #> map (Token.assign NONE);  | 
| 58068 | 613  | 
val thms =  | 
614  | 
map (fn (a, bs) => (Proof_Context.get_fact ctxt a, map prep_att bs)) xthms;  | 
|
615  | 
val facts =  | 
|
616  | 
Attrib.partial_evaluation ctxt [((Binding.name "dummy", []), thms)]  | 
|
617  | 
|> map (fn (_, bs) => ((Binding.empty, [Attrib.internal (K attribute)]), bs));  | 
|
| 60211 | 618  | 
|
| 58068 | 619  | 
fun decl phi =  | 
620  | 
Context.mapping I init #>  | 
|
621  | 
Attrib.generic_notes "" (Attrib.transform_facts phi facts) #> snd;  | 
|
| 60211 | 622  | 
|
| 60212 | 623  | 
val modifier_report =  | 
624  | 
(Position.set_range (Token.range_of modifier_toks),  | 
|
625  | 
Markup.properties (Position.def_properties_of pos)  | 
|
626  | 
(Markup.entity Markup.method_modifierN ""));  | 
|
| 60211 | 627  | 
val _ =  | 
| 60212 | 628  | 
Context_Position.reports ctxt (modifier_report :: Token.reports_of_value tok0);  | 
| 60211 | 629  | 
val _ = Token.assign (SOME (Token.Declaration decl)) tok0;  | 
| 58068 | 630  | 
in decl end);  | 
631  | 
in (Morphism.form decl context, decl) end));  | 
|
| 5824 | 632  | 
|
| 30540 | 633  | 
in  | 
634  | 
||
| 
59982
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
635  | 
fun sections ss =  | 
| 
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
636  | 
Args.context :|-- (fn ctxt =>  | 
| 
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
637  | 
if Config.get ctxt old_section_parser then old_sections ss  | 
| 
 
f402fd001429
option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
 
wenzelm 
parents: 
59981 
diff
changeset
 | 
638  | 
else Scan.repeat (sect (Scan.first ss)) >> K ());  | 
| 5824 | 639  | 
|
| 7268 | 640  | 
end;  | 
641  | 
||
| 5824 | 642  | 
|
| 30515 | 643  | 
(* extra rule methods *)  | 
644  | 
||
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53171 
diff
changeset
 | 
645  | 
fun xrule_meth meth =  | 
| 36950 | 646  | 
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
 | 
647  | 
(fn (n, ths) => fn ctxt => meth ctxt n ths);  | 
| 30515 | 648  | 
|
649  | 
||
| 
55761
 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 
wenzelm 
parents: 
55742 
diff
changeset
 | 
650  | 
(* text range *)  | 
| 
 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 
wenzelm 
parents: 
55742 
diff
changeset
 | 
651  | 
|
| 
 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 
wenzelm 
parents: 
55742 
diff
changeset
 | 
652  | 
type text_range = text * Position.range;  | 
| 
 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 
wenzelm 
parents: 
55742 
diff
changeset
 | 
653  | 
|
| 
 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 
wenzelm 
parents: 
55742 
diff
changeset
 | 
654  | 
fun text NONE = NONE  | 
| 
 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 
wenzelm 
parents: 
55742 
diff
changeset
 | 
655  | 
| text (SOME (txt, _)) = SOME txt;  | 
| 
 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 
wenzelm 
parents: 
55742 
diff
changeset
 | 
656  | 
|
| 
 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 
wenzelm 
parents: 
55742 
diff
changeset
 | 
657  | 
fun position NONE = Position.none  | 
| 
 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 
wenzelm 
parents: 
55742 
diff
changeset
 | 
658  | 
| position (SOME (_, (pos, _))) = pos;  | 
| 
 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 
wenzelm 
parents: 
55742 
diff
changeset
 | 
659  | 
|
| 
 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 
wenzelm 
parents: 
55742 
diff
changeset
 | 
660  | 
|
| 55795 | 661  | 
(* reports *)  | 
662  | 
||
663  | 
local  | 
|
664  | 
||
665  | 
fun keyword_positions (Source _) = []  | 
|
666  | 
| keyword_positions (Basic _) = []  | 
|
| 58005 | 667  | 
  | keyword_positions (Combinator (Combinator_Info {keywords}, _, texts)) =
 | 
668  | 
keywords @ maps keyword_positions texts;  | 
|
| 55795 | 669  | 
|
670  | 
in  | 
|
671  | 
||
672  | 
fun reports_of ((text, (pos, _)): text_range) =  | 
|
673  | 
(pos, Markup.language_method) ::  | 
|
| 
55917
 
5438ed05e1c9
special treatment of method combinators like Args.$$$ keywords, although parsed via Parse.$$$;
 
wenzelm 
parents: 
55828 
diff
changeset
 | 
674  | 
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
 | 
675  | 
(keyword_positions text);  | 
| 55795 | 676  | 
|
677  | 
val report = Position.reports o reports_of;  | 
|
678  | 
||
679  | 
end;  | 
|
680  | 
||
681  | 
||
| 59666 | 682  | 
(* parser *)  | 
683  | 
||
684  | 
local  | 
|
| 
27813
 
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27751 
diff
changeset
 | 
685  | 
|
| 
 
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27751 
diff
changeset
 | 
686  | 
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
 | 
687  | 
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
 | 
688  | 
|
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61812 
diff
changeset
 | 
689  | 
in  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61812 
diff
changeset
 | 
690  | 
|
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61812 
diff
changeset
 | 
691  | 
fun parser pri =  | 
| 59666 | 692  | 
let  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61812 
diff
changeset
 | 
693  | 
val meth_name = Parse.token Parse.xname;  | 
| 
27813
 
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27751 
diff
changeset
 | 
694  | 
|
| 59666 | 695  | 
fun meth5 x =  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61812 
diff
changeset
 | 
696  | 
(meth_name >> (Source o single) ||  | 
| 59666 | 697  | 
Scan.ahead Parse.cartouche |-- Parse.not_eof >> (fn tok =>  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61812 
diff
changeset
 | 
698  | 
        Source (Token.make_src ("cartouche", Position.none) [tok])) ||
 | 
| 59666 | 699  | 
      Parse.$$$ "(" |-- Parse.!!! (meth0 --| Parse.$$$ ")")) x
 | 
700  | 
and meth4 x =  | 
|
701  | 
(meth5 -- Parse.position (Parse.$$$ "?")  | 
|
702  | 
>> (fn (m, (_, pos)) => Combinator (combinator_info [pos], Try, [m])) ||  | 
|
703  | 
meth5 -- Parse.position (Parse.$$$ "+")  | 
|
704  | 
>> (fn (m, (_, pos)) => Combinator (combinator_info [pos], Repeat1, [m])) ||  | 
|
| 
59914
 
d1ddcd8df4e4
proper treatment of internal method name as already checked Token.src;
 
wenzelm 
parents: 
59909 
diff
changeset
 | 
705  | 
meth5 -- (Parse.position (Parse.$$$ "[") --  | 
| 
 
d1ddcd8df4e4
proper treatment of internal method name as already checked Token.src;
 
wenzelm 
parents: 
59909 
diff
changeset
 | 
706  | 
Scan.optional Parse.nat 1 -- Parse.position (Parse.$$$ "]"))  | 
| 59666 | 707  | 
>> (fn (m, (((_, pos1), n), (_, pos2))) =>  | 
708  | 
Combinator (combinator_info [pos1, pos2], Select_Goals n, [m])) ||  | 
|
709  | 
meth5) x  | 
|
710  | 
and meth3 x =  | 
|
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61812 
diff
changeset
 | 
711  | 
(meth_name ::: Parse.args1 is_symid_meth >> Source ||  | 
| 59666 | 712  | 
meth4) x  | 
713  | 
and meth2 x =  | 
|
714  | 
(Parse.enum1_positions "," meth3  | 
|
715  | 
>> (fn ([m], _) => m | (ms, ps) => Combinator (combinator_info ps, Then, ms))) x  | 
|
716  | 
and meth1 x =  | 
|
717  | 
(Parse.enum1_positions ";" meth2  | 
|
718  | 
>> (fn ([m], _) => m | (ms, ps) => Combinator (combinator_info ps, Then_All_New, ms))) x  | 
|
719  | 
and meth0 x =  | 
|
720  | 
(Parse.enum1_positions "|" meth1  | 
|
721  | 
>> (fn ([m], _) => m | (ms, ps) => Combinator (combinator_info ps, Orelse, ms))) x;  | 
|
| 59981 | 722  | 
|
723  | 
val meth =  | 
|
724  | 
nth [meth0, meth1, meth2, meth3, meth4, meth5] pri  | 
|
725  | 
        handle General.Subscript => raise Fail ("Bad method parser priority " ^ string_of_int pri);
 | 
|
726  | 
in Scan.trace meth >> (fn (m, toks) => (m, Token.range_of toks)) end;  | 
|
| 
27813
 
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27751 
diff
changeset
 | 
727  | 
|
| 59981 | 728  | 
val parse = parser 4;  | 
| 49866 | 729  | 
|
| 
55761
 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 
wenzelm 
parents: 
55742 
diff
changeset
 | 
730  | 
end;  | 
| 
27813
 
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27751 
diff
changeset
 | 
731  | 
|
| 
 
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27751 
diff
changeset
 | 
732  | 
|
| 18708 | 733  | 
(* theory setup *)  | 
| 5824 | 734  | 
|
| 53171 | 735  | 
val _ = Theory.setup  | 
| 56204 | 736  | 
 (setup @{binding fail} (Scan.succeed (K fail)) "force failure" #>
 | 
737  | 
  setup @{binding succeed} (Scan.succeed (K succeed)) "succeed" #>
 | 
|
| 60554 | 738  | 
  setup @{binding sleep} (Scan.lift Parse.real >> (fn s => fn _ => fn _ => sleep (seconds s)))
 | 
739  | 
"succeed after delay (in seconds)" #>  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
740  | 
  setup @{binding "-"} (Scan.succeed (K (SIMPLE_METHOD all_tac)))
 | 
| 60578 | 741  | 
"insert current facts, nothing else" #>  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
742  | 
  setup @{binding goal_cases} (Scan.lift (Scan.repeat Args.name_token) >> (fn names => fn _ =>
 | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
743  | 
CONTEXT_METHOD (fn _ => fn (ctxt, st) =>  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
744  | 
(case drop (Thm.nprems_of st) names of  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
745  | 
[] => NONE  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
746  | 
| bad =>  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
747  | 
if detect_closure_state st then NONE  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
748  | 
else  | 
| 61843 | 749  | 
            SOME (fn () => ("Excessive case name(s): " ^ commas_quote (map Token.content_of bad) ^
 | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
750  | 
Position.here (Position.set_range (Token.range_of bad)))))  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
751  | 
|> (fn SOME msg => Seq.single (Seq.Error msg)  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
752  | 
| NONE => goal_cases_tac (map Token.content_of names) (ctxt, st)))))  | 
| 61164 | 753  | 
"bind cases for goals" #>  | 
| 56204 | 754  | 
  setup @{binding insert} (Attrib.thms >> (K o insert))
 | 
| 60578 | 755  | 
"insert theorems, ignoring facts" #>  | 
| 58957 | 756  | 
  setup @{binding intro} (Attrib.thms >> (fn ths => fn ctxt => intro ctxt ths))
 | 
| 30515 | 757  | 
"repeatedly apply introduction rules" #>  | 
| 58957 | 758  | 
  setup @{binding elim} (Attrib.thms >> (fn ths => fn ctxt => elim ctxt ths))
 | 
| 30515 | 759  | 
"repeatedly apply elimination rules" #>  | 
| 56204 | 760  | 
  setup @{binding unfold} (Attrib.thms >> unfold_meth) "unfold definitions" #>
 | 
761  | 
  setup @{binding fold} (Attrib.thms >> fold_meth) "fold definitions" #>
 | 
|
762  | 
  setup @{binding atomize} (Scan.lift (Args.mode "full") >> atomize)
 | 
|
| 30515 | 763  | 
"present local premises as object-level statements" #>  | 
| 56204 | 764  | 
  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
 | 
765  | 
"apply some intro/elim rule" #>  | 
| 56204 | 766  | 
  setup @{binding erule} (xrule_meth erule) "apply rule in elimination manner (improper)" #>
 | 
767  | 
  setup @{binding drule} (xrule_meth drule) "apply rule in destruct manner (improper)" #>
 | 
|
768  | 
  setup @{binding frule} (xrule_meth frule) "apply rule in forward manner (improper)" #>
 | 
|
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59067 
diff
changeset
 | 
769  | 
  setup @{binding this} (Scan.succeed this) "apply current facts as rules" #>
 | 
| 56204 | 770  | 
  setup @{binding fact} (Attrib.thms >> fact) "composition by facts from context" #>
 | 
771  | 
  setup @{binding assumption} (Scan.succeed assumption)
 | 
|
| 30515 | 772  | 
"proof by assumption, preferring facts" #>  | 
| 56204 | 773  | 
  setup @{binding rename_tac} (Args.goal_spec -- Scan.lift (Scan.repeat1 Args.name) >>
 | 
| 52732 | 774  | 
(fn (quant, xs) => K (SIMPLE_METHOD'' quant (rename_tac xs))))  | 
| 30515 | 775  | 
"rename parameters of goal" #>  | 
| 56204 | 776  | 
  setup @{binding rotate_tac} (Args.goal_spec -- Scan.lift (Scan.optional Parse.int 1) >>
 | 
| 52732 | 777  | 
(fn (quant, i) => K (SIMPLE_METHOD'' quant (rotate_tac i))))  | 
| 30515 | 778  | 
"rotate assumptions of goal" #>  | 
| 
58018
 
beb4b7c0bb30
proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
 
wenzelm 
parents: 
58016 
diff
changeset
 | 
779  | 
  setup @{binding tactic} (parse_tactic >> (K o METHOD))
 | 
| 30515 | 780  | 
"ML tactic as proof method" #>  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
781  | 
  setup @{binding raw_tactic} (parse_tactic >> (fn tac => fn _ => CONTEXT_TACTIC o tac))
 | 
| 53171 | 782  | 
"ML tactic as raw proof method");  | 
| 5824 | 783  | 
|
784  | 
||
| 16145 | 785  | 
(*final declarations of this structure!*)  | 
786  | 
val unfold = unfold_meth;  | 
|
787  | 
val fold = fold_meth;  | 
|
788  | 
||
| 5824 | 789  | 
end;  | 
790  | 
||
| 61834 | 791  | 
structure Basic_Method: BASIC_METHOD = Method;  | 
792  | 
open Basic_Method;  | 
|
793  | 
||
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61834 
diff
changeset
 | 
794  | 
val CONTEXT_METHOD = Method.CONTEXT_METHOD;  | 
| 
30508
 
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
795  | 
val METHOD = Method.METHOD;  | 
| 
 
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
 
wenzelm 
parents: 
30466 
diff
changeset
 | 
796  | 
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
 | 
797  | 
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
 | 
798  | 
val SIMPLE_METHOD'' = Method.SIMPLE_METHOD'';  |