src/ZF/pair.ML
author nipkow
Mon, 06 Aug 2001 13:43:24 +0200
changeset 11464 ddea204de5bc
parent 9211 6236c5285bd8
child 12199 8213fd95acb5
permissions -rw-r--r--
turned translation for 1::nat into def. introduced 1' and replaced most occurrences of 1 by 1'.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1105
diff changeset
     1
(*  Title:      ZF/pair
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1105
diff changeset
     3
    Author:     Lawrence C Paulson, 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
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
Ordered pairs in Zermelo-Fraenkel Set Theory 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
(** Lemmas for showing that <a,b> uniquely determines a and b **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    11
Goal "{a} = {b} <-> a=b";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    12
by (resolve_tac [extension RS iff_trans] 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    13
by (Blast_tac 1) ;
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    14
qed "singleton_eq_iff";
822
8d5748202c53 Added Krzysztof's theorems singleton_eq_iff, fst_type, snd_type
lcp
parents: 782
diff changeset
    15
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    16
Goal "{a,b} = {c,d} <-> (a=c & b=d) | (a=d & b=c)";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    17
by (resolve_tac [extension RS iff_trans] 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    18
by (Blast_tac 1) ;
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    19
qed "doubleton_eq_iff";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    20
9211
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    21
Goalw [Pair_def] "<a,b> = <c,d> <-> a=c & b=d";
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    22
by (simp_tac (simpset() addsimps [doubleton_eq_iff]) 1);
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    23
by (Blast_tac 1) ;
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    24
qed "Pair_iff";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    25
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    26
Addsimps [Pair_iff];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    27
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    28
bind_thm ("Pair_inject", Pair_iff RS iffD1 RS conjE);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    29
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    30
AddSEs [Pair_inject];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    31
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    32
bind_thm ("Pair_inject1", Pair_iff RS iffD1 RS conjunct1);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    33
bind_thm ("Pair_inject2", Pair_iff RS iffD1 RS conjunct2);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    34
9211
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    35
Goalw [Pair_def]  "<a,b> ~= 0";
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    36
by (blast_tac (claset() addEs [equalityE]) 1) ;
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    37
qed "Pair_not_0";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    38
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    39
bind_thm ("Pair_neq_0", Pair_not_0 RS notE);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    40
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    41
AddSEs [Pair_neq_0, sym RS Pair_neq_0];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    42
9211
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    43
Goalw [Pair_def]  "<a,b>=a ==> P";
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    44
by (rtac (consI1 RS mem_asym RS FalseE) 1);
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    45
by (etac subst 1);
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    46
by (rtac consI1 1) ;
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    47
qed "Pair_neq_fst";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    48
9211
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    49
Goalw [Pair_def]  "<a,b>=b ==> P";
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    50
by (rtac (consI1 RS consI2 RS mem_asym RS FalseE) 1);
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    51
by (etac subst 1);
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    52
by (rtac (consI1 RS consI2) 1) ;
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    53
qed "Pair_neq_snd";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    54
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    55
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    56
(*** Sigma: Disjoint union of a family of sets
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    57
     Generalizes Cartesian product ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    58
9211
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    59
Goalw [Sigma_def]  "<a,b>: Sigma(A,B) <-> a:A & b:B(a)";
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    60
by (Blast_tac 1) ;
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    61
qed "Sigma_iff";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    62
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    63
Addsimps [Sigma_iff];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    64
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    65
Goal "[| a:A;  b:B(a) |] ==> <a,b> : Sigma(A,B)";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    66
by (Asm_simp_tac 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    67
qed "SigmaI";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    68
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    69
AddTCs [SigmaI];
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    70
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    71
bind_thm ("SigmaD1", Sigma_iff RS iffD1 RS conjunct1);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    72
bind_thm ("SigmaD2", Sigma_iff RS iffD1 RS conjunct2);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    73
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    74
(*The general elimination rule*)
9211
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    75
val major::prems= Goalw [Sigma_def] 
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    76
    "[| c: Sigma(A,B);  \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    77
\       !!x y.[| x:A;  y:B(x);  c=<x,y> |] ==> P \
9211
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    78
\    |] ==> P";
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    79
by (cut_facts_tac [major] 1);
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    80
by (REPEAT (eresolve_tac [UN_E, singletonE] 1 ORELSE ares_tac prems 1)) ;
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    81
qed "SigmaE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    82
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    83
val [major,minor]= Goal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    84
    "[| <a,b> : Sigma(A,B);    \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    85
\       [| a:A;  b:B(a) |] ==> P   \
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    86
\    |] ==> P";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    87
by (rtac minor 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    88
by (rtac (major RS SigmaD1) 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    89
by (rtac (major RS SigmaD2) 1) ;
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    90
qed "SigmaE2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    91
9211
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    92
val prems= Goalw [Sigma_def] 
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    93
    "[| A=A';  !!x. x:A' ==> B(x)=B'(x) |] ==> \
9211
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    94
\    Sigma(A,B) = Sigma(A',B')";
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    95
by (simp_tac (simpset() addsimps prems) 1) ;
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    96
qed "Sigma_cong";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    97
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    98
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    99
(*Sigma_cong, Pi_cong NOT given to Addcongs: they cause
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   100
  flex-flex pairs and the "Check your prover" error.  Most
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   101
  Sigmas and Pis are abbreviated as * or -> *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   102
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   103
AddSIs [SigmaI];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   104
AddSEs [SigmaE2, SigmaE];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   105
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   106
Goal "Sigma(0,B) = 0";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   107
by (Blast_tac 1) ;
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   108
qed "Sigma_empty1";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   109
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   110
Goal "A*0 = 0";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   111
by (Blast_tac 1) ;
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   112
qed "Sigma_empty2";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   113
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   114
Addsimps [Sigma_empty1, Sigma_empty2];
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   115
5264
7538fce1fe37 New result from AC directory
paulson
parents: 4477
diff changeset
   116
Goal "A*B=0 <-> A=0 | B=0";
7538fce1fe37 New result from AC directory
paulson
parents: 4477
diff changeset
   117
by (Blast_tac 1);
7538fce1fe37 New result from AC directory
paulson
parents: 4477
diff changeset
   118
qed "Sigma_empty_iff";
7538fce1fe37 New result from AC directory
paulson
parents: 4477
diff changeset
   119
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   120
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   121
(*** Projections: fst, snd ***)
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   122
9211
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
   123
Goalw [fst_def]  "fst(<a,b>) = a";
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
   124
by (Blast_tac 1) ;
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
   125
qed "fst_conv";
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   126
9211
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
   127
Goalw [snd_def]  "snd(<a,b>) = b";
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
   128
by (Blast_tac 1) ;
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
   129
qed "snd_conv";
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   130
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   131
Addsimps [fst_conv,snd_conv];
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   132
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   133
Goal "p:Sigma(A,B) ==> fst(p) : A";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   134
by (Auto_tac) ;
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   135
qed "fst_type";
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   136
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   137
Goal "p:Sigma(A,B) ==> snd(p) : B(fst(p))";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   138
by (Auto_tac) ;
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   139
qed "snd_type";
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   140
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   141
Goal "a: Sigma(A,B) ==> <fst(a),snd(a)> = a";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   142
by (Auto_tac) ;
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   143
qed "Pair_fst_snd_eq";
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   144
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   145
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   146
(*** Eliminator - split ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   147
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   148
(*A META-equality, so that it applies to higher types as well...*)
6112
5e4871c5136b datatype package improvements
paulson
parents: 5505
diff changeset
   149
Goalw [split_def] "split(%x y. c(x,y), <a,b>) == c(a,b)";
5e4871c5136b datatype package improvements
paulson
parents: 5505
diff changeset
   150
by (Simp_tac 1);
5e4871c5136b datatype package improvements
paulson
parents: 5505
diff changeset
   151
qed "split";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   152
Addsimps [split];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   153
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   154
val major::prems= Goal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   155
    "[|  p:Sigma(A,B);   \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   156
\        !!x y.[| x:A; y:B(x) |] ==> c(x,y):C(<x,y>) \
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   157
\    |] ==> split(%x y. c(x,y), p) : C(p)";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   158
by (rtac (major RS SigmaE) 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   159
by (asm_simp_tac (simpset() addsimps prems) 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   160
qed "split_type";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   161
AddTCs [split_type];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   162
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5264
diff changeset
   163
Goalw [split_def]
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5264
diff changeset
   164
  "u: A*B ==>   \
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   165
\       R(split(c,u)) <-> (ALL x:A. ALL y:B. u = <x,y> --> R(c(x,y)))";
4477
b3e5857d8d99 New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents: 4091
diff changeset
   166
by Auto_tac;
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 533
diff changeset
   167
qed "expand_split";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   168
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   169
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   170
(*** split for predicates: result type o ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   171
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5264
diff changeset
   172
Goalw [split_def] "R(a,b) ==> split(R, <a,b>)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   173
by (Asm_simp_tac 1);
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   174
qed "splitI";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   175
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5264
diff changeset
   176
val major::sigma::prems = Goalw [split_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1105
diff changeset
   177
    "[| split(R,z);  z:Sigma(A,B);                      \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1105
diff changeset
   178
\       !!x y. [| z = <x,y>;  R(x,y) |] ==> P           \
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   179
\    |] ==> P";
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   180
by (rtac (sigma RS SigmaE) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   181
by (cut_facts_tac [major] 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   182
by (REPEAT (ares_tac prems 1));
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   183
by (Asm_full_simp_tac 1);
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   184
qed "splitE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   185
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5264
diff changeset
   186
Goalw [split_def] "split(R,<a,b>) ==> R(a,b)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   187
by (Full_simp_tac 1);
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   188
qed "splitD";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   189
533
7357160bc56a ZF/pair.ML: moved some definitions here from simpdata.ML
lcp
parents: 437
diff changeset
   190
7357160bc56a ZF/pair.ML: moved some definitions here from simpdata.ML
lcp
parents: 437
diff changeset
   191