]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/ground/steps/rtc_isrt.ma
milestone update in basic_2, update in ground and static_2
[helm.git] / matita / matita / contribs / lambdadelta / ground / steps / rtc_isrt.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "ground/notation/relations/isredtype_2.ma".
16 include "ground/steps/rtc.ma".
17
18 (* RT-TRANSITION COUNTER ****************************************************)
19
20 definition isrt: relation2 nat rtc ≝ λts,c.
21                  ∃∃ri,rs. 〈ri,rs,0,ts〉 = c.
22
23 interpretation "test for costrained rt-transition counter (rtc)"
24    'IsRedType ts c = (isrt ts c).
25
26 (* Basic properties *********************************************************)
27
28 lemma isrt_00: 𝐑𝐓❪0,𝟘𝟘❫.
29 /2 width=3 by ex1_2_intro/ qed.
30
31 lemma isrt_10: 𝐑𝐓❪0,𝟙𝟘❫.
32 /2 width=3 by ex1_2_intro/ qed.
33
34 lemma isrt_01: 𝐑𝐓❪1,𝟘𝟙❫.
35 /2 width=3 by ex1_2_intro/ qed.
36
37 lemma isrt_eq_t_trans: ∀n,c1,c2. 𝐑𝐓❪n,c1❫ → rtc_eq_t c1 c2 → 𝐑𝐓❪n,c2❫.
38 #n #c1 #c2 * #ri1 #rs1 #H destruct
39 #H elim (rtc_eq_t_inv_dx … H) -H /2 width=3 by ex1_2_intro/
40 qed-.
41
42 (* Basic inversion properties ***********************************************)
43
44 lemma isrt_inv_00: ∀n. 𝐑𝐓❪n,𝟘𝟘❫ → 0 = n.
45 #n * #ri #rs #H destruct //
46 qed-.
47
48 lemma isrt_inv_10: ∀n. 𝐑𝐓❪n,𝟙𝟘❫ → 0 = n.
49 #n * #ri #rs #H destruct //
50 qed-.
51
52 lemma isrt_inv_01: ∀n. 𝐑𝐓❪n,𝟘𝟙❫ → 1 = n.
53 #n * #ri #rs #H destruct //
54 qed-.
55
56 (* Main inversion properties ************************************************)
57
58 theorem isrt_inj: ∀n1,n2,c. 𝐑𝐓❪n1,c❫ → 𝐑𝐓❪n2,c❫ → n1 = n2.
59 #n1 #n2 #c * #ri1 #rs1 #H1 * #ri2 #rs2 #H2 destruct //
60 qed-.
61
62 theorem isrt_mono: ∀n,c1,c2. 𝐑𝐓❪n,c1❫ → 𝐑𝐓❪n,c2❫ → rtc_eq_t c1 c2.
63 #n #c1 #c2 * #ri1 #rs1 #H1 * #ri2 #rs2 #H2 destruct //
64 qed-.