| author | wenzelm | 
| Fri, 30 Dec 2005 16:56:59 +0100 | |
| changeset 18525 | ce1ae48c320f | 
| parent 18374 | 598e7fd7438b | 
| child 18534 | 6799b38ed872 | 
| permissions | -rw-r--r-- | 
| 9938 | 1  | 
(* Title: Provers/classical.ML  | 
| 0 | 2  | 
ID: $Id$  | 
| 9938 | 3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 0 | 4  | 
Copyright 1992 University of Cambridge  | 
5  | 
||
6  | 
Theorem prover for classical reasoning, including predicate calculus, set  | 
|
7  | 
theory, etc.  | 
|
8  | 
||
| 
9563
 
216d053992a5
fixed classification of rules in atts and modifiers (final!?);
 
wenzelm 
parents: 
9513 
diff
changeset
 | 
9  | 
Rules must be classified as intro, elim, safe, hazardous (unsafe).  | 
| 0 | 10  | 
|
11  | 
A rule is unsafe unless it can be applied blindly without harmful results.  | 
|
12  | 
For a rule to be safe, its premises and conclusion should be logically  | 
|
13  | 
equivalent. There should be no variables in the premises that are not in  | 
|
14  | 
the conclusion.  | 
|
15  | 
*)  | 
|
16  | 
||
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
17  | 
(*higher precedence than := facilitates use of references*)  | 
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
18  | 
infix 4 addSIs addSEs addSDs addIs addEs addDs delrules  | 
| 4651 | 19  | 
addSWrapper delSWrapper addWrapper delWrapper  | 
| 
11181
 
d04f57b91166
renamed addaltern to addafter, addSaltern to addSafter
 
oheimb 
parents: 
11168 
diff
changeset
 | 
20  | 
addSbefore addSafter addbefore addafter  | 
| 5523 | 21  | 
addD2 addE2 addSD2 addSE2;  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
22  | 
|
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
23  | 
|
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
24  | 
(*should be a type abbreviation in signature CLASSICAL*)  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
25  | 
type netpair = (int * (bool * thm)) Net.net * (int * (bool * thm)) Net.net;  | 
| 4651 | 26  | 
type wrapper = (int -> tactic) -> (int -> tactic);  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
27  | 
|
| 0 | 28  | 
signature CLASSICAL_DATA =  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
29  | 
sig  | 
| 9171 | 30  | 
val make_elim : thm -> thm (* Tactic.make_elim or a classical version*)  | 
| 9938 | 31  | 
val mp : thm (* [| P-->Q; P |] ==> Q *)  | 
32  | 
val not_elim : thm (* [| ~P; P |] ==> R *)  | 
|
33  | 
val classical : thm (* (~P ==> P) ==> P *)  | 
|
34  | 
val sizef : thm -> int (* size function for BEST_FIRST *)  | 
|
| 0 | 35  | 
val hyp_subst_tacs: (int -> tactic) list  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
36  | 
end;  | 
| 0 | 37  | 
|
| 5841 | 38  | 
signature BASIC_CLASSICAL =  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
39  | 
sig  | 
| 0 | 40  | 
type claset  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
41  | 
val empty_cs: claset  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
42  | 
val print_cs: claset -> unit  | 
| 4380 | 43  | 
val print_claset: theory -> unit  | 
| 4653 | 44  | 
val rep_cs: (* BLAST_DATA in blast.ML dependent on this *)  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
45  | 
    claset -> {safeIs: thm list, safeEs: thm list,
 | 
| 9938 | 46  | 
hazIs: thm list, hazEs: thm list,  | 
| 10736 | 47  | 
swrappers: (string * wrapper) list,  | 
| 9938 | 48  | 
uwrappers: (string * wrapper) list,  | 
49  | 
safe0_netpair: netpair, safep_netpair: netpair,  | 
|
| 12401 | 50  | 
haz_netpair: netpair, dup_netpair: netpair,  | 
51  | 
xtra_netpair: ContextRules.netpair}  | 
|
| 9938 | 52  | 
val merge_cs : claset * claset -> claset  | 
53  | 
val addDs : claset * thm list -> claset  | 
|
54  | 
val addEs : claset * thm list -> claset  | 
|
55  | 
val addIs : claset * thm list -> claset  | 
|
56  | 
val addSDs : claset * thm list -> claset  | 
|
57  | 
val addSEs : claset * thm list -> claset  | 
|
58  | 
val addSIs : claset * thm list -> claset  | 
|
59  | 
val delrules : claset * thm list -> claset  | 
|
60  | 
val addSWrapper : claset * (string * wrapper) -> claset  | 
|
61  | 
val delSWrapper : claset * string -> claset  | 
|
62  | 
val addWrapper : claset * (string * wrapper) -> claset  | 
|
63  | 
val delWrapper : claset * string -> claset  | 
|
64  | 
val addSbefore : claset * (string * (int -> tactic)) -> claset  | 
|
| 
11181
 
d04f57b91166
renamed addaltern to addafter, addSaltern to addSafter
 
oheimb 
parents: 
11168 
diff
changeset
 | 
65  | 
val addSafter : claset * (string * (int -> tactic)) -> claset  | 
| 9938 | 66  | 
val addbefore : claset * (string * (int -> tactic)) -> claset  | 
| 
11181
 
d04f57b91166
renamed addaltern to addafter, addSaltern to addSafter
 
oheimb 
parents: 
11168 
diff
changeset
 | 
67  | 
val addafter : claset * (string * (int -> tactic)) -> claset  | 
| 5523 | 68  | 
val addD2 : claset * (string * thm) -> claset  | 
69  | 
val addE2 : claset * (string * thm) -> claset  | 
|
70  | 
val addSD2 : claset * (string * thm) -> claset  | 
|
71  | 
val addSE2 : claset * (string * thm) -> claset  | 
|
| 9938 | 72  | 
val appSWrappers : claset -> wrapper  | 
73  | 
val appWrappers : claset -> wrapper  | 
|
| 
982
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
74  | 
|
| 17880 | 75  | 
val change_claset_of: theory -> (claset -> claset) -> unit  | 
76  | 
val change_claset: (claset -> claset) -> unit  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
77  | 
val claset_of: theory -> claset  | 
| 17880 | 78  | 
val claset: unit -> claset  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
79  | 
val CLASET: (claset -> tactic) -> tactic  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
80  | 
val CLASET': (claset -> 'a -> tactic) -> 'a -> tactic  | 
| 15036 | 81  | 
val local_claset_of : Proof.context -> claset  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
82  | 
|
| 9938 | 83  | 
val fast_tac : claset -> int -> tactic  | 
84  | 
val slow_tac : claset -> int -> tactic  | 
|
85  | 
val weight_ASTAR : int ref  | 
|
86  | 
val astar_tac : claset -> int -> tactic  | 
|
87  | 
val slow_astar_tac : claset -> int -> tactic  | 
|
88  | 
val best_tac : claset -> int -> tactic  | 
|
89  | 
val first_best_tac : claset -> int -> tactic  | 
|
90  | 
val slow_best_tac : claset -> int -> tactic  | 
|
91  | 
val depth_tac : claset -> int -> int -> tactic  | 
|
92  | 
val deepen_tac : claset -> int -> int -> tactic  | 
|
| 
1587
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
93  | 
|
| 9938 | 94  | 
val contr_tac : int -> tactic  | 
95  | 
val dup_elim : thm -> thm  | 
|
96  | 
val dup_intr : thm -> thm  | 
|
97  | 
val dup_step_tac : claset -> int -> tactic  | 
|
98  | 
val eq_mp_tac : int -> tactic  | 
|
99  | 
val haz_step_tac : claset -> int -> tactic  | 
|
100  | 
val joinrules : thm list * thm list -> (bool * thm) list  | 
|
101  | 
val mp_tac : int -> tactic  | 
|
102  | 
val safe_tac : claset -> tactic  | 
|
103  | 
val safe_steps_tac : claset -> int -> tactic  | 
|
104  | 
val safe_step_tac : claset -> int -> tactic  | 
|
105  | 
val clarify_tac : claset -> int -> tactic  | 
|
106  | 
val clarify_step_tac : claset -> int -> tactic  | 
|
107  | 
val step_tac : claset -> int -> tactic  | 
|
108  | 
val slow_step_tac : claset -> int -> tactic  | 
|
109  | 
val swapify : thm list -> thm list  | 
|
110  | 
val swap_res_tac : thm list -> int -> tactic  | 
|
111  | 
val inst_step_tac : claset -> int -> tactic  | 
|
112  | 
val inst0_step_tac : claset -> int -> tactic  | 
|
113  | 
val instp_step_tac : claset -> int -> tactic  | 
|
| 1724 | 114  | 
|
| 9938 | 115  | 
val AddDs : thm list -> unit  | 
116  | 
val AddEs : thm list -> unit  | 
|
117  | 
val AddIs : thm list -> unit  | 
|
118  | 
val AddSDs : thm list -> unit  | 
|
119  | 
val AddSEs : thm list -> unit  | 
|
120  | 
val AddSIs : thm list -> unit  | 
|
121  | 
val Delrules : thm list -> unit  | 
|
122  | 
val Safe_tac : tactic  | 
|
123  | 
val Safe_step_tac : int -> tactic  | 
|
124  | 
val Clarify_tac : int -> tactic  | 
|
125  | 
val Clarify_step_tac : int -> tactic  | 
|
126  | 
val Step_tac : int -> tactic  | 
|
127  | 
val Fast_tac : int -> tactic  | 
|
128  | 
val Best_tac : int -> tactic  | 
|
129  | 
val Slow_tac : int -> tactic  | 
|
| 2066 | 130  | 
val Slow_best_tac : int -> tactic  | 
| 9938 | 131  | 
val Deepen_tac : int -> int -> tactic  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
132  | 
end;  | 
| 1724 | 133  | 
|
| 5841 | 134  | 
signature CLASSICAL =  | 
135  | 
sig  | 
|
136  | 
include BASIC_CLASSICAL  | 
|
| 18374 | 137  | 
val swap: thm (* ~P ==> (~Q ==> P) ==> Q *)  | 
| 15036 | 138  | 
val add_context_safe_wrapper: string * (Proof.context -> wrapper) -> theory -> theory  | 
139  | 
val del_context_safe_wrapper: string -> theory -> theory  | 
|
140  | 
val add_context_unsafe_wrapper: string * (Proof.context -> wrapper) -> theory -> theory  | 
|
141  | 
val del_context_unsafe_wrapper: string -> theory -> theory  | 
|
| 17880 | 142  | 
val get_claset: theory -> claset  | 
| 5841 | 143  | 
val print_local_claset: Proof.context -> unit  | 
144  | 
val get_local_claset: Proof.context -> claset  | 
|
145  | 
val put_local_claset: claset -> Proof.context -> Proof.context  | 
|
146  | 
val safe_dest_global: theory attribute  | 
|
147  | 
val safe_elim_global: theory attribute  | 
|
148  | 
val safe_intro_global: theory attribute  | 
|
| 6955 | 149  | 
val haz_dest_global: theory attribute  | 
150  | 
val haz_elim_global: theory attribute  | 
|
151  | 
val haz_intro_global: theory attribute  | 
|
| 9938 | 152  | 
val rule_del_global: theory attribute  | 
| 6955 | 153  | 
val safe_dest_local: Proof.context attribute  | 
154  | 
val safe_elim_local: Proof.context attribute  | 
|
155  | 
val safe_intro_local: Proof.context attribute  | 
|
| 5885 | 156  | 
val haz_dest_local: Proof.context attribute  | 
157  | 
val haz_elim_local: Proof.context attribute  | 
|
158  | 
val haz_intro_local: Proof.context attribute  | 
|
| 9938 | 159  | 
val rule_del_local: Proof.context attribute  | 
| 7272 | 160  | 
val cla_modifiers: (Args.T list -> (Method.modifier * Args.T list)) list  | 
| 7559 | 161  | 
val cla_meth: (claset -> tactic) -> thm list -> Proof.context -> Proof.method  | 
162  | 
val cla_meth': (claset -> int -> tactic) -> thm list -> Proof.context -> Proof.method  | 
|
| 15703 | 163  | 
val cla_method: (claset -> tactic) -> Method.src -> Proof.context -> Proof.method  | 
164  | 
val cla_method': (claset -> int -> tactic) -> Method.src -> Proof.context -> Proof.method  | 
|
| 5841 | 165  | 
val setup: (theory -> theory) list  | 
166  | 
end;  | 
|
167  | 
||
| 0 | 168  | 
|
| 5927 | 169  | 
functor ClassicalFun(Data: CLASSICAL_DATA): CLASSICAL =  | 
| 0 | 170  | 
struct  | 
171  | 
||
| 7354 | 172  | 
local open Data in  | 
| 0 | 173  | 
|
| 1800 | 174  | 
(*** Useful tactics for classical reasoning ***)  | 
| 0 | 175  | 
|
| 1524 | 176  | 
val imp_elim = (*cannot use bind_thm within a structure!*)  | 
| 9938 | 177  | 
  store_thm ("imp_elim", Data.make_elim mp);
 | 
