| author | huffman | 
| Tue, 21 Aug 2007 02:30:14 +0200 | |
| changeset 24376 | e403ab5c9415 | 
| parent 22261 | 9e185f78e7d4 | 
| child 24575 | 8b5c5eb7df87 | 
| permissions | -rw-r--r-- | 
| 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" | |
| 22261 
9e185f78e7d4
Adapted to changes in Transitive_Closure theory.
 berghofe parents: 
19769diff
changeset | 19 | "less_than == 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 | |
| 19769 
c40ce2de2020
Added [simp]-lemmas "in_inv_image" and "in_lex_prod" in the spirit of "in_measure".
 krauss parents: 
19623diff
changeset | 73 | lemma in_inv_image[simp]: "((x,y) : inv_image r f) = ((f x, f y) : r)" | 
| 
c40ce2de2020
Added [simp]-lemmas "in_inv_image" and "in_lex_prod" in the spirit of "in_measure".
 krauss parents: 
19623diff
changeset | 74 | by (auto simp:inv_image_def) | 
| 15346 | 75 | |
| 76 | subsubsection{*Finally, All Measures are Wellfounded.*}
 | |
| 77 | ||
| 19623 | 78 | lemma in_measure[simp]: "((x,y) : measure f) = (f x < f y)" | 
| 19769 
c40ce2de2020
Added [simp]-lemmas "in_inv_image" and "in_lex_prod" in the spirit of "in_measure".
 krauss parents: 
19623diff
changeset | 79 | by (simp add:measure_def) | 
| 19623 | 80 | |
| 15346 | 81 | lemma wf_measure [iff]: "wf (measure f)" | 
| 82 | apply (unfold measure_def) | |
| 83 | apply (rule wf_less_than [THEN wf_inv_image]) | |
| 84 | done | |
| 85 | ||
| 18277 | 86 | lemma measure_induct_rule [case_names less]: | 
| 87 | fixes f :: "'a \<Rightarrow> nat" | |
| 88 | assumes step: "\<And>x. (\<And>y. f y < f x \<Longrightarrow> P y) \<Longrightarrow> P x" | |
| 89 | shows "P a" | |
| 90 | proof - | |
| 91 | have "wf (measure f)" .. | |
| 92 | then show ?thesis | |
| 93 | proof induct | |
| 94 | case (less x) | |
| 95 | show ?case | |
| 96 | proof (rule step) | |
| 97 | fix y | |
| 98 | assume "f y < f x" | |
| 19623 | 99 | hence "(y, x) \<in> measure f" by simp | 
| 100 | thus "P y" by (rule less) | |
| 18277 | 101 | qed | 
| 102 | qed | |
| 103 | qed | |
| 104 | ||
| 105 | lemma measure_induct: | |
| 106 | fixes f :: "'a \<Rightarrow> nat" | |
| 107 | shows "(\<And>x. \<forall>y. f y < f x \<longrightarrow> P y \<Longrightarrow> P x) \<Longrightarrow> P a" | |
| 108 | by (rule measure_induct_rule [of f P a]) iprover | |
| 15346 | 109 | |
| 110 | ||
| 111 | subsection{*Other Ways of Constructing Wellfounded Relations*}
 | |
| 112 | ||
| 113 | text{*Wellfoundedness of lexicographic combinations*}
 | |
| 114 | lemma wf_lex_prod [intro!]: "[| wf(ra); wf(rb) |] ==> wf(ra <*lex*> rb)" | |
| 115 | apply (unfold wf_def lex_prod_def) | |
| 116 | apply (rule allI, rule impI) | |
| 117 | apply (simp (no_asm_use) only: split_paired_All) | |
| 118 | apply (drule spec, erule mp) | |
| 119 | apply (rule allI, rule impI) | |
| 120 | apply (drule spec, erule mp, blast) | |
| 121 | done | |
| 122 | ||
| 19769 
c40ce2de2020
Added [simp]-lemmas "in_inv_image" and "in_lex_prod" in the spirit of "in_measure".
 krauss parents: 
