From: Andrea Asperti <andrea.asperti@unibo.it> Date: Mon, 28 Oct 2013 14:37:01 +0000 (+0000) Subject: Splitted star X-Git-Tag: make_still_working~1059 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=4adf9860cd26175c4d73b73e8adbb3c6ceaa19c9;p=helm.git Splitted star --- diff --git a/matita/matita/lib/basics/star.ma b/matita/matita/lib/basics/star.ma index 00a23b633..391e08530 100644 --- a/matita/matita/lib/basics/star.ma +++ b/matita/matita/lib/basics/star.ma @@ -193,348 +193,3 @@ lemma WF_antimonotonic: âA,R,S. subR A R S â #H #Hind % #c #Rcb @Hind @subRS // qed. -(* added from λδ *) - -lemma TC_strap: âA. âR:relation A. âa1,a,a2. - R a1 a â TC ⦠R a a2 â TC ⦠R a1 a2. -/3 width=3/ qed. - -lemma TC_reflexive: âA,R. reflexive A R â reflexive A (TC ⦠R). -/2 width=1/ qed. - -lemma TC_star_ind: âA,R. reflexive A R â âa1. âP:predicate A. - P a1 â (âa,a2. TC ⦠R a1 a â R a a2 â P a â P a2) â - âa2. TC ⦠R a1 a2 â P a2. -#A #R #H #a1 #P #Ha1 #IHa1 #a2 #Ha12 elim Ha12 -a2 /3 width=4/ -qed-. - -inductive TC_dx (A:Type[0]) (R:relation A): A â A â Prop â - |inj_dx: âa,c. R a c â TC_dx A R a c - |step_dx : âa,b,c. R a b â TC_dx A R b c â TC_dx A R a c. - -lemma TC_dx_strap: âA. âR: relation A. - âa,b,c. TC_dx A R a b â R b c â TC_dx A R a c. -#A #R #a #b #c #Hab elim Hab -a -b /3 width=3/ -qed. - -lemma TC_to_TC_dx: âA. âR: relation A. - âa1,a2. TC ⦠R a1 a2 â TC_dx ⦠R a1 a2. -#A #R #a1 #a2 #Ha12 elim Ha12 -a2 /2 width=3/ -qed. - -lemma TC_dx_to_TC: âA. âR: relation A. - âa1,a2. TC_dx ⦠R a1 a2 â TC ⦠R a1 a2. -#A #R #a1 #a2 #Ha12 elim Ha12 -a1 -a2 /2 width=3/ -qed. - -fact TC_ind_dx_aux: âA,R,a2. âP:predicate A. - (âa1. R a1 a2 â P a1) â - (âa1,a. R a1 a â TC ⦠R a a2 â P a â P a1) â - âa1,a. TC ⦠R a1 a â a = a2 â P a1. -#A #R #a2 #P #H1 #H2 #a1 #a #Ha1 -elim (TC_to_TC_dx ???? Ha1) -a1 -a -[ #a #c #Hac #H destruct /2 width=1/ -| #a #b #c #Hab #Hbc #IH #H destruct /3 width=4/ -] -qed-. - -lemma TC_ind_dx: âA,R,a2. âP:predicate A. - (âa1. R a1 a2 â P a1) â - (âa1,a. R a1 a â TC ⦠R a a2 â P a â P a1) â - âa1. TC ⦠R a1 a2 â P a1. -#A #R #a2 #P #H1 #H2 #a1 #Ha12 -@(TC_ind_dx_aux ⦠H1 H2 ⦠Ha12) // -qed-. - -lemma TC_symmetric: âA,R. symmetric A R â symmetric A (TC ⦠R). -#A #R #HR #x #y #H @(TC_ind_dx ??????? H) -x /3 width=1/ /3 width=3/ -qed. - -lemma TC_star_ind_dx: âA,R. reflexive A R â - âa2. âP:predicate A. P a2 â - (âa1,a. R a1 a â TC ⦠R a a2 â P a â P a1) â - âa1. TC ⦠R a1 a2 â P a1. -#A #R #HR #a2 #P #Ha2 #H #a1 #Ha12 -@(TC_ind_dx ⦠P ? H ⦠Ha12) /3 width=4/ -qed-. - -lemma TC_Conf3: âA,B,S,R. Conf3 A B S R â Conf3 A B S (TC ⦠R). -#A #B #S #R #HSR #b #a1 #Ha1 #a2 #H elim H -a2 /2 width=3/ -qed. - -(* ************ confluence of star *****************) - -lemma star_strip: âA,R. confluent A R â - âa0,a1. star ⦠R a0 a1 â âa2. R a0 a2 â - ââa. R a1 a & star ⦠R a2 a. -#A #R #HR #a0 #a1 #H elim H -a1 /2 width=3/ -#a #a1 #_ #Ha1 #IHa0 #a2 #Ha02 -elim (IHa0 ⦠Ha02) -a0 #a0 #Ha0 #Ha20 -elim (HR ⦠Ha1 ⦠Ha0) -a /3 width=5/ -qed-. - -lemma star_confluent: âA,R. confluent A R â confluent A (star ⦠R). -#A #R #HR #a0 #a1 #H elim H -a1 /2 width=3/ -#a #a1 #_ #Ha1 #IHa0 #a2 #Ha02 -elim (IHa0 ⦠Ha02) -a0 #a0 #Ha0 #Ha20 -elim (star_strip ⦠HR ⦠Ha0 ⦠Ha1) -a /3 width=5/ -qed-. - -(* relations on unboxed pairs ***********************************************) - -inductive bi_TC (A,B) (R:bi_relation A B) (a:A) (b:B): relation2 A B â - |bi_inj : âc,d. R a b c d â bi_TC A B R a b c d - |bi_step: âc,d,e,f. bi_TC A B R a b c d â R c d e f â bi_TC A B R a b e f. - -lemma bi_TC_strap: âA,B. âR:bi_relation A B. âa1,a,a2,b1,b,b2. - R a1 b1 a b â bi_TC ⦠R a b a2 b2 â bi_TC ⦠R a1 b1 a2 b2. -#A #B #R #a1 #a #a2 #b1 #b #b2 #HR #H elim H -a2 -b2 /2 width=4/ /3 width=4/ -qed. - -lemma bi_TC_reflexive: âA,B,R. bi_reflexive A B R â - bi_reflexive ⦠(bi_TC ⦠R). -/2 width=1/ qed. - -inductive bi_TC_dx (A,B) (R:bi_relation A B): bi_relation A B â - |bi_inj_dx : âa1,a2,b1,b2. R a1 b1 a2 b2 â bi_TC_dx A B R a1 b1 a2 b2 - |bi_step_dx : âa1,a,a2,b1,b,b2. R a1 b1 a b â bi_TC_dx A B R a b a2 b2 â - bi_TC_dx A B R a1 b1 a2 b2. - -lemma bi_TC_dx_strap: âA,B. âR: bi_relation A B. - âa1,a,a2,b1,b,b2. bi_TC_dx A B R a1 b1 a b â - R a b a2 b2 â bi_TC_dx A B R a1 b1 a2 b2. -#A #B #R #a1 #a #a2 #b1 #b #b2 #H1 elim H1 -a -b /3 width=4/ -qed. - -lemma bi_TC_to_bi_TC_dx: âA,B. âR: bi_relation A B. - âa1,a2,b1,b2. bi_TC ⦠R a1 b1 a2 b2 â - bi_TC_dx ⦠R a1 b1 a2 b2. -#A #B #R #a1 #a2 #b1 #b2 #H12 elim H12 -a2 -b2 /2 width=4/ -qed. - -lemma bi_TC_dx_to_bi_TC: âA,B. âR: bi_relation A B. - âa1,a2,b1,b2. bi_TC_dx ⦠R a1 b1 a2 b2 â - bi_TC ⦠R a1 b1 a2 b2. -#A #B #R #a1 #a2 #b1 #b2 #H12 elim H12 -a1 -a2 -b1 -b2 /2 width=4/ -qed. - -fact bi_TC_ind_dx_aux: âA,B,R,a2,b2. âP:relation2 A B. - (âa1,b1. R a1 b1 a2 b2 â P a1 b1) â - (âa1,a,b1,b. R a1 b1 a b â bi_TC ⦠R a b a2 b2 â P a b â P a1 b1) â - âa1,a,b1,b. bi_TC ⦠R a1 b1 a b â a = a2 â b = b2 â P a1 b1. -#A #B #R #a2 #b2 #P #H1 #H2 #a1 #a #b1 #b #H1 -elim (bi_TC_to_bi_TC_dx ⦠a1 a b1 b H1) -a1 -a -b1 -b -[ #a1 #x #b1 #y #H1 #Hx #Hy destruct /2 width=1/ -| #a1 #a #x #b1 #b #y #H1 #H #IH #Hx #Hy destruct /3 width=5/ -] -qed-. - -lemma bi_TC_ind_dx: âA,B,R,a2,b2. âP:relation2 A B. - (âa1,b1. R a1 b1 a2 b2 â P a1 b1) â - (âa1,a,b1,b. R a1 b1 a b â bi_TC ⦠R a b a2 b2 â P a b â P a1 b1) â - âa1,b1. bi_TC ⦠R a1 b1 a2 b2 â P a1 b1. -#A #B #R #a2 #b2 #P #H1 #H2 #a1 #b1 #H12 -@(bi_TC_ind_dx_aux ?????? H1 H2 ⦠H12) // -qed-. - -lemma bi_TC_symmetric: âA,B,R. bi_symmetric A B R â - bi_symmetric A B (bi_TC ⦠R). -#A #B #R #HR #a1 #a2 #b1 #b2 #H21 -@(bi_TC_ind_dx ⦠a2 b2 H21) -a2 -b2 /3 width=1/ /3 width=4/ -qed. - -lemma bi_TC_transitive: âA,B,R. bi_transitive A B (bi_TC ⦠R). -#A #B #R #a1 #a #b1 #b #H elim H -a -b /2 width=4/ /3 width=4/ -qed. - -definition bi_Conf3: âA,B,C. relation3 A B C â predicate (bi_relation A B) â - λA,B,C,S,R. - âc,a1,b1. S a1 b1 c â âa2,b2. R a1 b1 a2 b2 â S a2 b2 c. - -lemma bi_TC_Conf3: âA,B,C,S,R. bi_Conf3 A B C S R â bi_Conf3 A B C S (bi_TC ⦠R). -#A #B #C #S #R #HSR #c #a1 #b1 #Hab1 #a2 #b2 #H elim H -a2 -b2 /2 width=4/ -qed. - -lemma bi_TC_star_ind: âA,B,R. bi_reflexive A B R â âa1,b1. âP:relation2 A B. - P a1 b1 â (âa,a2,b,b2. bi_TC ⦠R a1 b1 a b â R a b a2 b2 â P a b â P a2 b2) â - âa2,b2. bi_TC ⦠R a1 b1 a2 b2 â P a2 b2. -#A #B #R #HR #a1 #b1 #P #H1 #IH #a2 #b2 #H12 elim H12 -a2 -b2 /3 width=5/ -qed-. - -lemma bi_TC_star_ind_dx: âA,B,R. bi_reflexive A B R â - âa2,b2. âP:relation2 A B. P a2 b2 â - (âa1,a,b1,b. R a1 b1 a b â bi_TC ⦠R a b a2 b2 â P a b â P a1 b1) â - âa1,b1. bi_TC ⦠R a1 b1 a2 b2 â P a1 b1. -#A #B #R #HR #a2 #b2 #P #H2 #IH #a1 #b1 #H12 -@(bi_TC_ind_dx ⦠IH ⦠a1 b1 H12) /3 width=5/ -qed-. - -definition bi_star: âA,B,R. bi_relation A B â - λA,B,R. bi_RC A B (bi_TC ⦠R). - -lemma bi_star_bi_reflexive: âA,B,R. bi_reflexive A B (bi_star ⦠R). -/2 width=1/ qed. - -lemma bi_TC_to_bi_star: âA,B,R,a1,b1,a2,b2. - bi_TC A B R a1 b1 a2 b2 â bi_star A B R a1 b1 a2 b2. -/2 width=1/ qed. - -lemma bi_R_to_bi_star: âA,B,R,a1,b1,a2,b2. - R a1 b1 a2 b2 â bi_star A B R a1 b1 a2 b2. -/3 width=1/ qed. - -lemma bi_star_strap1: âA,B,R,a1,a,a2,b1,b,b2. bi_star A B R a1 b1 a b â - R a b a2 b2 â bi_star A B R a1 b1 a2 b2. -#A #B #R #a1 #a #a2 #b1 #b #b2 * -[ /3 width=4/ -| * #H1 #H2 destruct /2 width=1/ -] -qed. - -lemma bi_star_strap2: âA,B,R,a1,a,a2,b1,b,b2. R a1 b1 a b â - bi_star A B R a b a2 b2 â bi_star A B R a1 b1 a2 b2. -#A #B #R #a1 #a #a2 #b1 #b #b2 #H * -[ /3 width=4/ -| * #H1 #H2 destruct /2 width=1/ -] -qed. - -lemma bi_star_to_bi_TC_to_bi_TC: âA,B,R,a1,a,a2,b1,b,b2. bi_star A B R a1 b1 a b â - bi_TC A B R a b a2 b2 â bi_TC A B R a1 b1 a2 b2. -#A #B #R #a1 #a #a2 #b1 #b #b2 * -[ /2 width=4/ -| * #H1 #H2 destruct /2 width=1/ -] -qed. - -lemma bi_TC_to_bi_star_to_bi_TC: âA,B,R,a1,a,a2,b1,b,b2. bi_TC A B R a1 b1 a b â - bi_star A B R a b a2 b2 â bi_TC A B R a1 b1 a2 b2. -#A #B #R #a1 #a #a2 #b1 #b #b2 #H * -[ /2 width=4/ -| * #H1 #H2 destruct /2 width=1/ -] -qed. - -lemma bi_tansitive_bi_star: âA,B,R. bi_transitive A B (bi_star ⦠R). -#A #B #R #a1 #a #b1 #b #H #a2 #b2 * -[ /3 width=4/ -| * #H1 #H2 destruct /2 width=1/ -] -qed. - -lemma bi_star_ind: âA,B,R,a1,b1. âP:relation2 A B. P a1 b1 â - (âa,a2,b,b2. bi_star ⦠R a1 b1 a b â R a b a2 b2 â P a b â P a2 b2) â - âa2,b2. bi_star ⦠R a1 b1 a2 b2 â P a2 b2. -#A #B #R #a1 #b1 #P #H #IH #a2 #b2 * -[ #H12 elim H12 -a2 -b2 /2 width=5/ -H /3 width=5/ -| * #H1 #H2 destruct // -] -qed-. - -lemma bi_star_ind_dx: âA,B,R,a2,b2. âP:relation2 A B. P a2 b2 â - (âa1,a,b1,b. R a1 b1 a b â bi_star ⦠R a b a2 b2 â P a b â P a1 b1) â - âa1,b1. bi_star ⦠R a1 b1 a2 b2 â P a1 b1. -#A #B #R #a2 #b2 #P #H #IH #a1 #b1 * -[ #H12 @(bi_TC_ind_dx ⦠a1 b1 H12) -a1 -b1 /2 width=5/ -H /3 width=5/ -| * #H1 #H2 destruct // -] -qed-. - -(* relations on unboxed triples *********************************************) - -inductive tri_TC (A,B,C) (R:tri_relation A B C) (a1:A) (b1:B) (c1:C): relation3 A B C â - |tri_inj : âa2,b2,c2. R a1 b1 c1 a2 b2 c2 â tri_TC A B C R a1 b1 c1 a2 b2 c2 - |tri_step: âa,a2,b,b2,c,c2. - tri_TC A B C R a1 b1 c1 a b c â R a b c a2 b2 c2 â - tri_TC A B C R a1 b1 c1 a2 b2 c2. - -lemma tri_TC_strap: âA,B,C. âR:tri_relation A B C. âa1,a,a2,b1,b,b2,c1,c,c2. - R a1 b1 c1 a b c â tri_TC ⦠R a b c a2 b2 c2 â - tri_TC ⦠R a1 b1 c1 a2 b2 c2. -#A #B #C #R #a1 #a #a2 #b1 #b #b2 #c1 #c #c2 #HR #H elim H -a2 -b2 -c2 /2 width=5/ /3 width=5/ -qed. - -lemma tri_TC_reflexive: âA,B,C,R. tri_reflexive A B C R â - tri_reflexive ⦠(tri_TC ⦠R). -/2 width=1/ qed. - -inductive tri_TC_dx (A,B,C) (R:tri_relation A B C): tri_relation A B C â - |tri_inj_dx : âa1,a2,b1,b2,c1,c2. R a1 b1 c1 a2 b2 c2 â tri_TC_dx A B C R a1 b1 c1 a2 b2 c2 - |tri_step_dx : âa1,a,a2,b1,b,b2,c1,c,c2. - R a1 b1 c1 a b c â tri_TC_dx A B C R a b c a2 b2 c2 â - tri_TC_dx A B C R a1 b1 c1 a2 b2 c2. - -lemma tri_TC_dx_strap: âA,B,C. âR: tri_relation A B C. - âa1,a,a2,b1,b,b2,c1,c,c2. - tri_TC_dx A B C R a1 b1 c1 a b c â - R a b c a2 b2 c2 â tri_TC_dx A B C R a1 b1 c1 a2 b2 c2. -#A #B #C #R #a1 #a #a2 #b1 #b #b2 #c1 #c #c2 #H1 elim H1 -a -b -c /3 width=5/ -qed. - -lemma tri_TC_to_tri_TC_dx: âA,B,C. âR: tri_relation A B C. - âa1,a2,b1,b2,c1,c2. tri_TC ⦠R a1 b1 c1 a2 b2 c2 â - tri_TC_dx ⦠R a1 b1 c1 a2 b2 c2. -#A #B #C #R #a1 #a2 #b1 #b2 #c1 #c2 #H12 elim H12 -a2 -b2 -c2 /2 width=1/ /2 width=5/ -qed. - -lemma tri_TC_dx_to_tri_TC: âA,B,C. âR: tri_relation A B C. - âa1,a2,b1,b2,c1,c2. tri_TC_dx ⦠R a1 b1 c1 a2 b2 c2 â - tri_TC ⦠R a1 b1 c1 a2 b2 c2. -#A #B #C #R #a1 #a2 #b1 #b2 #c1 #c2 #H12 elim H12 -a1 -a2 -b1 -b2 -c1 -c2 -/2 width=1/ /2 width=5/ -qed. - -fact tri_TC_ind_dx_aux: âA,B,C,R,a2,b2,c2. âP:relation3 A B C. - (âa1,b1,c1. R a1 b1 c1 a2 b2 c2â P a1 b1 c1) â - (âa1,a,b1,b,c1,c. R a1 b1 c1 a b c â tri_TC ⦠R a b c a2 b2 c2 â P a b c â P a1 b1 c1) â - âa1,a,b1,b,c1,c. tri_TC ⦠R a1 b1 c1 a b c â a = a2 â b = b2 â c = c2 â P a1 b1 c1. -#A #B #C #R #a2 #b2 #c2 #P #H1 #H2 #a1 #a #b1 #b #c1 #c #H1 -elim (tri_TC_to_tri_TC_dx ⦠a1 a b1 b c1 c H1) -a1 -a -b1 -b -c1 -c -[ #a1 #x #b1 #y #c1 #z #H1 #Hx #Hy #Hz destruct /2 width=1/ -| #a1 #a #x #b1 #b #y #c1 #c #z #H1 #H #IH #Hx #Hy #Hz destruct /3 width=6/ -] -qed-. - -lemma tri_TC_ind_dx: âA,B,C,R,a2,b2,c2. âP:relation3 A B C. - (âa1,b1,c1. R a1 b1 c1 a2 b2 c2 â P a1 b1 c1) â - (âa1,a,b1,b,c1,c. R a1 b1 c1 a b c â tri_TC ⦠R a b c a2 b2 c2 â P a b c â P a1 b1 c1) â - âa1,b1,c1. tri_TC ⦠R a1 b1 c1 a2 b2 c2 â P a1 b1 c1. -#A #B #C #R #a2 #b2 #c2 #P #H1 #H2 #a1 #b1 #c1 #H12 -@(tri_TC_ind_dx_aux ???????? H1 H2 ⦠H12) // -qed-. - -lemma tri_TC_symmetric: âA,B,C,R. tri_symmetric A B C R â - tri_symmetric ⦠(tri_TC ⦠R). -#A #B #C #R #HR #a1 #a2 #b1 #b2 #c1 #c2 #H21 -@(tri_TC_ind_dx ⦠a2 b2 c2 H21) -a2 -b2 -c2 /3 width=1/ /3 width=5/ -qed. - -lemma tri_TC_transitive: âA,B,C,R. tri_transitive A B C (tri_TC ⦠R). -#A #B #C #R #a1 #a #b1 #b #c1 #c #H elim H -a -b -c /2 width=5/ /3 width=5/ -qed. - -definition tri_Conf4: âA,B,C,D. relation4 A B C D â predicate (tri_relation A B C) â - λA,B,C,D,S,R. - âd,a1,b1,c1. S a1 b1 c1 d â âa2,b2,c2. R a1 b1 c1 a2 b2 c2 â S a2 b2 c2 d. - -lemma tri_TC_Conf4: âA,B,C,D,S,R. - tri_Conf4 A B C D S R â tri_Conf4 A B C D S (tri_TC ⦠R). -#A #B #C #D #S #R #HSR #d #a1 #b1 #c1 #Habc1 #a2 #b2 #c2 #H elim H -a2 -b2 -c2 -/2 width=5/ -qed. - -lemma tri_TC_star_ind: âA,B,C,R. tri_reflexive A B C R â - âa1,b1,c1. âP:relation3 A B C. - P a1 b1 c1 â (âa,a2,b,b2,c,c2. tri_TC ⦠R a1 b1 c1 a b c â R a b c a2 b2 c2 â P a b c â P a2 b2 c2) â - âa2,b2,c2. tri_TC ⦠R a1 b1 c1 a2 b2 c2 â P a2 b2 c2. -#A #B #C #R #HR #a1 #b1 #c1 #P #H1 #IH #a2 #b2 #c2 #H12 elim H12 -a2 -b2 -c2 -/2 width=6/ /3 width=6/ -qed-. - -lemma tri_TC_star_ind_dx: âA,B,C,R. tri_reflexive A B C R â - âa2,b2,c2. âP:relation3 A B C. P a2 b2 c2 â - (âa1,a,b1,b,c1,c. R a1 b1 c1 a b c â tri_TC ⦠R a b c a2 b2 c2 â P a b c â P a1 b1 c1) â - âa1,b1,c1. tri_TC ⦠R a1 b1 c1 a2 b2 c2 â P a1 b1 c1. -#A #B #C #R #HR #a2 #b2 #c2 #P #H2 #IH #a1 #b1 #c1 #H12 -@(tri_TC_ind_dx ⦠IH ⦠a1 b1 c1 H12) /3 width=6/ -qed-. diff --git a/matita/matita/lib/basics/star1.ma b/matita/matita/lib/basics/star1.ma new file mode 100644 index 000000000..79da56c47 --- /dev/null +++ b/matita/matita/lib/basics/star1.ma @@ -0,0 +1,358 @@ +(* + ||M|| This file is part of HELM, an Hypertextual, Electronic + ||A|| Library of Mathematics, developed at the Computer Science + ||T|| Department of the University of Bologna, Italy. + ||I|| + ||T|| + ||A|| This file is distributed under the terms of the + \ / GNU General Public License Version 2 + \ / + V_______________________________________________________________ *) + +include "basics/star.ma". + +(* added from λδ *) + +lemma TC_strap: âA. âR:relation A. âa1,a,a2. + R a1 a â TC ⦠R a a2 â TC ⦠R a1 a2. +/3 width=3/ qed. + +lemma TC_reflexive: âA,R. reflexive A R â reflexive A (TC ⦠R). +/2 width=1/ qed. + +lemma TC_star_ind: âA,R. reflexive A R â âa1. âP:predicate A. + P a1 â (âa,a2. TC ⦠R a1 a â R a a2 â P a â P a2) â + âa2. TC ⦠R a1 a2 â P a2. +#A #R #H #a1 #P #Ha1 #IHa1 #a2 #Ha12 elim Ha12 -a2 /3 width=4/ +qed-. + +inductive TC_dx (A:Type[0]) (R:relation A): A â A â Prop â + |inj_dx: âa,c. R a c â TC_dx A R a c + |step_dx : âa,b,c. R a b â TC_dx A R b c â TC_dx A R a c. + +lemma TC_dx_strap: âA. âR: relation A. + âa,b,c. TC_dx A R a b â R b c â TC_dx A R a c. +#A #R #a #b #c #Hab elim Hab -a -b /3 width=3/ +qed. + +lemma TC_to_TC_dx: âA. âR: relation A. + âa1,a2. TC ⦠R a1 a2 â TC_dx ⦠R a1 a2. +#A #R #a1 #a2 #Ha12 elim Ha12 -a2 /2 width=3/ +qed. + +lemma TC_dx_to_TC: âA. âR: relation A. + âa1,a2. TC_dx ⦠R a1 a2 â TC ⦠R a1 a2. +#A #R #a1 #a2 #Ha12 elim Ha12 -a1 -a2 /2 width=3/ +qed. + +fact TC_ind_dx_aux: âA,R,a2. âP:predicate A. + (âa1. R a1 a2 â P a1) â + (âa1,a. R a1 a â TC ⦠R a a2 â P a â P a1) â + âa1,a. TC ⦠R a1 a â a = a2 â P a1. +#A #R #a2 #P #H1 #H2 #a1 #a #Ha1 +elim (TC_to_TC_dx ???? Ha1) -a1 -a +[ #a #c #Hac #H destruct /2 width=1/ +| #a #b #c #Hab #Hbc #IH #H destruct /3 width=4/ +] +qed-. + +lemma TC_ind_dx: âA,R,a2. âP:predicate A. + (âa1. R a1 a2 â P a1) â + (âa1,a. R a1 a â TC ⦠R a a2 â P a â P a1) â + âa1. TC ⦠R a1 a2 â P a1. +#A #R #a2 #P #H1 #H2 #a1 #Ha12 +@(TC_ind_dx_aux ⦠H1 H2 ⦠Ha12) // +qed-. + +lemma TC_symmetric: âA,R. symmetric A R â symmetric A (TC ⦠R). +#A #R #HR #x #y #H @(TC_ind_dx ??????? H) -x /3 width=1/ /3 width=3/ +qed. + +lemma TC_star_ind_dx: âA,R. reflexive A R â + âa2. âP:predicate A. P a2 â + (âa1,a. R a1 a â TC ⦠R a a2 â P a â P a1) â + âa1. TC ⦠R a1 a2 â P a1. +#A #R #HR #a2 #P #Ha2 #H #a1 #Ha12 +@(TC_ind_dx ⦠P ? H ⦠Ha12) /3 width=4/ +qed-. + +lemma TC_Conf3: âA,B,S,R. Conf3 A B S R â Conf3 A B S (TC ⦠R). +#A #B #S #R #HSR #b #a1 #Ha1 #a2 #H elim H -a2 /2 width=3/ +qed. + +(* ************ confluence of star *****************) + +lemma star_strip: âA,R. confluent A R â + âa0,a1. star ⦠R a0 a1 â âa2. R a0 a2 â + ââa. R a1 a & star ⦠R a2 a. +#A #R #HR #a0 #a1 #H elim H -a1 /2 width=3/ +#a #a1 #_ #Ha1 #IHa0 #a2 #Ha02 +elim (IHa0 ⦠Ha02) -a0 #a0 #Ha0 #Ha20 +elim (HR ⦠Ha1 ⦠Ha0) -a /3 width=5/ +qed-. + +lemma star_confluent: âA,R. confluent A R â confluent A (star ⦠R). +#A #R #HR #a0 #a1 #H elim H -a1 /2 width=3/ +#a #a1 #_ #Ha1 #IHa0 #a2 #Ha02 +elim (IHa0 ⦠Ha02) -a0 #a0 #Ha0 #Ha20 +elim (star_strip ⦠HR ⦠Ha0 ⦠Ha1) -a /3 width=5/ +qed-. + +(* relations on unboxed pairs ***********************************************) + +inductive bi_TC (A,B) (R:bi_relation A B) (a:A) (b:B): relation2 A B â + |bi_inj : âc,d. R a b c d â bi_TC A B R a b c d + |bi_step: âc,d,e,f. bi_TC A B R a b c d â R c d e f â bi_TC A B R a b e f. + +lemma bi_TC_strap: âA,B. âR:bi_relation A B. âa1,a,a2,b1,b,b2. + R a1 b1 a b â bi_TC ⦠R a b a2 b2 â bi_TC ⦠R a1 b1 a2 b2. +#A #B #R #a1 #a #a2 #b1 #b #b2 #HR #H elim H -a2 -b2 /2 width=4/ /3 width=4/ +qed. + +lemma bi_TC_reflexive: âA,B,R. bi_reflexive A B R â + bi_reflexive ⦠(bi_TC ⦠R). +/2 width=1/ qed. + +inductive bi_TC_dx (A,B) (R:bi_relation A B): bi_relation A B â + |bi_inj_dx : âa1,a2,b1,b2. R a1 b1 a2 b2 â bi_TC_dx A B R a1 b1 a2 b2 + |bi_step_dx : âa1,a,a2,b1,b,b2. R a1 b1 a b â bi_TC_dx A B R a b a2 b2 â + bi_TC_dx A B R a1 b1 a2 b2. + +lemma bi_TC_dx_strap: âA,B. âR: bi_relation A B. + âa1,a,a2,b1,b,b2. bi_TC_dx A B R a1 b1 a b â + R a b a2 b2 â bi_TC_dx A B R a1 b1 a2 b2. +#A #B #R #a1 #a #a2 #b1 #b #b2 #H1 elim H1 -a -b /3 width=4/ +qed. + +lemma bi_TC_to_bi_TC_dx: âA,B. âR: bi_relation A B. + âa1,a2,b1,b2. bi_TC ⦠R a1 b1 a2 b2 â + bi_TC_dx ⦠R a1 b1 a2 b2. +#A #B #R #a1 #a2 #b1 #b2 #H12 elim H12 -a2 -b2 /2 width=4/ +qed. + +lemma bi_TC_dx_to_bi_TC: âA,B. âR: bi_relation A B. + âa1,a2,b1,b2. bi_TC_dx ⦠R a1 b1 a2 b2 â + bi_TC ⦠R a1 b1 a2 b2. +#A #B #R #a1 #a2 #b1 #b2 #H12 elim H12 -a1 -a2 -b1 -b2 /2 width=4/ +qed. + +fact bi_TC_ind_dx_aux: âA,B,R,a2,b2. âP:relation2 A B. + (âa1,b1. R a1 b1 a2 b2 â P a1 b1) â + (âa1,a,b1,b. R a1 b1 a b â bi_TC ⦠R a b a2 b2 â P a b â P a1 b1) â + âa1,a,b1,b. bi_TC ⦠R a1 b1 a b â a = a2 â b = b2 â P a1 b1. +#A #B #R #a2 #b2 #P #H1 #H2 #a1 #a #b1 #b #H1 +elim (bi_TC_to_bi_TC_dx ⦠a1 a b1 b H1) -a1 -a -b1 -b +[ #a1 #x #b1 #y #H1 #Hx #Hy destruct /2 width=1/ +| #a1 #a #x #b1 #b #y #H1 #H #IH #Hx #Hy destruct /3 width=5/ +] +qed-. + +lemma bi_TC_ind_dx: âA,B,R,a2,b2. âP:relation2 A B. + (âa1,b1. R a1 b1 a2 b2 â P a1 b1) â + (âa1,a,b1,b. R a1 b1 a b â bi_TC ⦠R a b a2 b2 â P a b â P a1 b1) â + âa1,b1. bi_TC ⦠R a1 b1 a2 b2 â P a1 b1. +#A #B #R #a2 #b2 #P #H1 #H2 #a1 #b1 #H12 +@(bi_TC_ind_dx_aux ?????? H1 H2 ⦠H12) // +qed-. + +lemma bi_TC_symmetric: âA,B,R. bi_symmetric A B R â + bi_symmetric A B (bi_TC ⦠R). +#A #B #R #HR #a1 #a2 #b1 #b2 #H21 +@(bi_TC_ind_dx ⦠a2 b2 H21) -a2 -b2 /3 width=1/ /3 width=4/ +qed. + +lemma bi_TC_transitive: âA,B,R. bi_transitive A B (bi_TC ⦠R). +#A #B #R #a1 #a #b1 #b #H elim H -a -b /2 width=4/ /3 width=4/ +qed. + +definition bi_Conf3: âA,B,C. relation3 A B C â predicate (bi_relation A B) â + λA,B,C,S,R. + âc,a1,b1. S a1 b1 c â âa2,b2. R a1 b1 a2 b2 â S a2 b2 c. + +lemma bi_TC_Conf3: âA,B,C,S,R. bi_Conf3 A B C S R â bi_Conf3 A B C S (bi_TC ⦠R). +#A #B #C #S #R #HSR #c #a1 #b1 #Hab1 #a2 #b2 #H elim H -a2 -b2 /2 width=4/ +qed. + +lemma bi_TC_star_ind: âA,B,R. bi_reflexive A B R â âa1,b1. âP:relation2 A B. + P a1 b1 â (âa,a2,b,b2. bi_TC ⦠R a1 b1 a b â R a b a2 b2 â P a b â P a2 b2) â + âa2,b2. bi_TC ⦠R a1 b1 a2 b2 â P a2 b2. +#A #B #R #HR #a1 #b1 #P #H1 #IH #a2 #b2 #H12 elim H12 -a2 -b2 /3 width=5/ +qed-. + +lemma bi_TC_star_ind_dx: âA,B,R. bi_reflexive A B R â + âa2,b2. âP:relation2 A B. P a2 b2 â + (âa1,a,b1,b. R a1 b1 a b â bi_TC ⦠R a b a2 b2 â P a b â P a1 b1) â + âa1,b1. bi_TC ⦠R a1 b1 a2 b2 â P a1 b1. +#A #B #R #HR #a2 #b2 #P #H2 #IH #a1 #b1 #H12 +@(bi_TC_ind_dx ⦠IH ⦠a1 b1 H12) /3 width=5/ +qed-. + +definition bi_star: âA,B,R. bi_relation A B â + λA,B,R. bi_RC A B (bi_TC ⦠R). + +lemma bi_star_bi_reflexive: âA,B,R. bi_reflexive A B (bi_star ⦠R). +/2 width=1/ qed. + +lemma bi_TC_to_bi_star: âA,B,R,a1,b1,a2,b2. + bi_TC A B R a1 b1 a2 b2 â bi_star A B R a1 b1 a2 b2. +/2 width=1/ qed. + +lemma bi_R_to_bi_star: âA,B,R,a1,b1,a2,b2. + R a1 b1 a2 b2 â bi_star A B R a1 b1 a2 b2. +/3 width=1/ qed. + +lemma bi_star_strap1: âA,B,R,a1,a,a2,b1,b,b2. bi_star A B R a1 b1 a b â + R a b a2 b2 â bi_star A B R a1 b1 a2 b2. +#A #B #R #a1 #a #a2 #b1 #b #b2 * +[ /3 width=4/ +| * #H1 #H2 destruct /2 width=1/ +] +qed. + +lemma bi_star_strap2: âA,B,R,a1,a,a2,b1,b,b2. R a1 b1 a b â + bi_star A B R a b a2 b2 â bi_star A B R a1 b1 a2 b2. +#A #B #R #a1 #a #a2 #b1 #b #b2 #H * +[ /3 width=4/ +| * #H1 #H2 destruct /2 width=1/ +] +qed. + +lemma bi_star_to_bi_TC_to_bi_TC: âA,B,R,a1,a,a2,b1,b,b2. bi_star A B R a1 b1 a b â + bi_TC A B R a b a2 b2 â bi_TC A B R a1 b1 a2 b2. +#A #B #R #a1 #a #a2 #b1 #b #b2 * +[ /2 width=4/ +| * #H1 #H2 destruct /2 width=1/ +] +qed. + +lemma bi_TC_to_bi_star_to_bi_TC: âA,B,R,a1,a,a2,b1,b,b2. bi_TC A B R a1 b1 a b â + bi_star A B R a b a2 b2 â bi_TC A B R a1 b1 a2 b2. +#A #B #R #a1 #a #a2 #b1 #b #b2 #H * +[ /2 width=4/ +| * #H1 #H2 destruct /2 width=1/ +] +qed. + +lemma bi_tansitive_bi_star: âA,B,R. bi_transitive A B (bi_star ⦠R). +#A #B #R #a1 #a #b1 #b #H #a2 #b2 * +[ /3 width=4/ +| * #H1 #H2 destruct /2 width=1/ +] +qed. + +lemma bi_star_ind: âA,B,R,a1,b1. âP:relation2 A B. P a1 b1 â + (âa,a2,b,b2. bi_star ⦠R a1 b1 a b â R a b a2 b2 â P a b â P a2 b2) â + âa2,b2. bi_star ⦠R a1 b1 a2 b2 â P a2 b2. +#A #B #R #a1 #b1 #P #H #IH #a2 #b2 * +[ #H12 elim H12 -a2 -b2 /2 width=5/ -H /3 width=5/ +| * #H1 #H2 destruct // +] +qed-. + +lemma bi_star_ind_dx: âA,B,R,a2,b2. âP:relation2 A B. P a2 b2 â + (âa1,a,b1,b. R a1 b1 a b â bi_star ⦠R a b a2 b2 â P a b â P a1 b1) â + âa1,b1. bi_star ⦠R a1 b1 a2 b2 â P a1 b1. +#A #B #R #a2 #b2 #P #H #IH #a1 #b1 * +[ #H12 @(bi_TC_ind_dx ⦠a1 b1 H12) -a1 -b1 /2 width=5/ -H /3 width=5/ +| * #H1 #H2 destruct // +] +qed-. + +(* relations on unboxed triples *********************************************) + +inductive tri_TC (A,B,C) (R:tri_relation A B C) (a1:A) (b1:B) (c1:C): relation3 A B C â + |tri_inj : âa2,b2,c2. R a1 b1 c1 a2 b2 c2 â tri_TC A B C R a1 b1 c1 a2 b2 c2 + |tri_step: âa,a2,b,b2,c,c2. + tri_TC A B C R a1 b1 c1 a b c â R a b c a2 b2 c2 â + tri_TC A B C R a1 b1 c1 a2 b2 c2. + +lemma tri_TC_strap: âA,B,C. âR:tri_relation A B C. âa1,a,a2,b1,b,b2,c1,c,c2. + R a1 b1 c1 a b c â tri_TC ⦠R a b c a2 b2 c2 â + tri_TC ⦠R a1 b1 c1 a2 b2 c2. +#A #B #C #R #a1 #a #a2 #b1 #b #b2 #c1 #c #c2 #HR #H elim H -a2 -b2 -c2 /2 width=5/ /3 width=5/ +qed. + +lemma tri_TC_reflexive: âA,B,C,R. tri_reflexive A B C R â + tri_reflexive ⦠(tri_TC ⦠R). +/2 width=1/ qed. + +inductive tri_TC_dx (A,B,C) (R:tri_relation A B C): tri_relation A B C â + |tri_inj_dx : âa1,a2,b1,b2,c1,c2. R a1 b1 c1 a2 b2 c2 â tri_TC_dx A B C R a1 b1 c1 a2 b2 c2 + |tri_step_dx : âa1,a,a2,b1,b,b2,c1,c,c2. + R a1 b1 c1 a b c â tri_TC_dx A B C R a b c a2 b2 c2 â + tri_TC_dx A B C R a1 b1 c1 a2 b2 c2. + +lemma tri_TC_dx_strap: âA,B,C. âR: tri_relation A B C. + âa1,a,a2,b1,b,b2,c1,c,c2. + tri_TC_dx A B C R a1 b1 c1 a b c â + R a b c a2 b2 c2 â tri_TC_dx A B C R a1 b1 c1 a2 b2 c2. +#A #B #C #R #a1 #a #a2 #b1 #b #b2 #c1 #c #c2 #H1 elim H1 -a -b -c /3 width=5/ +qed. + +lemma tri_TC_to_tri_TC_dx: âA,B,C. âR: tri_relation A B C. + âa1,a2,b1,b2,c1,c2. tri_TC ⦠R a1 b1 c1 a2 b2 c2 â + tri_TC_dx ⦠R a1 b1 c1 a2 b2 c2. +#A #B #C #R #a1 #a2 #b1 #b2 #c1 #c2 #H12 elim H12 -a2 -b2 -c2 /2 width=1/ /2 width=5/ +qed. + +lemma tri_TC_dx_to_tri_TC: âA,B,C. âR: tri_relation A B C. + âa1,a2,b1,b2,c1,c2. tri_TC_dx ⦠R a1 b1 c1 a2 b2 c2 â + tri_TC ⦠R a1 b1 c1 a2 b2 c2. +#A #B #C #R #a1 #a2 #b1 #b2 #c1 #c2 #H12 elim H12 -a1 -a2 -b1 -b2 -c1 -c2 +/2 width=1/ /2 width=5/ +qed. + +fact tri_TC_ind_dx_aux: âA,B,C,R,a2,b2,c2. âP:relation3 A B C. + (âa1,b1,c1. R a1 b1 c1 a2 b2 c2â P a1 b1 c1) â + (âa1,a,b1,b,c1,c. R a1 b1 c1 a b c â tri_TC ⦠R a b c a2 b2 c2 â P a b c â P a1 b1 c1) â + âa1,a,b1,b,c1,c. tri_TC ⦠R a1 b1 c1 a b c â a = a2 â b = b2 â c = c2 â P a1 b1 c1. +#A #B #C #R #a2 #b2 #c2 #P #H1 #H2 #a1 #a #b1 #b #c1 #c #H1 +elim (tri_TC_to_tri_TC_dx ⦠a1 a b1 b c1 c H1) -a1 -a -b1 -b -c1 -c +[ #a1 #x #b1 #y #c1 #z #H1 #Hx #Hy #Hz destruct /2 width=1/ +| #a1 #a #x #b1 #b #y #c1 #c #z #H1 #H #IH #Hx #Hy #Hz destruct /3 width=6/ +] +qed-. + +lemma tri_TC_ind_dx: âA,B,C,R,a2,b2,c2. âP:relation3 A B C. + (âa1,b1,c1. R a1 b1 c1 a2 b2 c2 â P a1 b1 c1) â + (âa1,a,b1,b,c1,c. R a1 b1 c1 a b c â tri_TC ⦠R a b c a2 b2 c2 â P a b c â P a1 b1 c1) â + âa1,b1,c1. tri_TC ⦠R a1 b1 c1 a2 b2 c2 â P a1 b1 c1. +#A #B #C #R #a2 #b2 #c2 #P #H1 #H2 #a1 #b1 #c1 #H12 +@(tri_TC_ind_dx_aux ???????? H1 H2 ⦠H12) // +qed-. + +lemma tri_TC_symmetric: âA,B,C,R. tri_symmetric A B C R â + tri_symmetric ⦠(tri_TC ⦠R). +#A #B #C #R #HR #a1 #a2 #b1 #b2 #c1 #c2 #H21 +@(tri_TC_ind_dx ⦠a2 b2 c2 H21) -a2 -b2 -c2 /3 width=1/ /3 width=5/ +qed. + +lemma tri_TC_transitive: âA,B,C,R. tri_transitive A B C (tri_TC ⦠R). +#A #B #C #R #a1 #a #b1 #b #c1 #c #H elim H -a -b -c /2 width=5/ /3 width=5/ +qed. + +definition tri_Conf4: âA,B,C,D. relation4 A B C D â predicate (tri_relation A B C) â + λA,B,C,D,S,R. + âd,a1,b1,c1. S a1 b1 c1 d â âa2,b2,c2. R a1 b1 c1 a2 b2 c2 â S a2 b2 c2 d. + +lemma tri_TC_Conf4: âA,B,C,D,S,R. + tri_Conf4 A B C D S R â tri_Conf4 A B C D S (tri_TC ⦠R). +#A #B #C #D #S #R #HSR #d #a1 #b1 #c1 #Habc1 #a2 #b2 #c2 #H elim H -a2 -b2 -c2 +/2 width=5/ +qed. + +lemma tri_TC_star_ind: âA,B,C,R. tri_reflexive A B C R â + âa1,b1,c1. âP:relation3 A B C. + P a1 b1 c1 â (âa,a2,b,b2,c,c2. tri_TC ⦠R a1 b1 c1 a b c â R a b c a2 b2 c2 â P a b c â P a2 b2 c2) â + âa2,b2,c2. tri_TC ⦠R a1 b1 c1 a2 b2 c2 â P a2 b2 c2. +#A #B #C #R #HR #a1 #b1 #c1 #P #H1 #IH #a2 #b2 #c2 #H12 elim H12 -a2 -b2 -c2 +/2 width=6/ /3 width=6/ +qed-. + +lemma tri_TC_star_ind_dx: âA,B,C,R. tri_reflexive A B C R â + âa2,b2,c2. âP:relation3 A B C. P a2 b2 c2 â + (âa1,a,b1,b,c1,c. R a1 b1 c1 a b c â tri_TC ⦠R a b c a2 b2 c2 â P a b c â P a1 b1 c1) â + âa1,b1,c1. tri_TC ⦠R a1 b1 c1 a2 b2 c2 â P a1 b1 c1. +#A #B #C #R #HR #a2 #b2 #c2 #P #H2 #IH #a1 #b1 #c1 #H12 +@(tri_TC_ind_dx ⦠IH ⦠a1 b1 c1 H12) /3 width=6/ +qed-.