X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fstatic%2Ffle_fle.ma;h=18574be3dfe468ed418def66ee6c169a8255b4f0;hb=9323611e3819c1382b872a7ada00264991f36217;hp=e756b43f773e1d873c3f67fc8a583d7f481abe65;hpb=aff2f21cec89cf343297ae6bd720fc6573a3aba1;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/fle_fle.ma b/matita/matita/contribs/lambdadelta/basic_2/static/fle_fle.ma index e756b43f7..18574be3d 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/fle_fle.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/fle_fle.ma @@ -12,20 +12,30 @@ (* *) (**************************************************************************) -include "basic_2/static/frees_frees.ma". +include "basic_2/syntax/lveq_lveq.ma". include "basic_2/static/fle_fqup.ma". (* FREE VARIABLES INCLUSION FOR RESTRICTED CLOSURES *************************) (* Advanced inversion lemmas ************************************************) -lemma fle_inv_voids_sn_frees_dx: ∀L1,L2,T1,T2,n. ⦃ⓧ*[n]L1, T1⦄ ⊆ ⦃L2, T2⦄ → - |L1| = |L2| → ∀f2. L2 ⊢ 𝐅*⦃T2⦄ ≡ f2 → - ∃∃f1. ⓧ*[n]L1 ⊢ 𝐅*⦃T1⦄ ≡ f1 & ⫱*[n]f1 ⊆ f2. -#L1 #L2 #T1 #T2 #n #H #HL12 #f2 #Hf2 -elim (fle_inv_voids_sn … H HL12) -H -HL12 #f1 #g2 #Hf1 #Hg2 #Hfg -lapply (frees_mono … Hg2 … Hf2) -Hg2 -Hf2 #Hfg2 -lapply (sle_eq_repl_back2 … Hfg … Hfg2) -g2 +lemma fle_frees_trans: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊆ ⦃L2, T2⦄ → + ∀f2. L2 ⊢ 𝐅*⦃T2⦄ ≡ f2 → + ∃∃n1,n2,f1. L1 ⊢ 𝐅*⦃T1⦄ ≡ f1 & + L1 ≋ⓧ*[n1, n2] L2 & ⫱*[n1]f1 ⊆ ⫱*[n2]f2. +#L1 #L2 #T1 #T2 * #n1 #n2 #f1 #g2 #Hf1 #Hg2 #HL #Hn #f2 #Hf2 +lapply (frees_mono … Hg2 … Hf2) -Hg2 -Hf2 #Hgf2 +lapply (tls_eq_repl n2 … Hgf2) -Hgf2 #Hgf2 +lapply (sle_eq_repl_back2 … Hn … Hgf2) -g2 +/2 width=6 by ex3_3_intro/ +qed-. + +lemma fle_frees_trans_eq: ∀L1,L2. |L1| = |L2| → + ∀T1,T2. ⦃L1, T1⦄ ⊆ ⦃L2, T2⦄ → ∀f2. L2 ⊢ 𝐅*⦃T2⦄ ≡ f2 → + ∃∃f1. L1 ⊢ 𝐅*⦃T1⦄ ≡ f1 & f1 ⊆ f2. +#L1 #L2 #H1L #T1 #T2 #H2L #f2 #Hf2 +elim (fle_frees_trans … H2L … Hf2) -T2 #n1 #n2 #f1 #Hf1 #H2L #Hf12 +elim (lveq_inj_length … H2L) // -L2 #H1 #H2 destruct /2 width=3 by ex2_intro/ qed-. @@ -36,39 +46,53 @@ theorem fle_trans: bi_transitive … fle. /5 width=8 by frees_mono, sle_trans, sle_eq_repl_back2, ex3_2_intro/ qed-. *) -theorem fle_bind_sn: ∀L1,L2,V1,T1,T. ⦃L1, V1⦄ ⊆ ⦃L2, T⦄ → ⦃L1.ⓧ, T1⦄ ⊆ ⦃L2, T⦄ → - ∀p,I. ⦃L1, ⓑ{p,I}V1.T1⦄ ⊆ ⦃L2, T⦄. -#L1 #L2 #V1 #T1 #T * -L1 #f1 #x #L1 #n1 #Hf1 #Hx #HL12 #Hf1x ->voids_succ #H #p #I -elim (fle_inv_voids_sn_frees_dx … H … Hx) -H // #f2 #Hf2 -