X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground%2Fsteps%2Frtc_isrt.ma;h=7f9f2f74ad1d7e891662dcf7d0dd2abfeb9cacd1;hb=3c7b4071a9ac096b02334c1d47468776b948e2de;hp=17bba149600ee4e0686a9639f347a7bafa202277;hpb=68b4f2490c12139c03760b39895619e63b0f38c9;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/ground/steps/rtc_isrt.ma b/matita/matita/contribs/lambdadelta/ground/steps/rtc_isrt.ma index 17bba1496..7f9f2f74a 100644 --- a/matita/matita/contribs/lambdadelta/ground/steps/rtc_isrt.ma +++ b/matita/matita/contribs/lambdadelta/ground/steps/rtc_isrt.ma @@ -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-.