| 2469 |      1 | (*  Title:      ZF/AC/OrdQuant.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Authors:    Krzysztof Grabczewski and L C Paulson
 | 
|  |      4 | 
 | 
|  |      5 | Quantifiers and union operator for ordinals. 
 | 
|  |      6 | *)
 | 
|  |      7 | 
 | 
|  |      8 | (*** universal quantifier for ordinals ***)
 | 
|  |      9 | 
 | 
|  |     10 | qed_goalw "oallI" thy [oall_def]
 | 
|  |     11 |     "[| !!x. x<A ==> P(x) |] ==> ALL x<A. P(x)"
 | 
|  |     12 |  (fn prems=> [ (REPEAT (ares_tac (prems @ [allI,impI]) 1)) ]);
 | 
|  |     13 | 
 | 
|  |     14 | qed_goalw "ospec" thy [oall_def]
 | 
|  |     15 |     "[| ALL x<A. P(x);  x<A |] ==> P(x)"
 | 
|  |     16 |  (fn major::prems=>
 | 
|  |     17 |   [ (rtac (major RS spec RS mp) 1),
 | 
|  |     18 |     (resolve_tac prems 1) ]);
 | 
|  |     19 | 
 | 
|  |     20 | qed_goalw "oallE" thy [oall_def]
 | 
|  |     21 |     "[| ALL x<A. P(x);  P(x) ==> Q;  ~x<A ==> Q |] ==> Q"
 | 
|  |     22 |  (fn major::prems=>
 | 
|  |     23 |   [ (rtac (major RS allE) 1),
 | 
|  |     24 |     (REPEAT (eresolve_tac (prems@[asm_rl,impCE]) 1)) ]);
 | 
|  |     25 | 
 | 
|  |     26 | qed_goal "rev_oallE" thy
 | 
|  |     27 |     "[| ALL x<A. P(x);  ~x<A ==> Q;  P(x) ==> Q |] ==> Q"
 | 
|  |     28 |  (fn major::prems=>
 | 
|  |     29 |   [ (rtac (major RS oallE) 1),
 | 
|  |     30 |     (REPEAT (eresolve_tac prems 1)) ]);
 | 
|  |     31 | 
 | 
|  |     32 | (*Trival rewrite rule;   (ALL x<a.P)<->P holds only if a is not 0!*)
 | 
|  |     33 | qed_goal "oall_simp" thy "(ALL x<a. True) <-> True"
 | 
|  |     34 |  (fn _=> [ (REPEAT (ares_tac [TrueI,oallI,iffI] 1)) ]);
 | 
|  |     35 | 
 | 
|  |     36 | (*Congruence rule for rewriting*)
 | 
|  |     37 | qed_goalw "oall_cong" thy [oall_def]
 | 
|  |     38 |     "[| a=a';  !!x. x<a' ==> P(x) <-> P'(x) |] ==> oall(a,P) <-> oall(a',P')"
 | 
| 4091 |     39 |  (fn prems=> [ (simp_tac (simpset() addsimps prems) 1) ]);
 | 
| 2469 |     40 | 
 | 
|  |     41 | 
 | 
|  |     42 | (*** existential quantifier for ordinals ***)
 | 
|  |     43 | 
 | 
|  |     44 | qed_goalw "oexI" thy [oex_def]
 | 
|  |     45 |     "[| P(x);  x<A |] ==> EX x<A. P(x)"
 | 
|  |     46 |  (fn prems=> [ (REPEAT (ares_tac (prems @ [exI,conjI]) 1)) ]);
 | 
|  |     47 | 
 | 
|  |     48 | (*Not of the general form for such rules; ~EX has become ALL~ *)
 | 
|  |     49 | qed_goal "oexCI" thy 
 | 
| 3840 |     50 |    "[| ALL x<A. ~P(x) ==> P(a);  a<A |] ==> EX x<A. P(x)"
 | 
| 2469 |     51 |  (fn prems=>
 | 
|  |     52 |   [ (rtac classical 1),
 | 
|  |     53 |     (REPEAT (ares_tac (prems@[oexI,oallI,notI,notE]) 1)) ]);
 | 
|  |     54 | 
 | 
|  |     55 | qed_goalw "oexE" thy [oex_def]
 | 
|  |     56 |     "[| EX x<A. P(x);  !!x. [| x<A; P(x) |] ==> Q \
 | 