| 0 | 178  | 
|
| 10736 | 179  | 
(*Prove goal that assumes both P and ~P.  | 
| 4392 | 180  | 
No backtracking if it finds an equal assumption. Perhaps should call  | 
181  | 
ematch_tac instead of eresolve_tac, but then cannot prove ZF/cantor.*)  | 
|
| 10736 | 182  | 
val contr_tac = eresolve_tac [not_elim] THEN'  | 
| 4392 | 183  | 
(eq_assume_tac ORELSE' assume_tac);  | 
| 0 | 184  | 
|
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
185  | 
(*Finds P-->Q and P in the assumptions, replaces implication by Q.  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
186  | 
Could do the same thing for P<->Q and P... *)  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
187  | 
fun mp_tac i = eresolve_tac [not_elim, imp_elim] i THEN assume_tac i;  | 
| 0 | 188  | 
|
189  | 
(*Like mp_tac but instantiates no variables*)  | 
|
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
190  | 
fun eq_mp_tac i = ematch_tac [not_elim, imp_elim] i THEN eq_assume_tac i;  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
191  | 
|
| 1524 | 192  | 
val swap =  | 
193  | 
  store_thm ("swap", rule_by_tactic (etac thin_rl 1) (not_elim RS classical));
 | 
|
| 0 | 194  | 
|
195  | 
(*Creates rules to eliminate ~A, from rules to introduce A*)  | 
|
196  | 
fun swapify intrs = intrs RLN (2, [swap]);  | 
|
| 12401 | 197  | 
fun swapped x = Attrib.no_args (fn (x, th) => (x, th RSN (2, swap))) x;  | 
| 0 | 198  | 
|
199  | 
(*Uses introduction rules in the normal way, or on negated assumptions,  | 
|
200  | 
trying rules in order. *)  | 
|
| 10736 | 201  | 
fun swap_res_tac rls =  | 
| 54 | 202  | 
let fun addrl (rl,brls) = (false, rl) :: (true, rl RSN (2,swap)) :: brls  | 
| 10736 | 203  | 
in assume_tac ORELSE'  | 
204  | 
contr_tac ORELSE'  | 
|
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
205  | 
biresolve_tac (foldr addrl [] rls)  | 
| 0 | 206  | 
end;  | 
207  | 
||
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
208  | 
(*Duplication of hazardous rules, for complete provers*)  | 
| 
2689
 
6d3d893453de
dup_intr & dup_elim no longer call standard -- this
 
paulson 
parents: 
2630 
diff
changeset
 | 
209  | 
fun dup_intr th = zero_var_indexes (th RS classical);  | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
210  | 
|
| 6967 | 211  | 
fun dup_elim th =  | 
| 13525 | 212  | 
(case try (fn () =>  | 
213  | 
rule_by_tactic (TRYALL (etac revcut_rl))  | 
|
| 
17257
 
0ab67cb765da
introduced binding priority 1 for linear combinators etc.
 
haftmann 
parents: 
17084 
diff
changeset
 | 
214  | 
((th RSN (2, revcut_rl)) |> assumption 2 |> Seq.hd)) () of  | 
| 15531 | 215  | 
SOME th' => th'  | 
| 6967 | 216  | 
  | _ => error ("Bad format for elimination rule\n" ^ string_of_thm th));
 | 
| 0 | 217  | 
|
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
218  | 
|
| 1800 | 219  | 
(**** Classical rule sets ****)  | 
| 0 | 220  | 
|
221  | 
datatype claset =  | 
|
| 12401 | 222  | 
  CS of {safeIs         : thm list,                (*safe introduction rules*)
 | 
223  | 
safeEs : thm list, (*safe elimination rules*)  | 
|
224  | 
hazIs : thm list, (*unsafe introduction rules*)  | 
|
225  | 
hazEs : thm list, (*unsafe elimination rules*)  | 
|
226  | 
swrappers : (string * wrapper) list, (*for transforming safe_step_tac*)  | 
|
| 9938 | 227  | 
uwrappers : (string * wrapper) list, (*for transforming step_tac*)  | 
| 12401 | 228  | 
safe0_netpair : netpair, (*nets for trivial cases*)  | 
229  | 
safep_netpair : netpair, (*nets for >0 subgoals*)  | 
|
230  | 
haz_netpair : netpair, (*nets for unsafe rules*)  | 
|
231  | 
dup_netpair : netpair, (*nets for duplication*)  | 
|
232  | 
xtra_netpair : ContextRules.netpair}; (*nets for extra rules*)  | 
|
| 0 | 233  | 
|
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
234  | 
(*Desired invariants are  | 
| 9938 | 235  | 
safe0_netpair = build safe0_brls,  | 
236  | 
safep_netpair = build safep_brls,  | 
|
237  | 
haz_netpair = build (joinrules(hazIs, hazEs)),  | 
|
| 10736 | 238  | 
dup_netpair = build (joinrules(map dup_intr hazIs,  | 
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
239  | 
map dup_elim hazEs))  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
240  | 
|
| 10736 | 241  | 
where build = build_netpair(Net.empty,Net.empty),  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
242  | 
safe0_brls contains all brules that solve the subgoal, and  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
243  | 
safep_brls contains all brules that generate 1 or more new subgoals.  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
244  | 
The theorem lists are largely comments, though they are used in merge_cs and print_cs.  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
245  | 
Nets must be built incrementally, to save space and time.  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
246  | 
*)  | 
| 0 | 247  | 
|
| 6502 | 248  | 
val empty_netpair = (Net.empty, Net.empty);  | 
249  | 
||
| 10736 | 250  | 
val empty_cs =  | 
| 9938 | 251  | 
  CS{safeIs     = [],
 | 
252  | 
safeEs = [],  | 
|
253  | 
hazIs = [],  | 
|
254  | 
hazEs = [],  | 
|
| 4651 | 255  | 
swrappers = [],  | 
256  | 
uwrappers = [],  | 
|
| 6502 | 257  | 
safe0_netpair = empty_netpair,  | 
258  | 
safep_netpair = empty_netpair,  | 
|
259  | 
haz_netpair = empty_netpair,  | 
|
| 6955 | 260  | 
dup_netpair = empty_netpair,  | 
261  | 
xtra_netpair = empty_netpair};  | 
|
| 0 | 262  | 
|
| 15036 | 263  | 
fun print_cs (CS {safeIs, safeEs, hazIs, hazEs, swrappers, uwrappers, ...}) =
 | 
| 3546 | 264  | 
let val pretty_thms = map Display.pretty_thm in  | 
| 9760 | 265  | 
[Pretty.big_list "safe introduction rules (intro!):" (pretty_thms safeIs),  | 
266  | 
Pretty.big_list "introduction rules (intro):" (pretty_thms hazIs),  | 
|
267  | 
Pretty.big_list "safe elimination rules (elim!):" (pretty_thms safeEs),  | 
|
| 15036 | 268  | 
Pretty.big_list "elimination rules (elim):" (pretty_thms hazEs),  | 
269  | 
      Pretty.strs ("safe wrappers:" :: map #1 swrappers),
 | 
|
270  | 
      Pretty.strs ("unsafe wrappers:" :: map #1 uwrappers)]
 | 
|
| 8727 | 271  | 
|> Pretty.chunks |> Pretty.writeln  | 
| 3546 | 272  | 
end;  | 
| 0 | 273  | 
|
| 4653 | 274  | 
fun rep_cs (CS args) = args;  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
275  | 
|
| 10736 | 276  | 
local  | 
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
277  | 
fun wrap l tac = foldr (fn ((name,tacf),w) => tacf w) tac l;  | 
| 10736 | 278  | 
in  | 
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
279  | 
  fun appSWrappers (CS{swrappers,...}) = wrap swrappers;
 | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
280  | 
  fun appWrappers  (CS{uwrappers,...}) = wrap uwrappers;
 | 
| 4651 | 281  | 
end;  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
282  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
283  | 
|
| 1800 | 284  | 
(*** Adding (un)safe introduction or elimination rules.  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
285  | 
|
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
286  | 
In case of overlap, new rules are tried BEFORE old ones!!  | 
| 1800 | 287  | 
***)  | 
| 0 | 288  | 
|
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
289  | 
(*For use with biresolve_tac. Combines intro rules with swap to handle negated  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
290  | 
assumptions. Pairs elim rules with true. *)  | 
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
291  | 
fun joinrules (intrs, elims) =  | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
292  | 
(map (pair true) (elims @ swapify intrs) @ map (pair false) intrs);  | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
293  | 
|
| 12401 | 294  | 
fun joinrules' (intrs, elims) =  | 
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
295  | 
(map (pair true) elims @ map (pair false) intrs);  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
296  | 
|
| 10736 | 297  | 
(*Priority: prefer rules with fewest subgoals,  | 
| 1231 | 298  | 
then rules added most recently (preferring the head of the list).*)  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
299  | 
fun tag_brls k [] = []  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
300  | 
| tag_brls k (brl::brls) =  | 
| 10736 | 301  | 
(1000000*subgoals_of_brl brl + k, brl) ::  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
302  | 
tag_brls (k+1) brls;  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
303  | 
|
| 12401 | 304  | 
fun tag_brls' _ _ [] = []  | 
305  | 
| tag_brls' w k (brl::brls) = ((w, k), brl) :: tag_brls' w (k + 1) brls;  | 
|
| 10736 | 306  | 
|
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
307  | 
fun insert_tagged_list kbrls netpr = foldr Tactic.insert_tagged_brl netpr kbrls;  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
308  | 
|
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
309  | 
(*Insert into netpair that already has nI intr rules and nE elim rules.  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
310  | 
Count the intr rules double (to account for swapify). Negate to give the  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
311  | 
new insertions the lowest priority.*)  | 
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
312  | 
fun insert (nI, nE) = insert_tagged_list o (tag_brls (~(2*nI+nE))) o joinrules;  | 
| 12401 | 313  | 
fun insert' w (nI, nE) = insert_tagged_list o tag_brls' w (~(nI + nE)) o joinrules';  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
314  | 
|
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
315  | 
fun delete_tagged_list brls netpr = foldr Tactic.delete_tagged_brl netpr brls;  | 
| 12362 | 316  | 
fun delete x = delete_tagged_list (joinrules x);  | 
| 12401 | 317  | 
fun delete' x = delete_tagged_list (joinrules' x);  | 
| 1800 | 318  | 
|
| 
13105
 
3d1e7a199bdc
use eq_thm_prop instead of slightly inadequate eq_thm;
 
wenzelm 
parents: 
12401 
diff
changeset
 | 
319  | 
val mem_thm = gen_mem Drule.eq_thm_prop  | 
| 
 
3d1e7a199bdc
use eq_thm_prop instead of slightly inadequate eq_thm;
 
wenzelm 
parents: 
12401 
diff
changeset
 | 
320  | 
and rem_thm = gen_rem Drule.eq_thm_prop;  | 
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
321  | 
|
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
322  | 
(*Warn if the rule is already present ELSEWHERE in the claset. The addition  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
323  | 
is still allowed.*)  | 
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
324  | 
fun warn_dup th (CS{safeIs, safeEs, hazIs, hazEs, ...}) =
 | 
| 10736 | 325  | 
if mem_thm (th, safeIs) then  | 
| 9938 | 326  | 
         warning ("Rule already declared as safe introduction (intro!)\n" ^ string_of_thm th)
 | 
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
327  | 
else if mem_thm (th, safeEs) then  | 
| 9408 | 328  | 
         warning ("Rule already declared as safe elimination (elim!)\n" ^ string_of_thm th)
 | 
| 10736 | 329  | 
else if mem_thm (th, hazIs) then  | 
| 9760 | 330  | 
         warning ("Rule already declared as introduction (intro)\n" ^ string_of_thm th)
 | 
| 10736 | 331  | 
else if mem_thm (th, hazEs) then  | 
| 9760 | 332  | 
         warning ("Rule already declared as elimination (elim)\n" ^ string_of_thm th)
 | 
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
333  | 
else ();  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
334  | 
|
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
335  | 
|
| 1800 | 336  | 
(*** Safe rules ***)  | 
| 
982
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
337  | 
|
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
338  | 
fun addSI (cs as CS{safeIs, safeEs, hazIs, hazEs, swrappers, uwrappers,
 | 
| 9938 | 339  | 
safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair},  | 
340  | 
th) =  | 
|
| 10736 | 341  | 
if mem_thm (th, safeIs) then  | 
| 9938 | 342  | 
         (warning ("Ignoring duplicate safe introduction (intro!)\n" ^ string_of_thm th);
 | 
343  | 
cs)  | 
|
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
344  | 
else  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
345  | 
let val (safe0_rls, safep_rls) = (*0 subgoals vs 1 or more*)  | 
| 15570 | 346  | 
List.partition Thm.no_prems [th]  | 
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
347  | 
val nI = length safeIs + 1  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
348  | 
and nE = length safeEs  | 
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
349  | 
in warn_dup th cs;  | 
| 9938 | 350  | 
     CS{safeIs  = th::safeIs,
 | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
351  | 
safe0_netpair = insert (nI,nE) (safe0_rls, []) safe0_netpair,  | 
| 9938 | 352  | 
safep_netpair = insert (nI,nE) (safep_rls, []) safep_netpair,  | 
353  | 
safeEs = safeEs,  | 
|
354  | 
hazIs = hazIs,  | 
|
355  | 
hazEs = hazEs,  | 
|
356  | 
swrappers = swrappers,  | 
|
357  | 
uwrappers = uwrappers,  | 
|
358  | 
haz_netpair = haz_netpair,  | 
|
359  | 
dup_netpair = dup_netpair,  | 
|
| 12401 | 360  | 
xtra_netpair = insert' 0 (nI,nE) ([th], []) xtra_netpair}  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
361  | 
end;  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
362  | 
|
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
363  | 
fun addSE (cs as CS{safeIs, safeEs, hazIs, hazEs, swrappers, uwrappers,
 | 
| 9938 | 364  | 
safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair},  | 
365  | 
th) =  | 
|
| 10736 | 366  | 
if mem_thm (th, safeEs) then  | 
| 9938 | 367  | 
         (warning ("Ignoring duplicate safe elimination (elim!)\n" ^ string_of_thm th);
 | 
368  | 
cs)  | 
|
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
369  | 
else  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
370  | 
let val (safe0_rls, safep_rls) = (*0 subgoals vs 1 or more*)  | 
| 15570 | 371  | 
List.partition (fn rl => nprems_of rl=1) [th]  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
372  | 
val nI = length safeIs  | 
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
373  | 
and nE = length safeEs + 1  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
374  | 
in warn_dup th cs;  | 
| 9938 | 375  | 
     CS{safeEs  = th::safeEs,
 | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
376  | 
safe0_netpair = insert (nI,nE) ([], safe0_rls) safe0_netpair,  | 
| 9938 | 377  | 
safep_netpair = insert (nI,nE) ([], safep_rls) safep_netpair,  | 
378  | 
safeIs = safeIs,  | 
|
379  | 
hazIs = hazIs,  | 
|
380  | 
hazEs = hazEs,  | 
|
381  | 
swrappers = swrappers,  | 
|
382  | 
uwrappers = uwrappers,  | 
|
383  | 
haz_netpair = haz_netpair,  | 
|
384  | 
dup_netpair = dup_netpair,  | 
|
| 12401 | 385  | 
xtra_netpair = insert' 0 (nI,nE) ([], [th]) xtra_netpair}  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
386  | 
end;  | 
| 0 | 387  | 
|
| 15570 | 388  | 
fun rev_foldl f (e, l) = Library.foldl f (e, rev l);  | 
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
389  | 
|
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
390  | 
val op addSIs = rev_foldl addSI;  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
391  | 
val op addSEs = rev_foldl addSE;  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
392  | 
|
| 
17084
 
fb0a80aef0be
classical rules must have names for ATP integration
 
paulson 
parents: 
17057 
diff
changeset
 | 
393  | 
(*Give new theorem a name, if it has one already.*)  | 
| 
 
fb0a80aef0be
classical rules must have names for ATP integration
 
paulson 
parents: 
17057 
diff
changeset
 | 
394  | 
fun name_make_elim th =  | 
| 
 
fb0a80aef0be
classical rules must have names for ATP integration
 
paulson 
parents: 
17057 
diff
changeset
 | 
395  | 
case Thm.name_of_thm th of  | 
| 
 
fb0a80aef0be
classical rules must have names for ATP integration
 
paulson 
parents: 
17057 
diff
changeset
 | 
396  | 
"" => Data.make_elim th  | 
| 
 
fb0a80aef0be
classical rules must have names for ATP integration
 
paulson 
parents: 
17057 
diff
changeset
 | 
397  | 
| a => Thm.name_thm (a ^ "_dest", Data.make_elim th);  | 
| 
 
fb0a80aef0be
classical rules must have names for ATP integration
 
paulson 
parents: 
17057 
diff
changeset
 | 
398  | 
|
| 
 
fb0a80aef0be
classical rules must have names for ATP integration
 
paulson 
parents: 
17057 
diff
changeset
 | 
399  | 
fun cs addSDs ths = cs addSEs (map name_make_elim ths);  | 
| 0 | 400  | 
|
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
401  | 
|
| 1800 | 402  | 
(*** Hazardous (unsafe) rules ***)  | 
| 0 | 403  | 
|
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
404  | 
fun addI (cs as CS{safeIs, safeEs, hazIs, hazEs, swrappers, uwrappers,
 | 
| 9938 | 405  | 
safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair},  | 
406  | 
th)=  | 
|
| 10736 | 407  | 
if mem_thm (th, hazIs) then  | 
| 9938 | 408  | 
         (warning ("Ignoring duplicate introduction (intro)\n" ^ string_of_thm th);
 | 
409  | 
cs)  | 
|
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
410  | 
else  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
411  | 
let val nI = length hazIs + 1  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
412  | 
and nE = length hazEs  | 
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
413  | 
in warn_dup th cs;  | 
| 9938 | 414  | 
     CS{hazIs   = th::hazIs,
 | 
415  | 
haz_netpair = insert (nI,nE) ([th], []) haz_netpair,  | 
|
416  | 
dup_netpair = insert (nI,nE) (map dup_intr [th], []) dup_netpair,  | 
|
| 10736 | 417  | 
safeIs = safeIs,  | 
| 9938 | 418  | 
safeEs = safeEs,  | 
419  | 
hazEs = hazEs,  | 
|
420  | 
swrappers = swrappers,  | 
|
421  | 
uwrappers = uwrappers,  | 
|
422  | 
safe0_netpair = safe0_netpair,  | 
|
423  | 
safep_netpair = safep_netpair,  | 
|
| 12401 | 424  | 
xtra_netpair = insert' 1 (nI,nE) ([th], []) xtra_netpair}  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
425  | 
end;  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
426  | 
|
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
427  | 
fun addE (cs as CS{safeIs, safeEs, hazIs, hazEs, swrappers, uwrappers,
 | 
| 9938 | 428  | 
safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair},  | 
429  | 
th) =  | 
|
| 10736 | 430  | 
if mem_thm (th, hazEs) then  | 
| 9938 | 431  | 
         (warning ("Ignoring duplicate elimination (elim)\n" ^ string_of_thm th);
 | 
432  | 
cs)  | 
|
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
433  | 
else  | 
| 10736 | 434  | 
let val nI = length hazIs  | 
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
435  | 
and nE = length hazEs + 1  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
436  | 
in warn_dup th cs;  | 
| 9938 | 437  | 
     CS{hazEs   = th::hazEs,
 | 
438  | 
haz_netpair = insert (nI,nE) ([], [th]) haz_netpair,  | 
|
439  | 
dup_netpair = insert (nI,nE) ([], map dup_elim [th]) dup_netpair,  | 
|
| 10736 | 440  | 
safeIs = safeIs,  | 
| 9938 | 441  | 
safeEs = safeEs,  | 
442  | 
hazIs = hazIs,  | 
|
443  | 
swrappers = swrappers,  | 
|
444  | 
uwrappers = uwrappers,  | 
|
445  | 
safe0_netpair = safe0_netpair,  | 
|
446  | 
safep_netpair = safep_netpair,  | 
|
| 12401 | 447  | 
xtra_netpair = insert' 1 (nI,nE) ([], [th]) xtra_netpair}  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
448  | 
end;  | 
| 0 | 449  | 
|
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
450  | 
val op addIs = rev_foldl addI;  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
451  | 
val op addEs = rev_foldl addE;  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
452  | 
|
| 
17084
 
fb0a80aef0be
classical rules must have names for ATP integration
 
paulson 
parents: 
17057 
diff
changeset
 | 
453  | 
fun cs addDs ths = cs addEs (map name_make_elim ths);  | 
| 0 | 454  | 
|
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
455  | 
|
| 10736 | 456  | 
(*** Deletion of rules  | 
| 1800 | 457  | 
Working out what to delete, requires repeating much of the code used  | 
| 9938 | 458  | 
to insert.  | 
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
459  | 
Separate functions delSI, etc., are not exported; instead delrules  | 
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
460  | 
searches in all the lists and chooses the relevant delXX functions.  | 
| 1800 | 461  | 
***)  | 
462  | 
||
| 10736 | 463  | 
fun delSI th  | 
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
464  | 
          (cs as CS{safeIs, safeEs, hazIs, hazEs, swrappers, uwrappers,
 | 
| 9938 | 465  | 
safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair}) =  | 
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
466  | 
if mem_thm (th, safeIs) then  | 
| 15570 | 467  | 
let val (safe0_rls, safep_rls) = List.partition Thm.no_prems [th]  | 
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
468  | 
   in CS{safe0_netpair = delete (safe0_rls, []) safe0_netpair,
 | 
| 9938 | 469  | 
safep_netpair = delete (safep_rls, []) safep_netpair,  | 
470  | 
safeIs = rem_thm (safeIs,th),  | 
|
471  | 
safeEs = safeEs,  | 
|
472  | 
hazIs = hazIs,  | 
|
473  | 
hazEs = hazEs,  | 
|
474  | 
swrappers = swrappers,  | 
|
475  | 
uwrappers = uwrappers,  | 
|
476  | 
haz_netpair = haz_netpair,  | 
|
477  | 
dup_netpair = dup_netpair,  | 
|
| 12401 | 478  | 
xtra_netpair = delete' ([th], []) xtra_netpair}  | 
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
479  | 
end  | 
| 
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
480  | 
else cs;  | 
| 1800 | 481  | 
|
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
482  | 
fun delSE th  | 
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
483  | 
          (cs as CS{safeIs, safeEs, hazIs, hazEs, swrappers, uwrappers,
 | 
| 9938 | 484  | 
safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair}) =  | 
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
485  | 
if mem_thm (th, safeEs) then  | 
| 15570 | 486  | 
let val (safe0_rls, safep_rls) = List.partition (fn rl => nprems_of rl=1) [th]  | 
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
487  | 
   in CS{safe0_netpair = delete ([], safe0_rls) safe0_netpair,
 | 
| 9938 | 488  | 
safep_netpair = delete ([], safep_rls) safep_netpair,  | 
489  | 
safeIs = safeIs,  | 
|
490  | 
safeEs = rem_thm (safeEs,th),  | 
|
491  | 
hazIs = hazIs,  | 
|
492  | 
hazEs = hazEs,  | 
|
493  | 
swrappers = swrappers,  | 
|
494  | 
uwrappers = uwrappers,  | 
|
495  | 
haz_netpair = haz_netpair,  | 
|
496  | 
dup_netpair = dup_netpair,  | 
|
| 12401 | 497  | 
xtra_netpair = delete' ([], [th]) xtra_netpair}  | 
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
498  | 
end  | 
| 
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
499  | 
else cs;  | 
| 1800 | 500  | 
|
501  | 
||
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
502  | 
fun delI th  | 
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
503  | 
         (cs as CS{safeIs, safeEs, hazIs, hazEs, swrappers, uwrappers,
 | 
| 9938 | 504  | 
safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair}) =  | 
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
505  | 
if mem_thm (th, hazIs) then  | 
| 
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
506  | 
     CS{haz_netpair = delete ([th], []) haz_netpair,
 | 
| 9938 | 507  | 
dup_netpair = delete ([dup_intr th], []) dup_netpair,  | 
| 10736 | 508  | 
safeIs = safeIs,  | 
| 9938 | 509  | 
safeEs = safeEs,  | 
510  | 
hazIs = rem_thm (hazIs,th),  | 
|
511  | 
hazEs = hazEs,  | 
|
512  | 
swrappers = swrappers,  | 
|
513  | 
uwrappers = uwrappers,  | 
|
514  | 
safe0_netpair = safe0_netpair,  | 
|
515  | 
safep_netpair = safep_netpair,  | 
|
| 12401 | 516  | 
xtra_netpair = delete' ([th], []) xtra_netpair}  | 
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
517  | 
else cs;  | 
| 1800 | 518  | 
|
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
519  | 
fun delE th  | 
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
520  | 
         (cs as CS{safeIs, safeEs, hazIs, hazEs, swrappers, uwrappers,
 | 
| 9938 | 521  | 
safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair}) =  | 
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
522  | 
if mem_thm (th, hazEs) then  | 
| 
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
523  | 
     CS{haz_netpair = delete ([], [th]) haz_netpair,
 | 
| 9938 | 524  | 
dup_netpair = delete ([], [dup_elim th]) dup_netpair,  | 
| 10736 | 525  | 
safeIs = safeIs,  | 
| 9938 | 526  | 
safeEs = safeEs,  | 
527  | 
hazIs = hazIs,  | 
|
528  | 
hazEs = rem_thm (hazEs,th),  | 
|
529  | 
swrappers = swrappers,  | 
|
530  | 
uwrappers = uwrappers,  | 
|
531  | 
safe0_netpair = safe0_netpair,  | 
|
532  | 
safep_netpair = safep_netpair,  | 
|
| 12401 | 533  | 
xtra_netpair = delete' ([], [th]) xtra_netpair}  | 
| 6955 | 534  | 
else cs;  | 
535  | 
||
| 1800 | 536  | 
|
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
537  | 
(*Delete ALL occurrences of "th" in the claset (perhaps from several lists)*)  | 
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
538  | 
fun delrule (cs as CS{safeIs, safeEs, hazIs, hazEs, ...}, th) =
 | 
| 9938 | 539  | 
let val th' = Data.make_elim th in  | 
540  | 
if mem_thm (th, safeIs) orelse mem_thm (th, safeEs) orelse  | 
|
541  | 
mem_thm (th, hazIs) orelse mem_thm (th, hazEs) orelse  | 
|
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
542  | 
mem_thm (th', safeEs) orelse mem_thm (th', hazEs)  | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
543  | 
then delSI th (delSE th (delI th (delE th (delSE th' (delE th' cs)))))  | 
| 9938 | 544  | 
    else (warning ("Undeclared classical rule\n" ^ (string_of_thm th)); cs)
 | 
545  | 
end;  | 
|
| 1800 | 546  | 
|
| 15570 | 547  | 
val op delrules = Library.foldl delrule;  | 
| 1800 | 548  | 
|
549  | 
||
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
550  | 
(*** Modifying the wrapper tacticals ***)  | 
| 10736 | 551  | 
fun update_swrappers  | 
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
552  | 
(CS{safeIs, safeEs, hazIs, hazEs, swrappers, uwrappers,
 | 
| 6955 | 553  | 
safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair}) f =  | 
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
554  | 
 CS{safeIs = safeIs, safeEs = safeEs, hazIs = hazIs, hazEs = hazEs,
 | 
| 
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
555  | 
swrappers = f swrappers, uwrappers = uwrappers,  | 
| 
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
556  | 
safe0_netpair = safe0_netpair, safep_netpair = safep_netpair,  | 
| 6955 | 557  | 
haz_netpair = haz_netpair, dup_netpair = dup_netpair, xtra_netpair = xtra_netpair};  | 
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
558  | 
|
| 10736 | 559  | 
fun update_uwrappers  | 
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
560  | 
(CS{safeIs, safeEs, hazIs, hazEs, swrappers, uwrappers,
 | 
| 6955 | 561  | 
safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair}) f =  | 
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
562  | 
 CS{safeIs = safeIs, safeEs = safeEs, hazIs = hazIs, hazEs = hazEs,
 | 
| 
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
563  | 
swrappers = swrappers, uwrappers = f uwrappers,  | 
| 
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
564  | 
safe0_netpair = safe0_netpair, safep_netpair = safep_netpair,  | 
| 6955 | 565  | 
haz_netpair = haz_netpair, dup_netpair = dup_netpair, xtra_netpair = xtra_netpair};  | 
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
566  | 
|
| 
982
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
567  | 
|
| 4651 | 568  | 
(*Add/replace a safe wrapper*)  | 
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
569  | 
fun cs addSWrapper new_swrapper = update_swrappers cs (fn swrappers =>  | 
| 9721 | 570  | 
overwrite_warn (swrappers, new_swrapper)  | 
571  | 
       ("Overwriting safe wrapper " ^ fst new_swrapper));
 | 
|
| 4651 | 572  | 
|
573  | 
(*Add/replace an unsafe wrapper*)  | 
|
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
574  | 
fun cs addWrapper new_uwrapper = update_uwrappers cs (fn uwrappers =>  | 
| 9721 | 575  | 
overwrite_warn (uwrappers, new_uwrapper)  | 
| 9938 | 576  | 
        ("Overwriting unsafe wrapper "^fst new_uwrapper));
 | 
| 
982
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
577  | 
|
| 4651 | 578  | 
(*Remove a safe wrapper*)  | 
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
579  | 
fun cs delSWrapper name = update_swrappers cs (fn swrappers =>  | 
| 17795 | 580  | 
let val swrappers' = filter_out (equal name o fst) swrappers in  | 
| 15036 | 581  | 
if length swrappers <> length swrappers' then swrappers'  | 
582  | 
    else (warning ("No such safe wrapper in claset: "^ name); swrappers)
 | 
|
583  | 
end);  | 
|
| 
982
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
584  | 
|
| 4651 | 585  | 
(*Remove an unsafe wrapper*)  | 
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
586  | 
fun cs delWrapper name = update_uwrappers cs (fn uwrappers =>  | 
| 17795 | 587  | 
let val uwrappers' = filter_out (equal name o fst) uwrappers in  | 
| 15036 | 588  | 
if length uwrappers <> length uwrappers' then uwrappers'  | 
589  | 
    else (warning ("No such unsafe wrapper in claset: " ^ name); uwrappers)
 | 
|
590  | 
end);  | 
|
| 
982
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
591  | 
|
| 11168 | 592  | 
(* compose a safe tactic alternatively before/after safe_step_tac *)  | 
| 10736 | 593  | 
fun cs addSbefore (name, tac1) =  | 
| 5523 | 594  | 
cs addSWrapper (name, fn tac2 => tac1 ORELSE' tac2);  | 
| 
11181
 
d04f57b91166
renamed addaltern to addafter, addSaltern to addSafter
 
oheimb 
parents: 
11168 
diff
changeset
 | 
595  | 
fun cs addSafter (name, tac2) =  | 
| 5523 | 596  | 
cs addSWrapper (name, fn tac1 => tac1 ORELSE' tac2);  | 
| 
982
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
597  | 
|
| 11168 | 598  | 
(*compose a tactic alternatively before/after the step tactic *)  | 
| 10736 | 599  | 
fun cs addbefore (name, tac1) =  | 
| 5523 | 600  | 
cs addWrapper (name, fn tac2 => tac1 APPEND' tac2);  | 
| 
11181
 
d04f57b91166
renamed addaltern to addafter, addSaltern to addSafter
 
oheimb 
parents: 
11168 
diff
changeset
 | 
601  | 
fun cs addafter (name, tac2) =  | 
| 5523 | 602  | 
cs addWrapper (name, fn tac1 => tac1 APPEND' tac2);  | 
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
603  | 
|
| 10736 | 604  | 
fun cs addD2 (name, thm) =  | 
| 
11181
 
d04f57b91166
renamed addaltern to addafter, addSaltern to addSafter
 
oheimb 
parents: 
11168 
diff
changeset
 | 
605  | 
cs addafter (name, datac thm 1);  | 
| 10736 | 606  | 
fun cs addE2 (name, thm) =  | 
| 
11181
 
d04f57b91166
renamed addaltern to addafter, addSaltern to addSafter
 
oheimb 
parents: 
11168 
diff
changeset
 | 
607  | 
cs addafter (name, eatac thm 1);  | 
| 
 
d04f57b91166
renamed addaltern to addafter, addSaltern to addSafter
 
oheimb 
parents: 
11168 
diff
changeset
 | 
608  | 
fun cs addSD2 (name, thm) =  | 
| 
 
d04f57b91166
renamed addaltern to addafter, addSaltern to addSafter
 
oheimb 
parents: 
11168 
diff
changeset
 | 
609  | 
cs addSafter (name, dmatch_tac [thm] THEN' eq_assume_tac);  | 
| 
 
d04f57b91166
renamed addaltern to addafter, addSaltern to addSafter
 
oheimb 
parents: 
11168 
diff
changeset
 | 
610  | 
fun cs addSE2 (name, thm) =  | 
| 
 
d04f57b91166
renamed addaltern to addafter, addSaltern to addSafter
 
oheimb 
parents: 
11168 
diff
changeset
 | 
611  | 
cs addSafter (name, ematch_tac [thm] THEN' eq_assume_tac);  | 
| 
982
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
612  | 
|
| 1711 | 613  | 
(*Merge works by adding all new rules of the 2nd claset into the 1st claset.  | 
614  | 
Merging the term nets may look more efficient, but the rather delicate  | 
|
615  | 
treatment of priority might get muddled up.*)  | 
|
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
616  | 
fun merge_cs (cs as CS{safeIs, safeEs, hazIs, hazEs, ...},
 | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
617  | 
     CS{safeIs=safeIs2, safeEs=safeEs2, hazIs=hazIs2, hazEs=hazEs2, swrappers, uwrappers, ...}) =
 | 
| 
13105
 
3d1e7a199bdc
use eq_thm_prop instead of slightly inadequate eq_thm;
 
wenzelm 
parents: 
12401 
diff
changeset
 | 
618  | 
let val safeIs' = gen_rems Drule.eq_thm_prop (safeIs2,safeIs)  | 
| 
 
3d1e7a199bdc
use eq_thm_prop instead of slightly inadequate eq_thm;
 
wenzelm 
parents: 
12401 
diff
changeset
 | 
619  | 
val safeEs' = gen_rems Drule.eq_thm_prop (safeEs2,safeEs)  | 
| 
 
3d1e7a199bdc
use eq_thm_prop instead of slightly inadequate eq_thm;
 
wenzelm 
parents: 
12401 
diff
changeset
 | 
620  | 
val hazIs' = gen_rems Drule.eq_thm_prop (hazIs2, hazIs)  | 
| 
 
3d1e7a199bdc
use eq_thm_prop instead of slightly inadequate eq_thm;
 
wenzelm 
parents: 
12401 
diff
changeset
 | 
621  | 
val hazEs' = gen_rems Drule.eq_thm_prop (hazEs2, hazEs)  | 
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
622  | 
val cs1 = cs addSIs safeIs'  | 
| 9938 | 623  | 
addSEs safeEs'  | 
624  | 
addIs hazIs'  | 
|
625  | 
addEs hazEs'  | 
|
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
626  | 
val cs2 = update_swrappers cs1 (fn ws => merge_alists ws swrappers);  | 
| 
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
627  | 
val cs3 = update_uwrappers cs2 (fn ws => merge_alists ws uwrappers);  | 
| 10736 | 628  | 
in cs3  | 
| 1711 | 629  | 
end;  | 
630  | 
||
| 
982
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
631  | 
|
| 1800 | 632  | 
(**** Simple tactics for theorem proving ****)  | 
| 0 | 633  | 
|
634  | 
(*Attack subgoals using safe inferences -- matching, not resolution*)  | 
|
| 10736 | 635  | 
fun safe_step_tac (cs as CS{safe0_netpair,safep_netpair,...}) =
 | 
| 4651 | 636  | 
appSWrappers cs (FIRST' [  | 
| 9938 | 637  | 
eq_assume_tac,  | 
638  | 
eq_mp_tac,  | 
|
639  | 
bimatch_from_nets_tac safe0_netpair,  | 
|
640  | 
FIRST' hyp_subst_tacs,  | 
|
641  | 
bimatch_from_nets_tac safep_netpair]);  | 
|
| 0 | 642  | 
|
| 5757 | 643  | 
(*Repeatedly attack a subgoal using safe inferences -- it's deterministic!*)  | 
| 10736 | 644  | 
fun safe_steps_tac cs = REPEAT_DETERM1 o  | 
| 9938 | 645  | 
(fn i => COND (has_fewer_prems i) no_tac (safe_step_tac cs i));  | 
| 5757 | 646  | 
|
| 0 | 647  | 
(*Repeatedly attack subgoals using safe inferences -- it's deterministic!*)  | 
| 5757 | 648  | 
fun safe_tac cs = REPEAT_DETERM1 (FIRSTGOAL (safe_steps_tac cs));  | 
| 
747
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
649  | 
|
| 
3705
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
650  | 
|
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
651  | 
(*** Clarify_tac: do safe steps without causing branching ***)  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
652  | 
|
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
653  | 
fun nsubgoalsP n (k,brl) = (subgoals_of_brl brl = n);  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
654  | 
|
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
655  | 
(*version of bimatch_from_nets_tac that only applies rules that  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
656  | 
create precisely n subgoals.*)  | 
| 10736 | 657  | 
fun n_bimatch_from_nets_tac n =  | 
| 15570 | 658  | 
biresolution_from_nets_tac (Tactic.orderlist o List.filter (nsubgoalsP n)) true;  | 
| 
3705
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
659  | 
|
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
660  | 
fun eq_contr_tac i = ematch_tac [not_elim] i THEN eq_assume_tac i;  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
661  | 
val eq_assume_contr_tac = eq_assume_tac ORELSE' eq_contr_tac;  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
662  | 
|
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
663  | 
(*Two-way branching is allowed only if one of the branches immediately closes*)  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
664  | 
fun bimatch2_tac netpair i =  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
665  | 
n_bimatch_from_nets_tac 2 netpair i THEN  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
666  | 
(eq_assume_contr_tac i ORELSE eq_assume_contr_tac (i+1));  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
667  | 
|
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
668  | 
(*Attack subgoals using safe inferences -- matching, not resolution*)  | 
| 10736 | 669  | 
fun clarify_step_tac (cs as CS{safe0_netpair,safep_netpair,...}) =
 | 
| 4651 | 670  | 
appSWrappers cs (FIRST' [  | 
| 9938 | 671  | 
eq_assume_contr_tac,  | 
672  | 
bimatch_from_nets_tac safe0_netpair,  | 
|
673  | 
FIRST' hyp_subst_tacs,  | 
|
674  | 
n_bimatch_from_nets_tac 1 safep_netpair,  | 
|
| 
3705
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
675  | 
bimatch2_tac safep_netpair]);  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
676  | 
|
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
677  | 
fun clarify_tac cs = SELECT_GOAL (REPEAT_DETERM (clarify_step_tac cs 1));  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
678  | 
|
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
679  | 
|
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
680  | 
(*** Unsafe steps instantiate variables or lose information ***)  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
681  | 
|
| 4066 | 682  | 
(*Backtracking is allowed among the various these unsafe ways of  | 
683  | 
proving a subgoal. *)  | 
|
| 
747
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
684  | 
fun inst0_step_tac (CS{safe0_netpair,safep_netpair,...}) =
 | 
| 10736 | 685  | 
assume_tac APPEND'  | 
686  | 
contr_tac APPEND'  | 
|
| 
747
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
687  | 
biresolve_from_nets_tac safe0_netpair;  | 
| 
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
688  | 
|
| 4066 | 689  | 
(*These unsafe steps could generate more subgoals.*)  | 
| 
747
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
690  | 
fun instp_step_tac (CS{safep_netpair,...}) =
 | 
| 
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
691  | 
biresolve_from_nets_tac safep_netpair;  | 
| 0 | 692  | 
|
693  | 
(*These steps could instantiate variables and are therefore unsafe.*)  | 
|
| 
747
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
694  | 
fun inst_step_tac cs = inst0_step_tac cs APPEND' instp_step_tac cs;  | 
| 0 | 695  | 
|
| 10736 | 696  | 
fun haz_step_tac (CS{haz_netpair,...}) =
 | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
697  | 
biresolve_from_nets_tac haz_netpair;  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
698  | 
|
| 0 | 699  | 
(*Single step for the prover. FAILS unless it makes progress. *)  | 
| 10736 | 700  | 
fun step_tac cs i = safe_tac cs ORELSE appWrappers cs  | 
| 9938 | 701  | 
(inst_step_tac cs ORELSE' haz_step_tac cs) i;  | 
| 0 | 702  | 
|
703  | 
(*Using a "safe" rule to instantiate variables is unsafe. This tactic  | 
|
704  | 
allows backtracking from "safe" rules to "unsafe" rules here.*)  | 
|
| 10736 | 705  | 
fun slow_step_tac cs i = safe_tac cs ORELSE appWrappers cs  | 
| 9938 | 706  | 
(inst_step_tac cs APPEND' haz_step_tac cs) i;  | 
| 0 | 707  | 
|
| 1800 | 708  | 
(**** The following tactics all fail unless they solve one goal ****)  | 
| 0 | 709  | 
|
710  | 
(*Dumb but fast*)  | 
|
| 
10382
 
1fb807260ff1
atomize: all automated tactics that "solve" goals;
 
wenzelm 
parents: 
10309 
diff
changeset
 | 
711  | 
fun fast_tac cs =  | 
| 11754 | 712  | 
ObjectLogic.atomize_tac THEN' SELECT_GOAL (DEPTH_SOLVE (step_tac cs 1));  | 
| 0 | 713  | 
|
714  | 
(*Slower but smarter than fast_tac*)  | 
|
| 
10382
 
1fb807260ff1
atomize: all automated tactics that "solve" goals;
 
wenzelm 
parents: 
10309 
diff
changeset
 | 
715  | 
fun best_tac cs =  | 
| 11754 | 716  | 
ObjectLogic.atomize_tac THEN'  | 
| 0 | 717  | 
SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, sizef) (step_tac cs 1));  | 
718  | 
||
| 9402 | 719  | 
(*even a bit smarter than best_tac*)  | 
| 
10382
 
1fb807260ff1
atomize: all automated tactics that "solve" goals;
 
wenzelm 
parents: 
10309 
diff
changeset
 | 
720  | 
fun first_best_tac cs =  | 
| 11754 | 721  | 
ObjectLogic.atomize_tac THEN'  | 
| 9402 | 722  | 
SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, sizef) (FIRSTGOAL (step_tac cs)));  | 
723  | 
||
| 
10382
 
1fb807260ff1
atomize: all automated tactics that "solve" goals;
 
wenzelm 
parents: 
10309 
diff
changeset
 | 
724  | 
fun slow_tac cs =  | 
| 11754 | 725  | 
ObjectLogic.atomize_tac THEN'  | 
| 
10382
 
1fb807260ff1
atomize: all automated tactics that "solve" goals;
 
wenzelm 
parents: 
10309 
diff
changeset
 | 
726  | 
SELECT_GOAL (DEPTH_SOLVE (slow_step_tac cs 1));  | 
| 0 | 727  | 
|
| 
10382
 
1fb807260ff1
atomize: all automated tactics that "solve" goals;
 
wenzelm 
parents: 
10309 
diff
changeset
 | 
728  | 
fun slow_best_tac cs =  | 
| 11754 | 729  | 
ObjectLogic.atomize_tac THEN'  | 
| 0 | 730  | 
SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, sizef) (slow_step_tac cs 1));  | 
731  | 
||
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
732  | 
|
| 10736 | 733  | 
(***ASTAR with weight weight_ASTAR, by Norbert Voelker*)  | 
734  | 
val weight_ASTAR = ref 5;  | 
|
| 
1587
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
735  | 
|
| 
10382
 
1fb807260ff1
atomize: all automated tactics that "solve" goals;
 
wenzelm 
parents: 
10309 
diff
changeset
 | 
736  | 
fun astar_tac cs =  | 
| 11754 | 737  | 
ObjectLogic.atomize_tac THEN'  | 
| 
10382
 
1fb807260ff1
atomize: all automated tactics that "solve" goals;
 
wenzelm 
parents: 
10309 
diff
changeset
 | 
738  | 
SELECT_GOAL  | 
| 
 
1fb807260ff1
atomize: all automated tactics that "solve" goals;
 
wenzelm 
parents: 
10309 
diff
changeset
 | 
739  | 
(ASTAR (has_fewer_prems 1, fn lev => fn thm => size_of_thm thm + !weight_ASTAR * lev)  | 
| 
 
1fb807260ff1
atomize: all automated tactics that "solve" goals;
 
wenzelm 
parents: 
10309 
diff
changeset
 | 
740  | 
(step_tac cs 1));  | 
| 
1587
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
741  | 
|
| 10736 | 742  | 
fun slow_astar_tac cs =  | 
| 11754 | 743  | 
ObjectLogic.atomize_tac THEN'  | 
| 
10382
 
1fb807260ff1
atomize: all automated tactics that "solve" goals;
 
wenzelm 
parents: 
10309 
diff
changeset
 | 
744  | 
SELECT_GOAL  | 
| 
 
1fb807260ff1
atomize: all automated tactics that "solve" goals;
 
wenzelm 
parents: 
10309 
diff
changeset
 | 
745  | 
(ASTAR (has_fewer_prems 1, fn lev => fn thm => size_of_thm thm + !weight_ASTAR * lev)  | 
| 
 
1fb807260ff1
atomize: all automated tactics that "solve" goals;
 
wenzelm 
parents: 
10309 
diff
changeset
 | 
746  | 
(slow_step_tac cs 1));  | 
| 
1587
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
747  | 
|
| 1800 | 748  | 
(**** Complete tactic, loosely based upon LeanTaP. This tactic is the outcome  | 
| 
747
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
749  | 
of much experimentation! Changing APPEND to ORELSE below would prove  | 
| 
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
750  | 
easy theorems faster, but loses completeness -- and many of the harder  | 
| 1800 | 751  | 
theorems such as 43. ****)  | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
752  | 
|
| 
747
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
753  | 
(*Non-deterministic! Could always expand the first unsafe connective.  | 
| 
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
754  | 
That's hard to implement and did not perform better in experiments, due to  | 
| 
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
755  | 
greater search depth required.*)  | 
| 10736 | 756  | 
fun dup_step_tac (cs as (CS{dup_netpair,...})) =
 | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
757  | 
biresolve_from_nets_tac dup_netpair;  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
758  | 
|
| 5523 | 759  | 
(*Searching to depth m. A variant called nodup_depth_tac appears in clasimp.ML*)  | 
| 5757 | 760  | 
local  | 
| 10736 | 761  | 
fun slow_step_tac' cs = appWrappers cs  | 
| 9938 | 762  | 
(instp_step_tac cs APPEND' dup_step_tac cs);  | 
| 10736 | 763  | 
in fun depth_tac cs m i state = SELECT_GOAL  | 
764  | 
(safe_steps_tac cs 1 THEN_ELSE  | 
|
| 9938 | 765  | 
(DEPTH_SOLVE (depth_tac cs m 1),  | 
766  | 
inst0_step_tac cs 1 APPEND COND (K (m=0)) no_tac  | 
|
767  | 
(slow_step_tac' cs 1 THEN DEPTH_SOLVE (depth_tac cs (m-1) 1))  | 
|
| 5757 | 768  | 
)) i state;  | 
769  | 
end;  | 
|
| 
747
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
770  | 
|
| 10736 | 771  | 
(*Search, with depth bound m.  | 
| 2173 | 772  | 
This is the "entry point", which does safe inferences first.*)  | 
| 10736 | 773  | 
fun safe_depth_tac cs m =  | 
774  | 
SUBGOAL  | 
|
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
775  | 
(fn (prem,i) =>  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
776  | 
let val deti =  | 
| 9938 | 777  | 
(*No Vars in the goal? No need to backtrack between goals.*)  | 
778  | 
case term_vars prem of  | 
|
| 10736 | 779  | 
[] => DETERM  | 
| 9938 | 780  | 
| _::_ => I  | 
| 10736 | 781  | 
in SELECT_GOAL (TRY (safe_tac cs) THEN  | 
| 9938 | 782  | 
DEPTH_SOLVE (deti (depth_tac cs m 1))) i  | 
| 
747
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
783  | 
end);  | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
784  | 
|
| 
2868
 
17a23a62f82a
New DEEPEN allows giving an upper bound for deepen_tac
 
paulson 
parents: 
2813 
diff
changeset
 | 
785  | 
fun deepen_tac cs = DEEPEN (2,10) (safe_depth_tac cs);  | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
786  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
787  | 
|
| 1724 | 788  | 
|
| 15036 | 789  | 
(** context dependent claset components **)  | 
790  | 
||
791  | 
datatype context_cs = ContextCS of  | 
|
792  | 
 {swrappers: (string * (Proof.context -> wrapper)) list,
 | 
|
793  | 
uwrappers: (string * (Proof.context -> wrapper)) list};  | 
|
794  | 
||
795  | 
fun context_cs ctxt cs (ContextCS {swrappers, uwrappers}) =
 | 
|
796  | 
let  | 
|
797  | 
fun add_wrapper add (name, f) claset = add (claset, (name, f ctxt));  | 
|
798  | 
in  | 
|
799  | 
cs |> fold_rev (add_wrapper (op addSWrapper)) swrappers  | 
|
800  | 
|> fold_rev (add_wrapper (op addWrapper)) uwrappers  | 
|
801  | 
end;  | 
|
802  | 
||
803  | 
fun make_context_cs (swrappers, uwrappers) =  | 
|
804  | 
  ContextCS {swrappers = swrappers, uwrappers = uwrappers};
 | 
|
805  | 
||
806  | 
val empty_context_cs = make_context_cs ([], []);  | 
|
807  | 
||
808  | 
fun merge_context_cs (ctxt_cs1, ctxt_cs2) =  | 
|
809  | 
let  | 
|
810  | 
    val ContextCS {swrappers = swrappers1, uwrappers = uwrappers1} = ctxt_cs1;
 | 
|
811  | 
    val ContextCS {swrappers = swrappers2, uwrappers = uwrappers2} = ctxt_cs2;
 | 
|
812  | 
||
813  | 
val swrappers' = merge_alists swrappers1 swrappers2;  | 
|
814  | 
val uwrappers' = merge_alists uwrappers1 uwrappers2;  | 
|
815  | 
in make_context_cs (swrappers', uwrappers') end;  | 
|
816  | 
||
817  | 
||
818  | 
||
| 17880 | 819  | 
(** claset data **)  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
820  | 
|
| 17880 | 821  | 
(* global clasets *)  | 
| 1724 | 822  | 
|
| 16424 | 823  | 
structure GlobalClaset = TheoryDataFun  | 
824  | 
(struct  | 
|
| 7354 | 825  | 
val name = "Provers/claset";  | 
| 15036 | 826  | 
type T = claset ref * context_cs;  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
827  | 
|
| 15036 | 828  | 
val empty = (ref empty_cs, empty_context_cs);  | 
| 16424 | 829  | 
fun copy (ref cs, ctxt_cs) = (ref cs, ctxt_cs): T;  | 
830  | 
val extend = copy;  | 
|
831  | 
fun merge _ ((ref cs1, ctxt_cs1), (ref cs2, ctxt_cs2)) =  | 
|
| 15036 | 832  | 
(ref (merge_cs (cs1, cs2)), merge_context_cs (ctxt_cs1, ctxt_cs2));  | 
833  | 
fun print _ (ref cs, _) = print_cs cs;  | 
|
| 16424 | 834  | 
end);  | 
| 1724 | 835  | 
|
| 7354 | 836  | 
val print_claset = GlobalClaset.print;  | 
| 17880 | 837  | 
val get_claset = ! o #1 o GlobalClaset.get;  | 
838  | 
||
| 15036 | 839  | 
val get_context_cs = #2 o GlobalClaset.get o ProofContext.theory_of;  | 
840  | 
fun map_context_cs f = GlobalClaset.map (apsnd  | 
|
841  | 
  (fn ContextCS {swrappers, uwrappers} => make_context_cs (f (swrappers, uwrappers))));
 | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
842  | 
|
| 17880 | 843  | 
val change_claset_of = change o #1 o GlobalClaset.get;  | 
844  | 
fun change_claset f = change_claset_of (Context.the_context ()) f;  | 
|
| 1800 | 845  | 
|
| 17880 | 846  | 
fun claset_of thy =  | 
847  | 
let val (cs_ref, ctxt_cs) = GlobalClaset.get thy  | 
|
848  | 
in context_cs (Context.init_proof thy) (! cs_ref) (ctxt_cs) end;  | 
|
| 5028 | 849  | 
val claset = claset_of o Context.the_context;  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
850  | 
|
| 17880 | 851  | 
fun CLASET tacf st = tacf (claset_of (Thm.theory_of_thm st)) st;  | 
852  | 
fun CLASET' tacf i st = tacf (claset_of (Thm.theory_of_thm st)) i st;  | 
|
| 1724 | 853  | 
|
| 17880 | 854  | 
fun AddDs args = change_claset (fn cs => cs addDs args);  | 
855  | 
fun AddEs args = change_claset (fn cs => cs addEs args);  | 
|
856  | 
fun AddIs args = change_claset (fn cs => cs addIs args);  | 
|
857  | 
fun AddSDs args = change_claset (fn cs => cs addSDs args);  | 
|
858  | 
fun AddSEs args = change_claset (fn cs => cs addSEs args);  | 
|
859  | 
fun AddSIs args = change_claset (fn cs => cs addSIs args);  | 
|
860  | 
fun Delrules args = change_claset (fn cs => cs delrules args);  | 
|
| 
3727
 
ed63c05d7992
Added Safe_tac; all other default claset tactics now dereference "claset"
 
paulson 
parents: 
3705 
diff
changeset
 | 
861  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
862  | 
|
| 15036 | 863  | 
(* context dependent components *)  | 
864  | 
||
865  | 
fun add_context_safe_wrapper wrapper = map_context_cs (apfst (merge_alists [wrapper]));  | 
|
866  | 
fun del_context_safe_wrapper name = map_context_cs (apfst (filter_out (equal name o #1)));  | 
|
867  | 
||
868  | 
fun add_context_unsafe_wrapper wrapper = map_context_cs (apsnd (merge_alists [wrapper]));  | 
|
869  | 
fun del_context_unsafe_wrapper name = map_context_cs (apsnd (filter_out (equal name o #1)));  | 
|
870  | 
||
871  | 
||
| 5841 | 872  | 
(* proof data kind 'Provers/claset' *)  | 
873  | 
||
| 16424 | 874  | 
structure LocalClaset = ProofDataFun  | 
875  | 
(struct  | 
|
| 5841 | 876  | 
val name = "Provers/claset";  | 
877  | 
type T = claset;  | 
|
| 17880 | 878  | 
val init = get_claset;  | 
| 15036 | 879  | 
fun print ctxt cs = print_cs (context_cs ctxt cs (get_context_cs ctxt));  | 
| 16424 | 880  | 
end);  | 
| 5841 | 881  | 
|
882  | 
val print_local_claset = LocalClaset.print;  | 
|
883  | 
val get_local_claset = LocalClaset.get;  | 
|
884  | 
val put_local_claset = LocalClaset.put;  | 
|
885  | 
||
| 15036 | 886  | 
fun local_claset_of ctxt =  | 
887  | 
context_cs ctxt (get_local_claset ctxt) (get_context_cs ctxt);  | 
|
888  | 
||
| 5841 | 889  | 
|
| 5885 | 890  | 
(* attributes *)  | 
891  | 
||
| 17880 | 892  | 
fun change_global_cs f (thy, th) = (change_claset_of thy (fn cs => f (cs, [th])); (thy, th));  | 
893  | 
fun change_local_cs f (ctxt, th) = (LocalClaset.map (fn cs => f (cs, [th])) ctxt, th);  | 
|
| 5885 | 894  | 
|
895  | 
val safe_dest_global = change_global_cs (op addSDs);  | 
|
896  | 
val safe_elim_global = change_global_cs (op addSEs);  | 
|
897  | 
val safe_intro_global = change_global_cs (op addSIs);  | 
|
| 6955 | 898  | 
val haz_dest_global = change_global_cs (op addDs);  | 
899  | 
val haz_elim_global = change_global_cs (op addEs);  | 
|
900  | 
val haz_intro_global = change_global_cs (op addIs);  | 
|
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
901  | 
val rule_del_global = change_global_cs (op delrules) o ContextRules.rule_del_global;  | 
| 5885 | 902  | 
|
| 16806 | 903  | 
val safe_dest_local = change_local_cs (op addSDs);  | 
904  | 
val safe_elim_local = change_local_cs (op addSEs);  | 
|
905  | 
val safe_intro_local = change_local_cs (op addSIs);  | 
|
906  | 
val haz_dest_local = change_local_cs (op addDs);  | 
|
907  | 
val haz_elim_local = change_local_cs (op addEs);  | 
|
908  | 
val haz_intro_local = change_local_cs (op addIs);  | 
|
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
909  | 
val rule_del_local = change_local_cs (op delrules) o ContextRules.rule_del_local;  | 
| 5885 | 910  | 
|
911  | 
||
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
912  | 
(* tactics referring to the implicit claset *)  | 
| 1800 | 913  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
914  | 
(*the abstraction over the proof state delays the dereferencing*)  | 
| 9938 | 915  | 
fun Safe_tac st = safe_tac (claset()) st;  | 
916  | 
fun Safe_step_tac i st = safe_step_tac (claset()) i st;  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
917  | 
fun Clarify_step_tac i st = clarify_step_tac (claset()) i st;  | 
| 9938 | 918  | 
fun Clarify_tac i st = clarify_tac (claset()) i st;  | 
919  | 
fun Step_tac i st = step_tac (claset()) i st;  | 
|
920  | 
fun Fast_tac i st = fast_tac (claset()) i st;  | 
|
921  | 
fun Best_tac i st = best_tac (claset()) i st;  | 
|
922  | 
fun Slow_tac i st = slow_tac (claset()) i st;  | 
|
923  | 
fun Slow_best_tac i st = slow_best_tac (claset()) i st;  | 
|
924  | 
fun Deepen_tac m = deepen_tac (claset()) m;  | 
|
| 2066 | 925  | 
|
| 1800 | 926  | 
|
| 10736 | 927  | 
end;  | 
| 5841 | 928  | 
|
929  | 
||
930  | 
||
| 5885 | 931  | 
(** concrete syntax of attributes **)  | 
| 5841 | 932  | 
|
933  | 
(* add / del rules *)  | 
|
934  | 
||
935  | 
val introN = "intro";  | 
|
936  | 
val elimN = "elim";  | 
|
937  | 
val destN = "dest";  | 
|
| 9938 | 938  | 
val ruleN = "rule";  | 
| 5841 | 939  | 
|
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
940  | 
fun add_rule xtra haz safe = Attrib.syntax  | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
941  | 
(Scan.lift (Args.query |-- Scan.option Args.nat >> xtra || Args.bang >> K safe ||  | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
942  | 
Scan.succeed haz));  | 
| 5841 | 943  | 
|
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
944  | 
fun del_rule att = Attrib.syntax (Scan.lift Args.del >> K att);  | 
| 5841 | 945  | 
|
946  | 
||
947  | 
(* setup_attrs *)  | 
|
948  | 
||
| 
9941
 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 
wenzelm 
parents: 
9938 
diff
changeset
 | 
949  | 
fun elim_format x = Attrib.no_args (Drule.rule_attribute (K Data.make_elim)) x;  | 
| 9184 | 950  | 
|
| 5841 | 951  | 
val setup_attrs = Attrib.add_attributes  | 
| 
9941
 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 
wenzelm 
parents: 
9938 
diff
changeset
 | 
952  | 
 [("elim_format", (elim_format, elim_format),
 | 
| 
 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 
wenzelm 
parents: 
9938 
diff
changeset
 | 
953  | 
"destruct rule turned into elimination rule format (classical)"),  | 
| 12401 | 954  | 
  ("swapped", (swapped, swapped), "classical swap of introduction rule"),
 | 
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
955  | 
(destN,  | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
956  | 
(add_rule ContextRules.dest_query_global haz_dest_global safe_dest_global,  | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
957  | 
add_rule ContextRules.dest_query_local haz_dest_local safe_dest_local),  | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
958  | 
"declaration of destruction rule"),  | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
959  | 
(elimN,  | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
960  | 
(add_rule ContextRules.elim_query_global haz_elim_global safe_elim_global,  | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
961  | 
add_rule ContextRules.elim_query_local haz_elim_local safe_elim_local),  | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
962  | 
"declaration of elimination rule"),  | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
963  | 
(introN,  | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
964  | 
(add_rule ContextRules.intro_query_global haz_intro_global safe_intro_global,  | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
965  | 
add_rule ContextRules.intro_query_local haz_intro_local safe_intro_local),  | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
966  | 
"declaration of introduction rule"),  | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
967  | 
(ruleN, (del_rule rule_del_global, del_rule rule_del_local),  | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
968  | 
"remove declaration of intro/elim/dest rule")];  | 
| 5841 | 969  | 
|
970  | 
||
971  | 
||
| 7230 | 972  | 
(** proof methods **)  | 
973  | 
||
| 
14605
 
9de4d64eee3b
'instance' and intro_classes now handle general sorts;
 
wenzelm 
parents: 
13525 
diff
changeset
 | 
974  | 
fun METHOD_CLASET tac ctxt =  | 
| 15036 | 975  | 
Method.METHOD (tac ctxt (local_claset_of ctxt));  | 
| 5841 | 976  | 
|
| 8098 | 977  | 
fun METHOD_CLASET' tac ctxt =  | 
| 15036 | 978  | 
Method.METHOD (HEADGOAL o tac ctxt (local_claset_of ctxt));  | 
| 7230 | 979  | 
|
980  | 
||
981  | 
local  | 
|
982  | 
||
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
983  | 
fun some_rule_tac ctxt (CS {xtra_netpair, ...}) facts = SUBGOAL (fn (goal, i) =>
 | 
| 5841 | 984  | 
let  | 
| 12401 | 985  | 
val [rules1, rules2, rules4] = ContextRules.find_rules false facts goal ctxt;  | 
986  | 
val rules3 = ContextRules.find_rules_netpair true facts goal xtra_netpair;  | 
|
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
987  | 
val rules = rules1 @ rules2 @ rules3 @ rules4;  | 
| 18223 | 988  | 
val ruleq = Drule.multi_resolves facts rules;  | 
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
989  | 
in  | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
990  | 
Method.trace ctxt rules;  | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
991  | 
fn st => Seq.flat (Seq.map (fn rule => Tactic.rtac rule i st) ruleq)  | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
992  | 
end);  | 
| 5841 | 993  | 
|
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
994  | 
fun rule_tac [] ctxt cs facts = some_rule_tac ctxt cs facts  | 
| 10394 | 995  | 
| rule_tac rules _ _ facts = Method.rule_tac rules facts;  | 
| 7281 | 996  | 
|
| 
10382
 
1fb807260ff1
atomize: all automated tactics that "solve" goals;
 
wenzelm 
parents: 
10309 
diff
changeset
 | 
997  | 
fun default_tac rules ctxt cs facts =  | 
| 
14605
 
9de4d64eee3b
'instance' and intro_classes now handle general sorts;
 
wenzelm 
parents: 
13525 
diff
changeset
 | 
998  | 
HEADGOAL (rule_tac rules ctxt cs facts) ORELSE  | 
| 
10382
 
1fb807260ff1
atomize: all automated tactics that "solve" goals;
 
wenzelm 
parents: 
10309 
diff
changeset
 | 
999  | 
AxClass.default_intro_classes_tac facts;  | 
| 10309 | 1000  | 
|
| 7230 | 1001  | 
in  | 
| 7281 | 1002  | 
val rule = METHOD_CLASET' o rule_tac;  | 
| 
14605
 
9de4d64eee3b
'instance' and intro_classes now handle general sorts;
 
wenzelm 
parents: 
13525 
diff
changeset
 | 
1003  | 
val default = METHOD_CLASET o default_tac;  | 
| 7230 | 1004  | 
end;  | 
| 5841 | 1005  | 
|
1006  | 
||
| 7230 | 1007  | 
(* contradiction method *)  | 
| 6502 | 1008  | 
|
| 7425 | 1009  | 
val contradiction = Method.rule [Data.not_elim, Data.not_elim COMP Drule.swap_prems_rl];  | 
| 6502 | 1010  | 
|
1011  | 
||
1012  | 
(* automatic methods *)  | 
|
| 5841 | 1013  | 
|
| 5927 | 1014  | 
val cla_modifiers =  | 
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
1015  | 
[Args.$$$ destN -- Args.bang_colon >> K ((I, safe_dest_local): Method.modifier),  | 
| 10034 | 1016  | 
Args.$$$ destN -- Args.colon >> K (I, haz_dest_local),  | 
1017  | 
Args.$$$ elimN -- Args.bang_colon >> K (I, safe_elim_local),  | 
|
1018  | 
Args.$$$ elimN -- Args.colon >> K (I, haz_elim_local),  | 
|
1019  | 
Args.$$$ introN -- Args.bang_colon >> K (I, safe_intro_local),  | 
|
1020  | 
Args.$$$ introN -- Args.colon >> K (I, haz_intro_local),  | 
|
1021  | 
Args.del -- Args.colon >> K (I, rule_del_local)];  | 
|
| 5927 | 1022  | 
|
| 7559 | 1023  | 
fun cla_meth tac prems ctxt = Method.METHOD (fn facts =>  | 
| 15036 | 1024  | 
ALLGOALS (Method.insert_tac (prems @ facts)) THEN tac (local_claset_of ctxt));  | 
| 7132 | 1025  | 
|
| 7559 | 1026  | 
fun cla_meth' tac prems ctxt = Method.METHOD (fn facts =>  | 
| 15036 | 1027  | 
HEADGOAL (Method.insert_tac (prems @ facts) THEN' tac (local_claset_of ctxt)));  | 
| 5841 | 1028  | 
|
| 7559 | 1029  | 
val cla_method = Method.bang_sectioned_args cla_modifiers o cla_meth;  | 
1030  | 
val cla_method' = Method.bang_sectioned_args cla_modifiers o cla_meth';  | 
|
| 5841 | 1031  | 
|
1032  | 
||
1033  | 
||
1034  | 
(** setup_methods **)  | 
|
1035  | 
||
1036  | 
val setup_methods = Method.add_methods  | 
|
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
1037  | 
 [("default", Method.thms_ctxt_args default, "apply some intro/elim rule (potentially classical)"),
 | 
| 
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
1038  | 
  ("rule", Method.thms_ctxt_args rule, "apply some intro/elim rule (potentially classical)"),
 | 
| 6502 | 1039  | 
  ("contradiction", Method.no_args contradiction, "proof by contradiction"),
 | 
| 10821 | 1040  | 
  ("clarify", cla_method' (CHANGED_PROP oo clarify_tac), "repeatedly apply safe steps"),
 | 
| 7004 | 1041  | 
  ("fast", cla_method' fast_tac, "classical prover (depth-first)"),
 | 
| 9806 | 1042  | 
  ("slow", cla_method' slow_tac, "classical prover (slow depth-first)"),
 | 
| 9773 | 1043  | 
  ("best", cla_method' best_tac, "classical prover (best-first)"),
 | 
| 18015 | 1044  | 
  ("deepen", cla_method' (fn cs => deepen_tac cs 4), "classical prover (iterative deepening)"),
 | 
| 10821 | 1045  | 
  ("safe", cla_method (CHANGED_PROP o safe_tac), "classical prover (apply safe rules)")];
 | 
| 5841 | 1046  | 
|
1047  | 
||
1048  | 
||
1049  | 
(** theory setup **)  | 
|
1050  | 
||
| 16806 | 1051  | 
val setup = [GlobalClaset.init, LocalClaset.init, setup_attrs, setup_methods];  | 
| 5841 | 1052  | 
|
1053  | 
||
| 8667 | 1054  | 
|
1055  | 
(** outer syntax **)  | 
|
1056  | 
||
1057  | 
val print_clasetP =  | 
|
1058  | 
OuterSyntax.improper_command "print_claset" "print context of Classical Reasoner"  | 
|
| 17057 | 1059  | 
OuterKeyword.diag  | 
| 9513 | 1060  | 
(Scan.succeed (Toplevel.no_timing o Toplevel.unknown_context o (Toplevel.keep  | 
| 9010 | 1061  | 
(Toplevel.node_case print_claset (print_local_claset o Proof.context_of)))));  | 
| 8667 | 1062  | 
|
1063  | 
val _ = OuterSyntax.add_parsers [print_clasetP];  | 
|
1064  | 
||
1065  | 
||
| 5841 | 1066  | 
end;  |