| author | paulson | 
| Mon, 23 Sep 1996 18:12:45 +0200 | |
| changeset 2009 | 9023e474d22a | 
| parent 1985 | 84cf16192e03 | 
| child 2891 | d8f254ad1ab9 | 
| permissions | -rw-r--r-- | 
| 1465 | 1 | (* Title: HOL/univ | 
| 923 | 2 | ID: $Id$ | 
| 1465 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 923 | 4 | Copyright 1991 University of Cambridge | 
| 5 | ||
| 6 | For univ.thy | |
| 7 | *) | |
| 8 | ||
| 9 | open Univ; | |
| 10 | ||
| 11 | (** apfst -- can be used in similar type definitions **) | |
| 12 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 13 | goalw Univ.thy [apfst_def] "apfst f (a,b) = (f(a),b)"; | 
| 923 | 14 | by (rtac split 1); | 
| 976 
14b55f7fbf15
renamed theorem "apfst" to "apfst_conv" to avoid conflict with function
 clasohm parents: 
972diff
changeset | 15 | qed "apfst_conv"; | 
| 923 | 16 | |
| 17 | val [major,minor] = goal Univ.thy | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 18 | "[| q = apfst f p; !!x y. [| p = (x,y); q = (f(x),y) |] ==> R \ | 
| 923 | 19 | \ |] ==> R"; | 
| 20 | by (rtac PairE 1); | |
| 21 | by (rtac minor 1); | |
| 22 | by (assume_tac 1); | |
| 23 | by (rtac (major RS trans) 1); | |
| 24 | by (etac ssubst 1); | |
| 976 
14b55f7fbf15
renamed theorem "apfst" to "apfst_conv" to avoid conflict with function
 clasohm parents: 
972diff
changeset | 25 | by (rtac apfst_conv 1); | 
| 
14b55f7fbf15
renamed theorem "apfst" to "apfst_conv" to avoid conflict with function
 clasohm parents: 
972diff
changeset | 26 | qed "apfst_convE"; | 
| 923 | 27 | |
| 28 | (** Push -- an injection, analogous to Cons on lists **) | |
| 29 | ||
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 30 | val [major] = goalw Univ.thy [Push_def] "Push i f = Push j g ==> i=j"; | 
| 923 | 31 | by (rtac (major RS fun_cong RS box_equals RS Suc_inject) 1); | 
| 32 | by (rtac nat_case_0 1); | |
| 33 | by (rtac nat_case_0 1); | |
| 34 | qed "Push_inject1"; | |
| 35 | ||
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 36 | val [major] = goalw Univ.thy [Push_def] "Push i f = Push j g ==> f=g"; | 
| 923 | 37 | by (rtac (major RS fun_cong RS ext RS box_equals) 1); | 
| 38 | by (rtac (nat_case_Suc RS ext) 1); | |
| 39 | by (rtac (nat_case_Suc RS ext) 1); | |
| 40 | qed "Push_inject2"; | |
| 41 | ||
| 42 | val [major,minor] = goal Univ.thy | |
| 43 | "[| Push i f =Push j g; [| i=j; f=g |] ==> P \ | |
| 44 | \ |] ==> P"; | |
| 45 | by (rtac ((major RS Push_inject2) RS ((major RS Push_inject1) RS minor)) 1); | |
| 46 | qed "Push_inject"; | |
| 47 | ||
| 48 | val [major] = goalw Univ.thy [Push_def] "Push k f =(%z.0) ==> P"; | |
| 49 | by (rtac (major RS fun_cong RS box_equals RS Suc_neq_Zero) 1); | |
| 50 | by (rtac nat_case_0 1); | |
| 51 | by (rtac refl 1); | |
| 52 | qed "Push_neq_K0"; | |
| 53 | ||
| 54 | (*** Isomorphisms ***) | |
| 55 | ||
| 56 | goal Univ.thy "inj(Rep_Node)"; | |
| 1465 | 57 | by (rtac inj_inverseI 1); (*cannot combine by RS: multiple unifiers*) | 
| 923 | 58 | by (rtac Rep_Node_inverse 1); | 
| 59 | qed "inj_Rep_Node"; | |
| 60 | ||
| 61 | goal Univ.thy "inj_onto Abs_Node Node"; | |
| 62 | by (rtac inj_onto_inverseI 1); | |
| 63 | by (etac Abs_Node_inverse 1); | |
| 64 | qed "inj_onto_Abs_Node"; | |
| 65 | ||
| 66 | val Abs_Node_inject = inj_onto_Abs_Node RS inj_ontoD; | |
| 67 | ||
| 68 | ||
| 69 | (*** Introduction rules for Node ***) | |
| 70 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 71 | goalw Univ.thy [Node_def] "(%k. 0,a) : Node"; | 
| 1760 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 72 | by (Fast_tac 1); | 
| 923 | 73 | qed "Node_K0_I"; | 
| 74 | ||
| 75 | goalw Univ.thy [Node_def,Push_def] | |
| 76 | "!!p. p: Node ==> apfst (Push i) p : Node"; | |
| 1760 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 77 | by (fast_tac (!claset addSIs [apfst_conv, nat_case_Suc RS trans]) 1); | 
| 923 | 78 | qed "Node_Push_I"; | 
| 79 | ||
| 80 | ||
| 81 | (*** Distinctness of constructors ***) | |
| 82 | ||
| 83 | (** Scons vs Atom **) | |
| 84 | ||
| 85 | goalw Univ.thy [Atom_def,Scons_def,Push_Node_def] "(M$N) ~= Atom(a)"; | |
| 86 | by (rtac notI 1); | |
| 87 | by (etac (equalityD2 RS subsetD RS UnE) 1); | |
| 88 | by (rtac singletonI 1); | |
| 976 
14b55f7fbf15
renamed theorem "apfst" to "apfst_conv" to avoid conflict with function
 clasohm parents: 
