X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Fsteps%2Frtc.ma;h=1e08d2474be31943ed50dcfd6fd1e55c50051768;hp=62228d229e21af156c36c43980e8956f51a4fd48;hb=ff612dc35167ec0c145864c9aa8ae5e1ebe20a48;hpb=f6b452b9c9be141740d4058dfbcf81a4b75fd00b diff --git a/matita/matita/contribs/lambdadelta/ground_2/steps/rtc.ma b/matita/matita/contribs/lambdadelta/ground_2/steps/rtc.ma index 62228d229..1e08d2474 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/steps/rtc.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/steps/rtc.ma @@ -12,10 +12,10 @@ (* *) (**************************************************************************) -include "ground_2/notation/constructors/tuple_4.ma". -include "ground_2/notation/constructors/zerozero_0.ma". -include "ground_2/notation/constructors/zeroone_0.ma". -include "ground_2/notation/constructors/onezero_0.ma". +include "ground_2/notation/functions/tuple_4.ma". +include "ground_2/notation/functions/zerozero_0.ma". +include "ground_2/notation/functions/zeroone_0.ma". +include "ground_2/notation/functions/onezero_0.ma". include "ground_2/lib/arith.ma". (* RT-TRANSITION COUNTER ****************************************************) @@ -38,3 +38,29 @@ interpretation "one r-step (rtc)" interpretation "one t-step (rtc)" 'ZeroOne = (mk_rtc O O O (S O)). + +definition eq_f: relation rtc ≝ λc1,c2. ⊤. + +inductive eq_t: relation rtc ≝ +| eq_t_intro: ∀ri1,ri2,rs1,rs2,ti,ts. + eq_t (〈ri1, rs1, ti, ts〉) (〈ri2, rs2, ti, ts〉) +. + +(* Basic properties *********************************************************) + +lemma eq_t_refl: reflexive … eq_t. +* // qed. + +(* Basic inversion lemmas ***************************************************) + +fact eq_t_inv_dx_aux: ∀x,y. eq_t x y → + ∀ri1,rs1,ti,ts. x = 〈ri1,rs1,ti,ts〉 → + ∃∃ri2,rs2. y = 〈ri2,rs2,ti,ts〉. +#x #y * -x -y +#ri1 #ri #rs1 #rs #ti1 #ts1 #ri2 #rs2 #ti2 #ts2 #H destruct -ri2 -rs2 +/2 width=3 by ex1_2_intro/ +qed-. + +lemma eq_t_inv_dx: ∀ri1,rs1,ti,ts,y. eq_t (〈ri1,rs1,ti,ts〉) y → + ∃∃ri2,rs2. y = 〈ri2,rs2,ti,ts〉. +/2 width=5 by eq_t_inv_dx_aux/ qed-.