src/Sequents/prover.ML
author wenzelm
Wed, 31 Dec 2008 15:30:10 +0100
changeset 29269 5c25a2012975
parent 26928 ca87aff1ad2d
child 32091 30e2ffbba718
permissions -rw-r--r--
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML); tuned signature of structure Term;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
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: 6054
diff changeset
     9
(*Higher precedence than := facilitates use of references*)
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff 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
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
    12
structure Cla =
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
    13
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
    14
struct
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
    15
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    16
datatype pack = Pack of thm list * thm list;
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    17
7122
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
    18
val trace = ref false;
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
    19
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    20
(*A theorem pack has the form  (safe rules, unsafe rules)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    21
  An unsafe rule is incomplete or introduces variables in subgoals,
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    22
  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
    23
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    24
fun less (rl1,rl2) = (nprems_of rl1) < (nprems_of rl2);
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    25
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    26
val empty_pack = Pack([],[]);
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    27
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    28
fun warn_duplicates [] = []
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    29
  | warn_duplicates dups =
26928
ca87aff1ad2d structure Display: less pervasive operations;
wenzelm
parents: 22360
diff changeset
    30
      (warning (cat_lines ("Ignoring duplicate theorems:" :: map Display.string_of_thm dups));
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff 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   = 
22360
26ead7ed4f4b moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents: 21428
diff changeset
    34
    let val dups = warn_duplicates (gen_inter Thm.eq_thm_prop (ths,safes))
26ead7ed4f4b moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents: 21428
diff 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: 6054
diff changeset
    36
    in
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff 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: 6054
diff 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 = 
22360
26ead7ed4f4b moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents: 21428
diff changeset
    41
    let val dups = warn_duplicates (gen_inter Thm.eq_thm_prop (ths,unsafes))
26ead7ed4f4b moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents: 21428
diff 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: 6054
diff changeset
    43
    in
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    44
	Pack(safes, sort (make_ord less) (ths'@unsafes))
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    45
    end;
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    46
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff 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: 6054
diff changeset
    48
        Pack(sort (make_ord less) (safes@safes'), 
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff 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: 6054
diff changeset
    52
fun print_pack (Pack(safes,unsafes)) =
26928
ca87aff1ad2d structure Display: less pervasive operations;
wenzelm
parents: 22360
diff changeset
    53
    (writeln "Safe rules:";  Display.print_thms safes;
ca87aff1ad2d structure Display: less pervasive operations;
wenzelm
parents: 22360
diff changeset
    54
     writeln "Unsafe rules:"; Display.print_thms unsafes);
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    55
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    56
(*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: 6054
diff changeset
    57
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
    58
  | forms_of_seq (H $ u) = forms_of_seq u
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    59
  | forms_of_seq _ = [];
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    60
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    61
(*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
    62
  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
    63
  Assumes each formula in seqc is surrounded by sequence variables
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    64
  -- checks that each concl formula looks like some subgoal formula.
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    65
  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
    66
fun could_res (seqp,seqc) =
29269
5c25a2012975 moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
wenzelm
parents: 26928
diff changeset
    67
      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
    68
                              (forms_of_seq seqp))
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    69
             (forms_of_seq seqc);
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    70
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
(*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
    73
fun could_resolve_seq (prem,conc) =
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    74
  case (prem,conc) of
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    75
      (_ $ Abs(_,_,leftp) $ Abs(_,_,rightp),
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    76
       _ $ Abs(_,_,leftc) $ Abs(_,_,rightc)) =>
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    77
	  could_res (leftp,leftc) andalso could_res (rightp,rightc)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    78
    | (_ $ Abs(_,_,leftp) $ rightp,
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    79
       _ $ Abs(_,_,leftc) $ rightc) =>
29269
5c25a2012975 moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
wenzelm
parents: 26928
diff changeset
    80
	  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
    81
    | _ => false;
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    82
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
(*Like filt_resolve_tac, using could_resolve_seq
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    85
  Much faster than resolve_tac when there are many rules.
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    86
  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
    87
fun filseq_resolve_tac rules maxr = SUBGOAL(fn (prem,i) =>
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    88
  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
    89
  in  if length rls > maxr  then  no_tac
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    90
	  else (*((rtac derelict 1 THEN rtac impl 1
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    91
		 THEN (rtac identity 2 ORELSE rtac ll_mp 2)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    92
		 THEN rtac context1 1)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    93
		 ORELSE *) resolve_tac rls i
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    94
  end);
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    95
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
(*Predicate: does the rule have n premises? *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    98
fun has_prems n rule =  (nprems_of rule = n);
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    99
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   100
(*Continuation-style tactical for resolution.
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   101
  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
   102
  The resulting tactic, gtac, tries to resolve with rules.
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   103
  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
   104
  Else fails.  (Treatment of goals due to Ph. de Groote) 
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   105
  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
   106
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   107
(*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
   108
  The abstraction over state prevents needless divergence in recursion.
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   109
  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
   110
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   111
fun RESOLVE_THEN rules =
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   112
  let val [rls0,rls1,rls2] = partition_list has_prems 0 2 rules;
3538
ed9de44032e0 Removal of the tactical STATE
paulson
parents: 2073
diff changeset
   113
      fun tac nextac i state = state |>
ed9de44032e0 Removal of the tactical STATE
paulson
parents: 2073
diff changeset
   114
	     (filseq_resolve_tac rls0 9999 i 
ed9de44032e0 Removal of the tactical STATE
paulson
parents: 2073
diff changeset
   115
	      ORELSE
ed9de44032e0 Removal of the tactical STATE
paulson
parents: 2073
diff changeset
   116
	      (DETERM(filseq_resolve_tac rls1 9999 i) THEN  TRY(nextac i))
ed9de44032e0 Removal of the tactical STATE
paulson
parents: 2073
diff changeset
   117
	      ORELSE
ed9de44032e0 Removal of the tactical STATE
paulson
parents: 2073
diff changeset
   118
	      (DETERM(filseq_resolve_tac rls2 9999 i) THEN  TRY(nextac(i+1))
ed9de44032e0 Removal of the tactical STATE
paulson
parents: 2073
diff changeset
   119
					    THEN  TRY(nextac i)))
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   120
  in  tac  end;
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
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
(*repeated resolution applied to the designated goal*)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   125
fun reresolve_tac rules = 
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   126
  let val restac = RESOLVE_THEN rules;  (*preprocessing done now*)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   127
      fun gtac i = restac gtac i
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   128
  in  gtac  end; 
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   129
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   130
(*tries the safe rules repeatedly before the unsafe rules. *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   131
fun repeat_goal_tac (Pack(safes,unsafes)) = 
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   132
  let val restac  =    RESOLVE_THEN safes
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   133
      and lastrestac = RESOLVE_THEN unsafes;
6054
4a4f6ad607a1 added new arg for print_tac
paulson
parents: 4440
diff changeset
   134
      fun gtac i = restac gtac i  
7122
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   135
	           ORELSE  (if !trace then
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   136
				(print_tac "" THEN lastrestac gtac i)
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   137
			    else lastrestac gtac i)
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   138
  in  gtac  end; 
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   139
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
(*Tries safe rules only*)
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   142
fun safe_tac (Pack(safes,unsafes)) = reresolve_tac safes;
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   143
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   144
val safe_goal_tac = safe_tac;   (*backwards compatibility*)
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   145
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   146
(*Tries a safe rule or else a unsafe rule.  Single-step for tracing. *)
7122
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   147
fun step_tac (pack as Pack(safes,unsafes)) =
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   148
    safe_tac pack  ORELSE'
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   149
    filseq_resolve_tac unsafes 9999;
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   150
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
(* Tactic for reducing a goal, using Predicate Calculus rules.
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   153
   A decision procedure for Propositional Calculus, it is incomplete
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   154
   for Predicate-Calculus because of allL_thin and exR_thin.  
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   155
   Fails if it can do nothing.      *)
7122
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   156
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
   157
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
(*The following two tactics are analogous to those provided by 
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   160
  Provers/classical.  In fact, pc_tac is usually FASTER than fast_tac!*)
7122
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   161
fun fast_tac pack =
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   162
  SELECT_GOAL (DEPTH_SOLVE (step_tac pack 1));
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   163
7122
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   164
fun best_tac pack  = 
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   165
  SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, size_of_thm) 
7122
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   166
	       (step_tac pack 1));
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   167
7122
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   168
end;
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   169
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   170
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   171
open Cla;