X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fstatic%2Flfxs_lfxs.ma;h=cdf65c7c0e71cc2ea04c8b1da0bb9bb2aad5f628;hp=f827a1f20a9f566ab4e2630b0fb5508e11662c30;hb=990f97071a9939d47be16b36f6045d3b23f218e0;hpb=42705ef31dd3513a998533e02b5f20fb38dd4fb2 diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_lfxs.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_lfxs.ma index f827a1f20..cdf65c7c0 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_lfxs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_lfxs.ma @@ -15,8 +15,8 @@ include "basic_2/relocation/lexs_length.ma". include "basic_2/relocation/lexs_lexs.ma". include "basic_2/static/frees_drops.ma". -include "basic_2/static/fle_fle.ma". -include "basic_2/static/lfxs.ma". +include "basic_2/static/fsle_fsle.ma". +include "basic_2/static/lfxs_fsle.ma". (* GENERIC EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ****) @@ -27,13 +27,13 @@ lemma lfxs_inv_frees: ∀R,L1,L2,T. L1 ⪤*[R, T] L2 → #R #L1 #L2 #T * /3 width=6 by frees_mono, lexs_eq_repl_back/ qed-. -lemma frees_lexs_conf: ∀R. lfxs_fle_compatible R → +lemma frees_lexs_conf: ∀R. lfxs_fsle_compatible R → ∀L1,T,f1. L1 ⊢ 𝐅*⦃T⦄ ≡ f1 → ∀L2. L1 ⪤*[cext2 R, cfull, f1] L2 → ∃∃f2. L2 ⊢ 𝐅*⦃T⦄ ≡ f2 & f2 ⊆ f1. #R #HR #L1 #T #f1 #Hf1 #L2 #H1L lapply (HR L1 L2 T ?) /2 width=3 by ex2_intro/ #H2L -@(fle_frees_trans_eq … H2L … Hf1) /3 width=4 by lexs_fwd_length, sym_eq/ +@(fsle_frees_trans_eq … H2L … Hf1) /3 width=4 by lexs_fwd_length, sym_eq/ qed-. (* Properties with free variables inclusion for restricted closures *********) @@ -48,7 +48,7 @@ qed-. (* Advanced properties ******************************************************) -lemma lfxs_sym: ∀R. lfxs_fle_compatible R → +lemma lfxs_sym: ∀R. lfxs_fsle_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 @@ -67,7 +67,7 @@ elim (lexs_dec (cext2 R) cfull … L1 L2 f) qed-. lemma lfxs_pair_sn_split: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) → - lfxs_fle_compatible R1 → + lfxs_fsle_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 @@ -86,7 +86,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_fle_compatible R1 → + lfxs_fsle_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 @@ -102,7 +102,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_fle_compatible R1 → + lfxs_fsle_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 @@ -122,7 +122,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_fle_compatible R1 → + lfxs_fsle_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 @@ -169,8 +169,8 @@ lapply (lexs_fwd_bind … Hf2) -Hf2 #Hf2 elim (sor_isfin_ex f1 (⫱f2)) qed. theorem lfxs_conf: ∀R1,R2. - lfxs_fle_compatible R1 → - lfxs_fle_compatible R2 → + lfxs_fsle_compatible R1 → + lfxs_fsle_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