972diff
changeset | 89 | by (REPEAT (eresolve_tac [imageE, Abs_Node_inject RS apfst_convE, | 
| 1465 | 90 | Pair_inject, sym RS Push_neq_K0] 1 | 
| 923 | 91 | ORELSE resolve_tac [Node_K0_I, Rep_Node RS Node_Push_I] 1)); | 
| 92 | qed "Scons_not_Atom"; | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 93 | bind_thm ("Atom_not_Scons", Scons_not_Atom RS not_sym);
 | 
| 923 | 94 | |
| 95 | ||
| 96 | (*** Injectiveness ***) | |
| 97 | ||
| 98 | (** Atomic nodes **) | |
| 99 | ||
| 1563 | 100 | goalw Univ.thy [Atom_def, inj_def] "inj(Atom)"; | 
| 1760 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 101 | by (fast_tac (!claset addSIs [Node_K0_I] addSDs [Abs_Node_inject]) 1); | 
| 923 | 102 | qed "inj_Atom"; | 
| 103 | val Atom_inject = inj_Atom RS injD; | |
| 104 | ||
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 105 | goal Univ.thy "(Atom(a)=Atom(b)) = (a=b)"; | 
| 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 106 | by (fast_tac (!claset addSEs [Atom_inject]) 1); | 
| 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 107 | qed "Atom_Atom_eq"; | 
| 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 108 | AddIffs [Atom_Atom_eq]; | 
| 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 109 | |
| 923 | 110 | goalw Univ.thy [Leaf_def,o_def] "inj(Leaf)"; | 
| 111 | by (rtac injI 1); | |
| 112 | by (etac (Atom_inject RS Inl_inject) 1); | |
| 113 | qed "inj_Leaf"; | |
| 114 | ||
| 115 | val Leaf_inject = inj_Leaf RS injD; | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 116 | AddSDs [Leaf_inject]; | 
| 923 | 117 | |
| 118 | goalw Univ.thy [Numb_def,o_def] "inj(Numb)"; | |
| 119 | by (rtac injI 1); | |
| 120 | by (etac (Atom_inject RS Inr_inject) 1); | |
| 121 | qed "inj_Numb"; | |
| 122 | ||
| 123 | val Numb_inject = inj_Numb RS injD; | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 124 | AddSDs [Numb_inject]; | 
| 923 | 125 | |
| 126 | (** Injectiveness of Push_Node **) | |
| 127 | ||
| 128 | val [major,minor] = goalw Univ.thy [Push_Node_def] | |
| 129 | "[| Push_Node i m =Push_Node j n; [| i=j; m=n |] ==> P \ | |
| 130 | \ |] ==> P"; | |
| 976 
14b55f7fbf15
renamed theorem "apfst" to "apfst_conv" to avoid conflict with function
 clasohm parents: 
972diff
changeset | 131 | by (rtac (major RS Abs_Node_inject RS apfst_convE) 1); | 
| 923 | 132 | by (REPEAT (resolve_tac [Rep_Node RS Node_Push_I] 1)); | 
| 976 
14b55f7fbf15
renamed theorem "apfst" to "apfst_conv" to avoid conflict with function
 clasohm parents: 
