]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/steps/rtc_plus.ma
bug fix in the context reduction rule for cast (cpm)
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / steps / rtc_plus.ma
index b0b52924071710793e8a33dc291a772e0bf318a8..ad9a04f01171bfb1def1287d4947d0bdc4ca83cb 100644 (file)
@@ -83,13 +83,13 @@ qed-.
 lemma isrt_inv_plus_O_dx: ∀n,c1,c2. 𝐑𝐓⦃n, c1 + c2⦄ → 𝐑𝐓⦃0, c2⦄ → 𝐑𝐓⦃n, c1⦄.
 #n #c1 #c2 #H #H2
 elim (isrt_inv_plus … H) -H #n1 #n2 #Hn1 #Hn2 #H destruct
-lapply (isrt_mono … Hn2 H2) -c2 #H destruct //
+lapply (isrt_inj … Hn2 H2) -c2 #H destruct //
 qed-.
 
 lemma isrt_inv_plus_SO_dx: ∀n,c1,c2. 𝐑𝐓⦃n, c1 + c2⦄ → 𝐑𝐓⦃1, c2⦄ →
                            ∃∃m. 𝐑𝐓⦃m, c1⦄ & n = ⫯m.
 #n #c1 #c2 #H #H2
 elim (isrt_inv_plus … H) -H #n1 #n2 #Hn1 #Hn2 #H destruct
-lapply (isrt_mono … Hn2 H2) -c2 #H destruct
+lapply (isrt_inj … Hn2 H2) -c2 #H destruct
 /2 width=3 by ex2_intro/
 qed-.