| author | wenzelm | 
| Tue, 28 Oct 2014 09:57:12 +0100 | |
| changeset 58797 | 6d71f19a9fd6 | 
| parent 58068 | d6f29bf9b783 | 
| child 58837 | e84d900cd287 | 
| permissions | -rw-r--r-- | 
| 5824 | 1 | (* Title: Pure/Isar/method.ML | 
| 2 | Author: Markus Wenzel, TU Muenchen | |
| 3 | ||
| 17110 | 4 | Isar proof methods. | 
| 5824 | 5 | *) | 
| 6 | ||
| 7 | signature METHOD = | |
| 8 | sig | |
| 58002 | 9 | type method = thm list -> cases_tactic | 
| 18227 | 10 | val METHOD_CASES: (thm list -> cases_tactic) -> method | 
| 17110 | 11 | val METHOD: (thm list -> tactic) -> method | 
| 12 | val fail: method | |
| 13 | val succeed: method | |
| 14 | val insert_tac: thm list -> int -> tactic | |
| 15 | val insert: thm list -> method | |
| 16 | val insert_facts: method | |
| 17 | val SIMPLE_METHOD: tactic -> method | |
| 21592 | 18 | val SIMPLE_METHOD': (int -> tactic) -> method | 
| 19 | val SIMPLE_METHOD'': ((int -> tactic) -> tactic) -> (int -> tactic) -> method | |
| 52059 | 20 | val cheating: Proof.context -> bool -> method | 
| 17110 | 21 | val intro: thm list -> method | 
| 22 | val elim: thm list -> method | |
| 20289 | 23 | val unfold: thm list -> Proof.context -> method | 
| 24 | 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 | 25 | val atomize: bool -> Proof.context -> method | 
| 17110 | 26 | val this: method | 
| 20289 | 27 | 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 | 28 | 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 | 29 | val all_assm_tac: Proof.context -> tactic | 
| 20289 | 30 | val assumption: Proof.context -> method | 
| 46466 
61c7214b4885
tuned signature, according to actual usage of these operations;
 wenzelm parents: 
45375diff
changeset | 31 | val rule_trace: bool Config.T | 
| 20289 | 32 | 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 | 33 | 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 | 34 | val some_rule_tac: Proof.context -> thm list -> thm list -> int -> tactic | 
| 25270 | 35 | val intros_tac: thm list -> thm list -> tactic | 
| 36093 
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
 ballarin parents: 
33522diff
changeset | 36 | val try_intros_tac: thm list -> thm list -> tactic | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53171diff
changeset | 37 | 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 | 38 | 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 | 39 | 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 | 40 | 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 | 41 | val set_tactic: (morphism -> thm list -> tactic) -> Context.generic -> Context.generic | 
| 55765 | 42 | type combinator_info | 
| 43 | val no_combinator_info: combinator_info | |
| 58005 | 44 | datatype combinator = Then | Orelse | Try | Repeat1 | Select_Goals of int | 
| 5824 | 45 | datatype text = | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
58007diff
changeset | 46 | Source of Token.src | | 
| 32193 
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
 wenzelm parents: 
32091diff
changeset | 47 | Basic of Proof.context -> method | | 
| 58005 | 48 | Combinator of combinator_info * combinator * text list | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
58007diff
changeset | 49 | val map_source: (Token.src -> Token.src) -> text -> text | 
| 54883 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 wenzelm parents: 
54742diff
changeset | 50 | val primitive_text: (Proof.context -> thm -> thm) -> text | 
| 17857 | 51 | val succeed_text: text | 
| 17110 | 52 | val default_text: text | 
| 53 | val this_text: text | |
| 54 | val done_text: text | |
| 17356 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 55 | val sorry_text: bool -> text | 
| 32193 
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
 wenzelm parents: 
32091diff
changeset | 56 | val finish_text: text option * bool -> text | 
| 55742 
a989bdaf8121
modernized Method.check_name/check_source (with reports) vs. strict Method.the_method (without interning nor reports), e.g. relevant for semantic completion;
 wenzelm parents: 
55709diff
changeset | 57 | val print_methods: Proof.context -> unit | 
| 
a989bdaf8121
modernized Method.check_name/check_source (with reports) vs. strict Method.the_method (without interning nor reports), e.g. relevant for semantic completion;
 wenzelm parents: 
55709diff
changeset | 58 | val check_name: Proof.context -> xstring * Position.T -> string | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
58007diff
changeset | 59 | val method_syntax: (Proof.context -> method) context_parser -> | 
| 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
58007diff
changeset | 60 | Token.src -> Proof.context -> method | 
| 30512 | 61 | val setup: binding -> (Proof.context -> method) context_parser -> string -> theory -> theory | 
| 57935 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 wenzelm parents: 
57863diff
changeset | 62 | val local_setup: binding -> (Proof.context -> method) context_parser -> string -> | 
| 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 wenzelm parents: 
57863diff
changeset | 63 | local_theory -> local_theory | 
| 57941 
57200bdc2aa7
localized command 'method_setup' and 'attribute_setup';
 wenzelm parents: 
57938diff
changeset | 64 | val method_setup: bstring * Position.T -> Symbol_Pos.source -> string -> | 
| 
57200bdc2aa7
localized command 'method_setup' and 'attribute_setup';
 wenzelm parents: 
57938diff
changeset | 65 | local_theory -> local_theory | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
58007diff
changeset | 66 | val method: Proof.context -> Token.src -> Proof.context -> method | 
| 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
58007diff
changeset | 67 | val method_closure: Proof.context -> Token.src -> Token.src | 
| 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
58007diff
changeset | 68 | val method_cmd: Proof.context -> Token.src -> Proof.context -> method | 
| 58007 
671c607fb4af
just one context for Method.evaluate (in contrast to a989bdaf8121, but in accordance to old global situation);
 wenzelm parents: 
