src/Sequents/prover.ML
author nipkow
Mon, 06 Aug 2001 13:43:24 +0200
changeset 11464 ddea204de5bc
parent 7150 d203e2282789
child 12123 739eba13e2cd
permissions -rw-r--r--
turned translation for 1::nat into def. introduced 1' and replaced most occurrences of 1 by 1'.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7122
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
     1
(*  Title:      Sequents/prover
2073
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
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    10
(*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
    11
infix 4 add_safes add_unsafes;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    12
7122
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
    13
structure Cla =
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
    14
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
    15
struct
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
    16
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    17
datatype pack = Pack of thm list * thm list;
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    18
7122
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
    19
val trace = ref false;
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
    20
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    21
(*A theorem pack has the form  (safe rules, unsafe rules)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    22
  An unsafe rule is incomplete or introduces variables in subgoals,
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    23
  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
    24
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    25
fun less (rl1,rl2) = (nprems_of rl1) < (nprems_of rl2);
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    26
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    27
val empty_pack = Pack([],[]);
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    28
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    29
fun warn_duplicates [] = []
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    30
  | warn_duplicates dups =
7150
d203e2282789 cat_lines;
wenzelm
parents: 7122
diff changeset
    31
      (warning (cat_lines ("Ignoring duplicate theorems:" :: map string_of_thm dups));
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    32
       dups);
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    33
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    34
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
    35
    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
    36
	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
    37
    in
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    38
        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
    39
    end;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    40
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    41
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
    42
    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
    43
	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
    44
    in
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    45
	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
    46
    end;
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    47
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    48
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
    49
        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
    50
	     sort (make_ord less) (unsafes@unsafes'));
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    51
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    52
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    53
fun print_pack (Pack(safes,unsafes)) =
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    54
    (writeln "Safe rules:";  print_thms safes;
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
    55
     writeln "Unsafe rules:"; print_thms unsafes);
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff 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: 6054
diff 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) =
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    68
      forall (fn Qc => exists (fn Qp => could_unify (Qp,Qc)) 
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)) =>
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    78
	  could_res (leftp,leftc) andalso could_res (rightp,rightc)
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) =>
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    81
	  could_res (leftp,leftc)  andalso  could_unify (rightp,rightc)
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
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    91
	  else (*((rtac derelict 1 THEN rtac impl 1
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    92
		 THEN (rtac identity 2 ORELSE rtac ll_mp 2)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    93
		 THEN rtac context1 1)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    94
		 ORELSE *) resolve_tac rls i
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
ed9de44032e0 Removal of the tactical STATE
paulson
parents: 2073
diff changeset
   114
      fun tac nextac i state = state |>
ed9de44032e0 Removal of the tactical STATE
paulson
parents: 2073
diff changeset
   115
	     (filseq_resolve_tac rls0 9999 i 
ed9de44032e0 Removal of the tactical STATE
paulson
parents: 2073
diff changeset
   116
	      ORELSE
ed9de44032e0 Removal of the tactical STATE
paulson
parents: 2073
diff changeset
   117
	      (DETERM(filseq_resolve_tac rls1 9999 i) THEN  TRY(nextac i))
ed9de44032e0 Removal of the tactical STATE
paulson
parents: 2073
diff changeset
   118
	      ORELSE
ed9de44032e0 Removal of the tactical STATE
paulson
parents: 2073
diff changeset
   119
	      (DETERM(filseq_resolve_tac rls2 9999 i) THEN  TRY(nextac(i+1))
ed9de44032e0 Removal of the tactical STATE
paulson
parents: 2073
diff 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
4a4f6ad607a1 added new arg for print_tac
paulson
parents: 4440
diff changeset
   135
      fun gtac i = restac gtac i  
7122
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   136
	           ORELSE  (if !trace then
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   137
				(print_tac "" THEN lastrestac gtac i)
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff 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: 6054
diff 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: 6054
diff changeset
   144
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff 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
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   148
fun step_tac (pack as Pack(safes,unsafes)) =
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   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
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   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
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   162
fun fast_tac pack =
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   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
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   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) 
7122
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   167
	       (step_tac pack 1));
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   168
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   169
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   170
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   171
structure ProverArgs =
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   172
  struct
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   173
  val name = "Sequents/prover";
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   174
  type T = pack ref;
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   175
  val empty = ref empty_pack
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   176
  fun copy (ref pack) = ref pack;
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   177
  val prep_ext = copy;
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   178
  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
   179
  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
   180
  end;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   181
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   182
structure ProverData = TheoryDataFun(ProverArgs);
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
val prover_setup = [ProverData.init];
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   185
7122
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   186
val print_pack = ProverData.print;
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   187
val pack_ref_of_sg = ProverData.get_sg;
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   188
val pack_ref_of = ProverData.get;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   189
7122
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   190
(* access global pack *)
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   191
7122
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   192
val pack_of_sg = ! o pack_ref_of_sg;
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   193
val pack_of = pack_of_sg o sign_of;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   194
7122
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   195
val pack = pack_of o Context.the_context;
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   196
val pack_ref = 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
   197
7097
5ab37ed3d53c moved the modal prover to modal.ML; installed the prover using TheoryDataFun
paulson
parents: 6054
diff changeset
   198
7122
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   199
(* change global pack *)
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   200
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   201
fun change_pack f x = pack_ref () := (f (pack (), x));
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   202
7122
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   203
val Add_safes = change_pack (op add_safes);
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   204
val Add_unsafes = change_pack (op add_unsafes);
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   205
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   206
7122
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   207
fun Fast_tac st = fast_tac (pack()) st;
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   208
fun Step_tac st = step_tac (pack()) st;
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   209
fun Safe_tac st = safe_tac (pack()) st;
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   210
fun Pc_tac st   = pc_tac (pack()) st;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   211
7122
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   212
end;
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   213
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   214
87b233b31889 renamed ...thm_pack... to ...pack...
paulson
parents: 7097
diff changeset
   215
open Cla;