| 1459 |      1 | (*  Title:      FOLP/FOLP.ML
 | 
| 0 |      2 |     ID:         $Id$
 | 
| 1459 |      3 |     Author:     Martin D Coen, Cambridge University Computer Laboratory
 | 
| 0 |      4 |     Copyright   1991  University of Cambridge
 | 
|  |      5 | 
 | 
| 1142 |      6 | Tactics and lemmas for FOLP (Classical First-Order Logic with Proofs)
 | 
| 0 |      7 | *)
 | 
|  |      8 | 
 | 
|  |      9 | (*** Classical introduction rules for | and EX ***)
 | 
|  |     10 | 
 | 
| 9263 |     11 | val prems= goal FOLP.thy 
 | 
|  |     12 |    "(!!x. x:~Q ==> f(x):P) ==> ?p : P|Q";
 | 
|  |     13 | by (rtac classical 1);
 | 
|  |     14 | by (REPEAT (ares_tac (prems@[disjI1,notI]) 1));
 | 
|  |     15 | by (REPEAT (ares_tac (prems@[disjI2,notE]) 1)) ;
 | 
|  |     16 | qed "disjCI";
 | 
| 0 |     17 | 
 | 
|  |     18 | (*introduction rule involving only EX*)
 | 
| 9263 |     19 | val prems= goal FOLP.thy 
 | 
|  |     20 |    "( !!u. u:~(EX x. P(x)) ==> f(u):P(a)) ==> ?p : EX x. P(x)";
 | 
|  |     21 | by (rtac classical 1);
 | 
|  |     22 | by (eresolve_tac (prems RL [exI]) 1) ;
 | 
|  |     23 | qed "ex_classical";
 | 
| 0 |     24 | 
 | 
|  |     25 | (*version of above, simplifying ~EX to ALL~ *)
 | 
| 9263 |     26 | val [prem]= goal FOLP.thy 
 | 
|  |     27 |    "(!!u. u:ALL x. ~P(x) ==> f(u):P(a)) ==> ?p : EX x. P(x)";
 | 
|  |     28 | by (rtac ex_classical 1);
 | 
|  |     29 | by (resolve_tac [notI RS allI RS prem] 1);
 | 
|  |     30 | by (etac notE 1);
 | 
|  |     31 | by (etac exI 1) ;
 | 
|  |     32 | qed "exCI";
 | 
| 0 |     33 | 
 | 
|  |     34 | val excluded_middle = prove_goal FOLP.thy "?p : ~P | P"
 | 
|  |     35 |  (fn _=> [ rtac disjCI 1, assume_tac 1 ]);
 | 
|  |     36 | 
 | 
|  |     37 | 
 | 
|  |     38 | (*** Special elimination rules *)
 | 
|  |     39 | 
 | 
|  |     40 | 
 | 
|  |     41 | (*Classical implies (-->) elimination. *)
 | 
| 9263 |     42 | val major::prems= goal FOLP.thy 
 | 
|  |     43 |     "[| p:P-->Q;  !!x. x:~P ==> f(x):R;  !!y. y:Q ==> g(y):R |] ==> ?p : R";
 | 
|  |     44 | by (resolve_tac [excluded_middle RS disjE] 1);
 | 
|  |     45 | by (DEPTH_SOLVE (ares_tac (prems@[major RS mp]) 1)) ;
 | 
|  |     46 | qed "impCE";
 | 
| 0 |     47 | 
 | 
|  |     48 | (*Double negation law*)
 | 
| 9263 |     49 | Goal "p:~~P ==> ?p : P";
 | 
|  |     50 | by (rtac classical 1);
 | 
|  |     51 | by (etac notE 1);
 | 
|  |     52 | by (assume_tac 1);
 | 
|  |     53 | qed "notnotD";
 | 
| 0 |     54 | 
 | 
|  |     55 | 
 | 
|  |     56 | (*** Tactics for implication and contradiction ***)
 | 
|  |     57 | 
 | 
|  |     58 | (*Classical <-> elimination.  Proof substitutes P=Q in 
 | 
|  |     59 |     ~P ==> ~Q    and    P ==> Q  *)
 | 
| 9263 |     60 | val prems = goalw FOLP.thy [iff_def]
 | 
| 0 |     61 |     "[| p:P<->Q; !!x y.[| x:P; y:Q |] ==> f(x,y):R;  \
 | 
| 9263 |     62 | \                !!x y.[| x:~P; y:~Q |] ==> g(x,y):R |] ==> ?p : R";
 | 
|  |     63 | by (rtac conjE 1);
 | 
|  |     64 | by (REPEAT (DEPTH_SOLVE_1 (etac impCE 1
 | 
|  |     65 |                ORELSE  mp_tac 1  ORELSE  ares_tac prems 1))) ;
 | 
|  |     66 | qed "iffCE";
 | 
| 0 |     67 | 
 | 
|  |     68 | 
 | 
|  |     69 | (*Should be used as swap since ~P becomes redundant*)
 | 
| 9263 |     70 | val major::prems= goal FOLP.thy 
 | 
|  |     71 |    "p:~P ==> (!!x. x:~Q ==> f(x):P) ==> ?p : Q";
 | 
|  |     72 | by (rtac classical 1);
 | 
|  |     73 | by (rtac (major RS notE) 1);
 | 
|  |     74 | by (REPEAT (ares_tac prems 1)) ;
 | 
|  |     75 | qed "swap";
 | 
| 0 |     76 | 
 |