X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Fsteps%2Frtc_ist.ma;fp=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Fsteps%2Frtc_ist.ma;h=6b9b64681e4372887e234a10d5ac16ba21fae63a;hb=b4f76b0d8fa0e5365fb48e91474febe200b647a7;hp=0000000000000000000000000000000000000000;hpb=67fe9cec87e129a2a41c75d7ed8456a6f3314421;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/ground_2/steps/rtc_ist.ma b/matita/matita/contribs/lambdadelta/ground_2/steps/rtc_ist.ma new file mode 100644 index 000000000..6b9b64681 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/ground_2/steps/rtc_ist.ma @@ -0,0 +1,52 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "ground_2/notation/relations/istype_2.ma". +include "ground_2/steps/rtc.ma". + +(* T-TRANSITION COUNTER *****************************************************) + +definition ist: relation2 nat rtc ≝ + λts,c. 〈0,0,0,ts〉 = c. + +interpretation "test for t-transition counter (rtc)" + 'IsType ts c = (ist ts c). + +(* Basic properties *********************************************************) + +lemma isr_00: 𝐓⦃0,𝟘𝟘⦄. +// qed. + +lemma ist_01: 𝐓⦃1,𝟘𝟙⦄. +// qed. + +(* Basic inversion properties ***********************************************) + +lemma ist_inv_00: ∀n. 𝐓⦃n,𝟘𝟘⦄ → 0 = n. +#n #H destruct // +qed-. + +lemma ist_inv_01: ∀n. 𝐓⦃n,𝟘𝟙⦄ → 1 = n. +#n #H destruct // +qed-. + +(* Main inversion properties ************************************************) + +theorem ist_inj: ∀n1,n2,c. 𝐓⦃n1,c⦄ → 𝐓⦃n2,c⦄ → n1 = n2. +#n1 #n2 #c #H1 #H2 destruct // +qed-. + +theorem ist_mono: ∀n,c1,c2. 𝐓⦃n,c1⦄ → 𝐓⦃n,c2⦄ → c1 = c2. +#n #c1 #c2 #H1 #H2 destruct // +qed-.