58006diff
changeset | 69 | val evaluate: text -> Proof.context -> method | 
| 58048 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58034diff
changeset | 70 |   type modifier = {init: Proof.context -> Proof.context, attribute: attribute, pos: Position.T}
 | 
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58034diff
changeset | 71 | val modifier: attribute -> Position.T -> modifier | 
| 58029 
2137e60b6f6d
clarified Method.section: explicit declaration with static closure;
 wenzelm parents: 
58026diff
changeset | 72 | val section: modifier parser list -> declaration context_parser | 
| 
2137e60b6f6d
clarified Method.section: explicit declaration with static closure;
 wenzelm parents: 
58026diff
changeset | 73 | val sections: modifier parser list -> declaration list context_parser | 
| 49889 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49866diff
changeset | 74 | type text_range = text * Position.range | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49866diff
changeset | 75 | val text: text_range option -> text option | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49866diff
changeset | 76 | val position: text_range option -> Position.T | 
| 55795 | 77 | val reports_of: text_range -> Position.report list | 
| 78 | val report: text_range -> unit | |
| 55761 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 wenzelm parents: 
55742diff
changeset | 79 | val parse: text_range parser | 
| 5824 | 80 | end; | 
| 81 | ||
| 82 | structure Method: METHOD = | |
| 83 | struct | |
| 84 | ||
| 12324 
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
 wenzelm parents: 
12311diff
changeset | 85 | (** proof methods **) | 
| 
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
 wenzelm parents: 
12311diff
changeset | 86 | |
| 58006 | 87 | (* method *) | 
| 11731 | 88 | |
| 58002 | 89 | type method = thm list -> cases_tactic; | 
| 11731 | 90 | |
| 58030 | 91 | fun METHOD_CASES tac : method = fn facts => Seq.THEN (ALLGOALS Goal.conjunction_tac, tac facts); | 
| 92 | fun METHOD tac : method = fn facts => NO_CASES (ALLGOALS Goal.conjunction_tac THEN tac facts); | |
| 5824 | 93 | |
| 94 | val fail = METHOD (K no_tac); | |
| 95 | val succeed = METHOD (K all_tac); | |
| 96 | ||
| 97 | ||
| 17110 | 98 | (* insert facts *) | 
| 7419 | 99 | |
| 100 | local | |
| 5824 | 101 | |
| 21579 | 102 | fun cut_rule_tac rule = | 
| 52732 | 103 | rtac (Drule.forall_intr_vars rule COMP_INCR revcut_rl); | 
| 6981 | 104 | |
| 7419 | 105 | in | 
| 5824 | 106 | |
| 51552 
c713c9505f68
clarified Skip_Proof.cheat_tac: more standard tactic;
 wenzelm parents: 
51383diff
changeset | 107 | fun insert_tac [] _ = all_tac | 
| 7419 | 108 | | insert_tac facts i = EVERY (map (fn th => cut_rule_tac th i) facts); | 
| 6981 | 109 | |
| 7555 | 110 | val insert_facts = METHOD (ALLGOALS o insert_tac); | 
| 7664 | 111 | fun insert thms = METHOD (fn _ => ALLGOALS (insert_tac thms)); | 
| 7419 | 112 | |
| 9706 | 113 | fun SIMPLE_METHOD tac = METHOD (fn facts => ALLGOALS (insert_tac facts) THEN tac); | 
| 21592 | 114 | fun SIMPLE_METHOD'' quant tac = METHOD (fn facts => quant (insert_tac facts THEN' tac)); | 
| 115 | val SIMPLE_METHOD' = SIMPLE_METHOD'' HEADGOAL; | |
| 9706 | 116 | |
| 12324 
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
 wenzelm parents: 
12311diff
changeset | 117 | end; | 
| 
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
 wenzelm parents: 
12311diff
changeset | 118 | |
| 9706 | 119 | |
| 17356 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 120 | (* cheating *) | 
| 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 121 | |
| 52059 | 122 | fun cheating ctxt int = METHOD (fn _ => fn st => | 
| 123 | if int orelse Config.get ctxt quick_and_dirty then | |
| 51552 
c713c9505f68
clarified Skip_Proof.cheat_tac: more standard tactic;
 wenzelm parents: 
51383diff
changeset | 124 | ALLGOALS Skip_Proof.cheat_tac st | 
| 
c713c9505f68
clarified Skip_Proof.cheat_tac: more standard tactic;
 wenzelm parents: 
51383diff
changeset | 125 | else error "Cheating requires quick_and_dirty mode!"); | 
| 17356 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 126 | |
| 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 127 | |
| 17110 | 128 | (* unfold intro/elim rules *) | 
| 129 | ||
| 52732 | 130 | fun intro ths = SIMPLE_METHOD' (CHANGED_PROP o REPEAT_ALL_NEW (match_tac ths)); | 
| 131 | fun elim ths = SIMPLE_METHOD' (CHANGED_PROP o REPEAT_ALL_NEW (ematch_tac ths)); | |
| 17110 | 132 | |
| 133 | ||
| 12384 | 134 | (* unfold/fold definitions *) | 
| 135 | ||
| 35624 | 136 | fun unfold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (Local_Defs.unfold_tac ctxt ths)); | 
| 137 | fun fold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (Local_Defs.fold_tac ctxt ths)); | |
| 6532 | 138 | |
| 12384 | 139 | |
| 12829 | 140 | (* atomize rule statements *) | 
| 141 | ||
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53171diff
changeset | 142 | fun atomize false ctxt = | 
| 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53171diff
changeset | 143 | 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 | 144 | | atomize true ctxt = | 
| 58002 | 145 | NO_CASES o K (HEADGOAL (CHANGED_PROP o Object_Logic.full_atomize_tac ctxt)); | 
| 12829 | 146 | |
| 147 | ||
| 18039 | 148 | (* this -- resolve facts directly *) | 
| 12384 | 149 | |
| 52732 | 150 | val this = METHOD (EVERY o map (HEADGOAL o rtac)); | 
| 9484 | 151 | |
| 152 | ||
| 18039 | 153 | (* fact -- composition by facts from context *) | 
| 154 | ||
| 42360 | 155 | 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 | 156 | | fact rules ctxt = SIMPLE_METHOD' (Proof_Context.fact_tac ctxt rules); | 
| 18039 | 157 | |
| 158 | ||
| 17110 | 159 | (* assumption *) | 
| 7419 | 160 | |
| 161 | local | |
| 162 | ||
| 19778 | 163 | fun cond_rtac cond rule = SUBGOAL (fn (prop, i) => | 
| 164 | if cond (Logic.strip_assums_concl prop) | |
| 52732 | 165 | then rtac rule i else no_tac); | 
| 7419 | 166 | |
| 29857 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29301diff
changeset | 167 | in | 
| 
2cc976ed8a3c
FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
 kleing parents: 
