| author | wenzelm | 
| Wed, 14 Jul 1999 13:06:08 +0200 | |
| changeset 7004 | c799d0859638 | 
| parent 6967 | a3c163ed1e04 | 
| child 7106 | c47d94f61ced | 
| permissions | -rw-r--r-- | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
1  | 
(* Title: Provers/classical.ML  | 
| 0 | 2  | 
ID: $Id$  | 
3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
4  | 
Copyright 1992 University of Cambridge  | 
|
5  | 
||
6  | 
Theorem prover for classical reasoning, including predicate calculus, set  | 
|
7  | 
theory, etc.  | 
|
8  | 
||
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
9  | 
Rules must be classified as intr, 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*)  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
18  | 
infix 4 addSIs addSEs addSDs addIs addEs addDs delrules  | 
| 4651 | 19  | 
addSWrapper delSWrapper addWrapper delWrapper  | 
| 5523 | 20  | 
addSbefore addSaltern addbefore addaltern  | 
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  | 
|
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
28  | 
signature CLASET_THY_DATA =  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
29  | 
sig  | 
| 5001 | 30  | 
val clasetN: string  | 
31  | 
val clasetK: Object.kind  | 
|
32  | 
exception ClasetData of Object.T ref  | 
|
| 4854 | 33  | 
val setup: (theory -> theory) list  | 
| 6556 | 34  | 
val fix_methods: Object.T * (Object.T -> Object.T) * (Object.T -> Object.T) *  | 
| 5001 | 35  | 
(Object.T * Object.T -> Object.T) * (Sign.sg -> Object.T -> unit) -> unit  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
36  | 
end;  | 
| 
2868
 
17a23a62f82a
New DEEPEN allows giving an upper bound for deepen_tac
 
paulson 
parents: 
2813 
diff
changeset
 | 
37  | 
|
| 0 | 38  | 
signature CLASSICAL_DATA =  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
39  | 
sig  | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
40  | 
val mp : thm (* [| P-->Q; P |] ==> Q *)  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
41  | 
val not_elim : thm (* [| ~P; P |] ==> R *)  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
42  | 
val classical : thm (* (~P ==> P) ==> P *)  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
43  | 
val sizef : thm -> int (* size function for BEST_FIRST *)  | 
| 0 | 44  | 
val hyp_subst_tacs: (int -> tactic) list  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
45  | 
end;  | 
| 0 | 46  | 
|
| 5841 | 47  | 
signature BASIC_CLASSICAL =  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
48  | 
sig  | 
| 0 | 49  | 
type claset  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
50  | 
val empty_cs: claset  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
51  | 
val print_cs: claset -> unit  | 
| 4380 | 52  | 
val print_claset: theory -> unit  | 
| 4653 | 53  | 
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
 | 
54  | 
    claset -> {safeIs: thm list, safeEs: thm list,
 | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
55  | 
hazIs: thm list, hazEs: thm list,  | 
| 6955 | 56  | 
xtraIs: thm list, xtraEs: thm list,  | 
| 4651 | 57  | 
swrappers: (string * wrapper) list,  | 
58  | 
uwrappers: (string * wrapper) list,  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
59  | 
safe0_netpair: netpair, safep_netpair: netpair,  | 
| 6955 | 60  | 
haz_netpair: netpair, dup_netpair: netpair, xtra_netpair: netpair}  | 
| 1711 | 61  | 
val merge_cs : claset * claset -> claset  | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
62  | 
val addDs : claset * thm list -> claset  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
63  | 
val addEs : claset * thm list -> claset  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
64  | 
val addIs : claset * thm list -> claset  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
65  | 
val addSDs : claset * thm list -> claset  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
66  | 
val addSEs : claset * thm list -> claset  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
67  | 
val addSIs : claset * thm list -> claset  | 
| 1800 | 68  | 
val delrules : claset * thm list -> claset  | 
| 4651 | 69  | 
val addSWrapper : claset * (string * wrapper) -> claset  | 
70  | 
val delSWrapper : claset * string -> claset  | 
|
71  | 
val addWrapper : claset * (string * wrapper) -> claset  | 
|
72  | 
val delWrapper : claset * string -> claset  | 
|
73  | 
val addSbefore : claset * (string * (int -> tactic)) -> claset  | 
|
74  | 
val addSaltern : claset * (string * (int -> tactic)) -> claset  | 
|
75  | 
val addbefore : claset * (string * (int -> tactic)) -> claset  | 
|
76  | 
val addaltern : claset * (string * (int -> tactic)) -> claset  | 
|
| 5523 | 77  | 
val addD2 : claset * (string * thm) -> claset  | 
78  | 
val addE2 : claset * (string * thm) -> claset  | 
|
79  | 
val addSD2 : claset * (string * thm) -> claset  | 
|
80  | 
val addSE2 : claset * (string * thm) -> claset  | 
|
| 4765 | 81  | 
val appSWrappers : claset -> wrapper  | 
| 4651 | 82  | 
val appWrappers : claset -> wrapper  | 
| 5927 | 83  | 
val trace_rules : bool ref  | 
| 
982
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
84  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
85  | 
val claset_ref_of_sg: Sign.sg -> claset ref  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
86  | 
val claset_ref_of: theory -> claset ref  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
87  | 
val claset_of_sg: Sign.sg -> claset  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
88  | 
val claset_of: theory -> claset  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
89  | 
val CLASET: (claset -> tactic) -> tactic  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
90  | 
val CLASET': (claset -> 'a -> tactic) -> 'a -> tactic  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
91  | 
val claset: unit -> claset  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
92  | 
val claset_ref: unit -> claset ref  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
93  | 
|
| 
1587
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
94  | 
val fast_tac : claset -> int -> tactic  | 
| 
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
95  | 
val slow_tac : claset -> int -> tactic  | 
| 
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
96  | 
val weight_ASTAR : int ref  | 
| 
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
97  | 
val astar_tac : claset -> int -> tactic  | 
| 
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
98  | 
val slow_astar_tac : claset -> int -> tactic  | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
99  | 
val best_tac : claset -> int -> tactic  | 
| 
1587
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
100  | 
val slow_best_tac : claset -> int -> tactic  | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
101  | 
val depth_tac : claset -> int -> int -> tactic  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
102  | 
val deepen_tac : claset -> int -> int -> tactic  | 
| 
1587
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
103  | 
|
| 
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
104  | 
val contr_tac : int -> tactic  | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
105  | 
val dup_elim : thm -> thm  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
106  | 
val dup_intr : thm -> thm  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
107  | 
val dup_step_tac : claset -> int -> tactic  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
108  | 
val eq_mp_tac : int -> tactic  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
109  | 
val haz_step_tac : claset -> int -> tactic  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
110  | 
val joinrules : thm list * thm list -> (bool * thm) list  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
111  | 
val mp_tac : int -> tactic  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
112  | 
val safe_tac : claset -> tactic  | 
| 5757 | 113  | 
val safe_steps_tac : claset -> int -> tactic  | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
114  | 
val safe_step_tac : claset -> int -> tactic  | 
| 
3705
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
115  | 
val clarify_tac : claset -> int -> tactic  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
116  | 
val clarify_step_tac : claset -> int -> tactic  | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
117  | 
val step_tac : claset -> int -> tactic  | 
| 2630 | 118  | 
val slow_step_tac : claset -> int -> tactic  | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
119  | 
val swap : thm (* ~P ==> (~Q ==> P) ==> Q *)  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
120  | 
val swapify : thm list -> thm list  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
121  | 
val swap_res_tac : thm list -> int -> tactic  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
122  | 
val inst_step_tac : claset -> int -> tactic  | 
| 
747
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
123  | 
val inst0_step_tac : claset -> int -> tactic  | 
| 
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
124  | 
val instp_step_tac : claset -> int -> tactic  | 
| 1724 | 125  | 
|
126  | 
val AddDs : thm list -> unit  | 
|
127  | 
val AddEs : thm list -> unit  | 
|
128  | 
val AddIs : thm list -> unit  | 
|
129  | 
val AddSDs : thm list -> unit  | 
|
130  | 
val AddSEs : thm list -> unit  | 
|
131  | 
val AddSIs : thm list -> unit  | 
|
| 6955 | 132  | 
val AddXDs : thm list -> unit  | 
133  | 
val AddXEs : thm list -> unit  | 
|
134  | 
val AddXIs : thm list -> unit  | 
|
| 1807 | 135  | 
val Delrules : thm list -> unit  | 
| 
3727
 
ed63c05d7992
Added Safe_tac; all other default claset tactics now dereference "claset"
 
paulson 
parents: 
3705 
diff
changeset
 | 