19623diff
changeset | 123 | lemma in_lex_prod[simp]: | 
| 
c40ce2de2020
Added [simp]-lemmas "in_inv_image" and "in_lex_prod" in the spirit of "in_measure".
 krauss parents: 
19623diff
changeset | 124 | "(((a,b),(a',b')): r <*lex*> s) = ((a,a'): r \<or> (a = a' \<and> (b, b') : s))" | 
| 
c40ce2de2020
Added [simp]-lemmas "in_inv_image" and "in_lex_prod" in the spirit of "in_measure".
 krauss parents: 
19623diff
changeset | 125 | by (auto simp:lex_prod_def) | 
| 
c40ce2de2020
Added [simp]-lemmas "in_inv_image" and "in_lex_prod" in the spirit of "in_measure".
 krauss parents: 
19623diff
changeset | 126 | |
| 15346 | 127 | text{*Transitivity of WF combinators.*}
 | 
| 128 | lemma trans_lex_prod [intro!]: | |
| 129 | "[| trans R1; trans R2 |] ==> trans (R1 <*lex*> R2)" | |
| 130 | by (unfold trans_def lex_prod_def, blast) | |
| 131 | ||
| 132 | subsubsection{*Wellfoundedness of proper subset on finite sets.*}
 | |
| 133 | lemma wf_finite_psubset: "wf(finite_psubset)" | |
| 134 | apply (unfold finite_psubset_def) | |
| 135 | apply (rule wf_measure [THEN wf_subset]) | |
| 136 | apply (simp add: measure_def inv_image_def less_than_def less_def [symmetric]) | |
| 137 | apply (fast elim!: psubset_card_mono) | |
| 138 | done | |
| 139 | ||
| 140 | lemma trans_finite_psubset: "trans finite_psubset" | |
| 141 | by (simp add: finite_psubset_def psubset_def trans_def, blast) | |
| 142 | ||
| 143 | ||
| 144 | subsubsection{*Wellfoundedness of finite acyclic relations*}
 | |
| 145 | ||
| 146 | text{*This proof belongs in this theory because it needs Finite.*}
 | |
| 10213 | 147 | |
| 15346 | 148 | lemma finite_acyclic_wf [rule_format]: "finite r ==> acyclic r --> wf r" | 
| 149 | apply (erule finite_induct, blast) | |
| 150 | apply (simp (no_asm_simp) only: split_tupled_all) | |
| 151 | apply simp | |
| 152 | done | |
| 153 | ||
| 154 | lemma finite_acyclic_wf_converse: "[|finite r; acyclic r|] ==> wf (r^-1)" | |
| 155 | apply (erule finite_converse [THEN iffD2, THEN finite_acyclic_wf]) | |
| 156 | apply (erule acyclic_converse [THEN iffD2]) | |
| 157 | done | |
| 158 | ||
| 159 | lemma wf_iff_acyclic_if_finite: "finite r ==> wf r = acyclic r" | |
| 160 | by (blast intro: finite_acyclic_wf wf_acyclic) | |
| 161 | ||
| 162 | ||
| 15352 | 163 | subsubsection{*Wellfoundedness of @{term same_fst}*}
 | 
| 15346 | 164 | |
| 165 | lemma same_fstI [intro!]: | |
| 166 | "[| P x; (y',y) : R x |] ==> ((x,y'),(x,y)) : same_fst P R" | |
| 167 | by (simp add: same_fst_def) | |
| 168 | ||
| 169 | lemma wf_same_fst: | |
| 170 | assumes prem: "(!!x. P x ==> wf(R x))" | |
| 171 | shows "wf(same_fst P R)" | |
| 172 | apply (simp cong del: imp_cong add: wf_def same_fst_def) | |
| 173 | apply (intro strip) | |
| 174 | apply (rename_tac a b) | |
| 175 | apply (case_tac "wf (R a)") | |
| 176 | apply (erule_tac a = b in wf_induct, blast) | |
| 177 | apply (blast intro: prem) | |
| 178 | done | |
| 179 | ||
| 180 | ||
| 181 | subsection{*Weakly decreasing sequences (w.r.t. some well-founded order) 
 | |
| 182 | stabilize.*} | |
| 183 | ||
| 184 | text{*This material does not appear to be used any longer.*}
 | |
