X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fstatic_2%2Fsyntax%2Flveq_lveq.ma;h=83b9bc787722484386c60b71fe8d87fbe9e9cfe4;hp=6966715c4202ad3fb716de9cd84218785b8aecc9;hb=156d974ad89aa04a086fdf9d332c8b04adf279fd;hpb=8fe4dc148d50a0352313633bea61441bc817afbf diff --git a/matita/matita/contribs/lambdadelta/static_2/syntax/lveq_lveq.ma b/matita/matita/contribs/lambdadelta/static_2/syntax/lveq_lveq.ma index 6966715c4..83b9bc787 100644 --- a/matita/matita/contribs/lambdadelta/static_2/syntax/lveq_lveq.ma +++ b/matita/matita/contribs/lambdadelta/static_2/syntax/lveq_lveq.ma @@ -18,12 +18,12 @@ include "static_2/syntax/lveq_length.ma". (* Main inversion lemmas ****************************************************) -theorem lveq_inv_bind: ∀K1,K2. K1 ≋ⓧ*[0,0] K2 → +theorem lveq_inv_bind: ∀K1,K2. K1 ≋ⓧ*[𝟎,𝟎] K2 → ∀I1,I2,m1,m2. K1.ⓘ[I1] ≋ⓧ*[m1,m2] K2.ⓘ[I2] → - ∧∧ 0 = m1 & 0 = m2. + ∧∧ 𝟎 = m1 & 𝟎 = m2. #K1 #K2 #HK #I1 #I2 #m1 #m2 #H lapply (lveq_fwd_length_eq … HK) -HK #HK -elim (lveq_inj_length … H) -H normalize /3 width=1 by conj, eq_f/ +elim (lveq_inj_length … H) -H /3 width=1 by conj/ qed-. theorem lveq_inj: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1,n2] L2 → @@ -38,16 +38,17 @@ qed-. theorem lveq_inj_void_sn_ge: ∀K1,K2. |K2| ≤ |K1| → ∀n1,n2. K1 ≋ⓧ*[n1,n2] K2 → ∀m1,m2. K1.ⓧ ≋ⓧ*[m1,m2] K2 → - ∧∧ ↑n1 = m1 & 0 = m2 & 0 = n2. + ∧∧ ↑n1 = m1 & 𝟎 = m2 & 𝟎 = n2. #L1 #L2 #HL #n1 #n2 #Hn #m1 #m2 #Hm elim (lveq_fwd_length … Hn) -Hn #H1 #H2 destruct elim (lveq_fwd_length … Hm) -Hm #H1 #H2 destruct ->length_bind >nminus_succ_dx >(nle_inv_eq_zero_minus … HL) -/3 width=4 by nminus_plus_comm_23, and3_intro/ +>length_bind