X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fstatic_2%2Fstatic%2Ffsle.ma;fp=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fstatic_2%2Fstatic%2Ffsle.ma;h=bfafabbdbb0edfcefae8808a3b750f3ceb567374;hb=ff612dc35167ec0c145864c9aa8ae5e1ebe20a48;hp=0000000000000000000000000000000000000000;hpb=222044da28742b24584549ba86b1805a87def070;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/static_2/static/fsle.ma b/matita/matita/contribs/lambdadelta/static_2/static/fsle.ma new file mode 100644 index 000000000..bfafabbdb --- /dev/null +++ b/matita/matita/contribs/lambdadelta/static_2/static/fsle.ma @@ -0,0 +1,38 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "ground_2/relocation/rtmap_id.ma". +include "static_2/notation/relations/subseteq_4.ma". +include "static_2/syntax/lveq.ma". +include "static_2/static/frees.ma". + +(* FREE VARIABLES INCLUSION FOR RESTRICTED CLOSURES *************************) + +definition fsle: bi_relation lenv term ≝ λL1,T1,L2,T2. + ∃∃n1,n2,f1,f2. L1 ⊢ 𝐅*⦃T1⦄ ≘ f1 & L2 ⊢ 𝐅*⦃T2⦄ ≘ f2 & + L1 ≋ⓧ*[n1, n2] L2 & ⫱*[n1]f1 ⊆ ⫱*[n2]f2. + +interpretation "free variables inclusion (restricted closure)" + 'SubSetEq L1 T1 L2 T2 = (fsle L1 T1 L2 T2). + +interpretation "free variables inclusion (term)" + 'subseteq T1 T2 = (fsle LAtom T1 LAtom T2). + +(* Basic properties *********************************************************) + +lemma fsle_sort: ∀L,s1,s2. ⦃L, ⋆s1⦄ ⊆ ⦃L, ⋆s2⦄. +/3 width=8 by frees_sort, sle_refl, ex4_4_intro/ qed. + +lemma fsle_gref: ∀L,l1,l2. ⦃L, §l1⦄ ⊆ ⦃L, §l2⦄. +/3 width=8 by frees_gref, sle_refl, ex4_4_intro/ qed.