src/HOL/WF_Rel.ML
changeset 5069 3ea049f7979d
parent 4751 6fbd9838ccae
child 5143 b94cd208f073
equal deleted inserted replaced
5068:fb28eaa07e01 5069:3ea049f7979d
    11 
    11 
    12 (*----------------------------------------------------------------------------
    12 (*----------------------------------------------------------------------------
    13  * "Less than" on the natural numbers
    13  * "Less than" on the natural numbers
    14  *---------------------------------------------------------------------------*)
    14  *---------------------------------------------------------------------------*)
    15 
    15 
    16 goalw thy [less_than_def] "wf less_than"; 
    16 Goalw [less_than_def] "wf less_than"; 
    17 by (rtac (wf_pred_nat RS wf_trancl) 1);
    17 by (rtac (wf_pred_nat RS wf_trancl) 1);
    18 qed "wf_less_than";
    18 qed "wf_less_than";
    19 AddIffs [wf_less_than];
    19 AddIffs [wf_less_than];
    20 
    20 
    21 goalw thy [less_than_def] "trans less_than"; 
    21 Goalw [less_than_def] "trans less_than"; 
    22 by (rtac trans_trancl 1);
    22 by (rtac trans_trancl 1);
    23 qed "trans_less_than";
    23 qed "trans_less_than";
    24 AddIffs [trans_less_than];
    24 AddIffs [trans_less_than];
    25 
    25 
    26 goalw thy [less_than_def, less_def] "((x,y): less_than) = (x<y)"; 
    26 Goalw [less_than_def, less_def] "((x,y): less_than) = (x<y)"; 
    27 by (Simp_tac 1);
    27 by (Simp_tac 1);
    28 qed "less_than_iff";
    28 qed "less_than_iff";
    29 AddIffs [less_than_iff];
    29 AddIffs [less_than_iff];
    30 
    30 
    31 (*----------------------------------------------------------------------------
    31 (*----------------------------------------------------------------------------
    32  * The inverse image into a wellfounded relation is wellfounded.
    32  * The inverse image into a wellfounded relation is wellfounded.
    33  *---------------------------------------------------------------------------*)
    33  *---------------------------------------------------------------------------*)
    34 
    34 
    35 goal thy "!!r. wf(r) ==> wf(inv_image r (f::'a=>'b))"; 
    35 Goal "!!r. wf(r) ==> wf(inv_image r (f::'a=>'b))"; 
    36 by (full_simp_tac (simpset() addsimps [inv_image_def, wf_eq_minimal]) 1);
    36 by (full_simp_tac (simpset() addsimps [inv_image_def, wf_eq_minimal]) 1);
    37 by (Clarify_tac 1);
    37 by (Clarify_tac 1);
    38 by (subgoal_tac "? (w::'b). w : {w. ? (x::'a). x: Q & (f x = w)}" 1);
    38 by (subgoal_tac "? (w::'b). w : {w. ? (x::'a). x: Q & (f x = w)}" 1);
    39 by (blast_tac (claset() delrules [allE]) 2);
    39 by (blast_tac (claset() delrules [allE]) 2);
    40 by (etac allE 1);
    40 by (etac allE 1);
    41 by (mp_tac 1);
    41 by (mp_tac 1);
    42 by (Blast_tac 1);
    42 by (Blast_tac 1);
    43 qed "wf_inv_image";
    43 qed "wf_inv_image";
    44 AddSIs [wf_inv_image];
    44 AddSIs [wf_inv_image];
    45 
    45 
    46 goalw thy [trans_def,inv_image_def]
    46 Goalw [trans_def,inv_image_def]
    47     "!!r. trans r ==> trans (inv_image r f)";
    47     "!!r. trans r ==> trans (inv_image r f)";
    48 by (Simp_tac 1);
    48 by (Simp_tac 1);
    49 by (Blast_tac 1);
    49 by (Blast_tac 1);
    50 qed "trans_inv_image";
    50 qed "trans_inv_image";
    51 
    51 
    52 
    52 
    53 (*----------------------------------------------------------------------------
    53 (*----------------------------------------------------------------------------
    54  * All measures are wellfounded.
    54  * All measures are wellfounded.
    55  *---------------------------------------------------------------------------*)
    55  *---------------------------------------------------------------------------*)
    56 
    56 
    57 goalw thy [measure_def] "wf (measure f)";
    57 Goalw [measure_def] "wf (measure f)";
    58 by (rtac (wf_less_than RS wf_inv_image) 1);
    58 by (rtac (wf_less_than RS wf_inv_image) 1);
    59 qed "wf_measure";
    59 qed "wf_measure";
    60 AddIffs [wf_measure];
    60 AddIffs [wf_measure];
    61 
    61 
    62 val measure_induct = standard
    62 val measure_induct = standard
    80 AddSIs [wf_lex_prod];
    80 AddSIs [wf_lex_prod];
    81 
    81 
    82 (*---------------------------------------------------------------------------
    82 (*---------------------------------------------------------------------------
    83  * Transitivity of WF combinators.
    83  * Transitivity of WF combinators.
    84  *---------------------------------------------------------------------------*)
    84  *---------------------------------------------------------------------------*)
    85 goalw thy [trans_def, lex_prod_def]
    85 Goalw [trans_def, lex_prod_def]
    86     "!!R1 R2. [| trans R1; trans R2 |] ==> trans (R1 ** R2)";
    86     "!!R1 R2. [| trans R1; trans R2 |] ==> trans (R1 ** R2)";
    87 by (Simp_tac 1);
    87 by (Simp_tac 1);
    88 by (Blast_tac 1);
    88 by (Blast_tac 1);
    89 qed "trans_lex_prod";
    89 qed "trans_lex_prod";
    90 AddSIs [trans_lex_prod];
    90 AddSIs [trans_lex_prod];
    91 
    91 
    92 
    92 
    93 (*---------------------------------------------------------------------------
    93 (*---------------------------------------------------------------------------
    94  * Wellfoundedness of proper subset on finite sets.
    94  * Wellfoundedness of proper subset on finite sets.
    95  *---------------------------------------------------------------------------*)
    95  *---------------------------------------------------------------------------*)
    96 goalw thy [finite_psubset_def] "wf(finite_psubset)";
    96 Goalw [finite_psubset_def] "wf(finite_psubset)";
    97 by (rtac (wf_measure RS wf_subset) 1);
    97 by (rtac (wf_measure RS wf_subset) 1);
    98 by (simp_tac (simpset() addsimps [measure_def, inv_image_def, less_than_def,
    98 by (simp_tac (simpset() addsimps [measure_def, inv_image_def, less_than_def,
    99 				 symmetric less_def])1);
    99 				 symmetric less_def])1);
   100 by (fast_tac (claset() addSIs [psubset_card]) 1);
   100 by (fast_tac (claset() addSIs [psubset_card]) 1);
   101 qed "wf_finite_psubset";
   101 qed "wf_finite_psubset";
   102 
   102 
   103 goalw thy [finite_psubset_def, trans_def] "trans finite_psubset";
   103 Goalw [finite_psubset_def, trans_def] "trans finite_psubset";
   104 by (simp_tac (simpset() addsimps [psubset_def]) 1);
   104 by (simp_tac (simpset() addsimps [psubset_def]) 1);
   105 by (Blast_tac 1);
   105 by (Blast_tac 1);
   106 qed "trans_finite_psubset";
   106 qed "trans_finite_psubset";
   107 
   107 
   108 (*---------------------------------------------------------------------------
   108 (*---------------------------------------------------------------------------
   109  * Wellfoundedness of finite acyclic relations
   109  * Wellfoundedness of finite acyclic relations
   110  * Cannot go into WF because it needs Finite
   110  * Cannot go into WF because it needs Finite
   111  *---------------------------------------------------------------------------*)
   111  *---------------------------------------------------------------------------*)
   112 
   112 
   113 goal thy "!!r. finite r ==> acyclic r --> wf r";
   113 Goal "!!r. finite r ==> acyclic r --> wf r";
   114 by (etac finite_induct 1);
   114 by (etac finite_induct 1);
   115  by (Blast_tac 1);
   115  by (Blast_tac 1);
   116 by (split_all_tac 1);
   116 by (split_all_tac 1);
   117 by (Asm_full_simp_tac 1);
   117 by (Asm_full_simp_tac 1);
   118 qed_spec_mp "finite_acyclic_wf";
   118 qed_spec_mp "finite_acyclic_wf";
   120 qed_goal "finite_acyclic_wf_converse" thy 
   120 qed_goal "finite_acyclic_wf_converse" thy 
   121  "!!X. [|finite r; acyclic r|] ==> wf (r^-1)" (K [
   121  "!!X. [|finite r; acyclic r|] ==> wf (r^-1)" (K [
   122 	etac (finite_converse RS iffD2 RS finite_acyclic_wf) 1,
   122 	etac (finite_converse RS iffD2 RS finite_acyclic_wf) 1,
   123 	etac (acyclic_converse RS iffD2) 1]);
   123 	etac (acyclic_converse RS iffD2) 1]);
   124 
   124 
   125 goal thy "!!r. finite r ==> wf r = acyclic r";
   125 Goal "!!r. finite r ==> wf r = acyclic r";
   126 by (blast_tac (claset() addIs [finite_acyclic_wf,wf_acyclic]) 1);
   126 by (blast_tac (claset() addIs [finite_acyclic_wf,wf_acyclic]) 1);
   127 qed "wf_iff_acyclic_if_finite";
   127 qed "wf_iff_acyclic_if_finite";
   128 
   128 
   129 
   129 
   130 (*---------------------------------------------------------------------------
   130 (*---------------------------------------------------------------------------
   131  * A relation is wellfounded iff it has no infinite descending chain
   131  * A relation is wellfounded iff it has no infinite descending chain
   132  *---------------------------------------------------------------------------*)
   132  *---------------------------------------------------------------------------*)
   133 
   133 
   134 goalw thy [wf_eq_minimal RS eq_reflection]
   134 Goalw [wf_eq_minimal RS eq_reflection]
   135   "wf r = (~(? f. !i. (f(Suc i),f i) : r))";
   135   "wf r = (~(? f. !i. (f(Suc i),f i) : r))";
   136 by (rtac iffI 1);
   136 by (rtac iffI 1);
   137  by (rtac notI 1);
   137  by (rtac notI 1);
   138  by (etac exE 1);
   138  by (etac exE 1);
   139  by (eres_inst_tac [("x","{w. ? i. w=f i}")] allE 1);
   139  by (eres_inst_tac [("x","{w. ? i. w=f i}")] allE 1);