X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fstatic%2Flsubr_lsubr.ma;h=f0e171e0b2198c2564fbcfe41061d6fe2499a12e;hb=ae22633ad934bba5fb32143f1726cdfbd255b899;hp=9c3fec807ce2ac7b9364473a4552746dff90faf6;hpb=944b1f7b762774a6f8d99a2c2846f865b6788712;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lsubr_lsubr.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lsubr_lsubr.ma index 9c3fec807..f0e171e0b 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lsubr_lsubr.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/lsubr_lsubr.ma @@ -18,10 +18,10 @@ include "basic_2/static/lsubr.ma". (* Auxiliary inversion lemmas ***********************************************) -fact lsubr_inv_bind1_aux: ∀L1,L2. L1 ⊑ L2 → ∀I,K1,X. L1 = K1.ⓑ{I}X → +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 & + | ∃∃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/ @@ -30,12 +30,12 @@ fact lsubr_inv_bind1_aux: ∀L1,L2. L1 ⊑ L2 → ∀I,K1,X. L1 = K1.ⓑ{I}X → ] qed-. -lemma lsubr_inv_bind1: ∀I,K1,L2,X. K1.ⓑ{I}X ⊑ L2 → +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 & + | ∃∃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_bind1_aux/ qed-. +/2 width=3 by lsubr_inv_pair1_aux/ qed-. (* Main properties **********************************************************) @@ -44,10 +44,10 @@ theorem lsubr_trans: Transitive … lsubr. [ #L1 #X #H lapply (lsubr_inv_atom1 … H) -H // | #I #L1 #L #V #_ #IHL1 #X #H - elim (lsubr_inv_bind1 … H) -H // * - #L2 [2: #V2 #W2 ] #HL2 #H1 [ #H2 #H3 ] destruct /3 width=1 by lsubr_bind, lsubr_abst/ + 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_abst/ + #L2 #HL2 #H destruct /3 width=1 by lsubr_beta/ ] qed-.