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