author | wenzelm |
Thu, 29 Jun 2000 22:31:29 +0200 | |
changeset 9194 | a57987e0250b |
parent 8966 | 916966f68907 |
child 9222 | 92ad2341179d |
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 |
|
11 |
val print_methods: theory -> unit |
|
12 |
val Method: bstring -> (Args.src -> Proof.context -> Proof.method) -> string -> unit |
|
13 |
end; |
|
14 |
||
15 |
signature METHOD = |
|
16 |
sig |
|
17 |
include BASIC_METHOD |
|
8153
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
18 |
val print_global_rules: theory -> unit |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
19 |
val print_local_rules: Proof.context -> unit |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
20 |
val dest_global: theory attribute |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
21 |
val elim_global: theory attribute |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
22 |
val intro_global: theory attribute |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
23 |
val delrule_global: theory attribute |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
24 |
val dest_local: Proof.context attribute |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
25 |
val elim_local: Proof.context attribute |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
26 |
val intro_local: Proof.context attribute |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
27 |
val delrule_local: Proof.context attribute |
6091 | 28 |
val METHOD: (thm list -> tactic) -> Proof.method |
8372 | 29 |
val METHOD_CASES: |
30 |
(thm list -> thm -> (thm * (string * RuleCases.T) list) Seq.seq) -> Proof.method |
|
5824 | 31 |
val METHOD0: tactic -> Proof.method |
32 |
val fail: Proof.method |
|
33 |
val succeed: Proof.method |
|
8167 | 34 |
val defer: int option -> Proof.method |
35 |
val prefer: int -> Proof.method |
|
7419 | 36 |
val insert_tac: thm list -> int -> tactic |
7574 | 37 |
val insert: thm list -> Proof.method |
7555 | 38 |
val insert_facts: Proof.method |
7601 | 39 |
val unfold: thm list -> Proof.method |
7419 | 40 |
val fold: thm list -> Proof.method |
41 |
val multi_resolve: thm list -> thm -> thm Seq.seq |
|
42 |
val multi_resolves: thm list -> thm list -> thm Seq.seq |
|
8335 | 43 |
val resolveq_tac: thm Seq.seq -> int -> tactic |
8372 | 44 |
val resolveq_cases_tac: (thm * string list) Seq.seq |
45 |
-> int -> thm -> (thm * (string * RuleCases.T) list) Seq.seq |
|
6091 | 46 |
val rule_tac: thm list -> thm list -> int -> tactic |
47 |
val rule: thm list -> Proof.method |
|
7130 | 48 |
val erule: thm list -> Proof.method |
8220 | 49 |
val drule: thm list -> Proof.method |
50 |
val frule: thm list -> Proof.method |
|
8195 | 51 |
val this: Proof.method |
7555 | 52 |
val assumption: Proof.context -> Proof.method |
8351 | 53 |
val set_tactic: (Proof.context -> thm list -> tactic) -> unit |
54 |
val tactic: string -> Proof.context -> Proof.method |
|
5916 | 55 |
exception METHOD_FAIL of (string * Position.T) * exn |
8720 | 56 |
val help_methods: theory option -> Pretty.T list |
5824 | 57 |
val method: theory -> Args.src -> Proof.context -> Proof.method |
9194 | 58 |
val add_method: bstring * (Args.src -> Proof.context -> Proof.method) * string -> theory -> theory |
5824 | 59 |
val add_methods: (bstring * (Args.src -> Proof.context -> Proof.method) * string) list |
60 |
-> theory -> theory |
|
5884 | 61 |
val syntax: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) -> |
8282 | 62 |
Args.src -> Proof.context -> Proof.context * 'a |
8351 | 63 |
val simple_args: (Args.T list -> 'a * Args.T list) |
64 |
-> ('a -> Proof.context -> Proof.method) -> Args.src -> Proof.context -> Proof.method |
|
7555 | 65 |
val ctxt_args: (Proof.context -> Proof.method) -> Args.src -> Proof.context -> Proof.method |
5884 | 66 |
val no_args: Proof.method -> Args.src -> Proof.context -> Proof.method |
7268 | 67 |
type modifier |
7601 | 68 |
val sectioned_args: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) -> |
7268 | 69 |
(Args.T list -> modifier * Args.T list) list -> |
5884 | 70 |
('a -> Proof.context -> Proof.method) -> Args.src -> Proof.context -> Proof.method |
7601 | 71 |
val bang_sectioned_args: |
72 |
(Args.T list -> modifier * Args.T list) list -> |
|
7555 | 73 |
(thm list -> Proof.context -> Proof.method) -> Args.src -> Proof.context -> Proof.method |
7601 | 74 |
val only_sectioned_args: |
75 |
(Args.T list -> modifier * Args.T list) list -> |
|
5884 | 76 |
(Proof.context -> Proof.method) -> Args.src -> Proof.context -> Proof.method |
8093 | 77 |
val thms_ctxt_args: (thm list -> Proof.context -> Proof.method) |
78 |
-> Args.src -> Proof.context -> Proof.method |
|
6091 | 79 |
val thms_args: (thm list -> Proof.method) -> Args.src -> Proof.context -> Proof.method |
5824 | 80 |
datatype text = |
81 |
Basic of (Proof.context -> Proof.method) | |
|
82 |
Source of Args.src | |
|
83 |
Then of text list | |
|
84 |
Orelse of text list | |
|
85 |
Try of text | |
|
86 |
Repeat1 of text |
|
87 |
val refine: text -> Proof.state -> Proof.state Seq.seq |
|
8238 | 88 |
val refine_end: text -> Proof.state -> Proof.state Seq.seq |
5824 | 89 |
val proof: text option -> Proof.state -> Proof.state Seq.seq |
8966 | 90 |
val local_qed: bool -> text option |
6981 | 91 |
-> ({kind: string, name: string, thm: thm} -> unit) * (thm -> unit) |
6736 | 92 |
-> Proof.state -> Proof.state Seq.seq |
6981 | 93 |
val local_terminal_proof: text * text option |
94 |
-> ({kind: string, name: string, thm: thm} -> unit) * (thm -> unit) |
|
6736 | 95 |
-> Proof.state -> Proof.state Seq.seq |
6981 | 96 |
val local_default_proof: ({kind: string, name: string, thm: thm} -> unit) * (thm -> unit) |
6736 | 97 |
-> Proof.state -> Proof.state Seq.seq |
8966 | 98 |
val local_immediate_proof: ({kind: string, name: string, thm: thm} -> unit) * (thm -> unit) |
99 |
-> Proof.state -> Proof.state Seq.seq |
|
100 |
val local_done_proof: ({kind: string, name: string, thm: thm} -> unit) * (thm -> unit) |
|
101 |
-> Proof.state -> Proof.state Seq.seq |
|
102 |
val global_qed: bool -> text option |
|
103 |
-> Proof.state -> theory * {kind: string, name: string, thm: thm} |
|
6934 | 104 |
val global_terminal_proof: text * text option |
105 |
-> Proof.state -> theory * {kind: string, name: string, thm: thm} |
|
8966 | 106 |
val global_default_proof: Proof.state -> theory * {kind: string, name: string, thm: thm} |
6532 | 107 |
val global_immediate_proof: Proof.state -> theory * {kind: string, name: string, thm: thm} |
8966 | 108 |
val global_done_proof: Proof.state -> theory * {kind: string, name: string, thm: thm} |
5824 | 109 |
val setup: (theory -> theory) list |
110 |
end; |
|
111 |
||
112 |
structure Method: METHOD = |
|
113 |
struct |
|
114 |
||
115 |
||
8153
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
116 |
(** global and local rule data **) |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
117 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
118 |
fun prt_rules kind ths = |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
119 |
Pretty.writeln (Pretty.big_list ("standard " ^ kind ^ " rules:") (map Display.pretty_thm ths)); |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
120 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
121 |
fun print_rules (intro, elim) = |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
122 |
(prt_rules "introduction" intro; prt_rules "elimination" elim); |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
123 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
124 |
fun merge_rules (ths1, ths2) = Library.generic_merge Thm.eq_thm I I ths1 ths2; |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
125 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
126 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
127 |
(* theory data kind 'Isar/rules' *) |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
128 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
129 |
structure GlobalRulesArgs = |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
130 |
struct |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
131 |
val name = "Isar/rules"; |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
132 |
type T = thm list * thm list; |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
133 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
134 |
val empty = ([], []); |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
135 |
val copy = I; |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
136 |
val prep_ext = I; |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
137 |
fun merge ((intro1, elim1), (intro2, elim2)) = |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
138 |
(merge_rules (intro1, intro2), merge_rules (elim1, elim2)); |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
139 |
fun print _ = print_rules; |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
140 |
end; |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
141 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
142 |
structure GlobalRules = TheoryDataFun(GlobalRulesArgs); |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
143 |
val print_global_rules = GlobalRules.print; |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
144 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
145 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
146 |
(* proof data kind 'Isar/rules' *) |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
147 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
148 |
structure LocalRulesArgs = |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
149 |
struct |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
150 |
val name = "Isar/rules"; |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
151 |
type T = thm list * thm list; |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
152 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
153 |
val init = GlobalRules.get; |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
154 |
fun print _ = print_rules; |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
155 |
end; |
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 |
structure LocalRules = ProofDataFun(LocalRulesArgs); |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
158 |
val print_local_rules = LocalRules.print; |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
159 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
160 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
161 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
162 |
(** attributes **) |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
163 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
164 |
(* add rules *) |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
165 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
166 |
local |
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 |
fun add_rule thm rules = Library.gen_ins Thm.eq_thm (thm, rules); |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
169 |
fun del_rule thm rules = Library.gen_rem Thm.eq_thm (rules, thm); |
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 |
fun add_dest thm (intro, elim) = (intro, add_rule (Tactic.make_elim thm) elim); |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
172 |
fun add_elim thm (intro, elim) = (intro, add_rule thm elim); |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
173 |
fun add_intro thm (intro, elim) = (add_rule thm intro, elim); |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
174 |
fun delrule thm (intro, elim) = (del_rule thm intro, del_rule thm elim); |
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 |
fun mk_att f g (x, thm) = (f (g thm) x, thm); |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
177 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
178 |
in |
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 |
val dest_global = mk_att GlobalRules.map add_dest; |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
181 |
val elim_global = mk_att GlobalRules.map add_elim; |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
182 |
val intro_global = mk_att GlobalRules.map add_intro; |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
183 |
val delrule_global = mk_att GlobalRules.map delrule; |
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 |
val dest_local = mk_att LocalRules.map add_dest; |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
186 |
val elim_local = mk_att LocalRules.map add_elim; |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
187 |
val intro_local = mk_att LocalRules.map add_intro; |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
188 |
val delrule_local = mk_att LocalRules.map delrule; |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
189 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
190 |
end; |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
191 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
192 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
193 |
(* concrete syntax *) |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
194 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
195 |
val rule_atts = |
8519 | 196 |
[("dest", (Attrib.no_args dest_global, Attrib.no_args dest_local), "declare destruction rule"), |
197 |
("elim", (Attrib.no_args elim_global, Attrib.no_args elim_local), "declare elimination rule"), |
|
198 |
("intro", (Attrib.no_args intro_global, Attrib.no_args intro_local), "declare introduction rule"), |
|
199 |
("delrule", (Attrib.no_args delrule_global, Attrib.no_args delrule_local), "undeclare rule")]; |
|
8153
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
200 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
201 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
202 |
|
5824 | 203 |
(** proof methods **) |
204 |
||
8372 | 205 |
(* make methods *) |
5824 | 206 |
|
6849 | 207 |
val METHOD = Proof.method; |
8372 | 208 |
val METHOD_CASES = Proof.method_cases; |
209 |
||
8966 | 210 |
fun METHOD0 tac = METHOD (fn [] => tac | _ => error "Cannot handle current facts"); |
5824 | 211 |
|
212 |
||
213 |
(* primitive *) |
|
214 |
||
215 |
val fail = METHOD (K no_tac); |
|
216 |
val succeed = METHOD (K all_tac); |
|
217 |
||
218 |
||
8167 | 219 |
(* shuffle *) |
220 |
||
8240 | 221 |
fun prefer i = METHOD (K (Tactic.defer_tac i THEN PRIMITIVE (Thm.permute_prems 0 ~1))); |
8167 | 222 |
fun defer opt_i = METHOD (K (Tactic.defer_tac (if_none opt_i 1))); |
223 |
||
224 |
||
7419 | 225 |
(* insert *) |
226 |
||
227 |
local |
|
5824 | 228 |
|
6981 | 229 |
fun cut_rule_tac raw_rule = |
230 |
let |
|
231 |
val rule = Drule.forall_intr_vars raw_rule; |
|
232 |
val revcut_rl = Drule.incr_indexes_wrt [] [] [] [rule] Drule.revcut_rl; |
|
7555 | 233 |
in Tactic.rtac (rule COMP revcut_rl) end; |
6981 | 234 |
|
7419 | 235 |
in |
5824 | 236 |
|
7419 | 237 |
fun insert_tac [] i = all_tac |
238 |
| insert_tac facts i = EVERY (map (fn th => cut_rule_tac th i) facts); |
|
6981 | 239 |
|
7555 | 240 |
val insert_facts = METHOD (ALLGOALS o insert_tac); |
7664 | 241 |
fun insert thms = METHOD (fn _ => ALLGOALS (insert_tac thms)); |
7419 | 242 |
|
243 |
end; |
|
5824 | 244 |
|
245 |
||
7601 | 246 |
(* unfold / fold definitions *) |
6532 | 247 |
|
7601 | 248 |
fun unfold thms = METHOD (fn facts => ALLGOALS (insert_tac facts) THEN rewrite_goals_tac thms); |
7555 | 249 |
fun fold thms = METHOD (fn facts => ALLGOALS (insert_tac facts) THEN fold_goals_tac thms); |
6532 | 250 |
|
251 |
||
7419 | 252 |
(* multi_resolve *) |
253 |
||
254 |
local |
|
255 |
||
256 |
fun res th i rule = |
|
257 |
Thm.biresolution false [(false, th)] i rule handle THM _ => Seq.empty; |
|
258 |
||
259 |
fun multi_res _ [] rule = Seq.single rule |
|
260 |
| multi_res i (th :: ths) rule = Seq.flat (Seq.map (res th i) (multi_res (i + 1) ths rule)); |
|
261 |
||
262 |
in |
|
263 |
||
264 |
val multi_resolve = multi_res 1; |
|
8372 | 265 |
fun multi_resolves facts rules = Seq.flat (Seq.map (multi_resolve facts) (Seq.of_list rules)); |
7419 | 266 |
|
267 |
end; |
|
268 |
||
269 |
||
8372 | 270 |
(* general rule *) |
5824 | 271 |
|
8335 | 272 |
fun gen_resolveq_tac tac rules i st = |
8372 | 273 |
Seq.flat (Seq.map (fn rule => tac rule i st) rules); |
274 |
||
275 |
val resolveq_tac = gen_resolveq_tac Tactic.rtac; |
|
8335 | 276 |
|
8372 | 277 |
val resolveq_cases_tac = gen_resolveq_tac (fn (rule, cases) => fn i => fn st => |
278 |
Seq.map (rpair (RuleCases.make rule cases)) (Tactic.rtac rule i st)); |
|
8335 | 279 |
|
280 |
||
8372 | 281 |
(* simple rule *) |
282 |
||
7419 | 283 |
local |
5824 | 284 |
|
7130 | 285 |
fun gen_rule_tac tac rules [] = tac rules |
8372 | 286 |
| gen_rule_tac tac erules facts = gen_resolveq_tac (tac o single) (multi_resolves facts erules); |
7130 | 287 |
|
8671 | 288 |
fun gen_rule tac rules = METHOD (HEADGOAL o tac rules); |
8153
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
289 |
|
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
290 |
fun gen_rule' tac arg_rules ctxt = METHOD (fn facts => |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
291 |
let val rules = |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
292 |
if not (null arg_rules) then arg_rules |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
293 |
else if null facts then #1 (LocalRules.get ctxt) |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
294 |
else op @ (LocalRules.get ctxt); |
8671 | 295 |
in HEADGOAL (tac rules facts) end); |
8153
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
296 |
|
8220 | 297 |
fun setup raw_tac = |
298 |
let val tac = gen_rule_tac raw_tac |
|
299 |
in (tac, gen_rule tac, gen_rule' tac) end; |
|
300 |
||
7419 | 301 |
in |
302 |
||
8220 | 303 |
val (rule_tac, rule, some_rule) = setup Tactic.resolve_tac; |
304 |
val (erule_tac, erule, some_erule) = setup Tactic.eresolve_tac; |
|
305 |
val (drule_tac, drule, some_drule) = setup Tactic.dresolve_tac; |
|
306 |
val (frule_tac, frule, some_frule) = setup Tactic.forward_tac; |
|
5824 | 307 |
|
7419 | 308 |
end; |
309 |
||
310 |
||
8195 | 311 |
(* this *) |
312 |
||
8671 | 313 |
val this = METHOD (EVERY o map (HEADGOAL o Tactic.rtac)); |
8195 | 314 |
|
315 |
||
316 |
(* assumption *) |
|
7555 | 317 |
|
318 |
fun assm_tac ctxt = |
|
319 |
assume_tac APPEND' resolve_tac (filter Thm.no_prems (ProofContext.prems_of ctxt)); |
|
7419 | 320 |
|
7555 | 321 |
fun assumption_tac ctxt [] = assm_tac ctxt |
322 |
| assumption_tac _ [fact] = resolve_tac [fact] |
|
323 |
| assumption_tac _ _ = K no_tac; |
|
7419 | 324 |
|
8671 | 325 |
fun assumption ctxt = METHOD (HEADGOAL o assumption_tac ctxt); |
7419 | 326 |
|
327 |
||
8329 | 328 |
(* res_inst_tac emulations *) |
8238 | 329 |
|
8329 | 330 |
(*Note: insts refer to the implicit (!) goal context; use at your own risk*) |
8537 | 331 |
fun gen_res_inst tac (quant, (insts, thm)) = |
332 |
METHOD (fn facts => (quant (insert_tac facts THEN' tac insts thm))); |
|
8238 | 333 |
|
334 |
val res_inst = gen_res_inst Tactic.res_inst_tac; |
|
335 |
val eres_inst = gen_res_inst Tactic.eres_inst_tac; |
|
336 |
val dres_inst = gen_res_inst Tactic.dres_inst_tac; |
|
337 |
val forw_inst = gen_res_inst Tactic.forw_inst_tac; |
|
338 |
||
339 |
||
8329 | 340 |
(* simple Prolog interpreter *) |
341 |
||
342 |
fun prolog_tac rules facts = |
|
343 |
DEPTH_SOLVE_1 (HEADGOAL (Tactic.assume_tac APPEND' Tactic.resolve_tac (facts @ rules))); |
|
344 |
||
345 |
val prolog = METHOD o prolog_tac; |
|
346 |
||
347 |
||
8351 | 348 |
(* ML tactics *) |
349 |
||
350 |
val tactic_ref = ref ((fn _ => raise Match): Proof.context -> thm list -> tactic); |
|
351 |
fun set_tactic f = tactic_ref := f; |
|
352 |
||
353 |
fun tactic txt ctxt = METHOD (fn facts => |
|
8372 | 354 |
(Context.use_mltext |
8671 | 355 |
("let fun tactic (ctxt: PureIsar.Proof.context) (facts: thm list) : tactic = \ |
8613 | 356 |
\let val thm = PureIsar.ProofContext.get_thm ctxt and thms = PureIsar.ProofContext.get_thms ctxt in\n" |
8372 | 357 |
^ txt ^ |
8613 | 358 |
"\nend in PureIsar.Method.set_tactic tactic end") |
8372 | 359 |
false (Some (ProofContext.theory_of ctxt)); ! tactic_ref ctxt facts)); |
8351 | 360 |
|
361 |
||
5824 | 362 |
|
363 |
(** methods theory data **) |
|
364 |
||
365 |
(* data kind 'Isar/methods' *) |
|
366 |
||
367 |
structure MethodsDataArgs = |
|
368 |
struct |
|
369 |
val name = "Isar/methods"; |
|
370 |
type T = |
|
371 |
{space: NameSpace.T, |
|
372 |
meths: (((Args.src -> Proof.context -> Proof.method) * string) * stamp) Symtab.table}; |
|
373 |
||
374 |
val empty = {space = NameSpace.empty, meths = Symtab.empty}; |
|
6546 | 375 |
val copy = I; |
5824 | 376 |
val prep_ext = I; |
377 |
fun merge ({space = space1, meths = meths1}, {space = space2, meths = meths2}) = |
|
378 |
{space = NameSpace.merge (space1, space2), |
|
379 |
meths = Symtab.merge eq_snd (meths1, meths2) handle Symtab.DUPS dups => |
|
380 |
error ("Attempt to merge different versions of methods " ^ commas_quote dups)}; |
|
381 |
||
8720 | 382 |
fun pretty_meths verbose {space, meths} = |
5824 | 383 |
let |
384 |
fun prt_meth (name, ((_, comment), _)) = Pretty.block |
|
6849 | 385 |
[Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment]; |
5824 | 386 |
in |
8720 | 387 |
(if not verbose then [] else [Display.pretty_name_space ("method name space", space)]) @ |
388 |
[Pretty.big_list "methods:" (map prt_meth (NameSpace.cond_extern_table space meths))] |
|
5824 | 389 |
end; |
7367 | 390 |
|
8720 | 391 |
fun print _ = Pretty.writeln o Pretty.chunks o pretty_meths true; |
5824 | 392 |
end; |
393 |
||
394 |
structure MethodsData = TheoryDataFun(MethodsDataArgs); |
|
395 |
val print_methods = MethodsData.print; |
|
7611 | 396 |
|
8720 | 397 |
fun help_methods None = [Pretty.str "methods: (unkown theory context)"] |
398 |
| help_methods (Some thy) = MethodsDataArgs.pretty_meths false (MethodsData.get thy); |
|
5824 | 399 |
|
400 |
||
401 |
(* get methods *) |
|
402 |
||
5916 | 403 |
exception METHOD_FAIL of (string * Position.T) * exn; |
404 |
||
5824 | 405 |
fun method thy = |
406 |
let |
|
407 |
val {space, meths} = MethodsData.get thy; |
|
408 |
||
5884 | 409 |
fun meth src = |
410 |
let |
|
411 |
val ((raw_name, _), pos) = Args.dest_src src; |
|
412 |
val name = NameSpace.intern space raw_name; |
|
413 |
in |
|
5824 | 414 |
(case Symtab.lookup (meths, name) of |
415 |
None => error ("Unknown proof method: " ^ quote name ^ Position.str_of pos) |
|
5916 | 416 |
| Some ((mth, _), _) => transform_failure (curry METHOD_FAIL (name, pos)) (mth src)) |
5824 | 417 |
end; |
418 |
in meth end; |
|
419 |
||
420 |
||
9194 | 421 |
(* add_method(s) *) |
5824 | 422 |
|
423 |
fun add_methods raw_meths thy = |
|
424 |
let |
|
425 |
val full = Sign.full_name (Theory.sign_of thy); |
|
426 |
val new_meths = |
|
427 |
map (fn (name, f, comment) => (full name, ((f, comment), stamp ()))) raw_meths; |
|
428 |
||
429 |
val {space, meths} = MethodsData.get thy; |
|
430 |
val space' = NameSpace.extend (space, map fst new_meths); |
|
431 |
val meths' = Symtab.extend (meths, new_meths) handle Symtab.DUPS dups => |
|
432 |
error ("Duplicate declaration of method(s) " ^ commas_quote dups); |
|
433 |
in |
|
434 |
thy |> MethodsData.put {space = space', meths = meths'} |
|
435 |
end; |
|
436 |
||
9194 | 437 |
val add_method = add_methods o Library.single; |
438 |
||
5824 | 439 |
(*implicit version*) |
440 |
fun Method name meth cmt = Context.>> (add_methods [(name, meth, cmt)]); |
|
441 |
||
442 |
||
5884 | 443 |
|
444 |
(** method syntax **) |
|
5824 | 445 |
|
5884 | 446 |
(* basic *) |
447 |
||
448 |
fun syntax (scan: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list))) = |
|
449 |
Args.syntax "method" scan; |
|
5824 | 450 |
|
8351 | 451 |
fun simple_args scan f src ctxt : Proof.method = |
452 |
#2 (syntax (Scan.lift (scan >> (fn x => f x ctxt))) src ctxt); |
|
453 |
||
7555 | 454 |
fun ctxt_args (f: Proof.context -> Proof.method) src ctxt = |
8282 | 455 |
#2 (syntax (Scan.succeed (f ctxt)) src ctxt); |
7555 | 456 |
|
457 |
fun no_args m = ctxt_args (K m); |
|
5884 | 458 |
|
459 |
||
8351 | 460 |
|
5884 | 461 |
(* sections *) |
5824 | 462 |
|
7268 | 463 |
type modifier = (Proof.context -> Proof.context) * Proof.context attribute; |
464 |
||
465 |
local |
|
466 |
||
8381 | 467 |
fun sect ss = Scan.first (map Scan.lift ss); |
5884 | 468 |
fun thms ss = Scan.unless (sect ss) Attrib.local_thms; |
469 |
fun thmss ss = Scan.repeat (thms ss) >> flat; |
|
470 |
||
7268 | 471 |
fun apply (f, att) (ctxt, ths) = Thm.applys_attributes ((f ctxt, ths), [att]); |
5824 | 472 |
|
7268 | 473 |
fun section ss = (sect ss -- thmss ss) :-- (fn (m, ths) => Scan.depend (fn ctxt => |
474 |
Scan.succeed (apply m (ctxt, ths)))) >> #2; |
|
5884 | 475 |
|
7601 | 476 |
fun sectioned args ss = args -- Scan.repeat (section ss); |
5884 | 477 |
|
7268 | 478 |
in |
5824 | 479 |
|
5884 | 480 |
fun sectioned_args args ss f src ctxt = |
8282 | 481 |
let val (ctxt', (x, _)) = syntax (sectioned args ss) src ctxt |
5921 | 482 |
in f x ctxt' end; |
5884 | 483 |
|
7601 | 484 |
fun bang_sectioned_args ss f = sectioned_args Args.bang_facts ss f; |
485 |
fun only_sectioned_args ss f = sectioned_args (Scan.succeed ()) ss (fn () => f); |
|
7268 | 486 |
|
8093 | 487 |
fun thms_ctxt_args f = sectioned_args (thmss []) [] f; |
488 |
fun thms_args f = thms_ctxt_args (K o f); |
|
5824 | 489 |
|
7268 | 490 |
end; |
491 |
||
5824 | 492 |
|
8238 | 493 |
(* insts *) |
494 |
||
495 |
val insts = |
|
8537 | 496 |
Args.enum1 "and" (Scan.lift (Args.name -- Args.!!! (Args.$$$ "=" |-- Args.name))) -- |
8238 | 497 |
(Scan.lift (Args.$$$ "in") |-- Attrib.local_thm); |
498 |
||
8537 | 499 |
fun inst_args f = f oo (#2 oo syntax (Args.goal_spec HEADGOAL -- insts)); |
500 |
||
501 |
||
502 |
(* subgoal *) |
|
503 |
||
504 |
fun subgoal x = (Args.goal_spec HEADGOAL -- Scan.lift Args.name >> |
|
505 |
(fn (quant, s) => METHOD (fn facts => quant (insert_tac facts THEN' Tactic.subgoal_tac s)))) x; |
|
506 |
||
507 |
val subgoal_meth = #2 oo syntax subgoal; |
|
8238 | 508 |
|
509 |
||
5824 | 510 |
|
511 |
(** method text **) |
|
512 |
||
513 |
(* datatype text *) |
|
514 |
||
515 |
datatype text = |
|
516 |
Basic of (Proof.context -> Proof.method) | |
|
517 |
Source of Args.src | |
|
518 |
Then of text list | |
|
519 |
Orelse of text list | |
|
520 |
Try of text | |
|
521 |
Repeat1 of text; |
|
522 |
||
523 |
||
524 |
(* refine *) |
|
525 |
||
8238 | 526 |
fun gen_refine f text state = |
5824 | 527 |
let |
528 |
val thy = Proof.theory_of state; |
|
529 |
||
8238 | 530 |
fun eval (Basic mth) = f mth |
531 |
| eval (Source src) = f (method thy src) |
|
5824 | 532 |
| eval (Then txts) = Seq.EVERY (map eval txts) |
533 |
| eval (Orelse txts) = Seq.FIRST (map eval txts) |
|
534 |
| eval (Try txt) = Seq.TRY (eval txt) |
|
535 |
| eval (Repeat1 txt) = Seq.REPEAT1 (eval txt); |
|
536 |
in eval text state end; |
|
537 |
||
8238 | 538 |
val refine = gen_refine Proof.refine; |
539 |
val refine_end = gen_refine Proof.refine_end; |
|
6404 | 540 |
|
5824 | 541 |
|
6404 | 542 |
(* structured proof steps *) |
5824 | 543 |
|
7506 | 544 |
val default_text = Source (Args.src (("default", []), Position.none)); |
8195 | 545 |
val this_text = Basic (K this); |
8966 | 546 |
val done_text = Basic (K (METHOD0 all_tac)); |
7555 | 547 |
|
8966 | 548 |
fun close_text asm = Basic (fn ctxt => METHOD (K |
549 |
(FILTER Thm.no_prems ((if asm then ALLGOALS (assm_tac ctxt) else all_tac) THEN flexflex_tac)))); |
|
550 |
||
551 |
fun finish_text asm None = close_text asm |
|
552 |
| finish_text asm (Some txt) = Then [txt, close_text asm]; |
|
6872 | 553 |
|
5824 | 554 |
fun proof opt_text state = |
555 |
state |
|
556 |
|> Proof.assert_backward |
|
6404 | 557 |
|> refine (if_none opt_text default_text) |
8242 | 558 |
|> Seq.map (Proof.goal_facts (K [])) |
5824 | 559 |
|> Seq.map Proof.enter_forward; |
560 |
||
8966 | 561 |
fun local_qed asm opt_text = Proof.local_qed (refine (finish_text asm opt_text)); |
562 |
fun local_terminal_proof (text, opt_text) pr = |
|
563 |
Seq.THEN (proof (Some text), local_qed true opt_text pr); |
|
564 |
val local_default_proof = local_terminal_proof (default_text, None); |
|
8195 | 565 |
val local_immediate_proof = local_terminal_proof (this_text, None); |
8966 | 566 |
fun local_done_proof pr = Seq.THEN (proof (Some done_text), local_qed false None pr); |
5824 | 567 |
|
6872 | 568 |
|
8966 | 569 |
fun global_qeds asm opt_text = Proof.global_qed (refine (finish_text asm opt_text)); |
5824 | 570 |
|
8966 | 571 |
fun global_qed asm opt_text state = |
6872 | 572 |
state |
8966 | 573 |
|> global_qeds asm opt_text |
6872 | 574 |
|> Proof.check_result "Failed to finish proof" state |
575 |
|> Seq.hd; |
|
576 |
||
8966 | 577 |
fun global_term_proof asm (text, opt_text) state = |
6872 | 578 |
state |
579 |
|> proof (Some text) |
|
580 |
|> Proof.check_result "Terminal proof method failed" state |
|
8966 | 581 |
|> (Seq.flat o Seq.map (global_qeds asm opt_text)) |
6872 | 582 |
|> Proof.check_result "Failed to finish proof (after successful terminal method)" state |
583 |
|> Seq.hd; |
|
584 |
||
8966 | 585 |
val global_terminal_proof = global_term_proof true; |
6934 | 586 |
val global_default_proof = global_terminal_proof (default_text, None); |
8966 | 587 |
val global_immediate_proof = global_terminal_proof (this_text, None); |
588 |
val global_done_proof = global_term_proof false (done_text, None); |
|
5824 | 589 |
|
590 |
||
591 |
(** theory setup **) |
|
592 |
||
593 |
(* pure_methods *) |
|
594 |
||
595 |
val pure_methods = |
|
596 |
[("fail", no_args fail, "force failure"), |
|
597 |
("succeed", no_args succeed, "succeed"), |
|
7574 | 598 |
("-", no_args insert_facts, "do nothing, inserting current facts only"), |
7664 | 599 |
("insert", thms_args insert, "insert theorems, ignoring facts (improper!)"), |
7601 | 600 |
("unfold", thms_args unfold, "unfold definitions"), |
601 |
("fold", thms_args fold, "fold definitions"), |
|
8153
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
602 |
("default", thms_ctxt_args some_rule, "apply some rule"), |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
603 |
("rule", thms_ctxt_args some_rule, "apply some rule"), |
8220 | 604 |
("erule", thms_ctxt_args some_erule, "apply some rule in elimination manner (improper!)"), |
605 |
("drule", thms_ctxt_args some_drule, "apply some rule in destruct manner (improper!)"), |
|
606 |
("frule", thms_ctxt_args some_frule, "apply some rule in forward manner (improper!)"), |
|
8195 | 607 |
("this", no_args this, "apply current facts as rules"), |
8238 | 608 |
("assumption", ctxt_args assumption, "proof by assumption, preferring facts"), |
8537 | 609 |
("res_inst_tac", inst_args res_inst, "res_inst_tac emulation (dynamic instantiation!)"), |
610 |
("eres_inst_tac", inst_args eres_inst, "eres_inst_tac emulation (dynamic instantiation!)"), |
|
611 |
("dres_inst_tac", inst_args dres_inst, "dres_inst_tac emulation (dynamic instantiation!)"), |
|
612 |
("forw_inst_tac", inst_args forw_inst, "forw_inst_tac emulation (dynamic instantiation!)"), |
|
613 |
("subgoal_tac", subgoal_meth, "subgoal_tac emulation (dynamic instantiation!)"), |
|
8351 | 614 |
("prolog", thms_args prolog, "simple prolog interpreter"), |
615 |
("tactic", simple_args Args.name tactic, "ML tactic as proof method")]; |
|
5824 | 616 |
|
617 |
||
618 |
(* setup *) |
|
619 |
||
8153
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
620 |
val setup = |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
621 |
[GlobalRules.init, LocalRules.init, Attrib.add_attributes rule_atts, |
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
wenzelm
parents:
8093
diff
changeset
|
622 |
MethodsData.init, add_methods pure_methods]; |
5824 | 623 |
|
624 |
||
625 |
end; |
|
626 |
||
627 |
||
628 |
structure BasicMethod: BASIC_METHOD = Method; |
|
629 |
open BasicMethod; |