| author | wenzelm | 
| Fri, 04 Aug 2000 22:53:44 +0200 | |
| changeset 9523 | 232b09dba0fe | 
| parent 9211 | 6236c5285bd8 | 
| child 9907 | 473a6604da94 | 
| permissions | -rw-r--r-- | 
| 9211 | 1 | (* Title: ZF/QUniv | 
| 0 | 2 | ID: $Id$ | 
| 1461 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 0 | 4 | Copyright 1993 University of Cambridge | 
| 5 | ||
| 5809 | 6 | A small universe for lazy recursive types | 
| 0 | 7 | *) | 
| 8 | ||
| 838 
120edb26ee93
Moved Transset_includes_summands and Transset_sum_Int_subset
 lcp parents: 
803diff
changeset | 9 | (** Properties involving Transset and Sum **) | 
| 
120edb26ee93
Moved Transset_includes_summands and Transset_sum_Int_subset
 lcp parents: 
803diff
changeset | 10 | |
| 9211 | 11 | Goalw [sum_def] "[| Transset(C); A+B <= C |] ==> A <= C & B <= C"; | 
| 12 | by (dtac (Un_subset_iff RS iffD1) 1); | |
| 13 | by (blast_tac (claset() addDs [Transset_includes_range]) 1); | |
| 838 
120edb26ee93
Moved Transset_includes_summands and Transset_sum_Int_subset
 lcp parents: 
803diff
changeset | 14 | qed "Transset_includes_summands"; | 
| 
120edb26ee93
Moved Transset_includes_summands and Transset_sum_Int_subset
 lcp parents: 
803diff
changeset | 15 | |
| 9211 | 16 | Goalw [sum_def] "Transset(C) ==> (A+B) Int C <= (A Int C) + (B Int C)"; | 
| 2033 | 17 | by (stac Int_Un_distrib 1); | 
| 9211 | 18 | by (blast_tac (claset() addDs [Transset_Pair_D]) 1); | 
| 838 
120edb26ee93
Moved Transset_includes_summands and Transset_sum_Int_subset
 lcp parents: 
803diff
changeset | 19 | qed "Transset_sum_Int_subset"; | 
| 
120edb26ee93
Moved Transset_includes_summands and Transset_sum_Int_subset
 lcp parents: 
803diff
changeset | 20 | |
| 0 | 21 | (** Introduction and elimination rules avoid tiresome folding/unfolding **) | 
| 22 | ||
| 5067 | 23 | Goalw [quniv_def] | 
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5137diff
changeset | 24 | "X <= univ(eclose(A)) ==> X : quniv(A)"; | 
| 15 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 25 | by (etac PowI 1); | 
| 760 | 26 | qed "qunivI"; | 
| 0 | 27 | |
| 5067 | 28 | Goalw [quniv_def] | 
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5137diff
changeset | 29 | "X : quniv(A) ==> X <= univ(eclose(A))"; | 
| 15 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 30 | by (etac PowD 1); | 
| 760 | 31 | qed "qunivD"; | 
| 0 | 32 | |
| 5137 | 33 | Goalw [quniv_def] "A<=B ==> quniv(A) <= quniv(B)"; | 
| 0 | 34 | by (etac (eclose_mono RS univ_mono RS Pow_mono) 1); | 
| 760 | 35 | qed "quniv_mono"; | 
| 0 | 36 | |
| 37 | (*** Closure properties ***) | |
| 38 | ||
| 5067 | 39 | Goalw [quniv_def] "univ(eclose(A)) <= quniv(A)"; | 
| 0 | 40 | by (rtac (Transset_iff_Pow RS iffD1) 1); | 
| 41 | by (rtac (Transset_eclose RS Transset_univ) 1); | |
| 760 | 42 | qed "univ_eclose_subset_quniv"; | 
| 0 | 43 | |
| 170 
590c9d1e0d73
ZF/quniv/QPair_Int_Vset_subset_UN: new, isolates key argument of many
 lcp parents: 
129diff
changeset | 44 | (*Key property for proving A_subset_quniv; requires eclose in def of quniv*) | 
| 5067 | 45 | Goal "univ(A) <= quniv(A)"; | 
| 0 | 46 | by (rtac (arg_subset_eclose RS univ_mono RS subset_trans) 1); | 
| 47 | by (rtac univ_eclose_subset_quniv 1); | |
| 760 | 48 | qed "univ_subset_quniv"; | 
| 0 | 49 | |
| 803 
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
 lcp parents: 
