author | blanchet |
Fri, 14 Mar 2014 10:08:33 +0100 | |
changeset 56123 | a27859b0ef7d |
parent 55228 | 901a6696cdd8 |
child 56491 | a8ccf3d6a6e4 |
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 |
val extend = I; |
|
50 |
fun merge (Pack (safes, unsafes), Pack (safes', unsafes')) = |
|
51 |
Pack |
|
52 |
(sort_rules (Thm.merge_thms (safes, safes')), |
|
53 |
sort_rules (Thm.merge_thms (unsafes, unsafes'))); |
|
54 |
); |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
55 |
|
55228 | 56 |
val put_pack = Context.proof_map o Pack.put; |
57 |
val get_pack = Pack.get o Context.Proof; |
|
58 |
fun get_rules ctxt = let val Pack rules = get_pack ctxt in rules end; |
|
59 |
||
60 |
||
61 |
(* print pack *) |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
62 |
|
55228 | 63 |
fun pretty_pack ctxt = |
64 |
let val (safes, unsafes) = get_rules ctxt in |
|
65 |
Pretty.chunks |
|
66 |
[Pretty.big_list "safe rules:" (map (Display.pretty_thm ctxt) safes), |
|
67 |
Pretty.big_list "unsafe rules:" (map (Display.pretty_thm ctxt) unsafes)] |
|
68 |
end; |
|
7097
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents:
6054
diff
changeset
|
69 |
|
55228 | 70 |
val _ = |
71 |
Outer_Syntax.command @{command_spec "print_pack"} "print pack of classical rules" |
|
72 |
(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
|
73 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
74 |
|
55228 | 75 |
(* declare rules *) |
76 |
||
77 |
fun add_rule which th context = context |> Pack.map (fn Pack rules => |
|
78 |
Pack (rules |> which (fn ths => |
|
79 |
if member Thm.eq_thm_prop ths th then |
|
80 |
let |
|
81 |
val ctxt = Context.proof_of context; |
|
82 |
val _ = |
|
83 |
if Context_Position.is_visible ctxt then |
|
84 |
warning ("Ignoring duplicate theorems:\n" ^ Display.string_of_thm ctxt th) |
|
85 |
else (); |
|
86 |
in ths end |
|
87 |
else sort_rules (th :: ths)))); |
|
88 |
||
89 |
val add_safe = Context.proof_map o add_rule apfst; |
|
90 |
val add_unsafe = Context.proof_map o add_rule apsnd; |
|
91 |
||
92 |
||
93 |
(* attributes *) |
|
94 |
||
95 |
val safe_add = Thm.declaration_attribute (add_rule apfst); |
|
96 |
val unsafe_add = Thm.declaration_attribute (add_rule apsnd); |
|
97 |
||
98 |
val _ = Theory.setup |
|
99 |
(Attrib.setup @{binding safe} (Scan.succeed safe_add) "" #> |
|
100 |
Attrib.setup @{binding unsafe} (Scan.succeed unsafe_add) ""); |
|
101 |
||
102 |
||
103 |
(* proof method syntax *) |
|
104 |
||
105 |
fun method tac = |
|
106 |
Method.sections |
|
107 |
[Args.$$$ "add" -- Args.bang_colon >> K (I, safe_add), |
|
108 |
Args.$$$ "add" -- Args.colon >> K (I, unsafe_add)] |
|
109 |
>> K (SIMPLE_METHOD' o tac); |
|
110 |
||
111 |
||
112 |
(** tactics **) |
|
113 |
||
114 |
val trace = Attrib.setup_config_bool @{binding cla_trace} (K false); |
|
115 |
||
7097
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents:
6054
diff
changeset
|
116 |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
117 |
(*Returns the list of all formulas in the sequent*) |
38500 | 118 |
fun forms_of_seq (Const(@{const_name "SeqO'"},_) $ P $ u) = P :: forms_of_seq u |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
119 |
| forms_of_seq (H $ u) = forms_of_seq u |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
120 |
| forms_of_seq _ = []; |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
121 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
122 |
(*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
|
123 |
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
|
124 |
Assumes each formula in seqc is surrounded by sequence variables |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
125 |
-- checks that each concl formula looks like some subgoal formula. |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
126 |
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
|
127 |
fun could_res (seqp,seqc) = |
55228 | 128 |
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
|
129 |
(forms_of_seq seqp)) |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
130 |
(forms_of_seq seqc); |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
131 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
132 |
(*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
|
133 |
fun could_resolve_seq (prem,conc) = |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
134 |
case (prem,conc) of |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
135 |
(_ $ Abs(_,_,leftp) $ Abs(_,_,rightp), |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
136 |
_ $ Abs(_,_,leftc) $ Abs(_,_,rightc)) => |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32740
diff
changeset
|
137 |
could_res (leftp,leftc) andalso could_res (rightp,rightc) |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
138 |
| (_ $ Abs(_,_,leftp) $ rightp, |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
139 |
_ $ Abs(_,_,leftc) $ rightc) => |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32740
diff
changeset
|
140 |
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
|
141 |
| _ => false; |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
142 |
|
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 |
(*Like filt_resolve_tac, using could_resolve_seq |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
145 |
Much faster than resolve_tac when there are many rules. |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
146 |
Resolve subgoal i using the rules, unless more than maxr are compatible. *) |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
147 |
fun filseq_resolve_tac rules maxr = SUBGOAL(fn (prem,i) => |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
148 |
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
|
149 |
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
|
150 |
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
|
151 |
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
|
152 |
THEN rtac context1 1) |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32740
diff
changeset
|
153 |
ORELSE *) resolve_tac rls i |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
154 |
end); |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
155 |
|
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 |
(*Predicate: does the rule have n premises? *) |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
158 |
fun has_prems n rule = (nprems_of rule = n); |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
159 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
160 |
(*Continuation-style tactical for resolution. |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
161 |
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
|
162 |
The resulting tactic, gtac, tries to resolve with rules. |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
163 |
If successful, it recursively applies nextac to the new subgoals only. |
55228 | 164 |
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
|
165 |
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
|
166 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
167 |
(*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
|
168 |
The abstraction over state prevents needless divergence in recursion. |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
169 |
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
|
170 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
171 |
fun RESOLVE_THEN rules = |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
172 |
let val [rls0,rls1,rls2] = partition_list has_prems 0 2 rules; |
3538 | 173 |
fun tac nextac i state = state |> |
55228 | 174 |
(filseq_resolve_tac rls0 9999 i |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32740
diff
changeset
|
175 |
ORELSE |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32740
diff
changeset
|
176 |
(DETERM(filseq_resolve_tac rls1 9999 i) THEN TRY(nextac i)) |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32740
diff
changeset
|
177 |
ORELSE |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32740
diff
changeset
|
178 |
(DETERM(filseq_resolve_tac rls2 9999 i) THEN TRY(nextac(i+1)) |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32740
diff
changeset
|
179 |
THEN TRY(nextac i))) |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
180 |
in tac end; |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
181 |
|
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 |
(*repeated resolution applied to the designated goal*) |
55228 | 185 |
fun reresolve_tac rules = |
186 |
let |
|
187 |
val restac = RESOLVE_THEN rules; (*preprocessing done now*) |
|
188 |
fun gtac i = restac gtac i; |
|
189 |
in gtac end; |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
190 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
191 |
(*tries the safe rules repeatedly before the unsafe rules. *) |
55228 | 192 |
fun repeat_goal_tac ctxt = |
193 |
let |
|
194 |
val (safes, unsafes) = get_rules ctxt; |
|
195 |
val restac = RESOLVE_THEN safes; |
|
196 |
val lastrestac = RESOLVE_THEN unsafes; |
|
197 |
fun gtac i = |
|
198 |
restac gtac i ORELSE |
|
199 |
(if Config.get ctxt trace then print_tac "" THEN lastrestac gtac i |
|
200 |
else lastrestac gtac i) |
|
201 |
in gtac end; |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
202 |
|
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 |
(*Tries safe rules only*) |
55228 | 205 |
fun safe_tac ctxt = reresolve_tac (#1 (get_rules ctxt)); |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
206 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
207 |
(*Tries a safe rule or else a unsafe rule. Single-step for tracing. *) |
55228 | 208 |
fun step_tac ctxt = |
209 |
safe_tac ctxt ORELSE' filseq_resolve_tac (#2 (get_rules ctxt)) 9999; |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
210 |
|
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 |
(* Tactic for reducing a goal, using Predicate Calculus rules. |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
213 |
A decision procedure for Propositional Calculus, it is incomplete |
55228 | 214 |
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
|
215 |
Fails if it can do nothing. *) |
55228 | 216 |
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
|
217 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
218 |
|
55228 | 219 |
(*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
|
220 |
Provers/classical. In fact, pc_tac is usually FASTER than fast_tac!*) |
55228 | 221 |
fun fast_tac ctxt = |
222 |
SELECT_GOAL (DEPTH_SOLVE (step_tac ctxt 1)); |
|
223 |
||
224 |
fun best_tac ctxt = |
|
225 |
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
|
226 |
|
55228 | 227 |
val _ = Theory.setup |
228 |
(Method.setup @{binding safe} (method safe_tac) "" #> |
|
229 |
Method.setup @{binding step} (method step_tac) "" #> |
|
230 |
Method.setup @{binding pc} (method pc_tac) "" #> |
|
231 |
Method.setup @{binding fast} (method fast_tac) "" #> |
|
232 |
Method.setup @{binding best} (method best_tac) ""); |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
233 |
|
7122 | 234 |
end; |
235 |