972diff
changeset | 133 | by (etac (sym RS apfst_convE) 1); | 
| 923 | 134 | by (rtac minor 1); | 
| 135 | by (etac Pair_inject 1); | |
| 136 | by (etac (Push_inject1 RS sym) 1); | |
| 137 | by (rtac (inj_Rep_Node RS injD) 1); | |
| 138 | by (etac trans 1); | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 139 | by (safe_tac (!claset addSEs [Push_inject,sym])); | 
| 923 | 140 | qed "Push_Node_inject"; | 
| 141 | ||
| 142 | ||
| 143 | (** Injectiveness of Scons **) | |
| 144 | ||
| 145 | val [major] = goalw Univ.thy [Scons_def] "M$N <= M'$N' ==> M<=M'"; | |
| 146 | by (cut_facts_tac [major] 1); | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 147 | by (fast_tac (!claset addSEs [Push_Node_inject]) 1); | 
| 923 | 148 | qed "Scons_inject_lemma1"; | 
| 149 | ||
| 150 | val [major] = goalw Univ.thy [Scons_def] "M$N <= M'$N' ==> N<=N'"; | |
| 151 | by (cut_facts_tac [major] 1); | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 152 | by (fast_tac (!claset addSEs [Push_Node_inject]) 1); | 
| 923 | 153 | qed "Scons_inject_lemma2"; | 
| 154 | ||
| 155 | val [major] = goal Univ.thy "M$N = M'$N' ==> M=M'"; | |
| 156 | by (rtac (major RS equalityE) 1); | |
| 157 | by (REPEAT (ares_tac [equalityI, Scons_inject_lemma1] 1)); | |
| 158 | qed "Scons_inject1"; | |
| 159 | ||
| 160 | val [major] = goal Univ.thy "M$N = M'$N' ==> N=N'"; | |
| 161 | by (rtac (major RS equalityE) 1); | |
| 162 | by (REPEAT (ares_tac [equalityI, Scons_inject_lemma2] 1)); | |
| 163 | qed "Scons_inject2"; | |
| 164 | ||
| 165 | val [major,minor] = goal Univ.thy | |
| 166 | "[| M$N = M'$N'; [| M=M'; N=N' |] ==> P \ | |
| 167 | \ |] ==> P"; | |
| 168 | by (rtac ((major RS Scons_inject2) RS ((major RS Scons_inject1) RS minor)) 1); | |
| 169 | qed "Scons_inject"; | |
| 170 | ||
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 171 | AddSDs [Scons_inject]; | 
| 923 | 172 | |
| 173 | goal Univ.thy "(M$N = M'$N') = (M=M' & N=N')"; | |
| 1760 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 174 | by (fast_tac (!claset addSEs [Scons_inject]) 1); | 
| 923 | 175 | qed "Scons_Scons_eq"; | 
| 176 | ||
| 177 | (*** Distinctness involving Leaf and Numb ***) | |
| 178 | ||
| 179 | (** Scons vs Leaf **) | |
| 180 | ||
| 181 | goalw Univ.thy [Leaf_def,o_def] "(M$N) ~= Leaf(a)"; | |
| 182 | by (rtac Scons_not_Atom 1); | |
| 183 | qed "Scons_not_Leaf"; | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 184 | bind_thm ("Leaf_not_Scons", Scons_not_Leaf RS not_sym);
 | 
| 923 | 185 | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 186 | AddIffs [Scons_not_Leaf, Leaf_not_Scons]; | 
| 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 187 | |
| 923 | 188 | |
| 189 | (** Scons vs Numb **) | |
| 190 | ||
| 191 | goalw Univ.thy [Numb_def,o_def] "(M$N) ~= Numb(k)"; | |
| 192 | by (rtac Scons_not_Atom 1); | |
| 193 | qed "Scons_not_Numb"; | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 194 | bind_thm ("Numb_not_Scons", Scons_not_Numb RS not_sym);
 | 
| 923 | 195 | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 196 | AddIffs [Scons_not_Numb, Numb_not_Scons]; | 
| 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 197 | |
| 923 | 198 | |
| 199 | (** Leaf vs Numb **) | |
| 200 | ||
| 201 | goalw Univ.thy [Leaf_def,Numb_def] "Leaf(a) ~= Numb(k)"; | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 202 | by (simp_tac (!simpset addsimps [Inl_not_Inr]) 1); | 
| 923 | 203 | qed "Leaf_not_Numb"; | 
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 204 | bind_thm ("Numb_not_Leaf", Leaf_not_Numb RS not_sym);
 | 
| 923 | 205 | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 206 | AddIffs [Leaf_not_Numb, Numb_not_Leaf]; | 
| 923 | 207 | |
| 208 | ||
| 209 | (*** ndepth -- the depth of a node ***) | |
| 210 | ||
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 211 | Addsimps [apfst_conv]; | 
| 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 212 | AddIffs [Scons_not_Atom, Atom_not_Scons, Scons_Scons_eq]; | 
| 923 | 213 | |
| 214 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 215 | goalw Univ.thy [ndepth_def] "ndepth (Abs_Node((%k.0, x))) = 0"; | 
| 1485 
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
 nipkow parents: 
