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/xoa/ex_1_2.ma".
16 include "ground/notation/functions/tuple_4.ma".
17 include "ground/notation/functions/zerozero_0.ma".
18 include "ground/notation/functions/zeroone_0.ma".
19 include "ground/notation/functions/onezero_0.ma".
20 include "ground/generated/insert_eq_1.ma".
21 include "ground/arith/nat.ma".
23 (* RT-TRANSITION COUNTERS ***************************************************)
25 record rtc: Type[0] ≝ {
26 (* Note: inner r-steps *)
28 (* Note: spine r-steps *)
30 (* Note: inner t-steps *)
32 (* Note: spine t-steps *)
37 "construction (rt-transition counters)"
38 'Tuple ri rs ti ts = (mk_rtc ri rs ti ts).
41 "one structural step (rt-transition counters)"
42 'ZeroZero = (mk_rtc nzero nzero nzero nzero).
45 "one r-step (rt-transition counters)"
46 'OneZero = (mk_rtc nzero (ninj punit) nzero nzero).
49 "one t-step (rt-transition counters)"
50 'ZeroOne = (mk_rtc nzero nzero nzero (ninj punit)).
52 definition rtc_eq_f: relation rtc ≝ λc1,c2. ⊤.
54 inductive rtc_eq_t: relation rtc ≝
55 | eq_t_intro: ∀ri1,ri2,rs1,rs2,ti,ts.
56 rtc_eq_t (〈ri1,rs1,ti,ts〉) (〈ri2,rs2,ti,ts〉)
59 (* Basic constructions ******************************************************)
65 (* Basic inversions *********************************************************)
67 lemma rtc_eq_t_inv_dx:
68 ∀ri1,rs1,ti,ts,y. rtc_eq_t (〈ri1,rs1,ti,ts〉) y →
69 ∃∃ri2,rs2. y = 〈ri2,rs2,ti,ts〉.
70 #ri0 #rs0 #ti0 #ts0 #y @(insert_eq_1 … (〈ri0,rs0,ti0,ts0〉))
72 #ri1 #ri2 #rs1 #rs2 #ti1 #ts1 #H destruct -ri1 -rs1
73 /2 width=3 by ex1_2_intro/