# HG changeset patch # User wenzelm # Date 1463998696 -7200 # Node ID 02b8855917358ff9dda70cc8896f0c6ceda999af # Parent 932cf444f2fe6d661d3cae3b5d7dfd342f2c88ca misc tuning and modernization; diff -r 932cf444f2fe -r 02b885591735 src/HOL/Wellfounded.thy --- a/src/HOL/Wellfounded.thy Sat May 21 07:08:59 2016 +0200 +++ b/src/HOL/Wellfounded.thy Mon May 23 12:18:16 2016 +0200 @@ -14,34 +14,33 @@ subsection \Basic Definitions\ -definition wf :: "('a * 'a) set => bool" where - "wf r \ (!P. (!x. (!y. (y,x):r --> P(y)) --> P(x)) --> (!x. P(x)))" +definition wf :: "('a \ 'a) set \ bool" + where "wf r \ (\P. (\x. (\y. (y, x) \ r \ P y) \ P x) \ (\x. P x))" -definition wfP :: "('a => 'a => bool) => bool" where - "wfP r \ wf {(x, y). r x y}" +definition wfP :: "('a \ 'a \ bool) \ bool" + where "wfP r \ wf {(x, y). r x y}" lemma wfP_wf_eq [pred_set_conv]: "wfP (\x y. (x, y) \ r) = wf r" by (simp add: wfP_def) -lemma wfUNIVI: - "(!!P x. (ALL x. (ALL y. (y,x) : r --> P(y)) --> P(x)) ==> P(x)) ==> wf(r)" +lemma wfUNIVI: "(\P x. (\x. (\y. (y, x) \ r \ P y) \ P x) \ P x) \ wf r" unfolding wf_def by blast lemmas wfPUNIVI = wfUNIVI [to_pred] -text\Restriction to domain @{term A} and range @{term B}. If @{term r} is - well-founded over their intersection, then @{term "wf r"}\ -lemma wfI: - "[| r \ A \ B; - !!x P. [|\x. (\y. (y,x) : r --> P y) --> P x; x : A; x : B |] ==> P x |] - ==> wf r" - unfolding wf_def by blast +text \Restriction to domain \A\ and range \B\. + If \r\ is well-founded over their intersection, then \wf r\.\ +lemma wfI: + assumes "r \ A \ B" + and "\x P. \\x. (\y. (y, x) \ r \ P y) \ P x; x \ A; x \ B\ \ P x" + shows "wf r" + using assms unfolding wf_def by blast -lemma wf_induct: - "[| wf(r); - !!x.[| ALL y. (y,x): r --> P(y) |] ==> P(x) - |] ==> P(a)" - unfolding wf_def by blast +lemma wf_induct: + assumes "wf r" + and "\x. \y. (y, x) \ r \ P y \ P x" + shows "P a" + using assms unfolding wf_def by blast lemmas wfP_induct = wf_induct [to_pred] @@ -49,7 +48,7 @@ lemmas wfP_induct_rule = wf_induct_rule [to_pred, induct set: wfP] -lemma wf_not_sym: "wf r ==> (a, x) : r ==> (x, a) ~: r" +lemma wf_not_sym: "wf r \ (a, x) \ r \ (x, a) \ r" by (induct a arbitrary: x set: wf) blast lemma wf_asym: @@ -57,22 +56,25 @@ obtains "(x, a) \ r" by (drule wf_not_sym[OF assms]) -lemma wf_not_refl [simp]: "wf r ==> (a, a) ~: r" +lemma wf_not_refl [simp]: "wf r \ (a, a) \ r" by (blast elim: wf_asym) lemma wf_irrefl: assumes "wf r" obtains "(a, a) \ r" -by (drule wf_not_refl[OF assms]) + by (drule wf_not_refl[OF assms]) lemma wf_wellorderI: assumes wf: "wf {(x::'a::ord, y). x < y}" assumes lin: "OFCLASS('a::ord, linorder_class)" shows "OFCLASS('a::ord, wellorder_class)" -using lin by (rule wellorder_class.intro) - (rule class.wellorder_axioms.intro, rule wf_induct_rule [OF wf], simp) + using lin + apply (rule wellorder_class.intro) + apply (rule class.wellorder_axioms.intro) + apply (rule wf_induct_rule [OF wf]) + apply simp + done -lemma (in wellorder) wf: - "wf {(x, y). x < y}" -unfolding wf_def by (blast intro: less_induct) +lemma (in wellorder) wf: "wf {(x, y). x < y}" + unfolding wf_def by (blast intro: less_induct) subsection \Basic Results\ @@ -84,14 +86,13 @@ assumes a: "A \ R `` A" shows "A = {}" proof - - { fix x - from wf have "x \ A" - proof induct - fix x assume "\y. (y, x) \ R \ y \ A" - then have "x \ R `` A" by blast - with a show "x \ A" by blast - qed - } thus ?thesis by auto + from wf have "x \ A" for x + proof induct + fix x assume "\y. (y, x) \ R \ y \ A" + then have "x \ R `` A" by blast + with a show "x \ A" by blast + qed + then show ?thesis by auto qed lemma wfI_pf: @@ -105,7 +106,8 @@ with a show "P x" by blast qed -text\Minimal-element characterization of well-foundedness\ + +subsubsection \Minimal-element characterization of well-foundedness\ lemma wfE_min: assumes wf: "wf R" and Q: "x \ Q" @@ -120,14 +122,14 @@ assumes a: "\x Q. x \ Q \ \z\Q. \y. (y, z) \ R \ y \ Q" shows "wf R" proof (rule wfI_pf) - fix A assume b: "A \ R `` A" - { fix x assume "x \ A" - from a[OF this] b have "False" by blast - } - thus "A = {}" by blast + fix A + assume b: "A \ R `` A" + have False if "x \ A" for x + using a[OF that] b by blast + then show "A = {}" by blast qed -lemma wf_eq_minimal: "wf r = (\Q x. x\Q --> (\z\Q. \y. (y,z)\r --> y\Q))" +lemma wf_eq_minimal: "wf r \ (\Q x. x \ Q \ (\z\Q. \y. (y, z) \ r \ y \ Q))" apply auto apply (erule wfE_min, assumption, blast) apply (rule wfI_min, auto) @@ -135,51 +137,52 @@ lemmas wfP_eq_minimal = wf_eq_minimal [to_pred] -text\Well-foundedness of transitive closure\ + +subsubsection \Well-foundedness of transitive closure\ lemma wf_trancl: assumes "wf r" - shows "wf (r^+)" + shows "wf (r\<^sup>+)" proof - - { - fix P and x - assume induct_step: "!!x. (!!y. (y, x) : r^+ ==> P y) ==> P x" - have "P x" - proof (rule induct_step) - fix y assume "(y, x) : r^+" - with \wf r\ show "P y" - proof (induct x arbitrary: y) - case (less x) - note hyp = \\x' y'. (x', x) : r ==> (y', x') : r^+ ==> P y'\ - from \(y, x) : r^+\ show "P y" - proof cases - case base - show "P y" - proof (rule induct_step) - fix y' assume "(y', y) : r^+" - with \(y, x) : r\ show "P y'" by (rule hyp [of y y']) - qed - next - case step - then obtain x' where "(x', x) : r" and "(y, x') : r^+" by simp - then show "P y" by (rule hyp [of x' y]) + have "P x" if induct_step: "\x. (\y. (y, x) \ r\<^sup>+ \ P y) \ P x" for P x + proof (rule induct_step) + show "P y" if "(y, x) \ r\<^sup>+" for y + using \wf r\ and that + proof (induct x arbitrary: y) + case (less x) + note hyp = \\x' y'. (x', x) \ r \ (y', x') \ r\<^sup>+ \ P y'\ + from \(y, x) \ r\<^sup>+\ show "P y" + proof cases + case base + show "P y" + proof (rule induct_step) + fix y' + assume "(y', y) \ r\<^sup>+" + with \(y, x) \ r\ show "P y'" + by (rule hyp [of y y']) qed + next + case step + then obtain x' where "(x', x) \ r" and "(y, x') \ r\<^sup>+" + by simp + then show "P y" by (rule hyp [of x' y]) qed qed - } then show ?thesis unfolding wf_def by blast + qed + then show ?thesis unfolding wf_def by blast qed lemmas wfP_trancl = wf_trancl [to_pred] -lemma wf_converse_trancl: "wf (r^-1) ==> wf ((r^+)^-1)" +lemma wf_converse_trancl: "wf (r\) \ wf ((r\<^sup>+)\)" apply (subst trancl_converse [symmetric]) apply (erule wf_trancl) done text \Well-foundedness of subsets\ -lemma wf_subset: "[| wf(r); p<=r |] ==> wf(p)" - apply (simp (no_asm_use) add: wf_eq_minimal) +lemma wf_subset: "wf r \ p \ r \ wf p" + apply (simp add: wf_eq_minimal) apply fast done @@ -197,15 +200,15 @@ then show ?thesis by (simp add: bot_fun_def) qed -lemma wf_Int1: "wf r ==> wf (r Int r')" +lemma wf_Int1: "wf r \ wf (r Int r')" apply (erule wf_subset) apply (rule Int_lower1) done -lemma wf_Int2: "wf r ==> wf (r' Int r)" +lemma wf_Int2: "wf r \ wf (r' Int r)" apply (erule wf_subset) apply (rule Int_lower2) - done + done text \Exponentiation\ @@ -221,33 +224,34 @@ text \Well-foundedness of insert\ -lemma wf_insert [iff]: "wf(insert (y,x) r) = (wf(r) & (x,y) ~: r^*)" +lemma wf_insert [iff]: "wf (insert (y, x) r) \ wf r \ (x, y) \ r\<^sup>*" apply (rule iffI) apply (blast elim: wf_trancl [THEN wf_irrefl] - intro: rtrancl_into_trancl1 wf_subset + intro: rtrancl_into_trancl1 wf_subset rtrancl_mono [THEN [2] rev_subsetD]) apply (simp add: wf_eq_minimal, safe) -apply (rule allE, assumption, erule impE, blast) +apply (rule allE, assumption, erule impE, blast) apply (erule bexE) apply (rename_tac "a", case_tac "a = x") prefer 2 -apply blast -apply (case_tac "y:Q") +apply blast +apply (case_tac "y \ Q") prefer 2 apply blast -apply (rule_tac x = "{z. z:Q & (z,y) : r^*}" in allE) +apply (rule_tac x = "{z. z \ Q \ (z,y) \ r\<^sup>*}" in allE) apply assumption -apply (erule_tac V = "ALL Q. (EX x. x : Q) --> P Q" for P in thin_rl) - \\essential for speed\ -txt\Blast with new substOccur fails\ +apply (erule_tac V = "\Q. (\x. x \ Q) \ P Q" for P in thin_rl) + (*essential for speed*) +(*blast with new substOccur fails*) apply (fast intro: converse_rtrancl_into_rtrancl) done -text\Well-foundedness of image\ + +subsubsection \Well-foundedness of image\ -lemma wf_map_prod_image: "[| wf r; inj f |] ==> wf (map_prod f f ` r)" +lemma wf_map_prod_image: "wf r \ inj f \ wf (map_prod f f ` r)" apply (simp only: wf_eq_minimal, clarify) -apply (case_tac "EX p. f p : Q") -apply (erule_tac x = "{p. f p : Q}" in allE) +apply (case_tac "\p. f p \ Q") +apply (erule_tac x = "{p. f p \ Q}" in allE) apply (fast dest: inj_onD, blast) done @@ -259,25 +263,23 @@ assumes "R O S \ R" shows "wf (R \ S)" proof (rule wfI_min) - fix x :: 'a and Q + fix x :: 'a and Q let ?Q' = "{x \ Q. \y. (y, x) \ R \ y \ Q}" assume "x \ Q" obtain a where "a \ ?Q'" by (rule wfE_min [OF \wf R\ \x \ Q\]) blast - with \wf S\ - obtain z where "z \ ?Q'" and zmin: "\y. (y, z) \ S \ y \ ?Q'" by (erule wfE_min) - { + with \wf S\ obtain z where "z \ ?Q'" and zmin: "\y. (y, z) \ S \ y \ ?Q'" + by (erule wfE_min) + { fix y assume "(y, z) \ S" then have "y \ ?Q'" by (rule zmin) - have "y \ Q" - proof + proof assume "y \ Q" - with \y \ ?Q'\ - obtain w where "(w, y) \ R" and "w \ Q" by auto + with \y \ ?Q'\ obtain w where "(w, y) \ R" and "w \ Q" by auto from \(w, y) \ R\ \(y, z) \ S\ have "(w, z) \ R O S" by (rule relcompI) with \R O S \ R\ have "(w, z) \ R" .. - with \z \ ?Q'\ have "w \ Q" by blast + with \z \ ?Q'\ have "w \ Q" by blast with \w \ Q\ show False by contradiction qed } @@ -287,18 +289,21 @@ text \Well-foundedness of indexed union with disjoint domains and ranges\ -lemma wf_UN: "[| ALL i:I. wf(r i); - ALL i:I. ALL j:I. r i ~= r j --> Domain(r i) Int Range(r j) = {} - |] ==> wf(UN i:I. r i)" -apply (simp only: wf_eq_minimal, clarify) -apply (rename_tac A a, case_tac "EX i:I. EX a:A. EX b:A. (b,a) : r i") - prefer 2 - apply force -apply clarify -apply (drule bspec, assumption) -apply (erule_tac x="{a. a:A & (EX b:A. (b,a) : r i) }" in allE) -apply (blast elim!: allE) -done +lemma wf_UN: + assumes "\i\I. wf (r i)" + and "\i\I. \j\I. r i \ r j \ Domain (r i) \ Range (r j) = {}" + shows "wf (\i\I. r i)" + using assms + apply (simp only: wf_eq_minimal) + apply clarify + apply (rename_tac A a, case_tac "\i\I. \a\A. \b\A. (b, a) \ r i") + prefer 2 + apply force + apply clarify + apply (drule bspec, assumption) + apply (erule_tac x="{a. a \ A \ (\b\A. (b, a) \ r i) }" in allE) + apply (blast elim!: allE) + done lemma wfP_SUP: "\i. wfP (r i) \ \i j. r i \ r j \ inf (DomainP (r i)) (RangeP (r j)) = bot \ wfP (SUPREMUM UNIV r)" @@ -306,11 +311,11 @@ apply simp_all done -lemma wf_Union: - "[| ALL r:R. wf r; - ALL r:R. ALL s:R. r ~= s --> Domain r Int Range s = {} - |] ==> wf (\ R)" - using wf_UN[of R "\i. i"] by simp +lemma wf_Union: + assumes "\r\R. wf r" + and "\r\R. \s\R. r \ s \ Domain r \ Range s = {}" + shows "wf (\R)" + using assms wf_UN[of R "\i. i"] by simp (*Intuition: we find an (R u S)-min element of a nonempty subset A by case distinction. @@ -323,17 +328,16 @@ Pick an S-min element of A. In this case it must be an R-min element of A as well. *) -lemma wf_Un: - "[| wf r; wf s; Domain r Int Range s = {} |] ==> wf(r Un s)" - using wf_union_compatible[of s r] +lemma wf_Un: "wf r \ wf s \ Domain r \ Range s = {} \ wf (r \ s)" + using wf_union_compatible[of s r] by (auto simp: Un_ac) -lemma wf_union_merge: - "wf (R \ S) = wf (R O R \ S O R \ S)" (is "wf ?A = wf ?B") +lemma wf_union_merge: "wf (R \ S) = wf (R O R \ S O R \ S)" + (is "wf ?A = wf ?B") proof assume "wf ?A" - with wf_trancl have wfT: "wf (?A^+)" . - moreover have "?B \ ?A^+" + with wf_trancl have wfT: "wf (?A\<^sup>+)" . + moreover have "?B \ ?A\<^sup>+" by (subst trancl_unfold, subst trancl_unfold) blast ultimately show "wf ?B" by (rule wf_subset) next @@ -341,35 +345,34 @@ show "wf ?A" proof (rule wfI_min) - fix Q :: "'a set" and x + fix Q :: "'a set" and x assume "x \ Q" with \wf ?B\ - obtain z where "z \ Q" and "\y. (y, z) \ ?B \ y \ Q" + obtain z where "z \ Q" and "\y. (y, z) \ ?B \ y \ Q" by (erule wfE_min) then have A1: "\y. (y, z) \ R O R \ y \ Q" and A2: "\y. (y, z) \ S O R \ y \ Q" and A3: "\y. (y, z) \ S \ y \ Q" by auto - + show "\z\Q. \y. (y, z) \ ?A \ y \ Q" proof (cases "\y. (y, z) \ R \ y \ Q") case True with \z \ Q\ A3 show ?thesis by blast next - case False + case False then obtain z' where "z'\Q" "(z', z) \ R" by blast - have "\y. (y, z') \ ?A \ y \ Q" proof (intro allI impI) fix y assume "(y, z') \ ?A" then show "y \ Q" proof - assume "(y, z') \ R" + assume "(y, z') \ R" then have "(y, z) \ R O R" using \(z', z) \ R\ .. with A1 show "y \ Q" . next - assume "(y, z') \ S" + assume "(y, z') \ S" then have "(y, z) \ S O R" using \(z', z) \ R\ .. with A2 show "y \ Q" . qed @@ -443,22 +446,23 @@ subsection \Acyclic relations\ -lemma wf_acyclic: "wf r ==> acyclic r" +lemma wf_acyclic: "wf r \ acyclic r" apply (simp add: acyclic_def) apply (blast elim: wf_trancl [THEN wf_irrefl]) done lemmas wfP_acyclicP = wf_acyclic [to_pred] -text\Wellfoundedness of finite acyclic relations\ + +subsubsection \Wellfoundedness of finite acyclic relations\ -lemma finite_acyclic_wf [rule_format]: "finite r ==> acyclic r --> wf r" +lemma finite_acyclic_wf [rule_format]: "finite r \ acyclic r \ wf r" apply (erule finite_induct, blast) -apply (simp (no_asm_simp) only: split_tupled_all) +apply (simp only: split_tupled_all) apply simp done -lemma finite_acyclic_wf_converse: "[|finite r; acyclic r|] ==> wf (r^-1)" +lemma finite_acyclic_wf_converse: "finite r \ acyclic r \ wf (r\)" apply (erule finite_converse [THEN iffD2, THEN finite_acyclic_wf]) apply (erule acyclic_converse [THEN iffD2]) done @@ -477,44 +481,39 @@ with \finite r\ show ?thesis by (rule finite_acyclic_wf_converse) qed -lemma wf_iff_acyclic_if_finite: "finite r ==> wf r = acyclic r" +lemma wf_iff_acyclic_if_finite: "finite r \ wf r = acyclic r" by (blast intro: finite_acyclic_wf wf_acyclic) subsection \@{typ nat} is well-founded\ -lemma less_nat_rel: "op < = (\m n. n = Suc m)^++" +lemma less_nat_rel: "op < = (\m n. n = Suc m)\<^sup>+\<^sup>+" proof (rule ext, rule ext, rule iffI) fix n m :: nat - assume "m < n" - then show "(\m n. n = Suc m)^++ m n" + show "(\m n. n = Suc m)\<^sup>+\<^sup>+ m n" if "m < n" + using that proof (induct n) - case 0 then show ?case by auto + case 0 + then show ?case by auto next - case (Suc n) then show ?case + case (Suc n) + then show ?case by (auto simp add: less_Suc_eq_le le_less intro: tranclp.trancl_into_trancl) qed -next - fix n m :: nat - assume "(\m n. n = Suc m)^++ m n" - then show "m < n" - by (induct n) - (simp_all add: less_Suc_eq_le reflexive le_less) + show "m < n" if "(\m n. n = Suc m)\<^sup>+\<^sup>+ m n" + using that by (induct n) (simp_all add: less_Suc_eq_le reflexive le_less) qed -definition - pred_nat :: "(nat * nat) set" where - "pred_nat = {(m, n). n = Suc m}" +definition pred_nat :: "(nat \ nat) set" + where "pred_nat = {(m, n). n = Suc m}" -definition - less_than :: "(nat * nat) set" where - "less_than = pred_nat^+" +definition less_than :: "(nat \ nat) set" + where "less_than = pred_nat\<^sup>+" -lemma less_eq: "(m, n) \ pred_nat^+ \ m < n" +lemma less_eq: "(m, n) \ pred_nat\<^sup>+ \ m < n" unfolding less_nat_rel pred_nat_def trancl_def by simp -lemma pred_nat_trancl_eq_le: - "(m, n) \ pred_nat^* \ m \ n" +lemma pred_nat_trancl_eq_le: "(m, n) \ pred_nat\<^sup>* \ m \ n" unfolding less_eq rtrancl_eq_or_trancl by auto lemma wf_pred_nat: "wf pred_nat" @@ -528,7 +527,7 @@ lemma trans_less_than [iff]: "trans less_than" by (simp add: less_than_def) -lemma less_than_iff [iff]: "((x,y): less_than) = (x less_than) = (xAccessible Part\ text \ - Inductive definition of the accessible part @{term "acc r"} of a - relation; see also @{cite "paulin-tlca"}. + Inductive definition of the accessible part \acc r\ of a + relation; see also @{cite "paulin-tlca"}. \ -inductive_set - acc :: "('a * 'a) set => 'a set" - for r :: "('a * 'a) set" - where - accI: "(!!y. (y, x) : r ==> y : acc r) ==> x : acc r" +inductive_set acc :: "('a \ 'a) set \ 'a set" for r :: "('a \ 'a) set" + where accI: "(\y. (y, x) \ r \ y \ acc r) \ x \ acc r" -abbreviation - termip :: "('a => 'a => bool) => 'a => bool" where - "termip r \ accp (r\\)" +abbreviation termip :: "('a \ 'a \ bool) \ 'a \ bool" + where "termip r \ accp (r\\)" -abbreviation - termi :: "('a * 'a) set => 'a set" where - "termi r \ acc (r\)" +abbreviation termi :: "('a \ 'a) set \ 'a set" + where "termi r \ acc (r\)" lemmas accpI = accp.accI -lemma accp_eq_acc [code]: - "accp r = (\x. x \ Wellfounded.acc {(x, y). r x y})" +lemma accp_eq_acc [code]: "accp r = (\x. x \ Wellfounded.acc {(x, y). r x y})" by (simp add: acc_def) @@ -567,7 +560,7 @@ theorem accp_induct: assumes major: "accp r a" - assumes hyp: "!!x. accp r x ==> \y. r y x --> P y ==> P x" + assumes hyp: "\x. accp r x \ \y. r y x \ P y \ P x" shows "P a" apply (rule major [THEN accp.induct]) apply (rule hyp) @@ -578,7 +571,7 @@ lemmas accp_induct_rule = accp_induct [rule_format, induct set: accp] -theorem accp_downward: "accp r b ==> r a b ==> accp r a" +theorem accp_downward: "accp r b \ r a b \ accp r a" apply (erule accp.cases) apply fast done @@ -588,13 +581,11 @@ obtains z where "R z x" and "\ accp R z" proof - assume a: "\z. \R z x; \ accp R z\ \ thesis" - show thesis proof (cases "\z. R z x \ accp R z") case True - hence "\z. R z x \ accp R z" by auto - hence "accp R x" - by (rule accp.accI) + then have "\z. R z x \ accp R z" by auto + then have "accp R x" by (rule accp.accI) with na show thesis .. next case False then obtain z where "R z x" and "\ accp R z" @@ -603,24 +594,24 @@ qed qed -lemma accp_downwards_aux: "r\<^sup>*\<^sup>* b a ==> accp r a --> accp r b" +lemma accp_downwards_aux: "r\<^sup>*\<^sup>* b a \ accp r a \ accp r b" apply (erule rtranclp_induct) apply blast apply (blast dest: accp_downward) done -theorem accp_downwards: "accp r a ==> r\<^sup>*\<^sup>* b a ==> accp r b" +theorem accp_downwards: "accp r a \ r\<^sup>*\<^sup>* b a \ accp r b" apply (blast dest: accp_downwards_aux) done -theorem accp_wfPI: "\x. accp r x ==> wfP r" +theorem accp_wfPI: "\x. accp r x \ wfP r" apply (rule wfPUNIVI) apply (rule_tac P=P in accp_induct) apply blast apply blast done -theorem accp_wfPD: "wfP r ==> accp r x" +theorem accp_wfPD: "wfP r \ accp r x" apply (erule wfP_induct_rule) apply (rule accp.accI) apply blast @@ -699,10 +690,10 @@ text \Inverse Image\ -lemma wf_inv_image [simp,intro!]: "wf(r) ==> wf(inv_image r (f::'a=>'b))" +lemma wf_inv_image [simp,intro!]: "wf r \ wf (inv_image r f)" for f :: "'a \ 'b" apply (simp (no_asm_use) add: inv_image_def wf_eq_minimal) apply clarify -apply (subgoal_tac "EX (w::'b) . w : {w. EX (x::'a) . x: Q & (f x = w) }") +apply (subgoal_tac "\w::'b. w \ {w. \x::'a. x \ Q \ f x = w}") prefer 2 apply (blast del: allE) apply (erule allE) apply (erule (1) notE impE) @@ -711,10 +702,10 @@ text \Measure functions into @{typ nat}\ -definition measure :: "('a => nat) => ('a * 'a)set" -where "measure = inv_image less_than" +definition measure :: "('a \ nat) \ ('a \ 'a) set" + where "measure = inv_image less_than" -lemma in_measure[simp, code_unfold]: "((x,y) : measure f) = (f x < f y)" +lemma in_measure[simp, code_unfold]: "(x, y) \ measure f \ f x < f y" by (simp add:measure_def) lemma wf_measure [iff]: "wf (measure f)" @@ -722,8 +713,8 @@ apply (rule wf_less_than [THEN wf_inv_image]) done -lemma wf_if_measure: fixes f :: "'a \ nat" -shows "(!!x. P x \ f(g x) < f x) \ wf {(y,x). P x \ y = g x}" +lemma wf_if_measure: "(\x. P x \ f(g x) < f x) \ wf {(y,x). P x \ y = g x}" + for f :: "'a \ nat" apply(insert wf_measure[of f]) apply(simp only: measure_def inv_image_def less_than_def less_eq) apply(erule wf_subset) @@ -731,69 +722,66 @@ done -text\Lexicographic combinations\ +subsubsection \Lexicographic combinations\ -definition lex_prod :: "('a \'a) set \ ('b \ 'b) set \ (('a \ 'b) \ ('a \ 'b)) set" (infixr "<*lex*>" 80) where - "ra <*lex*> rb = {((a, b), (a', b')). (a, a') \ ra \ a = a' \ (b, b') \ rb}" +definition lex_prod :: "('a \'a) set \ ('b \ 'b) set \ (('a \ 'b) \ ('a \ 'b)) set" + (infixr "<*lex*>" 80) + where "ra <*lex*> rb = {((a, b), (a', b')). (a, a') \ ra \ a = a' \ (b, b') \ rb}" -lemma wf_lex_prod [intro!]: "[| wf(ra); wf(rb) |] ==> wf(ra <*lex*> rb)" -apply (unfold wf_def lex_prod_def) +lemma wf_lex_prod [intro!]: "wf ra \ wf rb \ wf (ra <*lex*> rb)" +apply (unfold wf_def lex_prod_def) apply (rule allI, rule impI) apply (simp (no_asm_use) only: split_paired_All) -apply (drule spec, erule mp) +apply (drule spec, erule mp) apply (rule allI, rule impI) -apply (drule spec, erule mp, blast) +apply (drule spec, erule mp, blast) done -lemma in_lex_prod[simp]: - "(((a,b),(a',b')): r <*lex*> s) = ((a,a'): r \ (a = a' \ (b, b') : s))" +lemma in_lex_prod[simp]: "((a, b), (a', b')) \ r <*lex*> s \ (a, a') \ r \ a = a' \ (b, b') \ s" by (auto simp:lex_prod_def) -text\@{term "op <*lex*>"} preserves transitivity\ +text \\<*lex*>\ preserves transitivity\ +lemma trans_lex_prod [intro!]: "trans R1 \ trans R2 \ trans (R1 <*lex*> R2)" + unfolding trans_def lex_prod_def by blast -lemma trans_lex_prod [intro!]: - "[| trans R1; trans R2 |] ==> trans (R1 <*lex*> R2)" -by (unfold trans_def lex_prod_def, blast) text \lexicographic combinations with measure functions\ -definition - mlex_prod :: "('a \ nat) \ ('a \ 'a) set \ ('a \ 'a) set" (infixr "<*mlex*>" 80) -where - "f <*mlex*> R = inv_image (less_than <*lex*> R) (%x. (f x, x))" +definition mlex_prod :: "('a \ nat) \ ('a \ 'a) set \ ('a \ 'a) set" (infixr "<*mlex*>" 80) + where "f <*mlex*> R = inv_image (less_than <*lex*> R) (\x. (f x, x))" lemma wf_mlex: "wf R \ wf (f <*mlex*> R)" -unfolding mlex_prod_def -by auto + unfolding mlex_prod_def + by auto lemma mlex_less: "f x < f y \ (x, y) \ f <*mlex*> R" -unfolding mlex_prod_def by simp + unfolding mlex_prod_def by simp lemma mlex_leq: "f x \ f y \ (x, y) \ R \ (x, y) \ f <*mlex*> R" -unfolding mlex_prod_def by auto + unfolding mlex_prod_def by auto text \proper subset relation on finite sets\ -definition finite_psubset :: "('a set * 'a set) set" -where "finite_psubset = {(A,B). A < B & finite B}" +definition finite_psubset :: "('a set \ 'a set) set" + where "finite_psubset = {(A,B). A < B \ finite B}" -lemma wf_finite_psubset[simp]: "wf(finite_psubset)" -apply (unfold finite_psubset_def) -apply (rule wf_measure [THEN wf_subset]) -apply (simp add: measure_def inv_image_def less_than_def less_eq) -apply (fast elim!: psubset_card_mono) -done +lemma wf_finite_psubset[simp]: "wf finite_psubset" + apply (unfold finite_psubset_def) + apply (rule wf_measure [THEN wf_subset]) + apply (simp add: measure_def inv_image_def less_than_def less_eq) + apply (fast elim!: psubset_card_mono) + done lemma trans_finite_psubset: "trans finite_psubset" -by (simp add: finite_psubset_def less_le trans_def, blast) + by (auto simp add: finite_psubset_def less_le trans_def) -lemma in_finite_psubset[simp]: "(A, B) \ finite_psubset = (A < B & finite B)" -unfolding finite_psubset_def by auto +lemma in_finite_psubset[simp]: "(A, B) \ finite_psubset \ A < B \ finite B" + unfolding finite_psubset_def by auto text \max- and min-extension of order to finite sets\ -inductive_set max_ext :: "('a \ 'a) set \ ('a set \ 'a set) set" -for R :: "('a \ 'a) set" +inductive_set max_ext :: "('a \ 'a) set \ ('a set \ 'a set) set" + for R :: "('a \ 'a) set" where max_extI[intro]: "finite X \ finite Y \ Y \ {} \ (\x. x \ X \ \y\Y. (x, y) \ R) \ (X, Y) \ max_ext R" @@ -801,10 +789,11 @@ assumes wf: "wf r" shows "wf (max_ext r)" proof (rule acc_wfI, intro allI) - fix M show "M \ acc (max_ext r)" (is "_ \ ?W") + fix M + show "M \ acc (max_ext r)" (is "_ \ ?W") proof cases assume "finite M" - thus ?thesis + then show ?thesis proof (induct M) show "{} \ ?W" by (rule accI) (auto elim: max_ext.cases) @@ -815,42 +804,37 @@ fix M a assume "M \ ?W" and [intro]: "finite M" assume hyp: "\b M. (b, a) \ r \ M \ ?W \ finite M \ insert b M \ ?W" - { - fix N M :: "'a set" - assume "finite N" "finite M" - then - have "\M \ ?W ; (\y. y \ N \ (y, a) \ r)\ \ N \ M \ ?W" - by (induct N arbitrary: M) (auto simp: hyp) - } - note add_less = this - + have add_less: "\M \ ?W ; (\y. y \ N \ (y, a) \ r)\ \ N \ M \ ?W" + if "finite N" "finite M" for N M :: "'a set" + using that by (induct N arbitrary: M) (auto simp: hyp) + show "insert a M \ ?W" proof (rule accI) - fix N assume Nless: "(N, insert a M) \ max_ext r" - hence asm1: "\x. x \ N \ (x, a) \ r \ (\y \ M. (x, y) \ r)" + fix N + assume Nless: "(N, insert a M) \ max_ext r" + then have *: "\x. x \ N \ (x, a) \ r \ (\y \ M. (x, y) \ r)" by (auto elim!: max_ext.cases) - let ?N1 = "{ n \ N. (n, a) \ r }" - let ?N2 = "{ n \ N. (n, a) \ r }" + let ?N1 = "{n \ N. (n, a) \ r}" + let ?N2 = "{n \ N. (n, a) \ r}" have N: "?N1 \ ?N2 = N" by (rule set_eqI) auto from Nless have "finite N" by (auto elim: max_ext.cases) then have finites: "finite ?N1" "finite ?N2" by auto - + have "?N2 \ ?W" proof cases assume [simp]: "M = {}" have Mw: "{} \ ?W" by (rule accI) (auto elim: max_ext.cases) - from asm1 have "?N2 = {}" by auto + from * have "?N2 = {}" by auto with Mw show "?N2 \ ?W" by (simp only:) next assume "M \ {}" - from asm1 finites have N2: "(?N2, M) \ max_ext r" + from * finites have N2: "(?N2, M) \ max_ext r" by (rule_tac max_extI[OF _ _ \M \ {}\]) auto - with \M \ ?W\ show "?N2 \ ?W" by (rule acc_downward) qed - with finites have "?N1 \ ?N2 \ ?W" + with finites have "?N1 \ ?N2 \ ?W" by (rule add_less) simp then show "N \ ?W" by (simp only: N) qed @@ -863,32 +847,30 @@ qed qed -lemma max_ext_additive: - "(A, B) \ max_ext R \ (C, D) \ max_ext R \ - (A \ C, B \ D) \ max_ext R" -by (force elim!: max_ext.cases) +lemma max_ext_additive: + "(A, B) \ max_ext R \ (C, D) \ max_ext R \ + (A \ C, B \ D) \ max_ext R" + by (force elim!: max_ext.cases) -definition min_ext :: "('a \ 'a) set \ ('a set \ 'a set) set" where - "min_ext r = {(X, Y) | X Y. X \ {} \ (\y \ Y. (\x \ X. (x, y) \ r))}" +definition min_ext :: "('a \ 'a) set \ ('a set \ 'a set) set" + where "min_ext r = {(X, Y) | X Y. X \ {} \ (\y \ Y. (\x \ X. (x, y) \ r))}" lemma min_ext_wf: assumes "wf r" shows "wf (min_ext r)" proof (rule wfI_min) - fix Q :: "'a set set" - fix x - assume nonempty: "x \ Q" - show "\m \ Q. (\ n. (n, m) \ min_ext r \ n \ Q)" - proof cases - assume "Q = {{}}" thus ?thesis by (simp add: min_ext_def) + show "\m \ Q. (\ n. (n, m) \ min_ext r \ n \ Q)" if nonempty: "x \ Q" + for Q :: "'a set set" and x + proof (cases "Q = {{}}") + case True + then show ?thesis by (simp add: min_ext_def) next - assume "Q \ {{}}" - with nonempty - obtain e x where "x \ Q" "e \ x" by force + case False + with nonempty obtain e x where "x \ Q" "e \ x" by force then have eU: "e \ \Q" by auto - with \wf r\ - obtain z where z: "z \ \Q" "\y. (y, z) \ r \ y \ \Q" + with \wf r\ + obtain z where z: "z \ \Q" "\y. (y, z) \ r \ y \ \Q" by (erule wfE_min) from z obtain m where "m \ Q" "z \ m" by auto from \m \ Q\ @@ -898,36 +880,38 @@ assume smaller: "(n, m) \ min_ext r" with \z \ m\ obtain y where y: "y \ n" "(y, z) \ r" by (auto simp: min_ext_def) then show "n \ Q" using z(2) by auto - qed + qed qed qed -text\Bounded increase must terminate:\ + +subsubsection \Bounded increase must terminate\ lemma wf_bounded_measure: -fixes ub :: "'a \ nat" and f :: "'a \ nat" -assumes "!!a b. (b,a) : r \ ub b \ ub a & ub a \ f b & f b > f a" -shows "wf r" -apply(rule wf_subset[OF wf_measure[of "%a. ub a - f a"]]) -apply (auto dest: assms) -done + fixes ub :: "'a \ nat" + and f :: "'a \ nat" + assumes "\a b. (b, a) \ r \ ub b \ ub a \ ub a \ f b \ f b > f a" + shows "wf r" + apply (rule wf_subset[OF wf_measure[of "\a. ub a - f a"]]) + apply (auto dest: assms) + done lemma wf_bounded_set: -fixes ub :: "'a \ 'b set" and f :: "'a \ 'b set" -assumes "!!a b. (b,a) : r \ - finite(ub a) & ub b \ ub a & ub a \ f b & f b \ f a" -shows "wf r" -apply(rule wf_bounded_measure[of r "%a. card(ub a)" "%a. card(f a)"]) -apply(drule assms) -apply (blast intro: card_mono finite_subset psubset_card_mono dest: psubset_eq[THEN iffD2]) -done + fixes ub :: "'a \ 'b set" + and f :: "'a \ 'b set" + assumes "\a b. (b,a) \ r \ finite (ub a) \ ub b \ ub a \ ub a \ f b \ f b \ f a" + shows "wf r" + apply(rule wf_bounded_measure[of r "\a. card(ub a)" "\a. card(f a)"]) + apply(drule assms) + apply (blast intro: card_mono finite_subset psubset_card_mono dest: psubset_eq[THEN iffD2]) + done lemma finite_subset_wf: assumes "finite A" shows "wf {(X,Y). X \ Y \ Y \ A}" proof (intro finite_acyclic_wf) have "{(X,Y). X \ Y \ Y \ A} \ Pow A \ Pow A" by blast - thus "finite {(X,Y). X \ Y \ Y \ A}" + then show "finite {(X,Y). X \ Y \ Y \ A}" by (rule finite_subset) (auto simp: assms finite_cartesian_product) next have "{(X, Y). X \ Y \ Y \ A}\<^sup>+ = {(X, Y). X \ Y \ Y \ A}"