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