X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_transition%2Fcpm_drops.ma;h=fbcd7958d4776e6ec180603c5e5601a0981d7e9b;hb=09b4420070d6a71990e16211e499b51dbb0742cb;hp=786d2f9ce164817a442252ff32db91230d63e882;hpb=bba53a83579540bc3925d47d679e2aad22e85755;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpm_drops.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpm_drops.ma index 786d2f9ce..fbcd7958d 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpm_drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpm_drops.ma @@ -42,8 +42,8 @@ lemma cpm_inv_atom1_drops: ∀n,h,I,G,L,T2. ⦃G, L⦄ ⊢ ⓪{I} ➡[n, h] T2 | ∃∃s. T2 = ⋆(next h s) & I = Sort s & n = 1 | ∃∃K,V,V2,i. ⬇*[i] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V ➡[n, h] V2 & ⬆*[⫯i] V2 ≡ T2 & I = LRef i - | ∃∃m,K,V,V2,i. ⬇*[i] L ≡ K.ⓛV & ⦃G, K⦄ ⊢ V ➡[m, h] V2 & - ⬆*[⫯i] V2 ≡ T2 & I = LRef i & n = ⫯m. + | ∃∃k,K,V,V2,i. ⬇*[i] L ≡ K.ⓛV & ⦃G, K⦄ ⊢ V ➡[k, h] V2 & + ⬆*[⫯i] V2 ≡ T2 & I = LRef i & n = ⫯k. #n #h #I #G #L #T2 * #c #Hc #H elim (cpg_inv_atom1_drops … H) -H * [ #H1 #H2 destruct lapply (isrt_inv_00 … Hc) -Hc /3 width=1 by or4_intro0, conj/ @@ -61,8 +61,8 @@ lemma cpm_inv_lref1_drops: ∀n,h,G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡[n, h] T2 → ∨∨ T2 = #i ∧ n = 0 | ∃∃K,V,V2. ⬇*[i] L ≡ K. ⓓV & ⦃G, K⦄ ⊢ V ➡[n, h] V2 & ⬆*[⫯i] V2 ≡ T2 - | ∃∃m,K,V,V2. ⬇*[i] L ≡ K. ⓛV & ⦃G, K⦄ ⊢ V ➡[m, h] V2 & - ⬆*[⫯i] V2 ≡ T2 & n = ⫯m. + | ∃∃k,K,V,V2. ⬇*[i] L ≡ K. ⓛV & ⦃G, K⦄ ⊢ V ➡[k, h] V2 & + ⬆*[⫯i] V2 ≡ T2 & n = ⫯k. #n #h #G #L #T2 #i * #c #Hc #H elim (cpg_inv_lref1_drops … H) -H * [ #H1 #H2 destruct lapply (isrt_inv_00 … Hc) -Hc /3 width=1 by or3_intro0, conj/