X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2A%2Fstatic%2Flsubr_lsubr.ma;fp=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2A%2Fstatic%2Flsubr_lsubr.ma;h=fbc688aa038f804ed3ac2eb545f2eb1af6312046;hb=d2545ffd201b1aa49887313791386add78fa8603;hp=0000000000000000000000000000000000000000;hpb=57ae1762497a5f3ea75740e2908e04adb8642cc2;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2A/static/lsubr_lsubr.ma b/matita/matita/contribs/lambdadelta/basic_2A/static/lsubr_lsubr.ma new file mode 100644 index 000000000..fbc688aa0 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2A/static/lsubr_lsubr.ma @@ -0,0 +1,53 @@ +(**************************************************************************) +(* ___ *) +(* ||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_2A/static/lsubr.ma". + +(* RESTRICTED LOCAL ENVIRONMENT REFINEMENT **********************************) + +(* Auxiliary inversion lemmas ***********************************************) + +fact lsubr_inv_pair1_aux: ∀L1,L2. L1 ⫃ L2 → ∀I,K1,X. L1 = K1.ⓑ{I}X → + ∨∨ L2 = ⋆ + | ∃∃K2. K1 ⫃ K2 & L2 = K2.ⓑ{I}X + | ∃∃K2,V,W. K1 ⫃ K2 & L2 = K2.ⓛW & + I = Abbr & X = ⓝW.V. +#L1 #L2 * -L1 -L2 +[ #L #J #K1 #X #H destruct /2 width=1 by or3_intro0/ +| #I #L1 #L2 #V #HL12 #J #K1 #X #H destruct /3 width=3 by or3_intro1, ex2_intro/ +| #L1 #L2 #V #W #HL12 #J #K1 #X #H destruct /3 width=6 by or3_intro2, ex4_3_intro/ +] +qed-. + +lemma lsubr_inv_pair1: ∀I,K1,L2,X. K1.ⓑ{I}X ⫃ L2 → + ∨∨ L2 = ⋆ + | ∃∃K2. K1 ⫃ K2 & L2 = K2.ⓑ{I}X + | ∃∃K2,V,W. K1 ⫃ K2 & L2 = K2.ⓛW & + I = Abbr & X = ⓝW.V. +/2 width=3 by lsubr_inv_pair1_aux/ qed-. + +(* Main properties **********************************************************) + +theorem lsubr_trans: Transitive … lsubr. +#L1 #L #H elim H -L1 -L +[ #L1 #X #H + lapply (lsubr_inv_atom1 … H) -H // +| #I #L1 #L #V #_ #IHL1 #X #H + elim (lsubr_inv_pair1 … H) -H // * + #L2 [2: #V2 #W2 ] #HL2 #H1 [ #H2 #H3 ] destruct /3 width=1 by lsubr_pair, lsubr_beta/ +| #L1 #L #V1 #W #_ #IHL1 #X #H + elim (lsubr_inv_abst1 … H) -H // * + #L2 #HL2 #H destruct /3 width=1 by lsubr_beta/ +] +qed-.