782diff
changeset | 50 | bind_thm ("univ_into_quniv", univ_subset_quniv RS subsetD);
 | 
| 0 | 51 | |
| 5067 | 52 | Goalw [quniv_def] "Pow(univ(A)) <= quniv(A)"; | 
| 0 | 53 | by (rtac (arg_subset_eclose RS univ_mono RS Pow_mono) 1); | 
| 760 | 54 | qed "Pow_univ_subset_quniv"; | 
| 0 | 55 | |
| 803 
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
 lcp parents: 
782diff
changeset | 56 | bind_thm ("univ_subset_into_quniv", 
 | 
| 1461 | 57 | PowI RS (Pow_univ_subset_quniv RS subsetD)); | 
| 0 | 58 | |
| 803 
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
 lcp parents: 
782diff
changeset | 59 | bind_thm ("zero_in_quniv", zero_in_univ RS univ_into_quniv);
 | 
| 
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
 lcp parents: 
782diff
changeset | 60 | bind_thm ("one_in_quniv", one_in_univ RS univ_into_quniv);
 | 
| 
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
 lcp parents: 
782diff
changeset | 61 | bind_thm ("two_in_quniv", two_in_univ RS univ_into_quniv);
 | 
| 0 | 62 | |
| 803 
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
 lcp parents: 
782diff
changeset | 63 | bind_thm ("A_subset_quniv",
 | 
| 1461 | 64 | [A_subset_univ, univ_subset_quniv] MRS subset_trans); | 
| 0 | 65 | |
| 66 | val A_into_quniv = A_subset_quniv RS subsetD; | |
| 67 | ||
| 68 | (*** univ(A) closure for Quine-inspired pairs and injections ***) | |
| 69 | ||
| 70 | (*Quine ordered pairs*) | |
| 5067 | 71 | Goalw [QPair_def] | 
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5137diff
changeset | 72 | "[| a <= univ(A); b <= univ(A) |] ==> <a;b> <= univ(A)"; | 
| 0 | 73 | by (REPEAT (ares_tac [sum_subset_univ] 1)); | 
| 760 | 74 | qed "QPair_subset_univ"; | 
| 0 | 75 | |
| 76 | (** Quine disjoint sum **) | |
| 77 | ||
| 5137 | 78 | Goalw [QInl_def] "a <= univ(A) ==> QInl(a) <= univ(A)"; | 
| 0 | 79 | by (etac (empty_subsetI RS QPair_subset_univ) 1); | 
| 760 | 80 | qed "QInl_subset_univ"; | 
| 0 | 81 | |
| 82 | val naturals_subset_nat = | |
| 83 | rewrite_rule [Transset_def] (Ord_nat RS Ord_is_Transset) | |
| 84 | RS bspec; | |
| 85 | ||
| 86 | val naturals_subset_univ = | |
| 87 | [naturals_subset_nat, nat_subset_univ] MRS subset_trans; | |
| 88 | ||
| 5137 | 89 | Goalw [QInr_def] "a <= univ(A) ==> QInr(a) <= univ(A)"; | 
| 0 | 90 | by (etac (nat_1I RS naturals_subset_univ RS QPair_subset_univ) 1); | 
| 760 | 91 | qed "QInr_subset_univ"; | 
| 0 | 92 | |
| 93 | (*** Closure for Quine-inspired products and sums ***) | |
| 94 | ||
| 95 | (*Quine ordered pairs*) | |
| 5067 | 96 | Goalw [quniv_def,QPair_def] | 
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5137diff
changeset | 97 | "[| a: quniv(A); b: quniv(A) |] ==> <a;b> : quniv(A)"; | 
| 0 | 98 | by (REPEAT (dtac PowD 1)); | 
| 99 | by (REPEAT (ares_tac [PowI, sum_subset_univ] 1)); | |
| 760 | 100 | qed "QPair_in_quniv"; | 
| 0 | 101 | |
| 5067 | 102 | Goal "quniv(A) <*> quniv(A) <= quniv(A)"; | 
| 0 | 103 | by (REPEAT (ares_tac [subsetI, QPair_in_quniv] 1 | 
| 104 | ORELSE eresolve_tac [QSigmaE, ssubst] 1)); | |
| 760 | 105 | qed "QSigma_quniv"; | 
| 0 | 106 | |
| 803 
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
 lcp parents: 
