src/CCL/wfd.ML
changeset 17456 bcf7544875b2
equal deleted inserted replaced
17455:151e76f0e3c7 17456:bcf7544875b2
       
     1 (*  Title:      CCL/Wfd.ML
       
     2     ID:         $Id$
       
     3 *)
       
     4 
       
     5 (***********)
       
     6 
       
     7 val [major,prem] = goalw (the_context ()) [Wfd_def]
       
     8     "[| Wfd(R);       \
       
     9 \       !!x.[| ALL y. <y,x>: R --> P(y) |] ==> P(x) |]  ==>  \
       
    10 \    P(a)";
       
    11 by (rtac (major RS spec RS mp RS spec RS CollectD) 1);
       
    12 by (fast_tac (set_cs addSIs [prem RS CollectI]) 1);
       
    13 qed "wfd_induct";
       
    14 
       
    15 val [p1,p2,p3] = goal (the_context ())
       
    16     "[| !!x y.<x,y> : R ==> Q(x); \
       
    17 \       ALL x. (ALL y. <y,x> : R --> y : P) --> x : P; \
       
    18 \       !!x. Q(x) ==> x:P |] ==> a:P";
       
    19 by (rtac (p2 RS  spec  RS mp) 1);
       
    20 by (fast_tac (set_cs addSIs [p1 RS p3]) 1);
       
    21 qed "wfd_strengthen_lemma";
       
    22 
       
    23 fun wfd_strengthen_tac s i = res_inst_tac [("Q",s)] wfd_strengthen_lemma i THEN
       
    24                              assume_tac (i+1);
       
    25 
       
    26 val wfd::prems = goal (the_context ()) "[| Wfd(r);  <a,x>:r;  <x,a>:r |] ==> P";
       
    27 by (subgoal_tac "ALL x. <a,x>:r --> <x,a>:r --> P" 1);
       
    28 by (fast_tac (FOL_cs addIs prems) 1);
       
    29 by (rtac (wfd RS  wfd_induct) 1);
       
    30 by (ALLGOALS (fast_tac (ccl_cs addSIs prems)));
       
    31 qed "wf_anti_sym";
       
    32 
       
    33 val prems = goal (the_context ()) "[| Wfd(r);  <a,a>: r |] ==> P";
       
    34 by (rtac wf_anti_sym 1);
       
    35 by (REPEAT (resolve_tac prems 1));
       
    36 qed "wf_anti_refl";
       
    37 
       
    38 (*** Irreflexive transitive closure ***)
       
    39 
       
    40 val [prem] = goal (the_context ()) "Wfd(R) ==> Wfd(R^+)";
       
    41 by (rewtac Wfd_def);
       
    42 by (REPEAT (ares_tac [allI,ballI,impI] 1));
       
    43 (*must retain the universal formula for later use!*)
       
    44 by (rtac allE 1 THEN assume_tac 1);
       
    45 by (etac mp 1);
       
    46 by (rtac (prem RS wfd_induct) 1);
       
    47 by (rtac (impI RS allI) 1);
       
    48 by (etac tranclE 1);
       
    49 by (fast_tac ccl_cs 1);
       
    50 by (etac (spec RS mp RS spec RS mp) 1);
       
    51 by (REPEAT (atac 1));
       
    52 qed "trancl_wf";
       
    53 
       
    54 (*** Lexicographic Ordering ***)
       
    55 
       
    56 Goalw [lex_def]
       
    57  "p : ra**rb <-> (EX a a' b b'. p = <<a,b>,<a',b'>> & (<a,a'> : ra | a=a' & <b,b'> : rb))";
       
    58 by (fast_tac ccl_cs 1);
       
    59 qed "lexXH";
       
    60 
       
    61 val prems = goal (the_context ())
       
    62  "<a,a'> : ra ==> <<a,b>,<a',b'>> : ra**rb";
       
    63 by (fast_tac (ccl_cs addSIs (prems @ [lexXH RS iffD2])) 1);
       
    64 qed "lexI1";
       
    65 
       
    66 val prems = goal (the_context ())
       
    67  "<b,b'> : rb ==> <<a,b>,<a,b'>> : ra**rb";
       
    68 by (fast_tac (ccl_cs addSIs (prems @ [lexXH RS iffD2])) 1);
       
    69 qed "lexI2";
       
    70 
       
    71 val major::prems = goal (the_context ())
       
    72  "[| p : ra**rb;  \
       
    73 \    !!a a' b b'.[| <a,a'> : ra; p=<<a,b>,<a',b'>> |] ==> R;  \
       
    74 \    !!a b b'.[| <b,b'> : rb;  p = <<a,b>,<a,b'>> |] ==> R  |] ==> \
       
    75 \ R";
       
    76 by (rtac (major RS (lexXH RS iffD1) RS exE) 1);
       
    77 by (REPEAT_SOME (eresolve_tac ([exE,conjE,disjE]@prems)));
       
    78 by (ALLGOALS (fast_tac ccl_cs));
       
    79 qed "lexE";
       
    80 
       
    81 val [major,minor] = goal (the_context ())
       
    82  "[| p : r**s;  !!a a' b b'. p = <<a,b>,<a',b'>> ==> P |] ==>P";
       
    83 by (rtac (major RS lexE) 1);
       
    84 by (ALLGOALS (fast_tac (set_cs addSEs [minor])));
       
    85 qed "lex_pair";
       
    86 
       
    87 val [wfa,wfb] = goal (the_context ())
       
    88  "[| Wfd(R); Wfd(S) |] ==> Wfd(R**S)";
       
    89 by (rewtac Wfd_def);
       
    90 by (safe_tac ccl_cs);
       
    91 by (wfd_strengthen_tac "%x. EX a b. x=<a,b>" 1);
       
    92 by (fast_tac (term_cs addSEs [lex_pair]) 1);
       
    93 by (subgoal_tac "ALL a b.<a,b>:P" 1);
       
    94 by (fast_tac ccl_cs 1);
       
    95 by (rtac (wfa RS wfd_induct RS allI) 1);
       
    96 by (rtac (wfb RS wfd_induct RS allI) 1);back();
       
    97 by (fast_tac (type_cs addSEs [lexE]) 1);
       
    98 qed "lex_wf";
       
    99 
       
   100 (*** Mapping ***)
       
   101 
       
   102 Goalw [wmap_def]
       
   103  "p : wmap(f,r) <-> (EX x y. p=<x,y>  &  <f(x),f(y)> : r)";
       
   104 by (fast_tac ccl_cs 1);
       
   105 qed "wmapXH";
       
   106 
       
   107 val prems = goal (the_context ())
       
   108  "<f(a),f(b)> : r ==> <a,b> : wmap(f,r)";
       
   109 by (fast_tac (ccl_cs addSIs (prems @ [wmapXH RS iffD2])) 1);
       
   110 qed "wmapI";
       
   111 
       
   112 val major::prems = goal (the_context ())
       
   113  "[| p : wmap(f,r);  !!a b.[| <f(a),f(b)> : r;  p=<a,b> |] ==> R |] ==> R";
       
   114 by (rtac (major RS (wmapXH RS iffD1) RS exE) 1);
       
   115 by (REPEAT_SOME (eresolve_tac ([exE,conjE,disjE]@prems)));
       
   116 by (ALLGOALS (fast_tac ccl_cs));
       
   117 qed "wmapE";
       
   118 
       
   119 val [wf] = goal (the_context ())
       
   120  "Wfd(r) ==> Wfd(wmap(f,r))";
       
   121 by (rewtac Wfd_def);
       
   122 by (safe_tac ccl_cs);
       
   123 by (subgoal_tac "ALL b. ALL a. f(a)=b-->a:P" 1);
       
   124 by (fast_tac ccl_cs 1);
       
   125 by (rtac (wf RS wfd_induct RS allI) 1);
       
   126 by (safe_tac ccl_cs);
       
   127 by (etac (spec RS mp) 1);
       
   128 by (safe_tac (ccl_cs addSEs [wmapE]));
       
   129 by (etac (spec RS mp RS spec RS mp) 1);
       
   130 by (assume_tac 1);
       
   131 by (rtac refl 1);
       
   132 qed "wmap_wf";
       
   133 
       
   134 (* Projections *)
       
   135 
       
   136 val prems = goal (the_context ()) "<xa,ya> : r ==> <<xa,xb>,<ya,yb>> : wmap(fst,r)";
       
   137 by (rtac wmapI 1);
       
   138 by (simp_tac (term_ss addsimps prems) 1);
       
   139 qed "wfstI";
       
   140 
       
   141 val prems = goal (the_context ()) "<xb,yb> : r ==> <<xa,xb>,<ya,yb>> : wmap(snd,r)";
       
   142 by (rtac wmapI 1);
       
   143 by (simp_tac (term_ss addsimps prems) 1);
       
   144 qed "wsndI";
       
   145 
       
   146 val prems = goal (the_context ()) "<xc,yc> : r ==> <<xa,<xb,xc>>,<ya,<yb,yc>>> : wmap(thd,r)";
       
   147 by (rtac wmapI 1);
       
   148 by (simp_tac (term_ss addsimps prems) 1);
       
   149 qed "wthdI";
       
   150 
       
   151 (*** Ground well-founded relations ***)
       
   152 
       
   153 val prems = goalw (the_context ()) [wf_def]
       
   154     "[| Wfd(r);  a : r |] ==> a : wf(r)";
       
   155 by (fast_tac (set_cs addSIs prems) 1);
       
   156 qed "wfI";
       
   157 
       
   158 val prems = goalw (the_context ()) [Wfd_def] "Wfd({})";
       
   159 by (fast_tac (set_cs addEs [EmptyXH RS iffD1 RS FalseE]) 1);
       
   160 qed "Empty_wf";
       
   161 
       
   162 val prems = goalw (the_context ()) [wf_def] "Wfd(wf(R))";
       
   163 by (res_inst_tac [("Q","Wfd(R)")] (excluded_middle RS disjE) 1);
       
   164 by (ALLGOALS (asm_simp_tac CCL_ss));
       
   165 by (rtac Empty_wf 1);
       
   166 qed "wf_wf";
       
   167 
       
   168 Goalw [NatPR_def]  "p : NatPR <-> (EX x:Nat. p=<x,succ(x)>)";
       
   169 by (fast_tac set_cs 1);
       
   170 qed "NatPRXH";
       
   171 
       
   172 Goalw [ListPR_def]  "p : ListPR(A) <-> (EX h:A. EX t:List(A).p=<t,h$t>)";
       
   173 by (fast_tac set_cs 1);
       
   174 qed "ListPRXH";
       
   175 
       
   176 val NatPRI = refl RS (bexI RS (NatPRXH RS iffD2));
       
   177 val ListPRI = refl RS (bexI RS (bexI RS (ListPRXH RS iffD2)));
       
   178 
       
   179 Goalw [Wfd_def]  "Wfd(NatPR)";
       
   180 by (safe_tac set_cs);
       
   181 by (wfd_strengthen_tac "%x. x:Nat" 1);
       
   182 by (fast_tac (type_cs addSEs [XH_to_E NatPRXH]) 1);
       
   183 by (etac Nat_ind 1);
       
   184 by (ALLGOALS (fast_tac (type_cs addEs [XH_to_E NatPRXH])));
       
   185 qed "NatPR_wf";
       
   186 
       
   187 Goalw [Wfd_def]  "Wfd(ListPR(A))";
       
   188 by (safe_tac set_cs);
       
   189 by (wfd_strengthen_tac "%x. x:List(A)" 1);
       
   190 by (fast_tac (type_cs addSEs [XH_to_E ListPRXH]) 1);
       
   191 by (etac List_ind 1);
       
   192 by (ALLGOALS (fast_tac (type_cs addEs [XH_to_E ListPRXH])));
       
   193 qed "ListPR_wf";