]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/ground_2/steps/rtc_isrt.ma
first definition of cpm:
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / 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_2/notation/relations/isredtype_2.ma".
16 include "ground_2/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 isr_00: 𝐑𝐓⦃0, 𝟘𝟘⦄.
29 /2 width=3 by ex1_2_intro/ qed.
30
31 lemma isr_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 (* Basic inversion properties ***********************************************)
38
39 lemma isrt_inv_00: ∀n. 𝐑𝐓⦃n, 𝟘𝟘⦄ → 0 = n.
40 #n * #ri #rs #H destruct //
41 qed-.
42
43 lemma isrt_inv_10: ∀n. 𝐑𝐓⦃n, 𝟙𝟘⦄ → 0 = n.
44 #n * #ri #rs #H destruct //
45 qed-.
46
47 lemma isrt_inv_01: ∀n. 𝐑𝐓⦃n, 𝟘𝟙⦄ → 1 = n.
48 #n * #ri #rs #H destruct //
49 qed-.
50
51 (* Main inversion properties ************************************************)
52
53 theorem isrt_mono: ∀n1,n2,c. 𝐑𝐓⦃n1, c⦄ → 𝐑𝐓⦃n2, c⦄ → n1 = n2.
54 #n1 #n2 #c * #ri1 #rs1 #H1 * #ri2 #rs2 #H2 destruct //
55 qed-.