1465diff
changeset | 216 | by (EVERY1[stac (Node_K0_I RS Abs_Node_inverse), stac split]); | 
| 923 | 217 | by (rtac Least_equality 1); | 
| 218 | by (rtac refl 1); | |
| 219 | by (etac less_zeroE 1); | |
| 220 | qed "ndepth_K0"; | |
| 221 | ||
| 222 | goal Univ.thy "k < Suc(LEAST x. f(x)=0) --> nat_case (Suc i) f k ~= 0"; | |
| 223 | by (nat_ind_tac "k" 1); | |
| 1264 | 224 | by (ALLGOALS Simp_tac); | 
| 923 | 225 | by (rtac impI 1); | 
| 226 | by (etac not_less_Least 1); | |
| 227 | qed "ndepth_Push_lemma"; | |
| 228 | ||
| 229 | goalw Univ.thy [ndepth_def,Push_Node_def] | |
| 230 | "ndepth (Push_Node i n) = Suc(ndepth(n))"; | |
| 231 | by (stac (Rep_Node RS Node_Push_I RS Abs_Node_inverse) 1); | |
| 232 | by (cut_facts_tac [rewrite_rule [Node_def] Rep_Node] 1); | |
| 1786 
8a31d85d27b8
best_tac, deepen_tac and safe_tac now also use default claset.
 berghofe parents: 
1761diff
changeset | 233 | by (safe_tac (!claset)); | 
| 1465 | 234 | by (etac ssubst 1); (*instantiates type variables!*) | 
| 1264 | 235 | by (Simp_tac 1); | 
| 923 | 236 | by (rtac Least_equality 1); | 
| 237 | by (rewtac Push_def); | |
| 238 | by (rtac (nat_case_Suc RS trans) 1); | |
| 239 | by (etac LeastI 1); | |
| 240 | by (etac (ndepth_Push_lemma RS mp) 1); | |
| 241 | qed "ndepth_Push_Node"; | |
| 242 | ||
| 243 | ||
| 244 | (*** ntrunc applied to the various node sets ***) | |
| 245 | ||
| 246 | goalw Univ.thy [ntrunc_def] "ntrunc 0 M = {}";
 | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 247 | by (Fast_tac 1); | 
| 923 | 248 | qed "ntrunc_0"; | 
| 249 | ||
| 250 | goalw Univ.thy [Atom_def,ntrunc_def] "ntrunc (Suc k) (Atom a) = Atom(a)"; | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 251 | by (fast_tac (!claset addss (!simpset addsimps [ndepth_K0])) 1); | 
| 923 | 252 | qed "ntrunc_Atom"; | 
| 253 | ||
| 254 | goalw Univ.thy [Leaf_def,o_def] "ntrunc (Suc k) (Leaf a) = Leaf(a)"; | |
| 255 | by (rtac ntrunc_Atom 1); | |
| 256 | qed "ntrunc_Leaf"; | |
| 257 | ||
| 258 | goalw Univ.thy [Numb_def,o_def] "ntrunc (Suc k) (Numb i) = Numb(i)"; | |
| 259 | by (rtac ntrunc_Atom 1); | |
| 260 | qed "ntrunc_Numb"; | |
| 261 | ||
| 262 | goalw Univ.thy [Scons_def,ntrunc_def] | |
| 263 | "ntrunc (Suc k) (M$N) = ntrunc k M $ ntrunc k N"; | |
| 1786 
8a31d85d27b8
best_tac, deepen_tac and safe_tac now also use default claset.
 berghofe parents: 
1761diff
changeset | 264 | by (safe_tac ((claset_of "Fun") addSIs [equalityI,imageI])); | 
| 923 | 265 | by (REPEAT (stac ndepth_Push_Node 3 THEN etac Suc_mono 3)); | 
| 266 | by (REPEAT (rtac Suc_less_SucD 1 THEN | |
| 1465 | 267 | rtac (ndepth_Push_Node RS subst) 1 THEN | 
| 268 | assume_tac 1)); | |
| 923 | 269 | qed "ntrunc_Scons"; | 
| 270 | ||
| 271 | (** Injection nodes **) | |
| 272 | ||
| 273 | goalw Univ.thy [In0_def] "ntrunc (Suc 0) (In0 M) = {}";
 | |
