src/Sequents/prover.ML
author paulson
Tue, 27 Jul 1999 19:01:46 +0200
changeset 7097 5ab37ed3d53c
parent 6054 4a4f6ad607a1
child 7122 87b233b31889
permissions -rw-r--r--
moved the modal prover to modal.ML; installed the prover using TheoryDataFun
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
     1
(*  Title:      LK/LK.ML
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
     2
    ID:         $Id$
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
     5
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
     6
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
     7
*)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
     8
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
     9
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    10
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    11
(*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
    12
infix 4 add_safes add_unsafes;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    13
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    14
datatype pack = Pack of thm list * thm list;
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    15
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    16
(*A theorem pack has the form  (safe rules, unsafe rules)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    17
  An unsafe rule is incomplete or introduces variables in subgoals,
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    18
  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
    19
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    20
fun less (rl1,rl2) = (nprems_of rl1) < (nprems_of rl2);
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    21
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    22
val empty_pack = Pack([],[]);
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    23
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    24
fun warn_duplicates [] = []
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    25
  | warn_duplicates dups =
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    26
      (warning (String.concat ("Ignoring duplicate theorems:\n"::
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    27
			       map (suffix "\n" o string_of_thm) dups));
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    28
       dups);
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    29
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    30
fun (Pack(safes,unsafes)) add_safes ths   = 
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    31
    let val dups = warn_duplicates (gen_inter eq_thm (ths,safes))
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    32
	val ths' = gen_rems eq_thm (ths,dups)
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    33
    in
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    34
        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
    35
    end;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    36
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    37
fun (Pack(safes,unsafes)) add_unsafes ths = 
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    38
    let val dups = warn_duplicates (gen_inter eq_thm (ths,unsafes))
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    39
	val ths' = gen_rems eq_thm (ths,dups)
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    40
    in
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    41
	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
    42
    end;
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    43
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    44
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
    45
        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
    46
	     sort (make_ord less) (unsafes@unsafes'));
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    47
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    48
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    49
fun print_pack (Pack(safes,unsafes)) =
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    50
    (writeln "Safe rules:";  print_thms safes;
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    51
     writeln "Unsafe rules:"; print_thms unsafes);
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    52
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    53
(*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
    54
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
    55
  | forms_of_seq (H $ u) = forms_of_seq u
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    56
  | forms_of_seq _ = [];
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    57
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    58
(*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
    59
  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
    60
  Assumes each formula in seqc is surrounded by sequence variables
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    61
  -- checks that each concl formula looks like some subgoal formula.
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    62
  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
    63
fun could_res (seqp,seqc) =
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    64
      forall (fn Qc => exists (fn Qp => could_unify (Qp,Qc)) 
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    65
                              (forms_of_seq seqp))
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    66
             (forms_of_seq seqc);
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    67
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    68
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    69
(*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
    70
fun could_resolve_seq (prem,conc) =
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    71
  case (prem,conc) of
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    72
      (_ $ Abs(_,_,leftp) $ Abs(_,_,rightp),
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    73
       _ $ Abs(_,_,leftc) $ Abs(_,_,rightc)) =>
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    74
	  could_res (leftp,leftc) andalso could_res (rightp,rightc)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    75
    | (_ $ Abs(_,_,leftp) $ rightp,
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    76
       _ $ Abs(_,_,leftc) $ rightc) =>
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    77
	  could_res (leftp,leftc)  andalso  could_unify (rightp,rightc)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    78
    | _ => false;
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    79
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    80
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    81
(*Like filt_resolve_tac, using could_resolve_seq
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    82
  Much faster than resolve_tac when there are many rules.
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    83
  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
    84
fun filseq_resolve_tac rules maxr = SUBGOAL(fn (prem,i) =>
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    85
  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
    86
  in  if length rls > maxr  then  no_tac
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    87
	  else (*((rtac derelict 1 THEN rtac impl 1
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    88
		 THEN (rtac identity 2 ORELSE rtac ll_mp 2)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    89
		 THEN rtac context1 1)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    90
		 ORELSE *) resolve_tac rls i
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    91
  end);
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    92
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    93
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    94
(*Predicate: does the rule have n premises? *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    95
fun has_prems n rule =  (nprems_of rule = n);
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
(*Continuation-style tactical for resolution.
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    98
  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
    99
  The resulting tactic, gtac, tries to resolve with rules.
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   100
  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
   101
  Else fails.  (Treatment of goals due to Ph. de Groote) 
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   102
  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
   103
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   104
(*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
   105
  The abstraction over state prevents needless divergence in recursion.
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   106
  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
   107
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   108
fun RESOLVE_THEN rules =
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   109
  let val [rls0,rls1,rls2] = partition_list has_prems 0 2 rules;
3538
ed9de44032e0 Removal of the tactical STATE
paulson
parents: 2073
diff changeset
   110
      fun tac nextac i state = state |>
ed9de44032e0 Removal of the tactical STATE
paulson
parents: 2073
diff changeset
   111
	     (filseq_resolve_tac rls0 9999 i 
ed9de44032e0 Removal of the tactical STATE
paulson
parents: 2073
diff changeset
   112
	      ORELSE
ed9de44032e0 Removal of the tactical STATE
paulson
parents: 2073
diff changeset
   113
	      (DETERM(filseq_resolve_tac rls1 9999 i) THEN  TRY(nextac i))
ed9de44032e0 Removal of the tactical STATE
paulson
parents: 2073
diff changeset
   114
	      ORELSE
ed9de44032e0 Removal of the tactical STATE
paulson
parents: 2073
diff changeset
   115
	      (DETERM(filseq_resolve_tac rls2 9999 i) THEN  TRY(nextac(i+1))
ed9de44032e0 Removal of the tactical STATE
paulson
parents: 2073
diff changeset
   116
					    THEN  TRY(nextac i)))
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   117
  in  tac  end;
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   118
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   119
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   120
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   121
(*repeated resolution applied to the designated goal*)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   122
fun reresolve_tac rules = 
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   123
  let val restac = RESOLVE_THEN rules;  (*preprocessing done now*)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   124
      fun gtac i = restac gtac i
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   125
  in  gtac  end; 
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   126
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   127
(*tries the safe rules repeatedly before the unsafe rules. *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   128
fun repeat_goal_tac (Pack(safes,unsafes)) = 
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   129
  let val restac  =    RESOLVE_THEN safes
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   130
      and lastrestac = RESOLVE_THEN unsafes;
6054
4a4f6ad607a1 added new arg for print_tac
paulson
parents: 4440
diff changeset
   131
      fun gtac i = restac gtac i  
4a4f6ad607a1 added new arg for print_tac
paulson
parents: 4440
diff changeset
   132
	           ORELSE  (print_tac "" THEN lastrestac gtac i)
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   133
  in  gtac  end; 
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   134
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   135
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   136
(*Tries safe rules only*)
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   137
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
   138
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   139
val safe_goal_tac = safe_tac;   (*backwards compatibility*)
2073
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 a safe rule or else a unsafe rule.  Single-step for tracing. *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   142
fun step_tac (thm_pack as Pack(safes,unsafes)) =
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   143
    safe_tac thm_pack  ORELSE'
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   144
    filseq_resolve_tac unsafes 9999;
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
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   147
(* Tactic for reducing a goal, using Predicate Calculus rules.
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   148
   A decision procedure for Propositional Calculus, it is incomplete
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   149
   for Predicate-Calculus because of allL_thin and exR_thin.  
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   150
   Fails if it can do nothing.      *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   151
fun pc_tac thm_pack = SELECT_GOAL (DEPTH_SOLVE (repeat_goal_tac thm_pack 1));
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
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   154
(*The following two tactics are analogous to those provided by 
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   155
  Provers/classical.  In fact, pc_tac is usually FASTER than fast_tac!*)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   156
fun fast_tac thm_pack =
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   157
  SELECT_GOAL (DEPTH_SOLVE (step_tac thm_pack 1));
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
fun best_tac thm_pack  = 
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   160
  SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, size_of_thm) 
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   161
	       (step_tac thm_pack 1));
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   162
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   163
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   164
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   165
structure ProverArgs =
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   166
  struct
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   167
  val name = "Sequents/prover";
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   168
  type T = pack ref;
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   169
  val empty = ref empty_pack
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   170
  fun copy (ref pack) = ref pack;
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   171
  val prep_ext = copy;
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   172
  fun merge (ref pack1, ref pack2) = ref (merge_pack (pack1, pack2));
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   173
  fun print _ (ref pack) = print_pack pack;
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   174
  end;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   175
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   176
structure ProverData = TheoryDataFun(ProverArgs);
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   177
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   178
val prover_setup = [ProverData.init];
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   179
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   180
val print_thm_pack = ProverData.print;
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   181
val thm_pack_ref_of_sg = ProverData.get_sg;
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   182
val thm_pack_ref_of = ProverData.get;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   183
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   184
(* access global thm_pack *)
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   185
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   186
val thm_pack_of_sg = ! o thm_pack_ref_of_sg;
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   187
val thm_pack_of = thm_pack_of_sg o sign_of;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   188
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   189
val thm_pack = thm_pack_of o Context.the_context;
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   190
val thm_pack_ref = thm_pack_ref_of_sg o sign_of o Context.the_context;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   191
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   192
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   193
(* change global thm_pack *)
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   194
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   195
fun change_thm_pack f x = thm_pack_ref () := (f (thm_pack (), x));
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   196
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   197
val Add_safes = change_thm_pack (op add_safes);
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   198
val Add_unsafes = change_thm_pack (op add_unsafes);
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   199