| author | wenzelm |
| Fri, 12 Oct 2001 12:09:38 +0200 | |
| changeset 11732 | 139aaced13f4 |
| parent 11731 | 1a0c1ef86518 |
| child 11765 | 4c45eb23ef68 |
| permissions | -rw-r--r-- |
| 5824 | 1 |
(* Title: Pure/Isar/method.ML |
2 |
ID: $Id$ |
|
3 |
Author: Markus Wenzel, TU Muenchen |
|
| 8807 | 4 |
License: GPL (GNU GENERAL PUBLIC LICENSE) |
| 5824 | 5 |
|
6 |
Proof methods. |
|
7 |
*) |
|
8 |
||
9 |
signature BASIC_METHOD = |
|
10 |
sig |
|
| 11731 | 11 |
val trace_rules: bool ref |
| 5824 | 12 |
val print_methods: theory -> unit |
13 |
val Method: bstring -> (Args.src -> Proof.context -> Proof.method) -> string -> unit |
|
14 |
end; |
|
15 |
||
16 |
signature METHOD = |
|
17 |
sig |
|
18 |
include BASIC_METHOD |
|
| 11731 | 19 |
val trace: thm list -> unit |
|
8153
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
20 |
val print_global_rules: theory -> unit |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
21 |
val print_local_rules: Proof.context -> unit |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
22 |
val dest_global: theory attribute |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
23 |
val elim_global: theory attribute |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
24 |
val intro_global: theory attribute |
| 9938 | 25 |
val rule_del_global: theory attribute |
|
8153
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
26 |
val dest_local: Proof.context attribute |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
27 |
val elim_local: Proof.context attribute |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
28 |
val intro_local: Proof.context attribute |
| 9938 | 29 |
val rule_del_local: Proof.context attribute |
| 6091 | 30 |
val METHOD: (thm list -> tactic) -> Proof.method |
| 8372 | 31 |
val METHOD_CASES: |
32 |
(thm list -> thm -> (thm * (string * RuleCases.T) list) Seq.seq) -> Proof.method |
|
| 9706 | 33 |
val SIMPLE_METHOD: tactic -> Proof.method |
34 |
val SIMPLE_METHOD': ((int -> tactic) -> tactic) -> (int -> tactic) -> Proof.method |
|
| 5824 | 35 |
val fail: Proof.method |
36 |
val succeed: Proof.method |
|
| 8167 | 37 |
val defer: int option -> Proof.method |
38 |
val prefer: int -> Proof.method |
|
| 7419 | 39 |
val insert_tac: thm list -> int -> tactic |
| 7574 | 40 |
val insert: thm list -> Proof.method |
| 7555 | 41 |
val insert_facts: Proof.method |
| 7601 | 42 |
val unfold: thm list -> Proof.method |
| 7419 | 43 |
val fold: thm list -> Proof.method |
| 9484 | 44 |
val atomize_tac: thm list -> int -> tactic |
| 9485 | 45 |
val atomize_goal: thm list -> int -> thm -> thm |
| 10907 | 46 |
val atomize_strip_tac: thm list * thm list -> int -> tactic |
| 7419 | 47 |
val multi_resolve: thm list -> thm -> thm Seq.seq |
48 |
val multi_resolves: thm list -> thm list -> thm Seq.seq |
|
| 6091 | 49 |
val rule_tac: thm list -> thm list -> int -> tactic |
| 10309 | 50 |
val some_rule_tac: thm list -> Proof.context -> thm list -> int -> tactic |
| 6091 | 51 |
val rule: thm list -> Proof.method |
| 10744 | 52 |
val erule: int -> thm list -> Proof.method |
53 |
val drule: int -> thm list -> Proof.method |
|
54 |
val frule: int -> thm list -> Proof.method |
|
| 8195 | 55 |
val this: Proof.method |
| 7555 | 56 |
val assumption: Proof.context -> Proof.method |
| 8351 | 57 |
val set_tactic: (Proof.context -> thm list -> tactic) -> unit |
58 |
val tactic: string -> Proof.context -> Proof.method |
|
| 5916 | 59 |
exception METHOD_FAIL of (string * Position.T) * exn |
| 5824 | 60 |
val method: theory -> Args.src -> Proof.context -> Proof.method |
| 9539 | 61 |
val add_method: bstring * (Args.src -> Proof.context -> Proof.method) * string |
62 |
-> theory -> theory |
|
| 5824 | 63 |
val add_methods: (bstring * (Args.src -> Proof.context -> Proof.method) * string) list |
64 |
-> theory -> theory |
|
| 5884 | 65 |
val syntax: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) -> |
| 8282 | 66 |
Args.src -> Proof.context -> Proof.context * 'a |
| 8351 | 67 |
val simple_args: (Args.T list -> 'a * Args.T list) |
68 |
-> ('a -> Proof.context -> Proof.method) -> Args.src -> Proof.context -> Proof.method
|
|
| 7555 | 69 |
val ctxt_args: (Proof.context -> Proof.method) -> Args.src -> Proof.context -> Proof.method |
| 5884 | 70 |
val no_args: Proof.method -> Args.src -> Proof.context -> Proof.method |
| 7268 | 71 |
type modifier |
| 7601 | 72 |
val sectioned_args: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) -> |
| 7268 | 73 |
(Args.T list -> modifier * Args.T list) list -> |
| 9864 | 74 |
('a -> Proof.context -> 'b) -> Args.src -> Proof.context -> 'b
|
| 7601 | 75 |
val bang_sectioned_args: |
76 |
(Args.T list -> modifier * Args.T list) list -> |
|
| 9864 | 77 |
(thm list -> Proof.context -> 'a) -> Args.src -> Proof.context -> 'a |
| 9777 | 78 |
val bang_sectioned_args': |
79 |
(Args.T list -> modifier * Args.T list) list -> |
|
80 |
(Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) -> |
|
| 9864 | 81 |
('a -> thm list -> Proof.context -> 'b) -> Args.src -> Proof.context -> 'b
|
| 7601 | 82 |
val only_sectioned_args: |
83 |
(Args.T list -> modifier * Args.T list) list -> |
|
| 9864 | 84 |
(Proof.context -> 'a) -> Args.src -> Proof.context -> 'a |
85 |
val thms_ctxt_args: (thm list -> Proof.context -> 'a) -> Args.src -> Proof.context -> 'a |
|
86 |
val thms_args: (thm list -> 'a) -> Args.src -> Proof.context -> 'a |
|
87 |
val thm_args: (thm -> 'a) -> Args.src -> Proof.context -> 'a |
|
| 5824 | 88 |
datatype text = |
89 |
Basic of (Proof.context -> Proof.method) | |
|
90 |
Source of Args.src | |
|
91 |
Then of text list | |
|
92 |
Orelse of text list | |
|
93 |
Try of text | |
|
94 |
Repeat1 of text |
|
95 |
val refine: text -> Proof.state -> Proof.state Seq.seq |
|
| 8238 | 96 |
val refine_end: text -> Proof.state -> Proof.state Seq.seq |
| 5824 | 97 |
val proof: text option -> Proof.state -> Proof.state Seq.seq |
| 8966 | 98 |
val local_qed: bool -> text option |
| 6981 | 99 |
-> ({kind: string, name: string, thm: thm} -> unit) * (thm -> unit)
|
| 6736 | 100 |
-> Proof.state -> Proof.state Seq.seq |
| 6981 | 101 |
val local_terminal_proof: text * text option |
102 |
-> ({kind: string, name: string, thm: thm} -> unit) * (thm -> unit)
|
|
| 6736 | 103 |
-> Proof.state -> Proof.state Seq.seq |
| 6981 | 104 |
val local_default_proof: ({kind: string, name: string, thm: thm} -> unit) * (thm -> unit)
|
| 6736 | 105 |
-> Proof.state -> Proof.state Seq.seq |
| 8966 | 106 |
val local_immediate_proof: ({kind: string, name: string, thm: thm} -> unit) * (thm -> unit)
|
107 |
-> Proof.state -> Proof.state Seq.seq |
|
108 |
val local_done_proof: ({kind: string, name: string, thm: thm} -> unit) * (thm -> unit)
|
|
109 |
-> Proof.state -> Proof.state Seq.seq |
|
110 |
val global_qed: bool -> text option |
|
111 |
-> Proof.state -> theory * {kind: string, name: string, thm: thm}
|
|
| 6934 | 112 |
val global_terminal_proof: text * text option |
113 |
-> Proof.state -> theory * {kind: string, name: string, thm: thm}
|
|
| 8966 | 114 |
val global_default_proof: Proof.state -> theory * {kind: string, name: string, thm: thm}
|
| 6532 | 115 |
val global_immediate_proof: Proof.state -> theory * {kind: string, name: string, thm: thm}
|
| 8966 | 116 |
val global_done_proof: Proof.state -> theory * {kind: string, name: string, thm: thm}
|
| 9539 | 117 |
val goal_args: (Args.T list -> 'a * Args.T list) -> ('a -> int -> tactic)
|
118 |
-> Args.src -> Proof.context -> Proof.method |
|
119 |
val goal_args': (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) |
|
120 |
-> ('a -> int -> tactic) -> Args.src -> Proof.context -> Proof.method
|
|
| 5824 | 121 |
val setup: (theory -> theory) list |
122 |
end; |
|
123 |
||
124 |
structure Method: METHOD = |
|
125 |
struct |
|
126 |
||
127 |
||
| 11731 | 128 |
(** tracing *) |
129 |
||
130 |
val trace_rules = ref false; |
|
131 |
||
132 |
fun trace rules = |
|
133 |
if not (! trace_rules) then () |
|
134 |
else Pretty.writeln (Pretty.big_list "rules:" (map Display.pretty_thm rules)); |
|
135 |
||
136 |
||
137 |
||
|
8153
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
138 |
(** global and local rule data **) |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
139 |
|
| 10008 | 140 |
local |
141 |
fun prt_rules kind sg ths = |
|
142 |
Pretty.writeln (Pretty.big_list ("standard " ^ kind ^ " rules:")
|
|
143 |
(map (Display.pretty_thm_sg sg) ths)); |
|
144 |
in |
|
145 |
fun print_rules sg (intro, elim) = |
|
146 |
(prt_rules "introduction" sg intro; prt_rules "elimination" sg elim); |
|
147 |
end; |
|
|
8153
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
148 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
149 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
150 |
(* theory data kind 'Isar/rules' *) |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
151 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
152 |
structure GlobalRulesArgs = |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
153 |
struct |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
154 |
val name = "Isar/rules"; |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
155 |
type T = thm list * thm list; |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
156 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
157 |
val empty = ([], []); |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
158 |
val copy = I; |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
159 |
val prep_ext = I; |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
160 |
fun merge ((intro1, elim1), (intro2, elim2)) = |
| 9418 | 161 |
(Drule.merge_rules (intro1, intro2), Drule.merge_rules (elim1, elim2)); |
| 10008 | 162 |
val print = print_rules; |
|
8153
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
163 |
end; |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
164 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
165 |
structure GlobalRules = TheoryDataFun(GlobalRulesArgs); |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
166 |
val print_global_rules = GlobalRules.print; |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
167 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
168 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
169 |
(* proof data kind 'Isar/rules' *) |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
170 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
171 |
structure LocalRulesArgs = |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
172 |
struct |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
173 |
val name = "Isar/rules"; |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
174 |
type T = thm list * thm list; |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
175 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
176 |
val init = GlobalRules.get; |
| 10008 | 177 |
val print = print_rules o ProofContext.sign_of; |
|
8153
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
178 |
end; |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
179 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
180 |
structure LocalRules = ProofDataFun(LocalRulesArgs); |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
181 |
val print_local_rules = LocalRules.print; |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
182 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
183 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
184 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
185 |
(** attributes **) |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
186 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
187 |
(* add rules *) |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
188 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
189 |
local |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
190 |
|
| 9938 | 191 |
fun add_dest th (intro, elim) = (intro, Drule.add_rules [Tactic.make_elim th] elim); |
192 |
fun add_elim th (intro, elim) = (intro, Drule.add_rules [th] elim); |
|
193 |
fun add_intro th (intro, elim) = (Drule.add_rules [th] intro, elim); |
|
|
8153
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
194 |
|
| 9938 | 195 |
fun del_rule th (intro, elim) = |
196 |
let |
|
197 |
val th' = Tactic.make_elim th; |
|
198 |
val del = Drule.del_rules [th'] o Drule.del_rules [th]; |
|
199 |
in (del intro, del elim) end; |
|
|
8153
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
200 |
|
| 9938 | 201 |
fun mk_att f g (x, th) = (f (g th) x, th); |
|
8153
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
202 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
203 |
in |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
204 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
205 |
val dest_global = mk_att GlobalRules.map add_dest; |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
206 |
val elim_global = mk_att GlobalRules.map add_elim; |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
207 |
val intro_global = mk_att GlobalRules.map add_intro; |
| 9938 | 208 |
val rule_del_global = mk_att GlobalRules.map del_rule; |
|
8153
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
209 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
210 |
val dest_local = mk_att LocalRules.map add_dest; |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
211 |
val elim_local = mk_att LocalRules.map add_elim; |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
212 |
val intro_local = mk_att LocalRules.map add_intro; |
| 9938 | 213 |
val rule_del_local = mk_att LocalRules.map del_rule; |
214 |
||
| 10034 | 215 |
fun del_args att = Attrib.syntax (Scan.lift Args.del >> K att); |
|
8153
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
216 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
217 |
end; |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
218 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
219 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
220 |
(* concrete syntax *) |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
221 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
222 |
val rule_atts = |
| 9900 | 223 |
[("dest", (Attrib.no_args dest_global, Attrib.no_args dest_local),
|
224 |
"declaration of destruction rule"), |
|
225 |
("elim", (Attrib.no_args elim_global, Attrib.no_args elim_local),
|
|
226 |
"declaration of elimination rule"), |
|
227 |
("intro", (Attrib.no_args intro_global, Attrib.no_args intro_local),
|
|
228 |
"declaration of introduction rule"), |
|
| 9938 | 229 |
("rule", (del_args rule_del_global, del_args rule_del_local),
|
| 9900 | 230 |
"remove declaration of elim/intro rule")]; |
|
8153
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
231 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
232 |
|
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
233 |
|
| 5824 | 234 |
(** proof methods **) |
235 |
||
| 8372 | 236 |
(* make methods *) |
| 5824 | 237 |
|
| 6849 | 238 |
val METHOD = Proof.method; |
| 8372 | 239 |
val METHOD_CASES = Proof.method_cases; |
240 |
||
| 5824 | 241 |
|
242 |
(* primitive *) |
|
243 |
||
244 |
val fail = METHOD (K no_tac); |
|
245 |
val succeed = METHOD (K all_tac); |
|
246 |
||
247 |
||
| 8167 | 248 |
(* shuffle *) |
249 |
||
| 8240 | 250 |
fun prefer i = METHOD (K (Tactic.defer_tac i THEN PRIMITIVE (Thm.permute_prems 0 ~1))); |
| 8167 | 251 |
fun defer opt_i = METHOD (K (Tactic.defer_tac (if_none opt_i 1))); |
252 |
||
253 |
||
| 7419 | 254 |
(* insert *) |
255 |
||
256 |
local |
|
| 5824 | 257 |
|
| 6981 | 258 |
fun cut_rule_tac raw_rule = |
259 |
let |
|
260 |
val rule = Drule.forall_intr_vars raw_rule; |
|
261 |
val revcut_rl = Drule.incr_indexes_wrt [] [] [] [rule] Drule.revcut_rl; |
|
| 7555 | 262 |
in Tactic.rtac (rule COMP revcut_rl) end; |
| 6981 | 263 |
|
| 7419 | 264 |
in |
| 5824 | 265 |
|
| 7419 | 266 |
fun insert_tac [] i = all_tac |
267 |
| insert_tac facts i = EVERY (map (fn th => cut_rule_tac th i) facts); |
|
| 6981 | 268 |
|
| 7555 | 269 |
val insert_facts = METHOD (ALLGOALS o insert_tac); |
| 7664 | 270 |
fun insert thms = METHOD (fn _ => ALLGOALS (insert_tac thms)); |
| 7419 | 271 |
|
272 |
end; |
|
| 5824 | 273 |
|
274 |
||
| 9706 | 275 |
(* simple methods *) |
276 |
||
277 |
fun SIMPLE_METHOD tac = METHOD (fn facts => ALLGOALS (insert_tac facts) THEN tac); |
|
278 |
fun SIMPLE_METHOD' quant tac = METHOD (fn facts => quant (insert_tac facts THEN' tac)); |
|
279 |
||
280 |
||
| 7601 | 281 |
(* unfold / fold definitions *) |
| 6532 | 282 |
|
| 10821 | 283 |
fun unfold thms = SIMPLE_METHOD (CHANGED_PROP (rewrite_goals_tac thms)); |
284 |
fun fold thms = SIMPLE_METHOD (CHANGED_PROP (fold_goals_tac thms)); |
|
| 9484 | 285 |
|
286 |
||
287 |
(* atomize meta-connectives *) |
|
288 |
||
| 10907 | 289 |
fun atomize_strip_tac (rews, strip) = |
| 10405 | 290 |
let val rew_tac = rewrite_goal_tac rews in |
291 |
fn i => fn thm => |
|
| 10907 | 292 |
if Logic.has_meta_prems (#prop (Thm.rep_thm thm)) i then |
293 |
(rew_tac i THEN REPEAT (match_tac strip i)) thm |
|
| 10405 | 294 |
else all_tac thm |
295 |
end; |
|
296 |
||
| 10907 | 297 |
fun atomize_tac rews = atomize_strip_tac (rews, []); |
298 |
||
| 10405 | 299 |
fun atomize_goal rews = |
300 |
let val tac = atomize_tac rews |
|
301 |
in fn i => fn thm => (case Seq.pull (tac i thm) of None => thm | Some (thm', _) => thm') end; |
|
| 6532 | 302 |
|
303 |
||
| 7419 | 304 |
(* multi_resolve *) |
305 |
||
306 |
local |
|
307 |
||
308 |
fun res th i rule = |
|
309 |
Thm.biresolution false [(false, th)] i rule handle THM _ => Seq.empty; |
|
310 |
||
311 |
fun multi_res _ [] rule = Seq.single rule |
|
312 |
| multi_res i (th :: ths) rule = Seq.flat (Seq.map (res th i) (multi_res (i + 1) ths rule)); |
|
313 |
||
314 |
in |
|
315 |
||
316 |
val multi_resolve = multi_res 1; |
|
| 8372 | 317 |
fun multi_resolves facts rules = Seq.flat (Seq.map (multi_resolve facts) (Seq.of_list rules)); |
| 7419 | 318 |
|
319 |
end; |
|
320 |
||
321 |
||
| 10744 | 322 |
(* basic rules *) |
| 8372 | 323 |
|
| 7419 | 324 |
local |
| 5824 | 325 |
|
|
10541
fdec07d4f047
resolveq(_cases)_tac moved to HOL/Tools/induct_method.ML;
wenzelm
parents:
10529
diff
changeset
|
326 |
fun gen_rule_tac tac rules [] i st = tac rules i st |
|
fdec07d4f047
resolveq(_cases)_tac moved to HOL/Tools/induct_method.ML;
wenzelm
parents:
10529
diff
changeset
|
327 |
| gen_rule_tac tac erules facts i st = |
|
fdec07d4f047
resolveq(_cases)_tac moved to HOL/Tools/induct_method.ML;
wenzelm
parents:
10529
diff
changeset
|
328 |
Seq.flat (Seq.map (fn rule => (tac o single) rule i st) (multi_resolves facts erules)); |
| 7130 | 329 |
|
| 10744 | 330 |
fun gen_arule_tac tac j rules facts = |
331 |
EVERY' (gen_rule_tac tac rules facts :: replicate j Tactic.assume_tac); |
|
332 |
||
| 10309 | 333 |
fun gen_some_rule_tac tac arg_rules ctxt facts = |
|
8153
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
334 |
let val rules = |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
335 |
if not (null arg_rules) then arg_rules |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
336 |
else if null facts then #1 (LocalRules.get ctxt) |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
337 |
else op @ (LocalRules.get ctxt); |
| 11731 | 338 |
in trace rules; tac rules facts end; |
| 10309 | 339 |
|
| 10744 | 340 |
fun meth tac x = METHOD (HEADGOAL o tac x); |
341 |
fun meth' tac x y = METHOD (HEADGOAL o tac x y); |
|
| 8220 | 342 |
|
| 7419 | 343 |
in |
344 |
||
| 10744 | 345 |
val rule_tac = gen_rule_tac Tactic.resolve_tac; |
346 |
val rule = meth rule_tac; |
|
347 |
val some_rule_tac = gen_some_rule_tac rule_tac; |
|
348 |
val some_rule = meth' some_rule_tac; |
|
349 |
||
350 |
val erule = meth' (gen_arule_tac Tactic.eresolve_tac); |
|
351 |
val drule = meth' (gen_arule_tac Tactic.dresolve_tac); |
|
352 |
val frule = meth' (gen_arule_tac Tactic.forward_tac); |
|
| 5824 | 353 |
|
| 7419 | 354 |
end; |
355 |
||
356 |
||
| 8195 | 357 |
(* this *) |
358 |
||
| 8671 | 359 |
val this = METHOD (EVERY o map (HEADGOAL o Tactic.rtac)); |
| 8195 | 360 |
|
361 |
||
362 |
(* assumption *) |
|
| 7555 | 363 |
|
|
10378
98c95ebf804f
assumption / finish: handle non-atomic assumptions from context as well;
wenzelm
parents:
10309
diff
changeset
|
364 |
fun asm_tac ths = |
|
98c95ebf804f
assumption / finish: handle non-atomic assumptions from context as well;
wenzelm
parents:
10309
diff
changeset
|
365 |
foldr (op APPEND') (map (fn th => Tactic.rtac th THEN_ALL_NEW assume_tac) ths, K no_tac); |
|
98c95ebf804f
assumption / finish: handle non-atomic assumptions from context as well;
wenzelm
parents:
10309
diff
changeset
|
366 |
|
| 10405 | 367 |
fun assm_tac ctxt = |
368 |
assume_tac APPEND' |
|
369 |
asm_tac (ProofContext.prems_of ctxt) APPEND' |
|
370 |
Tactic.rtac Drule.reflexive_thm; |
|
| 7419 | 371 |
|
| 7555 | 372 |
fun assumption_tac ctxt [] = assm_tac ctxt |
|
10378
98c95ebf804f
assumption / finish: handle non-atomic assumptions from context as well;
wenzelm
parents:
10309
diff
changeset
|
373 |
| assumption_tac _ [fact] = asm_tac [fact] |
| 7555 | 374 |
| assumption_tac _ _ = K no_tac; |
| 7419 | 375 |
|
| 8671 | 376 |
fun assumption ctxt = METHOD (HEADGOAL o assumption_tac ctxt); |
| 7419 | 377 |
|
378 |
||
| 9539 | 379 |
(* res_inst_tac etc. *) |
| 8238 | 380 |
|
| 9539 | 381 |
(*Note: insts refer to the implicit (!!) goal context; use at your own risk*) |
|
9565
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
wenzelm
parents:
9539
diff
changeset
|
382 |
fun gen_res_inst _ tac (quant, ([], thms)) = |
|
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
wenzelm
parents:
9539
diff
changeset
|
383 |
METHOD (fn facts => (quant (insert_tac facts THEN' tac thms))) |
|
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
wenzelm
parents:
9539
diff
changeset
|
384 |
| gen_res_inst tac _ (quant, (insts, [thm])) = |
|
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
wenzelm
parents:
9539
diff
changeset
|
385 |
METHOD (fn facts => (quant (insert_tac facts THEN' tac insts thm))) |
|
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
wenzelm
parents:
9539
diff
changeset
|
386 |
| gen_res_inst _ _ _ = error "Cannot have instantiations with multiple rules"; |
| 8238 | 387 |
|
|
9565
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
wenzelm
parents:
9539
diff
changeset
|
388 |
val res_inst = gen_res_inst Tactic.res_inst_tac Tactic.resolve_tac; |
|
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
wenzelm
parents:
9539
diff
changeset
|
389 |
val eres_inst = gen_res_inst Tactic.eres_inst_tac Tactic.eresolve_tac; |
|
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
wenzelm
parents:
9539
diff
changeset
|
390 |
val dres_inst = gen_res_inst Tactic.dres_inst_tac Tactic.dresolve_tac; |
|
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
wenzelm
parents:
9539
diff
changeset
|
391 |
val forw_inst = gen_res_inst Tactic.forw_inst_tac Tactic.forward_tac; |
|
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
wenzelm
parents:
9539
diff
changeset
|
392 |
val cut_inst = gen_res_inst Tactic.cut_inst_tac Tactic.cut_facts_tac; |
| 8238 | 393 |
|
394 |
||
| 8329 | 395 |
(* simple Prolog interpreter *) |
396 |
||
397 |
fun prolog_tac rules facts = |
|
398 |
DEPTH_SOLVE_1 (HEADGOAL (Tactic.assume_tac APPEND' Tactic.resolve_tac (facts @ rules))); |
|
399 |
||
400 |
val prolog = METHOD o prolog_tac; |
|
401 |
||
402 |
||
| 8351 | 403 |
(* ML tactics *) |
404 |
||
405 |
val tactic_ref = ref ((fn _ => raise Match): Proof.context -> thm list -> tactic); |
|
406 |
fun set_tactic f = tactic_ref := f; |
|
407 |
||
408 |
fun tactic txt ctxt = METHOD (fn facts => |
|
| 9631 | 409 |
(Context.use_mltext |
410 |
("let fun tactic (ctxt: PureIsar.Proof.context) (facts: thm list) : tactic = \
|
|
411 |
\let val thm = PureIsar.ProofContext.get_thm_closure ctxt\n\ |
|
412 |
\ and thms = PureIsar.ProofContext.get_thms_closure ctxt in\n" |
|
413 |
^ txt ^ |
|
414 |
"\nend in PureIsar.Method.set_tactic tactic end") |
|
415 |
false None; |
|
416 |
Context.setmp (Some (ProofContext.theory_of ctxt)) (! tactic_ref ctxt) facts)); |
|
| 8351 | 417 |
|
418 |
||
| 5824 | 419 |
|
420 |
(** methods theory data **) |
|
421 |
||
422 |
(* data kind 'Isar/methods' *) |
|
423 |
||
424 |
structure MethodsDataArgs = |
|
425 |
struct |
|
426 |
val name = "Isar/methods"; |
|
427 |
type T = |
|
428 |
{space: NameSpace.T,
|
|
429 |
meths: (((Args.src -> Proof.context -> Proof.method) * string) * stamp) Symtab.table}; |
|
430 |
||
431 |
val empty = {space = NameSpace.empty, meths = Symtab.empty};
|
|
| 6546 | 432 |
val copy = I; |
| 5824 | 433 |
val prep_ext = I; |
434 |
fun merge ({space = space1, meths = meths1}, {space = space2, meths = meths2}) =
|
|
435 |
{space = NameSpace.merge (space1, space2),
|
|
436 |
meths = Symtab.merge eq_snd (meths1, meths2) handle Symtab.DUPS dups => |
|
437 |
error ("Attempt to merge different versions of methods " ^ commas_quote dups)};
|
|
438 |
||
| 9222 | 439 |
fun print _ {space, meths} =
|
| 5824 | 440 |
let |
441 |
fun prt_meth (name, ((_, comment), _)) = Pretty.block |
|
| 6849 | 442 |
[Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment]; |
| 5824 | 443 |
in |
| 8720 | 444 |
[Pretty.big_list "methods:" (map prt_meth (NameSpace.cond_extern_table space meths))] |
| 9222 | 445 |
|> Pretty.chunks |> Pretty.writeln |
| 5824 | 446 |
end; |
447 |
end; |
|
448 |
||
449 |
structure MethodsData = TheoryDataFun(MethodsDataArgs); |
|
450 |
val print_methods = MethodsData.print; |
|
| 7611 | 451 |
|
| 5824 | 452 |
|
453 |
(* get methods *) |
|
454 |
||
| 5916 | 455 |
exception METHOD_FAIL of (string * Position.T) * exn; |
456 |
||
| 5824 | 457 |
fun method thy = |
458 |
let |
|
459 |
val {space, meths} = MethodsData.get thy;
|
|
460 |
||
| 5884 | 461 |
fun meth src = |
462 |
let |
|
463 |
val ((raw_name, _), pos) = Args.dest_src src; |
|
464 |
val name = NameSpace.intern space raw_name; |
|
465 |
in |
|
| 5824 | 466 |
(case Symtab.lookup (meths, name) of |
467 |
None => error ("Unknown proof method: " ^ quote name ^ Position.str_of pos)
|
|
| 5916 | 468 |
| Some ((mth, _), _) => transform_failure (curry METHOD_FAIL (name, pos)) (mth src)) |
| 5824 | 469 |
end; |
470 |
in meth end; |
|
471 |
||
472 |
||
| 9194 | 473 |
(* add_method(s) *) |
| 5824 | 474 |
|
475 |
fun add_methods raw_meths thy = |
|
476 |
let |
|
477 |
val full = Sign.full_name (Theory.sign_of thy); |
|
478 |
val new_meths = |
|
479 |
map (fn (name, f, comment) => (full name, ((f, comment), stamp ()))) raw_meths; |
|
480 |
||
481 |
val {space, meths} = MethodsData.get thy;
|
|
482 |
val space' = NameSpace.extend (space, map fst new_meths); |
|
483 |
val meths' = Symtab.extend (meths, new_meths) handle Symtab.DUPS dups => |
|
484 |
error ("Duplicate declaration of method(s) " ^ commas_quote dups);
|
|
485 |
in |
|
486 |
thy |> MethodsData.put {space = space', meths = meths'}
|
|
487 |
end; |
|
488 |
||
| 9194 | 489 |
val add_method = add_methods o Library.single; |
490 |
||
| 5824 | 491 |
(*implicit version*) |
492 |
fun Method name meth cmt = Context.>> (add_methods [(name, meth, cmt)]); |
|
493 |
||
494 |
||
| 5884 | 495 |
|
496 |
(** method syntax **) |
|
| 5824 | 497 |
|
| 5884 | 498 |
(* basic *) |
499 |
||
500 |
fun syntax (scan: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list))) = |
|
501 |
Args.syntax "method" scan; |
|
| 5824 | 502 |
|
| 8351 | 503 |
fun simple_args scan f src ctxt : Proof.method = |
504 |
#2 (syntax (Scan.lift (scan >> (fn x => f x ctxt))) src ctxt); |
|
505 |
||
| 7555 | 506 |
fun ctxt_args (f: Proof.context -> Proof.method) src ctxt = |
| 8282 | 507 |
#2 (syntax (Scan.succeed (f ctxt)) src ctxt); |
| 7555 | 508 |
|
509 |
fun no_args m = ctxt_args (K m); |
|
| 5884 | 510 |
|
511 |
||
| 8351 | 512 |
|
| 5884 | 513 |
(* sections *) |
| 5824 | 514 |
|
| 7268 | 515 |
type modifier = (Proof.context -> Proof.context) * Proof.context attribute; |
516 |
||
517 |
local |
|
518 |
||
| 8381 | 519 |
fun sect ss = Scan.first (map Scan.lift ss); |
| 5884 | 520 |
fun thms ss = Scan.unless (sect ss) Attrib.local_thms; |
521 |
fun thmss ss = Scan.repeat (thms ss) >> flat; |
|
522 |
||
| 7268 | 523 |
fun apply (f, att) (ctxt, ths) = Thm.applys_attributes ((f ctxt, ths), [att]); |
| 5824 | 524 |
|
| 7268 | 525 |
fun section ss = (sect ss -- thmss ss) :-- (fn (m, ths) => Scan.depend (fn ctxt => |
526 |
Scan.succeed (apply m (ctxt, ths)))) >> #2; |
|
| 5884 | 527 |
|
| 7601 | 528 |
fun sectioned args ss = args -- Scan.repeat (section ss); |
| 5884 | 529 |
|
| 7268 | 530 |
in |
| 5824 | 531 |
|
| 5884 | 532 |
fun sectioned_args args ss f src ctxt = |
| 8282 | 533 |
let val (ctxt', (x, _)) = syntax (sectioned args ss) src ctxt |
| 5921 | 534 |
in f x ctxt' end; |
| 5884 | 535 |
|
| 7601 | 536 |
fun bang_sectioned_args ss f = sectioned_args Args.bang_facts ss f; |
| 9777 | 537 |
fun bang_sectioned_args' ss scan f = |
538 |
sectioned_args (Args.bang_facts -- scan >> swap) ss (uncurry f); |
|
| 7601 | 539 |
fun only_sectioned_args ss f = sectioned_args (Scan.succeed ()) ss (fn () => f); |
| 7268 | 540 |
|
| 8093 | 541 |
fun thms_ctxt_args f = sectioned_args (thmss []) [] f; |
542 |
fun thms_args f = thms_ctxt_args (K o f); |
|
| 9706 | 543 |
fun thm_args f = thms_args (fn [thm] => f thm | _ => error "Single theorem expected"); |
| 5824 | 544 |
|
| 7268 | 545 |
end; |
546 |
||
| 5824 | 547 |
|
| 9539 | 548 |
(* tactic syntax *) |
| 8238 | 549 |
|
| 10744 | 550 |
fun nat_thms_args f = uncurry f oo |
551 |
(#2 oo syntax (Scan.lift (Scan.optional (Args.parens Args.nat) 0) -- Attrib.local_thmss)); |
|
552 |
||
| 8238 | 553 |
val insts = |
| 9539 | 554 |
Scan.optional |
|
9565
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
wenzelm
parents:
9539
diff
changeset
|
555 |
(Args.enum1 "and" (Scan.lift (Args.name -- (Args.$$$ "=" |-- Args.!!! Args.name))) --| |
|
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
wenzelm
parents:
9539
diff
changeset
|
556 |
Scan.lift (Args.$$$ "in")) [] -- Attrib.local_thmss; |
| 8238 | 557 |
|
| 8537 | 558 |
fun inst_args f = f oo (#2 oo syntax (Args.goal_spec HEADGOAL -- insts)); |
559 |
||
560 |
||
| 9539 | 561 |
fun goal_args' args tac = #2 oo syntax (Args.goal_spec HEADGOAL -- args >> |
| 9706 | 562 |
(fn (quant, s) => SIMPLE_METHOD' quant (tac s))); |
| 8537 | 563 |
|
| 9539 | 564 |
fun goal_args args tac = goal_args' (Scan.lift args) tac; |
| 8238 | 565 |
|
566 |
||
| 5824 | 567 |
|
568 |
(** method text **) |
|
569 |
||
570 |
(* datatype text *) |
|
571 |
||
572 |
datatype text = |
|
573 |
Basic of (Proof.context -> Proof.method) | |
|
574 |
Source of Args.src | |
|
575 |
Then of text list | |
|
576 |
Orelse of text list | |
|
577 |
Try of text | |
|
578 |
Repeat1 of text; |
|
579 |
||
580 |
||
581 |
(* refine *) |
|
582 |
||
| 8238 | 583 |
fun gen_refine f text state = |
| 5824 | 584 |
let |
585 |
val thy = Proof.theory_of state; |
|
586 |
||
| 8238 | 587 |
fun eval (Basic mth) = f mth |
588 |
| eval (Source src) = f (method thy src) |
|
| 5824 | 589 |
| eval (Then txts) = Seq.EVERY (map eval txts) |
590 |
| eval (Orelse txts) = Seq.FIRST (map eval txts) |
|
591 |
| eval (Try txt) = Seq.TRY (eval txt) |
|
592 |
| eval (Repeat1 txt) = Seq.REPEAT1 (eval txt); |
|
593 |
in eval text state end; |
|
594 |
||
| 8238 | 595 |
val refine = gen_refine Proof.refine; |
596 |
val refine_end = gen_refine Proof.refine_end; |
|
| 6404 | 597 |
|
| 5824 | 598 |
|
| 6404 | 599 |
(* structured proof steps *) |
| 5824 | 600 |
|
| 7506 | 601 |
val default_text = Source (Args.src (("default", []), Position.none));
|
| 8195 | 602 |
val this_text = Basic (K this); |
| 9706 | 603 |
val done_text = Basic (K (SIMPLE_METHOD all_tac)); |
| 7555 | 604 |
|
| 8966 | 605 |
fun close_text asm = Basic (fn ctxt => METHOD (K |
606 |
(FILTER Thm.no_prems ((if asm then ALLGOALS (assm_tac ctxt) else all_tac) THEN flexflex_tac)))); |
|
607 |
||
608 |
fun finish_text asm None = close_text asm |
|
609 |
| finish_text asm (Some txt) = Then [txt, close_text asm]; |
|
| 6872 | 610 |
|
| 5824 | 611 |
fun proof opt_text state = |
612 |
state |
|
613 |
|> Proof.assert_backward |
|
| 6404 | 614 |
|> refine (if_none opt_text default_text) |
| 8242 | 615 |
|> Seq.map (Proof.goal_facts (K [])) |
| 5824 | 616 |
|> Seq.map Proof.enter_forward; |
617 |
||
| 8966 | 618 |
fun local_qed asm opt_text = Proof.local_qed (refine (finish_text asm opt_text)); |
619 |
fun local_terminal_proof (text, opt_text) pr = |
|
620 |
Seq.THEN (proof (Some text), local_qed true opt_text pr); |
|
621 |
val local_default_proof = local_terminal_proof (default_text, None); |
|
| 8195 | 622 |
val local_immediate_proof = local_terminal_proof (this_text, None); |
| 8966 | 623 |
fun local_done_proof pr = Seq.THEN (proof (Some done_text), local_qed false None pr); |
| 5824 | 624 |
|
| 6872 | 625 |
|
| 8966 | 626 |
fun global_qeds asm opt_text = Proof.global_qed (refine (finish_text asm opt_text)); |
| 5824 | 627 |
|
| 8966 | 628 |
fun global_qed asm opt_text state = |
| 6872 | 629 |
state |
| 8966 | 630 |
|> global_qeds asm opt_text |
| 6872 | 631 |
|> Proof.check_result "Failed to finish proof" state |
632 |
|> Seq.hd; |
|
633 |
||
| 8966 | 634 |
fun global_term_proof asm (text, opt_text) state = |
| 6872 | 635 |
state |
636 |
|> proof (Some text) |
|
637 |
|> Proof.check_result "Terminal proof method failed" state |
|
| 8966 | 638 |
|> (Seq.flat o Seq.map (global_qeds asm opt_text)) |
| 6872 | 639 |
|> Proof.check_result "Failed to finish proof (after successful terminal method)" state |
640 |
|> Seq.hd; |
|
641 |
||
| 8966 | 642 |
val global_terminal_proof = global_term_proof true; |
| 6934 | 643 |
val global_default_proof = global_terminal_proof (default_text, None); |
| 8966 | 644 |
val global_immediate_proof = global_terminal_proof (this_text, None); |
645 |
val global_done_proof = global_term_proof false (done_text, None); |
|
| 5824 | 646 |
|
647 |
||
648 |
(** theory setup **) |
|
649 |
||
| 9539 | 650 |
(* misc tactic emulations *) |
651 |
||
652 |
val subgoal_meth = goal_args (Scan.repeat1 Args.name) Tactic.subgoals_tac; |
|
653 |
val thin_meth = goal_args Args.name Tactic.thin_tac; |
|
654 |
val rename_meth = goal_args (Scan.repeat1 Args.name) Tactic.rename_params_tac; |
|
| 9631 | 655 |
val rotate_meth = goal_args (Scan.optional Args.int 1) Tactic.rotate_tac; |
| 9539 | 656 |
|
657 |
||
| 5824 | 658 |
(* pure_methods *) |
659 |
||
660 |
val pure_methods = |
|
661 |
[("fail", no_args fail, "force failure"),
|
|
662 |
("succeed", no_args succeed, "succeed"),
|
|
| 9587 | 663 |
("-", no_args insert_facts, "do nothing (insert current facts only)"),
|
| 9539 | 664 |
("insert", thms_args insert, "insert theorems, ignoring facts (improper)"),
|
| 7601 | 665 |
("unfold", thms_args unfold, "unfold definitions"),
|
666 |
("fold", thms_args fold, "fold definitions"),
|
|
|
8153
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
667 |
("rule", thms_ctxt_args some_rule, "apply some rule"),
|
| 10744 | 668 |
("erule", nat_thms_args erule, "apply rule in elimination manner (improper)"),
|
669 |
("drule", nat_thms_args drule, "apply rule in destruct manner (improper)"),
|
|
670 |
("frule", nat_thms_args frule, "apply rule in forward manner (improper)"),
|
|
| 8195 | 671 |
("this", no_args this, "apply current facts as rules"),
|
| 8238 | 672 |
("assumption", ctxt_args assumption, "proof by assumption, preferring facts"),
|
| 9539 | 673 |
("rule_tac", inst_args res_inst, "apply rule (dynamic instantiation!)"),
|
674 |
("erule_tac", inst_args eres_inst, "apply rule in elimination manner (dynamic instantiation!)"),
|
|
675 |
("drule_tac", inst_args dres_inst, "apply rule in destruct manner (dynamic instantiation!)"),
|
|
676 |
("frule_tac", inst_args forw_inst, "apply rule in forward manner (dynamic instantiation!)"),
|
|
677 |
("cut_tac", inst_args cut_inst, "cut rule (dynamic instantiation!)"),
|
|
|
9565
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
wenzelm
parents:
9539
diff
changeset
|
678 |
("subgoal_tac", subgoal_meth, "insert subgoal (dynamic instantiation!)"),
|
|
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
wenzelm
parents:
9539
diff
changeset
|
679 |
("thin_tac", thin_meth, "remove premise (dynamic instantiation!)"),
|
| 9631 | 680 |
("rename_tac", rename_meth, "rename parameters of goal (dynamic instantiation!)"),
|
681 |
("rotate_tac", rotate_meth, "rotate assumptions of goal"),
|
|
| 8351 | 682 |
("prolog", thms_args prolog, "simple prolog interpreter"),
|
683 |
("tactic", simple_args Args.name tactic, "ML tactic as proof method")];
|
|
| 5824 | 684 |
|
685 |
||
686 |
(* setup *) |
|
687 |
||
|
8153
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
688 |
val setup = |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
689 |
[GlobalRules.init, LocalRules.init, Attrib.add_attributes rule_atts, |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
690 |
MethodsData.init, add_methods pure_methods]; |
| 5824 | 691 |
|
692 |
||
693 |
end; |
|
694 |
||
695 |
||
696 |
structure BasicMethod: BASIC_METHOD = Method; |
|
697 |
open BasicMethod; |