X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frelocation%2Fgget_gget.ma;fp=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frelocation%2Fgget_gget.ma;h=6c1001b6af81d070fb76e3ac7c698d382869a596;hb=29973426e0227ee48368d1c24dc0c17bf2baef77;hp=0000000000000000000000000000000000000000;hpb=f95f6cb21b86f3dad114b21f687aa5df36088064;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/gget_gget.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/gget_gget.ma new file mode 100644 index 000000000..6c1001b6a --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/gget_gget.ma @@ -0,0 +1,40 @@ +(**************************************************************************) +(* ___ *) +(* ||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 "basic_2/relocation/gget.ma". + +(* GLOBAL ENVIRONMENT READING ***********************************************) + +(* Main properties **********************************************************) + +theorem gget_mono: ∀G,G1,e. ⇩[e] G ≡ G1 → ∀G2. ⇩[e] G ≡ G2 → G1 = G2. +#G #G1 #e #H elim H -G -G1 +[ #G #He #G2 #H + >(gget_inv_gt … H He) -H -He // +| #G #He #G2 #H + >(gget_inv_eq … H He) -H -He // +| #I #G #G1 #V #He #_ #IHG1 #G2 #H + lapply (gget_inv_lt … H He) -H -He /2 width=1/ +] +qed-. + +lemma gget_dec: ∀G1,G2,e. Decidable (⇩[e] G1 ≡ G2). +#G1 #G2 #e +elim (gget_total e G1) #G #HG1 +elim (genv_eq_dec G G2) #HG2 +[ destruct /2 width=1/ +| @or_intror #HG12 + lapply (gget_mono … HG1 … HG12) -HG1 -HG12 /2 width=1/ +] +qed-.