1 (* Title: FOLP/int-prover.ML
3 Author: Lawrence C Paulson, Cambridge University Computer Laboratory
4 Copyright 1992 University of Cambridge
6 A naive prover for intuitionistic logic
8 BEWARE OF NAME CLASHES WITH CLASSICAL TACTICS -- use IntPr.fast_tac ...
10 Completeness (for propositional logic) is proved in
13 Contraction-Free Sequent Calculi for IntPruitionistic Logic.
14 J. Symbolic Logic (in press)
17 signature INT_PROVER =
19 val best_tac: int -> tactic
20 val fast_tac: int -> tactic
21 val inst_step_tac: int -> tactic
22 val safe_step_tac: int -> tactic
23 val safe_brls: (bool * thm) list
25 val step_tac: int -> tactic
26 val haz_brls: (bool * thm) list
30 structure IntPr : INT_PROVER =
33 (*Negation is treated as a primitive symbol, with rules notI (introduction),
34 not_to_imp (converts the assumption ~P to P-->False), and not_impE
35 (handles double negations). Could instead rewrite by not_def as the first
36 step of an intuitionistic proof.
38 val safe_brls = sort (make_ord lessb)
39 [ (true,FalseE), (false,TrueI), (false,refl),
40 (false,impI), (false,notI), (false,allI),
41 (true,conjE), (true,exE),
42 (false,conjI), (true,conj_impE),
43 (true,disj_impE), (true,disjE),
44 (false,iffI), (true,iffE), (true,not_to_imp) ];
47 [ (false,disjI1), (false,disjI2), (false,exI),
48 (true,allE), (true,not_impE), (true,imp_impE), (true,iff_impE),
49 (true,all_impE), (true,ex_impE), (true,impE) ];
51 (*0 subgoals vs 1 or more: the p in safep is for positive*)
52 val (safe0_brls, safep_brls) =
53 partition (apl(0,op=) o subgoals_of_brl) safe_brls;
55 (*Attack subgoals using safe inferences*)
56 val safe_step_tac = FIRST' [uniq_assume_tac,
58 biresolve_tac safe0_brls,
60 biresolve_tac safep_brls] ;
62 (*Repeatedly attack subgoals using safe inferences*)
63 val safe_tac = DETERM (REPEAT_FIRST safe_step_tac);
65 (*These steps could instantiate variables and are therefore unsafe.*)
66 val inst_step_tac = assume_tac APPEND' mp_tac;
68 (*One safe or unsafe step. *)
69 fun step_tac i = FIRST [safe_tac, inst_step_tac i, biresolve_tac haz_brls i];
72 val fast_tac = SELECT_GOAL (DEPTH_SOLVE (step_tac 1));
74 (*Slower but smarter than fast_tac*)
76 SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, size_of_thm) (step_tac 1));