| author | wenzelm | 
| Sun, 22 Sep 2024 16:12:15 +0200 | |
| changeset 80923 | 6c9628a116cc | 
| parent 74561 | 8e6c973003c8 | 
| child 82804 | 070585eb5d54 | 
| permissions | -rw-r--r-- | 
| 
29269
 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 
wenzelm 
parents: 
26928 
diff
changeset
 | 
1  | 
(* Title: Sequents/prover.ML  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
3  | 
Copyright 1992 University of Cambridge  | 
| 
7097
 
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
 
paulson 
parents: 
6054 
diff
changeset
 | 
4  | 
|
| 
29269
 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 
wenzelm 
parents: 
26928 
diff
changeset
 | 
5  | 
Simple classical reasoner for the sequent calculus, based on "theorem packs".  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
6  | 
*)  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
7  | 
|
| 55228 | 8  | 
signature CLA =  | 
9  | 
sig  | 
|
10  | 
type pack  | 
|
11  | 
val empty_pack: pack  | 
|
12  | 
val get_pack: Proof.context -> pack  | 
|
13  | 
val put_pack: pack -> Proof.context -> Proof.context  | 
|
14  | 
val pretty_pack: Proof.context -> Pretty.T  | 
|
15  | 
val add_safe: thm -> Proof.context -> Proof.context  | 
|
16  | 
val add_unsafe: thm -> Proof.context -> Proof.context  | 
|
17  | 
val safe_add: attribute  | 
|
18  | 
val unsafe_add: attribute  | 
|
19  | 
val method: (Proof.context -> int -> tactic) -> (Proof.context -> Proof.method) context_parser  | 
|
20  | 
val trace: bool Config.T  | 
|
21  | 
val forms_of_seq: term -> term list  | 
|
22  | 
val safe_tac: Proof.context -> int -> tactic  | 
|
23  | 
val step_tac: Proof.context -> int -> tactic  | 
|
24  | 
val pc_tac: Proof.context -> int -> tactic  | 
|
25  | 
val fast_tac: Proof.context -> int -> tactic  | 
|
26  | 
val best_tac: Proof.context -> int -> tactic  | 
|
27  | 
end;  | 
|
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
28  | 
|
| 55228 | 29  | 
structure Cla: CLA =  | 
| 7122 | 30  | 
struct  | 
31  | 
||
| 55228 | 32  | 
(** rule declarations **)  | 
| 7122 | 33  | 
|
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
34  | 
(*A theorem pack has the form (safe rules, unsafe rules)  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
35  | 
An unsafe rule is incomplete or introduces variables in subgoals,  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
36  | 
and is tried only when the safe rules are not applicable. *)  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
37  | 
|
| 55228 | 38  | 
fun less (rl1, rl2) = Thm.nprems_of rl1 < Thm.nprems_of rl2;  | 
39  | 
val sort_rules = sort (make_ord less);  | 
|
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
40  | 
|
| 55228 | 41  | 
datatype pack = Pack of thm list * thm list;  | 
42  | 
||
43  | 
val empty_pack = Pack ([], []);  | 
|
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
44  | 
|
| 55228 | 45  | 
structure Pack = Generic_Data  | 
46  | 
(  | 
|
47  | 
type T = pack;  | 
|
48  | 
val empty = empty_pack;  | 
|
49  | 
fun merge (Pack (safes, unsafes), Pack (safes', unsafes')) =  | 
|
50  | 
Pack  | 
|
51  | 
(sort_rules (Thm.merge_thms (safes, safes')),  | 
|
52  | 
sort_rules (Thm.merge_thms (unsafes, unsafes')));  | 
|
53  | 
);  | 
|
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
54  | 
|
| 55228 | 55  | 
val put_pack = Context.proof_map o Pack.put;  | 
56  | 
val get_pack = Pack.get o Context.Proof;  | 
|
57  | 
fun get_rules ctxt = let val Pack rules = get_pack ctxt in rules end;  | 
|
58  | 
||
59  | 
||
60  | 
(* print pack *)  | 
|
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
61  | 
|
| 55228 | 62  | 
fun pretty_pack ctxt =  | 
63  | 
let val (safes, unsafes) = get_rules ctxt in  | 
|
64  | 
Pretty.chunks  | 
|
| 61268 | 65  | 
[Pretty.big_list "safe rules:" (map (Thm.pretty_thm ctxt) safes),  | 
66  | 
Pretty.big_list "unsafe rules:" (map (Thm.pretty_thm ctxt) unsafes)]  | 
|
| 55228 | 67  | 
end;  | 
| 
7097
 
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
 
paulson 
parents: 
6054 
diff
changeset
 | 
68  | 
|
| 55228 | 69  | 
val _ =  | 
| 69593 | 70  | 
Outer_Syntax.command \<^command_keyword>\<open>print_pack\<close> "print pack of classical rules"  | 
| 55228 | 71  | 
(Scan.succeed (Toplevel.keep (Pretty.writeln o pretty_pack o Toplevel.context_of)));  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
72  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
73  | 
|
| 55228 | 74  | 
(* declare rules *)  | 
75  | 
||
76  | 
fun add_rule which th context = context |> Pack.map (fn Pack rules =>  | 
|
77  | 
Pack (rules |> which (fn ths =>  | 
|
78  | 
if member Thm.eq_thm_prop ths th then  | 
|
79  | 
let  | 
|
80  | 
val _ =  | 
|
| 
57859
 
29e728588163
more careful treatment of context visibility for rule declarations (see also 39d9c7f175e0, e639d91d9073) -- avoid duplicate warnings;
 
wenzelm 
parents: 
56491 
diff
changeset
 | 
81  | 
(case context of  | 
| 
 
29e728588163
more careful treatment of context visibility for rule declarations (see also 39d9c7f175e0, e639d91d9073) -- avoid duplicate warnings;
 
wenzelm 
parents: 
56491 
diff
changeset
 | 
82  | 
Context.Proof ctxt =>  | 
| 
 
29e728588163
more careful treatment of context visibility for rule declarations (see also 39d9c7f175e0, e639d91d9073) -- avoid duplicate warnings;
 
wenzelm 
parents: 
56491 
diff
changeset
 | 
83  | 
if Context_Position.is_really_visible ctxt then  | 
| 61268 | 84  | 
                warning ("Ignoring duplicate theorems:\n" ^ Thm.string_of_thm ctxt th)
 | 
| 
57859
 
29e728588163
more careful treatment of context visibility for rule declarations (see also 39d9c7f175e0, e639d91d9073) -- avoid duplicate warnings;
 
wenzelm 
parents: 
56491 
diff
changeset
 | 
85  | 
else ()  | 
| 
 
29e728588163
more careful treatment of context visibility for rule declarations (see also 39d9c7f175e0, e639d91d9073) -- avoid duplicate warnings;
 
wenzelm 
parents: 
56491 
diff
changeset
 | 
86  | 
| _ => ());  | 
| 55228 | 87  | 
in ths end  | 
88  | 
else sort_rules (th :: ths))));  | 
|
89  | 
||
90  | 
val add_safe = Context.proof_map o add_rule apfst;  | 
|
91  | 
val add_unsafe = Context.proof_map o add_rule apsnd;  | 
|
92  | 
||
93  | 
||
94  | 
(* attributes *)  | 
|
95  | 
||
96  | 
val safe_add = Thm.declaration_attribute (add_rule apfst);  | 
|
97  | 
val unsafe_add = Thm.declaration_attribute (add_rule apsnd);  | 
|
98  | 
||
99  | 
val _ = Theory.setup  | 
|
| 69593 | 100  | 
(Attrib.setup \<^binding>\<open>safe\<close> (Scan.succeed safe_add) "" #>  | 
101  | 
Attrib.setup \<^binding>\<open>unsafe\<close> (Scan.succeed unsafe_add) "");  | 
|
| 55228 | 102  | 
|
103  | 
||
104  | 
(* proof method syntax *)  | 
|
105  | 
||
106  | 
fun method tac =  | 
|
107  | 
Method.sections  | 
|
| 64556 | 108  | 
[Args.$$$ "add" -- Args.bang_colon >> K (Method.modifier safe_add \<^here>),  | 
109  | 
Args.$$$ "add" -- Args.colon >> K (Method.modifier unsafe_add \<^here>)]  | 
|
| 55228 | 110  | 
>> K (SIMPLE_METHOD' o tac);  | 
111  | 
||
112  | 
||
113  | 
(** tactics **)  | 
|
114  | 
||
| 69593 | 115  | 
val trace = Attrib.setup_config_bool \<^binding>\<open>cla_trace\<close> (K false);  | 
| 55228 | 116  | 
|
| 
7097
 
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
 
paulson 
parents: 
6054 
diff
changeset
 | 
117  | 
|
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
118  | 
(*Returns the list of all formulas in the sequent*)  | 
| 74302 | 119  | 
fun forms_of_seq \<^Const_>\<open>SeqO' for P u\<close> = P :: forms_of_seq u  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
120  | 
| forms_of_seq (H $ u) = forms_of_seq u  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
121  | 
| forms_of_seq _ = [];  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
122  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
123  | 
(*Tests whether two sequences (left or right sides) could be resolved.  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
124  | 
seqp is a premise (subgoal), seqc is a conclusion of an object-rule.  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
125  | 
Assumes each formula in seqc is surrounded by sequence variables  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
126  | 
-- checks that each concl formula looks like some subgoal formula.  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
127  | 
It SHOULD check order as well, using recursion rather than forall/exists*)  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
128  | 
fun could_res (seqp,seqc) =  | 
| 55228 | 129  | 
forall (fn Qc => exists (fn Qp => Term.could_unify (Qp,Qc))  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
130  | 
(forms_of_seq seqp))  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
131  | 
(forms_of_seq seqc);  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
132  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
133  | 
(*Tests whether two sequents or pairs of sequents could be resolved*)  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
134  | 
fun could_resolve_seq (prem,conc) =  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
135  | 
case (prem,conc) of  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
136  | 
(_ $ Abs(_,_,leftp) $ Abs(_,_,rightp),  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
137  | 
_ $ Abs(_,_,leftc) $ Abs(_,_,rightc)) =>  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
138  | 
could_res (leftp,leftc) andalso could_res (rightp,rightc)  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
139  | 
| (_ $ Abs(_,_,leftp) $ rightp,  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
140  | 
_ $ Abs(_,_,leftc) $ rightc) =>  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
141  | 
could_res (leftp,leftc) andalso Term.could_unify (rightp,rightc)  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
142  | 
| _ => false;  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
143  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
144  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
145  | 
(*Like filt_resolve_tac, using could_resolve_seq  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
146  | 
Much faster than resolve_tac when there are many rules.  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
147  | 
Resolve subgoal i using the rules, unless more than maxr are compatible. *)  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58048 
diff
changeset
 | 
148  | 
fun filseq_resolve_tac ctxt rules maxr = SUBGOAL(fn (prem,i) =>  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
149  | 
let val rls = filter_thms could_resolve_seq (maxr+1, prem, rules)  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
150  | 
in if length rls > maxr then no_tac  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
151  | 
else (*((rtac derelict 1 THEN rtac impl 1  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
152  | 
THEN (rtac identity 2 ORELSE rtac ll_mp 2)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
153  | 
THEN rtac context1 1)  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58048 
diff
changeset
 | 
154  | 
ORELSE *) resolve_tac ctxt rls i  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
155  | 
end);  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
156  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
157  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
158  | 
(*Predicate: does the rule have n premises? *)  | 
| 59582 | 159  | 
fun has_prems n rule = (Thm.nprems_of rule = n);  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
160  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
161  | 
(*Continuation-style tactical for resolution.  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
162  | 
The list of rules is partitioned into 0, 1, 2 premises.  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
163  | 
The resulting tactic, gtac, tries to resolve with rules.  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
164  | 
If successful, it recursively applies nextac to the new subgoals only.  | 
| 55228 | 165  | 
Else fails. (Treatment of goals due to Ph. de Groote)  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
166  | 
Bind (RESOLVE_THEN rules) to a variable: it preprocesses the rules. *)  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
167  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
168  | 
(*Takes rule lists separated in to 0, 1, 2, >2 premises.  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
169  | 
The abstraction over state prevents needless divergence in recursion.  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
170  | 
The 9999 should be a parameter, to delay treatment of flexible goals. *)  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
171  | 
|
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58048 
diff
changeset
 | 
172  | 
fun RESOLVE_THEN ctxt rules =  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
173  | 
let val [rls0,rls1,rls2] = partition_list has_prems 0 2 rules;  | 
| 3538 | 174  | 
fun tac nextac i state = state |>  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58048 
diff
changeset
 | 
175  | 
(filseq_resolve_tac ctxt rls0 9999 i  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
176  | 
ORELSE  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58048 
diff
changeset
 | 
177  | 
(DETERM(filseq_resolve_tac ctxt rls1 9999 i) THEN TRY(nextac i))  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
178  | 
ORELSE  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58048 
diff
changeset
 | 
179  | 
(DETERM(filseq_resolve_tac ctxt rls2 9999 i) THEN TRY(nextac(i+1))  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
180  | 
THEN TRY(nextac i)))  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
181  | 
in tac end;  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
182  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
183  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
184  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
185  | 
(*repeated resolution applied to the designated goal*)  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58048 
diff
changeset
 | 
186  | 
fun reresolve_tac ctxt rules =  | 
| 55228 | 187  | 
let  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58048 
diff
changeset
 | 
188  | 
val restac = RESOLVE_THEN ctxt rules; (*preprocessing done now*)  | 
| 55228 | 189  | 
fun gtac i = restac gtac i;  | 
190  | 
in gtac end;  | 
|
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
191  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
192  | 
(*tries the safe rules repeatedly before the unsafe rules. *)  | 
| 55228 | 193  | 
fun repeat_goal_tac ctxt =  | 
194  | 
let  | 
|
195  | 
val (safes, unsafes) = get_rules ctxt;  | 
|
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58048 
diff
changeset
 | 
196  | 
val restac = RESOLVE_THEN ctxt safes;  | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58048 
diff
changeset
 | 
197  | 
val lastrestac = RESOLVE_THEN ctxt unsafes;  | 
| 55228 | 198  | 
fun gtac i =  | 
199  | 
restac gtac i ORELSE  | 
|
| 56491 | 200  | 
(if Config.get ctxt trace then print_tac ctxt "" THEN lastrestac gtac i  | 
| 55228 | 201  | 
else lastrestac gtac i)  | 
202  | 
in gtac end;  | 
|
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
203  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
204  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
205  | 
(*Tries safe rules only*)  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58048 
diff
changeset
 | 
206  | 
fun safe_tac ctxt = reresolve_tac ctxt (#1 (get_rules ctxt));  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
207  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
208  | 
(*Tries a safe rule or else a unsafe rule. Single-step for tracing. *)  | 
| 55228 | 209  | 
fun step_tac ctxt =  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58048 
diff
changeset
 | 
210  | 
safe_tac ctxt ORELSE' filseq_resolve_tac ctxt (#2 (get_rules ctxt)) 9999;  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
211  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
212  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
213  | 
(* Tactic for reducing a goal, using Predicate Calculus rules.  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
214  | 
A decision procedure for Propositional Calculus, it is incomplete  | 
| 55228 | 215  | 
for Predicate-Calculus because of allL_thin and exR_thin.  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
216  | 
Fails if it can do nothing. *)  | 
| 55228 | 217  | 
fun pc_tac ctxt = SELECT_GOAL (DEPTH_SOLVE (repeat_goal_tac ctxt 1));  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
218  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
219  | 
|
| 55228 | 220  | 
(*The following two tactics are analogous to those provided by  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
221  | 
Provers/classical. In fact, pc_tac is usually FASTER than fast_tac!*)  | 
| 55228 | 222  | 
fun fast_tac ctxt =  | 
223  | 
SELECT_GOAL (DEPTH_SOLVE (step_tac ctxt 1));  | 
|
224  | 
||
225  | 
fun best_tac ctxt =  | 
|
226  | 
SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, size_of_thm) (step_tac ctxt 1));  | 
|
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
227  | 
|
| 55228 | 228  | 
val _ = Theory.setup  | 
| 69593 | 229  | 
(Method.setup \<^binding>\<open>safe\<close> (method safe_tac) "" #>  | 
230  | 
Method.setup \<^binding>\<open>step\<close> (method step_tac) "" #>  | 
|
231  | 
Method.setup \<^binding>\<open>pc\<close> (method pc_tac) "" #>  | 
|
232  | 
Method.setup \<^binding>\<open>fast\<close> (method fast_tac) "" #>  | 
|
233  | 
Method.setup \<^binding>\<open>best\<close> (method best_tac) "");  | 
|
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
234  | 
|
| 7122 | 235  | 
end;  | 
236  |