| author | haftmann |
| Fri, 30 Nov 2007 20:13:03 +0100 | |
| changeset 25510 | 38c15efe603b |
| parent 25481 | aa16cd919dcc |
| child 25517 | 36d710d1dbce |
| 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:
19769
diff
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:
19623
diff
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:
19623
diff
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:
19623
diff
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 |
|
| 24853 | 110 |
(* Should go into Finite_Set, but needs measure. |
111 |
Maybe move Wf_Rel before Finite_Set and finite_psubset to Finite_set? |
|
112 |
*) |
|
113 |
lemma (in linorder) |
|
114 |
finite_linorder_induct[consumes 1, case_names empty insert]: |
|
115 |
"finite A \<Longrightarrow> P {} \<Longrightarrow>
|
|
| 25062 | 116 |
(!!A b. finite A \<Longrightarrow> ALL a:A. a < b \<Longrightarrow> P A \<Longrightarrow> P(insert b A)) |
| 24853 | 117 |
\<Longrightarrow> P A" |
118 |
proof (induct A rule: measure_induct[where f=card]) |
|
119 |
fix A :: "'a set" |
|
120 |
assume IH: "ALL B. card B < card A \<longrightarrow> finite B \<longrightarrow> P {} \<longrightarrow>
|
|
| 25062 | 121 |
(\<forall>A b. finite A \<longrightarrow> (\<forall>a\<in>A. a<b) \<longrightarrow> P A \<longrightarrow> P (insert b A)) |
| 24853 | 122 |
\<longrightarrow> P B" |
123 |
and "finite A" and "P {}"
|
|
| 25062 | 124 |
and step: "!!A b. \<lbrakk>finite A; \<forall>a\<in>A. a < b; P A\<rbrakk> \<Longrightarrow> P (insert b A)" |
| 24853 | 125 |
show "P A" |
126 |
proof cases |
|
127 |
assume "A = {}" thus "P A" using `P {}` by simp
|
|
128 |
next |
|
129 |
let ?B = "A - {Max A}" let ?A = "insert (Max A) ?B"
|
|
130 |
assume "A \<noteq> {}"
|
|
131 |
with `finite A` have "Max A : A" by auto |
|
132 |
hence A: "?A = A" using insert_Diff_single insert_absorb by auto |
|
133 |
note card_Diff1_less[OF `finite A` `Max A : A`] |
|
134 |
moreover have "finite ?B" using `finite A` by simp |
|
135 |
ultimately have "P ?B" using `P {}` step IH by blast
|
|
| 25062 | 136 |
moreover have "\<forall>a\<in>?B. a < Max A" |
| 24853 | 137 |
using Max_ge[OF `finite A` `A \<noteq> {}`] by fastsimp
|
138 |
ultimately show "P A" |
|
139 |
using A insert_Diff_single step[OF `finite ?B`] by fastsimp |
|
140 |
qed |
|
141 |
qed |
|
142 |
||
| 15346 | 143 |
|
144 |
subsection{*Other Ways of Constructing Wellfounded Relations*}
|
|
145 |
||
146 |
text{*Wellfoundedness of lexicographic combinations*}
|
|
147 |
lemma wf_lex_prod [intro!]: "[| wf(ra); wf(rb) |] ==> wf(ra <*lex*> rb)" |
|
148 |
apply (unfold wf_def lex_prod_def) |
|
149 |
apply (rule allI, rule impI) |
|
150 |
apply (simp (no_asm_use) only: split_paired_All) |
|
151 |
apply (drule spec, erule mp) |
|
152 |
apply (rule allI, rule impI) |
|
153 |
apply (drule spec, erule mp, blast) |
|
154 |
done |
|
155 |
||
|
19769
c40ce2de2020
Added [simp]-lemmas "in_inv_image" and "in_lex_prod" in the spirit of "in_measure".
krauss
parents:
19623
diff
changeset
|
156 |
lemma in_lex_prod[simp]: |
|
c40ce2de2020
Added [simp]-lemmas "in_inv_image" and "in_lex_prod" in the spirit of "in_measure".
krauss
parents:
19623
diff
changeset
|
157 |
"(((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:
19623
diff
changeset
|
158 |
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:
19623
diff
changeset
|
159 |
|
|
24575
8b5c5eb7df87
added "<*mlex*>" which lexicographically combines a measure function with a relation
krauss
parents:
22261
diff
changeset
|
160 |
text {* lexicographic combinations with measure functions *}
|
|
8b5c5eb7df87
added "<*mlex*>" which lexicographically combines a measure function with a relation
krauss
parents:
22261
diff
changeset
|
161 |
|
|
8b5c5eb7df87
added "<*mlex*>" which lexicographically combines a measure function with a relation
krauss
parents:
22261
diff
changeset
|
162 |
definition |
|
8b5c5eb7df87
added "<*mlex*>" which lexicographically combines a measure function with a relation
krauss
parents:
22261
diff
changeset
|
163 |
mlex_prod :: "('a \<Rightarrow> nat) \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set" (infixr "<*mlex*>" 80)
|
|
8b5c5eb7df87
added "<*mlex*>" which lexicographically combines a measure function with a relation
krauss
parents:
22261
diff
changeset
|
164 |
where |
|
8b5c5eb7df87
added "<*mlex*>" which lexicographically combines a measure function with a relation
krauss
parents:
22261
diff
changeset
|
165 |
"f <*mlex*> R = inv_image (less_than <*lex*> R) (%x. (f x, x))" |
|
8b5c5eb7df87
added "<*mlex*>" which lexicographically combines a measure function with a relation
krauss
parents:
22261
diff
changeset
|
166 |
|
|
8b5c5eb7df87
added "<*mlex*>" which lexicographically combines a measure function with a relation
krauss
parents:
22261
diff
changeset
|
167 |
lemma wf_mlex: "wf R \<Longrightarrow> wf (f <*mlex*> R)" |
|
8b5c5eb7df87
added "<*mlex*>" which lexicographically combines a measure function with a relation
krauss
parents:
22261
diff
changeset
|
168 |
unfolding mlex_prod_def |
|
8b5c5eb7df87
added "<*mlex*>" which lexicographically combines a measure function with a relation
krauss
parents:
22261
diff
changeset
|
169 |
by auto |
|
8b5c5eb7df87
added "<*mlex*>" which lexicographically combines a measure function with a relation
krauss
parents:
22261
diff
changeset
|
170 |
|
|
8b5c5eb7df87
added "<*mlex*>" which lexicographically combines a measure function with a relation
krauss
parents:
22261
diff
changeset
|
171 |
lemma mlex_less: "f x < f y \<Longrightarrow> (x, y) \<in> f <*mlex*> R" |
|
8b5c5eb7df87
added "<*mlex*>" which lexicographically combines a measure function with a relation
krauss
parents:
22261
diff
changeset
|
172 |
unfolding mlex_prod_def by simp |
|
8b5c5eb7df87
added "<*mlex*>" which lexicographically combines a measure function with a relation
krauss
parents:
22261
diff
changeset
|
173 |
|
|
8b5c5eb7df87
added "<*mlex*>" which lexicographically combines a measure function with a relation
krauss
parents:
22261
diff
changeset
|
174 |
lemma mlex_leq: "f x \<le> f y \<Longrightarrow> (x, y) \<in> R \<Longrightarrow> (x, y) \<in> f <*mlex*> R" |
|
8b5c5eb7df87
added "<*mlex*>" which lexicographically combines a measure function with a relation
krauss
parents:
22261
diff
changeset
|
175 |
unfolding mlex_prod_def by auto |
|
8b5c5eb7df87
added "<*mlex*>" which lexicographically combines a measure function with a relation
krauss
parents:
22261
diff
changeset
|
176 |
|
|
8b5c5eb7df87
added "<*mlex*>" which lexicographically combines a measure function with a relation
krauss
parents:
22261
diff
changeset
|
177 |
|
| 15346 | 178 |
text{*Transitivity of WF combinators.*}
|
179 |
lemma trans_lex_prod [intro!]: |
|
180 |
"[| trans R1; trans R2 |] ==> trans (R1 <*lex*> R2)" |
|
181 |
by (unfold trans_def lex_prod_def, blast) |
|
182 |
||
183 |
subsubsection{*Wellfoundedness of proper subset on finite sets.*}
|
|
184 |
lemma wf_finite_psubset: "wf(finite_psubset)" |
|
185 |
apply (unfold finite_psubset_def) |
|
186 |
apply (rule wf_measure [THEN wf_subset]) |
|
187 |
apply (simp add: measure_def inv_image_def less_than_def less_def [symmetric]) |
|
188 |
apply (fast elim!: psubset_card_mono) |
|
189 |
done |
|
190 |
||
191 |
lemma trans_finite_psubset: "trans finite_psubset" |
|
192 |
by (simp add: finite_psubset_def psubset_def trans_def, blast) |
|
193 |
||
194 |
||
195 |
subsubsection{*Wellfoundedness of finite acyclic relations*}
|
|
196 |
||
197 |
text{*This proof belongs in this theory because it needs Finite.*}
|
|
| 10213 | 198 |
|
| 15346 | 199 |
lemma finite_acyclic_wf [rule_format]: "finite r ==> acyclic r --> wf r" |
200 |
apply (erule finite_induct, blast) |
|
201 |
apply (simp (no_asm_simp) only: split_tupled_all) |
|
202 |
apply simp |
|
203 |
done |
|
204 |
||
205 |
lemma finite_acyclic_wf_converse: "[|finite r; acyclic r|] ==> wf (r^-1)" |
|
206 |
apply (erule finite_converse [THEN iffD2, THEN finite_acyclic_wf]) |
|
207 |
apply (erule acyclic_converse [THEN iffD2]) |
|
208 |
done |
|
209 |
||
210 |
lemma wf_iff_acyclic_if_finite: "finite r ==> wf r = acyclic r" |
|
211 |
by (blast intro: finite_acyclic_wf wf_acyclic) |
|
212 |
||
213 |
||
| 15352 | 214 |
subsubsection{*Wellfoundedness of @{term same_fst}*}
|
| 15346 | 215 |
|
216 |
lemma same_fstI [intro!]: |
|
217 |
"[| P x; (y',y) : R x |] ==> ((x,y'),(x,y)) : same_fst P R" |
|
218 |
by (simp add: same_fst_def) |
|
219 |
||
220 |
lemma wf_same_fst: |
|
221 |
assumes prem: "(!!x. P x ==> wf(R x))" |
|
222 |
shows "wf(same_fst P R)" |
|
223 |
apply (simp cong del: imp_cong add: wf_def same_fst_def) |
|
224 |
apply (intro strip) |
|
225 |
apply (rename_tac a b) |
|
226 |
apply (case_tac "wf (R a)") |
|
227 |
apply (erule_tac a = b in wf_induct, blast) |
|
228 |
apply (blast intro: prem) |
|
229 |
done |
|
230 |
||
231 |
||
232 |
subsection{*Weakly decreasing sequences (w.r.t. some well-founded order)
|
|
233 |
stabilize.*} |
|
234 |
||
235 |
text{*This material does not appear to be used any longer.*}
|
|
236 |
||
237 |
lemma lemma1: "[| ALL i. (f (Suc i), f i) : r^* |] ==> (f (i+k), f i) : r^*" |
|
238 |
apply (induct_tac "k", simp_all) |
|
239 |
apply (blast intro: rtrancl_trans) |
|
240 |
done |
|
241 |
||
242 |
lemma lemma2: "[| ALL i. (f (Suc i), f i) : r^*; wf (r^+) |] |
|
243 |
==> ALL m. f m = x --> (EX i. ALL k. f (m+i+k) = f (m+i))" |
|
244 |
apply (erule wf_induct, clarify) |
|
245 |
apply (case_tac "EX j. (f (m+j), f m) : r^+") |
|
246 |
apply clarify |
|
247 |
apply (subgoal_tac "EX i. ALL k. f ((m+j) +i+k) = f ( (m+j) +i) ") |
|
248 |
apply clarify |
|
249 |
apply (rule_tac x = "j+i" in exI) |
|
250 |
apply (simp add: add_ac, blast) |
|
251 |
apply (rule_tac x = 0 in exI, clarsimp) |
|
252 |
apply (drule_tac i = m and k = k in lemma1) |
|
253 |
apply (blast elim: rtranclE dest: rtrancl_into_trancl1) |
|
254 |
done |
|
255 |
||
256 |
lemma wf_weak_decr_stable: "[| ALL i. (f (Suc i), f i) : r^*; wf (r^+) |] |
|
257 |
==> EX i. ALL k. f (i+k) = f i" |
|
258 |
apply (drule_tac x = 0 in lemma2 [THEN spec], auto) |
|
259 |
done |
|
260 |
||
261 |
(* special case of the theorem above: <= *) |
|
262 |
lemma weak_decr_stable: |
|
263 |
"ALL i. f (Suc i) <= ((f i)::nat) ==> EX i. ALL k. f (i+k) = f i" |
|
264 |
apply (rule_tac r = pred_nat in wf_weak_decr_stable) |
|
265 |
apply (simp add: pred_nat_trancl_eq_le) |
|
266 |
apply (intro wf_trancl wf_pred_nat) |
|
267 |
done |
|
268 |
||
| 10213 | 269 |
end |