src/FOL/intprover.ML
author haftmann
Thu, 04 Oct 2007 19:41:55 +0200
changeset 24841 df8448bc7a8b
parent 21539 c5cf9243ad62
child 31974 e81979a703a4
permissions -rw-r--r--
concept for exceptions
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1459
d12da312eff4 expanded tabs
clasohm
parents: 1005
diff changeset
     1
(*  Title:      FOL/int-prover
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1459
d12da312eff4 expanded tabs
clasohm
parents: 1005
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
A naive prover for intuitionistic logic
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
2601
b301958c465d Renamed structure Int (intuitionistic prover) to IntPr to prevent clash
paulson
parents: 2572
diff changeset
     8
BEWARE OF NAME CLASHES WITH CLASSICAL TACTICS -- use IntPr.fast_tac ...
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
Completeness (for propositional logic) is proved in 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    12
Roy Dyckhoff.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
Contraction-Free Sequent Calculi for Intuitionistic Logic.
1005
65188e520024 Updated comments.
lcp
parents: 702
diff changeset
    14
J. Symbolic Logic  57(3), 1992, pages 795-807.
65188e520024 Updated comments.
lcp
parents: 702
diff changeset
    15
65188e520024 Updated comments.
lcp
parents: 702
diff changeset
    16
The approach was developed independently by Roy Dyckhoff and L C Paulson.
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    17
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    18
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
signature INT_PROVER = 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    20
  sig
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    21
  val best_tac: int -> tactic
5203
eb5a1511a07d A little quantifier duplication for IFOL
paulson
parents: 4440
diff changeset
    22
  val best_dup_tac: int -> tactic
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    23
  val fast_tac: int -> tactic
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    24
  val inst_step_tac: int -> tactic
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    25
  val safe_step_tac: int -> tactic
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    26
  val safe_brls: (bool * thm) list
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    27
  val safe_tac: tactic
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    28
  val step_tac: int -> tactic
5203
eb5a1511a07d A little quantifier duplication for IFOL
paulson
parents: 4440
diff changeset
    29
  val step_dup_tac: int -> tactic
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    30
  val haz_brls: (bool * thm) list
5203
eb5a1511a07d A little quantifier duplication for IFOL
paulson
parents: 4440
diff changeset
    31
  val haz_dup_brls: (bool * thm) list
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    32
  end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    33
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    34
2601
b301958c465d Renamed structure Int (intuitionistic prover) to IntPr to prevent clash
paulson
parents: 2572
diff changeset
    35
structure IntPr : INT_PROVER   = 
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    36
struct
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    37
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    38
(*Negation is treated as a primitive symbol, with rules notI (introduction),
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    39
  not_to_imp (converts the assumption ~P to P-->False), and not_impE
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    40
  (handles double negations).  Could instead rewrite by not_def as the first
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    41
  step of an intuitionistic proof.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    42
*)
4440
9ed4098074bc adapted to new sort function;
wenzelm
parents: 2601
diff changeset
    43
val safe_brls = sort (make_ord lessb)
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 17496
diff changeset
    44
    [ (true, thm "FalseE"), (false, thm "TrueI"), (false, thm "refl"),
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 17496
diff changeset
    45
      (false, thm "impI"), (false, thm "notI"), (false, thm "allI"),
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 17496
diff changeset
    46
      (true, thm "conjE"), (true, thm "exE"),
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 17496
diff changeset
    47
      (false, thm "conjI"), (true, thm "conj_impE"),
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 17496
diff changeset
    48
      (true, thm "disj_impE"), (true, thm "disjE"), 
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 17496
diff changeset
    49
      (false, thm "iffI"), (true, thm "iffE"), (true, thm "not_to_imp") ];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    50
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    51
val haz_brls =
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 17496
diff changeset
    52
    [ (false, thm "disjI1"), (false, thm "disjI2"), (false, thm "exI"), 
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 17496
diff changeset
    53
      (true, thm "allE"), (true, thm "not_impE"), (true, thm "imp_impE"), (true, thm "iff_impE"),
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 17496
diff changeset
    54
      (true, thm "all_impE"), (true, thm "ex_impE"), (true, thm "impE") ];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    55
5203
eb5a1511a07d A little quantifier duplication for IFOL
paulson
parents: 4440
diff changeset
    56
val haz_dup_brls =
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 17496
diff changeset
    57
    [ (false, thm "disjI1"), (false, thm "disjI2"), (false, thm "exI"),
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 17496
diff changeset
    58
      (true, thm "all_dupE"), (true, thm "not_impE"), (true, thm "imp_impE"), (true, thm "iff_impE"),
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 17496
diff changeset
    59
      (true, thm "all_impE"), (true, thm "ex_impE"), (true, thm "impE") ];
5203
eb5a1511a07d A little quantifier duplication for IFOL
paulson
parents: 4440
diff changeset
    60
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    61
(*0 subgoals vs 1 or more: the p in safep is for positive*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    62
val (safe0_brls, safep_brls) =
17496
26535df536ae slight adaptions to library changes
haftmann
parents: 15570
diff changeset
    63
    List.partition (curry (op =) 0 o subgoals_of_brl) safe_brls;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    64
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    65
(*Attack subgoals using safe inferences -- matching, not resolution*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    66
val safe_step_tac = FIRST' [eq_assume_tac,
1459
d12da312eff4 expanded tabs
clasohm
parents: 1005
diff changeset
    67
                            eq_mp_tac,
d12da312eff4 expanded tabs
clasohm
parents: 1005
diff changeset
    68
                            bimatch_tac safe0_brls,
d12da312eff4 expanded tabs
clasohm
parents: 1005
diff changeset
    69
                            hyp_subst_tac,
d12da312eff4 expanded tabs
clasohm
parents: 1005
diff changeset
    70
                            bimatch_tac safep_brls] ;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    71
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    72
(*Repeatedly attack subgoals using safe inferences -- it's deterministic!*)
702
98fc1a8e832a FOL/intprover/safe_tac: now uses REPEAT_DETERM_FIRST instead of REPEAT_DETERM
lcp
parents: 0
diff changeset
    73
val safe_tac = REPEAT_DETERM_FIRST safe_step_tac;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    74
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    75
(*These steps could instantiate variables and are therefore unsafe.*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    76
val inst_step_tac =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    77
  assume_tac APPEND' mp_tac APPEND' 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    78
  biresolve_tac (safe0_brls @ safep_brls);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    79
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    80
(*One safe or unsafe step. *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    81
fun step_tac i = FIRST [safe_tac, inst_step_tac i, biresolve_tac haz_brls i];
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    82
5203
eb5a1511a07d A little quantifier duplication for IFOL
paulson
parents: 4440
diff changeset
    83
fun step_dup_tac i = FIRST [safe_tac, inst_step_tac i, 
eb5a1511a07d A little quantifier duplication for IFOL
paulson
parents: 4440
diff changeset
    84
			    biresolve_tac haz_dup_brls i];
eb5a1511a07d A little quantifier duplication for IFOL
paulson
parents: 4440
diff changeset
    85
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    86
(*Dumb but fast*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    87
val fast_tac = SELECT_GOAL (DEPTH_SOLVE (step_tac 1));
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    88
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    89
(*Slower but smarter than fast_tac*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    90
val best_tac = 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    91
  SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, size_of_thm) (step_tac 1));
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    92
5203
eb5a1511a07d A little quantifier duplication for IFOL
paulson
parents: 4440
diff changeset
    93
(*Uses all_dupE: allows multiple use of universal assumptions.  VERY slow.*)
eb5a1511a07d A little quantifier duplication for IFOL
paulson
parents: 4440
diff changeset
    94
val best_dup_tac = 
eb5a1511a07d A little quantifier duplication for IFOL
paulson
parents: 4440
diff changeset
    95
  SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, size_of_thm) (step_dup_tac 1));
eb5a1511a07d A little quantifier duplication for IFOL
paulson
parents: 4440
diff changeset
    96
eb5a1511a07d A little quantifier duplication for IFOL
paulson
parents: 4440
diff changeset
    97
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    98
end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    99