src/ZF/pair.ML
author paulson
Wed, 28 Jun 2000 12:34:08 +0200
changeset 9180 3bda56c0d70d
parent 6153 bff90585cce5
child 9211 6236c5285bd8
permissions -rw-r--r--
tidying and unbatchifying
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
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5264
diff changeset
    21
qed_goalw "Pair_iff" thy [Pair_def]
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    22
    "<a,b> = <c,d> <-> a=c & b=d"
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
    23
 (fn _=> [ (simp_tac (simpset() addsimps [doubleton_eq_iff]) 1),
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    24
           (Blast_tac 1) ]);
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
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5264
diff changeset
    35
qed_goalw "Pair_not_0" thy [Pair_def] "<a,b> ~= 0"
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
    36
 (fn _ => [ (blast_tac (claset() addEs [equalityE]) 1) ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    37
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    38
bind_thm ("Pair_neq_0", Pair_not_0 RS notE);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    39
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    40
AddSEs [Pair_neq_0, sym RS Pair_neq_0];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    41
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5264
diff changeset
    42
qed_goalw "Pair_neq_fst" thy [Pair_def] "<a,b>=a ==> P"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    43
 (fn [major]=>
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
    44
  [ (rtac (consI1 RS mem_asym RS FalseE) 1),
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    45
    (rtac (major RS subst) 1),
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    46
    (rtac consI1 1) ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    47
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5264
diff changeset
    48
qed_goalw "Pair_neq_snd" thy [Pair_def] "<a,b>=b ==> P"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    49
 (fn [major]=>
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
    50
  [ (rtac (consI1 RS consI2 RS mem_asym RS FalseE) 1),
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    51
    (rtac (major RS subst) 1),
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    52
    (rtac (consI1 RS consI2) 1) ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    53
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    54
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    55
(*** Sigma: Disjoint union of a family of sets
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    56
     Generalizes Cartesian product ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    57
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5264
diff changeset
    58
qed_goalw "Sigma_iff" thy [Sigma_def] "<a,b>: Sigma(A,B) <-> a:A & b:B(a)"
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    59
 (fn _ => [ Blast_tac 1 ]);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    60
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    61
Addsimps [Sigma_iff];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    62
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    63
Goal "[| a:A;  b:B(a) |] ==> <a,b> : Sigma(A,B)";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    64
by (Asm_simp_tac 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    65
qed "SigmaI";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    66
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    67
AddTCs [SigmaI];
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    68
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    69
bind_thm ("SigmaD1", Sigma_iff RS iffD1 RS conjunct1);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    70
bind_thm ("SigmaD2", Sigma_iff RS iffD1 RS conjunct2);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    71
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    72
(*The general elimination rule*)
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5264
diff changeset
    73
qed_goalw "SigmaE" thy [Sigma_def]
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    74
    "[| c: Sigma(A,B);  \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    75
\       !!x y.[| x:A;  y:B(x);  c=<x,y> |] ==> P \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    76
\    |] ==> P"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    77
 (fn major::prems=>
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    78
  [ (cut_facts_tac [major] 1),
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    79
    (REPEAT (eresolve_tac [UN_E, singletonE] 1 ORELSE ares_tac prems 1)) ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    80
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    81
val [major,minor]= Goal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    82
    "[| <a,b> : Sigma(A,B);    \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    83
\       [| a:A;  b:B(a) |] ==> P   \
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    84
\    |] ==> P";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    85
by (rtac minor 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    86
by (rtac (major RS SigmaD1) 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    87
by (rtac (major RS SigmaD2) 1) ;
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
    88
qed "SigmaE2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    89
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5264
diff changeset
    90
qed_goalw "Sigma_cong" thy [Sigma_def]
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    91
    "[| A=A';  !!x. x:A' ==> B(x)=B'(x) |] ==> \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    92
\    Sigma(A,B) = Sigma(A',B')"
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
    93
 (fn prems=> [ (simp_tac (simpset() addsimps prems) 1) ]);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    94
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    95
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    96
(*Sigma_cong, Pi_cong NOT given to Addcongs: they cause
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    97
  flex-flex pairs and the "Check your prover" error.  Most
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    98
  Sigmas and Pis are abbreviated as * or -> *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    99
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   100
AddSIs [SigmaI];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   101
AddSEs [SigmaE2, SigmaE];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   102
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   103
Goal "Sigma(0,B) = 0";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   104
by (Blast_tac 1) ;
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   105
qed "Sigma_empty1";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   106
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   107
Goal "A*0 = 0";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   108
by (Blast_tac 1) ;
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   109
qed "Sigma_empty2";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   110
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   111
Addsimps [Sigma_empty1, Sigma_empty2];
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   112
5264
7538fce1fe37 New result from AC directory
paulson
parents: 4477
diff changeset
   113
Goal "A*B=0 <-> A=0 | B=0";
7538fce1fe37 New result from AC directory
paulson
parents: 4477
diff changeset
   114
by (Blast_tac 1);
7538fce1fe37 New result from AC directory
paulson
parents: 4477
diff changeset
   115
qed "Sigma_empty_iff";
7538fce1fe37 New result from AC directory
paulson
parents: 4477
diff changeset
   116
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   117
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   118
(*** Projections: fst, snd ***)
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   119
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5264
diff changeset
   120
qed_goalw "fst_conv" thy [fst_def] "fst(<a,b>) = a"
5505
paulson
parents: 5325
diff changeset
   121
 (fn _=> [ (Blast_tac 1) ]);
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   122
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5264
diff changeset
   123
qed_goalw "snd_conv" thy [snd_def] "snd(<a,b>) = b"
5505
paulson
parents: 5325
diff changeset
   124
 (fn _=> [ (Blast_tac 1) ]);
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   125
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   126
Addsimps [fst_conv,snd_conv];
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   127
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   128
Goal "p:Sigma(A,B) ==> fst(p) : A";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   129
by (Auto_tac) ;
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   130
qed "fst_type";
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   131
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   132
Goal "p:Sigma(A,B) ==> snd(p) : B(fst(p))";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   133
by (Auto_tac) ;
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   134
qed "snd_type";
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   135
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   136
Goal "a: Sigma(A,B) ==> <fst(a),snd(a)> = a";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   137
by (Auto_tac) ;
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   138
qed "Pair_fst_snd_eq";
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   139
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   140
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   141
(*** Eliminator - split ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   142
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   143
(*A META-equality, so that it applies to higher types as well...*)
6112
5e4871c5136b datatype package improvements
paulson
parents: 5505
diff changeset
   144
Goalw [split_def] "split(%x y. c(x,y), <a,b>) == c(a,b)";
5e4871c5136b datatype package improvements
paulson
parents: 5505
diff changeset
   145
by (Simp_tac 1);
5e4871c5136b datatype package improvements
paulson
parents: 5505
diff changeset
   146
qed "split";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   147
Addsimps [split];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   148
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   149
val major::prems= Goal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   150
    "[|  p:Sigma(A,B);   \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   151
\        !!x y.[| x:A; y:B(x) |] ==> c(x,y):C(<x,y>) \
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   152
\    |] ==> split(%x y. c(x,y), p) : C(p)";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   153
by (rtac (major RS SigmaE) 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   154
by (asm_simp_tac (simpset() addsimps prems) 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 6153
diff changeset
   155
qed "split_type";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   156
AddTCs [split_type];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   157
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5264
diff changeset
   158
Goalw [split_def]
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5264
diff changeset
   159
  "u: A*B ==>   \
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   160
\       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
   161
by Auto_tac;
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 533
diff changeset
   162
qed "expand_split";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   163
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 6
diff changeset
   164
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   165
(*** split for predicates: result type o ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   166
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5264
diff changeset
   167
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
   168
by (Asm_simp_tac 1);
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   169
qed "splitI";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   170
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5264
diff changeset
   171
val major::sigma::prems = Goalw [split_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1105
diff changeset
   172
    "[| split(R,z);  z:Sigma(A,B);                      \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1105
diff changeset
   173
\       !!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
   174
\    |] ==> P";
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   175
by (rtac (sigma RS SigmaE) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   176
by (cut_facts_tac [major] 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   177
by (REPEAT (ares_tac prems 1));
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   178
by (Asm_full_simp_tac 1);
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   179
qed "splitE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   180
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5264
diff changeset
   181
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
   182
by (Full_simp_tac 1);
1105
136b05aa77ed Modified proofs for (q)split, fst, snd for new
lcp
parents: 860
diff changeset
   183
qed "splitD";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   184
533
7357160bc56a ZF/pair.ML: moved some definitions here from simpdata.ML
lcp
parents: 437
diff changeset
   185
7357160bc56a ZF/pair.ML: moved some definitions here from simpdata.ML
lcp
parents: 437
diff changeset
   186