X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fstatic%2Flfxs_fsle.ma;h=011748dbfd44b20e4c68115c2a9645b2151ce0ca;hp=cc7218982c531e63403457063615dcccef4cb256;hb=f7296f9cf2ee73465a374942c46b138f35c42ccb;hpb=e8971d3db8935436e6dddc27fe1ae168c7f3b315 diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_fsle.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_fsle.ma index cc7218982..011748dbf 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_fsle.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_fsle.ma @@ -13,21 +13,24 @@ (**************************************************************************) include "basic_2/relocation/lexs_length.ma". -include "basic_2/static/frees_drops.ma". include "basic_2/static/fsle_fsle.ma". +include "basic_2/static/lfxs_drops.ma". include "basic_2/static/lfxs_lfxs.ma". (* GENERIC EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ****) -definition R_fsle_compatible: predicate (relation3 …) ≝ λRN. +definition R_fsge_compatible: predicate (relation3 …) ≝ λRN. ∀L,T1,T2. RN L T1 T2 → ⦃L, T2⦄ ⊆ ⦃L, T1⦄. -definition lfxs_fsle_compatible: predicate (relation3 …) ≝ λRN. +definition lfxs_fsge_compatible: predicate (relation3 …) ≝ λRN. ∀L1,L2,T. L1 ⪤*[RN, T] L2 → ⦃L2, T⦄ ⊆ ⦃L1, T⦄. +definition lfxs_fsle_compatible: predicate (relation3 …) ≝ λRN. + ∀L1,L2,T. L1 ⪤*[RN, T] L2 → ⦃L1, T⦄ ⊆ ⦃L2, T⦄. + (* Basic inversions with free variables inclusion for restricted closures ***) -lemma frees_lexs_conf: ∀R. lfxs_fsle_compatible R → +lemma frees_lexs_conf: ∀R. lfxs_fsge_compatible R → ∀L1,T,f1. L1 ⊢ 𝐅*⦃T⦄ ≡ f1 → ∀L2. L1 ⪤*[cext2 R, cfull, f1] L2 → ∃∃f2. L2 ⊢ 𝐅*⦃T⦄ ≡ f2 & f2 ⊆ f1. @@ -39,14 +42,14 @@ qed-. (* Properties with free variables inclusion for restricted closures *********) (* Note: we just need lveq_inv_refl: ∀L,n1,n2. L ≋ⓧ*[n1, n2] L → ∧∧ 0 = n1 & 0 = n2 *) -lemma fsle_lfxs_trans: ∀R,L1,T1,T2. ⦃L1, T1⦄ ⊆ ⦃L1, T2⦄ → +lemma fsge_lfxs_trans: ∀R,L1,T1,T2. ⦃L1, T1⦄ ⊆ ⦃L1, T2⦄ → ∀L2. L1 ⪤*[R, T2] L2 → L1 ⪤*[R, T1] L2. #R #L1 #T1 #T2 * #n1 #n2 #f1 #f2 #Hf1 #Hf2 #Hn #Hf #L2 #HL12 elim (lveq_inj_length … Hn ?) // #H1 #H2 destruct /4 width=5 by lfxs_inv_frees, sle_lexs_trans, ex2_intro/ qed-. -lemma lfxs_sym: ∀R. lfxs_fsle_compatible R → +lemma lfxs_sym: ∀R. lfxs_fsge_compatible R → (∀L1,L2,T1,T2. R L1 T1 T2 → R L2 T2 T1) → ∀T. symmetric … (lfxs R T). #R #H1R #H2R #T #L1 #L2 @@ -56,7 +59,7 @@ elim (frees_lexs_conf … Hf1 … HL12) -Hf1 // qed-. lemma lfxs_pair_sn_split: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) → - lfxs_fsle_compatible R1 → + lfxs_fsge_compatible R1 → ∀L1,L2,V. L1 ⪤*[R1, V] L2 → ∀I,T. ∃∃L. L1 ⪤*[R1, ②{I}V.T] L & L ⪤*[R2, V] L2. #R1 #R2 #HR1 #HR2 #HR #L1 #L2 #V * #f #Hf #HL12 * [ #p ] #I #T @@ -75,7 +78,7 @@ elim (frees_lexs_conf … Hf … H) -Hf -H qed-. lemma lfxs_flat_dx_split: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) → - lfxs_fsle_compatible R1 → + lfxs_fsge_compatible R1 → ∀L1,L2,T. L1 ⪤*[R1, T] L2 → ∀I,V. ∃∃L. L1 ⪤*[R1, ⓕ{I}V.T] L & L ⪤*[R2, T] L2. #R1 #R2 #HR1 #HR2 #HR #L1 #L2 #T * #f #Hf #HL12 #I #V @@ -91,7 +94,7 @@ elim (frees_lexs_conf … Hf … H) -Hf -H qed-. lemma lfxs_bind_dx_split: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) → - lfxs_fsle_compatible R1 → + lfxs_fsge_compatible R1 → ∀I,L1,L2,V1,T. L1.ⓑ{I}V1 ⪤*[R1, T] L2 → ∀p. ∃∃L,V. L1 ⪤*[R1, ⓑ{p,I}V1.T] L & L.ⓑ{I}V ⪤*[R2, T] L2 & R1 L1 V1 V. #R1 #R2 #HR1 #HR2 #HR #I #L1 #L2 #V1 #T * #f #Hf #HL12 #p @@ -111,7 +114,7 @@ elim (frees_lexs_conf … Hf … H0) -Hf -H0 qed-. lemma lfxs_bind_dx_split_void: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) → - lfxs_fsle_compatible R1 → + lfxs_fsge_compatible R1 → ∀L1,L2,T. L1.ⓧ ⪤*[R1, T] L2 → ∀p,I,V. ∃∃L. L1 ⪤*[R1, ⓑ{p,I}V.T] L & L.ⓧ ⪤*[R2, T] L2. #R1 #R2 #HR1 #HR2 #HR #L1 #L2 #T * #f #Hf #HL12 #p #I #V @@ -133,8 +136,8 @@ qed-. (* Main properties with free variables inclusion for restricted closures ****) theorem lfxs_conf: ∀R1,R2. - lfxs_fsle_compatible R1 → - lfxs_fsle_compatible R2 → + lfxs_fsge_compatible R1 → + lfxs_fsge_compatible R2 → R_confluent2_lfxs R1 R2 R1 R2 → ∀T. confluent2 … (lfxs R1 T) (lfxs R2 T). #R1 #R2 #HR1 #HR2 #HR12 #T #L0 #L1 * #f1 #Hf1 #HL01 #L2 * #f #Hf #HL02 @@ -160,3 +163,14 @@ elim (lexs_conf … HL01 … HL02) /2 width=3 by ex2_intro/ [ | -HL01 -HL02 ] ] ] qed-. + +theorem lfxs_trans_fsle: ∀R1,R2,R3. + lfxs_fsle_compatible R1 → lfxs_transitive_next R1 R2 R3 → + ∀L1,L,T. L1 ⪤*[R1, T] L → + ∀L2. L ⪤*[R2, T] L2 → L1 ⪤*[R3, T] L2. +#R1 #R2 #R3 #H1R #H2R #L1 #L #T #H +lapply (H1R … H) -H1R #H0 +cases H -H #f1 #Hf1 #HL1 #L2 * #f2 #Hf2 #HL2 +lapply (fsle_inv_frees_eq … H0 … Hf1 … Hf2) -H0 -Hf2 +/4 width=14 by lexs_trans_gen, lexs_fwd_length, sle_lexs_trans, ex2_intro/ +qed-.