]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/syntax/ext2_tc.ma
renaming in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / syntax / ext2_tc.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/lib/star.ma".
16 include "basic_2/syntax/ext2.ma".
17
18 (* EXTENSION TO BINDERS OF A RELATION FOR TERMS *****************************)
19
20 (* Properties with transitive closure ***************************************)
21
22 lemma ext2_tc_pair: ∀R,I,V1,V2. TC … R V1 V2 →
23                     TC … (ext2 R) (BPair I V1) (BPair I V2).
24 #R #I #V1 #V2 #H elim H -H -V2
25 /3 width=3 by ext2_pair, step, inj/
26 qed.
27
28 lemma ext2_tc_inj: ∀R,I1,I2. ext2 R I1 I2 → ext2 (TC … R) I1 I2.
29 #R #I1 #I2 * -I1 -I2
30 /3 width=1 by ext2_unit, ext2_pair, inj/
31 qed.
32
33 (* Main properties with transitive closure **********************************)
34
35 theorem ext2_tc_step: ∀R,I1,I. ext2 (TC … R) I1 I →
36                       ∀I2. ext2 R I I2 → ext2 (TC … R) I1 I2.
37 #R #I1 #I * -I1 -I
38 [ #I #Z #H >(ext2_inv_unit_sn … H) -Z /2 width=1 by ext2_unit/
39 | #I #V1 #V #HV1 #Z #H
40   elim (ext2_inv_pair_sn … H) -H #V2 #HV2 #H destruct
41   /3 width=3 by ext2_pair, step/
42 ]
43 qed-.
44
45 (* Advanced properties with transitive closure ******************************)
46
47 lemma ext2_tc: ∀R,I1,I2. TC … (ext2 R) I1 I2 → ext2 (TC … R) I1 I2.
48 #R #I1 #I2 #H elim H -I2
49 /2 width=3 by ext2_tc_step, ext2_tc_inj/
50 qed.
51
52 (* Advanced inversion lemmas with transitive closure ************************)
53
54 lemma ext2_inv_tc: ∀R,I1,I2. ext2 (TC … R) I1 I2 → TC … (ext2 R) I1 I2.
55 #R #I1 #I2 * -I1 -I2
56 /3 width=1 by ext2_tc_pair, ext2_unit, inj/
57 qed-.