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