]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground/steps/rtc_isrt.ma
update in bin
[helm.git] / matita / matita / contribs / lambdadelta / ground / steps / rtc_isrt.ma
index 17bba149600ee4e0686a9639f347a7bafa202277..66717209694982a291a1b5888c87fccf98667f99 100644 (file)
@@ -20,7 +20,7 @@ include "ground/steps/rtc.ma".
 definition isrt: relation2 nat rtc ≝ λts,c.
                  ∃∃ri,rs. 〈ri,rs,0,ts〉 = c.
 
-interpretation "test for costrained rt-transition counter (rtc)"
+interpretation "test for constrained rt-transition counter (rtc)"
    'IsRedType ts c = (isrt ts c).
 
 (* Basic properties *********************************************************)
@@ -34,9 +34,9 @@ lemma isrt_10: 𝐑𝐓❪0,𝟙𝟘❫.
 lemma isrt_01: 𝐑𝐓❪1,𝟘𝟙❫.
 /2 width=3 by ex1_2_intro/ qed.
 
-lemma isrt_eq_t_trans: ∀n,c1,c2. 𝐑𝐓❪n,c1❫ → eq_t c1 c2 → 𝐑𝐓❪n,c2❫.
+lemma isrt_eq_t_trans: ∀n,c1,c2. 𝐑𝐓❪n,c1❫ → rtc_eq_t c1 c2 → 𝐑𝐓❪n,c2❫.
 #n #c1 #c2 * #ri1 #rs1 #H destruct
-#H elim (eq_t_inv_dx … H) -H /2 width=3 by ex1_2_intro/
+#H elim (rtc_eq_t_inv_dx … H) -H /2 width=3 by ex1_2_intro/
 qed-.
 
 (* Basic inversion properties ***********************************************)
@@ -59,6 +59,6 @@ theorem isrt_inj: ∀n1,n2,c. 𝐑𝐓❪n1,c❫ → 𝐑𝐓❪n2,c❫ → n1 =
 #n1 #n2 #c * #ri1 #rs1 #H1 * #ri2 #rs2 #H2 destruct //
 qed-.
 
-theorem isrt_mono: ∀n,c1,c2. 𝐑𝐓❪n,c1❫ → 𝐑𝐓❪n,c2❫ → eq_t c1 c2.
+theorem isrt_mono: ∀n,c1,c2. 𝐑𝐓❪n,c1❫ → 𝐑𝐓❪n,c2❫ → rtc_eq_t c1 c2.
 #n #c1 #c2 * #ri1 #rs1 #H1 * #ri2 #rs2 #H2 destruct //
 qed-.