| 
15346
 | 
     1  | 
(*  ID:   $Id$
  | 
| 
10213
 | 
     2  | 
    Author:     Konrad Slind
  | 
| 
 | 
     3  | 
    Copyright   1995 TU Munich
  | 
| 
 | 
     4  | 
*)
  | 
| 
 | 
     5  | 
  | 
| 
15346
 | 
     6  | 
header {*Well-founded Relations*}
 | 
| 
 | 
     7  | 
  | 
| 
 | 
     8  | 
theory Wellfounded_Relations
  | 
| 
 | 
     9  | 
imports Finite_Set
  | 
| 
 | 
    10  | 
begin
  | 
| 
 | 
    11  | 
  | 
| 
 | 
    12  | 
text{*Derived WF relations such as inverse image, lexicographic product and
 | 
| 
 | 
    13  | 
measure. The simple relational product, in which @{term "(x',y')"} precedes
 | 
| 
 | 
    14  | 
@{term "(x,y)"} if @{term "x'<x"} and @{term "y'<y"}, is a subset of the
 | 
| 
 | 
    15  | 
lexicographic product, and therefore does not need to be defined separately.*}
  | 
| 
10213
 | 
    16  | 
  | 
| 
 | 
    17  | 
constdefs
  | 
| 
 | 
    18  | 
 less_than :: "(nat*nat)set"
  | 
| 
15346
 | 
    19  | 
    "less_than == trancl pred_nat"
  | 
| 
10213
 | 
    20  | 
  | 
| 
 | 
    21  | 
 measure   :: "('a => nat) => ('a * 'a)set"
 | 
| 
15346
 | 
    22  | 
    "measure == inv_image less_than"
  | 
| 
10213
 | 
    23  | 
  | 
| 
 | 
    24  | 
 lex_prod  :: "[('a*'a)set, ('b*'b)set] => (('a*'b)*('a*'b))set"
 | 
| 
 | 
    25  | 
               (infixr "<*lex*>" 80)
  | 
| 
15346
 | 
    26  | 
    "ra <*lex*> rb == {((a,b),(a',b')). (a,a') : ra | a=a' & (b,b') : rb}"
 | 
| 
 | 
    27  | 
  | 
| 
 | 
    28  | 
 finite_psubset  :: "('a set * 'a set) set"
 | 
| 
 | 
    29  | 
   --{* finite proper subset*}
 | 
| 
 | 
    30  | 
    "finite_psubset == {(A,B). A < B & finite B}"
 | 
| 
 | 
    31  | 
  | 
| 
 | 
    32  | 
 same_fst :: "('a => bool) => ('a => ('b * 'b)set) => (('a*'b)*('a*'b))set"
 | 
| 
 | 
    33  | 
    "same_fst P R == {((x',y'),(x,y)) . x'=x & P x & (y',y) : R x}"
 | 
| 
 | 
    34  | 
   --{*For @{text rec_def} declarations where the first n parameters
 | 
| 
 | 
    35  | 
       stay unchanged in the recursive call. 
  | 
| 
 | 
    36  | 
       See @{text "Library/While_Combinator.thy"} for an application.*}
 | 
| 
 | 
    37  | 
  | 
| 
 | 
    38  | 
  | 
| 
 | 
    39  | 
  | 
| 
 | 
    40  | 
  | 
| 
 | 
    41  | 
subsection{*Measure Functions make Wellfounded Relations*}
 | 
| 
 | 
    42  | 
  | 
| 
 | 
    43  | 
