| 
7355
 | 
     1  | 
(*  Title:      FOL/FOL_lemmas1.ML
  | 
| 
 | 
     2  | 
    ID:         $Id$
  | 
| 
 | 
     3  | 
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
  | 
| 
 | 
     4  | 
    Copyright   1991  University of Cambridge
  | 
| 
 | 
     5  | 
  | 
| 
 | 
     6  | 
Tactics and lemmas for theory FOL (classical First-Order Logic).
  | 
| 
 | 
     7  | 
*)
  | 
| 
 | 
     8  | 
  | 
| 
 | 
     9  | 
val classical = thm "classical";
  | 
| 
7422
 | 
    10  | 
bind_thm ("ccontr", FalseE RS classical);
 | 
| 
7355
 | 
    11  | 
  | 
| 
 | 
    12  | 
  | 
| 
 | 
    13  | 
(*** Classical introduction rules for | and EX ***)
  | 
| 
 | 
    14  | 
  | 
| 
9264
 | 
    15  | 
val prems = Goal "(~Q ==> P) ==> P|Q";
  | 
| 
 | 
    16  | 
by (rtac classical 1);
  | 
| 
 | 
    17  | 
by (REPEAT (ares_tac (prems@[disjI1,notI]) 1));
  | 
| 
 | 
    18  | 
by (REPEAT (ares_tac (prems@[disjI2,notE]) 1)) ;
  | 
| 
 | 
    19  | 
qed "disjCI";
  | 
| 
7355
 | 
    20  | 
  | 
| 
 | 
    21  | 
(*introduction rule involving only EX*)
  | 
| 
9264
 | 
    22  | 
val prems = Goal "( ~(EX x. P(x)) ==> P(a)) ==> EX x. P(x)";
  | 
| 
 | 
    23  | 
by (rtac classical 1);
  | 
| 
 | 
    24  | 
by (eresolve_tac (prems RL [exI]) 1) ;
  | 
| 
 | 
    25  | 
qed "ex_classical";
  | 
| 
7355
 | 
    26  | 
  | 
| 
 | 
    27  | 
(*version of above, simplifying ~EX to ALL~ *)
  | 
| 
9264
 | 
    28  | 
val [prem]= Goal "(ALL x. ~P(x) ==> P(a)) ==> EX x. P(x)";
  | 
| 
 | 
    29  | 
by (rtac ex_classical 1);
  | 
| 
 | 
    30  | 
by (resolve_tac [notI RS allI RS prem] 1);
  | 
| 
 | 
    31  | 
by (etac notE 1);
  | 
| 
 | 
    32  | 
by (etac exI 1) ;
  | 
| 
 | 
    33  | 
qed "exCI";
  | 
| 
7355
 | 
    34  | 
  | 
| 
9264
 | 
    35  | 
Goal"~P | P";
  | 
| 
 | 
    36  | 
by (rtac disjCI 1);
  | 
| 
 | 
    37  | 
by (assume_tac 1) ;
  | 
| 
 | 
    38  | 
qed "excluded_middle";
  | 
| 
7355
 | 
    39  | 
  | 
| 
 | 
    40  | 
(*For disjunctive case analysis*)
  | 
| 
 | 
    41  | 
fun excluded_middle_tac sP =
  | 
| 
 | 
    42  | 
    res_inst_tac [("Q",sP)] (excluded_middle RS disjE);
 | 
| 
 | 
    43  | 
  | 
| 
9264
 | 
    44  | 
val [p1,p2] = Goal"[| P ==> Q; ~P ==> Q |] ==> Q";
  | 
| 
 | 
    45  | 
by (rtac (excluded_middle RS disjE) 1);
  | 
| 
 | 
    46  | 
by (etac p2 1);
  | 
| 
 | 
    47  | 
by (etac p1 1);
  | 
| 
 | 
    48  | 
qed "case_split_thm";
  | 
| 
7355
 | 
    49  | 
  | 