| 1264 | 274 | by (simp_tac (!simpset addsimps [ntrunc_Scons,ntrunc_0]) 1); | 
| 923 | 275 | by (rewtac Scons_def); | 
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 276 | by (Fast_tac 1); | 
| 923 | 277 | qed "ntrunc_one_In0"; | 
| 278 | ||
| 279 | goalw Univ.thy [In0_def] | |
| 280 | "ntrunc (Suc (Suc k)) (In0 M) = In0 (ntrunc (Suc k) M)"; | |
| 1264 | 281 | by (simp_tac (!simpset addsimps [ntrunc_Scons,ntrunc_Numb]) 1); | 
| 923 | 282 | qed "ntrunc_In0"; | 
| 283 | ||
| 284 | goalw Univ.thy [In1_def] "ntrunc (Suc 0) (In1 M) = {}";
 | |
| 1264 | 285 | by (simp_tac (!simpset addsimps [ntrunc_Scons,ntrunc_0]) 1); | 
| 923 | 286 | by (rewtac Scons_def); | 
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 287 | by (Fast_tac 1); | 
| 923 | 288 | qed "ntrunc_one_In1"; | 
| 289 | ||
| 290 | goalw Univ.thy [In1_def] | |
| 291 | "ntrunc (Suc (Suc k)) (In1 M) = In1 (ntrunc (Suc k) M)"; | |
| 1264 | 292 | by (simp_tac (!simpset addsimps [ntrunc_Scons,ntrunc_Numb]) 1); | 
| 923 | 293 | qed "ntrunc_In1"; | 
| 294 | ||
| 295 | ||
| 296 | (*** Cartesian Product ***) | |
| 297 | ||
| 298 | goalw Univ.thy [uprod_def] "!!M N. [| M:A; N:B |] ==> (M$N) : A<*>B"; | |
| 299 | by (REPEAT (ares_tac [singletonI,UN_I] 1)); | |
| 300 | qed "uprodI"; | |
| 301 | ||
| 302 | (*The general elimination rule*) | |
| 303 | val major::prems = goalw Univ.thy [uprod_def] | |
| 304 | "[| c : A<*>B; \ | |
| 305 | \ !!x y. [| x:A; y:B; c=x$y |] ==> P \ | |
| 306 | \ |] ==> P"; | |
| 307 | by (cut_facts_tac [major] 1); | |
| 308 | by (REPEAT (eresolve_tac [asm_rl,singletonE,UN_E] 1 | |
| 309 | ORELSE resolve_tac prems 1)); | |
| 310 | qed "uprodE"; | |
| 311 | ||
| 312 | (*Elimination of a pair -- introduces no eigenvariables*) | |
| 313 | val prems = goal Univ.thy | |
| 314 | "[| (M$N) : A<*>B; [| M:A; N:B |] ==> P \ | |
| 315 | \ |] ==> P"; | |
| 316 | by (rtac uprodE 1); | |
| 317 | by (REPEAT (ares_tac prems 1 ORELSE eresolve_tac [Scons_inject,ssubst] 1)); | |
| 318 | qed "uprodE2"; | |
| 319 | ||
| 320 | ||
| 321 | (*** Disjoint Sum ***) | |
| 322 | ||
| 323 | goalw Univ.thy [usum_def] "!!M. M:A ==> In0(M) : A<+>B"; | |
| 1760 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 324 | by (Fast_tac 1); | 
| 923 | 325 | qed "usum_In0I"; | 
| 326 | ||
| 327 | goalw Univ.thy [usum_def] "!!N. N:B ==> In1(N) : A<+>B"; | |
| 1760 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 328 | by (Fast_tac 1); | 
| 923 | 329 | qed "usum_In1I"; | 
| 330 | ||
| 331 | val major::prems = goalw Univ.thy [usum_def] | |
| 332 | "[| u : A<+>B; \ | |
| 333 | \ !!x. [| x:A; u=In0(x) |] ==> P; \ | |
| 334 | \ !!y. [| y:B; u=In1(y) |] ==> P \ | |
| 335 | \ |] ==> P"; | |
| 336 | by (rtac (major RS UnE) 1); | |
| 337 | by (REPEAT (rtac refl 1 | |
| 338 | ORELSE eresolve_tac (prems@[imageE,ssubst]) 1)); | |
| 339 | qed "usumE"; | |
| 340 | ||
| 341 | ||
| 342 | (** Injection **) | |
| 343 | ||
| 344 | goalw Univ.thy [In0_def,In1_def] "In0(M) ~= In1(N)"; | |
| 345 | by (rtac notI 1); | |
| 346 | by (etac (Scons_inject1 RS Numb_inject RS Zero_neq_Suc) 1); | |
| 347 | qed "In0_not_In1"; | |
| 348 | ||
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 349 | bind_thm ("In1_not_In0", In0_not_In1 RS not_sym);
 | 
