diff -r a962f349c8c9 -r a95e7432d86c src/HOL/Transitive_Closure.thy --- a/src/HOL/Transitive_Closure.thy Wed Jul 06 14:09:13 2016 +0200 +++ b/src/HOL/Transitive_Closure.thy Wed Jul 06 20:19:51 2016 +0200 @@ -65,67 +65,65 @@ subsection \Reflexive closure\ -lemma refl_reflcl[simp]: "refl(r^=)" -by(simp add:refl_on_def) +lemma refl_reflcl[simp]: "refl (r\<^sup>=)" + by (simp add: refl_on_def) -lemma antisym_reflcl[simp]: "antisym(r^=) = antisym r" -by(simp add:antisym_def) +lemma antisym_reflcl[simp]: "antisym (r\<^sup>=) = antisym r" + by (simp add: antisym_def) -lemma trans_reflclI[simp]: "trans r \ trans(r^=)" -unfolding trans_def by blast +lemma trans_reflclI[simp]: "trans r \ trans (r\<^sup>=)" + unfolding trans_def by blast -lemma reflclp_idemp [simp]: "(P^==)^== = P^==" -by blast +lemma reflclp_idemp [simp]: "(P\<^sup>=\<^sup>=)\<^sup>=\<^sup>= = P\<^sup>=\<^sup>=" + by blast + subsection \Reflexive-transitive closure\ lemma reflcl_set_eq [pred_set_conv]: "(sup (\x y. (x, y) \ r) op =) = (\x y. (x, y) \ r \ Id)" by (auto simp add: fun_eq_iff) -lemma r_into_rtrancl [intro]: "!!p. p \ r ==> p \ r^*" +lemma r_into_rtrancl [intro]: "\p. p \ r \ p \ r\<^sup>*" \ \\rtrancl\ of \r\ contains \r\\ apply (simp only: split_tupled_all) apply (erule rtrancl_refl [THEN rtrancl_into_rtrancl]) done -lemma r_into_rtranclp [intro]: "r x y ==> r^** x y" +lemma r_into_rtranclp [intro]: "r x y \ r\<^sup>*\<^sup>* x y" \ \\rtrancl\ of \r\ contains \r\\ by (erule rtranclp.rtrancl_refl [THEN rtranclp.rtrancl_into_rtrancl]) -lemma rtranclp_mono: "r \ s ==> r^** \ s^**" +lemma rtranclp_mono: "r \ s \ r\<^sup>*\<^sup>* \ s\<^sup>*\<^sup>*" \ \monotonicity of \rtrancl\\ apply (rule predicate2I) apply (erule rtranclp.induct) apply (rule_tac [2] rtranclp.rtrancl_into_rtrancl, blast+) done -lemma mono_rtranclp[mono]: - "(\a b. x a b \ y a b) \ x^** a b \ y^** a b" +lemma mono_rtranclp[mono]: "(\a b. x a b \ y a b) \ x\<^sup>*\<^sup>* a b \ y\<^sup>*\<^sup>* a b" using rtranclp_mono[of x y] by auto lemmas rtrancl_mono = rtranclp_mono [to_set] theorem rtranclp_induct [consumes 1, case_names base step, induct set: rtranclp]: - assumes a: "r^** a b" - and cases: "P a" "!!y z. [| r^** a y; r y z; P y |] ==> P z" - shows "P b" using a - by (induct x\a b) (rule cases)+ + assumes a: "r\<^sup>*\<^sup>* a b" + and cases: "P a" "\y z. r\<^sup>*\<^sup>* a y \ r y z \ P y \ P z" + shows "P b" + using a by (induct x\a b) (rule cases)+ lemmas rtrancl_induct [induct set: rtrancl] = rtranclp_induct [to_set] lemmas rtranclp_induct2 = - rtranclp_induct[of _ "(ax,ay)" "(bx,by)", split_rule, - consumes 1, case_names refl step] + rtranclp_induct[of _ "(ax,ay)" "(bx,by)", split_rule, consumes 1, case_names refl step] lemmas rtrancl_induct2 = - rtrancl_induct[of "(ax,ay)" "(bx,by)", split_format (complete), - consumes 1, case_names refl step] + rtrancl_induct[of "(ax,ay)" "(bx,by)", split_format (complete), consumes 1, case_names refl step] -lemma refl_rtrancl: "refl (r^*)" -by (unfold refl_on_def) fast +lemma refl_rtrancl: "refl (r\<^sup>*)" + unfolding refl_on_def by fast text \Transitivity of transitive closure.\ -lemma trans_rtrancl: "trans (r^*)" +lemma trans_rtrancl: "trans (r\<^sup>*)" proof (rule transI) fix x y z assume "(x, y) \ r\<^sup>*" @@ -144,42 +142,40 @@ lemmas rtrancl_trans = trans_rtrancl [THEN transD] lemma rtranclp_trans: - assumes xy: "r^** x y" - and yz: "r^** y z" - shows "r^** x z" using yz xy - by induct iprover+ + assumes "r\<^sup>*\<^sup>* x y" + and "r\<^sup>*\<^sup>* y z" + shows "r\<^sup>*\<^sup>* x z" + using assms(2,1) by induct iprover+ lemma rtranclE [cases set: rtrancl]: - assumes major: "(a::'a, b) : r^*" + fixes a b :: 'a + assumes major: "(a, b) \ r\<^sup>*" obtains (base) "a = b" - | (step) y where "(a, y) : r^*" and "(y, b) : r" + | (step) y where "(a, y) \ r\<^sup>*" and "(y, b) \ r" \ \elimination of \rtrancl\ -- by induction on a special formula\ - apply (subgoal_tac "(a::'a) = b | (EX y. (a,y) : r^* & (y,b) : r)") + apply (subgoal_tac "a = b \ (\y. (a, y) \ r\<^sup>* \ (y, b) \ r)") apply (rule_tac [2] major [THEN rtrancl_induct]) prefer 2 apply blast prefer 2 apply blast apply (erule asm_rl exE disjE conjE base step)+ done -lemma rtrancl_Int_subset: "[| Id \ s; (r^* \ s) O r \ s|] ==> r^* \ s" +lemma rtrancl_Int_subset: "Id \ s \ (r\<^sup>* \ s) O r \ s \ r\<^sup>* \ s" apply (rule subsetI) apply auto apply (erule rtrancl_induct) apply auto done -lemma converse_rtranclp_into_rtranclp: - "r a b \ r\<^sup>*\<^sup>* b c \ r\<^sup>*\<^sup>* a c" +lemma converse_rtranclp_into_rtranclp: "r a b \ r\<^sup>*\<^sup>* b c \ r\<^sup>*\<^sup>* a c" by (rule rtranclp_trans) iprover+ lemmas converse_rtrancl_into_rtrancl = converse_rtranclp_into_rtranclp [to_set] -text \ - \medskip More @{term "r^*"} equations and inclusions. -\ +text \\<^medskip> More @{term "r\<^sup>*"} equations and inclusions.\ -lemma rtranclp_idemp [simp]: "(r^**)^** = r^**" +lemma rtranclp_idemp [simp]: "(r\<^sup>*\<^sup>*)\<^sup>*\<^sup>* = r\<^sup>*\<^sup>*" apply (auto intro!: order_antisym) apply (erule rtranclp_induct) apply (rule rtranclp.rtrancl_refl) @@ -188,18 +184,18 @@ lemmas rtrancl_idemp [simp] = rtranclp_idemp [to_set] -lemma rtrancl_idemp_self_comp [simp]: "R^* O R^* = R^*" +lemma rtrancl_idemp_self_comp [simp]: "R\<^sup>* O R\<^sup>* = R\<^sup>*" apply (rule set_eqI) apply (simp only: split_tupled_all) apply (blast intro: rtrancl_trans) done -lemma rtrancl_subset_rtrancl: "r \ s^* ==> r^* \ s^*" +lemma rtrancl_subset_rtrancl: "r \ s\<^sup>* \ r\<^sup>* \ s\<^sup>*" apply (drule rtrancl_mono) apply simp done -lemma rtranclp_subset: "R \ S ==> S \ R^** ==> S^** = R^**" +lemma rtranclp_subset: "R \ S \ S \ R\<^sup>*\<^sup>* \ S\<^sup>*\<^sup>* = R\<^sup>*\<^sup>*" apply (drule rtranclp_mono) apply (drule rtranclp_mono) apply simp @@ -207,17 +203,17 @@ lemmas rtrancl_subset = rtranclp_subset [to_set] -lemma rtranclp_sup_rtranclp: "(sup (R^**) (S^**))^** = (sup R S)^**" -by (blast intro!: rtranclp_subset intro: rtranclp_mono [THEN predicate2D]) +lemma rtranclp_sup_rtranclp: "(sup (R\<^sup>*\<^sup>*) (S\<^sup>*\<^sup>*))\<^sup>*\<^sup>* = (sup R S)\<^sup>*\<^sup>*" + by (blast intro!: rtranclp_subset intro: rtranclp_mono [THEN predicate2D]) lemmas rtrancl_Un_rtrancl = rtranclp_sup_rtranclp [to_set] -lemma rtranclp_reflclp [simp]: "(R^==)^** = R^**" -by (blast intro!: rtranclp_subset) +lemma rtranclp_reflclp [simp]: "(R\<^sup>=\<^sup>=)\<^sup>*\<^sup>* = R\<^sup>*\<^sup>*" + by (blast intro!: rtranclp_subset) lemmas rtrancl_reflcl [simp] = rtranclp_reflclp [to_set] -lemma rtrancl_r_diff_Id: "(r - Id)^* = r^*" +lemma rtrancl_r_diff_Id: "(r - Id)\<^sup>* = r\<^sup>*" apply (rule sym) apply (rule rtrancl_subset, blast, clarify) apply (rename_tac a b) @@ -226,39 +222,35 @@ apply blast done -lemma rtranclp_r_diff_Id: "(inf r op ~=)^** = r^**" +lemma rtranclp_r_diff_Id: "(inf r op \)\<^sup>*\<^sup>* = r\<^sup>*\<^sup>*" apply (rule sym) apply (rule rtranclp_subset) apply blast+ done theorem rtranclp_converseD: - assumes r: "(r^--1)^** x y" - shows "r^** y x" -proof - - from r show ?thesis - by induct (iprover intro: rtranclp_trans dest!: conversepD)+ -qed + assumes "(r\\)\<^sup>*\<^sup>* x y" + shows "r\<^sup>*\<^sup>* y x" + using assms by induct (iprover intro: rtranclp_trans dest!: conversepD)+ lemmas rtrancl_converseD = rtranclp_converseD [to_set] theorem rtranclp_converseI: - assumes "r^** y x" - shows "(r^--1)^** x y" - using assms - by induct (iprover intro: rtranclp_trans conversepI)+ + assumes "r\<^sup>*\<^sup>* y x" + shows "(r\\)\<^sup>*\<^sup>* x y" + using assms by induct (iprover intro: rtranclp_trans conversepI)+ lemmas rtrancl_converseI = rtranclp_converseI [to_set] -lemma rtrancl_converse: "(r^-1)^* = (r^*)^-1" +lemma rtrancl_converse: "(r^-1)\<^sup>* = (r\<^sup>*)^-1" by (fast dest!: rtrancl_converseD intro!: rtrancl_converseI) -lemma sym_rtrancl: "sym r ==> sym (r^*)" +lemma sym_rtrancl: "sym r \ sym (r\<^sup>*)" by (simp only: sym_conv_converse_eq rtrancl_converse [symmetric]) theorem converse_rtranclp_induct [consumes 1, case_names base step]: - assumes major: "r^** a b" - and cases: "P b" "!!y z. [| r y z; r^** z b; P z |] ==> P y" + assumes major: "r\<^sup>*\<^sup>* a b" + and cases: "P b" "\y z. r y z \ r\<^sup>*\<^sup>* z b \ P z \ P y" shows "P a" using rtranclp_converseI [OF major] by induct (iprover intro: cases dest!: conversepD rtranclp_converseD)+ @@ -266,19 +258,17 @@ lemmas converse_rtrancl_induct = converse_rtranclp_induct [to_set] lemmas converse_rtranclp_induct2 = - converse_rtranclp_induct [of _ "(ax,ay)" "(bx,by)", split_rule, - consumes 1, case_names refl step] + converse_rtranclp_induct [of _ "(ax,ay)" "(bx,by)", split_rule, consumes 1, case_names refl step] lemmas converse_rtrancl_induct2 = converse_rtrancl_induct [of "(ax,ay)" "(bx,by)", split_format (complete), - consumes 1, case_names refl step] + consumes 1, case_names refl step] lemma converse_rtranclpE [consumes 1, case_names base step]: - assumes major: "r^** x z" - and cases: "x=z ==> P" - "!!y. [| r x y; r^** y z |] ==> P" + assumes major: "r\<^sup>*\<^sup>* x z" + and cases: "x = z \ P" "\y. r x y \ r\<^sup>*\<^sup>* y z \ P" shows P - apply (subgoal_tac "x = z | (EX y. r x y & r^** y z)") + apply (subgoal_tac "x = z \ (\y. r x y \ r\<^sup>*\<^sup>* y z)") apply (rule_tac [2] major [THEN converse_rtranclp_induct]) prefer 2 apply iprover prefer 2 apply iprover @@ -291,41 +281,42 @@ lemmas converse_rtranclE2 = converse_rtranclE [of "(xa,xb)" "(za,zb)", split_rule] -lemma r_comp_rtrancl_eq: "r O r^* = r^* O r" +lemma r_comp_rtrancl_eq: "r O r\<^sup>* = r\<^sup>* O r" by (blast elim: rtranclE converse_rtranclE intro: rtrancl_into_rtrancl converse_rtrancl_into_rtrancl) -lemma rtrancl_unfold: "r^* = Id Un r^* O r" +lemma rtrancl_unfold: "r\<^sup>* = Id \ r\<^sup>* O r" by (auto intro: rtrancl_into_rtrancl elim: rtranclE) lemma rtrancl_Un_separatorE: - "(a,b) : (P \ Q)^* \ \x y. (a,x) : P^* \ (x,y) : Q \ x=y \ (a,b) : P^*" -apply (induct rule:rtrancl.induct) - apply blast -apply (blast intro:rtrancl_trans) -done + "(a, b) \ (P \ Q)\<^sup>* \ \x y. (a, x) \ P\<^sup>* \ (x, y) \ Q \ x = y \ (a, b) \ P\<^sup>*" + apply (induct rule:rtrancl.induct) + apply blast + apply (blast intro:rtrancl_trans) + done lemma rtrancl_Un_separator_converseE: - "(a,b) : (P \ Q)^* \ \x y. (x,b) : P^* \ (y,x) : Q \ y=x \ (a,b) : P^*" -apply (induct rule:converse_rtrancl_induct) - apply blast -apply (blast intro:rtrancl_trans) -done + "(a, b) \ (P \ Q)\<^sup>* \ \x y. (x, b) \ P\<^sup>* \ (y, x) \ Q \ y = x \ (a, b) \ P\<^sup>*" + apply (induct rule:converse_rtrancl_induct) + apply blast + apply (blast intro:rtrancl_trans) + done lemma Image_closed_trancl: - assumes "r `` X \ X" shows "r\<^sup>* `` X = X" + assumes "r `` X \ X" + shows "r\<^sup>* `` X = X" proof - - from assms have **: "{y. \x\X. (x, y) \ r} \ X" by auto - have "\x y. (y, x) \ r\<^sup>* \ y \ X \ x \ X" + from assms have **: "{y. \x\X. (x, y) \ r} \ X" + by auto + have "x \ X" if 1: "(y, x) \ r\<^sup>*" and 2: "y \ X" for x y proof - - fix x y - assume *: "y \ X" - assume "(y, x) \ r\<^sup>*" - then show "x \ X" + from 1 show "x \ X" proof induct - case base show ?case by (fact *) + case base + show ?case by (fact 2) next - case step with ** show ?case by auto + case step + with ** show ?case by auto qed qed then show ?thesis by auto @@ -334,31 +325,30 @@ subsection \Transitive closure\ -lemma trancl_mono: "!!p. p \ r^+ ==> r \ s ==> p \ s^+" +lemma trancl_mono: "\p. p \ r\<^sup>+ \ r \ s \ p \ s\<^sup>+" apply (simp add: split_tupled_all) apply (erule trancl.induct) apply (iprover dest: subsetD)+ done -lemma r_into_trancl': "!!p. p : r ==> p : r^+" +lemma r_into_trancl': "\p. p \ r \ p \ r\<^sup>+" by (simp only: split_tupled_all) (erule r_into_trancl) -text \ - \medskip Conversions between \trancl\ and \rtrancl\. -\ +text \\<^medskip> Conversions between \trancl\ and \rtrancl\.\ -lemma tranclp_into_rtranclp: "r^++ a b ==> r^** a b" +lemma tranclp_into_rtranclp: "r\<^sup>+\<^sup>+ a b \ r\<^sup>*\<^sup>* a b" by (erule tranclp.induct) iprover+ lemmas trancl_into_rtrancl = tranclp_into_rtranclp [to_set] -lemma rtranclp_into_tranclp1: assumes r: "r^** a b" - shows "!!c. r b c ==> r^++ a c" using r - by induct iprover+ +lemma rtranclp_into_tranclp1: + assumes "r\<^sup>*\<^sup>* a b" + shows "r b c \ r\<^sup>+\<^sup>+ a c" + using assms by (induct arbitrary: c) iprover+ lemmas rtrancl_into_trancl1 = rtranclp_into_tranclp1 [to_set] -lemma rtranclp_into_tranclp2: "[| r a b; r^** b c |] ==> r^++ a c" +lemma rtranclp_into_tranclp2: "r a b \ r\<^sup>*\<^sup>* b c \ r\<^sup>+\<^sup>+ a c" \ \intro rule from \r\ and \rtrancl\\ apply (erule rtranclp.cases) apply iprover @@ -370,26 +360,23 @@ text \Nice induction rule for \trancl\\ lemma tranclp_induct [consumes 1, case_names base step, induct pred: tranclp]: - assumes a: "r^++ a b" - and cases: "!!y. r a y ==> P y" - "!!y z. r^++ a y ==> r y z ==> P y ==> P z" - shows "P b" using a - by (induct x\a b) (iprover intro: cases)+ + assumes a: "r\<^sup>+\<^sup>+ a b" + and cases: "\y. r a y \ P y" "\y z. r\<^sup>+\<^sup>+ a y \ r y z \ P y \ P z" + shows "P b" + using a by (induct x\a b) (iprover intro: cases)+ lemmas trancl_induct [induct set: trancl] = tranclp_induct [to_set] lemmas tranclp_induct2 = - tranclp_induct [of _ "(ax,ay)" "(bx,by)", split_rule, - consumes 1, case_names base step] + tranclp_induct [of _ "(ax,ay)" "(bx,by)", split_rule, consumes 1, case_names base step] lemmas trancl_induct2 = trancl_induct [of "(ax,ay)" "(bx,by)", split_format (complete), consumes 1, case_names base step] lemma tranclp_trans_induct: - assumes major: "r^++ x y" - and cases: "!!x y. r x y ==> P x y" - "!!x y z. [| r^++ x y; P x y; r^++ y z; P y z |] ==> P x z" + assumes major: "r\<^sup>+\<^sup>+ x y" + and cases: "\x y. r x y \ P x y" "\x y z. r\<^sup>+\<^sup>+ x y \ P x y \ r\<^sup>+\<^sup>+ y z \ P y z \ P x z" shows "P x y" \ \Another induction rule for trancl, incorporating transitivity\ by (iprover intro: major [THEN tranclp_induct] cases) @@ -397,49 +384,49 @@ lemmas trancl_trans_induct = tranclp_trans_induct [to_set] lemma tranclE [cases set: trancl]: - assumes "(a, b) : r^+" + assumes "(a, b) \ r\<^sup>+" obtains - (base) "(a, b) : r" - | (step) c where "(a, c) : r^+" and "(c, b) : r" + (base) "(a, b) \ r" + | (step) c where "(a, c) \ r\<^sup>+" and "(c, b) \ r" using assms by cases simp_all -lemma trancl_Int_subset: "[| r \ s; (r^+ \ s) O r \ s|] ==> r^+ \ s" +lemma trancl_Int_subset: "r \ s \ (r\<^sup>+ \ s) O r \ s \ r\<^sup>+ \ s" apply (rule subsetI) apply auto apply (erule trancl_induct) apply auto done -lemma trancl_unfold: "r^+ = r Un r^+ O r" +lemma trancl_unfold: "r\<^sup>+ = r \ r\<^sup>+ O r" by (auto intro: trancl_into_trancl elim: tranclE) -text \Transitivity of @{term "r^+"}\ -lemma trans_trancl [simp]: "trans (r^+)" +text \Transitivity of @{term "r\<^sup>+"}\ +lemma trans_trancl [simp]: "trans (r\<^sup>+)" proof (rule transI) fix x y z - assume "(x, y) \ r^+" - assume "(y, z) \ r^+" - then show "(x, z) \ r^+" + assume "(x, y) \ r\<^sup>+" + assume "(y, z) \ r\<^sup>+" + then show "(x, z) \ r\<^sup>+" proof induct case (base u) - from \(x, y) \ r^+\ and \(y, u) \ r\ - show "(x, u) \ r^+" .. + from \(x, y) \ r\<^sup>+\ and \(y, u) \ r\ + show "(x, u) \ r\<^sup>+" .. next case (step u v) - from \(x, u) \ r^+\ and \(u, v) \ r\ - show "(x, v) \ r^+" .. + from \(x, u) \ r\<^sup>+\ and \(u, v) \ r\ + show "(x, v) \ r\<^sup>+" .. qed qed lemmas trancl_trans = trans_trancl [THEN transD] lemma tranclp_trans: - assumes xy: "r^++ x y" - and yz: "r^++ y z" - shows "r^++ x z" using yz xy - by induct iprover+ + assumes "r\<^sup>+\<^sup>+ x y" + and "r\<^sup>+\<^sup>+ y z" + shows "r\<^sup>+\<^sup>+ x z" + using assms(2,1) by induct iprover+ -lemma trancl_id [simp]: "trans r \ r^+ = r" +lemma trancl_id [simp]: "trans r \ r\<^sup>+ = r" apply auto apply (erule trancl_induct) apply assumption @@ -448,18 +435,18 @@ done lemma rtranclp_tranclp_tranclp: - assumes "r^** x y" - shows "!!z. r^++ y z ==> r^++ x z" using assms - by induct (iprover intro: tranclp_trans)+ + assumes "r\<^sup>*\<^sup>* x y" + shows "\z. r\<^sup>+\<^sup>+ y z \ r\<^sup>+\<^sup>+ x z" + using assms by induct (iprover intro: tranclp_trans)+ lemmas rtrancl_trancl_trancl = rtranclp_tranclp_tranclp [to_set] -lemma tranclp_into_tranclp2: "r a b ==> r^++ b c ==> r^++ a c" +lemma tranclp_into_tranclp2: "r a b \ r\<^sup>+\<^sup>+ b c \ r\<^sup>+\<^sup>+ a c" by (erule tranclp_trans [OF tranclp.r_into_trancl]) lemmas trancl_into_trancl2 = tranclp_into_tranclp2 [to_set] -lemma tranclp_converseI: "(r^++)^--1 x y ==> (r^--1)^++ x y" +lemma tranclp_converseI: "(r\<^sup>+\<^sup>+)\\ x y \ (r\\)\<^sup>+\<^sup>+ x y" apply (drule conversepD) apply (erule tranclp_induct) apply (iprover intro: conversepI tranclp_trans)+ @@ -467,7 +454,7 @@ lemmas trancl_converseI = tranclp_converseI [to_set] -lemma tranclp_converseD: "(r^--1)^++ x y ==> (r^++)^--1 x y" +lemma tranclp_converseD: "(r\\)\<^sup>+\<^sup>+ x y \ (r\<^sup>+\<^sup>+)\\ x y" apply (rule conversepI) apply (erule tranclp_induct) apply (iprover dest: conversepD intro: tranclp_trans)+ @@ -475,19 +462,17 @@ lemmas trancl_converseD = tranclp_converseD [to_set] -lemma tranclp_converse: "(r^--1)^++ = (r^++)^--1" - by (fastforce simp add: fun_eq_iff - intro!: tranclp_converseI dest!: tranclp_converseD) +lemma tranclp_converse: "(r\\)\<^sup>+\<^sup>+ = (r\<^sup>+\<^sup>+)\\" + by (fastforce simp add: fun_eq_iff intro!: tranclp_converseI dest!: tranclp_converseD) lemmas trancl_converse = tranclp_converse [to_set] -lemma sym_trancl: "sym r ==> sym (r^+)" +lemma sym_trancl: "sym r \ sym (r\<^sup>+)" by (simp only: sym_conv_converse_eq trancl_converse [symmetric]) lemma converse_tranclp_induct [consumes 1, case_names base step]: - assumes major: "r^++ a b" - and cases: "!!y. r y b ==> P(y)" - "!!y z.[| r y z; r^++ z b; P(z) |] ==> P(y)" + assumes major: "r\<^sup>+\<^sup>+ a b" + and cases: "\y. r y b \ P y" "\y z. r y z \ r\<^sup>+\<^sup>+ z b \ P z \ P y" shows "P a" apply (rule tranclp_induct [OF tranclp_converseI, OF conversepI, OF major]) apply (rule cases) @@ -497,7 +482,7 @@ lemmas converse_trancl_induct = converse_tranclp_induct [to_set] -lemma tranclpD: "R^++ x y ==> EX z. R x z \ R^** z y" +lemma tranclpD: "R\<^sup>+\<^sup>+ x y \ \z. R x z \ R\<^sup>*\<^sup>* z y" apply (erule converse_tranclp_induct) apply auto apply (blast intro: rtranclp_trans) @@ -507,48 +492,48 @@ lemma converse_tranclpE: assumes major: "tranclp r x z" - assumes base: "r x z ==> P" - assumes step: "\ y. [| r x y; tranclp r y z |] ==> P" + and base: "r x z \ P" + and step: "\ y. r x y \ tranclp r y z \ P" shows P proof - - from tranclpD[OF major] - obtain y where "r x y" and "rtranclp r y z" by iprover + from tranclpD [OF major] obtain y where "r x y" and "rtranclp r y z" + by iprover from this(2) show P proof (cases rule: rtranclp.cases) case rtrancl_refl - with \r x y\ base show P by iprover + with \r x y\ base show P + by iprover next case rtrancl_into_rtrancl from this have "tranclp r y z" by (iprover intro: rtranclp_into_tranclp1) - with \r x y\ step show P by iprover + with \r x y\ step show P + by iprover qed qed lemmas converse_tranclE = converse_tranclpE [to_set] -lemma tranclD2: - "(x, y) \ R\<^sup>+ \ \z. (x, z) \ R\<^sup>* \ (z, y) \ R" +lemma tranclD2: "(x, y) \ R\<^sup>+ \ \z. (x, z) \ R\<^sup>* \ (z, y) \ R" by (blast elim: tranclE intro: trancl_into_rtrancl) -lemma irrefl_tranclI: "r^-1 \ r^* = {} ==> (x, x) \ r^+" +lemma irrefl_tranclI: "r\ \ r\<^sup>* = {} \ (x, x) \ r\<^sup>+" by (blast elim: tranclE dest: trancl_into_rtrancl) -lemma irrefl_trancl_rD: "\x. (x, x) \ r^+ \ (x, y) \ r \ x \ y" +lemma irrefl_trancl_rD: "\x. (x, x) \ r\<^sup>+ \ (x, y) \ r \ x \ y" by (blast dest: r_into_trancl) -lemma trancl_subset_Sigma_aux: - "(a, b) \ r^* ==> r \ A \ A ==> a = b \ a \ A" +lemma trancl_subset_Sigma_aux: "(a, b) \ r\<^sup>* \ r \ A \ A \ a = b \ a \ A" by (induct rule: rtrancl_induct) auto -lemma trancl_subset_Sigma: "r \ A \ A ==> r^+ \ A \ A" +lemma trancl_subset_Sigma: "r \ A \ A \ r\<^sup>+ \ A \ A" apply (rule subsetI) apply (simp only: split_tupled_all) apply (erule tranclE) apply (blast dest!: trancl_into_rtrancl trancl_subset_Sigma_aux)+ done -lemma reflclp_tranclp [simp]: "(r^++)^== = r^**" +lemma reflclp_tranclp [simp]: "(r\<^sup>+\<^sup>+)\<^sup>=\<^sup>= = r\<^sup>*\<^sup>*" apply (safe intro!: order_antisym) apply (erule tranclp_into_rtranclp) apply (blast elim: rtranclp.cases dest: rtranclp_into_tranclp1) @@ -556,7 +541,7 @@ lemmas reflcl_trancl [simp] = reflclp_tranclp [to_set] -lemma trancl_reflcl [simp]: "(r^=)^+ = r^*" +lemma trancl_reflcl [simp]: "(r\<^sup>=)\<^sup>+ = r\<^sup>*" apply safe apply (drule trancl_into_rtrancl, simp) apply (erule rtranclE, safe) @@ -565,32 +550,30 @@ apply (erule rtrancl_reflcl [THEN equalityD2, THEN subsetD], fast) done -lemma rtrancl_trancl_reflcl [code]: "r^* = (r^+)^=" +lemma rtrancl_trancl_reflcl [code]: "r\<^sup>* = (r\<^sup>+)\<^sup>=" by simp -lemma trancl_empty [simp]: "{}^+ = {}" +lemma trancl_empty [simp]: "{}\<^sup>+ = {}" by (auto elim: trancl_induct) -lemma rtrancl_empty [simp]: "{}^* = Id" +lemma rtrancl_empty [simp]: "{}\<^sup>* = Id" by (rule subst [OF reflcl_trancl]) simp -lemma rtranclpD: "R^** a b ==> a = b \ a \ b \ R^++ a b" -by (force simp add: reflclp_tranclp [symmetric] simp del: reflclp_tranclp) +lemma rtranclpD: "R\<^sup>*\<^sup>* a b \ a = b \ a \ b \ R\<^sup>+\<^sup>+ a b" + by (force simp add: reflclp_tranclp [symmetric] simp del: reflclp_tranclp) lemmas rtranclD = rtranclpD [to_set] -lemma rtrancl_eq_or_trancl: - "(x,y) \ R\<^sup>* = (x=y \ x\y \ (x,y) \ R\<^sup>+)" +lemma rtrancl_eq_or_trancl: "(x,y) \ R\<^sup>* \ x = y \ x \ y \ (x, y) \ R\<^sup>+" by (fast elim: trancl_into_rtrancl dest: rtranclD) -lemma trancl_unfold_right: "r^+ = r^* O r" -by (auto dest: tranclD2 intro: rtrancl_into_trancl1) +lemma trancl_unfold_right: "r\<^sup>+ = r\<^sup>* O r" + by (auto dest: tranclD2 intro: rtrancl_into_trancl1) -lemma trancl_unfold_left: "r^+ = r O r^*" -by (auto dest: tranclD intro: rtrancl_into_trancl2) +lemma trancl_unfold_left: "r\<^sup>+ = r O r\<^sup>*" + by (auto dest: tranclD intro: rtrancl_into_trancl2) -lemma trancl_insert: - "(insert (y, x) r)^+ = r^+ \ {(a, b). (a, y) \ r^* \ (x, b) \ r^*}" +lemma trancl_insert: "(insert (y, x) r)\<^sup>+ = r\<^sup>+ \ {(a, b). (a, y) \ r\<^sup>* \ (x, b) \ r\<^sup>*}" \ \primitive recursion for \trancl\ over finite relations\ apply (rule equalityI) apply (rule subsetI) @@ -603,62 +586,60 @@ done lemma trancl_insert2: - "(insert (a,b) r)^+ = r^+ \ {(x,y). ((x,a) : r^+ \ x=a) \ ((b,y) \ r^+ \ y=b)}" -by(auto simp add: trancl_insert rtrancl_eq_or_trancl) + "(insert (a, b) r)\<^sup>+ = r\<^sup>+ \ {(x, y). ((x, a) \ r\<^sup>+ \ x = a) \ ((b, y) \ r\<^sup>+ \ y = b)}" + by (auto simp add: trancl_insert rtrancl_eq_or_trancl) -lemma rtrancl_insert: - "(insert (a,b) r)^* = r^* \ {(x,y). (x,a) : r^* \ (b,y) \ r^*}" -using trancl_insert[of a b r] -by(simp add: rtrancl_trancl_reflcl del: reflcl_trancl) blast +lemma rtrancl_insert: "(insert (a,b) r)\<^sup>* = r\<^sup>* \ {(x, y). (x, a) \ r\<^sup>* \ (b, y) \ r\<^sup>*}" + using trancl_insert[of a b r] + by (simp add: rtrancl_trancl_reflcl del: reflcl_trancl) blast text \Simplifying nested closures\ -lemma rtrancl_trancl_absorb[simp]: "(R^*)^+ = R^*" -by (simp add: trans_rtrancl) +lemma rtrancl_trancl_absorb[simp]: "(R\<^sup>*)\<^sup>+ = R\<^sup>*" + by (simp add: trans_rtrancl) -lemma trancl_rtrancl_absorb[simp]: "(R^+)^* = R^*" -by (subst reflcl_trancl[symmetric]) simp +lemma trancl_rtrancl_absorb[simp]: "(R\<^sup>+)\<^sup>* = R\<^sup>*" + by (subst reflcl_trancl[symmetric]) simp -lemma rtrancl_reflcl_absorb[simp]: "(R^*)^= = R^*" -by auto +lemma rtrancl_reflcl_absorb[simp]: "(R\<^sup>*)\<^sup>= = R\<^sup>*" + by auto text \\Domain\ and \Range\\ -lemma Domain_rtrancl [simp]: "Domain (R^*) = UNIV" +lemma Domain_rtrancl [simp]: "Domain (R\<^sup>*) = UNIV" by blast -lemma Range_rtrancl [simp]: "Range (R^*) = UNIV" +lemma Range_rtrancl [simp]: "Range (R\<^sup>*) = UNIV" by blast -lemma rtrancl_Un_subset: "(R^* \ S^*) \ (R Un S)^*" +lemma rtrancl_Un_subset: "(R\<^sup>* \ S\<^sup>*) \ (R \ S)\<^sup>*" by (rule rtrancl_Un_rtrancl [THEN subst]) fast -lemma in_rtrancl_UnI: "x \ R^* \ x \ S^* ==> x \ (R \ S)^*" +lemma in_rtrancl_UnI: "x \ R\<^sup>* \ x \ S\<^sup>* \ x \ (R \ S)\<^sup>*" by (blast intro: subsetD [OF rtrancl_Un_subset]) -lemma trancl_domain [simp]: "Domain (r^+) = Domain r" +lemma trancl_domain [simp]: "Domain (r\<^sup>+) = Domain r" by (unfold Domain_unfold) (blast dest: tranclD) -lemma trancl_range [simp]: "Range (r^+) = Range r" +lemma trancl_range [simp]: "Range (r\<^sup>+) = Range r" unfolding Domain_converse [symmetric] by (simp add: trancl_converse [symmetric]) -lemma Not_Domain_rtrancl: - "x ~: Domain R ==> ((x, y) : R^*) = (x = y)" +lemma Not_Domain_rtrancl: "x \ Domain R \ (x, y) \ R\<^sup>* \ x = y" apply auto apply (erule rev_mp) apply (erule rtrancl_induct) apply auto done -lemma trancl_subset_Field2: "r^+ <= Field r \ Field r" +lemma trancl_subset_Field2: "r\<^sup>+ \ Field r \ Field r" apply clarify apply (erule trancl_induct) apply (auto simp add: Field_def) done -lemma finite_trancl[simp]: "finite (r^+) = finite r" +lemma finite_trancl[simp]: "finite (r\<^sup>+) = finite r" apply auto prefer 2 apply (rule trancl_subset_Field2 [THEN finite_subset]) @@ -672,8 +653,7 @@ be merged with main body.\ lemma single_valued_confluent: - "\ single_valued r; (x,y) \ r^*; (x,z) \ r^* \ - \ (y,z) \ r^* \ (z,y) \ r^*" + "single_valued r \ (x, y) \ r\<^sup>* \ (x, z) \ r\<^sup>* \ (y, z) \ r\<^sup>* \ (z, y) \ r\<^sup>*" apply (erule rtrancl_induct) apply simp apply (erule disjE) @@ -681,18 +661,16 @@ apply(blast intro:rtrancl_trans) done -lemma r_r_into_trancl: "(a, b) \ R ==> (b, c) \ R ==> (a, c) \ R^+" +lemma r_r_into_trancl: "(a, b) \ R \ (b, c) \ R \ (a, c) \ R\<^sup>+" by (fast intro: trancl_trans) -lemma trancl_into_trancl [rule_format]: - "(a, b) \ r\<^sup>+ ==> (b, c) \ r --> (a,c) \ r\<^sup>+" - apply (erule trancl_induct) +lemma trancl_into_trancl: "(a, b) \ r\<^sup>+ \ (b, c) \ r \ (a, c) \ r\<^sup>+" + apply (induct rule: trancl_induct) apply (fast intro: r_r_into_trancl) apply (fast intro: r_r_into_trancl trancl_trans) done -lemma tranclp_rtranclp_tranclp: - "r\<^sup>+\<^sup>+ a b ==> r\<^sup>*\<^sup>* b c ==> r\<^sup>+\<^sup>+ a c" +lemma tranclp_rtranclp_tranclp: "r\<^sup>+\<^sup>+ a b \ r\<^sup>*\<^sup>* b c \ r\<^sup>+\<^sup>+ a c" apply (drule tranclpD) apply (elim exE conjE) apply (drule rtranclp_trans, assumption) @@ -715,37 +693,39 @@ declare trancl_into_rtrancl [elim] + subsection \The power operation on relations\ -text \\R ^^ n = R O ... O R\, the n-fold composition of \R\\ +text \\R ^^ n = R O \ O R\, the n-fold composition of \R\\ overloading - relpow == "compow :: nat \ ('a \ 'a) set \ ('a \ 'a) set" - relpowp == "compow :: nat \ ('a \ 'a \ bool) \ ('a \ 'a \ bool)" + relpow \ "compow :: nat \ ('a \ 'a) set \ ('a \ 'a) set" + relpowp \ "compow :: nat \ ('a \ 'a \ bool) \ ('a \ 'a \ bool)" begin -primrec relpow :: "nat \ ('a \ 'a) set \ ('a \ 'a) set" where - "relpow 0 R = Id" - | "relpow (Suc n) R = (R ^^ n) O R" +primrec relpow :: "nat \ ('a \ 'a) set \ ('a \ 'a) set" +where + "relpow 0 R = Id" +| "relpow (Suc n) R = (R ^^ n) O R" -primrec relpowp :: "nat \ ('a \ 'a \ bool) \ ('a \ 'a \ bool)" where - "relpowp 0 R = HOL.eq" - | "relpowp (Suc n) R = (R ^^ n) OO R" +primrec relpowp :: "nat \ ('a \ 'a \ bool) \ ('a \ 'a \ bool)" +where + "relpowp 0 R = HOL.eq" +| "relpowp (Suc n) R = (R ^^ n) OO R" end lemma relpowp_relpow_eq [pred_set_conv]: - fixes R :: "'a rel" - shows "(\x y. (x, y) \ R) ^^ n = (\x y. (x, y) \ R ^^ n)" + "(\x y. (x, y) \ R) ^^ n = (\x y. (x, y) \ R ^^ n)" for R :: "'a rel" by (induct n) (simp_all add: relcompp_relcomp_eq) -text \for code generation\ +text \For code generation:\ -definition relpow :: "nat \ ('a \ 'a) set \ ('a \ 'a) set" where - relpow_code_def [code_abbrev]: "relpow = compow" +definition relpow :: "nat \ ('a \ 'a) set \ ('a \ 'a) set" + where relpow_code_def [code_abbrev]: "relpow = compow" -definition relpowp :: "nat \ ('a \ 'a \ bool) \ ('a \ 'a \ bool)" where - relpowp_code_def [code_abbrev]: "relpowp = compow" +definition relpowp :: "nat \ ('a \ 'a \ bool) \ ('a \ 'a \ bool)" + where relpowp_code_def [code_abbrev]: "relpowp = compow" lemma [code]: "relpow (Suc n) R = (relpow n R) O R" @@ -760,54 +740,40 @@ hide_const (open) relpow hide_const (open) relpowp -lemma relpow_1 [simp]: - fixes R :: "('a \ 'a) set" - shows "R ^^ 1 = R" +lemma relpow_1 [simp]: "R ^^ 1 = R" for R :: "('a \ 'a) set" by simp -lemma relpowp_1 [simp]: - fixes P :: "'a \ 'a \ bool" - shows "P ^^ 1 = P" +lemma relpowp_1 [simp]: "P ^^ 1 = P" for P :: "'a \ 'a \ bool" by (fact relpow_1 [to_pred]) -lemma relpow_0_I: - "(x, x) \ R ^^ 0" +lemma relpow_0_I: "(x, x) \ R ^^ 0" by simp -lemma relpowp_0_I: - "(P ^^ 0) x x" +lemma relpowp_0_I: "(P ^^ 0) x x" by (fact relpow_0_I [to_pred]) -lemma relpow_Suc_I: - "(x, y) \ R ^^ n \ (y, z) \ R \ (x, z) \ R ^^ Suc n" +lemma relpow_Suc_I: "(x, y) \ R ^^ n \ (y, z) \ R \ (x, z) \ R ^^ Suc n" by auto -lemma relpowp_Suc_I: - "(P ^^ n) x y \ P y z \ (P ^^ Suc n) x z" +lemma relpowp_Suc_I: "(P ^^ n) x y \ P y z \ (P ^^ Suc n) x z" by (fact relpow_Suc_I [to_pred]) -lemma relpow_Suc_I2: - "(x, y) \ R \ (y, z) \ R ^^ n \ (x, z) \ R ^^ Suc n" +lemma relpow_Suc_I2: "(x, y) \ R \ (y, z) \ R ^^ n \ (x, z) \ R ^^ Suc n" by (induct n arbitrary: z) (simp, fastforce) -lemma relpowp_Suc_I2: - "P x y \ (P ^^ n) y z \ (P ^^ Suc n) x z" +lemma relpowp_Suc_I2: "P x y \ (P ^^ n) y z \ (P ^^ Suc n) x z" by (fact relpow_Suc_I2 [to_pred]) -lemma relpow_0_E: - "(x, y) \ R ^^ 0 \ (x = y \ P) \ P" +lemma relpow_0_E: "(x, y) \ R ^^ 0 \ (x = y \ P) \ P" by simp -lemma relpowp_0_E: - "(P ^^ 0) x y \ (x = y \ Q) \ Q" +lemma relpowp_0_E: "(P ^^ 0) x y \ (x = y \ Q) \ Q" by (fact relpow_0_E [to_pred]) -lemma relpow_Suc_E: - "(x, z) \ R ^^ Suc n \ (\y. (x, y) \ R ^^ n \ (y, z) \ R \ P) \ P" +lemma relpow_Suc_E: "(x, z) \ R ^^ Suc n \ (\y. (x, y) \ R ^^ n \ (y, z) \ R \ P) \ P" by auto -lemma relpowp_Suc_E: - "(P ^^ Suc n) x z \ (\y. (P ^^ n) x y \ P y z \ Q) \ Q" +lemma relpowp_Suc_E: "(P ^^ Suc n) x z \ (\y. (P ^^ n) x y \ P y z \ Q) \ Q" by (fact relpow_Suc_E [to_pred]) lemma relpow_E: @@ -822,31 +788,25 @@ \ Q" by (fact relpow_E [to_pred]) -lemma relpow_Suc_D2: - "(x, z) \ R ^^ Suc n \ (\y. (x, y) \ R \ (y, z) \ R ^^ n)" +lemma relpow_Suc_D2: "(x, z) \ R ^^ Suc n \ (\y. (x, y) \ R \ (y, z) \ R ^^ n)" apply (induct n arbitrary: x z) apply (blast intro: relpow_0_I elim: relpow_0_E relpow_Suc_E) apply (blast intro: relpow_Suc_I elim: relpow_0_E relpow_Suc_E) done -lemma relpowp_Suc_D2: - "(P ^^ Suc n) x z \ \y. P x y \ (P ^^ n) y z" +lemma relpowp_Suc_D2: "(P ^^ Suc n) x z \ \y. P x y \ (P ^^ n) y z" by (fact relpow_Suc_D2 [to_pred]) -lemma relpow_Suc_E2: - "(x, z) \ R ^^ Suc n \ (\y. (x, y) \ R \ (y, z) \ R ^^ n \ P) \ P" +lemma relpow_Suc_E2: "(x, z) \ R ^^ Suc n \ (\y. (x, y) \ R \ (y, z) \ R ^^ n \ P) \ P" by (blast dest: relpow_Suc_D2) -lemma relpowp_Suc_E2: - "(P ^^ Suc n) x z \ (\y. P x y \ (P ^^ n) y z \ Q) \ Q" +lemma relpowp_Suc_E2: "(P ^^ Suc n) x z \ (\y. P x y \ (P ^^ n) y z \ Q) \ Q" by (fact relpow_Suc_E2 [to_pred]) -lemma relpow_Suc_D2': - "\x y z. (x, y) \ R ^^ n \ (y, z) \ R \ (\w. (x, w) \ R \ (w, z) \ R ^^ n)" +lemma relpow_Suc_D2': "\x y z. (x, y) \ R ^^ n \ (y, z) \ R \ (\w. (x, w) \ R \ (w, z) \ R ^^ n)" by (induct n) (simp_all, blast) -lemma relpowp_Suc_D2': - "\x y z. (P ^^ n) x y \ P y z \ (\w. P x w \ (P ^^ n) w z)" +lemma relpowp_Suc_D2': "\x y z. (P ^^ n) x y \ P y z \ (\w. P x w \ (P ^^ n) w z)" by (fact relpow_Suc_D2' [to_pred]) lemma relpow_E2: @@ -864,83 +824,78 @@ \ Q" by (fact relpow_E2 [to_pred]) -lemma relpow_add: "R ^^ (m+n) = R^^m O R^^n" +lemma relpow_add: "R ^^ (m + n) = R^^m O R^^n" by (induct n) auto lemma relpowp_add: "P ^^ (m + n) = P ^^ m OO P ^^ n" by (fact relpow_add [to_pred]) lemma relpow_commute: "R O R ^^ n = R ^^ n O R" - by (induct n) (simp, simp add: O_assoc [symmetric]) + by (induct n) (simp_all add: O_assoc [symmetric]) lemma relpowp_commute: "P OO P ^^ n = P ^^ n OO P" by (fact relpow_commute [to_pred]) -lemma relpow_empty: - "0 < n \ ({} :: ('a \ 'a) set) ^^ n = {}" +lemma relpow_empty: "0 < n \ ({} :: ('a \ 'a) set) ^^ n = {}" by (cases n) auto -lemma relpowp_bot: - "0 < n \ (\ :: 'a \ 'a \ bool) ^^ n = \" +lemma relpowp_bot: "0 < n \ (\ :: 'a \ 'a \ bool) ^^ n = \" by (fact relpow_empty [to_pred]) lemma rtrancl_imp_UN_relpow: - assumes "p \ R^*" + assumes "p \ R\<^sup>*" shows "p \ (\n. R ^^ n)" proof (cases p) case (Pair x y) - with assms have "(x, y) \ R^*" by simp + with assms have "(x, y) \ R\<^sup>*" by simp then have "(x, y) \ (\n. R ^^ n)" proof induct - case base show ?case by (blast intro: relpow_0_I) + case base + show ?case by (blast intro: relpow_0_I) next - case step then show ?case by (blast intro: relpow_Suc_I) + case step + then show ?case by (blast intro: relpow_Suc_I) qed with Pair show ?thesis by simp qed lemma rtranclp_imp_Sup_relpowp: - assumes "(P^**) x y" + assumes "(P\<^sup>*\<^sup>*) x y" shows "(\n. P ^^ n) x y" using assms and rtrancl_imp_UN_relpow [of "(x, y)", to_pred] by simp lemma relpow_imp_rtrancl: assumes "p \ R ^^ n" - shows "p \ R^*" + shows "p \ R\<^sup>*" proof (cases p) case (Pair x y) with assms have "(x, y) \ R ^^ n" by simp - then have "(x, y) \ R^*" proof (induct n arbitrary: x y) - case 0 then show ?case by simp + then have "(x, y) \ R\<^sup>*" proof (induct n arbitrary: x y) + case 0 + then show ?case by simp next - case Suc then show ?case + case Suc + then show ?case by (blast elim: relpow_Suc_E intro: rtrancl_into_rtrancl) qed with Pair show ?thesis by simp qed -lemma relpowp_imp_rtranclp: - assumes "(P ^^ n) x y" - shows "(P^**) x y" - using assms and relpow_imp_rtrancl [of "(x, y)", to_pred] by simp +lemma relpowp_imp_rtranclp: "(P ^^ n) x y \ (P\<^sup>*\<^sup>*) x y" + using relpow_imp_rtrancl [of "(x, y)", to_pred] by simp -lemma rtrancl_is_UN_relpow: - "R^* = (\n. R ^^ n)" +lemma rtrancl_is_UN_relpow: "R\<^sup>* = (\n. R ^^ n)" by (blast intro: rtrancl_imp_UN_relpow relpow_imp_rtrancl) -lemma rtranclp_is_Sup_relpowp: - "P^** = (\n. P ^^ n)" +lemma rtranclp_is_Sup_relpowp: "P\<^sup>*\<^sup>* = (\n. P ^^ n)" using rtrancl_is_UN_relpow [to_pred, of P] by auto -lemma rtrancl_power: - "p \ R^* \ (\n. p \ R ^^ n)" +lemma rtrancl_power: "p \ R\<^sup>* \ (\n. p \ R ^^ n)" by (simp add: rtrancl_is_UN_relpow) -lemma rtranclp_power: - "(P^**) x y \ (\n. (P ^^ n) x y)" +lemma rtranclp_power: "(P\<^sup>*\<^sup>*) x y \ (\n. (P ^^ n) x y)" by (simp add: rtranclp_is_Sup_relpowp) -lemma trancl_power: - "p \ R^+ \ (\n > 0. p \ R ^^ n)" +lemma trancl_power: "p \ R\<^sup>+ \ (\n > 0. p \ R ^^ n)" apply (cases p) apply simp apply (rule iffI) @@ -956,187 +911,204 @@ apply (drule rtrancl_into_trancl1) apply auto done -lemma tranclp_power: - "(P^++) x y \ (\n > 0. (P ^^ n) x y)" +lemma tranclp_power: "(P\<^sup>+\<^sup>+) x y \ (\n > 0. (P ^^ n) x y)" using trancl_power [to_pred, of P "(x, y)"] by simp -lemma rtrancl_imp_relpow: - "p \ R^* \ \n. p \ R ^^ n" +lemma rtrancl_imp_relpow: "p \ R\<^sup>* \ \n. p \ R ^^ n" by (auto dest: rtrancl_imp_UN_relpow) -lemma rtranclp_imp_relpowp: - "(P^**) x y \ \n. (P ^^ n) x y" +lemma rtranclp_imp_relpowp: "(P\<^sup>*\<^sup>*) x y \ \n. (P ^^ n) x y" by (auto dest: rtranclp_imp_Sup_relpowp) -text\By Sternagel/Thiemann:\ -lemma relpow_fun_conv: - "((a,b) \ R ^^ n) = (\f. f 0 = a \ f n = b \ (\i R))" +text \By Sternagel/Thiemann:\ +lemma relpow_fun_conv: "(a, b) \ R ^^ n \ (\f. f 0 = a \ f n = b \ (\i R))" proof (induct n arbitrary: b) - case 0 show ?case by auto + case 0 + show ?case by auto next case (Suc n) show ?case proof (simp add: relcomp_unfold Suc) - show "(\y. (\f. f 0 = a \ f n = y \ (\i R)) \ (y,b) \ R) - = (\f. f 0 = a \ f(Suc n) = b \ (\i R))" + show "(\y. (\f. f 0 = a \ f n = y \ (\i R)) \ (y,b) \ R) \ + (\f. f 0 = a \ f(Suc n) = b \ (\i R))" (is "?l = ?r") proof assume ?l - then obtain c f where 1: "f 0 = a" "f n = c" "\i. i < n \ (f i, f (Suc i)) \ R" "(c,b) \ R" by auto + then obtain c f + where 1: "f 0 = a" "f n = c" "\i. i < n \ (f i, f (Suc i)) \ R" "(c,b) \ R" + by auto let ?g = "\ m. if m = Suc n then b else f m" - show ?r by (rule exI[of _ ?g], simp add: 1) + show ?r by (rule exI[of _ ?g]) (simp add: 1) next assume ?r - then obtain f where 1: "f 0 = a" "b = f (Suc n)" "\i. i < Suc n \ (f i, f (Suc i)) \ R" by auto + then obtain f where 1: "f 0 = a" "b = f (Suc n)" "\i. i < Suc n \ (f i, f (Suc i)) \ R" + by auto show ?l by (rule exI[of _ "f n"], rule conjI, rule exI[of _ f], insert 1, auto) qed qed qed -lemma relpowp_fun_conv: - "(P ^^ n) x y \ (\f. f 0 = x \ f n = y \ (\i (\f. f 0 = x \ f n = y \ (\i0" -shows "R^^k \ (UN n:{n. 0 ?r") -proof- - { fix a b k - have "(a,b) : R^^(Suc k) \ EX n. 0 {}" by auto - with card_0_eq[OF \finite R\] have "card R >= Suc 0" by auto - thus ?case using 0 by force + fixes R :: "('a \ 'a) set" + assumes "finite R" and "k > 0" + shows "R^^k \ (\n\{n. 0 < n \ n \ card R}. R^^n)" (is "_ \ ?r") +proof - + have "(a, b) \ R^^(Suc k) \ \n. 0 < n \ n \ card R \ (a, b) \ R^^n" for a b k + proof (induct k arbitrary: b) + case 0 + then have "R \ {}" by auto + with card_0_eq[OF \finite R\] have "card R \ Suc 0" by auto + then show ?case using 0 by force + next + case (Suc k) + then obtain a' where "(a, a') \ R^^(Suc k)" and "(a', b) \ R" + by auto + from Suc(1)[OF \(a, a') \ R^^(Suc k)\] obtain n where "n \ card R" and "(a, a') \ R ^^ n" + by auto + have "(a, b) \ R^^(Suc n)" + using \(a, a') \ R^^n\ and \(a', b)\ R\ by auto + from \n \ card R\ consider "n < card R" | "n = card R" by force + then show ?case + proof cases + case 1 + then show ?thesis + using \(a, b) \ R^^(Suc n)\ Suc_leI[OF \n < card R\] by blast next - case (Suc k) - then obtain a' where "(a,a') : R^^(Suc k)" and "(a',b) : R" by auto - from Suc(1)[OF \(a,a') : R^^(Suc k)\] - obtain n where "n \ card R" and "(a,a') \ R ^^ n" by auto - have "(a,b) : R^^(Suc n)" using \(a,a') \ R^^n\ and \(a',b)\ R\ by auto - { assume "n < card R" - hence ?case using \(a,b): R^^(Suc n)\ Suc_leI[OF \n < card R\] by blast - } moreover - { assume "n = card R" - from \(a,b) \ R ^^ (Suc n)\[unfolded relpow_fun_conv] - obtain f where "f 0 = a" and "f(Suc n) = b" - and steps: "\i. i <= n \ (f i, f (Suc i)) \ R" by auto - let ?p = "%i. (f i, f(Suc i))" - let ?N = "{i. i \ n}" - have "?p ` ?N <= R" using steps by auto - from card_mono[OF assms(1) this] - have "card(?p ` ?N) <= card R" . - also have "\ < card ?N" using \n = card R\ by simp - finally have "~ inj_on ?p ?N" by(rule pigeonhole) - then obtain i j where i: "i <= n" and j: "j <= n" and ij: "i \ j" and - pij: "?p i = ?p j" by(auto simp: inj_on_def) - let ?i = "min i j" let ?j = "max i j" - have i: "?i <= n" and j: "?j <= n" and pij: "?p ?i = ?p ?j" - and ij: "?i < ?j" - using i j ij pij unfolding min_def max_def by auto - from i j pij ij obtain i j where i: "i<=n" and j: "j<=n" and ij: "i R ^^ ?n" unfolding relpow_fun_conv - proof (rule exI[of _ ?g], intro conjI impI allI) - show "?g ?n = b" using \f(Suc n) = b\ j ij by auto + case 2 + from \(a, b) \ R ^^ (Suc n)\ [unfolded relpow_fun_conv] + obtain f where "f 0 = a" and "f (Suc n) = b" + and steps: "\i. i \ n \ (f i, f (Suc i)) \ R" by auto + let ?p = "\i. (f i, f(Suc i))" + let ?N = "{i. i \ n}" + have "?p ` ?N \ R" + using steps by auto + from card_mono[OF assms(1) this] have "card (?p ` ?N) \ card R" . + also have "\ < card ?N" + using \n = card R\ by simp + finally have "\ inj_on ?p ?N" + by (rule pigeonhole) + then obtain i j where i: "i \ n" and j: "j \ n" and ij: "i \ j" and pij: "?p i = ?p j" + by (auto simp: inj_on_def) + let ?i = "min i j" + let ?j = "max i j" + have i: "?i \ n" and j: "?j \ n" and pij: "?p ?i = ?p ?j" and ij: "?i < ?j" + using i j ij pij unfolding min_def max_def by auto + from i j pij ij obtain i j where i: "i \ n" and j: "j \ n" and ij: "i < j" + and pij: "?p i = ?p j" + by blast + let ?g = "\l. if l \ i then f l else f (l + (j - i))" + let ?n = "Suc (n - (j - i))" + have abl: "(a, b) \ R ^^ ?n" + unfolding relpow_fun_conv + proof (rule exI[of _ ?g], intro conjI impI allI) + show "?g ?n = b" + using \f(Suc n) = b\ j ij by auto + next + fix k + assume "k < ?n" + show "(?g k, ?g (Suc k)) \ R" + proof (cases "k < i") + case True + with i have "k \ n" + by auto + from steps[OF this] show ?thesis + using True by simp next - fix k assume "k < ?n" - show "(?g k, ?g (Suc k)) \ R" - proof (cases "k < i") + case False + then have "i \ k" by auto + show ?thesis + proof (cases "k = i") case True - with i have "k <= n" by auto - from steps[OF this] show ?thesis using True by simp + then show ?thesis + using ij pij steps[OF i] by simp next case False - hence "i \ k" by auto + with \i \ k\ have "i < k" by auto + then have small: "k + (j - i) \ n" + using \k by arith show ?thesis - proof (cases "k = i") - case True - thus ?thesis using ij pij steps[OF i] by simp - next - case False - with \i \ k\ have "i < k" by auto - hence small: "k + (j - i) <= n" using \k by arith - show ?thesis using steps[OF small] \i by auto - qed + using steps[OF small] \i by auto qed - qed (simp add: \f 0 = a\) - moreover have "?n <= n" using i j ij by arith - ultimately have ?case using \n = card R\ by blast - } - ultimately show ?case using \n \ card R\ by force + qed + qed (simp add: \f 0 = a\) + moreover have "?n \ n" + using i j ij by arith + ultimately show ?thesis + using \n = card R\ by blast qed - } - thus ?thesis using gr0_implies_Suc[OF \k>0\] by auto + qed + then show ?thesis + using gr0_implies_Suc[OF \k > 0\] by auto qed lemma relpow_finite_bounded: -assumes "finite(R :: ('a*'a)set)" -shows "R^^k \ (UN n:{n. n <= card R}. R^^n)" -apply(cases k) - apply force -using relpow_finite_bounded1[OF assms, of k] by auto + fixes R :: "('a \ 'a) set" + assumes "finite R" + shows "R^^k \ (UN n:{n. n \ card R}. R^^n)" + apply (cases k) + apply force + using relpow_finite_bounded1[OF assms, of k] + apply auto + done -lemma rtrancl_finite_eq_relpow: - "finite R \ R^* = (UN n : {n. n <= card R}. R^^n)" -by(fastforce simp: rtrancl_power dest: relpow_finite_bounded) +lemma rtrancl_finite_eq_relpow: "finite R \ R\<^sup>* = (\n\{n. n \ card R}. R^^n)" + by (fastforce simp: rtrancl_power dest: relpow_finite_bounded) -lemma trancl_finite_eq_relpow: - "finite R \ R^+ = (UN n : {n. 0 < n & n <= card R}. R^^n)" -apply(auto simp add: trancl_power) -apply(auto dest: relpow_finite_bounded1) -done +lemma trancl_finite_eq_relpow: "finite R \ R\<^sup>+ = (\n\{n. 0 < n \ n \ card R}. R^^n)" + apply (auto simp: trancl_power) + apply (auto dest: relpow_finite_bounded1) + done lemma finite_relcomp[simp,intro]: -assumes "finite R" and "finite S" -shows "finite(R O S)" + assumes "finite R" and "finite S" + shows "finite (R O S)" proof- have "R O S = (\(x, y)\R. \(u, v)\S. if u = y then {(x, v)} else {})" by (force simp add: split_def image_constant_conv split: if_splits) - then show ?thesis using assms by clarsimp + then show ?thesis + using assms by clarsimp qed -lemma finite_relpow[simp,intro]: - assumes "finite(R :: ('a*'a)set)" shows "n>0 \ finite(R^^n)" -apply(induct n) - apply simp -apply(case_tac n) - apply(simp_all add: assms) -done +lemma finite_relpow [simp, intro]: + fixes R :: "('a \ 'a) set" + assumes "finite R" + shows "n > 0 \ finite (R^^n)" + apply (induct n) + apply simp + apply (case_tac n) + apply (simp_all add: assms) + done lemma single_valued_relpow: - fixes R :: "('a * 'a) set" + fixes R :: "('a \ 'a) set" shows "single_valued R \ single_valued (R ^^ n)" -apply (induct n arbitrary: R) -apply simp_all -apply (rule single_valuedI) -apply (fast dest: single_valuedD elim: relpow_Suc_E) -done + apply (induct n arbitrary: R) + apply simp_all + apply (rule single_valuedI) + apply (fast dest: single_valuedD elim: relpow_Suc_E) + done subsection \Bounded transitive closure\ definition ntrancl :: "nat \ ('a \ 'a) set \ ('a \ 'a) set" -where - "ntrancl n R = (\i\{i. 0 < i \ i \ Suc n}. R ^^ i)" + where "ntrancl n R = (\i\{i. 0 < i \ i \ Suc n}. R ^^ i)" -lemma ntrancl_Zero [simp, code]: - "ntrancl 0 R = R" +lemma ntrancl_Zero [simp, code]: "ntrancl 0 R = R" proof show "R \ ntrancl 0 R" unfolding ntrancl_def by fastforce next - { - fix i have "0 < i \ i \ Suc 0 \ i = 1" by auto - } - from this show "ntrancl 0 R \ R" + have "0 < i \ i \ Suc 0 \ i = 1" for i + by auto + then show "ntrancl 0 R \ R" unfolding ntrancl_def by auto qed -lemma ntrancl_Suc [simp]: - "ntrancl (Suc n) R = ntrancl n R O (Id \ R)" +lemma ntrancl_Suc [simp]: "ntrancl (Suc n) R = ntrancl n R O (Id \ R)" proof { fix a b @@ -1159,75 +1131,67 @@ from this c2 show ?thesis by fastforce qed } - from this show "ntrancl (Suc n) R \ ntrancl n R O (Id \ R)" + then show "ntrancl (Suc n) R \ ntrancl n R O (Id \ R)" by auto -next show "ntrancl n R O (Id \ R) \ ntrancl (Suc n) R" unfolding ntrancl_def by fastforce qed -lemma [code]: - "ntrancl (Suc n) r = (let r' = ntrancl n r in r' Un r' O r)" -unfolding Let_def by auto +lemma [code]: "ntrancl (Suc n) r = (let r' = ntrancl n r in r' \ r' O r)" + by (auto simp: Let_def) -lemma finite_trancl_ntranl: - "finite R \ trancl R = ntrancl (card R - 1) R" +lemma finite_trancl_ntranl: "finite R \ trancl R = ntrancl (card R - 1) R" by (cases "card R") (auto simp add: trancl_finite_eq_relpow relpow_empty ntrancl_def) subsection \Acyclic relations\ -definition acyclic :: "('a * 'a) set => bool" where - "acyclic r \ (!x. (x,x) ~: r^+)" +definition acyclic :: "('a \ 'a) set \ bool" + where "acyclic r \ (\x. (x,x) \ r\<^sup>+)" -abbreviation acyclicP :: "('a => 'a => bool) => bool" where - "acyclicP r \ acyclic {(x, y). r x y}" +abbreviation acyclicP :: "('a \ 'a \ bool) \ bool" + where "acyclicP r \ acyclic {(x, y). r x y}" -lemma acyclic_irrefl [code]: - "acyclic r \ irrefl (r^+)" +lemma acyclic_irrefl [code]: "acyclic r \ irrefl (r\<^sup>+)" by (simp add: acyclic_def irrefl_def) -lemma acyclicI: "ALL x. (x, x) ~: r^+ ==> acyclic r" +lemma acyclicI: "\x. (x, x) \ r\<^sup>+ \ acyclic r" by (simp add: acyclic_def) lemma (in order) acyclicI_order: assumes *: "\a b. (a, b) \ r \ f b < f a" shows "acyclic r" proof - - { fix a b assume "(a, b) \ r\<^sup>+" - then have "f b < f a" - by induct (auto intro: * less_trans) } + have "f b < f a" if "(a, b) \ r\<^sup>+" for a b + using that by induct (auto intro: * less_trans) then show ?thesis by (auto intro!: acyclicI) qed -lemma acyclic_insert [iff]: - "acyclic(insert (y,x) r) = (acyclic r & (x,y) ~: r^*)" -apply (simp add: acyclic_def trancl_insert) -apply (blast intro: rtrancl_trans) -done +lemma acyclic_insert [iff]: "acyclic (insert (y, x) r) \ acyclic r \ (x, y) \ r\<^sup>*" + apply (simp add: acyclic_def trancl_insert) + apply (blast intro: rtrancl_trans) + done -lemma acyclic_converse [iff]: "acyclic(r^-1) = acyclic r" -by (simp add: acyclic_def trancl_converse) +lemma acyclic_converse [iff]: "acyclic (r\) \ acyclic r" + by (simp add: acyclic_def trancl_converse) lemmas acyclicP_converse [iff] = acyclic_converse [to_pred] -lemma acyclic_impl_antisym_rtrancl: "acyclic r ==> antisym(r^*)" -apply (simp add: acyclic_def antisym_def) -apply (blast elim: rtranclE intro: rtrancl_into_trancl1 rtrancl_trancl_trancl) -done +lemma acyclic_impl_antisym_rtrancl: "acyclic r \ antisym (r\<^sup>*)" + apply (simp add: acyclic_def antisym_def) + apply (blast elim: rtranclE intro: rtrancl_into_trancl1 rtrancl_trancl_trancl) + done (* Other direction: acyclic = no loops antisym = only self loops -Goalw [acyclic_def,antisym_def] "antisym( r^* ) ==> acyclic(r - Id) -==> antisym( r^* ) = acyclic(r - Id)"; +Goalw [acyclic_def,antisym_def] "antisym( r\<^sup>* ) \ acyclic(r - Id) +\ antisym( r\<^sup>* ) = acyclic(r - Id)"; *) -lemma acyclic_subset: "[| acyclic s; r <= s |] ==> acyclic r" -apply (simp add: acyclic_def) -apply (blast intro: trancl_mono) -done +lemma acyclic_subset: "acyclic s \ r \ s \ acyclic r" + unfolding acyclic_def by (blast intro: trancl_mono) subsection \Setup of transitivity reasoner\ @@ -1246,14 +1210,16 @@ val rtrancl_trans = @{thm rtrancl_trans}; fun decomp (@{const Trueprop} $ t) = - let fun dec (Const (@{const_name Set.member}, _) $ (Const (@{const_name Pair}, _) $ a $ b) $ rel ) = - let fun decr (Const (@{const_name rtrancl}, _ ) $ r) = (r,"r*") - | decr (Const (@{const_name trancl}, _ ) $ r) = (r,"r+") - | decr r = (r,"r"); - val (rel,r) = decr (Envir.beta_eta_contract rel); - in SOME (a,b,rel,r) end - | dec _ = NONE - in dec t end + let + fun dec (Const (@{const_name Set.member}, _) $ (Const (@{const_name Pair}, _) $ a $ b) $ rel) = + let + fun decr (Const (@{const_name rtrancl}, _ ) $ r) = (r,"r*") + | decr (Const (@{const_name trancl}, _ ) $ r) = (r,"r+") + | decr r = (r,"r"); + val (rel,r) = decr (Envir.beta_eta_contract rel); + in SOME (a,b,rel,r) end + | dec _ = NONE + in dec t end | decomp _ = NONE; ); @@ -1269,14 +1235,16 @@ val rtrancl_trans = @{thm rtranclp_trans}; fun decomp (@{const Trueprop} $ t) = - let fun dec (rel $ a $ b) = - let fun decr (Const (@{const_name rtranclp}, _ ) $ r) = (r,"r*") - | decr (Const (@{const_name tranclp}, _ ) $ r) = (r,"r+") - | decr r = (r,"r"); - val (rel,r) = decr rel; - in SOME (a, b, rel, r) end - | dec _ = NONE - in dec t end + let + fun dec (rel $ a $ b) = + let + fun decr (Const (@{const_name rtranclp}, _ ) $ r) = (r,"r*") + | decr (Const (@{const_name tranclp}, _ ) $ r) = (r,"r+") + | decr r = (r,"r"); + val (rel,r) = decr rel; + in SOME (a, b, rel, r) end + | dec _ = NONE + in dec t end | decomp _ = NONE; ); \