| 
 | 
    50  | 
(*HOL's more natural case analysis tactic*)
  | 
| 
 | 
    51  | 
fun case_tac a = res_inst_tac [("P",a)] case_split_thm;
 | 
| 
 | 
    52  | 
  | 
| 
 | 
    53  | 
  | 
| 
 | 
    54  | 
(*** Special elimination rules *)
  | 
| 
 | 
    55  | 
  | 
| 
 | 
    56  | 
  | 
| 
 | 
    57  | 
(*Classical implies (-->) elimination. *)
  | 
| 
9264
 | 
    58  | 
val major::prems = Goal "[| P-->Q;  ~P ==> R;  Q ==> R |] ==> R";
  | 
| 
 | 
    59  | 
by (resolve_tac [excluded_middle RS disjE] 1);
  | 
| 
 | 
    60  | 
by (DEPTH_SOLVE (ares_tac (prems@[major RS mp]) 1)) ;
  | 
| 
 | 
    61  | 
qed "impCE";
  | 
| 
7355
 | 
    62  | 
  | 
| 
 | 
    63  | 
(*This version of --> elimination works on Q before P.  It works best for
  | 
| 
 | 
    64  | 
  those cases in which P holds "almost everywhere".  Can't install as
  | 
| 
 | 
    65  | 
  default: would break old proofs.*)
  | 
| 
9264
 | 
    66  | 
val major::prems = Goal "[| P-->Q;  Q ==> R;  ~P ==> R |] ==> R";
  | 
| 
 | 
    67  | 
by (resolve_tac [excluded_middle RS disjE] 1);
  | 
| 
 | 
    68  | 
by (DEPTH_SOLVE (ares_tac (prems@[major RS mp]) 1)) ;
  | 
| 
 | 
    69  | 
qed "impCE'";
  | 
| 
7355
 | 
    70  | 
  | 
| 
 | 
    71  | 
(*Double negation law*)
  | 
| 
9264
 | 
    72  | 
Goal"~~P ==> P";
  | 
| 
 | 
    73  | 
by (rtac classical 1);
  | 
| 
 | 
    74  | 
by (etac notE 1);
  | 
| 
 | 
    75  | 
by (assume_tac 1);
  | 
| 
 | 
    76  | 
qed "notnotD";
  | 
| 
7355
 | 
    77  | 
  | 
| 
9264
 | 
    78  | 
val [p1,p2] = Goal"[| Q; ~ P ==> ~ Q |] ==> P";
  | 
| 
 | 
    79  | 
by (rtac classical 1);
  | 
| 
 | 
    80  | 
by (dtac p2 1);
  | 
| 
 | 
    81  | 
by (etac notE 1);
  | 
| 
 | 
    82  | 
by (rtac p1 1);
  | 
| 
 | 
    83  | 
qed "contrapos2";
  | 
| 
7355
 | 
    84  | 
  | 
| 
 | 
    85  | 
(*** Tactics for implication and contradiction ***)
  | 
| 
 | 
    86  | 
  | 
| 
 | 
    87  | 
(*Classical <-> elimination.  Proof substitutes P=Q in 
  | 
| 
 | 
    88  | 
    ~P ==> ~Q    and    P ==> Q  *)
  | 
| 
9264
 | 
    89  | 
val major::prems = 
  | 
| 
 | 
    90  | 
Goalw  [iff_def] "[| P<->Q;  [| P; Q |] ==> R;  [| ~P; ~Q |] ==> R |] ==> R";
  | 
| 
 | 
    91  | 
by (rtac (major RS conjE) 1);
  | 
| 
 | 
    92  | 
by (REPEAT_FIRST (etac impCE));
  | 
| 
 | 
    93  | 
by (REPEAT (DEPTH_SOLVE_1 (mp_tac 1  ORELSE  ares_tac prems 1)));
  | 
| 
 | 
    94  | 
qed "iffCE";
  | 
| 
 | 
    95  | 
  |