| 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 350 | |
| 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 351 | AddIffs [In0_not_In1, In1_not_In0]; | 
| 923 | 352 | |
| 353 | val [major] = goalw Univ.thy [In0_def] "In0(M) = In0(N) ==> M=N"; | |
| 354 | by (rtac (major RS Scons_inject2) 1); | |
| 355 | qed "In0_inject"; | |
| 356 | ||
| 357 | val [major] = goalw Univ.thy [In1_def] "In1(M) = In1(N) ==> M=N"; | |
| 358 | by (rtac (major RS Scons_inject2) 1); | |
| 359 | qed "In1_inject"; | |
| 360 | ||
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 361 | AddSDs [In0_inject, In1_inject]; | 
| 923 | 362 | |
| 363 | (*** proving equality of sets and functions using ntrunc ***) | |
| 364 | ||
| 365 | goalw Univ.thy [ntrunc_def] "ntrunc k M <= M"; | |
| 1760 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 366 | by (Fast_tac 1); | 
| 923 | 367 | qed "ntrunc_subsetI"; | 
| 368 | ||
| 369 | val [major] = goalw Univ.thy [ntrunc_def] | |
| 370 | "(!!k. ntrunc k M <= N) ==> M<=N"; | |
| 1760 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 371 | by (fast_tac (!claset addIs [less_add_Suc1, less_add_Suc2, | 
| 1465 | 372 | major RS subsetD]) 1); | 
| 923 | 373 | qed "ntrunc_subsetD"; | 
| 374 | ||
| 375 | (*A generalized form of the take-lemma*) | |
| 376 | val [major] = goal Univ.thy "(!!k. ntrunc k M = ntrunc k N) ==> M=N"; | |
| 377 | by (rtac equalityI 1); | |
| 378 | by (ALLGOALS (rtac ntrunc_subsetD)); | |
| 379 | by (ALLGOALS (rtac (ntrunc_subsetI RSN (2, subset_trans)))); | |
| 380 | by (rtac (major RS equalityD1) 1); | |
| 381 | by (rtac (major RS equalityD2) 1); | |
| 382 | qed "ntrunc_equality"; | |
| 383 | ||
| 384 | val [major] = goalw Univ.thy [o_def] | |
| 385 | "[| !!k. (ntrunc(k) o h1) = (ntrunc(k) o h2) |] ==> h1=h2"; | |
| 386 | by (rtac (ntrunc_equality RS ext) 1); | |
| 387 | by (rtac (major RS fun_cong) 1); | |
| 388 | qed "ntrunc_o_equality"; | |
| 389 | ||
| 390 | (*** Monotonicity ***) | |
| 391 | ||
| 392 | goalw Univ.thy [uprod_def] "!!A B. [| A<=A'; B<=B' |] ==> A<*>B <= A'<*>B'"; | |
| 1760 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 393 | by (Fast_tac 1); | 
| 923 | 394 | qed "uprod_mono"; | 
| 395 | ||
| 396 | goalw Univ.thy [usum_def] "!!A B. [| A<=A'; B<=B' |] ==> A<+>B <= A'<+>B'"; | |
| 1760 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 397 | by (Fast_tac 1); | 
| 923 | 398 | qed "usum_mono"; | 
| 399 | ||
| 400 | goalw Univ.thy [Scons_def] "!!M N. [| M<=M'; N<=N' |] ==> M$N <= M'$N'"; | |
| 1760 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 401 | by (Fast_tac 1); | 
| 923 | 402 | qed "Scons_mono"; | 
| 403 | ||
| 404 | goalw Univ.thy [In0_def] "!!M N. M<=N ==> In0(M) <= In0(N)"; | |
| 405 | by (REPEAT (ares_tac [subset_refl,Scons_mono] 1)); | |
| 406 | qed "In0_mono"; | |
| 407 | ||
| 408 | goalw Univ.thy [In1_def] "!!M N. M<=N ==> In1(M) <= In1(N)"; | |
| 409 | by (REPEAT (ares_tac [subset_refl,Scons_mono] 1)); | |
| 410 | qed "In1_mono"; | |
| 411 | ||
| 412 | ||
| 413 | (*** Split and Case ***) | |
| 414 | ||
| 415 | goalw Univ.thy [Split_def] "Split c (M$N) = c M N"; | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 416 | by (fast_tac (!claset addIs [select_equality]) 1); | 
| 923 | 417 | qed "Split"; | 
| 418 | ||
| 419 | goalw Univ.thy [Case_def] "Case c d (In0 M) = c(M)"; | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 420 | by (fast_tac (!claset addIs [select_equality]) 1); | 
| 923 | 421 | qed "Case_In0"; | 
| 422 | ||
| 423 | goalw Univ.thy [Case_def] "Case c d (In1 N) = d(N)"; | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 424 | by (fast_tac (!claset addIs [select_equality]) 1); | 
| 923 | 425 | qed "Case_In1"; | 
| 426 | ||
| 427 | (**** UN x. B(x) rules ****) | |
| 428 | ||
| 429 | goalw Univ.thy [ntrunc_def] "ntrunc k (UN x.f(x)) = (UN x. ntrunc k (f x))"; | |
| 1761 
29e08d527ba1
Removed equalityI from some proofs (because it is now included
 berghofe parents: 
1760diff
changeset | 430 | by (Fast_tac 1); | 
| 923 | 431 | qed "ntrunc_UN1"; | 
| 432 | ||
| 433 | goalw Univ.thy [Scons_def] "(UN x.f(x)) $ M = (UN x. f(x) $ M)"; | |
| 1761 
29e08d527ba1
Removed equalityI from some proofs (because it is now included
 berghofe parents: 
1760diff
changeset | 434 | by (Fast_tac 1); | 
| 923 | 435 | qed "Scons_UN1_x"; | 
| 436 | ||
| 437 | goalw Univ.thy [Scons_def] "M $ (UN x.f(x)) = (UN x. M $ f(x))"; | |
| 1761 
29e08d527ba1
Removed equalityI from some proofs (because it is now included
 berghofe parents: 
1760diff
changeset | 438 | by (Fast_tac 1); | 
| 923 | 439 | qed "Scons_UN1_y"; | 
| 440 | ||
| 441 | goalw Univ.thy [In0_def] "In0(UN x.f(x)) = (UN x. In0(f(x)))"; | |
| 1465 | 442 | by (rtac Scons_UN1_y 1); | 
| 923 | 443 | qed "In0_UN1"; | 
| 444 | ||
| 445 | goalw Univ.thy [In1_def] "In1(UN x.f(x)) = (UN x. In1(f(x)))"; | |
| 1465 | 446 | by (rtac Scons_UN1_y 1); | 
| 923 | 447 | qed "In1_UN1"; | 
| 448 | ||
| 449 | ||
| 450 | (*** Equality : the diagonal relation ***) | |
| 451 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 452 | goalw Univ.thy [diag_def] "!!a A. [| a=b; a:A |] ==> (a,b) : diag(A)"; | 
| 1760 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 453 | by (Fast_tac 1); | 
| 923 | 454 | qed "diag_eqI"; | 
| 455 | ||
| 456 | val diagI = refl RS diag_eqI |> standard; | |
| 457 | ||
| 458 | (*The general elimination rule*) | |
| 459 | val major::prems = goalw Univ.thy [diag_def] | |
| 460 | "[| c : diag(A); \ | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 461 | \ !!x y. [| x:A; c = (x,x) |] ==> P \ | 
| 923 | 462 | \ |] ==> P"; | 
| 463 | by (rtac (major RS UN_E) 1); | |
| 464 | by (REPEAT (eresolve_tac [asm_rl,singletonE] 1 ORELSE resolve_tac prems 1)); | |
| 465 | qed "diagE"; | |
| 466 | ||
| 467 | (*** Equality for Cartesian Product ***) | |
| 468 | ||
| 469 | goalw Univ.thy [dprod_def] | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 470 | "!!r s. [| (M,M'):r; (N,N'):s |] ==> (M$N, M'$N') : r<**>s"; | 
| 1760 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 471 | by (Fast_tac 1); | 
| 923 | 472 | qed "dprodI"; | 
| 473 | ||
| 474 | (*The general elimination rule*) | |
| 475 | val major::prems = goalw Univ.thy [dprod_def] | |
| 476 | "[| c : r<**>s; \ | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 477 | \ !!x y x' y'. [| (x,x') : r; (y,y') : s; c = (x$y,x'$y') |] ==> P \ | 
| 923 | 478 | \ |] ==> P"; | 
| 479 | by (cut_facts_tac [major] 1); | |
| 480 | by (REPEAT_FIRST (eresolve_tac [asm_rl, UN_E, mem_splitE, singletonE])); | |
| 481 | by (REPEAT (ares_tac prems 1 ORELSE hyp_subst_tac 1)); | |
| 482 | qed "dprodE"; | |
| 483 | ||
| 484 | ||
| 485 | (*** Equality for Disjoint Sum ***) | |
| 486 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 487 | goalw Univ.thy [dsum_def] "!!r. (M,M'):r ==> (In0(M), In0(M')) : r<++>s"; | 
| 1760 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 488 | by (Fast_tac 1); | 
| 923 | 489 | qed "dsum_In0I"; | 
| 490 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 491 | goalw Univ.thy [dsum_def] "!!r. (N,N'):s ==> (In1(N), In1(N')) : r<++>s"; | 
| 1760 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 492 | by (Fast_tac 1); | 
| 923 | 493 | qed "dsum_In1I"; | 
| 494 | ||
| 495 | val major::prems = goalw Univ.thy [dsum_def] | |
| 496 | "[| w : r<++>s; \ | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 497 | \ !!x x'. [| (x,x') : r; w = (In0(x), In0(x')) |] ==> P; \ | 
| 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 498 | \ !!y y'. [| (y,y') : s; w = (In1(y), In1(y')) |] ==> P \ | 
| 923 | 499 | \ |] ==> P"; | 
| 500 | by (cut_facts_tac [major] 1); | |
| 501 | by (REPEAT_FIRST (eresolve_tac [asm_rl, UN_E, UnE, mem_splitE, singletonE])); | |
| 502 | by (DEPTH_SOLVE (ares_tac prems 1 ORELSE hyp_subst_tac 1)); | |
| 503 | qed "dsumE"; | |
| 504 | ||
| 505 | ||
| 1760 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 506 | AddSIs [diagI, uprodI, dprodI]; | 
| 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 507 | AddIs [usum_In0I, usum_In1I, dsum_In0I, dsum_In1I]; | 
| 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 508 | AddSEs [diagE, uprodE, dprodE, usumE, dsumE]; | 
| 923 | 509 | |
| 510 | (*** Monotonicity ***) | |
| 511 | ||
| 512 | goal Univ.thy "!!r s. [| r<=r'; s<=s' |] ==> r<**>s <= r'<**>s'"; | |
| 1760 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 513 | by (Fast_tac 1); | 
| 923 | 514 | qed "dprod_mono"; | 
| 515 | ||
| 516 | goal Univ.thy "!!r s. [| r<=r'; s<=s' |] ==> r<++>s <= r'<++>s'"; | |
| 1760 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 517 | by (Fast_tac 1); | 
| 923 | 518 | qed "dsum_mono"; | 
| 519 | ||
| 520 | ||
| 521 | (*** Bounding theorems ***) | |
| 522 | ||
| 1642 | 523 | goal Univ.thy "diag(A) <= A Times A"; | 
| 1760 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 524 | by (Fast_tac 1); | 
| 923 | 525 | qed "diag_subset_Sigma"; | 
| 526 | ||
| 1642 | 527 | goal Univ.thy "((A Times B) <**> (C Times D)) <= (A<*>C) Times (B<*>D)"; | 
| 1760 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 528 | by (Fast_tac 1); | 
| 923 | 529 | qed "dprod_Sigma"; | 
| 530 | ||
| 531 | val dprod_subset_Sigma = [dprod_mono, dprod_Sigma] MRS subset_trans |>standard; | |
| 532 | ||
| 533 | (*Dependent version*) | |
| 534 | goal Univ.thy | |
| 535 | "(Sigma A B <**> Sigma C D) <= Sigma (A<*>C) (Split(%x y. B(x)<*>D(y)))"; | |
| 1786 
8a31d85d27b8
best_tac, deepen_tac and safe_tac now also use default claset.
 berghofe parents: 
1761diff
changeset | 536 | by (safe_tac (!claset)); | 
| 923 | 537 | by (stac Split 1); | 
| 1760 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 538 | by (Fast_tac 1); | 
| 923 | 539 | qed "dprod_subset_Sigma2"; | 
| 540 | ||
| 1642 | 541 | goal Univ.thy "(A Times B <++> C Times D) <= (A<+>C) Times (B<+>D)"; | 
| 1760 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 542 | by (Fast_tac 1); | 
| 923 | 543 | qed "dsum_Sigma"; | 
| 544 | ||
| 545 | val dsum_subset_Sigma = [dsum_mono, dsum_Sigma] MRS subset_trans |> standard; | |
| 546 | ||
| 547 | ||
| 548 | (*** Domain ***) | |
| 549 | ||
| 550 | goal Univ.thy "fst `` diag(A) = A"; | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 551 | by (Fast_tac 1); | 
| 923 | 552 | qed "fst_image_diag"; | 
| 553 | ||
| 554 | goal Univ.thy "fst `` (r<**>s) = (fst``r) <*> (fst``s)"; | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 555 | by (Fast_tac 1); | 
| 923 | 556 | qed "fst_image_dprod"; | 
| 557 | ||
| 558 | goal Univ.thy "fst `` (r<++>s) = (fst``r) <+> (fst``s)"; | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1786diff
changeset | 559 | by (Fast_tac 1); | 
| 923 | 560 | qed "fst_image_dsum"; | 
| 561 | ||
| 1264 | 562 | Addsimps [fst_image_diag, fst_image_dprod, fst_image_dsum]; |