author | wenzelm |
Sun, 16 Jul 2023 19:38:12 +0200 | |
changeset 78371 | 928e031b7c52 |
parent 74302 | 6bc96f31cafd |
permissions | -rw-r--r-- |
24584 | 1 |
(* Title: Sequents/modal.ML |
7096 | 2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
3 |
Copyright 1992 University of Cambridge |
|
4 |
||
29269
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
wenzelm
parents:
24584
diff
changeset
|
5 |
Simple modal reasoner. |
7096 | 6 |
*) |
7 |
||
8 |
signature MODAL_PROVER_RULE = |
|
9 |
sig |
|
59164 | 10 |
val rewrite_rls : thm list |
11 |
val safe_rls : thm list |
|
12 |
val unsafe_rls : thm list |
|
13 |
val bound_rls : thm list |
|
14 |
val aside_rls : thm list |
|
7096 | 15 |
end; |
16 |
||
59164 | 17 |
signature MODAL_PROVER = |
7096 | 18 |
sig |
59164 | 19 |
val rule_tac : Proof.context -> thm list -> int ->tactic |
20 |
val step_tac : Proof.context -> int -> tactic |
|
21 |
val solven_tac : Proof.context -> int -> int -> tactic |
|
22 |
val solve_tac : Proof.context -> int -> tactic |
|
7096 | 23 |
end; |
24 |
||
59164 | 25 |
functor Modal_ProverFun (Modal_Rule: MODAL_PROVER_RULE) : MODAL_PROVER = |
7096 | 26 |
struct |
27 |
||
28 |
(*Returns the list of all formulas in the sequent*) |
|
74302 | 29 |
fun forms_of_seq \<^Const_>\<open>SeqO' for P u\<close> = P :: forms_of_seq u |
7096 | 30 |
| forms_of_seq (H $ u) = forms_of_seq u |
31 |
| forms_of_seq _ = []; |
|
32 |
||
33 |
(*Tests whether two sequences (left or right sides) could be resolved. |
|
34 |
seqp is a premise (subgoal), seqc is a conclusion of an object-rule. |
|
35 |
Assumes each formula in seqc is surrounded by sequence variables |
|
36 |
-- checks that each concl formula looks like some subgoal formula.*) |
|
37 |
fun could_res (seqp,seqc) = |
|
59164 | 38 |
forall (fn Qc => exists (fn Qp => Term.could_unify (Qp,Qc)) |
7096 | 39 |
(forms_of_seq seqp)) |
40 |
(forms_of_seq seqc); |
|
41 |
||
42 |
(*Tests whether two sequents G|-H could be resolved, comparing each side.*) |
|
43 |
fun could_resolve_seq (prem,conc) = |
|
44 |
case (prem,conc) of |
|
45 |
(_ $ Abs(_,_,leftp) $ Abs(_,_,rightp), |
|
46 |
_ $ Abs(_,_,leftc) $ Abs(_,_,rightc)) => |
|
47 |
could_res (leftp,leftc) andalso could_res (rightp,rightc) |
|
48 |
| _ => false; |
|
49 |
||
50 |
(*Like filt_resolve_tac, using could_resolve_seq |
|
51 |
Much faster than resolve_tac when there are many rules. |
|
52 |
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:
59164
diff
changeset
|
53 |
fun filseq_resolve_tac ctxt rules maxr = SUBGOAL(fn (prem,i) => |
7096 | 54 |
let val rls = filter_thms could_resolve_seq (maxr+1, prem, rules) |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59164
diff
changeset
|
55 |
in if length rls > maxr then no_tac else resolve_tac ctxt rls i |
7096 | 56 |
end); |
57 |
||
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59164
diff
changeset
|
58 |
fun fresolve_tac ctxt rls n = filseq_resolve_tac ctxt rls 999 n; |
7096 | 59 |
|
60 |
(* NB No back tracking possible with aside rules *) |
|
61 |
||
59164 | 62 |
val aside_net = Tactic.build_net Modal_Rule.aside_rls; |
63 |
fun aside_tac ctxt n = DETERM (REPEAT (filt_resolve_from_net_tac ctxt 999 aside_net n)); |
|
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59164
diff
changeset
|
64 |
fun rule_tac ctxt rls n = fresolve_tac ctxt rls n THEN aside_tac ctxt n; |
7096 | 65 |
|
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59164
diff
changeset
|
66 |
fun fres_safe_tac ctxt = fresolve_tac ctxt Modal_Rule.safe_rls; |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59164
diff
changeset
|
67 |
fun fres_unsafe_tac ctxt = fresolve_tac ctxt Modal_Rule.unsafe_rls THEN' aside_tac ctxt; |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59164
diff
changeset
|
68 |
fun fres_bound_tac ctxt = fresolve_tac ctxt Modal_Rule.bound_rls; |
7096 | 69 |
|
70 |
fun UPTOGOAL n tf = let fun tac i = if i<n then all_tac |
|
71 |
else tf(i) THEN tac(i-1) |
|
59582 | 72 |
in fn st => tac (Thm.nprems_of st) st end; |
7096 | 73 |
|
74 |
(* Depth first search bounded by d *) |
|
59164 | 75 |
fun solven_tac ctxt d n st = st |> |
76 |
(if d < 0 then no_tac |
|
59582 | 77 |
else if Thm.nprems_of st = 0 then all_tac |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59164
diff
changeset
|
78 |
else (DETERM(fres_safe_tac ctxt n) THEN UPTOGOAL n (solven_tac ctxt d)) ORELSE |
59164 | 79 |
((fres_unsafe_tac ctxt n THEN UPTOGOAL n (solven_tac ctxt d)) APPEND |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59164
diff
changeset
|
80 |
(fres_bound_tac ctxt n THEN UPTOGOAL n (solven_tac ctxt (d - 1))))); |
7096 | 81 |
|
59164 | 82 |
fun solve_tac ctxt d = |
83 |
rewrite_goals_tac ctxt Modal_Rule.rewrite_rls THEN solven_tac ctxt d 1; |
|
7096 | 84 |
|
59164 | 85 |
fun step_tac ctxt n = |
86 |
COND (has_fewer_prems 1) all_tac |
|
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59164
diff
changeset
|
87 |
(DETERM(fres_safe_tac ctxt n) ORELSE |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59164
diff
changeset
|
88 |
(fres_unsafe_tac ctxt n APPEND fres_bound_tac ctxt n)); |
7096 | 89 |
|
90 |
end; |