subsubsection{*`Less than' on the natural numbers*}
 | 
| 
 | 
    44  | 
  | 
| 
 | 
    45  | 
lemma wf_less_than [iff]: "wf less_than"
  | 
| 
 | 
    46  | 
by (simp add: less_than_def wf_pred_nat [THEN wf_trancl])
  | 
| 
 | 
    47  | 
  | 
| 
 | 
    48  | 
lemma trans_less_than [iff]: "trans less_than"
  | 
| 
 | 
    49  | 
by (simp add: less_than_def trans_trancl)
  | 
| 
 | 
    50  | 
  | 
| 
 | 
    51  | 
lemma less_than_iff [iff]: "((x,y): less_than) = (x<y)"
  | 
| 
 | 
    52  | 
by (simp add: less_than_def less_def)
  | 
| 
 | 
    53  | 
  | 
| 
 | 
    54  | 
lemma full_nat_induct:
  | 
| 
 | 
    55  | 
  assumes ih: "(!!n. (ALL m. Suc m <= n --> P m) ==> P n)"
  | 
| 
 | 
    56  | 
  shows "P n"
  | 
| 
 | 
    57  | 
apply (rule wf_less_than [THEN wf_induct])
  | 
| 
 | 
    58  | 
apply (rule ih, auto)
  | 
| 
 | 
    59  | 
done
  | 
| 
 | 
    60  | 
  | 
| 
 | 
    61  | 
subsubsection{*The Inverse Image into a Wellfounded Relation is Wellfounded.*}
 | 
| 
 | 
    62  | 
  | 
| 
 | 
    63  | 
lemma wf_inv_image [simp,intro!]: "wf(r) ==> wf(inv_image r (f::'a=>'b))"
  | 
| 
 | 
    64  | 
apply (simp (no_asm_use) add: inv_image_def wf_eq_minimal)
  | 
| 
 | 
    65  | 
apply clarify
  | 
| 
 | 
    66  | 
apply (subgoal_tac "EX (w::'b) . w : {w. EX (x::'a) . x: Q & (f x = w) }")
 | 
| 
 | 
    67  | 
prefer 2 apply (blast del: allE)
  | 
| 
 | 
    68  | 
apply (erule allE)
  | 
| 
 | 
    69  | 
apply (erule (1) notE impE)
  | 
| 
 | 
    70  | 
apply blast
  | 
| 
 | 
    71  | 
done
  | 
| 
10213
 | 
    72  | 
  | 
| 
15346
 | 
    73  | 
  | 
| 
 | 
    74  | 
subsubsection{*Finally, All Measures are Wellfounded.*}
 | 
| 
 | 
    75  | 
  | 
| 
 | 
    76  | 
lemma wf_measure [iff]: "wf (measure f)"
  | 
| 
 | 
    77  | 
apply (unfold measure_def)
  | 
| 
 | 
    78  | 
apply (rule wf_less_than [THEN wf_inv_image])
  | 
| 
 | 
    79  | 
done
  | 
| 
 | 
    80  | 
  | 
| 
18277
 | 
    81  | 
lemma measure_induct_rule [case_names less]:
  | 
| 
 | 
    82  | 
  fixes f :: "'a \<Rightarrow> nat"
  | 
| 
 | 
    83  | 
  assumes step: "\<And>x. (\<And>y. f y < f x \<Longrightarrow> P y) \<Longrightarrow> P x"
  | 
| 
 | 
    84  | 
  shows "P a"
  | 
| 
 | 
    85  | 
proof -
  | 
| 
 | 
    86  | 
  have "wf (measure f)" ..
  | 
| 
 | 
    87  | 
  then show ?thesis
  | 
| 
 | 
    88  | 
  proof induct
  | 
| 
 | 
    89  | 
    case (less x)
  | 
| 
 | 
    90  | 
    show ?case
  | 
| 
 | 
    91  | 
    proof (rule step)
  | 
| 
 | 
    92  | 
      fix y
  | 
| 
 | 
    93  | 
      assume "f y < f x"
  | 
| 
 | 
    94  | 
      then have "(y, x) \<in> measure f"
  | 
| 
 | 
    95  | 
        by (simp add: measure_def inv_image_def)
  | 
| 
 | 
    96  | 
      then show "P y" by (rule less)
  | 
| 
 | 
    97  | 
    qed
  | 
| 
 | 
    98  | 
  qed
  | 
| 
 | 
    99  | 
qed
  | 
| 
 | 
   100  | 
  | 
| 
 | 
   101  | 
lemma measure_induct:
  | 
| 
 | 
   102  | 
  fixes f :: "'a \<Rightarrow> nat"
  | 
| 
 | 
   103  | 
  shows "(\<And>x. \<forall>y. f y < f x \<longrightarrow> P y \<Longrightarrow> P x) \<Longrightarrow> P a"
  | 
| 
 | 
   104  | 
  by (rule measure_induct_rule [of f P a]) iprover
  | 
| 
15346
 | 
   105  | 
  | 
| 
 | 
   106  | 
  | 
| 
 | 
   107  | 
subsection{*Other Ways of Constructing Wellfounded Relations*}
 | 
| 
 | 
   108  | 
  | 
| 
 | 
   109  | 
text{*Wellfoundedness of lexicographic combinations*}
 | 
| 
 | 
   110  | 
lemma wf_lex_prod [intro!]: "[| wf(ra); wf(rb) |] ==> wf(ra <*lex*> rb)"
  | 
| 
 | 
   111  | 
apply (unfold wf_def lex_prod_def) 
  | 
| 
 | 
   112  | 
apply (rule allI, rule impI)
  | 
| 
 | 
   113  | 
apply (simp (no_asm_use) only: split_paired_All)
  | 
| 
 | 
   114  | 
apply (drule spec, erule mp) 
  | 
| 
 | 
   115  | 
apply (rule allI, rule impI)
  | 
| 
 | 
   116  | 
apply (drule spec, erule mp, blast) 
  | 
| 
 | 
   117  | 
done
  | 
| 
 | 
   118  | 
  | 
| 
 | 
   119  | 
text{*Transitivity of WF combinators.*}
 | 
| 
 | 
   120  | 
lemma trans_lex_prod [intro!]: 
  | 
| 
 | 
   121  | 
    "[| trans R1; trans R2 |] ==> trans (R1 <*lex*> R2)"
  | 
| 
 | 
   122  | 
by (unfold trans_def lex_prod_def, blast) 
  | 
| 
 | 
   123  | 
  | 
| 
 | 
   124  | 
subsubsection{*Wellfoundedness of proper subset on finite sets.*}
 | 
| 
 | 
   125  | 
lemma wf_finite_psubset: "wf(finite_psubset)"
  | 
| 
 | 
   126  | 
apply (unfold finite_psubset_def)
  | 
| 
 | 
   127  | 
apply (rule wf_measure [THEN wf_subset])
  | 
| 
 | 
   128  | 
apply (simp add: measure_def inv_image_def less_than_def less_def [symmetric])
  | 
| 
 | 
   129  | 
apply (fast elim!: psubset_card_mono)
  | 
| 
 | 
   130  | 
done
  | 
| 
 | 
   131  | 
  | 
| 
 | 
   132  | 
lemma trans_finite_psubset: "trans finite_psubset"
  | 
| 
 | 
   133  | 
by (simp add: finite_psubset_def psubset_def trans_def, blast)
  | 
| 
 | 
   134  | 
  | 
| 
 | 
   135  | 
  | 
| 
 | 
   136  | 
subsubsection{*Wellfoundedness of finite acyclic relations*}
 | 
| 
 | 
   137  | 
  | 
| 
 | 
   138  | 
text{*This proof belongs in this theory because it needs Finite.*}
 | 
| 
10213
 | 
   139  | 
  | 
| 
15346
 | 
   140  | 
lemma finite_acyclic_wf [rule_format]: "finite r ==> acyclic r --> wf r"
  | 
| 
 | 
   141  | 
apply (erule finite_induct, blast)
  | 
| 
 | 
   142  | 
apply (simp (no_asm_simp) only: split_tupled_all)
  | 
| 
 | 
   143  | 
apply simp
  | 
| 
 | 
   144  | 
done
  | 
| 
 | 
   145  | 
  | 
| 
 | 
   146  | 
lemma finite_acyclic_wf_converse: "[|finite r; acyclic r|] ==> wf (r^-1)"
  | 
| 
 | 
   147  | 
apply (erule finite_converse [THEN iffD2, THEN finite_acyclic_wf])
  | 
| 
 | 
   148  | 
apply (erule acyclic_converse [THEN iffD2])
  | 
| 
 | 
   149  | 
done
  | 
| 
 | 
   150  | 
  | 
| 
 | 
   151  | 
lemma wf_iff_acyclic_if_finite: "finite r ==> wf r = acyclic r"
  | 
| 
 | 
   152  | 
by (blast intro: finite_acyclic_wf wf_acyclic)
  | 
| 
 | 
   153  | 
  | 
| 
 | 
   154  | 
  | 
| 
15352
 | 
   155  | 
subsubsection{*Wellfoundedness of @{term same_fst}*}
 | 
| 
15346
 | 
   156  | 
  | 
| 
 | 
   157  | 
lemma same_fstI [intro!]:
  | 
| 
 | 
   158  | 
     "[| P x; (y',y) : R x |] ==> ((x,y'),(x,y)) : same_fst P R"
  | 
| 
 | 
   159  | 
by (simp add: same_fst_def)
  | 
| 
 | 
   160  | 
  | 
| 
 | 
   161  | 
lemma wf_same_fst:
  | 
| 
 | 
   162  | 
  assumes prem: "(!!x. P x ==> wf(R x))"
  | 
| 
 | 
   163  | 
  shows "wf(same_fst P R)"
  | 
| 
 | 
   164  | 
apply (simp cong del: imp_cong add: wf_def same_fst_def)
  | 
| 
 | 
   165  | 
apply (intro strip)
  | 
| 
 | 
   166  | 
apply (rename_tac a b)
  | 
| 
 | 
   167  | 
apply (case_tac "wf (R a)")
  | 
| 
 | 
   168  | 
 apply (erule_tac a = b in wf_induct, blast)
  | 
| 
 | 
   169  | 
apply (blast intro: prem)
  | 
| 
 | 
   170  | 
done
  | 
| 
 | 
   171  | 
  | 
| 
 | 
   172  | 
  | 
| 
 | 
   173  | 
subsection{*Weakly decreasing sequences (w.r.t. some well-founded order) 
 | 
| 
 | 
   174  | 
   stabilize.*}
  | 
| 
 | 
   175  | 
  | 
| 
 | 
   176  | 
text{*This material does not appear to be used any longer.*}
 | 
| 
 | 
   177  | 
  | 
| 
 | 
   178  | 
lemma lemma1: "[| ALL i. (f (Suc i), f i) : r^* |] ==> (f (i+k), f i) : r^*"
  | 
| 
 | 
   179  | 
apply (induct_tac "k", simp_all)
  | 
| 
 | 
   180  | 
apply (blast intro: rtrancl_trans)
  | 
| 
 | 
   181  | 
done
  | 
| 
 | 
   182  | 
  | 
| 
 | 
   183  | 
lemma lemma2: "[| ALL i. (f (Suc i), f i) : r^*; wf (r^+) |]  
  | 
| 
 | 
   184  | 
      ==> ALL m. f m = x --> (EX i. ALL k. f (m+i+k) = f (m+i))"
  | 
| 
 | 
   185  | 
apply (erule wf_induct, clarify)
  | 
| 
 | 
   186  | 
apply (case_tac "EX j. (f (m+j), f m) : r^+")
  | 
| 
 | 
   187  | 
 apply clarify
  | 
| 
 | 
   188  | 
 apply (subgoal_tac "EX i. ALL k. f ((m+j) +i+k) = f ( (m+j) +i) ")
  | 
| 
 | 
   189  | 
  apply clarify
  | 
| 
 | 
   190  | 
  apply (rule_tac x = "j+i" in exI)
  | 
| 
 | 
   191  | 
  apply (simp add: add_ac, blast)
  | 
| 
 | 
   192  | 
apply (rule_tac x = 0 in exI, clarsimp)
  | 
| 
 | 
   193  | 
apply (drule_tac i = m and k = k in lemma1)
  | 
| 
 | 
   194  | 
apply (blast elim: rtranclE dest: rtrancl_into_trancl1)
  | 
| 
 | 
   195  | 
done
  | 
| 
 | 
   196  | 
  | 
| 
 | 
   197  | 
lemma wf_weak_decr_stable: "[| ALL i. (f (Suc i), f i) : r^*; wf (r^+) |]  
  | 
| 
 | 
   198  | 
      ==> EX i. ALL k. f (i+k) = f i"
  | 
| 
 | 
   199  | 
apply (drule_tac x = 0 in lemma2 [THEN spec], auto)
  | 
| 
 | 
   200  | 
done
  | 
| 
 | 
   201  | 
  | 
| 
 | 
   202  | 
(* special case of the theorem above: <= *)
  | 
| 
 | 
   203  | 
lemma weak_decr_stable:
  | 
| 
 | 
   204  | 
     "ALL i. f (Suc i) <= ((f i)::nat) ==> EX i. ALL k. f (i+k) = f i"
  | 
| 
 | 
   205  | 
apply (rule_tac r = pred_nat in wf_weak_decr_stable)
  | 
| 
 | 
   206  | 
apply (simp add: pred_nat_trancl_eq_le)
  | 
| 
 | 
   207  | 
apply (intro wf_trancl wf_pred_nat)
  | 
| 
 | 
   208  | 
done
  | 
| 
 | 
   209  | 
  | 
| 
 | 
   210  | 
  | 
| 
 | 
   211  | 
ML
  | 
| 
 | 
   212  | 
{*
 | 
| 
 | 
   213  | 
val less_than_def = thm "less_than_def";
  | 
| 
 | 
   214  | 
val measure_def = thm "measure_def";
  | 
| 
 | 
   215  | 
val lex_prod_def = thm "lex_prod_def";
  | 
| 
 | 
   216  | 
val finite_psubset_def = thm "finite_psubset_def";
  | 
| 
 | 
   217  | 
  | 
| 
 | 
   218  | 
val wf_less_than = thm "wf_less_than";
  | 
| 
 | 
   219  | 
val trans_less_than = thm "trans_less_than";
  | 
| 
 | 
   220  | 
val less_than_iff = thm "less_than_iff";
  | 
| 
 | 
   221  | 
val full_nat_induct = thm "full_nat_induct";
  | 
| 
 | 
   222  | 
val wf_inv_image = thm "wf_inv_image";
  | 
| 
 | 
   223  | 
val wf_measure = thm "wf_measure";
  | 
| 
 | 
   224  | 
val measure_induct = thm "measure_induct";
  | 
| 
 | 
   225  | 
val wf_lex_prod = thm "wf_lex_prod";
  | 
| 
 | 
   226  | 
val trans_lex_prod = thm "trans_lex_prod";
  | 
| 
 | 
   227  | 
val wf_finite_psubset = thm "wf_finite_psubset";
  | 
| 
 | 
   228  | 
val trans_finite_psubset = thm "trans_finite_psubset";
  | 
| 
 | 
   229  | 
val finite_acyclic_wf = thm "finite_acyclic_wf";
  | 
| 
 | 
   230  | 
val finite_acyclic_wf_converse = thm "finite_acyclic_wf_converse";
  | 
| 
 | 
   231  | 
val wf_iff_acyclic_if_finite = thm "wf_iff_acyclic_if_finite";
  | 
| 
 | 
   232  | 
val wf_weak_decr_stable = thm "wf_weak_decr_stable";
  | 
| 
 | 
   233  | 
val weak_decr_stable = thm "weak_decr_stable";
  | 
| 
 | 
   234  | 
val same_fstI = thm "same_fstI";
  | 
| 
 | 
   235  | 
val wf_same_fst = thm "wf_same_fst";
  | 
| 
 | 
   236  | 
*}
  | 
| 
 | 
   237  | 
  | 
| 
10213
 | 
   238  | 
  | 
| 
 | 
   239  | 
end
  |