3 Author: Lawrence C Paulson, Cambridge University Computer Laboratory
4 Author: Konrad Slind, Alexander Krauss
5 Copyright 1992-2008 University of Cambridge and TU Muenchen
8 header {*Well-founded Recursion*}
11 imports Finite_Set Nat
12 uses ("Tools/function_package/size.ML")
15 subsection {* Basic Definitions *}
18 wfrec_rel :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => 'a => 'b => bool"
19 for R :: "('a * 'a) set"
20 and F :: "('a => 'b) => 'a => 'b"
22 wfrecI: "ALL z. (z, x) : R --> wfrec_rel R F z (g z) ==>
23 wfrec_rel R F x (F g x)"
26 wf :: "('a * 'a)set => bool"
27 "wf(r) == (!P. (!x. (!y. (y,x):r --> P(y)) --> P(x)) --> (!x. P(x)))"
29 wfP :: "('a => 'a => bool) => bool"
30 "wfP r == wf {(x, y). r x y}"
32 acyclic :: "('a*'a)set => bool"
33 "acyclic r == !x. (x,x) ~: r^+"
35 cut :: "('a => 'b) => ('a * 'a)set => 'a => 'a => 'b"
36 "cut f r x == (%y. if (y,x):r then f y else arbitrary)"
38 adm_wf :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => bool"
39 "adm_wf R F == ALL f g x.
40 (ALL z. (z, x) : R --> f z = g z) --> F f x = F g x"
42 wfrec :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => 'a => 'b"
43 [code func del]: "wfrec R F == %x. THE y. wfrec_rel R (%f x. F (cut f R x) x) x y"
45 abbreviation acyclicP :: "('a => 'a => bool) => bool" where
46 "acyclicP r == acyclic {(x, y). r x y}"
48 lemma wfP_wf_eq [pred_set_conv]: "wfP (\<lambda>x y. (x, y) \<in> r) = wf r"
49 by (simp add: wfP_def)
52 "(!!P x. (ALL x. (ALL y. (y,x) : r --> P(y)) --> P(x)) ==> P(x)) ==> wf(r)"
53 unfolding wf_def by blast
55 lemmas wfPUNIVI = wfUNIVI [to_pred]
57 text{*Restriction to domain @{term A} and range @{term B}. If @{term r} is
58 well-founded over their intersection, then @{term "wf r"}*}
60 "[| r \<subseteq> A <*> B;
61 !!x P. [|\<forall>x. (\<forall>y. (y,x) : r --> P y) --> P x; x : A; x : B |] ==> P x |]
63 unfolding wf_def by blast
67 !!x.[| ALL y. (y,x): r --> P(y) |] ==> P(x)
69 unfolding wf_def by blast
71 lemmas wfP_induct = wf_induct [to_pred]
73 lemmas wf_induct_rule = wf_induct [rule_format, consumes 1, case_names less, induct set: wf]
75 lemmas wfP_induct_rule = wf_induct_rule [to_pred, induct set: wfP]
77 lemma wf_not_sym: "wf r ==> (a, x) : r ==> (x, a) ~: r"
78 by (induct a arbitrary: x set: wf) blast
80 (* [| wf r; ~Z ==> (a,x) : r; (x,a) ~: r ==> Z |] ==> Z *)
81 lemmas wf_asym = wf_not_sym [elim_format]
83 lemma wf_not_refl [simp]: "wf r ==> (a, a) ~: r"
84 by (blast elim: wf_asym)
86 (* [| wf r; (a,a) ~: r ==> PROP W |] ==> PROP W *)
87 lemmas wf_irrefl = wf_not_refl [elim_format]
90 assumes wf: "wf {(x::'a::ord, y). x < y}"
91 assumes lin: "OFCLASS('a::ord, linorder_class)"
92 shows "OFCLASS('a::ord, wellorder_class)"
93 using lin by (rule wellorder_class.intro)
94 (blast intro: wellorder_axioms.intro wf_induct_rule [OF wf])
96 lemma (in wellorder) wf:
98 unfolding wf_def by (blast intro: less_induct)
101 subsection {* Basic Results *}
103 text{*transitive closure of a well-founded relation is well-founded! *}
110 assume induct_step: "!!x. (!!y. (y, x) : r^+ ==> P y) ==> P x"
112 proof (rule induct_step)
113 fix y assume "(y, x) : r^+"
114 with `wf r` show "P y"
115 proof (induct x arbitrary: y)
117 note hyp = `\<And>x' y'. (x', x) : r ==> (y', x') : r^+ ==> P y'`
118 from `(y, x) : r^+` show "P y"
122 proof (rule induct_step)
123 fix y' assume "(y', y) : r^+"
124 with `(y, x) : r` show "P y'" by (rule hyp [of y y'])
128 then obtain x' where "(x', x) : r" and "(y, x') : r^+" by simp
129 then show "P y" by (rule hyp [of x' y])
133 } then show ?thesis unfolding wf_def by blast
136 lemmas wfP_trancl = wf_trancl [to_pred]
138 lemma wf_converse_trancl: "wf (r^-1) ==> wf ((r^+)^-1)"
139 apply (subst trancl_converse [symmetric])
140 apply (erule wf_trancl)
144 text{*Minimal-element characterization of well-foundedness*}
145 lemma wf_eq_minimal: "wf r = (\<forall>Q x. x\<in>Q --> (\<exists>z\<in>Q. \<forall>y. (y,z)\<in>r --> y\<notin>Q))"
146 proof (intro iffI strip)
147 fix Q :: "'a set" and x
148 assume "wf r" and "x \<in> Q"
149 then show "\<exists>z\<in>Q. \<forall>y. (y, z) \<in> r \<longrightarrow> y \<notin> Q"
151 by (blast dest: spec [of _ "%x. x\<in>Q \<longrightarrow> (\<exists>z\<in>Q. \<forall>y. (y,z) \<in> r \<longrightarrow> y\<notin>Q)"])
153 assume 1: "\<forall>Q x. x \<in> Q \<longrightarrow> (\<exists>z\<in>Q. \<forall>y. (y, z) \<in> r \<longrightarrow> y \<notin> Q)"
156 fix P :: "'a \<Rightarrow> bool" and x
157 assume 2: "\<forall>x. (\<forall>y. (y, x) \<in> r \<longrightarrow> P y) \<longrightarrow> P x"
158 let ?Q = "{x. \<not> P x}"
159 have "x \<in> ?Q \<longrightarrow> (\<exists>z \<in> ?Q. \<forall>y. (y, z) \<in> r \<longrightarrow> y \<notin> ?Q)"
160 by (rule 1 [THEN spec, THEN spec])
161 then have "\<not> P x \<longrightarrow> (\<exists>z. \<not> P z \<and> (\<forall>y. (y, z) \<in> r \<longrightarrow> P y))" by simp
162 with 2 have "\<not> P x \<longrightarrow> (\<exists>z. \<not> P z \<and> P z)" by fast
163 then show "P x" by simp
168 assumes "wf R" "x \<in> Q"
169 obtains z where "z \<in> Q" "\<And>y. (y, z) \<in> R \<Longrightarrow> y \<notin> Q"
170 using assms unfolding wf_eq_minimal by blast
173 "(\<And>x Q. x \<in> Q \<Longrightarrow> \<exists>z\<in>Q. \<forall>y. (y, z) \<in> R \<longrightarrow> y \<notin> Q)
174 \<Longrightarrow> wf R"
175 unfolding wf_eq_minimal by blast
177 lemmas wfP_eq_minimal = wf_eq_minimal [to_pred]
179 text {* Well-foundedness of subsets *}
180 lemma wf_subset: "[| wf(r); p<=r |] ==> wf(p)"
181 apply (simp (no_asm_use) add: wf_eq_minimal)
185 lemmas wfP_subset = wf_subset [to_pred]
187 text {* Well-foundedness of the empty relation *}
188 lemma wf_empty [iff]: "wf({})"
189 by (simp add: wf_def)
191 lemmas wfP_empty [iff] =
192 wf_empty [to_pred bot_empty_eq2, simplified bot_fun_eq bot_bool_eq]
194 lemma wf_Int1: "wf r ==> wf (r Int r')"
195 apply (erule wf_subset)
196 apply (rule Int_lower1)
199 lemma wf_Int2: "wf r ==> wf (r' Int r)"
200 apply (erule wf_subset)
201 apply (rule Int_lower2)
204 text{*Well-foundedness of insert*}
205 lemma wf_insert [iff]: "wf(insert (y,x) r) = (wf(r) & (x,y) ~: r^*)"
207 apply (blast elim: wf_trancl [THEN wf_irrefl]
208 intro: rtrancl_into_trancl1 wf_subset
209 rtrancl_mono [THEN [2] rev_subsetD])
210 apply (simp add: wf_eq_minimal, safe)
211 apply (rule allE, assumption, erule impE, blast)
213 apply (rename_tac "a", case_tac "a = x")
216 apply (case_tac "y:Q")
218 apply (rule_tac x = "{z. z:Q & (z,y) : r^*}" in allE)
220 apply (erule_tac V = "ALL Q. (EX x. x : Q) --> ?P Q" in thin_rl)
221 --{*essential for speed*}
222 txt{*Blast with new substOccur fails*}
223 apply (fast intro: converse_rtrancl_into_rtrancl)
226 text{*Well-foundedness of image*}
227 lemma wf_prod_fun_image: "[| wf r; inj f |] ==> wf(prod_fun f f ` r)"
228 apply (simp only: wf_eq_minimal, clarify)
229 apply (case_tac "EX p. f p : Q")
230 apply (erule_tac x = "{p. f p : Q}" in allE)
231 apply (fast dest: inj_onD, blast)
235 subsection {* Well-Foundedness Results for Unions *}
237 lemma wf_union_compatible:
238 assumes "wf R" "wf S"
239 assumes "S O R \<subseteq> R"
240 shows "wf (R \<union> S)"
243 let ?Q' = "{x \<in> Q. \<forall>y. (y, x) \<in> R \<longrightarrow> y \<notin> Q}"
245 obtain a where "a \<in> ?Q'"
246 by (rule wfE_min [OF `wf R` `x \<in> Q`]) blast
248 obtain z where "z \<in> ?Q'" and zmin: "\<And>y. (y, z) \<in> S \<Longrightarrow> y \<notin> ?Q'" by (erule wfE_min)
250 fix y assume "(y, z) \<in> S"
251 then have "y \<notin> ?Q'" by (rule zmin)
256 with `y \<notin> ?Q'`
257 obtain w where "(w, y) \<in> R" and "w \<in> Q" by auto
258 from `(w, y) \<in> R` `(y, z) \<in> S` have "(w, z) \<in> S O R" by (rule rel_compI)
259 with `S O R \<subseteq> R` have "(w, z) \<in> R" ..
260 with `z \<in> ?Q'` have "w \<notin> Q" by blast
261 with `w \<in> Q` show False by contradiction
264 with `z \<in> ?Q'` show "\<exists>z\<in>Q. \<forall>y. (y, z) \<in> R \<union> S \<longrightarrow> y \<notin> Q" by blast
268 text {* Well-foundedness of indexed union with disjoint domains and ranges *}
270 lemma wf_UN: "[| ALL i:I. wf(r i);
271 ALL i:I. ALL j:I. r i ~= r j --> Domain(r i) Int Range(r j) = {}
272 |] ==> wf(UN i:I. r i)"
273 apply (simp only: wf_eq_minimal, clarify)
274 apply (rename_tac A a, case_tac "EX i:I. EX a:A. EX b:A. (b,a) : r i")
278 apply (drule bspec, assumption)
279 apply (erule_tac x="{a. a:A & (EX b:A. (b,a) : r i) }" in allE)
280 apply (blast elim!: allE)
283 lemmas wfP_SUP = wf_UN [where I=UNIV and r="\<lambda>i. {(x, y). r i x y}",
284 to_pred SUP_UN_eq2 bot_empty_eq pred_equals_eq, simplified, standard]
288 ALL r:R. ALL s:R. r ~= s --> Domain r Int Range s = {}
290 apply (simp add: Union_def)
291 apply (blast intro: wf_UN)
294 (*Intuition: we find an (R u S)-min element of a nonempty subset A
296 1. There is a step a -R-> b with a,b : A.
297 Pick an R-min element z of the (nonempty) set {a:A | EX b:A. a -R-> b}.
298 By definition, there is z':A s.t. z -R-> z'. Because z is R-min in the
299 subset, z' must be R-min in A. Because z' has an R-predecessor, it cannot
300 have an S-successor and is thus S-min in A as well.
301 2. There is no such step.
302 Pick an S-min element of A. In this case it must be an R-min
303 element of A as well.
307 "[| wf r; wf s; Domain r Int Range s = {} |] ==> wf(r Un s)"
308 using wf_union_compatible[of s r]
309 by (auto simp: Un_ac)
311 lemma wf_union_merge:
312 "wf (R \<union> S) = wf (R O R \<union> R O S \<union> S)" (is "wf ?A = wf ?B")
315 with wf_trancl have wfT: "wf (?A^+)" .
316 moreover have "?B \<subseteq> ?A^+"
317 by (subst trancl_unfold, subst trancl_unfold) blast
318 ultimately show "wf ?B" by (rule wf_subset)
324 fix Q :: "'a set" and x
328 obtain z where "z \<in> Q" and "\<And>y. (y, z) \<in> ?B \<Longrightarrow> y \<notin> Q"
330 then have A1: "\<And>y. (y, z) \<in> R O R \<Longrightarrow> y \<notin> Q"
331 and A2: "\<And>y. (y, z) \<in> R O S \<Longrightarrow> y \<notin> Q"
332 and A3: "\<And>y. (y, z) \<in> S \<Longrightarrow> y \<notin> Q"
335 show "\<exists>z\<in>Q. \<forall>y. (y, z) \<in> ?A \<longrightarrow> y \<notin> Q"
336 proof (cases "\<forall>y. (y, z) \<in> R \<longrightarrow> y \<notin> Q")
338 with `z \<in> Q` A3 show ?thesis by blast
341 then obtain z' where "z'\<in>Q" "(z', z) \<in> R" by blast
343 have "\<forall>y. (y, z') \<in> ?A \<longrightarrow> y \<notin> Q"
344 proof (intro allI impI)
345 fix y assume "(y, z') \<in> ?A"
346 then show "y \<notin> Q"
348 assume "(y, z') \<in> R"
349 then have "(y, z) \<in> R O R" using `(z', z) \<in> R` ..
350 with A1 show "y \<notin> Q" .
352 assume "(y, z') \<in> S"
353 then have "(y, z) \<in> R O S" using `(z', z) \<in> R` ..
354 with A2 show "y \<notin> Q" .
357 with `z' \<in> Q` show ?thesis ..
362 lemma wf_comp_self: "wf R = wf (R O R)" -- {* special case *}
363 by (rule wf_union_merge [where S = "{}", simplified])
366 subsubsection {* acyclic *}
368 lemma acyclicI: "ALL x. (x, x) ~: r^+ ==> acyclic r"
369 by (simp add: acyclic_def)
371 lemma wf_acyclic: "wf r ==> acyclic r"
372 apply (simp add: acyclic_def)
373 apply (blast elim: wf_trancl [THEN wf_irrefl])
376 lemmas wfP_acyclicP = wf_acyclic [to_pred]
378 lemma acyclic_insert [iff]:
379 "acyclic(insert (y,x) r) = (acyclic r & (x,y) ~: r^*)"
380 apply (simp add: acyclic_def trancl_insert)
381 apply (blast intro: rtrancl_trans)
384 lemma acyclic_converse [iff]: "acyclic(r^-1) = acyclic r"
385 by (simp add: acyclic_def trancl_converse)
387 lemmas acyclicP_converse [iff] = acyclic_converse [to_pred]
389 lemma acyclic_impl_antisym_rtrancl: "acyclic r ==> antisym(r^*)"
390 apply (simp add: acyclic_def antisym_def)
391 apply (blast elim: rtranclE intro: rtrancl_into_trancl1 rtrancl_trancl_trancl)
396 antisym = only self loops
397 Goalw [acyclic_def,antisym_def] "antisym( r^* ) ==> acyclic(r - Id)
398 ==> antisym( r^* ) = acyclic(r - Id)";
401 lemma acyclic_subset: "[| acyclic s; r <= s |] ==> acyclic r"
402 apply (simp add: acyclic_def)
403 apply (blast intro: trancl_mono)
406 text{* Wellfoundedness of finite acyclic relations*}
408 lemma finite_acyclic_wf [rule_format]: "finite r ==> acyclic r --> wf r"
409 apply (erule finite_induct, blast)
410 apply (simp (no_asm_simp) only: split_tupled_all)
414 lemma finite_acyclic_wf_converse: "[|finite r; acyclic r|] ==> wf (r^-1)"
415 apply (erule finite_converse [THEN iffD2, THEN finite_acyclic_wf])
416 apply (erule acyclic_converse [THEN iffD2])
419 lemma wf_iff_acyclic_if_finite: "finite r ==> wf r = acyclic r"
420 by (blast intro: finite_acyclic_wf wf_acyclic)
423 subsection{*Well-Founded Recursion*}
427 lemma cuts_eq: "(cut f r x = cut g r x) = (ALL y. (y,x):r --> f(y)=g(y))"
428 by (simp add: expand_fun_eq cut_def)
430 lemma cut_apply: "(x,a):r ==> (cut f r a)(x) = f(x)"
431 by (simp add: cut_def)
433 text{*Inductive characterization of wfrec combinator; for details see:
434 John Harrison, "Inductive definitions: automation and application"*}
436 lemma wfrec_unique: "[| adm_wf R F; wf R |] ==> EX! y. wfrec_rel R F x y"
437 apply (simp add: adm_wf_def)
438 apply (erule_tac a=x in wf_induct)
440 apply (rule_tac g = "%x. THE y. wfrec_rel R F x y" in wfrec_rel.wfrecI)
441 apply (fast dest!: theI')
442 apply (erule wfrec_rel.cases, simp)
443 apply (erule allE, erule allE, erule allE, erule mp)
444 apply (fast intro: the_equality [symmetric])
447 lemma adm_lemma: "adm_wf R (%f x. F (cut f R x) x)"
448 apply (simp add: adm_wf_def)
450 apply (rule cuts_eq [THEN iffD2, THEN subst], assumption)
454 lemma wfrec: "wf(r) ==> wfrec r H a = H (cut (wfrec r H) r a) a"
455 apply (simp add: wfrec_def)
456 apply (rule adm_lemma [THEN wfrec_unique, THEN the1_equality], assumption)
457 apply (rule wfrec_rel.wfrecI)
459 apply (erule adm_lemma [THEN wfrec_unique, THEN theI'])
462 subsection {* Code generator setup *}
465 "wfrec" ("\<module>wfrec?")
467 fun wfrec f x = f (wfrec f) x;
471 subsection {* @{typ nat} is well-founded *}
473 lemma less_nat_rel: "op < = (\<lambda>m n. n = Suc m)^++"
474 proof (rule ext, rule ext, rule iffI)
477 then show "(\<lambda>m n. n = Suc m)^++ m n"
479 case 0 then show ?case by auto
481 case (Suc n) then show ?case
482 by (auto simp add: less_Suc_eq_le le_less intro: tranclp.trancl_into_trancl)
486 assume "(\<lambda>m n. n = Suc m)^++ m n"
489 (simp_all add: less_Suc_eq_le reflexive le_less)
493 pred_nat :: "(nat * nat) set" where
494 "pred_nat = {(m, n). n = Suc m}"
497 less_than :: "(nat * nat) set" where
498 "less_than = pred_nat^+"
500 lemma less_eq: "(m, n) \<in> pred_nat^+ \<longleftrightarrow> m < n"
501 unfolding less_nat_rel pred_nat_def trancl_def by simp
503 lemma pred_nat_trancl_eq_le:
504 "(m, n) \<in> pred_nat^* \<longleftrightarrow> m \<le> n"
505 unfolding less_eq rtrancl_eq_or_trancl by auto
507 lemma wf_pred_nat: "wf pred_nat"
508 apply (unfold wf_def pred_nat_def, clarify)
509 apply (induct_tac x, blast+)
512 lemma wf_less_than [iff]: "wf less_than"
513 by (simp add: less_than_def wf_pred_nat [THEN wf_trancl])
515 lemma trans_less_than [iff]: "trans less_than"
516 by (simp add: less_than_def trans_trancl)
518 lemma less_than_iff [iff]: "((x,y): less_than) = (x<y)"
519 by (simp add: less_than_def less_eq)
521 lemma wf_less: "wf {(x, y::nat). x < y}"
522 using wf_less_than by (simp add: less_than_def less_eq [symmetric])
525 subsection {* Accessible Part *}
528 Inductive definition of the accessible part @{term "acc r"} of a
529 relation; see also \cite{paulin-tlca}.
533 acc :: "('a * 'a) set => 'a set"
534 for r :: "('a * 'a) set"
536 accI: "(!!y. (y, x) : r ==> y : acc r) ==> x : acc r"
539 termip :: "('a => 'a => bool) => 'a => bool" where
540 "termip r == accp (r\<inverse>\<inverse>)"
543 termi :: "('a * 'a) set => 'a set" where
544 "termi r == acc (r\<inverse>)"
546 lemmas accpI = accp.accI
548 text {* Induction rules *}
551 assumes major: "accp r a"
552 assumes hyp: "!!x. accp r x ==> \<forall>y. r y x --> P y ==> P x"
554 apply (rule major [THEN accp.induct])
556 apply (rule accp.accI)
561 theorems accp_induct_rule = accp_induct [rule_format, induct set: accp]
563 theorem accp_downward: "accp r b ==> r a b ==> accp r a"
564 apply (erule accp.cases)
569 assumes na: "\<not> accp R x"
570 obtains z where "R z x" and "\<not> accp R z"
572 assume a: "\<And>z. \<lbrakk>R z x; \<not> accp R z\<rbrakk> \<Longrightarrow> thesis"
575 proof (cases "\<forall>z. R z x \<longrightarrow> accp R z")
577 hence "\<And>z. R z x \<Longrightarrow> accp R z" by auto
580 with na show thesis ..
582 case False then obtain z where "R z x" and "\<not> accp R z"
588 lemma accp_downwards_aux: "r\<^sup>*\<^sup>* b a ==> accp r a --> accp r b"
589 apply (erule rtranclp_induct)
591 apply (blast dest: accp_downward)
594 theorem accp_downwards: "accp r a ==> r\<^sup>*\<^sup>* b a ==> accp r b"
595 apply (blast dest: accp_downwards_aux)
598 theorem accp_wfPI: "\<forall>x. accp r x ==> wfP r"
599 apply (rule wfPUNIVI)
600 apply (induct_tac P x rule: accp_induct)
605 theorem accp_wfPD: "wfP r ==> accp r x"
606 apply (erule wfP_induct_rule)
607 apply (rule accp.accI)
611 theorem wfP_accp_iff: "wfP r = (\<forall>x. accp r x)"
612 apply (blast intro: accp_wfPI dest: accp_wfPD)
616 text {* Smaller relations have bigger accessible parts: *}
619 assumes sub: "R1 \<le> R2"
620 shows "accp R2 \<le> accp R1"
621 proof (rule predicate1I)
622 fix x assume "accp R2 x"
623 then show "accp R1 x"
626 assume ih: "\<And>y. R2 y x \<Longrightarrow> accp R1 y"
627 with sub show "accp R1 x"
628 by (blast intro: accp.accI)
633 text {* This is a generalized induction theorem that works on
634 subsets of the accessible part. *}
636 lemma accp_subset_induct:
637 assumes subset: "D \<le> accp R"
638 and dcl: "\<And>x z. \<lbrakk>D x; R z x\<rbrakk> \<Longrightarrow> D z"
640 and istep: "\<And>x. \<lbrakk>D x; (\<And>z. R z x \<Longrightarrow> P z)\<rbrakk> \<Longrightarrow> P x"
643 from subset and `D x`
645 then show "P x" using `D x`
649 and "\<And>y. R y x \<Longrightarrow> D y \<Longrightarrow> P y"
650 with dcl and istep show "P x" by blast
655 text {* Set versions of the above theorems *}
657 lemmas acc_induct = accp_induct [to_set]
659 lemmas acc_induct_rule = acc_induct [rule_format, induct set: acc]
661 lemmas acc_downward = accp_downward [to_set]
663 lemmas not_acc_down = not_accp_down [to_set]
665 lemmas acc_downwards_aux = accp_downwards_aux [to_set]
667 lemmas acc_downwards = accp_downwards [to_set]
669 lemmas acc_wfI = accp_wfPI [to_set]
671 lemmas acc_wfD = accp_wfPD [to_set]
673 lemmas wf_acc_iff = wfP_accp_iff [to_set]
675 lemmas acc_subset = accp_subset [to_set pred_subset_eq]
677 lemmas acc_subset_induct = accp_subset_induct [to_set pred_subset_eq]
680 subsection {* Tools for building wellfounded relations *}
682 text {* Inverse Image *}
684 lemma wf_inv_image [simp,intro!]: "wf(r) ==> wf(inv_image r (f::'a=>'b))"
685 apply (simp (no_asm_use) add: inv_image_def wf_eq_minimal)
687 apply (subgoal_tac "EX (w::'b) . w : {w. EX (x::'a) . x: Q & (f x = w) }")
688 prefer 2 apply (blast del: allE)
690 apply (erule (1) notE impE)
694 lemma in_inv_image[simp]: "((x,y) : inv_image r f) = ((f x, f y) : r)"
695 by (auto simp:inv_image_def)
697 text {* Measure functions into @{typ nat} *}
699 definition measure :: "('a => nat) => ('a * 'a)set"
700 where "measure == inv_image less_than"
702 lemma in_measure[simp]: "((x,y) : measure f) = (f x < f y)"
703 by (simp add:measure_def)
705 lemma wf_measure [iff]: "wf (measure f)"
706 apply (unfold measure_def)
707 apply (rule wf_less_than [THEN wf_inv_image])
710 text{* Lexicographic combinations *}
713 lex_prod :: "[('a*'a)set, ('b*'b)set] => (('a*'b)*('a*'b))set"
714 (infixr "<*lex*>" 80)
716 "ra <*lex*> rb == {((a,b),(a',b')). (a,a') : ra | a=a' & (b,b') : rb}"
718 lemma wf_lex_prod [intro!]: "[| wf(ra); wf(rb) |] ==> wf(ra <*lex*> rb)"
719 apply (unfold wf_def lex_prod_def)
720 apply (rule allI, rule impI)
721 apply (simp (no_asm_use) only: split_paired_All)
722 apply (drule spec, erule mp)
723 apply (rule allI, rule impI)
724 apply (drule spec, erule mp, blast)
727 lemma in_lex_prod[simp]:
728 "(((a,b),(a',b')): r <*lex*> s) = ((a,a'): r \<or> (a = a' \<and> (b, b') : s))"
729 by (auto simp:lex_prod_def)
731 text{* @{term "op <*lex*>"} preserves transitivity *}
733 lemma trans_lex_prod [intro!]:
734 "[| trans R1; trans R2 |] ==> trans (R1 <*lex*> R2)"
735 by (unfold trans_def lex_prod_def, blast)
737 text {* lexicographic combinations with measure functions *}
740 mlex_prod :: "('a \<Rightarrow> nat) \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set" (infixr "<*mlex*>" 80)
742 "f <*mlex*> R = inv_image (less_than <*lex*> R) (%x. (f x, x))"
744 lemma wf_mlex: "wf R \<Longrightarrow> wf (f <*mlex*> R)"
745 unfolding mlex_prod_def
748 lemma mlex_less: "f x < f y \<Longrightarrow> (x, y) \<in> f <*mlex*> R"
749 unfolding mlex_prod_def by simp
751 lemma mlex_leq: "f x \<le> f y \<Longrightarrow> (x, y) \<in> R \<Longrightarrow> (x, y) \<in> f <*mlex*> R"
752 unfolding mlex_prod_def by auto
754 text {* proper subset relation on finite sets *}
756 definition finite_psubset :: "('a set * 'a set) set"
757 where "finite_psubset == {(A,B). A < B & finite B}"
759 lemma wf_finite_psubset: "wf(finite_psubset)"
760 apply (unfold finite_psubset_def)
761 apply (rule wf_measure [THEN wf_subset])
762 apply (simp add: measure_def inv_image_def less_than_def less_eq)
763 apply (fast elim!: psubset_card_mono)
766 lemma trans_finite_psubset: "trans finite_psubset"
767 by (simp add: finite_psubset_def less_le trans_def, blast)
772 text {*Wellfoundedness of @{text same_fst}*}
775 same_fst :: "('a => bool) => ('a => ('b * 'b)set) => (('a*'b)*('a*'b))set"
777 "same_fst P R == {((x',y'),(x,y)) . x'=x & P x & (y',y) : R x}"
778 --{*For @{text rec_def} declarations where the first n parameters
779 stay unchanged in the recursive call.
780 See @{text "Library/While_Combinator.thy"} for an application.*}
782 lemma same_fstI [intro!]:
783 "[| P x; (y',y) : R x |] ==> ((x,y'),(x,y)) : same_fst P R"
784 by (simp add: same_fst_def)
787 assumes prem: "(!!x. P x ==> wf(R x))"
788 shows "wf(same_fst P R)"
789 apply (simp cong del: imp_cong add: wf_def same_fst_def)
791 apply (rename_tac a b)
792 apply (case_tac "wf (R a)")
793 apply (erule_tac a = b in wf_induct, blast)
794 apply (blast intro: prem)
798 subsection{*Weakly decreasing sequences (w.r.t. some well-founded order)
801 text{*This material does not appear to be used any longer.*}
803 lemma lemma1: "[| ALL i. (f (Suc i), f i) : r^* |] ==> (f (i+k), f i) : r^*"
804 apply (induct_tac "k", simp_all)
805 apply (blast intro: rtrancl_trans)
808 lemma lemma2: "[| ALL i. (f (Suc i), f i) : r^*; wf (r^+) |]
809 ==> ALL m. f m = x --> (EX i. ALL k. f (m+i+k) = f (m+i))"
810 apply (erule wf_induct, clarify)
811 apply (case_tac "EX j. (f (m+j), f m) : r^+")
813 apply (subgoal_tac "EX i. ALL k. f ((m+j) +i+k) = f ( (m+j) +i) ")
815 apply (rule_tac x = "j+i" in exI)
816 apply (simp add: add_ac, blast)
817 apply (rule_tac x = 0 in exI, clarsimp)
818 apply (drule_tac i = m and k = k in lemma1)
819 apply (blast elim: rtranclE dest: rtrancl_into_trancl1)
822 lemma wf_weak_decr_stable: "[| ALL i. (f (Suc i), f i) : r^*; wf (r^+) |]
823 ==> EX i. ALL k. f (i+k) = f i"
824 apply (drule_tac x = 0 in lemma2 [THEN spec], auto)
827 (* special case of the theorem above: <= *)
828 lemma weak_decr_stable:
829 "ALL i. f (Suc i) <= ((f i)::nat) ==> EX i. ALL k. f (i+k) = f i"
830 apply (rule_tac r = pred_nat in wf_weak_decr_stable)
831 apply (simp add: pred_nat_trancl_eq_le)
832 apply (intro wf_trancl wf_pred_nat)
836 subsection {* size of a datatype value *}
838 use "Tools/function_package/size.ML"
842 lemma size_bool [code func]:
843 "size (b\<Colon>bool) = 0" by (cases b) auto
845 lemma nat_size [simp, code func]: "size (n\<Colon>nat) = n"
846 by (induct n) simp_all
848 declare "prod.size" [noatp]