X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fstatic%2Ffle_fle.ma;h=daa61f75a8120112e3aaed003a9933f3017d867d;hb=42705ef31dd3513a998533e02b5f20fb38dd4fb2;hp=4301456f5871f052dd4409d823b138ea6ab17218;hpb=26d2ecb945a881c61d03f3c259996374209f5d7f;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 4301456f5..daa61f75a 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/fle_fle.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/fle_fle.ma @@ -13,7 +13,6 @@ (**************************************************************************) include "basic_2/syntax/lveq_lveq.ma". -include "basic_2/static/frees_frees.ma". include "basic_2/static/fle_fqup.ma". (* FREE VARIABLES INCLUSION FOR RESTRICTED CLOSURES *************************) @@ -31,13 +30,41 @@ 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-. + (* Main properties **********************************************************) -(* -theorem fle_trans: bi_transitive … fle. -#L1 #L #T1 #T * #f1 #f #HT1 #HT #Hf1 #L2 #T2 * #g #f2 #Hg #HT2 #Hf2 -/5 width=8 by frees_mono, sle_trans, sle_eq_repl_back2, ex3_2_intro/ + +theorem fle_trans_sn: ∀L1,L2,T1,T. ⦃L1, T1⦄ ⊆ ⦃L2, T⦄ → + ∀T2. ⦃L2, T⦄ ⊆ ⦃L2, T2⦄ → ⦃L1, T1⦄ ⊆ ⦃L2, T2⦄. +#L1 #L2 #T1 #T +* #m1 #m0 #g1 #g0 #Hg1 #Hg0 #Hm #Hg +#T2 +* #n0 #n2 #f0 #f2 #Hf0 #Hf2 #Hn #Hf +lapply (frees_mono … Hf0 … Hg0) -Hf0 -Hg0 #Hfg0 +elim (lveq_inj_length … Hn) // -Hn #H1 #H2 destruct +lapply (sle_eq_repl_back1 … Hf … Hfg0) -f0 +/4 width=10 by sle_tls, sle_trans, ex4_4_intro/ qed-. -*) + +theorem fle_trans_dx: ∀L1,T1,T. ⦃L1, T1⦄ ⊆ ⦃L1, T⦄ → + ∀L2,T2. ⦃L1, T⦄ ⊆ ⦃L2, T2⦄ → ⦃L1, T1⦄ ⊆ ⦃L2, T2⦄. +#L1 #T1 #T +* #m1 #m0 #g1 #g0 #Hg1 #Hg0 #Hm #Hg +#L2 #T2 +* #n0 #n2 #f0 #f2 #Hf0 #Hf2 #Hn #Hf +lapply (frees_mono … Hg0 … Hf0) -Hg0 -Hf0 #Hgf0 +elim (lveq_inj_length … Hm) // -Hm #H1 #H2 destruct +lapply (sle_eq_repl_back2 … Hg … Hgf0) -g0 +/4 width=10 by sle_tls, sle_trans, ex4_4_intro/ +qed-. + theorem fle_bind_sn_ge: ∀L1,L2. |L2| ≤ |L1| → ∀V1,T1,T. ⦃L1, V1⦄ ⊆ ⦃L2, T⦄ → ⦃L1.ⓧ, T1⦄ ⊆ ⦃L2, T⦄ → ∀p,I. ⦃L1, ⓑ{p,I}V1.T1⦄ ⊆ ⦃L2, T⦄. @@ -57,21 +84,33 @@ elim (lveq_inj … H1n1 … H1n2) -H1n2 #H1 #H2 destruct elim (sor_isfin_ex f1 f2) /2 width=3 by frees_fwd_isfin/ #f #Hf #_ /4 width=12 by frees_flat, sor_inv_sle, sor_tls, ex4_4_intro/ qed. -(* + +theorem fle_bind_eq: ∀L1,L2. |L1| = |L2| → ∀V1,V2. ⦃L1, V1⦄ ⊆ ⦃L2, V2⦄ → + ∀I2,T1,T2. ⦃L1.ⓧ, T1⦄ ⊆ ⦃L2.ⓑ{I2}V2, T2⦄ → + ∀p,I1. ⦃L1, ⓑ{p,I1}V1.T1⦄ ⊆ ⦃L2, ⓑ{p,I2}V2.T2⦄. +#L1 #L2 #HL #V1 #V2 +* #n1 #m1 #f1 #g1 #Hf1 #Hg1 #H1L #Hfg1 #I2 #T1 #T2 +* #n2 #m2 #f2 #g2 #Hf2 #Hg2 #H2L #Hfg2 #p #I1 +elim (lveq_inj_length … H1L) // #H1 #H2 destruct +elim (lveq_inj_length … H2L) // -HL -H2L #H1 #H2 destruct +elim (sor_isfin_ex f1 (⫱f2)) /3 width=3 by frees_fwd_isfin, isfin_tl/ #f #Hf #_ +elim (sor_isfin_ex g1 (⫱g2)) /3 width=3 by frees_fwd_isfin, isfin_tl/ #g #Hg #_ +/4 width=15 by frees_bind_void, frees_bind, monotonic_sle_sor, sle_tl, ex4_4_intro/ +qed. + theorem fle_bind: ∀L1,L2,V1,V2. ⦃L1, V1⦄ ⊆ ⦃L2, V2⦄ → ∀I1,I2,T1,T2. ⦃L1.ⓑ{I1}V1, T1⦄ ⊆ ⦃L2.ⓑ{I2}V2, T2⦄ → ∀p. ⦃L1, ⓑ{p,I1}V1.T1⦄ ⊆ ⦃L2, ⓑ{p,I2}V2.T2⦄. -#L1 #L2 #V1 #V2 #HV #I1 #I2 #T1 #T2 #HT #p -@fle_bind_sn -[ @fle_bind_dx_sn // -| @fle_bind_dx_dx - - +#L1 #L2 #V1 #V2 +* #n1 #m1 #f1 #g1 #Hf1 #Hg1 #H1L #Hfg1 #I1 #I2 #T1 #T2 +* #n2 #m2 #f2 #g2 #Hf2 #Hg2 #H2L #Hfg2 #p +elim (lveq_inv_pair_pair … H2L) -H2L #H2L #H1 #H2 destruct +elim (lveq_inj … H2L … H1L) -H1L #H1 #H2 destruct elim (sor_isfin_ex f1 (⫱f2)) /3 width=3 by frees_fwd_isfin, isfin_tl/ #f #Hf #_ elim (sor_isfin_ex g1 (⫱g2)) /3 width=3 by frees_fwd_isfin, isfin_tl/ #g #Hg #_ -/4 width=12 by frees_bind, monotonic_sle_sor, sle_tl, ex3_2_intro/ +/4 width=15 by frees_bind, monotonic_sle_sor, sle_tl, ex4_4_intro/ qed. -*) + theorem fle_flat: ∀L1,L2,V1,V2. ⦃L1, V1⦄ ⊆ ⦃L2, V2⦄ → ∀T1,T2. ⦃L1, T1⦄ ⊆ ⦃L2, T2⦄ → ∀I1,I2. ⦃L1, ⓕ{I1}V1.T1⦄ ⊆ ⦃L2, ⓕ{I2}V2.T2⦄.