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