|  |     57 | \    |] ==> Q"
 | 
|  |     58 |  (fn major::prems=>
 | 
|  |     59 |   [ (rtac (major RS exE) 1),
 | 
|  |     60 |     (REPEAT (eresolve_tac (prems @ [asm_rl,conjE]) 1)) ]);
 | 
|  |     61 | 
 | 
|  |     62 | qed_goalw "oex_cong" thy [oex_def]
 | 
|  |     63 |     "[| a=a';  !!x. x<a' ==> P(x) <-> P'(x) \
 | 
|  |     64 | \    |] ==> oex(a,P) <-> oex(a',P')"
 | 
| 4091 |     65 |  (fn prems=> [ (simp_tac (simpset() addsimps prems addcongs [conj_cong]) 1) ]);
 | 
| 2469 |     66 | 
 | 
|  |     67 | 
 | 
|  |     68 | (*** Rules for Ordinal-Indexed Unions ***)
 | 
|  |     69 | 
 | 
|  |     70 | qed_goalw "OUN_I" thy [OUnion_def]
 | 
|  |     71 |         "!!i. [| a<i;  b: B(a) |] ==> b: (UN z<i. B(z))"
 | 
| 4091 |     72 |  (fn _=> [ fast_tac (claset() addSEs [ltE]) 1 ]);
 | 
| 2469 |     73 | 
 | 
|  |     74 | qed_goalw "OUN_E" thy [OUnion_def]
 | 
|  |     75 |     "[| b : (UN z<i. B(z));  !!a.[| b: B(a);  a<i |] ==> R |] ==> R"
 | 
|  |     76 |  (fn major::prems=>
 | 
|  |     77 |   [ (rtac (major RS CollectE) 1),
 | 
|  |     78 |     (rtac UN_E 1),
 | 
|  |     79 |     (REPEAT (ares_tac (ltI::prems) 1)) ]);
 | 
|  |     80 | 
 | 
|  |     81 | qed_goalw "OUN_iff" thy [oex_def]
 | 
|  |     82 |     "b : (UN x<i. B(x)) <-> (EX x<i. b : B(x))"
 | 
| 4091 |     83 |  (fn _=> [ (fast_tac (claset() addIs [OUN_I] addSEs [OUN_E]) 1) ]);
 | 
| 2469 |     84 | 
 | 
|  |     85 | qed_goal "OUN_cong" thy
 | 
| 3840 |     86 |     "[| i=j;  !!x. x<j ==> C(x)=D(x) |] ==> (UN x<i. C(x)) = (UN x<j. D(x))"
 | 
| 2469 |     87 |  (fn prems=>
 | 
|  |     88 |       [ rtac equality_iffI 1,
 | 
| 5529 |     89 |         simp_tac (simpset() addcongs [oex_cong] addsimps OUN_iff::prems) 1 ]);
 | 
| 2469 |     90 | 
 | 
|  |     91 | AddSIs [oallI];
 | 
|  |     92 | AddIs  [oexI, OUN_I];
 | 
|  |     93 | AddSEs [oexE, OUN_E];
 | 
|  |     94 | AddEs  [rev_oallE];
 | 
|  |     95 | 
 | 
| 6093 |     96 | val Ord_atomize = atomize (("OrdQuant.oall", [ospec])::ZF_conn_pairs, 
 | 
| 2469 |     97 |                            ZF_mem_pairs);
 | 
|  |     98 | 
 | 
| 5553 |     99 | simpset_ref() := simpset() setmksimps (map mk_eq o Ord_atomize o gen_all)
 | 
| 2469 |    100 |                         addsimps [oall_simp, ltD RS beta]
 | 
|  |    101 |                         addcongs [oall_cong, oex_cong, OUN_cong];
 | 
|  |    102 | 
 | 
|  |    103 | val major::prems = goalw thy [lt_def, oall_def]
 | 
|  |    104 |     "[| i<k;  !!x.[| x<k;  ALL y<x. P(y) |] ==> P(x) \
 | 
|  |    105 | \    |]  ==>  P(i)";
 | 
|  |    106 | by (rtac (major RS conjE) 1);
 | 
|  |    107 | by (etac Ord_induct 1 THEN assume_tac 1);
 | 
| 4091 |    108 | by (fast_tac (claset() addIs prems) 1);
 | 
| 2469 |    109 | qed "lt_induct";
 | 
|  |    110 | 
 |