29301diff
changeset | 168 | |
| 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 | 169 | fun assm_tac ctxt = | 
| 17110 | 170 | assume_tac APPEND' | 
| 23349 | 171 | Goal.assume_rule_tac ctxt APPEND' | 
| 19778 | 172 | cond_rtac (can Logic.dest_equals) Drule.reflexive_thm APPEND' | 
| 173 | cond_rtac (can Logic.dest_term) Drule.termI; | |
| 17110 | 174 | |
| 49846 
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
 wenzelm parents: 
48992diff
changeset | 175 | fun all_assm_tac ctxt = | 
| 
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
 wenzelm parents: 
48992diff
changeset | 176 | let | 
| 
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
 wenzelm parents: 
48992diff
changeset | 177 | fun tac i st = | 
| 
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
 wenzelm parents: 
48992diff
changeset | 178 | if i > Thm.nprems_of st then all_tac st | 
| 
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
 wenzelm parents: 
48992diff
changeset | 179 | 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 | 180 | 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 | 181 | |
| 23349 | 182 | 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 | 183 | (fn [] => assm_tac ctxt | 
| 23349 | 184 | | [fact] => solve_tac [fact] | 
| 185 | | _ => K no_tac)); | |
| 186 | ||
| 49846 
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
 wenzelm parents: 
48992diff
changeset | 187 | fun finish immed ctxt = | 
| 
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
 wenzelm parents: 
48992diff
changeset | 188 | METHOD (K ((if immed then all_assm_tac ctxt else all_tac) THEN flexflex_tac)); | 
| 7419 | 189 | |
| 190 | end; | |
| 191 | ||
| 192 | ||
| 17110 | 193 | (* rule etc. -- single-step refinements *) | 
| 12347 | 194 | |
| 56204 | 195 | val rule_trace = Attrib.setup_config_bool @{binding rule_trace} (fn _ => false);
 | 
| 12347 | 196 | |
| 17110 | 197 | fun trace ctxt rules = | 
| 41379 | 198 | if Config.get ctxt rule_trace andalso not (null rules) then | 
| 51584 | 199 | Pretty.big_list "rules:" (map (Display.pretty_thm_item ctxt) rules) | 
| 21962 | 200 | |> Pretty.string_of |> tracing | 
| 201 | else (); | |
| 12347 | 202 | |
| 203 | local | |
| 204 | ||
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53171diff
changeset | 205 | fun gen_rule_tac tac ctxt rules facts = | 
| 18841 | 206 | (fn i => fn st => | 
| 207 | if null facts then tac rules i st | |
| 208 | else Seq.maps (fn rule => (tac o single) rule i st) (Drule.multi_resolves facts rules)) | |
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53171diff
changeset | 209 | THEN_ALL_NEW Goal.norm_hhf_tac ctxt; | 
| 7130 | 210 | |
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53171diff
changeset | 211 | fun gen_arule_tac tac ctxt j rules facts = | 
| 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53171diff
changeset | 212 | EVERY' (gen_rule_tac tac ctxt rules facts :: replicate j assume_tac); | 
| 10744 | 213 | |
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53171diff
changeset | 214 | fun gen_some_rule_tac tac ctxt arg_rules facts = SUBGOAL (fn (goal, i) => | 
| 11785 | 215 | let | 
| 216 | val rules = | |
| 217 | if not (null arg_rules) then arg_rules | |
| 33369 | 218 | else flat (Context_Rules.find_rules false facts goal ctxt) | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53171diff
changeset | 219 | in trace ctxt rules; tac ctxt rules facts i end); | 
| 10309 | 220 | |
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53171diff
changeset | 221 | 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 | 222 | fun meth' tac x y z = METHOD (HEADGOAL o tac x y z); | 
| 8220 | 223 | |
| 7419 | 224 | in | 
| 225 | ||
| 52732 | 226 | val rule_tac = gen_rule_tac resolve_tac; | 
| 10744 | 227 | val rule = meth rule_tac; | 
| 228 | 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 | 229 | val some_rule = meth some_rule_tac; | 
| 10744 | 230 | |
| 52732 | 231 | val erule = meth' (gen_arule_tac eresolve_tac); | 
| 232 | val drule = meth' (gen_arule_tac dresolve_tac); | |
| 233 | val frule = meth' (gen_arule_tac forward_tac); | |
| 5824 | 234 | |
| 7419 | 235 | end; | 
| 236 | ||
| 237 | ||
| 25270 | 238 | (* intros_tac -- pervasive search spanned by intro rules *) | 
| 239 | ||
| 36093 
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
 ballarin parents: 
33522diff
changeset | 240 | fun gen_intros_tac goals intros facts = | 
| 
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
 ballarin parents: 
33522diff
changeset | 241 | goals (insert_tac facts THEN' | 
| 25270 | 242 | REPEAT_ALL_NEW (resolve_tac intros)) | 
| 243 | THEN Tactic.distinct_subgoals_tac; | |
| 244 | ||
| 36093 
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
 ballarin parents: 
33522diff
changeset | 245 | 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 | 246 | val try_intros_tac = gen_intros_tac TRYALL; | 
| 25270 | 247 | |
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
37198diff
changeset | 248 | |
| 58016 | 249 | |
| 250 | (** method syntax **) | |
| 251 | ||
| 252 | (* context data *) | |
| 8351 | 253 | |
| 58016 | 254 | structure Data = Generic_Data | 
| 26472 
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
 wenzelm parents: 
