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