diff --git a/src/HOL/Relation.thy b/src/HOL/Relation.thy --- a/src/HOL/Relation.thy +++ b/src/HOL/Relation.thy @@ -1,1446 +1,1446 @@ (* Title: HOL/Relation.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Stefan Berghofer, TU Muenchen Author: Martin Desharnais, MPI-INF Saarbruecken *) section \Relations -- as sets of pairs, and binary predicates\ theory Relation imports Finite_Set begin text \A preliminary: classical rules for reasoning on predicates\ declare predicate1I [Pure.intro!, intro!] declare predicate1D [Pure.dest, dest] declare predicate2I [Pure.intro!, intro!] declare predicate2D [Pure.dest, dest] declare bot1E [elim!] declare bot2E [elim!] declare top1I [intro!] declare top2I [intro!] declare inf1I [intro!] declare inf2I [intro!] declare inf1E [elim!] declare inf2E [elim!] declare sup1I1 [intro?] declare sup2I1 [intro?] declare sup1I2 [intro?] declare sup2I2 [intro?] declare sup1E [elim!] declare sup2E [elim!] declare sup1CI [intro!] declare sup2CI [intro!] declare Inf1_I [intro!] declare INF1_I [intro!] declare Inf2_I [intro!] declare INF2_I [intro!] declare Inf1_D [elim] declare INF1_D [elim] declare Inf2_D [elim] declare INF2_D [elim] declare Inf1_E [elim] declare INF1_E [elim] declare Inf2_E [elim] declare INF2_E [elim] declare Sup1_I [intro] declare SUP1_I [intro] declare Sup2_I [intro] declare SUP2_I [intro] declare Sup1_E [elim!] declare SUP1_E [elim!] declare Sup2_E [elim!] declare SUP2_E [elim!] subsection \Fundamental\ subsubsection \Relations as sets of pairs\ type_synonym 'a rel = "('a \ 'a) set" lemma subrelI: "(\x y. (x, y) \ r \ (x, y) \ s) \ r \ s" \ \Version of @{thm [source] subsetI} for binary relations\ by auto lemma lfp_induct2: "(a, b) \ lfp f \ mono f \ (\a b. (a, b) \ f (lfp f \ {(x, y). P x y}) \ P a b) \ P a b" \ \Version of @{thm [source] lfp_induct} for binary relations\ using lfp_induct_set [of "(a, b)" f "case_prod P"] by auto subsubsection \Conversions between set and predicate relations\ lemma pred_equals_eq [pred_set_conv]: "(\x. x \ R) = (\x. x \ S) \ R = S" by (simp add: set_eq_iff fun_eq_iff) lemma pred_equals_eq2 [pred_set_conv]: "(\x y. (x, y) \ R) = (\x y. (x, y) \ S) \ R = S" by (simp add: set_eq_iff fun_eq_iff) lemma pred_subset_eq [pred_set_conv]: "(\x. x \ R) \ (\x. x \ S) \ R \ S" by (simp add: subset_iff le_fun_def) lemma pred_subset_eq2 [pred_set_conv]: "(\x y. (x, y) \ R) \ (\x y. (x, y) \ S) \ R \ S" by (simp add: subset_iff le_fun_def) lemma bot_empty_eq [pred_set_conv]: "\ = (\x. x \ {})" by (auto simp add: fun_eq_iff) lemma bot_empty_eq2 [pred_set_conv]: "\ = (\x y. (x, y) \ {})" by (auto simp add: fun_eq_iff) lemma top_empty_eq: "\ = (\x. x \ UNIV)" by (auto simp add: fun_eq_iff) lemma top_empty_eq2: "\ = (\x y. (x, y) \ UNIV)" by (auto simp add: fun_eq_iff) lemma inf_Int_eq [pred_set_conv]: "(\x. x \ R) \ (\x. x \ S) = (\x. x \ R \ S)" by (simp add: inf_fun_def) lemma inf_Int_eq2 [pred_set_conv]: "(\x y. (x, y) \ R) \ (\x y. (x, y) \ S) = (\x y. (x, y) \ R \ S)" by (simp add: inf_fun_def) lemma sup_Un_eq [pred_set_conv]: "(\x. x \ R) \ (\x. x \ S) = (\x. x \ R \ S)" by (simp add: sup_fun_def) lemma sup_Un_eq2 [pred_set_conv]: "(\x y. (x, y) \ R) \ (\x y. (x, y) \ S) = (\x y. (x, y) \ R \ S)" by (simp add: sup_fun_def) lemma INF_INT_eq [pred_set_conv]: "(\i\S. (\x. x \ r i)) = (\x. x \ (\i\S. r i))" by (simp add: fun_eq_iff) lemma INF_INT_eq2 [pred_set_conv]: "(\i\S. (\x y. (x, y) \ r i)) = (\x y. (x, y) \ (\i\S. r i))" by (simp add: fun_eq_iff) lemma SUP_UN_eq [pred_set_conv]: "(\i\S. (\x. x \ r i)) = (\x. x \ (\i\S. r i))" by (simp add: fun_eq_iff) lemma SUP_UN_eq2 [pred_set_conv]: "(\i\S. (\x y. (x, y) \ r i)) = (\x y. (x, y) \ (\i\S. r i))" by (simp add: fun_eq_iff) lemma Inf_INT_eq [pred_set_conv]: "\S = (\x. x \ (\(Collect ` S)))" by (simp add: fun_eq_iff) lemma INF_Int_eq [pred_set_conv]: "(\i\S. (\x. x \ i)) = (\x. x \ \S)" by (simp add: fun_eq_iff) lemma Inf_INT_eq2 [pred_set_conv]: "\S = (\x y. (x, y) \ (\(Collect ` case_prod ` S)))" by (simp add: fun_eq_iff) lemma INF_Int_eq2 [pred_set_conv]: "(\i\S. (\x y. (x, y) \ i)) = (\x y. (x, y) \ \S)" by (simp add: fun_eq_iff) lemma Sup_SUP_eq [pred_set_conv]: "\S = (\x. x \ \(Collect ` S))" by (simp add: fun_eq_iff) lemma SUP_Sup_eq [pred_set_conv]: "(\i\S. (\x. x \ i)) = (\x. x \ \S)" by (simp add: fun_eq_iff) lemma Sup_SUP_eq2 [pred_set_conv]: "\S = (\x y. (x, y) \ (\(Collect ` case_prod ` S)))" by (simp add: fun_eq_iff) lemma SUP_Sup_eq2 [pred_set_conv]: "(\i\S. (\x y. (x, y) \ i)) = (\x y. (x, y) \ \S)" by (simp add: fun_eq_iff) subsection \Properties of relations\ subsubsection \Reflexivity\ definition refl_on :: "'a set \ 'a rel \ bool" where "refl_on A r \ r \ A \ A \ (\x\A. (x, x) \ r)" abbreviation refl :: "'a rel \ bool" \ \reflexivity over a type\ where "refl \ refl_on UNIV" definition reflp_on :: "'a set \ ('a \ 'a \ bool) \ bool" where "reflp_on A R \ (\x\A. R x x)" abbreviation reflp :: "('a \ 'a \ bool) \ bool" where "reflp \ reflp_on UNIV" lemma reflp_def[no_atp]: "reflp R \ (\x. R x x)" by (simp add: reflp_on_def) text \@{thm [source] reflp_def} is for backward compatibility.\ lemma reflp_refl_eq [pred_set_conv]: "reflp (\x y. (x, y) \ r) \ refl r" by (simp add: refl_on_def reflp_def) lemma refl_onI [intro?]: "r \ A \ A \ (\x. x \ A \ (x, x) \ r) \ refl_on A r" unfolding refl_on_def by (iprover intro!: ballI) lemma reflp_onI: "(\x. x \ A \ R x x) \ reflp_on A R" by (simp add: reflp_on_def) lemma reflpI[intro?]: "(\x. R x x) \ reflp R" by (rule reflp_onI) lemma refl_onD: "refl_on A r \ a \ A \ (a, a) \ r" unfolding refl_on_def by blast lemma refl_onD1: "refl_on A r \ (x, y) \ r \ x \ A" unfolding refl_on_def by blast lemma refl_onD2: "refl_on A r \ (x, y) \ r \ y \ A" unfolding refl_on_def by blast lemma reflp_onD: "reflp_on A R \ x \ A \ R x x" by (simp add: reflp_on_def) lemma reflpD[dest?]: "reflp R \ R x x" by (simp add: reflp_onD) lemma reflpE: assumes "reflp r" obtains "r x x" using assms by (auto dest: refl_onD simp add: reflp_def) lemma reflp_on_subset: "reflp_on A R \ B \ A \ reflp_on B R" by (auto intro: reflp_onI dest: reflp_onD) lemma refl_on_Int: "refl_on A r \ refl_on B s \ refl_on (A \ B) (r \ s)" unfolding refl_on_def by blast lemma reflp_on_inf: "reflp_on A R \ reflp_on B S \ reflp_on (A \ B) (R \ S)" by (auto intro: reflp_onI dest: reflp_onD) lemma reflp_inf: "reflp r \ reflp s \ reflp (r \ s)" by (rule reflp_on_inf[of UNIV _ UNIV, unfolded Int_absorb]) lemma refl_on_Un: "refl_on A r \ refl_on B s \ refl_on (A \ B) (r \ s)" unfolding refl_on_def by blast lemma reflp_on_sup: "reflp_on A R \ reflp_on B S \ reflp_on (A \ B) (R \ S)" by (auto intro: reflp_onI dest: reflp_onD) lemma reflp_sup: "reflp r \ reflp s \ reflp (r \ s)" by (rule reflp_on_sup[of UNIV _ UNIV, unfolded Un_absorb]) lemma refl_on_INTER: "\x\S. refl_on (A x) (r x) \ refl_on (\(A ` S)) (\(r ` S))" unfolding refl_on_def by fast lemma reflp_on_Inf: "\x\S. reflp_on (A x) (R x) \ reflp_on (\(A ` S)) (\(R ` S))" by (auto intro: reflp_onI dest: reflp_onD) lemma refl_on_UNION: "\x\S. refl_on (A x) (r x) \ refl_on (\(A ` S)) (\(r ` S))" unfolding refl_on_def by blast lemma reflp_on_Sup: "\x\S. reflp_on (A x) (R x) \ reflp_on (\(A ` S)) (\(R ` S))" by (auto intro: reflp_onI dest: reflp_onD) lemma refl_on_empty [simp]: "refl_on {} {}" by (simp add: refl_on_def) lemma reflp_on_empty [simp]: "reflp_on {} R" by (auto intro: reflp_onI) lemma refl_on_singleton [simp]: "refl_on {x} {(x, x)}" by (blast intro: refl_onI) lemma refl_on_def' [nitpick_unfold, code]: "refl_on A r \ (\(x, y) \ r. x \ A \ y \ A) \ (\x \ A. (x, x) \ r)" by (auto intro: refl_onI dest: refl_onD refl_onD1 refl_onD2) lemma reflp_on_equality [simp]: "reflp_on A (=)" by (simp add: reflp_on_def) lemma reflp_on_mono: "reflp_on A R \ (\x y. x \ A \ y \ A \ R x y \ Q x y) \ reflp_on A Q" by (auto intro: reflp_onI dest: reflp_onD) lemma reflp_mono: "reflp R \ (\x y. R x y \ Q x y) \ reflp Q" by (rule reflp_on_mono[of UNIV R Q]) simp_all lemma (in preorder) reflp_on_le[simp]: "reflp_on A (\)" by (simp add: reflp_onI) lemma (in preorder) reflp_on_ge[simp]: "reflp_on A (\)" by (simp add: reflp_onI) subsubsection \Irreflexivity\ definition irrefl_on :: "'a set \ 'a rel \ bool" where "irrefl_on A r \ (\a \ A. (a, a) \ r)" abbreviation irrefl :: "'a rel \ bool" where "irrefl \ irrefl_on UNIV" definition irreflp_on :: "'a set \ ('a \ 'a \ bool) \ bool" where "irreflp_on A R \ (\a \ A. \ R a a)" abbreviation irreflp :: "('a \ 'a \ bool) \ bool" where "irreflp \ irreflp_on UNIV" lemma irrefl_def[no_atp]: "irrefl r \ (\a. (a, a) \ r)" by (simp add: irrefl_on_def) lemma irreflp_def[no_atp]: "irreflp R \ (\a. \ R a a)" by (simp add: irreflp_on_def) text \@{thm [source] irrefl_def} and @{thm [source] irreflp_def} are for backward compatibility.\ lemma irreflp_on_irrefl_on_eq [pred_set_conv]: "irreflp_on A (\a b. (a, b) \ r) \ irrefl_on A r" by (simp add: irrefl_on_def irreflp_on_def) lemmas irreflp_irrefl_eq = irreflp_on_irrefl_on_eq[of UNIV] lemma irrefl_onI: "(\a. a \ A \ (a, a) \ r) \ irrefl_on A r" by (simp add: irrefl_on_def) lemma irreflI[intro?]: "(\a. (a, a) \ r) \ irrefl r" by (rule irrefl_onI[of UNIV, simplified]) lemma irreflp_onI: "(\a. a \ A \ \ R a a) \ irreflp_on A R" - by (simp add: irreflp_on_def) + by (rule irrefl_onI[to_pred]) lemma irreflpI[intro?]: "(\a. \ R a a) \ irreflp R" - by (rule irreflp_onI[of UNIV, simplified]) + by (rule irreflI[to_pred]) lemma irrefl_onD: "irrefl_on A r \ a \ A \ (a, a) \ r" by (simp add: irrefl_on_def) lemma irreflD: "irrefl r \ (x, x) \ r" by (rule irrefl_onD[of UNIV, simplified]) lemma irreflp_onD: "irreflp_on A R \ a \ A \ \ R a a" - by (simp add: irreflp_on_def) + by (rule irrefl_onD[to_pred]) lemma irreflpD: "irreflp R \ \ R x x" - by (rule irreflp_onD[of UNIV, simplified]) + by (rule irreflD[to_pred]) lemma irrefl_on_distinct [code]: "irrefl_on A r \ (\(a, b) \ r. a \ A \ b \ A \ a \ b)" by (auto simp add: irrefl_on_def) lemmas irrefl_distinct = irrefl_on_distinct \ \For backward compatibility\ lemma irrefl_on_subset: "irrefl_on A r \ B \ A \ irrefl_on B r" by (auto simp: irrefl_on_def) lemma irreflp_on_subset: "irreflp_on A R \ B \ A \ irreflp_on B R" by (auto simp: irreflp_on_def) lemma (in preorder) irreflp_on_less[simp]: "irreflp_on A (<)" by (simp add: irreflp_onI) lemma (in preorder) irreflp_on_greater[simp]: "irreflp_on A (>)" by (simp add: irreflp_onI) subsubsection \Asymmetry\ inductive asym :: "'a rel \ bool" where asymI: "(\a b. (a, b) \ R \ (b, a) \ R) \ asym R" inductive asymp :: "('a \ 'a \ bool) \ bool" where asympI: "(\a b. R a b \ \ R b a) \ asymp R" lemma asymp_asym_eq [pred_set_conv]: "asymp (\a b. (a, b) \ R) \ asym R" by (auto intro!: asymI asympI elim: asym.cases asymp.cases simp add: irreflp_irrefl_eq) lemma asymD: "\asym R; (x,y) \ R\ \ (y,x) \ R" by (simp add: asym.simps) lemma asympD: "asymp R \ R x y \ \ R y x" by (rule asymD[to_pred]) lemma asym_iff: "asym R \ (\x y. (x,y) \ R \ (y,x) \ R)" by (blast intro: asymI dest: asymD) lemma (in preorder) asymp_less[simp]: "asymp (<)" by (auto intro: asympI dual_order.asym) lemma (in preorder) asymp_greater[simp]: "asymp (>)" by (auto intro: asympI dual_order.asym) subsubsection \Symmetry\ definition sym :: "'a rel \ bool" where "sym r \ (\x y. (x, y) \ r \ (y, x) \ r)" definition symp :: "('a \ 'a \ bool) \ bool" where "symp r \ (\x y. r x y \ r y x)" lemma symp_sym_eq [pred_set_conv]: "symp (\x y. (x, y) \ r) \ sym r" by (simp add: sym_def symp_def) lemma symI [intro?]: "(\a b. (a, b) \ r \ (b, a) \ r) \ sym r" by (unfold sym_def) iprover lemma sympI [intro?]: "(\a b. r a b \ r b a) \ symp r" by (fact symI [to_pred]) lemma symE: assumes "sym r" and "(b, a) \ r" obtains "(a, b) \ r" using assms by (simp add: sym_def) lemma sympE: assumes "symp r" and "r b a" obtains "r a b" using assms by (rule symE [to_pred]) lemma symD [dest?]: assumes "sym r" and "(b, a) \ r" shows "(a, b) \ r" using assms by (rule symE) lemma sympD [dest?]: assumes "symp r" and "r b a" shows "r a b" using assms by (rule symD [to_pred]) lemma sym_Int: "sym r \ sym s \ sym (r \ s)" by (fast intro: symI elim: symE) lemma symp_inf: "symp r \ symp s \ symp (r \ s)" by (fact sym_Int [to_pred]) lemma sym_Un: "sym r \ sym s \ sym (r \ s)" by (fast intro: symI elim: symE) lemma symp_sup: "symp r \ symp s \ symp (r \ s)" by (fact sym_Un [to_pred]) lemma sym_INTER: "\x\S. sym (r x) \ sym (\(r ` S))" by (fast intro: symI elim: symE) lemma symp_INF: "\x\S. symp (r x) \ symp (\(r ` S))" by (fact sym_INTER [to_pred]) lemma sym_UNION: "\x\S. sym (r x) \ sym (\(r ` S))" by (fast intro: symI elim: symE) lemma symp_SUP: "\x\S. symp (r x) \ symp (\(r ` S))" by (fact sym_UNION [to_pred]) subsubsection \Antisymmetry\ definition antisym :: "'a rel \ bool" where "antisym r \ (\x y. (x, y) \ r \ (y, x) \ r \ x = y)" definition antisymp :: "('a \ 'a \ bool) \ bool" where "antisymp r \ (\x y. r x y \ r y x \ x = y)" lemma antisymp_antisym_eq [pred_set_conv]: "antisymp (\x y. (x, y) \ r) \ antisym r" by (simp add: antisym_def antisymp_def) lemma antisymI [intro?]: "(\x y. (x, y) \ r \ (y, x) \ r \ x = y) \ antisym r" unfolding antisym_def by iprover lemma antisympI [intro?]: "(\x y. r x y \ r y x \ x = y) \ antisymp r" by (fact antisymI [to_pred]) lemma antisymD [dest?]: "antisym r \ (a, b) \ r \ (b, a) \ r \ a = b" unfolding antisym_def by iprover lemma antisympD [dest?]: "antisymp r \ r a b \ r b a \ a = b" by (fact antisymD [to_pred]) lemma antisym_subset: "r \ s \ antisym s \ antisym r" unfolding antisym_def by blast lemma antisymp_less_eq: "r \ s \ antisymp s \ antisymp r" by (fact antisym_subset [to_pred]) lemma antisym_empty [simp]: "antisym {}" unfolding antisym_def by blast lemma antisym_bot [simp]: "antisymp \" by (fact antisym_empty [to_pred]) lemma antisymp_equality [simp]: "antisymp HOL.eq" by (auto intro: antisympI) lemma antisym_singleton [simp]: "antisym {x}" by (blast intro: antisymI) lemma antisym_if_asym: "asym r \ antisym r" by (auto intro: antisymI elim: asym.cases) lemma antisymp_if_asymp: "asymp R \ antisymp R" by (rule antisym_if_asym[to_pred]) lemma (in preorder) antisymp_less[simp]: "antisymp (<)" by (rule antisymp_if_asymp[OF asymp_less]) lemma (in preorder) antisymp_greater[simp]: "antisymp (>)" by (rule antisymp_if_asymp[OF asymp_greater]) lemma (in order) antisymp_le[simp]: "antisymp (\)" by (simp add: antisympI) lemma (in order) antisymp_ge[simp]: "antisymp (\)" by (simp add: antisympI) subsubsection \Transitivity\ definition trans :: "'a rel \ bool" where "trans r \ (\x y z. (x, y) \ r \ (y, z) \ r \ (x, z) \ r)" definition transp :: "('a \ 'a \ bool) \ bool" where "transp r \ (\x y z. r x y \ r y z \ r x z)" lemma transp_trans_eq [pred_set_conv]: "transp (\x y. (x, y) \ r) \ trans r" by (simp add: trans_def transp_def) lemma transI [intro?]: "(\x y z. (x, y) \ r \ (y, z) \ r \ (x, z) \ r) \ trans r" by (unfold trans_def) iprover lemma transpI [intro?]: "(\x y z. r x y \ r y z \ r x z) \ transp r" by (fact transI [to_pred]) lemma transE: assumes "trans r" and "(x, y) \ r" and "(y, z) \ r" obtains "(x, z) \ r" using assms by (unfold trans_def) iprover lemma transpE: assumes "transp r" and "r x y" and "r y z" obtains "r x z" using assms by (rule transE [to_pred]) lemma transD [dest?]: assumes "trans r" and "(x, y) \ r" and "(y, z) \ r" shows "(x, z) \ r" using assms by (rule transE) lemma transpD [dest?]: assumes "transp r" and "r x y" and "r y z" shows "r x z" using assms by (rule transD [to_pred]) lemma trans_Int: "trans r \ trans s \ trans (r \ s)" by (fast intro: transI elim: transE) lemma transp_inf: "transp r \ transp s \ transp (r \ s)" by (fact trans_Int [to_pred]) lemma trans_INTER: "\x\S. trans (r x) \ trans (\(r ` S))" by (fast intro: transI elim: transD) lemma transp_INF: "\x\S. transp (r x) \ transp (\(r ` S))" by (fact trans_INTER [to_pred]) lemma trans_join [code]: "trans r \ (\(x, y1) \ r. \(y2, z) \ r. y1 = y2 \ (x, z) \ r)" by (auto simp add: trans_def) lemma transp_trans: "transp r \ trans {(x, y). r x y}" by (simp add: trans_def transp_def) lemma transp_equality [simp]: "transp (=)" by (auto intro: transpI) lemma trans_empty [simp]: "trans {}" by (blast intro: transI) lemma transp_empty [simp]: "transp (\x y. False)" using trans_empty[to_pred] by (simp add: bot_fun_def) lemma trans_singleton [simp]: "trans {(a, a)}" by (blast intro: transI) lemma transp_singleton [simp]: "transp (\x y. x = a \ y = a)" by (simp add: transp_def) lemma asym_if_irrefl_and_trans: "irrefl R \ trans R \ asym R" by (auto intro: asymI dest: transD irreflD) lemma asymp_if_irreflp_and_transp: "irreflp R \ transp R \ asymp R" by (rule asym_if_irrefl_and_trans[to_pred]) context preorder begin lemma transp_le[simp]: "transp (\)" by(auto simp add: transp_def intro: order_trans) lemma transp_less[simp]: "transp (<)" by(auto simp add: transp_def intro: less_trans) lemma transp_ge[simp]: "transp (\)" by(auto simp add: transp_def intro: order_trans) lemma transp_gr[simp]: "transp (>)" by(auto simp add: transp_def intro: less_trans) end subsubsection \Totality\ definition total_on :: "'a set \ 'a rel \ bool" where "total_on A r \ (\x\A. \y\A. x \ y \ (x, y) \ r \ (y, x) \ r)" abbreviation total :: "'a rel \ bool" where "total \ total_on UNIV" definition totalp_on :: "'a set \ ('a \ 'a \ bool) \ bool" where "totalp_on A R \ (\x \ A. \y \ A. x \ y \ R x y \ R y x)" abbreviation totalp :: "('a \ 'a \ bool) \ bool" where "totalp \ totalp_on UNIV" lemma totalp_on_refl_on_eq[pred_set_conv]: "totalp_on A (\x y. (x, y) \ r) \ total_on A r" by (simp add: totalp_on_def total_on_def) lemma total_onI [intro?]: "(\x y. x \ A \ y \ A \ x \ y \ (x, y) \ r \ (y, x) \ r) \ total_on A r" unfolding total_on_def by blast lemma totalI: "(\x y. x \ y \ (x, y) \ r \ (y, x) \ r) \ total r" by (rule total_onI) lemma totalp_onI: "(\x y. x \ A \ y \ A \ x \ y \ R x y \ R y x) \ totalp_on A R" - by (simp add: totalp_on_def) + by (rule total_onI[to_pred]) lemma totalpI: "(\x y. x \ y \ R x y \ R y x) \ totalp R" - by (rule totalp_onI) + by (rule totalI[to_pred]) lemma totalp_onD: "totalp_on A R \ x \ A \ y \ A \ x \ y \ R x y \ R y x" by (simp add: totalp_on_def) lemma totalpD: "totalp R \ x \ y \ R x y \ R y x" by (simp add: totalp_onD) lemma total_on_subset: "total_on A r \ B \ A \ total_on B r" by (auto simp: total_on_def) lemma totalp_on_subset: "totalp_on A R \ B \ A \ totalp_on B R" by (auto intro: totalp_onI dest: totalp_onD) lemma total_on_empty [simp]: "total_on {} r" by (simp add: total_on_def) lemma totalp_on_empty [simp]: "totalp_on {} R" by (simp add: totalp_on_def) lemma total_on_singleton [simp]: "total_on {x} r" by (simp add: total_on_def) lemma totalp_on_singleton [simp]: "totalp_on {x} R" by (simp add: totalp_on_def) lemma (in linorder) totalp_on_less[simp]: "totalp_on A (<)" by (auto intro: totalp_onI) lemma (in linorder) totalp_on_greater[simp]: "totalp_on A (>)" by (auto intro: totalp_onI) lemma (in linorder) totalp_on_le[simp]: "totalp_on A (\)" by (rule totalp_onI, rule linear) lemma (in linorder) totalp_on_ge[simp]: "totalp_on A (\)" by (rule totalp_onI, rule linear) subsubsection \Single valued relations\ definition single_valued :: "('a \ 'b) set \ bool" where "single_valued r \ (\x y. (x, y) \ r \ (\z. (x, z) \ r \ y = z))" definition single_valuedp :: "('a \ 'b \ bool) \ bool" where "single_valuedp r \ (\x y. r x y \ (\z. r x z \ y = z))" lemma single_valuedp_single_valued_eq [pred_set_conv]: "single_valuedp (\x y. (x, y) \ r) \ single_valued r" by (simp add: single_valued_def single_valuedp_def) lemma single_valuedp_iff_Uniq: "single_valuedp r \ (\x. \\<^sub>\\<^sub>1y. r x y)" unfolding Uniq_def single_valuedp_def by auto lemma single_valuedI: "(\x y. (x, y) \ r \ (\z. (x, z) \ r \ y = z)) \ single_valued r" unfolding single_valued_def by blast lemma single_valuedpI: "(\x y. r x y \ (\z. r x z \ y = z)) \ single_valuedp r" by (fact single_valuedI [to_pred]) lemma single_valuedD: "single_valued r \ (x, y) \ r \ (x, z) \ r \ y = z" by (simp add: single_valued_def) lemma single_valuedpD: "single_valuedp r \ r x y \ r x z \ y = z" by (fact single_valuedD [to_pred]) lemma single_valued_empty [simp]: "single_valued {}" by (simp add: single_valued_def) lemma single_valuedp_bot [simp]: "single_valuedp \" by (fact single_valued_empty [to_pred]) lemma single_valued_subset: "r \ s \ single_valued s \ single_valued r" unfolding single_valued_def by blast lemma single_valuedp_less_eq: "r \ s \ single_valuedp s \ single_valuedp r" by (fact single_valued_subset [to_pred]) subsection \Relation operations\ subsubsection \The identity relation\ definition Id :: "'a rel" where "Id = {p. \x. p = (x, x)}" lemma IdI [intro]: "(a, a) \ Id" by (simp add: Id_def) lemma IdE [elim!]: "p \ Id \ (\x. p = (x, x) \ P) \ P" unfolding Id_def by (iprover elim: CollectE) lemma pair_in_Id_conv [iff]: "(a, b) \ Id \ a = b" unfolding Id_def by blast lemma refl_Id: "refl Id" by (simp add: refl_on_def) lemma antisym_Id: "antisym Id" \ \A strange result, since \Id\ is also symmetric.\ by (simp add: antisym_def) lemma sym_Id: "sym Id" by (simp add: sym_def) lemma trans_Id: "trans Id" by (simp add: trans_def) lemma single_valued_Id [simp]: "single_valued Id" by (unfold single_valued_def) blast lemma irrefl_diff_Id [simp]: "irrefl (r - Id)" by (simp add: irrefl_def) lemma trans_diff_Id: "trans r \ antisym r \ trans (r - Id)" unfolding antisym_def trans_def by blast lemma total_on_diff_Id [simp]: "total_on A (r - Id) = total_on A r" by (simp add: total_on_def) lemma Id_fstsnd_eq: "Id = {x. fst x = snd x}" by force subsubsection \Diagonal: identity over a set\ definition Id_on :: "'a set \ 'a rel" where "Id_on A = (\x\A. {(x, x)})" lemma Id_on_empty [simp]: "Id_on {} = {}" by (simp add: Id_on_def) lemma Id_on_eqI: "a = b \ a \ A \ (a, b) \ Id_on A" by (simp add: Id_on_def) lemma Id_onI [intro!]: "a \ A \ (a, a) \ Id_on A" by (rule Id_on_eqI) (rule refl) lemma Id_onE [elim!]: "c \ Id_on A \ (\x. x \ A \ c = (x, x) \ P) \ P" \ \The general elimination rule.\ unfolding Id_on_def by (iprover elim!: UN_E singletonE) lemma Id_on_iff: "(x, y) \ Id_on A \ x = y \ x \ A" by blast lemma Id_on_def' [nitpick_unfold]: "Id_on {x. A x} = Collect (\(x, y). x = y \ A x)" by auto lemma Id_on_subset_Times: "Id_on A \ A \ A" by blast lemma refl_on_Id_on: "refl_on A (Id_on A)" by (rule refl_onI [OF Id_on_subset_Times Id_onI]) lemma antisym_Id_on [simp]: "antisym (Id_on A)" unfolding antisym_def by blast lemma sym_Id_on [simp]: "sym (Id_on A)" by (rule symI) clarify lemma trans_Id_on [simp]: "trans (Id_on A)" by (fast intro: transI elim: transD) lemma single_valued_Id_on [simp]: "single_valued (Id_on A)" unfolding single_valued_def by blast subsubsection \Composition\ inductive_set relcomp :: "('a \ 'b) set \ ('b \ 'c) set \ ('a \ 'c) set" (infixr "O" 75) for r :: "('a \ 'b) set" and s :: "('b \ 'c) set" where relcompI [intro]: "(a, b) \ r \ (b, c) \ s \ (a, c) \ r O s" notation relcompp (infixr "OO" 75) lemmas relcomppI = relcompp.intros text \ For historic reasons, the elimination rules are not wholly corresponding. Feel free to consolidate this. \ inductive_cases relcompEpair: "(a, c) \ r O s" inductive_cases relcomppE [elim!]: "(r OO s) a c" lemma relcompE [elim!]: "xz \ r O s \ (\x y z. xz = (x, z) \ (x, y) \ r \ (y, z) \ s \ P) \ P" apply (cases xz) apply simp apply (erule relcompEpair) apply iprover done lemma R_O_Id [simp]: "R O Id = R" by fast lemma Id_O_R [simp]: "Id O R = R" by fast lemma relcomp_empty1 [simp]: "{} O R = {}" by blast lemma relcompp_bot1 [simp]: "\ OO R = \" by (fact relcomp_empty1 [to_pred]) lemma relcomp_empty2 [simp]: "R O {} = {}" by blast lemma relcompp_bot2 [simp]: "R OO \ = \" by (fact relcomp_empty2 [to_pred]) lemma O_assoc: "(R O S) O T = R O (S O T)" by blast lemma relcompp_assoc: "(r OO s) OO t = r OO (s OO t)" by (fact O_assoc [to_pred]) lemma trans_O_subset: "trans r \ r O r \ r" by (unfold trans_def) blast lemma transp_relcompp_less_eq: "transp r \ r OO r \ r " by (fact trans_O_subset [to_pred]) lemma relcomp_mono: "r' \ r \ s' \ s \ r' O s' \ r O s" by blast lemma relcompp_mono: "r' \ r \ s' \ s \ r' OO s' \ r OO s " by (fact relcomp_mono [to_pred]) lemma relcomp_subset_Sigma: "r \ A \ B \ s \ B \ C \ r O s \ A \ C" by blast lemma relcomp_distrib [simp]: "R O (S \ T) = (R O S) \ (R O T)" by auto lemma relcompp_distrib [simp]: "R OO (S \ T) = R OO S \ R OO T" by (fact relcomp_distrib [to_pred]) lemma relcomp_distrib2 [simp]: "(S \ T) O R = (S O R) \ (T O R)" by auto lemma relcompp_distrib2 [simp]: "(S \ T) OO R = S OO R \ T OO R" by (fact relcomp_distrib2 [to_pred]) lemma relcomp_UNION_distrib: "s O \(r ` I) = (\i\I. s O r i) " by auto lemma relcompp_SUP_distrib: "s OO \(r ` I) = (\i\I. s OO r i)" by (fact relcomp_UNION_distrib [to_pred]) lemma relcomp_UNION_distrib2: "\(r ` I) O s = (\i\I. r i O s) " by auto lemma relcompp_SUP_distrib2: "\(r ` I) OO s = (\i\I. r i OO s)" by (fact relcomp_UNION_distrib2 [to_pred]) lemma single_valued_relcomp: "single_valued r \ single_valued s \ single_valued (r O s)" unfolding single_valued_def by blast lemma relcomp_unfold: "r O s = {(x, z). \y. (x, y) \ r \ (y, z) \ s}" by (auto simp add: set_eq_iff) lemma relcompp_apply: "(R OO S) a c \ (\b. R a b \ S b c)" unfolding relcomp_unfold [to_pred] .. lemma eq_OO: "(=) OO R = R" by blast lemma OO_eq: "R OO (=) = R" by blast subsubsection \Converse\ inductive_set converse :: "('a \ 'b) set \ ('b \ 'a) set" ("(_\)" [1000] 999) for r :: "('a \ 'b) set" where "(a, b) \ r \ (b, a) \ r\" notation conversep ("(_\\)" [1000] 1000) notation (ASCII) converse ("(_^-1)" [1000] 999) and conversep ("(_^--1)" [1000] 1000) lemma converseI [sym]: "(a, b) \ r \ (b, a) \ r\" by (fact converse.intros) lemma conversepI (* CANDIDATE [sym] *): "r a b \ r\\ b a" by (fact conversep.intros) lemma converseD [sym]: "(a, b) \ r\ \ (b, a) \ r" by (erule converse.cases) iprover lemma conversepD (* CANDIDATE [sym] *): "r\\ b a \ r a b" by (fact converseD [to_pred]) lemma converseE [elim!]: "yx \ r\ \ (\x y. yx = (y, x) \ (x, y) \ r \ P) \ P" \ \More general than \converseD\, as it ``splits'' the member of the relation.\ apply (cases yx) apply simp apply (erule converse.cases) apply iprover done lemmas conversepE [elim!] = conversep.cases lemma converse_iff [iff]: "(a, b) \ r\ \ (b, a) \ r" by (auto intro: converseI) lemma conversep_iff [iff]: "r\\ a b = r b a" by (fact converse_iff [to_pred]) lemma converse_converse [simp]: "(r\)\ = r" by (simp add: set_eq_iff) lemma conversep_conversep [simp]: "(r\\)\\ = r" by (fact converse_converse [to_pred]) lemma converse_empty[simp]: "{}\ = {}" by auto lemma converse_UNIV[simp]: "UNIV\ = UNIV" by auto lemma converse_relcomp: "(r O s)\ = s\ O r\" by blast lemma converse_relcompp: "(r OO s)\\ = s\\ OO r\\" by (iprover intro: order_antisym conversepI relcomppI elim: relcomppE dest: conversepD) lemma converse_Int: "(r \ s)\ = r\ \ s\" by blast lemma converse_meet: "(r \ s)\\ = r\\ \ s\\" by (simp add: inf_fun_def) (iprover intro: conversepI ext dest: conversepD) lemma converse_Un: "(r \ s)\ = r\ \ s\" by blast lemma converse_join: "(r \ s)\\ = r\\ \ s\\" by (simp add: sup_fun_def) (iprover intro: conversepI ext dest: conversepD) lemma converse_INTER: "(\(r ` S))\ = (\x\S. (r x)\)" by fast lemma converse_UNION: "(\(r ` S))\ = (\x\S. (r x)\)" by blast lemma converse_mono[simp]: "r\ \ s \ \ r \ s" by auto lemma conversep_mono[simp]: "r\\ \ s \\ \ r \ s" by (fact converse_mono[to_pred]) lemma converse_inject[simp]: "r\ = s \ \ r = s" by auto lemma conversep_inject[simp]: "r\\ = s \\ \ r = s" by (fact converse_inject[to_pred]) lemma converse_subset_swap: "r \ s \ \ r \ \ s" by auto lemma conversep_le_swap: "r \ s \\ \ r \\ \ s" by (fact converse_subset_swap[to_pred]) lemma converse_Id [simp]: "Id\ = Id" by blast lemma converse_Id_on [simp]: "(Id_on A)\ = Id_on A" by blast lemma refl_on_converse [simp]: "refl_on A (r\) = refl_on A r" by (auto simp: refl_on_def) lemma reflp_on_conversp [simp]: "reflp_on A R\\ \ reflp_on A R" by (auto simp: reflp_on_def) lemma irrefl_on_converse [simp]: "irrefl_on A (r\) = irrefl_on A r" by (simp add: irrefl_on_def) lemma irreflp_on_converse [simp]: "irreflp_on A (r\\) = irreflp_on A r" by (rule irrefl_on_converse[to_pred]) lemma sym_converse [simp]: "sym (converse r) = sym r" unfolding sym_def by blast lemma antisym_converse [simp]: "antisym (converse r) = antisym r" unfolding antisym_def by blast lemma trans_converse [simp]: "trans (converse r) = trans r" unfolding trans_def by blast lemma sym_conv_converse_eq: "sym r \ r\ = r" unfolding sym_def by fast lemma sym_Un_converse: "sym (r \ r\)" unfolding sym_def by blast lemma sym_Int_converse: "sym (r \ r\)" unfolding sym_def by blast lemma total_on_converse [simp]: "total_on A (r\) = total_on A r" by (auto simp: total_on_def) lemma totalp_on_converse [simp]: "totalp_on A R\\ = totalp_on A R" by (rule total_on_converse[to_pred]) lemma finite_converse [iff]: "finite (r\) = finite r" unfolding converse_def conversep_iff using [[simproc add: finite_Collect]] by (auto elim: finite_imageD simp: inj_on_def) lemma card_inverse[simp]: "card (R\) = card R" proof - have *: "\R. prod.swap ` R = R\" by auto { assume "\finite R" hence ?thesis by auto } moreover { assume "finite R" with card_image_le[of R prod.swap] card_image_le[of "R\" prod.swap] have ?thesis by (auto simp: *) } ultimately show ?thesis by blast qed lemma conversep_noteq [simp]: "(\)\\ = (\)" by (auto simp add: fun_eq_iff) lemma conversep_eq [simp]: "(=)\\ = (=)" by (auto simp add: fun_eq_iff) lemma converse_unfold [code]: "r\ = {(y, x). (x, y) \ r}" by (simp add: set_eq_iff) subsubsection \Domain, range and field\ inductive_set Domain :: "('a \ 'b) set \ 'a set" for r :: "('a \ 'b) set" where DomainI [intro]: "(a, b) \ r \ a \ Domain r" lemmas DomainPI = Domainp.DomainI inductive_cases DomainE [elim!]: "a \ Domain r" inductive_cases DomainpE [elim!]: "Domainp r a" inductive_set Range :: "('a \ 'b) set \ 'b set" for r :: "('a \ 'b) set" where RangeI [intro]: "(a, b) \ r \ b \ Range r" lemmas RangePI = Rangep.RangeI inductive_cases RangeE [elim!]: "b \ Range r" inductive_cases RangepE [elim!]: "Rangep r b" definition Field :: "'a rel \ 'a set" where "Field r = Domain r \ Range r" lemma FieldI1: "(i, j) \ R \ i \ Field R" unfolding Field_def by blast lemma FieldI2: "(i, j) \ R \ j \ Field R" unfolding Field_def by auto lemma Domain_fst [code]: "Domain r = fst ` r" by force lemma Range_snd [code]: "Range r = snd ` r" by force lemma fst_eq_Domain: "fst ` R = Domain R" by force lemma snd_eq_Range: "snd ` R = Range R" by force lemma range_fst [simp]: "range fst = UNIV" by (auto simp: fst_eq_Domain) lemma range_snd [simp]: "range snd = UNIV" by (auto simp: snd_eq_Range) lemma Domain_empty [simp]: "Domain {} = {}" by auto lemma Range_empty [simp]: "Range {} = {}" by auto lemma Field_empty [simp]: "Field {} = {}" by (simp add: Field_def) lemma Domain_empty_iff: "Domain r = {} \ r = {}" by auto lemma Range_empty_iff: "Range r = {} \ r = {}" by auto lemma Domain_insert [simp]: "Domain (insert (a, b) r) = insert a (Domain r)" by blast lemma Range_insert [simp]: "Range (insert (a, b) r) = insert b (Range r)" by blast lemma Field_insert [simp]: "Field (insert (a, b) r) = {a, b} \ Field r" by (auto simp add: Field_def) lemma Domain_iff: "a \ Domain r \ (\y. (a, y) \ r)" by blast lemma Range_iff: "a \ Range r \ (\y. (y, a) \ r)" by blast lemma Domain_Id [simp]: "Domain Id = UNIV" by blast lemma Range_Id [simp]: "Range Id = UNIV" by blast lemma Domain_Id_on [simp]: "Domain (Id_on A) = A" by blast lemma Range_Id_on [simp]: "Range (Id_on A) = A" by blast lemma Domain_Un_eq: "Domain (A \ B) = Domain A \ Domain B" by blast lemma Range_Un_eq: "Range (A \ B) = Range A \ Range B" by blast lemma Field_Un [simp]: "Field (r \ s) = Field r \ Field s" by (auto simp: Field_def) lemma Domain_Int_subset: "Domain (A \ B) \ Domain A \ Domain B" by blast lemma Range_Int_subset: "Range (A \ B) \ Range A \ Range B" by blast lemma Domain_Diff_subset: "Domain A - Domain B \ Domain (A - B)" by blast lemma Range_Diff_subset: "Range A - Range B \ Range (A - B)" by blast lemma Domain_Union: "Domain (\S) = (\A\S. Domain A)" by blast lemma Range_Union: "Range (\S) = (\A\S. Range A)" by blast lemma Field_Union [simp]: "Field (\R) = \(Field ` R)" by (auto simp: Field_def) lemma Domain_converse [simp]: "Domain (r\) = Range r" by auto lemma Range_converse [simp]: "Range (r\) = Domain r" by blast lemma Field_converse [simp]: "Field (r\) = Field r" by (auto simp: Field_def) lemma Domain_Collect_case_prod [simp]: "Domain {(x, y). P x y} = {x. \y. P x y}" by auto lemma Range_Collect_case_prod [simp]: "Range {(x, y). P x y} = {y. \x. P x y}" by auto lemma finite_Domain: "finite r \ finite (Domain r)" by (induct set: finite) auto lemma finite_Range: "finite r \ finite (Range r)" by (induct set: finite) auto lemma finite_Field: "finite r \ finite (Field r)" by (simp add: Field_def finite_Domain finite_Range) lemma Domain_mono: "r \ s \ Domain r \ Domain s" by blast lemma Range_mono: "r \ s \ Range r \ Range s" by blast lemma mono_Field: "r \ s \ Field r \ Field s" by (auto simp: Field_def Domain_def Range_def) lemma Domain_unfold: "Domain r = {x. \y. (x, y) \ r}" by blast lemma Field_square [simp]: "Field (x \ x) = x" unfolding Field_def by blast subsubsection \Image of a set under a relation\ definition Image :: "('a \ 'b) set \ 'a set \ 'b set" (infixr "``" 90) where "r `` s = {y. \x\s. (x, y) \ r}" lemma Image_iff: "b \ r``A \ (\x\A. (x, b) \ r)" by (simp add: Image_def) lemma Image_singleton: "r``{a} = {b. (a, b) \ r}" by (simp add: Image_def) lemma Image_singleton_iff [iff]: "b \ r``{a} \ (a, b) \ r" by (rule Image_iff [THEN trans]) simp lemma ImageI [intro]: "(a, b) \ r \ a \ A \ b \ r``A" unfolding Image_def by blast lemma ImageE [elim!]: "b \ r `` A \ (\x. (x, b) \ r \ x \ A \ P) \ P" unfolding Image_def by (iprover elim!: CollectE bexE) lemma rev_ImageI: "a \ A \ (a, b) \ r \ b \ r `` A" \ \This version's more effective when we already have the required \a\\ by blast lemma Image_empty1 [simp]: "{} `` X = {}" by auto lemma Image_empty2 [simp]: "R``{} = {}" by blast lemma Image_Id [simp]: "Id `` A = A" by blast lemma Image_Id_on [simp]: "Id_on A `` B = A \ B" by blast lemma Image_Int_subset: "R `` (A \ B) \ R `` A \ R `` B" by blast lemma Image_Int_eq: "single_valued (converse R) \ R `` (A \ B) = R `` A \ R `` B" by (auto simp: single_valued_def) lemma Image_Un: "R `` (A \ B) = R `` A \ R `` B" by blast lemma Un_Image: "(R \ S) `` A = R `` A \ S `` A" by blast lemma Image_subset: "r \ A \ B \ r``C \ B" by (iprover intro!: subsetI elim!: ImageE dest!: subsetD SigmaD2) lemma Image_eq_UN: "r``B = (\y\ B. r``{y})" \ \NOT suitable for rewriting\ by blast lemma Image_mono: "r' \ r \ A' \ A \ (r' `` A') \ (r `` A)" by blast lemma Image_UN: "r `` (\(B ` A)) = (\x\A. r `` (B x))" by blast lemma UN_Image: "(\i\I. X i) `` S = (\i\I. X i `` S)" by auto lemma Image_INT_subset: "(r `` (\(B ` A))) \ (\x\A. r `` (B x))" by blast text \Converse inclusion requires some assumptions\ lemma Image_INT_eq: assumes "single_valued (r\)" and "A \ {}" shows "r `` (\(B ` A)) = (\x\A. r `` B x)" proof(rule equalityI, rule Image_INT_subset) show "(\x\A. r `` B x) \ r `` \ (B ` A)" proof fix x assume "x \ (\x\A. r `` B x)" then show "x \ r `` \ (B ` A)" using assms unfolding single_valued_def by simp blast qed qed lemma Image_subset_eq: "r``A \ B \ A \ - ((r\) `` (- B))" by blast lemma Image_Collect_case_prod [simp]: "{(x, y). P x y} `` A = {y. \x\A. P x y}" by auto lemma Sigma_Image: "(SIGMA x:A. B x) `` X = (\x\X \ A. B x)" by auto lemma relcomp_Image: "(X O Y) `` Z = Y `` (X `` Z)" by auto lemma finite_Image[simp]: assumes "finite R" shows "finite (R `` A)" by(rule finite_subset[OF _ finite_Range[OF assms]]) auto subsubsection \Inverse image\ definition inv_image :: "'b rel \ ('a \ 'b) \ 'a rel" where "inv_image r f = {(x, y). (f x, f y) \ r}" definition inv_imagep :: "('b \ 'b \ bool) \ ('a \ 'b) \ 'a \ 'a \ bool" where "inv_imagep r f = (\x y. r (f x) (f y))" lemma [pred_set_conv]: "inv_imagep (\x y. (x, y) \ r) f = (\x y. (x, y) \ inv_image r f)" by (simp add: inv_image_def inv_imagep_def) lemma sym_inv_image: "sym r \ sym (inv_image r f)" unfolding sym_def inv_image_def by blast lemma trans_inv_image: "trans r \ trans (inv_image r f)" unfolding trans_def inv_image_def by (simp (no_asm)) blast lemma total_inv_image: "\inj f; total r\ \ total (inv_image r f)" unfolding inv_image_def total_on_def by (auto simp: inj_eq) lemma asym_inv_image: "asym R \ asym (inv_image R f)" by (simp add: inv_image_def asym_iff) lemma in_inv_image[simp]: "(x, y) \ inv_image r f \ (f x, f y) \ r" by (auto simp: inv_image_def) lemma converse_inv_image[simp]: "(inv_image R f)\ = inv_image (R\) f" unfolding inv_image_def converse_unfold by auto lemma in_inv_imagep [simp]: "inv_imagep r f x y = r (f x) (f y)" by (simp add: inv_imagep_def) subsubsection \Powerset\ definition Powp :: "('a \ bool) \ 'a set \ bool" where "Powp A = (\B. \x \ B. A x)" lemma Powp_Pow_eq [pred_set_conv]: "Powp (\x. x \ A) = (\x. x \ Pow A)" by (auto simp add: Powp_def fun_eq_iff) lemmas Powp_mono [mono] = Pow_mono [to_pred] subsubsection \Expressing relation operations via \<^const>\Finite_Set.fold\\ lemma Id_on_fold: assumes "finite A" shows "Id_on A = Finite_Set.fold (\x. Set.insert (Pair x x)) {} A" proof - interpret comp_fun_commute "\x. Set.insert (Pair x x)" by standard auto from assms show ?thesis unfolding Id_on_def by (induct A) simp_all qed lemma comp_fun_commute_Image_fold: "comp_fun_commute (\(x,y) A. if x \ S then Set.insert y A else A)" proof - interpret comp_fun_idem Set.insert by (fact comp_fun_idem_insert) show ?thesis by standard (auto simp: fun_eq_iff comp_fun_commute split: prod.split) qed lemma Image_fold: assumes "finite R" shows "R `` S = Finite_Set.fold (\(x,y) A. if x \ S then Set.insert y A else A) {} R" proof - interpret comp_fun_commute "(\(x,y) A. if x \ S then Set.insert y A else A)" by (rule comp_fun_commute_Image_fold) have *: "\x F. Set.insert x F `` S = (if fst x \ S then Set.insert (snd x) (F `` S) else (F `` S))" by (force intro: rev_ImageI) show ?thesis using assms by (induct R) (auto simp: *) qed lemma insert_relcomp_union_fold: assumes "finite S" shows "{x} O S \ X = Finite_Set.fold (\(w,z) A'. if snd x = w then Set.insert (fst x,z) A' else A') X S" proof - interpret comp_fun_commute "\(w,z) A'. if snd x = w then Set.insert (fst x,z) A' else A'" proof - interpret comp_fun_idem Set.insert by (fact comp_fun_idem_insert) show "comp_fun_commute (\(w,z) A'. if snd x = w then Set.insert (fst x,z) A' else A')" by standard (auto simp add: fun_eq_iff split: prod.split) qed have *: "{x} O S = {(x', z). x' = fst x \ (snd x, z) \ S}" by (auto simp: relcomp_unfold intro!: exI) show ?thesis unfolding * using \finite S\ by (induct S) (auto split: prod.split) qed lemma insert_relcomp_fold: assumes "finite S" shows "Set.insert x R O S = Finite_Set.fold (\(w,z) A'. if snd x = w then Set.insert (fst x,z) A' else A') (R O S) S" proof - have "Set.insert x R O S = ({x} O S) \ (R O S)" by auto then show ?thesis by (auto simp: insert_relcomp_union_fold [OF assms]) qed lemma comp_fun_commute_relcomp_fold: assumes "finite S" shows "comp_fun_commute (\(x,y) A. Finite_Set.fold (\(w,z) A'. if y = w then Set.insert (x,z) A' else A') A S)" proof - have *: "\a b A. Finite_Set.fold (\(w, z) A'. if b = w then Set.insert (a, z) A' else A') A S = {(a,b)} O S \ A" by (auto simp: insert_relcomp_union_fold[OF assms] cong: if_cong) show ?thesis by standard (auto simp: *) qed lemma relcomp_fold: assumes "finite R" "finite S" shows "R O S = Finite_Set.fold (\(x,y) A. Finite_Set.fold (\(w,z) A'. if y = w then Set.insert (x,z) A' else A') A S) {} R" proof - interpret commute_relcomp_fold: comp_fun_commute "(\(x, y) A. Finite_Set.fold (\(w, z) A'. if y = w then insert (x, z) A' else A') A S)" by (fact comp_fun_commute_relcomp_fold[OF \finite S\]) from assms show ?thesis by (induct R) (auto simp: comp_fun_commute_relcomp_fold insert_relcomp_fold cong: if_cong) qed end diff --git a/src/HOL/Wellfounded.thy b/src/HOL/Wellfounded.thy --- a/src/HOL/Wellfounded.thy +++ b/src/HOL/Wellfounded.thy @@ -1,1020 +1,1020 @@ (* Title: HOL/Wellfounded.thy Author: Tobias Nipkow Author: Lawrence C Paulson Author: Konrad Slind Author: Alexander Krauss Author: Andrei Popescu, TU Muenchen *) section \Well-founded Recursion\ theory Wellfounded imports Transitive_Closure begin 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 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. (\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 \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: 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] lemmas wf_induct_rule = wf_induct [rule_format, consumes 1, case_names less, induct set: wf] lemmas wfP_induct_rule = wf_induct_rule [to_pred, induct set: wfP] lemma wf_not_sym: "wf r \ (a, x) \ r \ (x, a) \ r" by (induct a arbitrary: x set: wf) blast lemma wf_asym: assumes "wf r" "(a, x) \ r" obtains "(x, a) \ r" by (drule wf_not_sym[OF assms]) lemma wf_imp_asym: "wf r \ asym r" by (auto intro: asymI elim: wf_asym) lemma wfP_imp_asymp: "wfP r \ asymp r" by (rule wf_imp_asym[to_pred]) 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]) lemma wf_imp_irrefl: assumes "wf r" shows "irrefl r" using wf_irrefl [OF assms] by (auto simp add: irrefl_def) lemma wfP_imp_irreflp: "wfP r \ irreflp r" - by (rule wf_imp_irrefl[to_pred, folded top_set_def]) + by (rule wf_imp_irrefl[to_pred]) lemma wf_wellorderI: assumes wf: "wf {(x::'a::ord, y). x < y}" and lin: "OFCLASS('a::ord, linorder_class)" shows "OFCLASS('a::ord, wellorder_class)" apply (rule wellorder_class.intro [OF lin]) apply (simp add: wellorder_class.intro class.wellorder_axioms.intro wf_induct_rule [OF wf]) done lemma (in wellorder) wf: "wf {(x, y). x < y}" unfolding wf_def by (blast intro: less_induct) lemma (in wellorder) wfP_less[simp]: "wfP (<)" by (simp add: wf wfP_def) subsection \Basic Results\ text \Point-free characterization of well-foundedness\ lemma wfE_pf: assumes wf: "wf R" and a: "A \ R `` A" shows "A = {}" proof - 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: assumes a: "\A. A \ R `` A \ A = {}" shows "wf R" proof (rule wfUNIVI) fix P :: "'a \ bool" and x let ?A = "{x. \ P x}" assume "\x. (\y. (y, x) \ R \ P y) \ P x" then have "?A \ R `` ?A" by blast with a show "P x" by blast qed subsubsection \Minimal-element characterization of well-foundedness\ lemma wfE_min: assumes wf: "wf R" and Q: "x \ Q" obtains z where "z \ Q" "\y. (y, z) \ R \ y \ Q" using Q wfE_pf[OF wf, of Q] by blast lemma wfE_min': "wf R \ Q \ {} \ (\z. z \ Q \ (\y. (y, z) \ R \ y \ Q) \ thesis) \ thesis" using wfE_min[of R _ Q] by blast lemma wfI_min: 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" 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))" apply (rule iffI) apply (blast intro: elim!: wfE_min) by (rule wfI_min) auto lemmas wfP_eq_minimal = wf_eq_minimal [to_pred] subsubsection \Well-foundedness of transitive closure\ lemma wf_trancl: assumes "wf r" shows "wf (r\<^sup>+)" proof - 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 qed then show ?thesis unfolding wf_def by blast qed lemmas wfP_trancl = wf_trancl [to_pred] 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" by (simp add: wf_eq_minimal) fast lemmas wfP_subset = wf_subset [to_pred] text \Well-foundedness of the empty relation\ lemma wf_empty [iff]: "wf {}" by (simp add: wf_def) lemma wfP_empty [iff]: "wfP (\x y. False)" proof - have "wfP bot" by (fact wf_empty[to_pred bot_empty_eq2]) then show ?thesis by (simp add: bot_fun_def) qed lemma wf_Int1: "wf r \ wf (r \ r')" by (erule wf_subset) (rule Int_lower1) lemma wf_Int2: "wf r \ wf (r' \ r)" by (erule wf_subset) (rule Int_lower2) text \Exponentiation.\ lemma wf_exp: assumes "wf (R ^^ n)" shows "wf R" proof (rule wfI_pf) fix A assume "A \ R `` A" then have "A \ (R ^^ n) `` A" by (induct n) force+ with \wf (R ^^ n)\ show "A = {}" by (rule wfE_pf) qed text \Well-foundedness of \insert\.\ lemma wf_insert [iff]: "wf (insert (y,x) r) \ wf r \ (x,y) \ r\<^sup>*" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (blast elim: wf_trancl [THEN wf_irrefl] intro: rtrancl_into_trancl1 wf_subset rtrancl_mono [THEN subsetD]) next assume R: ?rhs then have R': "Q \ {} \ (\z\Q. \y. (y, z) \ r \ y \ Q)" for Q by (auto simp: wf_eq_minimal) show ?lhs unfolding wf_eq_minimal proof clarify fix Q :: "'a set" and q assume "q \ Q" then obtain a where "a \ Q" and a: "\y. (y, a) \ r \ y \ Q" using R by (auto simp: wf_eq_minimal) show "\z\Q. \y'. (y', z) \ insert (y, x) r \ y' \ Q" proof (cases "a=x") case True show ?thesis proof (cases "y \ Q") case True then obtain z where "z \ Q" "(z, y) \ r\<^sup>*" "\z'. (z', z) \ r \ z' \ Q \ (z', y) \ r\<^sup>*" using R' [of "{z \ Q. (z,y) \ r\<^sup>*}"] by auto then have "\y'. (y', z) \ insert (y, x) r \ y' \ Q" using R by(blast intro: rtrancl_trans)+ then show ?thesis by (rule bexI) fact next case False then show ?thesis using a \a \ Q\ by blast qed next case False with a \a \ Q\ show ?thesis by blast qed qed qed subsubsection \Well-foundedness of image\ lemma wf_map_prod_image_Dom_Ran: fixes r:: "('a \ 'a) set" and f:: "'a \ 'b" assumes wf_r: "wf r" and inj: "\ a a'. a \ Domain r \ a' \ Range r \ f a = f a' \ a = a'" shows "wf (map_prod f f ` r)" proof (unfold wf_eq_minimal, clarify) fix B :: "'b set" and b::"'b" assume "b \ B" define A where "A = f -` B \ Domain r" show "\z\B. \y. (y, z) \ map_prod f f ` r \ y \ B" proof (cases "A = {}") case False then obtain a0 where "a0 \ A" and "\a. (a, a0) \ r \ a \ A" using wfE_min[OF wf_r] by auto thus ?thesis using inj unfolding A_def by (intro bexI[of _ "f a0"]) auto qed (use \b \ B\ in \unfold A_def, auto\) qed lemma wf_map_prod_image: "wf r \ inj f \ wf (map_prod f f ` r)" by(rule wf_map_prod_image_Dom_Ran) (auto dest: inj_onD) subsection \Well-Foundedness Results for Unions\ lemma wf_union_compatible: assumes "wf R" "wf S" assumes "R O S \ R" shows "wf (R \ S)" proof (rule wfI_min) 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) have "y \ Q" if "(y, z) \ S" for y proof from that have "y \ ?Q'" by (rule zmin) assume "y \ Q" 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 \w \ Q\ show False by contradiction qed with \z \ ?Q'\ show "\z\Q. \y. (y, z) \ R \ S \ y \ Q" by blast qed text \Well-foundedness of indexed union with disjoint domains and ranges.\ lemma wf_UN: assumes r: "\i. i \ I \ wf (r i)" and disj: "\i j. \i \ I; j \ I; r i \ r j\ \ Domain (r i) \ Range (r j) = {}" shows "wf (\i\I. r i)" unfolding wf_eq_minimal proof clarify fix A and a :: "'b" assume "a \ A" show "\z\A. \y. (y, z) \ \(r ` I) \ y \ A" proof (cases "\i\I. \a\A. \b\A. (b, a) \ r i") case True then obtain i b c where ibc: "i \ I" "b \ A" "c \ A" "(c,b) \ r i" by blast have ri: "\Q. Q \ {} \ \z\Q. \y. (y, z) \ r i \ y \ Q" using r [OF \i \ I\] unfolding wf_eq_minimal by auto show ?thesis using ri [of "{a. a \ A \ (\b\A. (b, a) \ r i) }"] ibc disj by blast next case False with \a \ A\ show ?thesis by blast qed qed lemma wfP_SUP: "\i. wfP (r i) \ \i j. r i \ r j \ inf (Domainp (r i)) (Rangep (r j)) = bot \ wfP (\(range r))" by (rule wf_UN[to_pred]) simp_all 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 text \ Intuition: We find an \R \ S\-min element of a nonempty subset \A\ by case distinction. \<^enum> There is a step \a \R\ b\ with \a, b \ A\. Pick an \R\-min element \z\ of the (nonempty) set \{a\A | \b\A. a \R\ b}\. By definition, there is \z' \ A\ s.t. \z \R\ z'\. Because \z\ is \R\-min in the subset, \z'\ must be \R\-min in \A\. Because \z'\ has an \R\-predecessor, it cannot have an \S\-successor and is thus \S\-min in \A\ as well. \<^enum> There is no such step. 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 \ 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") proof assume "wf ?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 assume "wf ?B" show "wf ?A" proof (rule wfI_min) fix Q :: "'a set" and x assume "x \ Q" with \wf ?B\ obtain z where "z \ Q" and "\y. (y, z) \ ?B \ y \ Q" by (erule wfE_min) then have 1: "\y. (y, z) \ R O R \ y \ Q" and 2: "\y. (y, z) \ S O R \ y \ Q" and 3: "\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\ 3 show ?thesis by blast next 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" then have "(y, z) \ R O R" using \(z', z) \ R\ .. with 1 show "y \ Q" . next assume "(y, z') \ S" then have "(y, z) \ S O R" using \(z', z) \ R\ .. with 2 show "y \ Q" . qed qed with \z' \ Q\ show ?thesis .. qed qed qed lemma wf_comp_self: "wf R \ wf (R O R)" \ \special case\ by (rule wf_union_merge [where S = "{}", simplified]) subsection \Well-Foundedness of Composition\ text \Bachmair and Dershowitz 1986, Lemma 2. [Provided by Tjark Weber]\ lemma qc_wf_relto_iff: assumes "R O S \ (R \ S)\<^sup>* O R" \ \R quasi-commutes over S\ shows "wf (S\<^sup>* O R O S\<^sup>*) \ wf R" (is "wf ?S \ _") proof show "wf R" if "wf ?S" proof - have "R \ ?S" by auto with wf_subset [of ?S] that show "wf R" by auto qed next show "wf ?S" if "wf R" proof (rule wfI_pf) fix A assume A: "A \ ?S `` A" let ?X = "(R \ S)\<^sup>* `` A" have *: "R O (R \ S)\<^sup>* \ (R \ S)\<^sup>* O R" proof - have "(x, z) \ (R \ S)\<^sup>* O R" if "(y, z) \ (R \ S)\<^sup>*" and "(x, y) \ R" for x y z using that proof (induct y z) case rtrancl_refl then show ?case by auto next case (rtrancl_into_rtrancl a b c) then have "(x, c) \ ((R \ S)\<^sup>* O (R \ S)\<^sup>*) O R" using assms by blast then show ?case by simp qed then show ?thesis by auto qed then have "R O S\<^sup>* \ (R \ S)\<^sup>* O R" using rtrancl_Un_subset by blast then have "?S \ (R \ S)\<^sup>* O (R \ S)\<^sup>* O R" by (simp add: relcomp_mono rtrancl_mono) also have "\ = (R \ S)\<^sup>* O R" by (simp add: O_assoc[symmetric]) finally have "?S O (R \ S)\<^sup>* \ (R \ S)\<^sup>* O R O (R \ S)\<^sup>*" by (simp add: O_assoc[symmetric] relcomp_mono) also have "\ \ (R \ S)\<^sup>* O (R \ S)\<^sup>* O R" using * by (simp add: relcomp_mono) finally have "?S O (R \ S)\<^sup>* \ (R \ S)\<^sup>* O R" by (simp add: O_assoc[symmetric]) then have "(?S O (R \ S)\<^sup>*) `` A \ ((R \ S)\<^sup>* O R) `` A" by (simp add: Image_mono) moreover have "?X \ (?S O (R \ S)\<^sup>*) `` A" using A by (auto simp: relcomp_Image) ultimately have "?X \ R `` ?X" by (auto simp: relcomp_Image) then have "?X = {}" using \wf R\ by (simp add: wfE_pf) moreover have "A \ ?X" by auto ultimately show "A = {}" by simp qed qed corollary wf_relcomp_compatible: assumes "wf R" and "R O S \ S O R" shows "wf (S O R)" proof - have "R O S \ (R \ S)\<^sup>* O R" using assms by blast then have "wf (S\<^sup>* O R O S\<^sup>*)" by (simp add: assms qc_wf_relto_iff) then show ?thesis by (rule Wellfounded.wf_subset) blast qed subsection \Acyclic relations\ lemma wf_acyclic: "wf r \ acyclic r" by (simp add: acyclic_def) (blast elim: wf_trancl [THEN wf_irrefl]) lemmas wfP_acyclicP = wf_acyclic [to_pred] subsubsection \Wellfoundedness of finite acyclic relations\ lemma finite_acyclic_wf: assumes "finite r" "acyclic r" shows "wf r" using assms proof (induction r rule: finite_induct) case (insert x r) then show ?case by (cases x) simp qed simp 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 text \ Observe that the converse of an irreflexive, transitive, and finite relation is again well-founded. Thus, we may employ it for well-founded induction. \ lemma wf_converse: assumes "irrefl r" and "trans r" and "finite r" shows "wf (r\)" proof - have "acyclic r" using \irrefl r\ and \trans r\ by (simp add: irrefl_def acyclic_irrefl) with \finite r\ show ?thesis by (rule finite_acyclic_wf_converse) qed 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: "(<) = (\m n. n = Suc m)\<^sup>+\<^sup>+" proof (rule ext, rule ext, rule iffI) fix n m :: nat 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 next case (Suc n) then show ?case by (auto simp add: less_Suc_eq_le le_less intro: tranclp.trancl_into_trancl) qed 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 less_than :: "(nat \ nat) set" where "less_than = pred_nat\<^sup>+" 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\<^sup>* \ m \ n" unfolding less_eq rtrancl_eq_or_trancl by auto lemma wf_pred_nat: "wf pred_nat" unfolding wf_def proof clarify fix P x assume "\x'. (\y. (y, x') \ pred_nat \ P y) \ P x'" then show "P x" unfolding pred_nat_def by (induction x) blast+ qed lemma wf_less_than [iff]: "wf less_than" by (simp add: less_than_def wf_pred_nat [THEN wf_trancl]) lemma trans_less_than [iff]: "trans less_than" by (simp add: less_than_def) lemma less_than_iff [iff]: "((x,y) \ less_than) = (xAccessible Part\ text \ 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" abbreviation termip :: "('a \ 'a \ bool) \ 'a \ bool" where "termip r \ accp (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})" by (simp add: acc_def) text \Induction rules\ theorem accp_induct: assumes major: "accp r a" 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) apply (rule accp.accI) apply auto done lemmas accp_induct_rule = accp_induct [rule_format, induct set: accp] theorem accp_downward: "accp r b \ r a b \ accp r a" by (cases rule: accp.cases) lemma not_accp_down: assumes na: "\ accp R x" 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 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" by auto with a show thesis . qed qed lemma accp_downwards_aux: "r\<^sup>*\<^sup>* b a \ accp r a \ accp r b" by (erule rtranclp_induct) (blast dest: accp_downward)+ theorem accp_downwards: "accp r a \ r\<^sup>*\<^sup>* b a \ accp r b" by (blast dest: accp_downwards_aux) theorem accp_wfPI: "\x. accp r x \ wfP r" proof (rule wfPUNIVI) fix P x assume "\x. accp r x" "\x. (\y. r y x \ P y) \ P x" then show "P x" using accp_induct[where P = P] by blast qed theorem accp_wfPD: "wfP r \ accp r x" apply (erule wfP_induct_rule) apply (rule accp.accI) apply blast done theorem wfP_accp_iff: "wfP r = (\x. accp r x)" by (blast intro: accp_wfPI dest: accp_wfPD) text \Smaller relations have bigger accessible parts:\ lemma accp_subset: assumes "R1 \ R2" shows "accp R2 \ accp R1" proof (rule predicate1I) fix x assume "accp R2 x" then show "accp R1 x" proof (induct x) fix x assume "\y. R2 y x \ accp R1 y" with assms show "accp R1 x" by (blast intro: accp.accI) qed qed text \This is a generalized induction theorem that works on subsets of the accessible part.\ lemma accp_subset_induct: assumes subset: "D \ accp R" and dcl: "\x z. D x \ R z x \ D z" and "D x" and istep: "\x. D x \ (\z. R z x \ P z) \ P x" shows "P x" proof - from subset and \D x\ have "accp R x" .. then show "P x" using \D x\ proof (induct x) fix x assume "D x" and "\y. R y x \ D y \ P y" with dcl and istep show "P x" by blast qed qed text \Set versions of the above theorems\ lemmas acc_induct = accp_induct [to_set] lemmas acc_induct_rule = acc_induct [rule_format, induct set: acc] lemmas acc_downward = accp_downward [to_set] lemmas not_acc_down = not_accp_down [to_set] lemmas acc_downwards_aux = accp_downwards_aux [to_set] lemmas acc_downwards = accp_downwards [to_set] lemmas acc_wfI = accp_wfPI [to_set] lemmas acc_wfD = accp_wfPD [to_set] lemmas wf_acc_iff = wfP_accp_iff [to_set] lemmas acc_subset = accp_subset [to_set] lemmas acc_subset_induct = accp_subset_induct [to_set] subsection \Tools for building wellfounded relations\ text \Inverse Image\ lemma wf_inv_image [simp,intro!]: fixes f :: "'a \ 'b" assumes "wf r" shows "wf (inv_image r f)" proof - have "\x P. x \ P \ \z\P. \y. (f y, f z) \ r \ y \ P" proof - fix P and x::'a assume "x \ P" then obtain w where w: "w \ {w. \x::'a. x \ P \ f x = w}" by auto have *: "\Q u. u \ Q \ \z\Q. \y. (y, z) \ r \ y \ Q" using assms by (auto simp add: wf_eq_minimal) show "\z\P. \y. (f y, f z) \ r \ y \ P" using * [OF w] by auto qed then show ?thesis by (clarsimp simp: inv_image_def wf_eq_minimal) qed subsubsection \Conversion to a known well-founded relation\ lemma wf_if_convertible_to_wf: fixes r :: "'a rel" and s :: "'b rel" and f :: "'a \ 'b" assumes "wf s" and convertible: "\x y. (x, y) \ r \ (f x, f y) \ s" shows "wf r" proof (rule wfI_min[of r]) fix x :: 'a and Q :: "'a set" assume "x \ Q" then obtain y where "y \ Q" and "\z. (f z, f y) \ s \ z \ Q" by (auto elim: wfE_min[OF wf_inv_image[of s f, OF \wf s\], unfolded in_inv_image]) thus "\z \ Q. \y. (y, z) \ r \ y \ Q" by (auto intro: convertible) qed lemma wfP_if_convertible_to_wfP: "wfP S \ (\x y. R x y \ S (f x) (f y)) \ wfP R" using wf_if_convertible_to_wf[to_pred, of S R f] by simp text \Converting to @{typ nat} is a very common special case that might be found more easily by Sledgehammer.\ lemma wfP_if_convertible_to_nat: fixes f :: "_ \ nat" shows "(\x y. R x y \ f x < f y) \ wfP R" by (rule wfP_if_convertible_to_wfP[of "(<) :: nat \ nat \ bool", simplified]) subsubsection \Measure functions into \<^typ>\nat\\ 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" by (simp add:measure_def) lemma wf_measure [iff]: "wf (measure f)" unfolding measure_def by (rule wf_less_than [THEN wf_inv_image]) lemma wf_if_measure: "(\x. P x \ f(g x) < f x) \ wf {(y,x). P x \ y = g x}" for f :: "'a \ nat" using wf_measure[of f] unfolding measure_def inv_image_def less_than_def less_eq by (rule wf_subset) auto 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}" 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) lemma wf_lex_prod [intro!]: assumes "wf ra" "wf rb" shows "wf (ra <*lex*> rb)" proof (rule wfI) fix z :: "'a \ 'b" and P assume * [rule_format]: "\u. (\v. (v, u) \ ra <*lex*> rb \ P v) \ P u" obtain x y where zeq: "z = (x,y)" by fastforce have "P(x,y)" using \wf ra\ proof (induction x arbitrary: y rule: wf_induct_rule) case (less x) note lessx = less show ?case using \wf rb\ less proof (induction y rule: wf_induct_rule) case (less y) show ?case by (force intro: * less.IH lessx) qed qed then show "P z" by (simp add: zeq) qed auto text \\<*lex*>\ preserves transitivity\ lemma trans_lex_prod [simp,intro!]: "trans R1 \ trans R2 \ trans (R1 <*lex*> R2)" unfolding trans_def lex_prod_def by blast lemma total_on_lex_prod [simp]: "total_on A r \ total_on B s \ total_on (A \ B) (r <*lex*> s)" by (auto simp: total_on_def) lemma asym_lex_prod: "\asym R; asym S\ \ asym (R <*lex*> S)" by (auto simp add: asym_iff lex_prod_def) lemma total_lex_prod [simp]: "total r \ total s \ total (r <*lex*> s)" by (auto simp: total_on_def) 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))" lemma wf_mlex: "wf R \ wf (f <*mlex*> R)" and mlex_less: "f x < f y \ (x, y) \ f <*mlex*> R" and mlex_leq: "f x \ f y \ (x, y) \ R \ (x, y) \ f <*mlex*> R" and mlex_iff: "(x, y) \ f <*mlex*> R \ f x < f y \ f x = f y \ (x, y) \ R" by (auto simp: mlex_prod_def) text \Proper subset relation on finite sets.\ 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 trans_finite_psubset: "trans finite_psubset" by (auto simp: 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 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" where max_extI[intro]: "finite X \ finite Y \ Y \ {} \ (\x. x \ X \ \y\Y. (x, y) \ R) \ (X, Y) \ max_ext R" lemma max_ext_wf: assumes wf: "wf r" shows "wf (max_ext r)" proof (rule acc_wfI, intro allI) show "M \ acc (max_ext r)" (is "_ \ ?W") for M proof (induct M rule: infinite_finite_induct) case empty show ?case by (rule accI) (auto elim: max_ext.cases) next case (insert a M) from wf \M \ ?W\ \finite M\ show "insert a M \ ?W" proof (induct arbitrary: M) fix M a assume "M \ ?W" assume [intro]: "finite M" assume hyp: "\b M. (b, a) \ r \ M \ ?W \ finite M \ insert b M \ ?W" 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" 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}" 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 "M = {}") case [simp]: True have Mw: "{} \ ?W" by (rule accI) (auto elim: max_ext.cases) from * have "?N2 = {}" by auto with Mw show "?N2 \ ?W" by (simp only:) next case False from * finites have N2: "(?N2, M) \ max_ext r" using max_extI[OF _ _ \M \ {}\, where ?X = ?N2] by auto with \M \ ?W\ show "?N2 \ ?W" by (rule acc_downward) qed with finites have "?N1 \ ?N2 \ ?W" by (rule add_less) simp then show "N \ ?W" by (simp only: N) qed qed next case infinite show ?case by (rule accI) (auto elim: max_ext.cases simp: infinite) 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) 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) 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 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" by (erule wfE_min) from z obtain m where "m \ Q" "z \ m" by auto from \m \ Q\ show ?thesis proof (intro rev_bexI allI impI) fix n assume smaller: "(n, m) \ min_ext r" with \z \ m\ obtain y where "y \ n" "(y, z) \ r" by (auto simp: min_ext_def) with z(2) show "n \ Q" by auto qed qed qed 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" by (rule wf_subset[OF wf_measure[of "\a. ub a - f a"]]) (auto dest: assms) 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 lemma finite_subset_wf: assumes "finite A" shows "wf {(X, Y). X \ Y \ Y \ A}" by (rule wf_subset[OF wf_finite_psubset[unfolded finite_psubset_def]]) (auto intro: finite_subset[OF _ assms]) hide_const (open) acc accp end