26463diff
changeset | 255 | ( | 
| 58016 | 256 | type T = | 
| 257 | ((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: 
58016diff
changeset | 258 | (morphism -> thm list -> tactic) option; (*ML tactic*) | 
| 58016 | 259 | val empty : T = (Name_Space.empty_table "method", NONE); | 
| 260 | val extend = I; | |
| 261 | fun merge ((tab, tac), (tab', tac')) : T = | |
| 262 | (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: 
26463diff
changeset | 263 | ); | 
| 
9afdd61cf528
ml_tactic: non-critical version via proof data and thread data;
 wenzelm parents: 
26463diff
changeset | 264 | |
| 58016 | 265 | val get_methods = fst o Data.get; | 
| 266 | val map_methods = Data.map o apfst; | |
| 267 | ||
| 268 | ||
| 269 | (* ML tactic *) | |
| 270 | ||
| 271 | val set_tactic = Data.map o apsnd o K o SOME; | |
| 272 | ||
| 273 | fun the_tactic context = | |
| 274 | (case snd (Data.get context) of | |
| 275 | SOME tac => tac | |
| 276 | | NONE => raise Fail "Undefined ML tactic"); | |
| 8351 | 277 | |
| 58018 
beb4b7c0bb30
proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
 wenzelm parents: 
58016diff
changeset | 278 | 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 | 279 | 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 | 280 | 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 | 281 | let | 
| 
beb4b7c0bb30
proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
 wenzelm parents: 
58016diff
changeset | 282 | val context' = context |> | 
| 
beb4b7c0bb30
proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
 wenzelm parents: 
58016diff
changeset | 283 | ML_Context.expression (#pos source) | 
| 58024 | 284 | "fun tactic (morphism: Morphism.morphism) (facts: thm list) : tactic" | 
| 58018 
beb4b7c0bb30
proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
 wenzelm parents: 
58016diff
changeset | 285 | "Method.set_tactic tactic" (ML_Lex.read_source false source); | 
| 
beb4b7c0bb30
proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
 wenzelm parents: 
58016diff
changeset | 286 | 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 | 287 | in | 
| 
beb4b7c0bb30
proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
 wenzelm parents: 
58016diff
changeset | 288 | fn phi => | 
| 
beb4b7c0bb30
proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
 wenzelm parents: 
58016diff
changeset | 289 | 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: 
58016diff
changeset | 290 | end)) >> (fn decl => Morphism.form (the_tactic (Morphism.form decl context)))); | 
| 8351 | 291 | |
| 292 | ||
| 17110 | 293 | (* method text *) | 
| 294 | ||
| 55765 | 295 | datatype combinator_info = Combinator_Info of {keywords: Position.T list};
 | 
| 296 | fun combinator_info keywords = Combinator_Info {keywords = keywords};
 | |
| 297 | val no_combinator_info = combinator_info []; | |
| 298 | ||
| 58005 | 299 | datatype combinator = Then | Orelse | Try | Repeat1 | Select_Goals of int; | 
| 300 | ||
| 17110 | 301 | datatype text = | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
58007diff
changeset | 302 | Source of Token.src | | 
| 32193 
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
 wenzelm parents: 
32091diff
changeset | 303 | Basic of Proof.context -> method | | 
| 58005 | 304 | Combinator of combinator_info * combinator * text list; | 
| 55765 | 305 | |
| 58006 | 306 | fun map_source f (Source src) = Source (f src) | 
| 307 | | map_source _ (Basic meth) = Basic meth | |
| 308 | | map_source f (Combinator (info, comb, txts)) = Combinator (info, comb, map (map_source f) txts); | |
| 309 | ||
| 54883 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 wenzelm parents: 
54742diff
changeset | 310 | 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 | 311 | val succeed_text = Basic (K succeed); | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
58007diff
changeset | 312 | val default_text = Source (Token.src ("default", Position.none) []);
 | 
| 32193 
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
 wenzelm parents: 
32091diff
changeset | 313 | val this_text = Basic (K this); | 
| 
c314b4836031
basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
 wenzelm parents: 
32091diff
changeset | 314 | val done_text = Basic (K (SIMPLE_METHOD all_tac)); | 
| 52059 | 315 | fun sorry_text int = Basic (fn ctxt => cheating ctxt int); | 
| 17110 | 316 | |
| 49846 
8fae089f5a0c
refined Proof.the_finished_goal with more informative error;
 wenzelm parents: 
48992diff
changeset | 317 | fun finish_text (NONE, immed) = Basic (finish immed) | 
| 58005 | 318 | | finish_text (SOME txt, immed) = | 
| 319 | Combinator (no_combinator_info, Then, [txt, Basic (finish immed)]); | |
| 17110 | 320 | |
| 321 | ||
| 322 | (* method definitions *) | |
| 5824 | 323 | |
| 57937 | 324 | fun transfer_methods ctxt = | 
| 57935 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 wenzelm parents: 
57863diff
changeset | 325 | let | 
| 58016 | 326 | val meths0 = get_methods (Context.Theory (Proof_Context.theory_of ctxt)); | 
| 327 | val meths' = Name_Space.merge_tables (meths0, get_methods (Context.Proof ctxt)); | |
| 328 | 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: 
55709diff
changeset | 329 | |
| 
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 | 330 | fun print_methods ctxt = | 
| 22846 | 331 | let | 
| 58016 | 332 | val meths = get_methods (Context.Proof ctxt); | 
| 50301 | 333 | fun prt_meth (name, (_, "")) = Pretty.mark_str name | 
| 42813 
6c841fa92fa2
optional description for 'attribute_setup' and 'method_setup';
 wenzelm parents: 
42616diff
changeset | 334 | | prt_meth (name, (_, comment)) = | 
| 50301 | 335 | Pretty.block | 
| 336 | (Pretty.mark_str name :: Pretty.str ":" :: Pretty.brk 2 :: Pretty.text comment); | |
| 22846 | 337 | in | 
| 56052 | 338 | [Pretty.big_list "methods:" (map prt_meth (Name_Space.markup_table ctxt meths))] | 
| 56334 
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
 wenzelm parents: 
56278diff
changeset | 339 | |> Pretty.writeln_chunks | 
| 22846 | 340 | end; | 
| 7611 | 341 | |
| 57935 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 wenzelm parents: 
57863diff
changeset | 342 | |
| 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 wenzelm parents: 
57863diff
changeset | 343 | (* define *) | 
| 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 wenzelm parents: 
57863diff
changeset | 344 | |
| 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 wenzelm parents: 
57863diff
changeset | 345 | fun define_global binding meth comment thy = | 
| 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 wenzelm parents: 
57863diff
changeset | 346 | let | 
| 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 wenzelm parents: 
57863diff
changeset | 347 | val context = Context.Theory thy; | 
| 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 wenzelm parents: 
57863diff
changeset | 348 | val (name, meths') = | 
| 58016 | 349 | Name_Space.define context true (binding, (meth, comment)) (get_methods context); | 
| 350 | in (name, Context.the_theory (map_methods (K meths') context)) end; | |
| 57935 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 wenzelm parents: 
57863diff
changeset | 351 | |
| 57941 
57200bdc2aa7
localized command 'method_setup' and 'attribute_setup';
 wenzelm parents: 
57938diff
changeset | 352 | fun define binding meth comment = | 
| 
57200bdc2aa7
localized command 'method_setup' and 'attribute_setup';
 wenzelm parents: 
57938diff
changeset | 353 | Local_Theory.background_theory_result (define_global binding meth comment) | 
| 
57200bdc2aa7
localized command 'method_setup' and 'attribute_setup';
 wenzelm parents: 
57938diff
changeset | 354 | #-> (fn name => | 
| 
57200bdc2aa7
localized command 'method_setup' and 'attribute_setup';
 wenzelm parents: 
57938diff
changeset | 355 | Local_Theory.map_contexts (K transfer_methods) | 
| 58016 | 356 | #> Local_Theory.generic_alias map_methods binding name | 
| 57941 
57200bdc2aa7
localized command 'method_setup' and 'attribute_setup';
 wenzelm parents: 
57938diff
changeset | 357 | #> pair name); | 
| 31304 | 358 | |
| 55997 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55917diff
changeset | 359 | |
| 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55917diff
changeset | 360 | (* check *) | 
| 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55917diff
changeset | 361 | |
| 58016 | 362 | fun check_name ctxt = | 
| 363 | let val context = Context.Proof ctxt | |
| 364 | in #1 o Name_Space.check context (get_methods context) end; | |
| 365 | ||
| 366 | fun check_src ctxt src = | |
| 367 | Token.check_src ctxt (get_methods (Context.Proof ctxt)) src; | |
| 55997 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55917diff
changeset | 368 | |
| 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55917diff
changeset | 369 | |
| 30512 | 370 | (* method setup *) | 
| 371 | ||
| 57935 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 wenzelm parents: 
57863diff
changeset | 372 | fun method_syntax scan src ctxt : method = | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
58007diff
changeset | 373 | 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 | 374 | |
| 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 wenzelm parents: 
57863diff
changeset | 375 | fun setup binding scan comment = define_global binding (method_syntax scan) comment #> snd; | 
| 
c578f3a37a67
localized method definitions (see also f14c1248d064);
 wenzelm parents: 
57863diff
changeset | 376 | 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 | 377 | |
| 55828 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
55795diff
changeset | 378 | fun method_setup name source cmt = | 
| 57941 
57200bdc2aa7
localized command 'method_setup' and 'attribute_setup';
 wenzelm parents: 
57938diff
changeset | 379 |   (ML_Lex.read Position.none ("(" ^ ML_Syntax.make_binding name ^ ", ") @
 | 
| 
57200bdc2aa7
localized command 'method_setup' and 'attribute_setup';
 wenzelm parents: 
57938diff
changeset | 380 | ML_Lex.read_source false source @ | 
| 
57200bdc2aa7
localized command 'method_setup' and 'attribute_setup';
 wenzelm parents: 
57938diff
changeset | 381 |     ML_Lex.read Position.none (", " ^ ML_Syntax.print_string cmt ^ ")"))
 | 
| 
57200bdc2aa7
localized command 'method_setup' and 'attribute_setup';
 wenzelm parents: 
57938diff
changeset | 382 | |> ML_Context.expression (#pos source) | 
| 30544 | 383 | "val (name, scan, comment): binding * (Proof.context -> Proof.method) context_parser * string" | 
| 57941 
57200bdc2aa7
localized command 'method_setup' and 'attribute_setup';
 wenzelm parents: 
57938diff
changeset | 384 | "Context.map_proof (Method.local_setup name scan comment)" | 
| 
57200bdc2aa7
localized command 'method_setup' and 'attribute_setup';
 wenzelm parents: 
57938diff
changeset | 385 | |> Context.proof_map; | 
| 17356 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 386 | |
| 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 387 | |
| 58006 | 388 | (* prepare methods *) | 
| 389 | ||
| 390 | fun method ctxt = | |
| 58016 | 391 | let val table = get_methods (Context.Proof ctxt) | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
58007diff
changeset | 392 | in fn src => #1 (Name_Space.get table (#1 (Token.name_of_src src))) src end; | 
| 58006 | 393 | |
| 394 | fun method_closure ctxt0 src0 = | |
| 395 | let | |
| 58026 | 396 | val (src1, _) = check_src ctxt0 src0; | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
58007diff
changeset | 397 | val src2 = Token.init_assignable_src src1; | 
| 58006 | 398 | val ctxt = Context_Position.not_really ctxt0; | 
| 399 | val _ = Seq.pull (method ctxt src2 ctxt [] (Goal.protect 0 Drule.dummy_thm)); | |
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
58007diff
changeset | 400 | in Token.closure_src src2 end; | 
| 58006 | 401 | |
| 402 | fun method_cmd ctxt = method ctxt o method_closure ctxt; | |
| 403 | ||
| 404 | ||
| 58003 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 405 | (* evaluate method text *) | 
| 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 406 | |
| 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 407 | local | 
| 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 408 | |
| 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 409 | fun APPEND_CASES (meth: cases_tactic) (cases, st) = | 
| 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 410 | meth st |> Seq.map (fn (cases', st') => (cases @ cases', st')); | 
| 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 411 | |
| 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 412 | fun BYPASS_CASES (tac: tactic) (cases, st) = | 
| 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 413 | tac st |> Seq.map (pair cases); | 
| 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 414 | |
| 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 415 | val op THEN = Seq.THEN; | 
| 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 416 | |
| 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 417 | fun SELECT_GOALS n method = | 
| 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 418 | BYPASS_CASES | 
| 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 419 | (ALLGOALS Goal.conjunction_tac THEN PRIMITIVE (Goal.restrict 1 n) THEN Goal.conjunction_tac 1) | 
| 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 420 | THEN method | 
| 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 421 | THEN BYPASS_CASES (PRIMITIVE (Goal.unrestrict 1)); | 
| 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 422 | |
| 58005 | 423 | fun COMBINATOR1 comb [meth] = comb meth | 
| 424 | | COMBINATOR1 _ _ = raise Fail "Method combinator requires exactly one argument"; | |
| 425 | ||
| 426 | fun combinator Then = Seq.EVERY | |
| 427 | | combinator Orelse = Seq.FIRST | |
| 428 | | combinator Try = COMBINATOR1 Seq.TRY | |
| 429 | | combinator Repeat1 = COMBINATOR1 Seq.REPEAT1 | |
| 430 | | combinator (Select_Goals n) = COMBINATOR1 (SELECT_GOALS n); | |
| 58003 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 431 | |
| 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 432 | in | 
| 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 433 | |
| 58007 
671c607fb4af
just one context for Method.evaluate (in contrast to a989bdaf8121, but in accordance to old global situation);
 wenzelm parents: 
58006diff
changeset | 434 | fun evaluate text ctxt = | 
| 58003 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 435 | let | 
| 58007 
671c607fb4af
just one context for Method.evaluate (in contrast to a989bdaf8121, but in accordance to old global situation);
 wenzelm parents: 
58006diff
changeset | 436 | fun eval (Basic meth) = APPEND_CASES o meth ctxt | 
| 
671c607fb4af
just one context for Method.evaluate (in contrast to a989bdaf8121, but in accordance to old global situation);
 wenzelm parents: 
58006diff
changeset | 437 | | eval (Source src) = APPEND_CASES o method_cmd ctxt src ctxt | 
| 58005 | 438 | | eval (Combinator (_, c, txts)) = | 
| 439 | let | |
| 440 | 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: 
58006diff
changeset | 441 | 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: 
58006diff
changeset | 442 | 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: 
58002diff
changeset | 443 | val meth = eval text; | 
| 58007 
671c607fb4af
just one context for Method.evaluate (in contrast to a989bdaf8121, but in accordance to old global situation);
 wenzelm parents: 
58006diff
changeset | 444 | in fn facts => fn st => meth facts ([], st) end; | 
| 58003 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 445 | |
| 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 446 | end; | 
| 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 447 | |
| 
250ecd2502ad
clarifed Method.evaluate: turn text into semantic method (like Basic);
 wenzelm parents: 
58002diff
changeset | 448 | |
| 5884 | 449 | |
| 17110 | 450 | (** concrete syntax **) | 
| 5824 | 451 | |
| 58048 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58034diff
changeset | 452 | (* type modifier *) | 
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58034diff
changeset | 453 | |
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58034diff
changeset | 454 | type modifier = | 
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58034diff
changeset | 455 |   {init: Proof.context -> Proof.context, attribute: attribute, pos: Position.T};
 | 
| 5824 | 456 | |
| 58048 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58034diff
changeset | 457 | fun modifier attribute pos : modifier = {init = I, attribute = attribute, pos = pos};
 | 
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58034diff
changeset | 458 | |
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58034diff
changeset | 459 | |
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58034diff
changeset | 460 | (* sections *) | 
| 7268 | 461 | |
| 462 | local | |
| 463 | ||
| 58048 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58034diff
changeset | 464 | fun sect (modifier : modifier parser) = Scan.depend (fn context => | 
| 58068 | 465 | Scan.ahead Parse.not_eof -- modifier -- Scan.repeat (Scan.unless modifier Parse.xthm) >> | 
| 466 |     (fn ((tok, {init, attribute, pos}), xthms) =>
 | |
| 58029 
2137e60b6f6d
clarified Method.section: explicit declaration with static closure;
 wenzelm parents: 
58026diff
changeset | 467 | let | 
| 58068 | 468 | val decl = | 
| 469 | (case Token.get_value tok of | |
| 470 | SOME (Token.Declaration decl) => decl | |
| 471 | | _ => | |
| 472 | let | |
| 473 | val ctxt = Context.proof_of context; | |
| 474 | fun prep_att src = | |
| 475 | let | |
| 476 | val src' = Attrib.check_src ctxt src; | |
| 477 | val _ = List.app (Token.assign NONE) (Token.args_of_src src'); | |
| 478 | in src' end; | |
| 479 | val thms = | |
| 480 | map (fn (a, bs) => (Proof_Context.get_fact ctxt a, map prep_att bs)) xthms; | |
| 481 | val facts = | |
| 482 | Attrib.partial_evaluation ctxt [((Binding.name "dummy", []), thms)] | |
| 483 | |> map (fn (_, bs) => ((Binding.empty, [Attrib.internal (K attribute)]), bs)); | |
| 484 | val _ = | |
| 485 | Context_Position.report ctxt (Token.pos_of tok) | |
| 486 | (Markup.entity Markup.method_modifierN "" | |
| 487 | |> Markup.properties (Position.def_properties_of pos)); | |
| 488 | fun decl phi = | |
| 489 | Context.mapping I init #> | |
| 490 | Attrib.generic_notes "" (Attrib.transform_facts phi facts) #> snd; | |
| 491 | val _ = Token.assign (SOME (Token.Declaration decl)) tok; | |
| 492 | in decl end); | |
| 493 | in (Morphism.form decl context, decl) end)); | |
| 5824 | 494 | |
| 30540 | 495 | in | 
| 496 | ||
| 58029 
2137e60b6f6d
clarified Method.section: explicit declaration with static closure;
 wenzelm parents: 
58026diff
changeset | 497 | val section = sect o Scan.first; | 
| 
2137e60b6f6d
clarified Method.section: explicit declaration with static closure;
 wenzelm parents: 
58026diff
changeset | 498 | val sections = Scan.repeat o section; | 
| 5824 | 499 | |
| 7268 | 500 | end; | 
| 501 | ||
| 5824 | 502 | |
| 30515 | 503 | (* extra rule methods *) | 
| 504 | ||
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53171diff
changeset | 505 | fun xrule_meth meth = | 
| 36950 | 506 | 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 | 507 | (fn (n, ths) => fn ctxt => meth ctxt n ths); | 
| 30515 | 508 | |
| 509 | ||
| 55761 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 wenzelm parents: 
55742diff
changeset | 510 | (* text range *) | 
| 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 wenzelm parents: 
55742diff
changeset | 511 | |
| 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 wenzelm parents: 
55742diff
changeset | 512 | type text_range = text * Position.range; | 
| 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 wenzelm parents: 
55742diff
changeset | 513 | |
| 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 wenzelm parents: 
55742diff
changeset | 514 | fun text NONE = NONE | 
| 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 wenzelm parents: 
55742diff
changeset | 515 | | text (SOME (txt, _)) = SOME txt; | 
| 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 wenzelm parents: 
55742diff
changeset | 516 | |
| 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 wenzelm parents: 
55742diff
changeset | 517 | fun position NONE = Position.none | 
| 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 wenzelm parents: 
55742diff
changeset | 518 | | position (SOME (_, (pos, _))) = pos; | 
| 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 wenzelm parents: 
55742diff
changeset | 519 | |
| 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 wenzelm parents: 
55742diff
changeset | 520 | |
| 55795 | 521 | (* reports *) | 
| 522 | ||
| 523 | local | |
| 524 | ||
| 525 | fun keyword_positions (Source _) = [] | |
| 526 | | keyword_positions (Basic _) = [] | |
| 58005 | 527 |   | keyword_positions (Combinator (Combinator_Info {keywords}, _, texts)) =
 | 
| 528 | keywords @ maps keyword_positions texts; | |
| 55795 | 529 | |
| 530 | in | |
| 531 | ||
| 532 | fun reports_of ((text, (pos, _)): text_range) = | |
| 533 | (pos, Markup.language_method) :: | |
| 55917 
5438ed05e1c9
special treatment of method combinators like Args.$$$ keywords, although parsed via Parse.$$$;
 wenzelm parents: 
55828diff
changeset | 534 | 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 | 535 | (keyword_positions text); | 
| 55795 | 536 | |
| 537 | val report = Position.reports o reports_of; | |
| 538 | ||
| 539 | end; | |
| 540 | ||
| 541 | ||
| 27813 
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
 wenzelm parents: 
27751diff
changeset | 542 | (* outer parser *) | 
| 
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
 wenzelm parents: 
27751diff
changeset | 543 | |
| 
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
 wenzelm parents: 
27751diff
changeset | 544 | 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 | 545 | 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 | 546 | |
| 
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
 wenzelm parents: 
27751diff
changeset | 547 | local | 
| 
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
 wenzelm parents: 
27751diff
changeset | 548 | |
| 
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
 wenzelm parents: 
27751diff
changeset | 549 | fun meth4 x = | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
58007diff
changeset | 550 | (Parse.position Parse.xname >> (fn name => Source (Token.src name [])) || | 
| 55048 
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
 wenzelm parents: 
54883diff
changeset | 551 | Scan.ahead Parse.cartouche |-- Parse.not_eof >> (fn tok => | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
58007diff
changeset | 552 |     Source (Token.src ("cartouche", Token.pos_of tok) [tok])) ||
 | 
| 36950 | 553 |   Parse.$$$ "(" |-- Parse.!!! (meth0 --| Parse.$$$ ")")) x
 | 
| 27813 
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
 wenzelm parents: 
27751diff
changeset | 554 | and meth3 x = | 
| 55765 | 555 | (meth4 -- Parse.position (Parse.$$$ "?") | 
| 58005 | 556 | >> (fn (m, (_, pos)) => Combinator (combinator_info [pos], Try, [m])) || | 
| 55765 | 557 | meth4 -- Parse.position (Parse.$$$ "+") | 
| 58005 | 558 | >> (fn (m, (_, pos)) => Combinator (combinator_info [pos], Repeat1, [m])) || | 
| 55765 | 559 | meth4 -- | 
| 560 | (Parse.position (Parse.$$$ "[") -- Scan.optional Parse.nat 1 -- Parse.position (Parse.$$$ "]")) | |
| 561 | >> (fn (m, (((_, pos1), n), (_, pos2))) => | |
| 58005 | 562 | Combinator (combinator_info [pos1, pos2], Select_Goals n, [m])) || | 
| 27813 
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
 wenzelm parents: 
27751diff
changeset | 563 | meth4) x | 
| 
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
 wenzelm parents: 
27751diff
changeset | 564 | and meth2 x = | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
58007diff
changeset | 565 | (Parse.position Parse.xname -- Parse.args1 is_symid_meth >> (Source o uncurry Token.src) || | 
| 27813 
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
 wenzelm parents: 
27751diff
changeset | 566 | meth3) x | 
| 55765 | 567 | and meth1 x = | 
| 568 | (Parse.enum1_positions "," meth2 | |
| 58005 | 569 | >> (fn ([m], _) => m | (ms, ps) => Combinator (combinator_info ps, Then, ms))) x | 
| 55765 | 570 | and meth0 x = | 
| 571 | (Parse.enum1_positions "|" meth1 | |
| 58005 | 572 | >> (fn ([m], _) => m | (ms, ps) => Combinator (combinator_info ps, Orelse, ms))) x; | 
| 27813 
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
 wenzelm parents: 
27751diff
changeset | 573 | |
| 49866 | 574 | in | 
| 575 | ||
| 576 | val parse = | |
| 55709 
4e5a83a46ded
clarified Token.range_of in accordance to Symbol_Pos.range;
 wenzelm parents: 
55708diff
changeset | 577 | Scan.trace meth3 >> (fn (m, toks) => (m, Token.range_of toks)); | 
| 49866 | 578 | |
| 55761 
213b9811f59f
method language markup, e.g. relevant to prevent outer keyword completion;
 wenzelm parents: 
55742diff
changeset | 579 | end; | 
| 27813 
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
 wenzelm parents: 
27751diff
changeset | 580 | |
| 
96fbe385a0d0
unified Args.T with OuterLex.token, renamed some operations;
 wenzelm parents: 
27751diff
changeset | 581 | |
| 18708 | 582 | (* theory setup *) | 
| 5824 | 583 | |
| 53171 | 584 | val _ = Theory.setup | 
| 56204 | 585 |  (setup @{binding fail} (Scan.succeed (K fail)) "force failure" #>
 | 
| 586 |   setup @{binding succeed} (Scan.succeed (K succeed)) "succeed" #>
 | |
| 587 |   setup @{binding "-"} (Scan.succeed (K insert_facts))
 | |
| 30515 | 588 | "do nothing (insert current facts only)" #> | 
| 56204 | 589 |   setup @{binding insert} (Attrib.thms >> (K o insert))
 | 
| 30515 | 590 | "insert theorems, ignoring facts (improper)" #> | 
| 56204 | 591 |   setup @{binding intro} (Attrib.thms >> (K o intro))
 | 
| 30515 | 592 | "repeatedly apply introduction rules" #> | 
| 56204 | 593 |   setup @{binding elim} (Attrib.thms >> (K o elim))
 | 
| 30515 | 594 | "repeatedly apply elimination rules" #> | 
| 56204 | 595 |   setup @{binding unfold} (Attrib.thms >> unfold_meth) "unfold definitions" #>
 | 
| 596 |   setup @{binding fold} (Attrib.thms >> fold_meth) "fold definitions" #>
 | |
| 597 |   setup @{binding atomize} (Scan.lift (Args.mode "full") >> atomize)
 | |
| 30515 | 598 | "present local premises as object-level statements" #> | 
| 56204 | 599 |   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 | 600 | "apply some intro/elim rule" #> | 
| 56204 | 601 |   setup @{binding erule} (xrule_meth erule) "apply rule in elimination manner (improper)" #>
 | 
| 602 |   setup @{binding drule} (xrule_meth drule) "apply rule in destruct manner (improper)" #>
 | |
| 603 |   setup @{binding frule} (xrule_meth frule) "apply rule in forward manner (improper)" #>
 | |
| 604 |   setup @{binding this} (Scan.succeed (K this)) "apply current facts as rules" #>
 | |
| 605 |   setup @{binding fact} (Attrib.thms >> fact) "composition by facts from context" #>
 | |
| 606 |   setup @{binding assumption} (Scan.succeed assumption)
 | |
| 30515 | 607 | "proof by assumption, preferring facts" #> | 
| 56204 | 608 |   setup @{binding rename_tac} (Args.goal_spec -- Scan.lift (Scan.repeat1 Args.name) >>
 | 
| 52732 | 609 | (fn (quant, xs) => K (SIMPLE_METHOD'' quant (rename_tac xs)))) | 
| 30515 | 610 | "rename parameters of goal" #> | 
| 56204 | 611 |   setup @{binding rotate_tac} (Args.goal_spec -- Scan.lift (Scan.optional Parse.int 1) >>
 | 
| 52732 | 612 | (fn (quant, i) => K (SIMPLE_METHOD'' quant (rotate_tac i)))) | 
| 30515 | 613 | "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 | 614 |   setup @{binding tactic} (parse_tactic >> (K o METHOD))
 | 
| 30515 | 615 | "ML tactic as proof method" #> | 
| 58018 
beb4b7c0bb30
proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
 wenzelm parents: 
58016diff
changeset | 616 |   setup @{binding raw_tactic} (parse_tactic >> (fn tac => fn _ => NO_CASES o tac))
 | 
| 53171 | 617 | "ML tactic as raw proof method"); | 
| 5824 | 618 | |
| 619 | ||
| 16145 | 620 | (*final declarations of this structure!*) | 
| 621 | val unfold = unfold_meth; | |
| 622 | val fold = fold_meth; | |
| 623 | ||
| 5824 | 624 | end; | 
| 625 | ||
| 30508 
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
 wenzelm parents: 
30466diff
changeset | 626 | val METHOD_CASES = Method.METHOD_CASES; | 
| 
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
 wenzelm parents: 
30466diff
changeset | 627 | val METHOD = Method.METHOD; | 
| 
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
 wenzelm parents: 
30466diff
changeset | 628 | val SIMPLE_METHOD = Method.SIMPLE_METHOD; | 
| 
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
 wenzelm parents: 
30466diff
changeset | 629 | val SIMPLE_METHOD' = Method.SIMPLE_METHOD'; | 
| 
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
 wenzelm parents: 
30466diff
changeset | 630 | val SIMPLE_METHOD'' = Method.SIMPLE_METHOD''; | 
| 
958cc116d03b
tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
 wenzelm parents: 
30466diff
changeset | 631 |