]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/csx_alt.ma
- bug fix in the induction for the closure property
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / csx_alt.ma
index 7c95fed6ed5cbf535c7ee4e551289536367163c7..ccd336acbcf7d58e248d617a72b251be143fa3f2 100644 (file)
@@ -55,7 +55,7 @@ lemma csxa_cpxs_trans: ∀h,g,G,L,T1. ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T1 →
                        ∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T2.
 #h #g #G #L #T1 #H elim H -T1 #T1 #HT1 #IHT1 #T2 #HLT12
 @csxa_intro #T #HLT2 #HT2
-elim (term_eq_dec T1 T2) #HT12
+elim (eq_term_dec T1 T2) #HT12
 [ -IHT1 -HLT12 destruct /3 width=1/
 | -HT1 -HT2 /3 width=4/
 qed.
@@ -69,7 +69,7 @@ lemma csxa_intro_cpx: ∀h,g,G,L,T1. (
 [ -H #H destruct #H
   elim H //
 | #T0 #T #HLT1 #HLT2 #IHT #HT10 #HT12 destruct
-  elim (term_eq_dec T0 T) #HT0
+  elim (eq_term_dec T0 T) #HT0
   [ -HLT1 -HLT2 -H /3 width=1/
   | -IHT -HT12 /4 width=3/
   ]