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