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/tuple_4.ma".
16 include "ground_2/notation/functions/zerozero_0.ma".
17 include "ground_2/notation/functions/zeroone_0.ma".
18 include "ground_2/notation/functions/onezero_0.ma".
19 include "ground_2/lib/arith.ma".
21 (* RT-TRANSITION COUNTER ****************************************************)
23 record rtc: Type[0] ≝ {
24 ri: nat; (* Note: inner r-steps *)
25 rs: nat; (* Note: spine r-steps *)
26 ti: nat; (* Note: inner t-steps *)
27 ts: nat (* Note: spine t-steps *)
30 interpretation "constructor (rtc)"
31 'Tuple ri rs ti ts = (mk_rtc ri rs ti ts).
33 interpretation "one structural step (rtc)"
34 'ZeroZero = (mk_rtc O O O O).
36 interpretation "one r-step (rtc)"
37 'OneZero = (mk_rtc O (S O) O O).
39 interpretation "one t-step (rtc)"
40 'ZeroOne = (mk_rtc O O O (S O)).
42 definition eq_f: relation rtc ≝ λc1,c2. ⊤.
44 inductive eq_t: relation rtc ≝
45 | eq_t_intro: ∀ri1,ri2,rs1,rs2,ti,ts.
46 eq_t (〈ri1,rs1,ti,ts〉) (〈ri2,rs2,ti,ts〉)
49 (* Basic properties *********************************************************)
51 lemma eq_t_refl: reflexive … eq_t.
54 (* Basic inversion lemmas ***************************************************)
56 fact eq_t_inv_dx_aux: ∀x,y. eq_t x y →
57 ∀ri1,rs1,ti,ts. x = 〈ri1,rs1,ti,ts〉 →
58 ∃∃ri2,rs2. y = 〈ri2,rs2,ti,ts〉.
60 #ri1 #ri #rs1 #rs #ti1 #ts1 #ri2 #rs2 #ti2 #ts2 #H destruct -ri2 -rs2
61 /2 width=3 by ex1_2_intro/
64 lemma eq_t_inv_dx: ∀ri1,rs1,ti,ts,y. eq_t (〈ri1,rs1,ti,ts〉) y →
65 ∃∃ri2,rs2. y = 〈ri2,rs2,ti,ts〉.
66 /2 width=5 by eq_t_inv_dx_aux/ qed-.