782diff
changeset | 107 | bind_thm ("QSigma_subset_quniv",
 | 
| 1461 | 108 | [QSigma_mono, QSigma_quniv] MRS subset_trans); | 
| 0 | 109 | |
| 110 | (*The opposite inclusion*) | |
| 5067 | 111 | Goalw [quniv_def,QPair_def] | 
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5137diff
changeset | 112 | "<a;b> : quniv(A) ==> a: quniv(A) & b: quniv(A)"; | 
| 129 | 113 | by (etac ([Transset_eclose RS Transset_univ, PowD] MRS | 
| 1461 | 114 | Transset_includes_summands RS conjE) 1); | 
| 0 | 115 | by (REPEAT (ares_tac [conjI,PowI] 1)); | 
| 760 | 116 | qed "quniv_QPair_D"; | 
| 0 | 117 | |
| 803 
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
 lcp parents: 
782diff
changeset | 118 | bind_thm ("quniv_QPair_E", quniv_QPair_D RS conjE);
 | 
| 0 | 119 | |
| 5067 | 120 | Goal "<a;b> : quniv(A) <-> a: quniv(A) & b: quniv(A)"; | 
| 0 | 121 | by (REPEAT (ares_tac [iffI, QPair_in_quniv, quniv_QPair_D] 1 | 
| 122 | ORELSE etac conjE 1)); | |
| 760 | 123 | qed "quniv_QPair_iff"; | 
| 0 | 124 | |
| 125 | ||
| 126 | (** Quine disjoint sum **) | |
| 127 | ||
| 5137 | 128 | Goalw [QInl_def] "a: quniv(A) ==> QInl(a) : quniv(A)"; | 
| 0 | 129 | by (REPEAT (ares_tac [zero_in_quniv,QPair_in_quniv] 1)); | 
| 760 | 130 | qed "QInl_in_quniv"; | 
| 0 | 131 | |
| 5137 | 132 | Goalw [QInr_def] "b: quniv(A) ==> QInr(b) : quniv(A)"; | 
| 0 | 133 | by (REPEAT (ares_tac [one_in_quniv, QPair_in_quniv] 1)); | 
| 760 | 134 | qed "QInr_in_quniv"; | 
| 0 | 135 | |
| 5067 | 136 | Goal "quniv(C) <+> quniv(C) <= quniv(C)"; | 
| 0 | 137 | by (REPEAT (ares_tac [subsetI, QInl_in_quniv, QInr_in_quniv] 1 | 
| 138 | ORELSE eresolve_tac [qsumE, ssubst] 1)); | |
| 760 | 139 | qed "qsum_quniv"; | 
| 0 | 140 | |
| 803 
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
 lcp parents: 
782diff
changeset | 141 | bind_thm ("qsum_subset_quniv", [qsum_mono, qsum_quniv] MRS subset_trans);
 | 
| 
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
 lcp parents: 
782diff
changeset | 142 | |
| 0 | 143 | |
| 144 | (*** The natural numbers ***) | |
| 145 | ||
| 803 
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
 lcp parents: 
782diff
changeset | 146 | bind_thm ("nat_subset_quniv",
 | 
| 1461 | 147 | [nat_subset_univ, univ_subset_quniv] MRS subset_trans); | 
| 0 | 148 | |
| 149 | (* n:nat ==> n:quniv(A) *) | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 150 | bind_thm ("nat_into_quniv", (nat_subset_quniv RS subsetD));
 | 
| 0 | 151 | |
| 803 
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
 lcp parents: 
782diff
changeset | 152 | bind_thm ("bool_subset_quniv",
 | 
| 1461 | 153 | [bool_subset_univ, univ_subset_quniv] MRS subset_trans); | 
| 0 | 154 | |
| 803 
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
 lcp parents: 
782diff
changeset | 155 | bind_thm ("bool_into_quniv", bool_subset_quniv RS subsetD);
 | 