136  | 
val Safe_tac : tactic  | 
| 1814 | 137  | 
val Safe_step_tac : int -> tactic  | 
| 
3705
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
138  | 
val Clarify_tac : int -> tactic  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
139  | 
val Clarify_step_tac : int -> tactic  | 
| 1800 | 140  | 
val Step_tac : int -> tactic  | 
| 1724 | 141  | 
val Fast_tac : int -> tactic  | 
| 1800 | 142  | 
val Best_tac : int -> tactic  | 
| 2066 | 143  | 
val Slow_tac : int -> tactic  | 
144  | 
val Slow_best_tac : int -> tactic  | 
|
| 1800 | 145  | 
val Deepen_tac : int -> int -> tactic  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
146  | 
end;  | 
| 1724 | 147  | 
|
| 5841 | 148  | 
signature CLASSICAL =  | 
149  | 
sig  | 
|
150  | 
include BASIC_CLASSICAL  | 
|
151  | 
val print_local_claset: Proof.context -> unit  | 
|
152  | 
val get_local_claset: Proof.context -> claset  | 
|
153  | 
val put_local_claset: claset -> Proof.context -> Proof.context  | 
|
154  | 
val safe_dest_global: theory attribute  | 
|
155  | 
val safe_elim_global: theory attribute  | 
|
156  | 
val safe_intro_global: theory attribute  | 
|
| 6955 | 157  | 
val haz_dest_global: theory attribute  | 
158  | 
val haz_elim_global: theory attribute  | 
|
159  | 
val haz_intro_global: theory attribute  | 
|
160  | 
val xtra_dest_global: theory attribute  | 
|
161  | 
val xtra_elim_global: theory attribute  | 
|
162  | 
val xtra_intro_global: theory attribute  | 
|
| 5885 | 163  | 
val delrule_global: theory attribute  | 
| 6955 | 164  | 
val safe_dest_local: Proof.context attribute  | 
165  | 
val safe_elim_local: Proof.context attribute  | 
|
166  | 
val safe_intro_local: Proof.context attribute  | 
|
| 5885 | 167  | 
val haz_dest_local: Proof.context attribute  | 
168  | 
val haz_elim_local: Proof.context attribute  | 
|
169  | 
val haz_intro_local: Proof.context attribute  | 
|
| 6955 | 170  | 
val xtra_dest_local: Proof.context attribute  | 
171  | 
val xtra_elim_local: Proof.context attribute  | 
|
172  | 
val xtra_intro_local: Proof.context attribute  | 
|
| 5885 | 173  | 
val delrule_local: Proof.context attribute  | 
| 5927 | 174  | 
val cla_modifiers: (Args.T list -> (Proof.context attribute * Args.T list)) list  | 
175  | 
val cla_method: (claset -> tactic) -> Args.src -> Proof.context -> Proof.method  | 
|
176  | 
val cla_method': (claset -> int -> tactic) -> Args.src -> Proof.context -> Proof.method  | 
|
| 6096 | 177  | 
val single_tac: claset -> thm list -> int -> tactic  | 
| 5841 | 178  | 
val setup: (theory -> theory) list  | 
179  | 
end;  | 
|
180  | 
||
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
181  | 
structure ClasetThyData: CLASET_THY_DATA =  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
182  | 
struct  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
183  | 
|
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
184  | 
(* data kind claset -- forward declaration *)  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
185  | 
|
| 5001 | 186  | 
val clasetN = "Provers/claset";  | 
187  | 
val clasetK = Object.kind clasetN;  | 
|
188  | 
exception ClasetData of Object.T ref;  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
189  | 
|
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
190  | 
local  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
191  | 
fun undef _ = raise Match;  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
192  | 
|
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
193  | 
val empty_ref = ref ERROR;  | 
| 6556 | 194  | 
val copy_fn = ref (undef: Object.T -> Object.T);  | 
| 5001 | 195  | 
val prep_ext_fn = ref (undef: Object.T -> Object.T);  | 
196  | 
val merge_fn = ref (undef: Object.T * Object.T -> Object.T);  | 
|
197  | 
val print_fn = ref (undef: Sign.sg -> Object.T -> unit);  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
198  | 
|
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
199  | 
val empty = ClasetData empty_ref;  | 
| 6556 | 200  | 
fun copy exn = ! copy_fn exn;  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
201  | 
fun prep_ext exn = ! prep_ext_fn exn;  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
202  | 
fun merge exn = ! merge_fn exn;  | 
| 4259 | 203  | 
fun print sg exn = ! print_fn sg exn;  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
204  | 
in  | 
| 6556 | 205  | 
val setup = [Theory.init_data clasetK (empty, copy, prep_ext, merge, print)];  | 
206  | 
fun fix_methods (e, cp, ext, mrg, prt) =  | 
|
207  | 
(empty_ref := e; copy_fn := cp; prep_ext_fn := ext; merge_fn := mrg; print_fn := prt);  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
208  | 
end;  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
209  | 
|
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
210  | 
|
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
211  | 
end;  | 
| 0 | 212  | 
|
213  | 
||
| 5927 | 214  | 
functor ClassicalFun(Data: CLASSICAL_DATA): CLASSICAL =  | 
| 0 | 215  | 
struct  | 
216  | 
||
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
217  | 
local open ClasetThyData Data in  | 
| 0 | 218  | 
|
| 1800 | 219  | 
(*** Useful tactics for classical reasoning ***)  | 
| 0 | 220  | 
|
| 1524 | 221  | 
val imp_elim = (*cannot use bind_thm within a structure!*)  | 
222  | 
  store_thm ("imp_elim", make_elim mp);
 | 
