X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Fsteps%2Frtc_isrt.ma;h=45be4c4adb03cd6593477afe3ebeca5cebf271d9;hp=9c988abea63f85bd33a65e7d05b191bfe96b1244;hb=f308429a0fde273605a2330efc63268b4ac36c99;hpb=87f57ddc367303c33e19c83cd8989cd561f3185b diff --git a/matita/matita/contribs/lambdadelta/ground_2/steps/rtc_isrt.ma b/matita/matita/contribs/lambdadelta/ground_2/steps/rtc_isrt.ma index 9c988abea..45be4c4ad 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/steps/rtc_isrt.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/steps/rtc_isrt.ma @@ -18,47 +18,47 @@ include "ground_2/steps/rtc.ma". (* RT-TRANSITION COUNTER ****************************************************) definition isrt: relation2 nat rtc ≝ λts,c. - ∃∃ri,rs. 〈ri, rs, 0, ts〉 = c. + ∃∃ri,rs. 〈ri,rs,0,ts〉 = c. interpretation "test for costrained rt-transition counter (rtc)" 'IsRedType ts c = (isrt ts c). (* Basic properties *********************************************************) -lemma isr_00: 𝐑𝐓⦃0, 𝟘𝟘⦄. +lemma isr_00: 𝐑𝐓⦃0,𝟘𝟘⦄. /2 width=3 by ex1_2_intro/ qed. -lemma isr_10: 𝐑𝐓⦃0, 𝟙𝟘⦄. +lemma isr_10: 𝐑𝐓⦃0,𝟙𝟘⦄. /2 width=3 by ex1_2_intro/ qed. -lemma isrt_01: 𝐑𝐓⦃1, 𝟘𝟙⦄. +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⦄ → 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/ qed-. (* Basic inversion properties ***********************************************) -lemma isrt_inv_00: ∀n. 𝐑𝐓⦃n, 𝟘𝟘⦄ → 0 = n. +lemma isrt_inv_00: ∀n. 𝐑𝐓⦃n,𝟘𝟘⦄ → 0 = n. #n * #ri #rs #H destruct // qed-. -lemma isrt_inv_10: ∀n. 𝐑𝐓⦃n, 𝟙𝟘⦄ → 0 = n. +lemma isrt_inv_10: ∀n. 𝐑𝐓⦃n,𝟙𝟘⦄ → 0 = n. #n * #ri #rs #H destruct // qed-. -lemma isrt_inv_01: ∀n. 𝐑𝐓⦃n, 𝟘𝟙⦄ → 1 = n. +lemma isrt_inv_01: ∀n. 𝐑𝐓⦃n,𝟘𝟙⦄ → 1 = n. #n * #ri #rs #H destruct // qed-. (* Main inversion properties ************************************************) -theorem isrt_inj: ∀n1,n2,c. 𝐑𝐓⦃n1, c⦄ → 𝐑𝐓⦃n2, c⦄ → n1 = n2. +theorem isrt_inj: ∀n1,n2,c. 𝐑𝐓⦃n1,c⦄ → 𝐑𝐓⦃n2,c⦄ → n1 = n2. #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⦄ → eq_t c1 c2. #n #c1 #c2 * #ri1 #rs1 #H1 * #ri2 #rs2 #H2 destruct // qed-.