paulson@7122: (* Title: Sequents/prover paulson@2073: ID: $Id$ paulson@2073: Author: Lawrence C Paulson, Cambridge University Computer Laboratory paulson@2073: Copyright 1992 University of Cambridge paulson@7097: paulson@7097: Simple classical reasoner for the sequent calculus, based on "theorem packs" paulson@2073: *) paulson@2073: paulson@2073: paulson@7097: (*Higher precedence than := facilitates use of references*) paulson@7097: infix 4 add_safes add_unsafes; paulson@2073: paulson@7122: structure Cla = paulson@7122: paulson@7122: struct paulson@7122: paulson@2073: datatype pack = Pack of thm list * thm list; paulson@2073: paulson@7122: val trace = ref false; paulson@7122: paulson@2073: (*A theorem pack has the form (safe rules, unsafe rules) paulson@2073: An unsafe rule is incomplete or introduces variables in subgoals, paulson@2073: and is tried only when the safe rules are not applicable. *) paulson@2073: paulson@2073: fun less (rl1,rl2) = (nprems_of rl1) < (nprems_of rl2); paulson@2073: paulson@2073: val empty_pack = Pack([],[]); paulson@2073: paulson@7097: fun warn_duplicates [] = [] paulson@7097: | warn_duplicates dups = wenzelm@7150: (warning (cat_lines ("Ignoring duplicate theorems:" :: map string_of_thm dups)); paulson@7097: dups); paulson@2073: paulson@2073: fun (Pack(safes,unsafes)) add_safes ths = wenzelm@13105: let val dups = warn_duplicates (gen_inter Drule.eq_thm_prop (ths,safes)) wenzelm@13105: val ths' = gen_rems Drule.eq_thm_prop (ths,dups) paulson@7097: in paulson@7097: Pack(sort (make_ord less) (ths'@safes), unsafes) paulson@7097: end; paulson@2073: paulson@2073: fun (Pack(safes,unsafes)) add_unsafes ths = wenzelm@13105: let val dups = warn_duplicates (gen_inter Drule.eq_thm_prop (ths,unsafes)) wenzelm@13105: val ths' = gen_rems Drule.eq_thm_prop (ths,dups) paulson@7097: in paulson@7097: Pack(safes, sort (make_ord less) (ths'@unsafes)) paulson@7097: end; paulson@7097: paulson@7097: fun merge_pack (Pack(safes,unsafes), Pack(safes',unsafes')) = paulson@7097: Pack(sort (make_ord less) (safes@safes'), paulson@7097: sort (make_ord less) (unsafes@unsafes')); paulson@2073: paulson@2073: paulson@7097: fun print_pack (Pack(safes,unsafes)) = paulson@7097: (writeln "Safe rules:"; print_thms safes; paulson@7097: writeln "Unsafe rules:"; print_thms unsafes); paulson@7097: paulson@2073: (*Returns the list of all formulas in the sequent*) paulson@7097: fun forms_of_seq (Const("SeqO'",_) $ P $ u) = P :: forms_of_seq u paulson@2073: | forms_of_seq (H $ u) = forms_of_seq u paulson@2073: | forms_of_seq _ = []; paulson@2073: paulson@2073: (*Tests whether two sequences (left or right sides) could be resolved. paulson@2073: seqp is a premise (subgoal), seqc is a conclusion of an object-rule. paulson@2073: Assumes each formula in seqc is surrounded by sequence variables paulson@2073: -- checks that each concl formula looks like some subgoal formula. paulson@2073: It SHOULD check order as well, using recursion rather than forall/exists*) paulson@2073: fun could_res (seqp,seqc) = paulson@2073: forall (fn Qc => exists (fn Qp => could_unify (Qp,Qc)) paulson@2073: (forms_of_seq seqp)) paulson@2073: (forms_of_seq seqc); paulson@2073: paulson@2073: paulson@2073: (*Tests whether two sequents or pairs of sequents could be resolved*) paulson@2073: fun could_resolve_seq (prem,conc) = paulson@2073: case (prem,conc) of paulson@2073: (_ $ Abs(_,_,leftp) $ Abs(_,_,rightp), paulson@2073: _ $ Abs(_,_,leftc) $ Abs(_,_,rightc)) => paulson@2073: could_res (leftp,leftc) andalso could_res (rightp,rightc) paulson@2073: | (_ $ Abs(_,_,leftp) $ rightp, paulson@2073: _ $ Abs(_,_,leftc) $ rightc) => paulson@2073: could_res (leftp,leftc) andalso could_unify (rightp,rightc) paulson@2073: | _ => false; paulson@2073: paulson@2073: paulson@2073: (*Like filt_resolve_tac, using could_resolve_seq paulson@2073: Much faster than resolve_tac when there are many rules. paulson@2073: Resolve subgoal i using the rules, unless more than maxr are compatible. *) paulson@2073: fun filseq_resolve_tac rules maxr = SUBGOAL(fn (prem,i) => paulson@2073: let val rls = filter_thms could_resolve_seq (maxr+1, prem, rules) paulson@2073: in if length rls > maxr then no_tac paulson@2073: else (*((rtac derelict 1 THEN rtac impl 1 paulson@2073: THEN (rtac identity 2 ORELSE rtac ll_mp 2) paulson@2073: THEN rtac context1 1) paulson@2073: ORELSE *) resolve_tac rls i paulson@2073: end); paulson@2073: paulson@2073: paulson@2073: (*Predicate: does the rule have n premises? *) paulson@2073: fun has_prems n rule = (nprems_of rule = n); paulson@2073: paulson@2073: (*Continuation-style tactical for resolution. paulson@2073: The list of rules is partitioned into 0, 1, 2 premises. paulson@2073: The resulting tactic, gtac, tries to resolve with rules. paulson@2073: If successful, it recursively applies nextac to the new subgoals only. paulson@2073: Else fails. (Treatment of goals due to Ph. de Groote) paulson@2073: Bind (RESOLVE_THEN rules) to a variable: it preprocesses the rules. *) paulson@2073: paulson@2073: (*Takes rule lists separated in to 0, 1, 2, >2 premises. paulson@2073: The abstraction over state prevents needless divergence in recursion. paulson@2073: The 9999 should be a parameter, to delay treatment of flexible goals. *) paulson@2073: paulson@2073: fun RESOLVE_THEN rules = paulson@2073: let val [rls0,rls1,rls2] = partition_list has_prems 0 2 rules; paulson@3538: fun tac nextac i state = state |> paulson@3538: (filseq_resolve_tac rls0 9999 i paulson@3538: ORELSE paulson@3538: (DETERM(filseq_resolve_tac rls1 9999 i) THEN TRY(nextac i)) paulson@3538: ORELSE paulson@3538: (DETERM(filseq_resolve_tac rls2 9999 i) THEN TRY(nextac(i+1)) paulson@3538: THEN TRY(nextac i))) paulson@2073: in tac end; paulson@2073: paulson@2073: paulson@2073: paulson@2073: (*repeated resolution applied to the designated goal*) paulson@2073: fun reresolve_tac rules = paulson@2073: let val restac = RESOLVE_THEN rules; (*preprocessing done now*) paulson@2073: fun gtac i = restac gtac i paulson@2073: in gtac end; paulson@2073: paulson@2073: (*tries the safe rules repeatedly before the unsafe rules. *) paulson@2073: fun repeat_goal_tac (Pack(safes,unsafes)) = paulson@2073: let val restac = RESOLVE_THEN safes paulson@2073: and lastrestac = RESOLVE_THEN unsafes; paulson@6054: fun gtac i = restac gtac i paulson@7122: ORELSE (if !trace then paulson@7122: (print_tac "" THEN lastrestac gtac i) paulson@7122: else lastrestac gtac i) paulson@2073: in gtac end; paulson@2073: paulson@2073: paulson@2073: (*Tries safe rules only*) paulson@7097: fun safe_tac (Pack(safes,unsafes)) = reresolve_tac safes; paulson@7097: paulson@7097: val safe_goal_tac = safe_tac; (*backwards compatibility*) paulson@2073: paulson@2073: (*Tries a safe rule or else a unsafe rule. Single-step for tracing. *) paulson@7122: fun step_tac (pack as Pack(safes,unsafes)) = paulson@7122: safe_tac pack ORELSE' paulson@2073: filseq_resolve_tac unsafes 9999; paulson@2073: paulson@2073: paulson@2073: (* Tactic for reducing a goal, using Predicate Calculus rules. paulson@2073: A decision procedure for Propositional Calculus, it is incomplete paulson@2073: for Predicate-Calculus because of allL_thin and exR_thin. paulson@2073: Fails if it can do nothing. *) paulson@7122: fun pc_tac pack = SELECT_GOAL (DEPTH_SOLVE (repeat_goal_tac pack 1)); paulson@2073: paulson@2073: paulson@2073: (*The following two tactics are analogous to those provided by paulson@2073: Provers/classical. In fact, pc_tac is usually FASTER than fast_tac!*) paulson@7122: fun fast_tac pack = paulson@7122: SELECT_GOAL (DEPTH_SOLVE (step_tac pack 1)); paulson@2073: paulson@7122: fun best_tac pack = paulson@2073: SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, size_of_thm) paulson@7122: (step_tac pack 1)); paulson@2073: paulson@2073: paulson@2073: paulson@7097: structure ProverArgs = paulson@7097: struct paulson@7097: val name = "Sequents/prover"; paulson@7097: type T = pack ref; paulson@7097: val empty = ref empty_pack paulson@7097: fun copy (ref pack) = ref pack; paulson@7097: val prep_ext = copy; paulson@7097: fun merge (ref pack1, ref pack2) = ref (merge_pack (pack1, pack2)); paulson@7097: fun print _ (ref pack) = print_pack pack; paulson@7097: end; paulson@2073: paulson@7097: structure ProverData = TheoryDataFun(ProverArgs); paulson@2073: paulson@7097: val prover_setup = [ProverData.init]; paulson@2073: paulson@7122: val print_pack = ProverData.print; paulson@7122: val pack_ref_of_sg = ProverData.get_sg; paulson@7122: val pack_ref_of = ProverData.get; paulson@2073: paulson@7122: (* access global pack *) paulson@2073: paulson@7122: val pack_of_sg = ! o pack_ref_of_sg; paulson@7122: val pack_of = pack_of_sg o sign_of; paulson@2073: paulson@7122: val pack = pack_of o Context.the_context; paulson@7122: val pack_ref = pack_ref_of_sg o sign_of o Context.the_context; paulson@2073: paulson@7097: paulson@7122: (* change global pack *) paulson@7122: paulson@7122: fun change_pack f x = pack_ref () := (f (pack (), x)); paulson@2073: paulson@7122: val Add_safes = change_pack (op add_safes); paulson@7122: val Add_unsafes = change_pack (op add_unsafes); paulson@7122: paulson@2073: paulson@7122: fun Fast_tac st = fast_tac (pack()) st; paulson@7122: fun Step_tac st = step_tac (pack()) st; paulson@7122: fun Safe_tac st = safe_tac (pack()) st; paulson@7122: fun Pc_tac st = pc_tac (pack()) st; paulson@2073: paulson@7122: end; paulson@7122: paulson@7122: paulson@7122: open Cla;