|
| 0 | 223  | 
|
| 4392 | 224  | 
(*Prove goal that assumes both P and ~P.  | 
225  | 
No backtracking if it finds an equal assumption. Perhaps should call  | 
|
226  | 
ematch_tac instead of eresolve_tac, but then cannot prove ZF/cantor.*)  | 
|
227  | 
val contr_tac = eresolve_tac [not_elim] THEN'  | 
|
228  | 
(eq_assume_tac ORELSE' assume_tac);  | 
|
| 0 | 229  | 
|
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
230  | 
(*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
 | 
231  | 
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
 | 
232  | 
fun mp_tac i = eresolve_tac [not_elim, imp_elim] i THEN assume_tac i;  | 
| 0 | 233  | 
|
234  | 
(*Like mp_tac but instantiates no variables*)  | 
|
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
235  | 
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
 | 
236  | 
|
| 1524 | 237  | 
val swap =  | 
238  | 
  store_thm ("swap", rule_by_tactic (etac thin_rl 1) (not_elim RS classical));
 | 
|
| 0 | 239  | 
|
240  | 
(*Creates rules to eliminate ~A, from rules to introduce A*)  | 
|
241  | 
fun swapify intrs = intrs RLN (2, [swap]);  | 
|
242  | 
||
243  | 
(*Uses introduction rules in the normal way, or on negated assumptions,  | 
|
244  | 
trying rules in order. *)  | 
|
245  | 
fun swap_res_tac rls =  | 
|
| 54 | 246  | 
let fun addrl (rl,brls) = (false, rl) :: (true, rl RSN (2,swap)) :: brls  | 
247  | 
in assume_tac ORELSE'  | 
|
248  | 
contr_tac ORELSE'  | 
|
249  | 
biresolve_tac (foldr addrl (rls,[]))  | 
|
| 0 | 250  | 
end;  | 
251  | 
||
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
252  | 
(*Duplication of hazardous rules, for complete provers*)  | 
| 
2689
 
6d3d893453de
dup_intr & dup_elim no longer call standard -- this
 
paulson 
parents: 
2630 
diff
changeset
 | 
253  | 
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
 | 
254  | 
|
| 6967 | 255  | 
fun dup_elim th =  | 
256  | 
(case try  | 
|
257  | 
(rule_by_tactic (TRYALL (etac revcut_rl)))  | 
|
258  | 
(th RSN (2, revcut_rl) |> assumption 2 |> Seq.hd) of  | 
|
259  | 
Some th' => th'  | 
|
260  | 
  | _ => error ("Bad format for elimination rule\n" ^ string_of_thm th));
 | 
|
| 0 | 261  | 
|
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
262  | 
|
| 1800 | 263  | 
(**** Classical rule sets ****)  | 
| 0 | 264  | 
|
265  | 
datatype claset =  | 
|
| 
982
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
266  | 
  CS of {safeIs		: thm list,		(*safe introduction rules*)
 | 
| 
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
267  | 
safeEs : thm list, (*safe elimination rules*)  | 
| 
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
268  | 
hazIs : thm list, (*unsafe introduction rules*)  | 
| 
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
269  | 
hazEs : thm list, (*unsafe elimination rules*)  | 
| 6955 | 270  | 
xtraIs : thm list, (*extra introduction rules*)  | 
271  | 
xtraEs : thm list, (*extra elimination rules*)  | 
|
| 4651 | 272  | 
swrappers : (string * wrapper) list, (*for transf. safe_step_tac*)  | 
273  | 
uwrappers : (string * wrapper) list, (*for transforming step_tac*)  | 
|
| 
982
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
274  | 
safe0_netpair : netpair, (*nets for trivial cases*)  | 
| 
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
275  | 
safep_netpair : netpair, (*nets for >0 subgoals*)  | 
| 
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
276  | 
haz_netpair : netpair, (*nets for unsafe rules*)  | 
| 6955 | 277  | 
dup_netpair : netpair, (*nets for duplication*)  | 
278  | 
xtra_netpair : netpair}; (*nets for extra rules*)  | 
|
| 0 | 279  | 
|
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
280  | 
(*Desired invariants are  | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
281  | 
safe0_netpair = build safe0_brls,  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
282  | 
safep_netpair = build safep_brls,  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
283  | 
haz_netpair = build (joinrules(hazIs, hazEs)),  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
284  | 
dup_netpair = build (joinrules(map dup_intr hazIs,  | 
| 6955 | 285  | 
map dup_elim hazEs)),  | 
286  | 
xtra_netpair = build (joinrules(xtraIs, xtraEs))}  | 
|
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
287  | 
|
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
288  | 
where build = build_netpair(Net.empty,Net.empty),  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
289  | 
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
 | 
290  | 
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
 | 
291  | 
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
 | 
292  | 
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
 | 
293  | 
*)  | 
| 0 | 294  | 
|
| 6502 | 295  | 
val empty_netpair = (Net.empty, Net.empty);  | 
296  | 
||
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
297  | 
val empty_cs =  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
298  | 
  CS{safeIs	= [],
 | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
299  | 
safeEs = [],  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
300  | 
hazIs = [],  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
301  | 
hazEs = [],  | 
| 6955 | 302  | 
xtraIs = [],  | 
303  | 
xtraEs = [],  | 
|
| 4651 | 304  | 
swrappers = [],  | 
305  | 
uwrappers = [],  | 
|
| 6502 | 306  | 
safe0_netpair = empty_netpair,  | 
307  | 
safep_netpair = empty_netpair,  | 
|
308  | 
haz_netpair = empty_netpair,  | 
|
| 6955 | 309  | 
dup_netpair = empty_netpair,  | 
310  | 
xtra_netpair = empty_netpair};  | 
|
| 0 | 311  | 
|
| 6955 | 312  | 
fun print_cs (CS {safeIs, safeEs, hazIs, hazEs, xtraIs, xtraEs, ...}) =
 | 
| 3546 | 313  | 
let val pretty_thms = map Display.pretty_thm in  | 
314  | 
Pretty.writeln (Pretty.big_list "safe introduction rules:" (pretty_thms safeIs));  | 
|
| 4624 | 315  | 
Pretty.writeln (Pretty.big_list "unsafe introduction rules:" (pretty_thms hazIs));  | 
| 6955 | 316  | 
Pretty.writeln (Pretty.big_list "extra introduction rules:" (pretty_thms xtraIs));  | 
| 4625 | 317  | 
Pretty.writeln (Pretty.big_list "safe elimination rules:" (pretty_thms safeEs));  | 
| 6955 | 318  | 
Pretty.writeln (Pretty.big_list "unsafe elimination rules:" (pretty_thms hazEs));  | 
319  | 
Pretty.writeln (Pretty.big_list "extra elimination rules:" (pretty_thms xtraEs))  | 
|
| 3546 | 320  | 
end;  | 
| 0 | 321  | 
|
| 4653 | 322  | 
fun rep_cs (CS args) = args;  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
323  | 
|
| 4651 | 324  | 
local  | 
325  | 
fun calc_wrap l tac = foldr (fn ((name,tacf),w) => tacf w) (l, tac);  | 
|
326  | 
in  | 
|
327  | 
  fun appSWrappers (CS{swrappers,...}) = calc_wrap swrappers;
 | 
|
328  | 
  fun appWrappers  (CS{uwrappers,...}) = calc_wrap uwrappers;
 | 
|
329  | 
end;  | 
|
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
330  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
331  | 
|
| 1800 | 332  | 
(*** Adding (un)safe introduction or elimination rules.  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
333  | 
|
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
334  | 
In case of overlap, new rules are tried BEFORE old ones!!  | 
| 1800 | 335  | 
***)  | 
| 0 | 336  | 
|
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
337  | 
(*For use with biresolve_tac. Combines intr rules with swap to handle negated  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
338  | 
assumptions. Pairs elim rules with true. *)  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
339  | 
fun joinrules (intrs,elims) =  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
340  | 
(map (pair true) (elims @ swapify intrs) @  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
341  | 
map (pair false) intrs);  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
342  | 
|
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
343  | 
(*Priority: prefer rules with fewest subgoals,  | 
| 1231 | 344  | 
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
 | 
345  | 
fun tag_brls k [] = []  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
346  | 
| tag_brls k (brl::brls) =  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
347  | 
(1000000*subgoals_of_brl brl + k, brl) ::  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
348  | 
tag_brls (k+1) brls;  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
349  | 
|
| 1800 | 350  | 
fun insert_tagged_list kbrls netpr = foldr insert_tagged_brl (kbrls, netpr);  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
351  | 
|
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
352  | 
(*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
 | 
353  | 
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
 | 
354  | 
new insertions the lowest priority.*)  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
355  | 
fun insert (nI,nE) = insert_tagged_list o (tag_brls (~(2*nI+nE))) o joinrules;  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
356  | 
|
| 1800 | 357  | 
fun delete_tagged_list brls netpr = foldr delete_tagged_brl (brls, netpr);  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
358  | 
|
| 1800 | 359  | 
val delete = delete_tagged_list o joinrules;  | 
360  | 
||
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
361  | 
val mem_thm = gen_mem eq_thm  | 
| 
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
362  | 
and rem_thm = gen_rem eq_thm;  | 
| 
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
363  | 
|
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
364  | 
(*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
 | 
365  | 
is still allowed.*)  | 
| 6955 | 366  | 
fun warn_dup th (CS{safeIs, safeEs, hazIs, hazEs, xtraIs, xtraEs, ...}) = 
 | 
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
367  | 
if mem_thm (th, safeIs) then  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
368  | 
	 warning ("Rule already in claset as Safe Intr\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
 | 
369  | 
else if mem_thm (th, safeEs) then  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
370  | 
         warning ("Rule already in claset as Safe Elim\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
 | 
371  | 
else if mem_thm (th, hazIs) then  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
372  | 
         warning ("Rule already in claset as unsafe Intr\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
 | 
373  | 
else if mem_thm (th, hazEs) then  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
374  | 
         warning ("Rule already in claset as unsafe Elim\n" ^ string_of_thm th)
 | 
| 6955 | 375  | 
else if mem_thm (th, xtraIs) then  | 
376  | 
         warning ("Rule already in claset as extra Intr\n" ^ string_of_thm th)
 | 
|
377  | 
else if mem_thm (th, xtraEs) then  | 
|
378  | 
         warning ("Rule already in claset as extra Elim\n" ^ string_of_thm th)
 | 
|
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
379  | 
else ();  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
380  | 
|
| 1800 | 381  | 
(*** Safe rules ***)  | 
| 
982
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
382  | 
|
| 6955 | 383  | 
fun addSI (cs as CS{safeIs, safeEs, hazIs, hazEs, xtraIs, xtraEs, swrappers, uwrappers, 
 | 
384  | 
safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair},  | 
|
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
385  | 
th) =  | 
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
386  | 
if mem_thm (th, safeIs) then  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
387  | 
	 (warning ("Ignoring duplicate Safe Intr\n" ^ string_of_thm th);
 | 
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
388  | 
cs)  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
389  | 
else  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
390  | 
let val (safe0_rls, safep_rls) = (*0 subgoals vs 1 or more*)  | 
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
391  | 
partition (fn rl => nprems_of rl=0) [th]  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
392  | 
val nI = length safeIs + 1  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
393  | 
and nE = length safeEs  | 
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
394  | 
in warn_dup th cs;  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
395  | 
     CS{safeIs	= th::safeIs,
 | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
396  | 
safe0_netpair = insert (nI,nE) (safe0_rls, []) safe0_netpair,  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
397  | 
safep_netpair = insert (nI,nE) (safep_rls, []) safep_netpair,  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
398  | 
safeEs = safeEs,  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
399  | 
hazIs = hazIs,  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
400  | 
hazEs = hazEs,  | 
| 6955 | 401  | 
xtraIs = xtraIs,  | 
402  | 
xtraEs = xtraEs,  | 
|
| 4651 | 403  | 
swrappers = swrappers,  | 
404  | 
uwrappers = uwrappers,  | 
|
| 2630 | 405  | 
haz_netpair = haz_netpair,  | 
| 6955 | 406  | 
dup_netpair = dup_netpair,  | 
407  | 
xtra_netpair = xtra_netpair}  | 
|
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
408  | 
end;  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
409  | 
|
| 6955 | 410  | 
fun addSE (cs as CS{safeIs, safeEs, hazIs, hazEs, xtraIs, xtraEs, swrappers, uwrappers, 
 | 
411  | 
safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair},  | 
|
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
412  | 
th) =  | 
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
413  | 
if mem_thm (th, safeEs) then  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
414  | 
	 (warning ("Ignoring duplicate Safe Elim\n" ^ string_of_thm th);
 | 
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
415  | 
cs)  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
416  | 
else  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
417  | 
let val (safe0_rls, safep_rls) = (*0 subgoals vs 1 or more*)  | 
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
418  | 
partition (fn rl => nprems_of rl=1) [th]  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
419  | 
val nI = length safeIs  | 
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
420  | 
and nE = length safeEs + 1  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
421  | 
in warn_dup th cs;  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
422  | 
     CS{safeEs	= th::safeEs,
 | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
423  | 
safe0_netpair = insert (nI,nE) ([], safe0_rls) safe0_netpair,  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
424  | 
safep_netpair = insert (nI,nE) ([], safep_rls) safep_netpair,  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
425  | 
safeIs = safeIs,  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
426  | 
hazIs = hazIs,  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
427  | 
hazEs = hazEs,  | 
| 6955 | 428  | 
xtraIs = xtraIs,  | 
429  | 
xtraEs = xtraEs,  | 
|
| 4651 | 430  | 
swrappers = swrappers,  | 
431  | 
uwrappers = uwrappers,  | 
|
| 2630 | 432  | 
haz_netpair = haz_netpair,  | 
| 6955 | 433  | 
dup_netpair = dup_netpair,  | 
434  | 
xtra_netpair = xtra_netpair}  | 
|
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
435  | 
end;  | 
| 0 | 436  | 
|
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
437  | 
fun rev_foldl f (e, l) = foldl f (e, rev l);  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
438  | 
|
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
439  | 
val op addSIs = rev_foldl addSI;  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
440  | 
val op addSEs = rev_foldl addSE;  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
441  | 
|
| 0 | 442  | 
fun cs addSDs ths = cs addSEs (map make_elim ths);  | 
443  | 
||
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
444  | 
|
| 1800 | 445  | 
(*** Hazardous (unsafe) rules ***)  | 
| 0 | 446  | 
|
| 6955 | 447  | 
fun addI (cs as CS{safeIs, safeEs, hazIs, hazEs, xtraIs, xtraEs, swrappers, uwrappers, 
 | 
448  | 
safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair},  | 
|
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
449  | 
th)=  | 
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
450  | 
if mem_thm (th, hazIs) then  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
451  | 
	 (warning ("Ignoring duplicate unsafe Intr\n" ^ string_of_thm th);
 | 
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
452  | 
cs)  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
453  | 
else  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
454  | 
let val nI = length hazIs + 1  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
455  | 
and nE = length hazEs  | 
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
456  | 
in warn_dup th cs;  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
457  | 
     CS{hazIs	= th::hazIs,
 | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
458  | 
haz_netpair = insert (nI,nE) ([th], []) haz_netpair,  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
459  | 
dup_netpair = insert (nI,nE) (map dup_intr [th], []) dup_netpair,  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
460  | 
safeIs = safeIs,  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
461  | 
safeEs = safeEs,  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
462  | 
hazEs = hazEs,  | 
| 6955 | 463  | 
xtraIs = xtraIs,  | 
464  | 
xtraEs = xtraEs,  | 
|
| 4651 | 465  | 
swrappers = swrappers,  | 
466  | 
uwrappers = uwrappers,  | 
|
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
467  | 
safe0_netpair = safe0_netpair,  | 
| 6955 | 468  | 
safep_netpair = safep_netpair,  | 
469  | 
xtra_netpair = xtra_netpair}  | 
|
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
470  | 
end;  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
471  | 
|
| 6955 | 472  | 
fun addE (cs as CS{safeIs, safeEs, hazIs, hazEs, xtraIs, xtraEs, swrappers, uwrappers, 
 | 
473  | 
safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair},  | 
|
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
474  | 
th) =  | 
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
475  | 
if mem_thm (th, hazEs) then  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
476  | 
	 (warning ("Ignoring duplicate unsafe Elim\n" ^ string_of_thm th);
 | 
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
477  | 
cs)  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
478  | 
else  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
479  | 
let val nI = length hazIs  | 
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
480  | 
and nE = length hazEs + 1  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
481  | 
in warn_dup th cs;  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
482  | 
     CS{hazEs	= th::hazEs,
 | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
483  | 
haz_netpair = insert (nI,nE) ([], [th]) haz_netpair,  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
484  | 
dup_netpair = insert (nI,nE) ([], map dup_elim [th]) dup_netpair,  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
485  | 
safeIs = safeIs,  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
486  | 
safeEs = safeEs,  | 
| 
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
487  | 
hazIs = hazIs,  | 
| 6955 | 488  | 
xtraIs = xtraIs,  | 
489  | 
xtraEs = xtraEs,  | 
|
| 4651 | 490  | 
swrappers = swrappers,  | 
491  | 
uwrappers = uwrappers,  | 
|
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
492  | 
safe0_netpair = safe0_netpair,  | 
| 6955 | 493  | 
safep_netpair = safep_netpair,  | 
494  | 
xtra_netpair = xtra_netpair}  | 
|
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
495  | 
end;  | 
| 0 | 496  | 
|
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
497  | 
val op addIs = rev_foldl addI;  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
498  | 
val op addEs = rev_foldl addE;  | 
| 
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
499  | 
|
| 0 | 500  | 
fun cs addDs ths = cs addEs (map make_elim ths);  | 
501  | 
||
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
502  | 
|
| 6955 | 503  | 
(*** Extra (single step) rules ***)  | 
504  | 
||
505  | 
fun addXI (cs as CS{safeIs, safeEs, hazIs, hazEs, xtraIs, xtraEs, swrappers, uwrappers,
 | 
|
506  | 
safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair},  | 
|
507  | 
th)=  | 
|
508  | 
if mem_thm (th, xtraIs) then  | 
|
509  | 
	 (warning ("Ignoring duplicate extra Intr\n" ^ string_of_thm th);
 | 
|
510  | 
cs)  | 
|
511  | 
else  | 
|
512  | 
let val nI = length xtraIs + 1  | 
|
513  | 
and nE = length xtraEs  | 
|
514  | 
in warn_dup th cs;  | 
|
515  | 
     CS{xtraIs	= th::xtraIs,
 | 
|
516  | 
xtra_netpair = insert (nI,nE) ([th], []) xtra_netpair,  | 
|
517  | 
safeIs = safeIs,  | 
|
518  | 
safeEs = safeEs,  | 
|
519  | 
hazIs = hazIs,  | 
|
520  | 
hazEs = hazEs,  | 
|
521  | 
xtraEs = xtraEs,  | 
|
522  | 
swrappers = swrappers,  | 
|
523  | 
uwrappers = uwrappers,  | 
|
524  | 
safe0_netpair = safe0_netpair,  | 
|
525  | 
safep_netpair = safep_netpair,  | 
|
526  | 
haz_netpair = haz_netpair,  | 
|
527  | 
dup_netpair = dup_netpair}  | 
|
528  | 
end;  | 
|
529  | 
||
530  | 
fun addXE (cs as CS{safeIs, safeEs, hazIs, hazEs, xtraIs, xtraEs, swrappers, uwrappers, 
 | 
|
531  | 
safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair},  | 
|
532  | 
th) =  | 
|
533  | 
if mem_thm (th, xtraEs) then  | 
|
534  | 
	 (warning ("Ignoring duplicate extra Elim\n" ^ string_of_thm th);
 | 
|
535  | 
cs)  | 
|
536  | 
else  | 
|
537  | 
let val nI = length xtraIs  | 
|
538  | 
and nE = length xtraEs + 1  | 
|
539  | 
in warn_dup th cs;  | 
|
540  | 
     CS{xtraEs	= th::xtraEs,
 | 
|
541  | 
xtra_netpair = insert (nI,nE) ([], [th]) xtra_netpair,  | 
|
542  | 
safeIs = safeIs,  | 
|
543  | 
safeEs = safeEs,  | 
|
544  | 
hazIs = hazIs,  | 
|
545  | 
hazEs = hazEs,  | 
|
546  | 
xtraIs = xtraIs,  | 
|
547  | 
swrappers = swrappers,  | 
|
548  | 
uwrappers = uwrappers,  | 
|
549  | 
safe0_netpair = safe0_netpair,  | 
|
550  | 
safep_netpair = safep_netpair,  | 
|
551  | 
haz_netpair = haz_netpair,  | 
|
552  | 
dup_netpair = dup_netpair}  | 
|
553  | 
end;  | 
|
554  | 
||
555  | 
infix 4 addXIs addXEs addXDs;  | 
|
556  | 
||
557  | 
val op addXIs = rev_foldl addXI;  | 
|
558  | 
val op addXEs = rev_foldl addXE;  | 
|
559  | 
||
560  | 
fun cs addXDs ths = cs addXEs (map make_elim ths);  | 
|
561  | 
||
562  | 
||
| 1800 | 563  | 
(*** Deletion of rules  | 
564  | 
Working out what to delete, requires repeating much of the code used  | 
|
565  | 
to insert.  | 
|
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
566  | 
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
 | 
567  | 
searches in all the lists and chooses the relevant delXX functions.  | 
| 1800 | 568  | 
***)  | 
569  | 
||
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
570  | 
fun delSI th  | 
| 6955 | 571  | 
          (cs as CS{safeIs, safeEs, hazIs, hazEs, xtraIs, xtraEs, swrappers, uwrappers,
 | 
572  | 
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
 | 
573  | 
if mem_thm (th, safeIs) then  | 
| 
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
574  | 
let val (safe0_rls, safep_rls) = partition (fn rl => nprems_of rl=0) [th]  | 
| 
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
575  | 
   in CS{safe0_netpair = delete (safe0_rls, []) safe0_netpair,
 | 
| 
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
576  | 
safep_netpair = delete (safep_rls, []) safep_netpair,  | 
| 
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
577  | 
safeIs = rem_thm (safeIs,th),  | 
| 
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
578  | 
safeEs = safeEs,  | 
| 
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
579  | 
hazIs = hazIs,  | 
| 
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
580  | 
hazEs = hazEs,  | 
| 6955 | 581  | 
xtraIs = xtraIs,  | 
582  | 
xtraEs = xtraEs,  | 
|
| 4651 | 583  | 
swrappers = swrappers,  | 
584  | 
uwrappers = uwrappers,  | 
|
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
585  | 
haz_netpair = haz_netpair,  | 
| 6955 | 586  | 
dup_netpair = dup_netpair,  | 
587  | 
xtra_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
 | 
588  | 
end  | 
| 
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
589  | 
else cs;  | 
| 1800 | 590  | 
|
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
591  | 
fun delSE th  | 
| 6955 | 592  | 
          (cs as CS{safeIs, safeEs, hazIs, hazEs, xtraIs, xtraEs, swrappers, uwrappers, 
 | 
593  | 
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
 | 
594  | 
if mem_thm (th, safeEs) then  | 
| 
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
595  | 
let val (safe0_rls, safep_rls) = partition (fn rl => nprems_of rl=1) [th]  | 
| 
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
596  | 
   in CS{safe0_netpair = delete ([], safe0_rls) safe0_netpair,
 | 
| 
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
597  | 
safep_netpair = delete ([], safep_rls) safep_netpair,  | 
| 
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
598  | 
safeIs = safeIs,  | 
| 
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
599  | 
safeEs = rem_thm (safeEs,th),  | 
| 
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
600  | 
hazIs = hazIs,  | 
| 
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
601  | 
hazEs = hazEs,  | 
| 6955 | 602  | 
xtraIs = xtraIs,  | 
603  | 
xtraEs = xtraEs,  | 
|
| 4651 | 604  | 
swrappers = swrappers,  | 
605  | 
uwrappers = uwrappers,  | 
|
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
606  | 
haz_netpair = haz_netpair,  | 
| 6955 | 607  | 
dup_netpair = dup_netpair,  | 
608  | 
xtra_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
 | 
609  | 
end  | 
| 
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
610  | 
else cs;  | 
| 1800 | 611  | 
|
612  | 
||
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
613  | 
fun delI th  | 
| 6955 | 614  | 
         (cs as CS{safeIs, safeEs, hazIs, hazEs, xtraIs, xtraEs, swrappers, uwrappers, 
 | 
615  | 
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
 | 
616  | 
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
 | 
617  | 
     CS{haz_netpair = delete ([th], []) haz_netpair,
 | 
| 1800 | 618  | 
dup_netpair = delete ([dup_intr th], []) dup_netpair,  | 
619  | 
safeIs = safeIs,  | 
|
620  | 
safeEs = safeEs,  | 
|
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
621  | 
hazIs = rem_thm (hazIs,th),  | 
| 1800 | 622  | 
hazEs = hazEs,  | 
| 6955 | 623  | 
xtraIs = xtraIs,  | 
624  | 
xtraEs = xtraEs,  | 
|
| 4651 | 625  | 
swrappers = swrappers,  | 
626  | 
uwrappers = uwrappers,  | 
|
| 1800 | 627  | 
safe0_netpair = safe0_netpair,  | 
| 6955 | 628  | 
safep_netpair = safep_netpair,  | 
629  | 
xtra_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
 | 
630  | 
else cs;  | 
| 1800 | 631  | 
|
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
632  | 
fun delE th  | 
| 6955 | 633  | 
	 (cs as CS{safeIs, safeEs, hazIs, hazEs, xtraIs, xtraEs, swrappers, uwrappers, 
 | 
634  | 
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
 | 
635  | 
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
 | 
636  | 
     CS{haz_netpair = delete ([], [th]) haz_netpair,
 | 
| 1800 | 637  | 
dup_netpair = delete ([], [dup_elim th]) dup_netpair,  | 
638  | 
safeIs = safeIs,  | 
|
639  | 
safeEs = safeEs,  | 
|
640  | 
hazIs = hazIs,  | 
|
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
641  | 
hazEs = rem_thm (hazEs,th),  | 
| 6955 | 642  | 
xtraIs = xtraIs,  | 
643  | 
xtraEs = xtraEs,  | 
|
644  | 
swrappers = swrappers,  | 
|
645  | 
uwrappers = uwrappers,  | 
|
646  | 
safe0_netpair = safe0_netpair,  | 
|
647  | 
safep_netpair = safep_netpair,  | 
|
648  | 
xtra_netpair = xtra_netpair}  | 
|
649  | 
else cs;  | 
|
650  | 
||
651  | 
||
652  | 
fun delXI th  | 
|
653  | 
         (cs as CS{safeIs, safeEs, hazIs, hazEs, xtraIs, xtraEs, swrappers, uwrappers, 
 | 
|
654  | 
safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair}) =  | 
|
655  | 
if mem_thm (th, xtraIs) then  | 
|
656  | 
     CS{xtra_netpair = delete ([th], []) xtra_netpair,
 | 
|
657  | 
safeIs = safeIs,  | 
|
658  | 
safeEs = safeEs,  | 
|
659  | 
hazIs = hazIs,  | 
|
660  | 
hazEs = hazEs,  | 
|
661  | 
xtraIs = rem_thm (xtraIs,th),  | 
|
662  | 
xtraEs = xtraEs,  | 
|
| 4651 | 663  | 
swrappers = swrappers,  | 
664  | 
uwrappers = uwrappers,  | 
|
| 1800 | 665  | 
safe0_netpair = safe0_netpair,  | 
| 6955 | 666  | 
safep_netpair = safep_netpair,  | 
667  | 
haz_netpair = haz_netpair,  | 
|
668  | 
dup_netpair = dup_netpair}  | 
|
669  | 
else cs;  | 
|
670  | 
||
671  | 
fun delXE th  | 
|
672  | 
	 (cs as CS{safeIs, safeEs, hazIs, hazEs, xtraIs, xtraEs, swrappers, uwrappers, 
 | 
|
673  | 
safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair}) =  | 
|
674  | 
if mem_thm (th, xtraEs) then  | 
|
675  | 
     CS{xtra_netpair = delete ([], [th]) xtra_netpair,
 | 
|
676  | 
safeIs = safeIs,  | 
|
677  | 
safeEs = safeEs,  | 
|
678  | 
hazIs = hazIs,  | 
|
679  | 
hazEs = hazEs,  | 
|
680  | 
xtraIs = xtraIs,  | 
|
681  | 
xtraEs = rem_thm (xtraEs,th),  | 
|
682  | 
swrappers = swrappers,  | 
|
683  | 
uwrappers = uwrappers,  | 
|
684  | 
safe0_netpair = safe0_netpair,  | 
|
685  | 
safep_netpair = safep_netpair,  | 
|
686  | 
haz_netpair = haz_netpair,  | 
|
687  | 
dup_netpair = dup_netpair}  | 
|
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
688  | 
else cs;  | 
| 1800 | 689  | 
|
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
690  | 
(*Delete ALL occurrences of "th" in the claset (perhaps from several lists)*)  | 
| 6955 | 691  | 
fun delrule (cs as CS{safeIs, safeEs, hazIs, hazEs, xtraIs, xtraEs, ...}, th) =
 | 
| 
2813
 
cc4c816dafdc
delrules now deletes ALL occurrences of a rule, since it may appear in any of
 
paulson 
parents: 
2689 
diff
changeset
 | 
692  | 
if mem_thm (th, safeIs) orelse mem_thm (th, safeEs) orelse  | 
| 6955 | 693  | 
mem_thm (th, hazIs) orelse mem_thm (th, hazEs) orelse  | 
694  | 
mem_thm (th, xtraIs) orelse mem_thm (th, xtraEs)  | 
|
695  | 
then delSI th (delSE th (delI th (delE th (delXI th (delXE th cs)))))  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
696  | 
       else (warning ("Rule not in claset\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
 | 
697  | 
cs);  | 
| 1800 | 698  | 
|
699  | 
val op delrules = foldl delrule;  | 
|
700  | 
||
701  | 
||
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
702  | 
(*** Modifying the wrapper tacticals ***)  | 
| 
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
703  | 
fun update_swrappers  | 
| 6955 | 704  | 
(CS{safeIs, safeEs, hazIs, hazEs, xtraIs, xtraEs, swrappers, uwrappers, 
 | 
705  | 
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
 | 
706  | 
 CS{safeIs = safeIs, safeEs = safeEs, hazIs = hazIs, hazEs = hazEs,
 | 
| 6955 | 707  | 
xtraIs = xtraIs, xtraEs = xtraEs,  | 
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
708  | 
swrappers = f swrappers, uwrappers = uwrappers,  | 
| 
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
709  | 
safe0_netpair = safe0_netpair, safep_netpair = safep_netpair,  | 
| 6955 | 710  | 
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
 | 
711  | 
|
| 
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
712  | 
fun update_uwrappers  | 
| 6955 | 713  | 
(CS{safeIs, safeEs, hazIs, hazEs, xtraIs, xtraEs, swrappers, uwrappers, 
 | 
714  | 
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
 | 
715  | 
 CS{safeIs = safeIs, safeEs = safeEs, hazIs = hazIs, hazEs = hazEs,
 | 
| 6955 | 716  | 
xtraIs = xtraIs, xtraEs = xtraEs,  | 
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
717  | 
swrappers = swrappers, uwrappers = f uwrappers,  | 
| 
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
718  | 
safe0_netpair = safe0_netpair, safep_netpair = safep_netpair,  | 
| 6955 | 719  | 
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
 | 
720  | 
|
| 
982
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
721  | 
|
| 4651 | 722  | 
(*Add/replace a safe wrapper*)  | 
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
723  | 
fun cs addSWrapper new_swrapper = update_swrappers cs (fn swrappers =>  | 
| 
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
724  | 
(case assoc_string (swrappers,(fst new_swrapper)) of None =>()  | 
| 4651 | 725  | 
	   | Some x => warning("overwriting safe wrapper "^fst new_swrapper); 
 | 
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
726  | 
overwrite (swrappers, new_swrapper)));  | 
| 4651 | 727  | 
|
728  | 
(*Add/replace an unsafe wrapper*)  | 
|
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
729  | 
fun cs addWrapper new_uwrapper = update_uwrappers cs (fn uwrappers =>  | 
| 
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
730  | 
(case assoc_string (uwrappers,(fst new_uwrapper)) of None =>()  | 
| 4651 | 731  | 
	   | Some x => warning ("overwriting unsafe wrapper "^fst new_uwrapper);
 | 
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
732  | 
overwrite (uwrappers, new_uwrapper)));  | 
| 
982
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
733  | 
|
| 4651 | 734  | 
(*Remove a safe wrapper*)  | 
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
735  | 
fun cs delSWrapper name = update_swrappers cs (fn swrappers =>  | 
| 
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
736  | 
let val (del,rest) = partition (fn (n,_) => n=name) swrappers  | 
| 
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
737  | 
    in if null del then (warning ("No such safe wrapper in claset: "^ name); 
 | 
| 
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
738  | 
swrappers) else rest end);  | 
| 
982
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
739  | 
|
| 4651 | 740  | 
(*Remove an unsafe wrapper*)  | 
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
741  | 
fun cs delWrapper name = update_uwrappers cs (fn uwrappers =>  | 
| 
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
742  | 
let val (del,rest) = partition (fn (n,_) => n=name) uwrappers  | 
| 
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
743  | 
    in if null del then (warning ("No such unsafe wrapper in claset: " ^ name);
 | 
| 
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
744  | 
uwrappers) else rest end);  | 
| 
982
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
745  | 
|
| 2630 | 746  | 
(*compose a safe tactic sequentially before/alternatively after safe_step_tac*)  | 
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
747  | 
fun cs addSbefore (name, tac1) =  | 
| 5523 | 748  | 
cs addSWrapper (name, fn tac2 => tac1 ORELSE' tac2);  | 
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
749  | 
fun cs addSaltern (name, tac2) =  | 
| 5523 | 750  | 
cs addSWrapper (name, fn tac1 => tac1 ORELSE' tac2);  | 
| 
982
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
751  | 
|
| 2630 | 752  | 
(*compose a tactic sequentially before/alternatively after the step tactic*)  | 
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
753  | 
fun cs addbefore (name, tac1) =  | 
| 5523 | 754  | 
cs addWrapper (name, fn tac2 => tac1 APPEND' tac2);  | 
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
755  | 
fun cs addaltern (name, tac2) =  | 
| 5523 | 756  | 
cs addWrapper (name, fn tac1 => tac1 APPEND' tac2);  | 
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
757  | 
|
| 5523 | 758  | 
fun cs addD2 (name, thm) =  | 
759  | 
cs addaltern (name, dtac thm THEN' atac);  | 
|
760  | 
fun cs addE2 (name, thm) =  | 
|
761  | 
cs addaltern (name, etac thm THEN' atac);  | 
|
762  | 
fun cs addSD2 (name, thm) =  | 
|
763  | 
cs addSaltern (name, dmatch_tac [thm] THEN' eq_assume_tac);  | 
|
764  | 
fun cs addSE2 (name, thm) =  | 
|
765  | 
cs addSaltern (name, ematch_tac [thm] THEN' eq_assume_tac);  | 
|
| 
982
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
766  | 
|
| 1711 | 767  | 
(*Merge works by adding all new rules of the 2nd claset into the 1st claset.  | 
768  | 
Merging the term nets may look more efficient, but the rather delicate  | 
|
769  | 
treatment of priority might get muddled up.*)  | 
|
770  | 
fun merge_cs  | 
|
| 6955 | 771  | 
    (cs as CS{safeIs, safeEs, hazIs, hazEs, xtraIs, xtraEs, ...},
 | 
| 4765 | 772  | 
     CS{safeIs=safeIs2, safeEs=safeEs2, hazIs=hazIs2, hazEs=hazEs2,
 | 
| 6955 | 773  | 
xtraIs=xtraIs2, xtraEs=xtraEs2, swrappers, uwrappers, ...}) =  | 
| 1711 | 774  | 
let val safeIs' = gen_rems eq_thm (safeIs2,safeIs)  | 
775  | 
val safeEs' = gen_rems eq_thm (safeEs2,safeEs)  | 
|
| 2630 | 776  | 
val hazIs' = gen_rems eq_thm ( hazIs2, hazIs)  | 
777  | 
val hazEs' = gen_rems eq_thm ( hazEs2, hazEs)  | 
|
| 6955 | 778  | 
val xtraIs' = gen_rems eq_thm (xtraIs2, xtraIs)  | 
779  | 
val xtraEs' = gen_rems eq_thm (xtraEs2, xtraEs)  | 
|
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
780  | 
val cs1 = cs addSIs safeIs'  | 
| 4765 | 781  | 
addSEs safeEs'  | 
782  | 
addIs hazIs'  | 
|
783  | 
addEs hazEs'  | 
|
| 6955 | 784  | 
addXIs xtraIs'  | 
785  | 
addXEs xtraEs'  | 
|
| 
4767
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
786  | 
val cs2 = update_swrappers cs1 (fn ws => merge_alists ws swrappers);  | 
| 
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
787  | 
val cs3 = update_uwrappers cs2 (fn ws => merge_alists ws uwrappers);  | 
| 
 
b9f3468c6ee2
introduced functions for updating the wrapper lists
 
oheimb 
parents: 
4765 
diff
changeset
 | 
788  | 
in cs3  | 
| 1711 | 789  | 
end;  | 
790  | 
||
| 
982
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
791  | 
|
| 1800 | 792  | 
(**** Simple tactics for theorem proving ****)  | 
| 0 | 793  | 
|
794  | 
(*Attack subgoals using safe inferences -- matching, not resolution*)  | 
|
| 2630 | 795  | 
fun safe_step_tac (cs as CS{safe0_netpair,safep_netpair,...}) = 
 | 
| 4651 | 796  | 
appSWrappers cs (FIRST' [  | 
| 2630 | 797  | 
eq_assume_tac,  | 
798  | 
eq_mp_tac,  | 
|
799  | 
bimatch_from_nets_tac safe0_netpair,  | 
|
800  | 
FIRST' hyp_subst_tacs,  | 
|
801  | 
bimatch_from_nets_tac safep_netpair]);  | 
|
| 0 | 802  | 
|
| 5757 | 803  | 
(*Repeatedly attack a subgoal using safe inferences -- it's deterministic!*)  | 
804  | 
fun safe_steps_tac cs = REPEAT_DETERM1 o  | 
|
805  | 
(fn i => COND (has_fewer_prems i) no_tac (safe_step_tac cs i));  | 
|
806  | 
||
| 0 | 807  | 
(*Repeatedly attack subgoals using safe inferences -- it's deterministic!*)  | 
| 5757 | 808  | 
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
 | 
809  | 
|
| 
3705
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
810  | 
|
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
811  | 
(*** 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
 | 
812  | 
|
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
813  | 
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
 | 
814  | 
|
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
815  | 
(*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
 | 
816  | 
create precisely n subgoals.*)  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
817  | 
fun n_bimatch_from_nets_tac n =  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
818  | 
biresolution_from_nets_tac (orderlist o filter (nsubgoalsP n)) true;  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
819  | 
|
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
820  | 
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
 | 
821  | 
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
 | 
822  | 
|
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
823  | 
(*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
 | 
824  | 
fun bimatch2_tac netpair i =  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
825  | 
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
 | 
826  | 
(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
 | 
827  | 
|
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
828  | 
(*Attack subgoals using safe inferences -- matching, not resolution*)  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
829  | 
fun clarify_step_tac (cs as CS{safe0_netpair,safep_netpair,...}) = 
 | 
| 4651 | 830  | 
appSWrappers cs (FIRST' [  | 
| 
3705
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
831  | 
eq_assume_contr_tac,  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
832  | 
bimatch_from_nets_tac safe0_netpair,  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
833  | 
FIRST' hyp_subst_tacs,  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
834  | 
n_bimatch_from_nets_tac 1 safep_netpair,  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
835  | 
bimatch2_tac safep_netpair]);  | 
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
836  | 
|
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
837  | 
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
 | 
838  | 
|
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
839  | 
|
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
840  | 
(*** 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
 | 
841  | 
|
| 4066 | 842  | 
(*Backtracking is allowed among the various these unsafe ways of  | 
843  | 
proving a subgoal. *)  | 
|
| 
747
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
844  | 
fun inst0_step_tac (CS{safe0_netpair,safep_netpair,...}) =
 | 
| 
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
845  | 
assume_tac APPEND'  | 
| 
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
846  | 
contr_tac APPEND'  | 
| 
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
847  | 
biresolve_from_nets_tac safe0_netpair;  | 
| 
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
848  | 
|
| 4066 | 849  | 
(*These unsafe steps could generate more subgoals.*)  | 
| 
747
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
850  | 
fun instp_step_tac (CS{safep_netpair,...}) =
 | 
| 
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
851  | 
biresolve_from_nets_tac safep_netpair;  | 
| 0 | 852  | 
|
853  | 
(*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
 | 
854  | 
fun inst_step_tac cs = inst0_step_tac cs APPEND' instp_step_tac cs;  | 
| 0 | 855  | 
|
| 
982
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
856  | 
fun haz_step_tac (CS{haz_netpair,...}) = 
 | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
857  | 
biresolve_from_nets_tac haz_netpair;  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
858  | 
|
| 0 | 859  | 
(*Single step for the prover. FAILS unless it makes progress. *)  | 
| 5523 | 860  | 
fun step_tac cs i = safe_tac cs ORELSE appWrappers cs  | 
861  | 
(inst_step_tac cs ORELSE' haz_step_tac cs) i;  | 
|
| 0 | 862  | 
|
863  | 
(*Using a "safe" rule to instantiate variables is unsafe. This tactic  | 
|
864  | 
allows backtracking from "safe" rules to "unsafe" rules here.*)  | 
|
| 5523 | 865  | 
fun slow_step_tac cs i = safe_tac cs ORELSE appWrappers cs  | 
866  | 
(inst_step_tac cs APPEND' haz_step_tac cs) i;  | 
|
| 0 | 867  | 
|
| 1800 | 868  | 
(**** The following tactics all fail unless they solve one goal ****)  | 
| 0 | 869  | 
|
870  | 
(*Dumb but fast*)  | 
|
871  | 
fun fast_tac cs = SELECT_GOAL (DEPTH_SOLVE (step_tac cs 1));  | 
|
872  | 
||
873  | 
(*Slower but smarter than fast_tac*)  | 
|
874  | 
fun best_tac cs =  | 
|
875  | 
SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, sizef) (step_tac cs 1));  | 
|
876  | 
||
877  | 
fun slow_tac cs = SELECT_GOAL (DEPTH_SOLVE (slow_step_tac cs 1));  | 
|
878  | 
||
879  | 
fun slow_best_tac cs =  | 
|
880  | 
SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, sizef) (slow_step_tac cs 1));  | 
|
881  | 
||
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
882  | 
|
| 1800 | 883  | 
(***ASTAR with weight weight_ASTAR, by Norbert Voelker*)  | 
| 
1587
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
884  | 
val weight_ASTAR = ref 5;  | 
| 
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
885  | 
|
| 
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
886  | 
fun astar_tac cs =  | 
| 
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
887  | 
SELECT_GOAL ( ASTAR (has_fewer_prems 1  | 
| 
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
888  | 
, fn level =>(fn thm =>size_of_thm thm + !weight_ASTAR *level))  | 
| 
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
889  | 
(step_tac cs 1));  | 
| 
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
890  | 
|
| 
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
891  | 
fun slow_astar_tac cs =  | 
| 
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
892  | 
SELECT_GOAL ( ASTAR (has_fewer_prems 1  | 
| 
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
893  | 
, fn level =>(fn thm =>size_of_thm thm + !weight_ASTAR *level))  | 
| 
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
894  | 
(slow_step_tac cs 1));  | 
| 
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
895  | 
|
| 1800 | 896  | 
(**** 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
 | 
897  | 
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
 | 
898  | 
easy theorems faster, but loses completeness -- and many of the harder  | 
| 1800 | 899  | 
theorems such as 43. ****)  | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
900  | 
|
| 
747
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
901  | 
(*Non-deterministic! Could always expand the first unsafe connective.  | 
| 
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
902  | 
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
 | 
903  | 
greater search depth required.*)  | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
904  | 
fun dup_step_tac (cs as (CS{dup_netpair,...})) = 
 | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
905  | 
biresolve_from_nets_tac dup_netpair;  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
906  | 
|
| 5523 | 907  | 
(*Searching to depth m. A variant called nodup_depth_tac appears in clasimp.ML*)  | 
| 5757 | 908  | 
local  | 
909  | 
fun slow_step_tac' cs = appWrappers cs  | 
|
910  | 
(instp_step_tac cs APPEND' dup_step_tac cs);  | 
|
911  | 
in fun depth_tac cs m i state = SELECT_GOAL  | 
|
912  | 
(safe_steps_tac cs 1 THEN_ELSE  | 
|
913  | 
(DEPTH_SOLVE (depth_tac cs m 1),  | 
|
914  | 
inst0_step_tac cs 1 APPEND COND (K (m=0)) no_tac  | 
|
915  | 
(slow_step_tac' cs 1 THEN DEPTH_SOLVE (depth_tac cs (m-1) 1))  | 
|
916  | 
)) i state;  | 
|
917  | 
end;  | 
|
| 
747
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
918  | 
|
| 2173 | 919  | 
(*Search, with depth bound m.  | 
920  | 
This is the "entry point", which does safe inferences first.*)  | 
|
| 
747
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
921  | 
fun safe_depth_tac cs m =  | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
922  | 
SUBGOAL  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
923  | 
(fn (prem,i) =>  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
924  | 
let val deti =  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
925  | 
(*No Vars in the goal? No need to backtrack between goals.*)  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
926  | 
case term_vars prem of  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
927  | 
[] => DETERM  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
928  | 
| _::_ => I  | 
| 
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
929  | 
in SELECT_GOAL (TRY (safe_tac cs) THEN  | 
| 
747
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
930  | 
DEPTH_SOLVE (deti (depth_tac cs m 1))) i  | 
| 
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
931  | 
end);  | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
932  | 
|
| 
2868
 
17a23a62f82a
New DEEPEN allows giving an upper bound for deepen_tac
 
paulson 
parents: 
2813 
diff
changeset
 | 
933  | 
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
 | 
934  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
935  | 
|
| 1724 | 936  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
937  | 
(** claset theory data **)  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
938  | 
|
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
939  | 
(* init data kind claset *)  | 
| 1724 | 940  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
941  | 
exception CSData of claset ref;  | 
| 1724 | 942  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
943  | 
local  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
944  | 
val empty = CSData (ref empty_cs);  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
945  | 
|
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
946  | 
(*create new references*)  | 
| 6556 | 947  | 
fun copy (ClasetData (ref (CSData (ref cs)))) =  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
948  | 
ClasetData (ref (CSData (ref cs)));  | 
| 6556 | 949  | 
val prep_ext = copy;  | 
| 1724 | 950  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
951  | 
fun merge (ClasetData (ref (CSData (ref cs1))), ClasetData (ref (CSData (ref cs2)))) =  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
952  | 
ClasetData (ref (CSData (ref (merge_cs (cs1, cs2)))));  | 
| 1724 | 953  | 
|
| 4259 | 954  | 
fun print (_: Sign.sg) (ClasetData (ref (CSData (ref cs)))) = print_cs cs;  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
955  | 
in  | 
| 6556 | 956  | 
val _ = fix_methods (empty, copy, prep_ext, merge, print);  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
957  | 
end;  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
958  | 
|
| 1724 | 959  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
960  | 
(* access claset *)  | 
| 1724 | 961  | 
|
| 5001 | 962  | 
val print_claset = Theory.print_data clasetK;  | 
| 4380 | 963  | 
|
| 5001 | 964  | 
val claset_ref_of_sg = Sign.get_data clasetK (fn ClasetData (ref (CSData r)) => r);  | 
| 1807 | 965  | 
|
| 6391 | 966  | 
val claset_ref_of = claset_ref_of_sg o Theory.sign_of;  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
967  | 
val claset_of_sg = ! o claset_ref_of_sg;  | 
| 6391 | 968  | 
val claset_of = claset_of_sg o Theory.sign_of;  | 
| 1800 | 969  | 
|
| 6391 | 970  | 
fun CLASET tacf state = tacf (claset_of_sg (Thm.sign_of_thm state)) state;  | 
971  | 
fun CLASET' tacf i state = tacf (claset_of_sg (Thm.sign_of_thm state)) i state;  | 
|
| 
3705
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
972  | 
|
| 5028 | 973  | 
val claset = claset_of o Context.the_context;  | 
| 6391 | 974  | 
val claset_ref = claset_ref_of_sg o Theory.sign_of o Context.the_context;  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
975  | 
|
| 
3705
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
976  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
977  | 
(* change claset *)  | 
| 1800 | 978  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
979  | 
fun change_claset f x = claset_ref () := (f (claset (), x));  | 
| 1724 | 980  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
981  | 
val AddDs = change_claset (op addDs);  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
982  | 
val AddEs = change_claset (op addEs);  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
983  | 
val AddIs = change_claset (op addIs);  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
984  | 
val AddSDs = change_claset (op addSDs);  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
985  | 
val AddSEs = change_claset (op addSEs);  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
986  | 
val AddSIs = change_claset (op addSIs);  | 
| 6955 | 987  | 
val AddXDs = change_claset (op addXDs);  | 
988  | 
val AddXEs = change_claset (op addXEs);  | 
|
989  | 
val AddXIs = change_claset (op addXIs);  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
990  | 
val Delrules = change_claset (op delrules);  | 
| 
3727
 
ed63c05d7992
Added Safe_tac; all other default claset tactics now dereference "claset"
 
paulson 
parents: 
3705 
diff
changeset
 | 
991  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
992  | 
|
| 5841 | 993  | 
(* proof data kind 'Provers/claset' *)  | 
994  | 
||
995  | 
structure LocalClasetArgs =  | 
|
996  | 
struct  | 
|
997  | 
val name = "Provers/claset";  | 
|
998  | 
type T = claset;  | 
|
999  | 
val init = claset_of;  | 
|
1000  | 
fun print _ cs = print_cs cs;  | 
|
1001  | 
end;  | 
|
1002  | 
||
1003  | 
structure LocalClaset = ProofDataFun(LocalClasetArgs);  | 
|
1004  | 
val print_local_claset = LocalClaset.print;  | 
|
1005  | 
val get_local_claset = LocalClaset.get;  | 
|
1006  | 
val put_local_claset = LocalClaset.put;  | 
|
1007  | 
||
1008  | 
||
| 5885 | 1009  | 
(* attributes *)  | 
1010  | 
||
1011  | 
fun change_global_cs f (thy, th) =  | 
|
1012  | 
let val r = claset_ref_of thy  | 
|
| 6096 | 1013  | 
in r := f (! r, [th]); (thy, th) end;  | 
| 5885 | 1014  | 
|
1015  | 
fun change_local_cs f (ctxt, th) =  | 
|
| 6096 | 1016  | 
let val cs = f (get_local_claset ctxt, [th])  | 
| 5885 | 1017  | 
in (put_local_claset cs ctxt, th) end;  | 
1018  | 
||
1019  | 
val safe_dest_global = change_global_cs (op addSDs);  | 
|
1020  | 
val safe_elim_global = change_global_cs (op addSEs);  | 
|
1021  | 
val safe_intro_global = change_global_cs (op addSIs);  | 
|
| 6955 | 1022  | 
val haz_dest_global = change_global_cs (op addDs);  | 
1023  | 
val haz_elim_global = change_global_cs (op addEs);  | 
|
1024  | 
val haz_intro_global = change_global_cs (op addIs);  | 
|
1025  | 
val xtra_dest_global = change_global_cs (op addXDs);  | 
|
1026  | 
val xtra_elim_global = change_global_cs (op addXEs);  | 
|
1027  | 
val xtra_intro_global = change_global_cs (op addXIs);  | 
|
| 5885 | 1028  | 
val delrule_global = change_global_cs (op delrules);  | 
1029  | 
||
| 6955 | 1030  | 
val safe_dest_local = change_local_cs (op addSDs);  | 
1031  | 
val safe_elim_local = change_local_cs (op addSEs);  | 
|
1032  | 
val safe_intro_local = change_local_cs (op addSIs);  | 
|
| 5885 | 1033  | 
val haz_dest_local = change_local_cs (op addDs);  | 
1034  | 
val haz_elim_local = change_local_cs (op addEs);  | 
|
1035  | 
val haz_intro_local = change_local_cs (op addIs);  | 
|
| 6955 | 1036  | 
val xtra_dest_local = change_local_cs (op addXDs);  | 
1037  | 
val xtra_elim_local = change_local_cs (op addXEs);  | 
|
1038  | 
val xtra_intro_local = change_local_cs (op addXIs);  | 
|
| 5885 | 1039  | 
val delrule_local = change_local_cs (op delrules);  | 
1040  | 
||
1041  | 
||
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
1042  | 
(* tactics referring to the implicit claset *)  | 
| 1800 | 1043  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
1044  | 
(*the abstraction over the proof state delays the dereferencing*)  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
1045  | 
fun Safe_tac st = safe_tac (claset()) st;  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
1046  | 
fun Safe_step_tac i st = safe_step_tac (claset()) i st;  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
1047  | 
fun Clarify_step_tac i st = clarify_step_tac (claset()) i st;  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
1048  | 
fun Clarify_tac i st = clarify_tac (claset()) i st;  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
1049  | 
fun Step_tac i st = step_tac (claset()) i st;  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
1050  | 
fun Fast_tac i st = fast_tac (claset()) i st;  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
1051  | 
fun Best_tac i st = best_tac (claset()) i st;  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
1052  | 
fun Slow_tac i st = slow_tac (claset()) i st;  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
1053  | 
fun Slow_best_tac i st = slow_best_tac (claset()) i st;  | 
| 
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
1054  | 
fun Deepen_tac m = deepen_tac (claset()) m;  | 
| 2066 | 1055  | 
|
| 1800 | 1056  | 
|
| 0 | 1057  | 
end;  | 
| 5841 | 1058  | 
|
1059  | 
||
1060  | 
||
| 5885 | 1061  | 
(** concrete syntax of attributes **)  | 
| 5841 | 1062  | 
|
1063  | 
(* add / del rules *)  | 
|
1064  | 
||
1065  | 
val introN = "intro";  | 
|
1066  | 
val elimN = "elim";  | 
|
1067  | 
val destN = "dest";  | 
|
1068  | 
val delN = "del";  | 
|
| 5885 | 1069  | 
val delruleN = "delrule";  | 
| 5841 | 1070  | 
|
| 5885 | 1071  | 
val bang = Args.$$$ "!";  | 
| 6955 | 1072  | 
val bbang = Args.$$$ "!!";  | 
| 5841 | 1073  | 
|
| 6955 | 1074  | 
fun cla_att change xtra haz safe = Attrib.syntax  | 
1075  | 
(Scan.lift ((bbang >> K xtra || bang >> K haz || Scan.succeed safe) >> change));  | 
|
| 5841 | 1076  | 
|
| 6955 | 1077  | 
fun cla_attr f g h = (cla_att change_global_cs f g h, cla_att change_local_cs f g h);  | 
| 5885 | 1078  | 
val del_attr = (Attrib.no_args delrule_global, Attrib.no_args delrule_local);  | 
| 5841 | 1079  | 
|
1080  | 
||
1081  | 
(* setup_attrs *)  | 
|
1082  | 
||
1083  | 
val setup_attrs = Attrib.add_attributes  | 
|
| 6955 | 1084  | 
[(destN, cla_attr (op addXDs) (op addDs) (op addSDs), "destruction rule"),  | 
1085  | 
(elimN, cla_attr (op addXEs) (op addEs) (op addSEs), "elimination rule"),  | 
|
1086  | 
(introN, cla_attr (op addXIs) (op addIs) (op addSIs), "introduction rule"),  | 
|
| 5885 | 1087  | 
(delruleN, del_attr, "delete rule")];  | 
| 5841 | 1088  | 
|
1089  | 
||
1090  | 
||
| 5885 | 1091  | 
(** single rule proof method **)  | 
| 5841 | 1092  | 
|
1093  | 
(* utils *)  | 
|
1094  | 
||
1095  | 
fun resolve_from_seq_tac rq i st = Seq.flat (Seq.map (fn r => rtac r i st) rq);  | 
|
1096  | 
fun order_rules xs = map snd (Tactic.orderlist xs);  | 
|
1097  | 
||
1098  | 
fun find_rules concl nets =  | 
|
1099  | 
let fun rules_of (inet, _) = order_rules (Net.unify_term inet concl)  | 
|
1100  | 
in flat (map rules_of nets) end;  | 
|
1101  | 
||
1102  | 
fun find_erules [] _ = []  | 
|
| 6955 | 1103  | 
| find_erules (fact :: _) nets =  | 
| 5841 | 1104  | 
let  | 
| 6502 | 1105  | 
fun may_unify net = Net.unify_term net o Logic.strip_assums_concl o #prop o Thm.rep_thm;  | 
| 6955 | 1106  | 
fun erules_of (_, enet) = order_rules (may_unify enet fact);  | 
| 6502 | 1107  | 
in flat (map erules_of nets) end;  | 
| 5841 | 1108  | 
|
1109  | 
||
1110  | 
(* trace rules *)  | 
|
1111  | 
||
1112  | 
val trace_rules = ref false;  | 
|
1113  | 
||
1114  | 
fun print_rules rules i =  | 
|
1115  | 
if not (! trace_rules) then ()  | 
|
1116  | 
else  | 
|
1117  | 
    Pretty.writeln (Pretty.big_list ("trying standard rule(s) on goal #" ^ string_of_int i ^ ":")
 | 
|
1118  | 
(map Display.pretty_thm rules));  | 
|
1119  | 
||
1120  | 
||
1121  | 
(* single_tac *)  | 
|
1122  | 
||
| 6502 | 1123  | 
val imp_elim_netpair = insert (0, 0) ([], [imp_elim]) empty_netpair;  | 
1124  | 
val not_elim_netpair = insert (0, 0) ([], [Data.not_elim]) empty_netpair;  | 
|
1125  | 
||
| 6096 | 1126  | 
fun single_tac cs facts =  | 
| 5841 | 1127  | 
let  | 
| 6955 | 1128  | 
    val CS {safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair, ...} = cs;
 | 
| 6502 | 1129  | 
val nets = [imp_elim_netpair, safe0_netpair, safep_netpair,  | 
| 6955 | 1130  | 
not_elim_netpair, haz_netpair, dup_netpair, xtra_netpair];  | 
| 6492 | 1131  | 
val erules = find_erules facts nets;  | 
| 5841 | 1132  | 
|
1133  | 
val tac = SUBGOAL (fn (goal, i) =>  | 
|
1134  | 
let  | 
|
1135  | 
val irules = find_rules (Logic.strip_assums_concl goal) nets;  | 
|
1136  | 
val rules = erules @ irules;  | 
|
1137  | 
val ruleq = Method.forward_chain facts rules;  | 
|
1138  | 
in  | 
|
1139  | 
print_rules rules i;  | 
|
1140  | 
fn st => Seq.flat (Seq.map (fn rule => Tactic.rtac rule i st) ruleq) end);  | 
|
1141  | 
in tac end;  | 
|
1142  | 
||
1143  | 
val single = Method.METHOD (FIRSTGOAL o (fn facts => CLASET' (fn cs => single_tac cs facts)));  | 
|
1144  | 
||
1145  | 
||
1146  | 
||
| 6502 | 1147  | 
(** more proof methods **)  | 
1148  | 
||
1149  | 
(* contradiction *)  | 
|
1150  | 
||
| 7004 | 1151  | 
val contradiction =  | 
1152  | 
Method.METHOD (fn facts => FIRSTGOAL (Method.same_tac facts THEN' contr_tac));  | 
|
| 6502 | 1153  | 
|
1154  | 
||
1155  | 
(* automatic methods *)  | 
|
| 5841 | 1156  | 
|
| 5927 | 1157  | 
val cla_modifiers =  | 
| 6955 | 1158  | 
[Args.$$$ destN -- bbang >> K xtra_dest_local,  | 
1159  | 
Args.$$$ destN -- bang >> K haz_dest_local,  | 
|
| 5927 | 1160  | 
Args.$$$ destN >> K safe_dest_local,  | 
| 6955 | 1161  | 
Args.$$$ elimN -- bbang >> K xtra_elim_local,  | 
| 5927 | 1162  | 
Args.$$$ elimN -- bang >> K haz_elim_local,  | 
1163  | 
Args.$$$ elimN >> K safe_elim_local,  | 
|
| 6955 | 1164  | 
Args.$$$ introN -- bbang >> K xtra_intro_local,  | 
| 5927 | 1165  | 
Args.$$$ introN -- bang >> K haz_intro_local,  | 
1166  | 
Args.$$$ introN >> K safe_intro_local,  | 
|
1167  | 
Args.$$$ delN >> K delrule_local];  | 
|
1168  | 
||
1169  | 
val cla_args = Method.only_sectioned_args cla_modifiers;  | 
|
| 5841 | 1170  | 
|
| 5885 | 1171  | 
fun cla_meth tac ctxt = Method.METHOD0 (tac (get_local_claset ctxt));  | 
1172  | 
fun cla_meth' tac ctxt = Method.METHOD0 (FIRSTGOAL (tac (get_local_claset ctxt)));  | 
|
| 5841 | 1173  | 
|
| 5885 | 1174  | 
val cla_method = cla_args o cla_meth;  | 
1175  | 
val cla_method' = cla_args o cla_meth';  | 
|
| 5841 | 1176  | 
|
1177  | 
||
1178  | 
||
1179  | 
(** setup_methods **)  | 
|
1180  | 
||
1181  | 
val setup_methods = Method.add_methods  | 
|
1182  | 
 [("single", Method.no_args single, "apply standard rule (single step)"),
 | 
|
1183  | 
  ("default", Method.no_args single, "apply standard rule (single step)"),
 | 
|
| 6502 | 1184  | 
  ("contradiction", Method.no_args contradiction, "proof by contradiction"),
 | 
| 5885 | 1185  | 
  ("safe_tac", cla_method safe_tac, "safe_tac"),
 | 
1186  | 
  ("safe_step", cla_method' safe_step_tac, "step_tac"),
 | 
|
| 7004 | 1187  | 
  ("fast", cla_method' fast_tac, "classical prover (depth-first)"),
 | 
1188  | 
  ("best", cla_method' best_tac, "classical prover (best-first)"),
 | 
|
1189  | 
  ("slow", cla_method' slow_tac, "classical prover (depth-first, more backtracking)"),
 | 
|
1190  | 
  ("slow_best", cla_method' slow_best_tac, "classical prover (best-first, more backtracking)")];
 | 
|
| 5841 | 1191  | 
|
1192  | 
||
1193  | 
||
1194  | 
(** theory setup **)  | 
|
1195  | 
||
1196  | 
(* FIXME claset theory data *)  | 
|
1197  | 
||
1198  | 
val setup = [LocalClaset.init, setup_attrs, setup_methods];  | 
|
1199  | 
||
1200  | 
||
1201  | 
end;  |