1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "ground_2/notation/functions/drop_1.ma".
16 include "ground_2/steps/rtc.ma".
18 (* RT-TRANSITION COUNTER ****************************************************)
20 definition shift (r:rtc): rtc ≝ match r with
21 [ mk_rtc ri rh ti th ⇒ 〈ri+rh, 0, ti+th, 0〉 ].
23 interpretation "shift (rtc)"
26 (* Basic properties *********************************************************)
28 lemma shift_refl: ∀ri,ti. 〈ri, 0, ti, 0〉 = ↓〈ri, 0, ti, 0〉.