src/FOLP/IFOLP.ML
author paulson
Thu, 06 Jul 2000 13:11:32 +0200
changeset 9263 53e09e592278
parent 3836 f1a1817659e6
child 15570 8d8c70b41bab
permissions -rw-r--r--
removal of batch style, and tidying
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1459
d12da312eff4 expanded tabs
clasohm
parents: 1142
diff changeset
     1
(*  Title:      FOLP/IFOLP.ML
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1459
d12da312eff4 expanded tabs
clasohm
parents: 1142
diff changeset
     3
    Author:     Martin D Coen, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
1142
eb0e2ff8f032 Corrected comments in headers
lcp
parents: 0
diff changeset
     6
Tactics and lemmas for IFOLP (Intuitionistic First-Order Logic with Proofs)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
(*** Sequent-style elimination rules for & --> and ALL ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    10
val prems= Goal 
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    11
    "[| p:P&Q; !!x y.[| x:P; y:Q |] ==> f(x,y):R |] ==> ?a:R";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    12
by (REPEAT (resolve_tac prems 1
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    13
   ORELSE (resolve_tac [conjunct1, conjunct2] 1 THEN resolve_tac prems 1))) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    14
qed "conjE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    16
val prems= Goal 
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    17
    "[| p:P-->Q;  q:P;  !!x. x:Q ==> r(x):R |] ==> ?p:R";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    18
by (REPEAT (resolve_tac (prems@[mp]) 1)) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    19
qed "impE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    20
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    21
val prems= Goal 
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    22
    "[| p:ALL x. P(x); !!y. y:P(x) ==> q(y):R |] ==> ?p:R";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    23
by (REPEAT (resolve_tac (prems@[spec]) 1)) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    24
qed "allE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    25
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    26
(*Duplicates the quantifier; for use with eresolve_tac*)
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    27
val prems= Goal 
3836
f1a1817659e6 fixed dots;
wenzelm
parents: 1459
diff changeset
    28
    "[| p:ALL x. P(x);  !!y z.[| y:P(x); z:ALL x. P(x) |] ==> q(y,z):R \
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    29
\    |] ==> ?p:R";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    30
by (REPEAT (resolve_tac (prems@[spec]) 1)) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    31
qed "all_dupE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    32
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    33
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    34
(*** Negation rules, which translate between ~P and P-->False ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    35
3836
f1a1817659e6 fixed dots;
wenzelm
parents: 1459
diff changeset
    36
val notI = prove_goalw IFOLP.thy [not_def]  "(!!x. x:P ==> q(x):False) ==> ?p:~P"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    37
 (fn prems=> [ (REPEAT (ares_tac (prems@[impI]) 1)) ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    38
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    39
val notE = prove_goalw IFOLP.thy [not_def] "[| p:~P;  q:P |] ==> ?p:R"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    40
 (fn prems=>
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    41
  [ (resolve_tac [mp RS FalseE] 1),
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    42
    (REPEAT (resolve_tac prems 1)) ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    43
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    44
(*This is useful with the special implication rules for each kind of P. *)
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    45
val prems= Goal 
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    46
    "[| p:~P;  !!x. x:(P-->False) ==> q(x):Q |] ==> ?p:Q";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    47
by (REPEAT (ares_tac (prems@[impI,notE]) 1)) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    48
qed "not_to_imp";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    49
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    50
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    51
(* For substitution int an assumption P, reduce Q to P-->Q, substitute into
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    52
   this implication, then apply impI to move P back into the assumptions.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    53
   To specify P use something like
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    54
      eres_inst_tac [ ("P","ALL y. ?S(x,y)") ] rev_mp 1   *)
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    55
Goal "[| p:P;  q:P --> Q |] ==> ?p:Q";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    56
by (REPEAT (ares_tac [mp] 1)) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    57
qed "rev_mp";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    58
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    59
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    60
(*Contrapositive of an inference rule*)
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    61
val [major,minor]= Goal "[| p:~Q;  !!y. y:P==>q(y):Q |] ==> ?a:~P";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    62
by (rtac (major RS notE RS notI) 1);
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    63
by (etac minor 1) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    64
qed "contrapos";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    65
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    66
(** Unique assumption tactic.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    67
    Ignores proof objects.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    68
    Fails unless one assumption is equal and exactly one is unifiable 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    69
**)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    70
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    71
local
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    72
    fun discard_proof (Const("Proof",_) $ P $ _) = P;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    73
in
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    74
val uniq_assume_tac =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    75
  SUBGOAL
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    76
    (fn (prem,i) =>
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    77
      let val hyps = map discard_proof (Logic.strip_assums_hyp prem)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    78
          and concl = discard_proof (Logic.strip_assums_concl prem)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    79
      in  
1459
d12da312eff4 expanded tabs
clasohm
parents: 1142
diff changeset
    80
          if exists (fn hyp => hyp aconv concl) hyps
d12da312eff4 expanded tabs
clasohm
parents: 1142
diff changeset
    81
          then case distinct (filter (fn hyp=> could_unify(hyp,concl)) hyps) of
d12da312eff4 expanded tabs
clasohm
parents: 1142
diff changeset
    82
                   [_] => assume_tac i
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    83
                 |  _  => no_tac
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    84
          else no_tac
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    85
      end);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    86
end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    87
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    88
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    89
(*** Modus Ponens Tactics ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    90
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    91
(*Finds P-->Q and P in the assumptions, replaces implication by Q *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    92
fun mp_tac i = eresolve_tac [notE,make_elim mp] i  THEN  assume_tac i;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    93
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    94
(*Like mp_tac but instantiates no variables*)
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
    95
fun int_uniq_mp_tac i = eresolve_tac [notE,impE] i  THEN  uniq_assume_tac i;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    96
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    97
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    98
(*** If-and-only-if ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    99
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   100
val iffI = prove_goalw IFOLP.thy [iff_def]
3836
f1a1817659e6 fixed dots;
wenzelm
parents: 1459
diff changeset
   101
   "[| !!x. x:P ==> q(x):Q;  !!x. x:Q ==> r(x):P |] ==> ?p:P<->Q"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   102
 (fn prems=> [ (REPEAT (ares_tac (prems@[conjI, impI]) 1)) ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   103
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   104
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   105
(*Observe use of rewrite_rule to unfold "<->" in meta-assumptions (prems) *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   106
val iffE = prove_goalw IFOLP.thy [iff_def]
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   107
    "[| p:P <-> Q;  !!x y.[| x:P-->Q; y:Q-->P |] ==> q(x,y):R |] ==> ?p:R"
1459
d12da312eff4 expanded tabs
clasohm
parents: 1142
diff changeset
   108
 (fn prems => [ (rtac conjE 1), (REPEAT (ares_tac prems 1)) ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   109
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   110
(* Destruct rules for <-> similar to Modus Ponens *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   111
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   112
val iffD1 = prove_goalw IFOLP.thy [iff_def] "[| p:P <-> Q;  q:P |] ==> ?p:Q"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   113
 (fn prems => [ (rtac (conjunct1 RS mp) 1), (REPEAT (ares_tac prems 1)) ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   114
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   115
val iffD2 = prove_goalw IFOLP.thy [iff_def] "[| p:P <-> Q;  q:Q |] ==> ?p:P"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   116
 (fn prems => [ (rtac (conjunct2 RS mp) 1), (REPEAT (ares_tac prems 1)) ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   117
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   118
Goal "?p:P <-> P";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   119
by (REPEAT (ares_tac [iffI] 1)) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   120
qed "iff_refl";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   121
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   122
Goal "p:Q <-> P ==> ?p:P <-> Q";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   123
by (etac iffE 1);
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   124
by (rtac iffI 1);
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   125
by (REPEAT (eresolve_tac [asm_rl,mp] 1)) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   126
qed "iff_sym";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   127
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   128
Goal "[| p:P <-> Q; q:Q<-> R |] ==> ?p:P <-> R";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   129
by (rtac iffI 1);
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   130
by (REPEAT (eresolve_tac [asm_rl,iffE] 1 ORELSE mp_tac 1)) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   131
qed "iff_trans";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   132
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   133
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   134
(*** Unique existence.  NOTE THAT the following 2 quantifications
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   135
   EX!x such that [EX!y such that P(x,y)]     (sequential)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   136
   EX!x,y such that P(x,y)                    (simultaneous)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   137
 do NOT mean the same thing.  The parser treats EX!x y.P(x,y) as sequential.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   138
***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   139
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   140
val prems = goalw IFOLP.thy [ex1_def]
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   141
    "[| p:P(a);  !!x u. u:P(x) ==> f(u) : x=a |] ==> ?p:EX! x. P(x)";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   142
by (REPEAT (ares_tac (prems@[exI,conjI,allI,impI]) 1)) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   143
qed "ex1I";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   144
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   145
val prems = goalw IFOLP.thy [ex1_def]
3836
f1a1817659e6 fixed dots;
wenzelm
parents: 1459
diff changeset
   146
    "[| p:EX! x. P(x);  \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   147
\       !!x u v. [| u:P(x);  v:ALL y. P(y) --> y=x |] ==> f(x,u,v):R |] ==>\
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   148
\    ?a : R";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   149
by (cut_facts_tac prems 1);
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   150
by (REPEAT (eresolve_tac [exE,conjE] 1 ORELSE ares_tac prems 1)) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   151
qed "ex1E";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   152
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   153
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   154
(*** <-> congruence rules for simplification ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   155
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   156
(*Use iffE on a premise.  For conj_cong, imp_cong, all_cong, ex_cong*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   157
fun iff_tac prems i =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   158
    resolve_tac (prems RL [iffE]) i THEN
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   159
    REPEAT1 (eresolve_tac [asm_rl,mp] i);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   160
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   161
val conj_cong = prove_goal IFOLP.thy 
3836
f1a1817659e6 fixed dots;
wenzelm
parents: 1459
diff changeset
   162
    "[| p:P <-> P';  !!x. x:P' ==> q(x):Q <-> Q' |] ==> ?p:(P&Q) <-> (P'&Q')"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   163
 (fn prems =>
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   164
  [ (cut_facts_tac prems 1),
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   165
    (REPEAT  (ares_tac [iffI,conjI] 1
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   166
      ORELSE  eresolve_tac [iffE,conjE,mp] 1
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   167
      ORELSE  iff_tac prems 1)) ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   168
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   169
val disj_cong = prove_goal IFOLP.thy 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   170
    "[| p:P <-> P';  q:Q <-> Q' |] ==> ?p:(P|Q) <-> (P'|Q')"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   171
 (fn prems =>
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   172
  [ (cut_facts_tac prems 1),
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   173
    (REPEAT  (eresolve_tac [iffE,disjE,disjI1,disjI2] 1
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   174
      ORELSE  ares_tac [iffI] 1
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   175
      ORELSE  mp_tac 1)) ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   176
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   177
val imp_cong = prove_goal IFOLP.thy 
3836
f1a1817659e6 fixed dots;
wenzelm
parents: 1459
diff changeset
   178
    "[| p:P <-> P';  !!x. x:P' ==> q(x):Q <-> Q' |] ==> ?p:(P-->Q) <-> (P'-->Q')"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   179
 (fn prems =>
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   180
  [ (cut_facts_tac prems 1),
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   181
    (REPEAT   (ares_tac [iffI,impI] 1
1459
d12da312eff4 expanded tabs
clasohm
parents: 1142
diff changeset
   182
      ORELSE  etac iffE 1
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   183
      ORELSE  mp_tac 1 ORELSE iff_tac prems 1)) ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   184
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   185
val iff_cong = prove_goal IFOLP.thy 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   186
    "[| p:P <-> P';  q:Q <-> Q' |] ==> ?p:(P<->Q) <-> (P'<->Q')"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   187
 (fn prems =>
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   188
  [ (cut_facts_tac prems 1),
1459
d12da312eff4 expanded tabs
clasohm
parents: 1142
diff changeset
   189
    (REPEAT   (etac iffE 1
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   190
      ORELSE  ares_tac [iffI] 1
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   191
      ORELSE  mp_tac 1)) ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   192
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   193
val not_cong = prove_goal IFOLP.thy 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   194
    "p:P <-> P' ==> ?p:~P <-> ~P'"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   195
 (fn prems =>
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   196
  [ (cut_facts_tac prems 1),
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   197
    (REPEAT   (ares_tac [iffI,notI] 1
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   198
      ORELSE  mp_tac 1
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   199
      ORELSE  eresolve_tac [iffE,notE] 1)) ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   200
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   201
val all_cong = prove_goal IFOLP.thy 
3836
f1a1817659e6 fixed dots;
wenzelm
parents: 1459
diff changeset
   202
    "(!!x. f(x):P(x) <-> Q(x)) ==> ?p:(ALL x. P(x)) <-> (ALL x. Q(x))"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   203
 (fn prems =>
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   204
  [ (REPEAT   (ares_tac [iffI,allI] 1
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   205
      ORELSE   mp_tac 1
1459
d12da312eff4 expanded tabs
clasohm
parents: 1142
diff changeset
   206
      ORELSE   etac allE 1 ORELSE iff_tac prems 1)) ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   207
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   208
val ex_cong = prove_goal IFOLP.thy 
3836
f1a1817659e6 fixed dots;
wenzelm
parents: 1459
diff changeset
   209
    "(!!x. f(x):P(x) <-> Q(x)) ==> ?p:(EX x. P(x)) <-> (EX x. Q(x))"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   210
 (fn prems =>
1459
d12da312eff4 expanded tabs
clasohm
parents: 1142
diff changeset
   211
  [ (REPEAT   (etac exE 1 ORELSE ares_tac [iffI,exI] 1
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   212
      ORELSE   mp_tac 1
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   213
      ORELSE   iff_tac prems 1)) ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   214
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   215
(*NOT PROVED
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   216
val ex1_cong = prove_goal IFOLP.thy 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   217
    "(!!x.f(x):P(x) <-> Q(x)) ==> ?p:(EX! x.P(x)) <-> (EX! x.Q(x))"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   218
 (fn prems =>
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   219
  [ (REPEAT   (eresolve_tac [ex1E, spec RS mp] 1 ORELSE ares_tac [iffI,ex1I] 1
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   220
      ORELSE   mp_tac 1
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   221
      ORELSE   iff_tac prems 1)) ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   222
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   223
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   224
(*** Equality rules ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   225
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   226
val refl = ieqI;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   227
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   228
val subst = prove_goal IFOLP.thy "[| p:a=b;  q:P(a) |] ==> ?p : P(b)"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   229
 (fn [prem1,prem2] => [ rtac (prem2 RS rev_mp) 1, (rtac (prem1 RS ieqE) 1), 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   230
                        rtac impI 1, atac 1 ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   231
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   232
Goal "q:a=b ==> ?c:b=a";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   233
by (etac subst 1);
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   234
by (rtac refl 1) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   235
qed "sym";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   236
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   237
Goal "[| p:a=b;  q:b=c |] ==> ?d:a=c";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   238
by (etac subst 1 THEN assume_tac 1); 
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   239
qed "trans";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   240
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   241
(** ~ b=a ==> ~ a=b **)
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   242
Goal "p:~ b=a ==> ?q:~ a=b";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   243
by (etac contrapos 1);
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   244
by (etac sym 1) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   245
qed "not_sym";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   246
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   247
(*calling "standard" reduces maxidx to 0*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   248
val ssubst = standard (sym RS subst);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   249
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   250
(*A special case of ex1E that would otherwise need quantifier expansion*)
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   251
Goal "[| p:EX! x. P(x);  q:P(a);  r:P(b) |] ==> ?d:a=b";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   252
by (etac ex1E 1);
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   253
by (rtac trans 1);
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   254
by (rtac sym 2);
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   255
by (REPEAT (eresolve_tac [asm_rl, spec RS mp] 1)) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   256
qed "ex1_equalsE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   257
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   258
(** Polymorphic congruence rules **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   259
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   260
Goal "[| p:a=b |]  ==>  ?d:t(a)=t(b)";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   261
by (etac ssubst 1);
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   262
by (rtac refl 1) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   263
qed "subst_context";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   264
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   265
Goal "[| p:a=b;  q:c=d |]  ==>  ?p:t(a,c)=t(b,d)";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   266
by (REPEAT (etac ssubst 1)); 
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   267
by (rtac refl 1) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   268
qed "subst_context2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   269
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   270
Goal "[| p:a=b;  q:c=d;  r:e=f |]  ==>  ?p:t(a,c,e)=t(b,d,f)";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   271
by (REPEAT (etac ssubst 1)); 
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   272
by (rtac refl 1) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   273
qed "subst_context3";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   274
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   275
(*Useful with eresolve_tac for proving equalties from known equalities.
1459
d12da312eff4 expanded tabs
clasohm
parents: 1142
diff changeset
   276
        a = b
d12da312eff4 expanded tabs
clasohm
parents: 1142
diff changeset
   277
        |   |
d12da312eff4 expanded tabs
clasohm
parents: 1142
diff changeset
   278
        c = d   *)
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   279
Goal "[| p:a=b;  q:a=c;  r:b=d |] ==> ?p:c=d";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   280
by (rtac trans 1);
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   281
by (rtac trans 1);
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   282
by (rtac sym 1);
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   283
by (REPEAT (assume_tac 1)) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   284
qed "box_equals";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   285
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   286
(*Dual of box_equals: for proving equalities backwards*)
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   287
Goal "[| p:a=c;  q:b=d;  r:c=d |] ==> ?p:a=b";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   288
by (rtac trans 1);
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   289
by (rtac trans 1);
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   290
by (REPEAT (eresolve_tac [asm_rl, sym] 1)) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   291
qed "simp_equals";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   292
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   293
(** Congruence rules for predicate letters **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   294
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   295
Goal "p:a=a' ==> ?p:P(a) <-> P(a')";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   296
by (rtac iffI 1);
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   297
by (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   298
qed "pred1_cong";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   299
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   300
Goal "[| p:a=a';  q:b=b' |] ==> ?p:P(a,b) <-> P(a',b')";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   301
by (rtac iffI 1);
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   302
by (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   303
qed "pred2_cong";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   304
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   305
Goal "[| p:a=a';  q:b=b';  r:c=c' |] ==> ?p:P(a,b,c) <-> P(a',b',c')";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   306
by (rtac iffI 1);
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   307
by (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   308
qed "pred3_cong";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   309
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   310
(*special cases for free variables P, Q, R, S -- up to 3 arguments*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   311
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   312
val pred_congs = 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   313
    flat (map (fn c => 
1459
d12da312eff4 expanded tabs
clasohm
parents: 1142
diff changeset
   314
               map (fn th => read_instantiate [("P",c)] th)
d12da312eff4 expanded tabs
clasohm
parents: 1142
diff changeset
   315
                   [pred1_cong,pred2_cong,pred3_cong])
d12da312eff4 expanded tabs
clasohm
parents: 1142
diff changeset
   316
               (explode"PQRS"));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   317
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   318
(*special case for the equality predicate!*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   319
val eq_cong = read_instantiate [("P","op =")] pred2_cong;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   320
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   321
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   322
(*** Simplifications of assumed implications.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   323
     Roy Dyckhoff has proved that conj_impE, disj_impE, and imp_impE
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   324
     used with mp_tac (restricted to atomic formulae) is COMPLETE for 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   325
     intuitionistic propositional logic.  See
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   326
   R. Dyckhoff, Contraction-free sequent calculi for intuitionistic logic
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   327
    (preprint, University of St Andrews, 1991)  ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   328
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   329
val major::prems= Goal 
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   330
    "[| p:(P&Q)-->S;  !!x. x:P-->(Q-->S) ==> q(x):R |] ==> ?p:R";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   331
by (REPEAT (ares_tac ([conjI, impI, major RS mp]@prems) 1)) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   332
qed "conj_impE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   333
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   334
val major::prems= Goal 
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   335
    "[| p:(P|Q)-->S;  !!x y.[| x:P-->S; y:Q-->S |] ==> q(x,y):R |] ==> ?p:R";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   336
by (DEPTH_SOLVE (ares_tac ([disjI1, disjI2, impI, major RS mp]@prems) 1)) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   337
qed "disj_impE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   338
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   339
(*Simplifies the implication.  Classical version is stronger. 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   340
  Still UNSAFE since Q must be provable -- backtracking needed.  *)
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   341
val major::prems= Goal 
3836
f1a1817659e6 fixed dots;
wenzelm
parents: 1459
diff changeset
   342
    "[| p:(P-->Q)-->S;  !!x y.[| x:P; y:Q-->S |] ==> q(x,y):Q;  !!x. x:S ==> r(x):R |] ==> \
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   343
\    ?p:R";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   344
by (REPEAT (ares_tac ([impI, major RS mp]@prems) 1)) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   345
qed "imp_impE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   346
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   347
(*Simplifies the implication.  Classical version is stronger. 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   348
  Still UNSAFE since ~P must be provable -- backtracking needed.  *)
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   349
val major::prems= Goal
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   350
    "[| p:~P --> S;  !!y. y:P ==> q(y):False;  !!y. y:S ==> r(y):R |] ==> ?p:R";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   351
by (REPEAT (ares_tac ([notI, impI, major RS mp]@prems) 1)) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   352
qed "not_impE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   353
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   354
(*Simplifies the implication.   UNSAFE.  *)
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   355
val major::prems= Goal 
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   356
    "[| p:(P<->Q)-->S;  !!x y.[| x:P; y:Q-->S |] ==> q(x,y):Q;  \
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   357
\       !!x y.[| x:Q; y:P-->S |] ==> r(x,y):P;  !!x. x:S ==> s(x):R |] ==> ?p:R";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   358
by (REPEAT (ares_tac ([iffI, impI, major RS mp]@prems) 1)) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   359
qed "iff_impE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   360
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   361
(*What if (ALL x.~~P(x)) --> ~~(ALL x.P(x)) is an assumption? UNSAFE*)
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   362
val major::prems= Goal 
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   363
    "[| p:(ALL x. P(x))-->S;  !!x. q:P(x);  !!y. y:S ==> r(y):R |] ==> ?p:R";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   364
by (REPEAT (ares_tac ([allI, impI, major RS mp]@prems) 1)) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   365
qed "all_impE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   366
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   367
(*Unsafe: (EX x.P(x))-->S  is equivalent to  ALL x.P(x)-->S.  *)
9263
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   368
val major::prems= Goal 
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   369
    "[| p:(EX x. P(x))-->S;  !!y. y:P(a)-->S ==> q(y):R |] ==> ?p:R";
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   370
by (REPEAT (ares_tac ([exI, impI, major RS mp]@prems) 1)) ;
53e09e592278 removal of batch style, and tidying
paulson
parents: 3836
diff changeset
   371
qed "ex_impE";