| 0 | 156 | |
| 157 | ||
| 158 | (*** Intersecting <a;b> with Vfrom... ***) | |
| 159 | ||
| 5067 | 160 | Goalw [QPair_def,sum_def] | 
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5137diff
changeset | 161 | "Transset(X) ==> \ | 
| 0 | 162 | \ <a;b> Int Vfrom(X, succ(i)) <= <a Int Vfrom(X,i); b Int Vfrom(X,i)>"; | 
| 2033 | 163 | by (stac Int_Un_distrib 1); | 
| 15 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 164 | by (rtac Un_mono 1); | 
| 0 | 165 | by (REPEAT (ares_tac [product_Int_Vfrom_subset RS subset_trans, | 
| 1461 | 166 | [Int_lower1, subset_refl] MRS Sigma_mono] 1)); | 
| 760 | 167 | qed "QPair_Int_Vfrom_succ_subset"; | 
| 0 | 168 | |
| 170 
590c9d1e0d73
ZF/quniv/QPair_Int_Vset_subset_UN: new, isolates key argument of many
 lcp parents: 
129diff
changeset | 169 | (**** "Take-lemma" rules for proving a=b by coinduction and c: quniv(A) ****) | 
| 0 | 170 | |
| 170 
590c9d1e0d73
ZF/quniv/QPair_Int_Vset_subset_UN: new, isolates key argument of many
 lcp parents: 
129diff
changeset | 171 | (*Rule for level i -- preserving the level, not decreasing it*) | 
| 0 | 172 | |
| 5067 | 173 | Goalw [QPair_def] | 
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5137diff
changeset | 174 | "Transset(X) ==> \ | 
| 0 | 175 | \ <a;b> Int Vfrom(X,i) <= <a Int Vfrom(X,i); b Int Vfrom(X,i)>"; | 
| 15 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 176 | by (etac (Transset_Vfrom RS Transset_sum_Int_subset) 1); | 
| 760 | 177 | qed "QPair_Int_Vfrom_subset"; | 
| 0 | 178 | |
| 170 
590c9d1e0d73
ZF/quniv/QPair_Int_Vset_subset_UN: new, isolates key argument of many
 lcp parents: 
129diff
changeset | 179 | (*[| a Int Vset(i) <= c; b Int Vset(i) <= d |] ==> <a;b> Int Vset(i) <= <c;d>*) | 
| 803 
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
 lcp parents: 
782diff
changeset | 180 | bind_thm ("QPair_Int_Vset_subset_trans", 
 | 
| 1461 | 181 | [Transset_0 RS QPair_Int_Vfrom_subset, QPair_mono] MRS subset_trans); | 
| 0 | 182 | |
| 5268 | 183 | Goal "Ord(i) ==> <a;b> Int Vset(i) <= (UN j:i. <a Int Vset(j); b Int Vset(j)>)"; | 
| 170 
590c9d1e0d73
ZF/quniv/QPair_Int_Vset_subset_UN: new, isolates key argument of many
 lcp parents: 
129diff
changeset | 184 | by (etac Ord_cases 1 THEN REPEAT_FIRST hyp_subst_tac); | 
| 
590c9d1e0d73
ZF/quniv/QPair_Int_Vset_subset_UN: new, isolates key argument of many
 lcp parents: 
129diff
changeset | 185 | (*0 case*) | 
| 2033 | 186 | by (stac Vfrom_0 1); | 
| 2469 | 187 | by (Fast_tac 1); | 
| 170 
590c9d1e0d73
ZF/quniv/QPair_Int_Vset_subset_UN: new, isolates key argument of many
 lcp parents: 
129diff
changeset | 188 | (*succ(j) case*) | 
| 
590c9d1e0d73
ZF/quniv/QPair_Int_Vset_subset_UN: new, isolates key argument of many
 lcp parents: 
129diff
changeset | 189 | by (rtac (Transset_0 RS QPair_Int_Vfrom_succ_subset RS subset_trans) 1); | 
| 
590c9d1e0d73
ZF/quniv/QPair_Int_Vset_subset_UN: new, isolates key argument of many
 lcp parents: 
129diff
changeset | 190 | by (rtac (succI1 RS UN_upper) 1); | 
| 
590c9d1e0d73
ZF/quniv/QPair_Int_Vset_subset_UN: new, isolates key argument of many
 lcp parents: 
129diff
changeset | 191 | (*Limit(i) case*) | 
| 5268 | 192 | by (asm_simp_tac | 
| 6070 | 193 | (simpset() addsimps [Limit_Vfrom_eq, Int_UN_distrib, | 
| 5268 | 194 | UN_mono, QPair_Int_Vset_subset_trans]) 1); | 
| 760 | 195 | qed "QPair_Int_Vset_subset_UN"; |