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