| author | huffman | 
| Mon, 08 Mar 2010 12:43:44 -0800 | |
| changeset 35660 | 8169419cd824 | 
| parent 33049 | c38f02fdf35d | 
| child 38500 | d5477ee35820 | 
| permissions | -rw-r--r-- | 
| 29269 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
26928diff
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: 
6054diff
changeset | 4 | |
| 29269 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
26928diff
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 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 8 | |
| 7097 
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
 paulson parents: 
6054diff
changeset | 9 | (*Higher precedence than := facilitates use of references*) | 
| 
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
 paulson parents: 
6054diff
changeset | 10 | infix 4 add_safes add_unsafes; | 
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 11 | |
| 7122 | 12 | structure Cla = | 
| 13 | struct | |
| 14 | ||
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 15 | datatype pack = Pack of thm list * thm list; | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 16 | |
| 32740 | 17 | val trace = Unsynchronized.ref false; | 
| 7122 | 18 | |
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 19 | (*A theorem pack has the form (safe rules, unsafe rules) | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 20 | An unsafe rule is incomplete or introduces variables in subgoals, | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 21 | 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 | 22 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 23 | fun less (rl1,rl2) = (nprems_of rl1) < (nprems_of rl2); | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 24 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 25 | val empty_pack = Pack([],[]); | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 26 | |
| 7097 
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
 paulson parents: 
6054diff
changeset | 27 | fun warn_duplicates [] = [] | 
| 
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
 paulson parents: 
6054diff
changeset | 28 | | warn_duplicates dups = | 
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
29269diff
changeset | 29 |       (warning (cat_lines ("Ignoring duplicate theorems:" ::
 | 
| 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
29269diff
changeset | 30 | map Display.string_of_thm_without_context dups)); | 
| 7097 
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
 paulson parents: 
6054diff
changeset | 31 | dups); | 
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 32 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 33 | fun (Pack(safes,unsafes)) add_safes ths = | 
| 33049 
c38f02fdf35d
curried inter as canonical list operation (beware of argument order)
 haftmann parents: 
33038diff
changeset | 34 | let val dups = warn_duplicates (inter Thm.eq_thm_prop ths safes) | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 35 | val ths' = subtract Thm.eq_thm_prop dups ths | 
| 7097 
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
 paulson parents: 
6054diff
changeset | 36 | in | 
| 
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
 paulson parents: 
6054diff
changeset | 37 | Pack(sort (make_ord less) (ths'@safes), unsafes) | 
| 
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
 paulson parents: 
6054diff
changeset | 38 | end; | 
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 39 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 40 | fun (Pack(safes,unsafes)) add_unsafes ths = | 
| 33049 
c38f02fdf35d
curried inter as canonical list operation (beware of argument order)
 haftmann parents: 
33038diff
changeset | 41 | let val dups = warn_duplicates (inter Thm.eq_thm_prop unsafes ths) | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 42 | val ths' = subtract Thm.eq_thm_prop dups ths | 
| 7097 
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
 paulson parents: 
6054diff
changeset | 43 | in | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 44 | Pack(safes, sort (make_ord less) (ths'@unsafes)) | 
| 7097 
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
 paulson parents: 
6054diff
changeset | 45 | end; | 
| 
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
 paulson parents: 
6054diff
changeset | 46 | |
| 
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
 paulson parents: 
6054diff
changeset | 47 | fun merge_pack (Pack(safes,unsafes), Pack(safes',unsafes')) = | 
| 
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
 paulson parents: 
6054diff
changeset | 48 | Pack(sort (make_ord less) (safes@safes'), | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 49 | sort (make_ord less) (unsafes@unsafes')); | 
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 50 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 51 | |
| 7097 
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
 paulson parents: 
6054diff
changeset | 52 | fun print_pack (Pack(safes,unsafes)) = | 
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
29269diff
changeset | 53 | writeln (cat_lines | 
| 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
29269diff
changeset | 54 | (["Safe rules:"] @ map Display.string_of_thm_without_context safes @ | 
| 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
29269diff
changeset | 55 | ["Unsafe rules:"] @ map Display.string_of_thm_without_context unsafes)); | 
| 7097 
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
 paulson parents: 
6054diff
changeset | 56 | |
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 57 | (*Returns the list of all formulas in the sequent*) | 
| 7097 
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
 paulson parents: 
6054diff
changeset | 58 | fun forms_of_seq (Const("SeqO'",_) $ P $ u) = P :: forms_of_seq u
 | 
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 59 | | forms_of_seq (H $ u) = forms_of_seq u | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 60 | | forms_of_seq _ = []; | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 61 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 62 | (*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 | 63 | 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 | 64 | Assumes each formula in seqc is surrounded by sequence variables | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 65 | -- checks that each concl formula looks like some subgoal formula. | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 66 | 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 | 67 | fun could_res (seqp,seqc) = | 
| 29269 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
26928diff
changeset | 68 | 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 | 69 | (forms_of_seq seqp)) | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 70 | (forms_of_seq seqc); | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 71 | |
| 
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 | (*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 | 74 | fun could_resolve_seq (prem,conc) = | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 75 | case (prem,conc) of | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 76 | (_ $ Abs(_,_,leftp) $ Abs(_,_,rightp), | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 77 | _ $ Abs(_,_,leftc) $ Abs(_,_,rightc)) => | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 78 | could_res (leftp,leftc) andalso could_res (rightp,rightc) | 
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 79 | | (_ $ Abs(_,_,leftp) $ rightp, | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 80 | _ $ Abs(_,_,leftc) $ rightc) => | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 81 | 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 | 82 | | _ => false; | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 83 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 84 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 85 | (*Like filt_resolve_tac, using could_resolve_seq | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 86 | Much faster than resolve_tac when there are many rules. | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 87 | 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 | 88 | fun filseq_resolve_tac rules maxr = SUBGOAL(fn (prem,i) => | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 89 | 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 | 90 | in if length rls > maxr then no_tac | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 91 | else (*((rtac derelict 1 THEN rtac impl 1 | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 92 | THEN (rtac identity 2 ORELSE rtac ll_mp 2) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 93 | THEN rtac context1 1) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 94 | ORELSE *) resolve_tac rls i | 
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 95 | end); | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 96 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 97 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 98 | (*Predicate: does the rule have n premises? *) | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 99 | fun has_prems n rule = (nprems_of rule = n); | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 100 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 101 | (*Continuation-style tactical for resolution. | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 102 | 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 | 103 | The resulting tactic, gtac, tries to resolve with rules. | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 104 | If successful, it recursively applies nextac to the new subgoals only. | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 105 | Else fails. (Treatment of goals due to Ph. de Groote) | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 106 | 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 | 107 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 108 | (*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 | 109 | The abstraction over state prevents needless divergence in recursion. | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 110 | 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 | 111 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 112 | fun RESOLVE_THEN rules = | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 113 | let val [rls0,rls1,rls2] = partition_list has_prems 0 2 rules; | 
| 3538 | 114 | fun tac nextac i state = state |> | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 115 | (filseq_resolve_tac rls0 9999 i | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 116 | ORELSE | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 117 | (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: 
32740diff
changeset | 118 | ORELSE | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 119 | (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: 
32740diff
changeset | 120 | THEN TRY(nextac i))) | 
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 121 | in tac end; | 
| 
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 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 124 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 125 | (*repeated resolution applied to the designated goal*) | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 126 | fun reresolve_tac rules = | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 127 | let val restac = RESOLVE_THEN rules; (*preprocessing done now*) | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 128 | fun gtac i = restac gtac i | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 129 | in gtac end; | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 130 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 131 | (*tries the safe rules repeatedly before the unsafe rules. *) | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 132 | fun repeat_goal_tac (Pack(safes,unsafes)) = | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 133 | let val restac = RESOLVE_THEN safes | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 134 | and lastrestac = RESOLVE_THEN unsafes; | 
| 6054 | 135 | fun gtac i = restac gtac i | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 136 | ORELSE (if !trace then | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 137 | (print_tac "" THEN lastrestac gtac i) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 138 | else lastrestac gtac i) | 
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 139 | in gtac end; | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 140 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 141 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 142 | (*Tries safe rules only*) | 
| 7097 
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
 paulson parents: 
6054diff
changeset | 143 | fun safe_tac (Pack(safes,unsafes)) = reresolve_tac safes; | 
| 
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
 paulson parents: 
6054diff
changeset | 144 | |
| 
5ab37ed3d53c
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
 paulson parents: 
6054diff
changeset | 145 | val safe_goal_tac = safe_tac; (*backwards compatibility*) | 
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 146 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 147 | (*Tries a safe rule or else a unsafe rule. Single-step for tracing. *) | 
| 7122 | 148 | fun step_tac (pack as Pack(safes,unsafes)) = | 
| 149 | safe_tac pack ORELSE' | |
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 150 | filseq_resolve_tac unsafes 9999; | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 151 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 152 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 153 | (* Tactic for reducing a goal, using Predicate Calculus rules. | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 154 | A decision procedure for Propositional Calculus, it is incomplete | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 155 | for Predicate-Calculus because of allL_thin and exR_thin. | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 156 | Fails if it can do nothing. *) | 
| 7122 | 157 | fun pc_tac pack = SELECT_GOAL (DEPTH_SOLVE (repeat_goal_tac pack 1)); | 
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 158 | |
| 
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 | (*The following two tactics are analogous to those provided by | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 161 | Provers/classical. In fact, pc_tac is usually FASTER than fast_tac!*) | 
| 7122 | 162 | fun fast_tac pack = | 
| 163 | SELECT_GOAL (DEPTH_SOLVE (step_tac pack 1)); | |
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 164 | |
| 7122 | 165 | fun best_tac pack = | 
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 166 | SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, size_of_thm) | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 167 | (step_tac pack 1)); | 
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 168 | |
| 7122 | 169 | end; | 
| 170 | ||
| 171 | ||
| 172 | open Cla; |