| 185 | ||
| 186 | lemma lemma1: "[| ALL i. (f (Suc i), f i) : r^* |] ==> (f (i+k), f i) : r^*" | |
| 187 | apply (induct_tac "k", simp_all) | |
| 188 | apply (blast intro: rtrancl_trans) | |
| 189 | done | |
| 190 | ||
| 191 | lemma lemma2: "[| ALL i. (f (Suc i), f i) : r^*; wf (r^+) |] | |
| 192 | ==> ALL m. f m = x --> (EX i. ALL k. f (m+i+k) = f (m+i))" | |
| 193 | apply (erule wf_induct, clarify) | |
| 194 | apply (case_tac "EX j. (f (m+j), f m) : r^+") | |
| 195 | apply clarify | |
| 196 | apply (subgoal_tac "EX i. ALL k. f ((m+j) +i+k) = f ( (m+j) +i) ") | |
| 197 | apply clarify | |
| 198 | apply (rule_tac x = "j+i" in exI) | |
| 199 | apply (simp add: add_ac, blast) | |
| 200 | apply (rule_tac x = 0 in exI, clarsimp) | |
| 201 | apply (drule_tac i = m and k = k in lemma1) | |
| 202 | apply (blast elim: rtranclE dest: rtrancl_into_trancl1) | |
| 203 | done | |
| 204 | ||
| 205 | lemma wf_weak_decr_stable: "[| ALL i. (f (Suc i), f i) : r^*; wf (r^+) |] | |
| 206 | ==> EX i. ALL k. f (i+k) = f i" | |
| 207 | apply (drule_tac x = 0 in lemma2 [THEN spec], auto) | |
| 208 | done | |
| 209 | ||
| 210 | (* special case of the theorem above: <= *) | |
| 211 | lemma weak_decr_stable: | |
| 212 | "ALL i. f (Suc i) <= ((f i)::nat) ==> EX i. ALL k. f (i+k) = f i" | |
| 213 | apply (rule_tac r = pred_nat in wf_weak_decr_stable) | |
| 214 | apply (simp add: pred_nat_trancl_eq_le) | |
| 215 | apply (intro wf_trancl wf_pred_nat) | |
| 216 | done | |
| 217 | ||
| 218 | ||
| 219 | ML | |
| 220 | {*
 | |
| 221 | val less_than_def = thm "less_than_def"; | |
| 222 | val measure_def = thm "measure_def"; | |
| 223 | val lex_prod_def = thm "lex_prod_def"; | |
| 224 | val finite_psubset_def = thm "finite_psubset_def"; | |
| 225 | ||
| 226 | val wf_less_than = thm "wf_less_than"; | |
| 227 | val trans_less_than = thm "trans_less_than"; | |
| 228 | val less_than_iff = thm "less_than_iff"; | |
| 229 | val full_nat_induct = thm "full_nat_induct"; | |
| 230 | val wf_inv_image = thm "wf_inv_image"; | |
| 231 | val wf_measure = thm "wf_measure"; | |
| 232 | val measure_induct = thm "measure_induct"; | |
| 233 | val wf_lex_prod = thm "wf_lex_prod"; | |
| 234 | val trans_lex_prod = thm "trans_lex_prod"; | |
| 235 | val wf_finite_psubset = thm "wf_finite_psubset"; | |
| 236 | val trans_finite_psubset = thm "trans_finite_psubset"; | |
| 237 | val finite_acyclic_wf = thm "finite_acyclic_wf"; | |
| 238 | val finite_acyclic_wf_converse = thm "finite_acyclic_wf_converse"; | |
| 239 | val wf_iff_acyclic_if_finite = thm "wf_iff_acyclic_if_finite"; | |
| 240 | val wf_weak_decr_stable = thm "wf_weak_decr_stable"; | |
| 241 | val weak_decr_stable = thm "weak_decr_stable"; | |
| 242 | val same_fstI = thm "same_fstI"; | |
| 243 | val wf_same_fst = thm "wf_same_fst"; | |
| 244 | *} | |
| 245 | ||
| 10213 | 246 | |
| 247 | end |