]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2A/substitution/lpx_sn_drop.ma
update in lambdadelta
[helm.git] / matita / matita / contribs / lambdadelta / basic_2A / substitution / lpx_sn_drop.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 "basic_2A/substitution/drop_lreq.ma".
16 include "basic_2A/substitution/lpx_sn.ma".
17
18 (* SN POINTWISE EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS *********)
19
20 (* Properties on dropping ****************************************************)
21
22 lemma lpx_sn_drop_conf: ∀R,L1,L2. lpx_sn R L1 L2 →
23                         ∀I,K1,V1,i. ⬇[i] L1 ≡ K1.ⓑ{I}V1 →
24                         ∃∃K2,V2. ⬇[i] L2 ≡ K2.ⓑ{I}V2 & lpx_sn R K1 K2 & R K1 V1 V2.
25 #R #L1 #L2 #H elim H -L1 -L2
26 [ #I0 #K0 #V0 #i #H elim (drop_inv_atom1 … H) -H #H destruct
27 | #I #K1 #K2 #V1 #V2 #HK12 #HV12 #IHK12 #I0 #K0 #V0 #i #H elim (drop_inv_O1_pair1 … H) * -H
28   [ -IHK12 #H1 #H2 destruct /3 width=5 by drop_pair, ex3_2_intro/
29   | -HK12 -HV12 #Hi #HK10 elim (IHK12 … HK10) -IHK12 -HK10
30     /3 width=5 by drop_drop_lt, ex3_2_intro/
31   ]
32 ]
33 qed-.
34
35 lemma lpx_sn_drop_trans: ∀R,L1,L2. lpx_sn R L1 L2 →
36                          ∀I,K2,V2,i. ⬇[i] L2 ≡ K2.ⓑ{I}V2 →
37                          ∃∃K1,V1. ⬇[i] L1 ≡ K1.ⓑ{I}V1 & lpx_sn R K1 K2 & R K1 V1 V2.
38 #R #L1 #L2 #H elim H -L1 -L2
39 [ #I0 #K0 #V0 #i #H elim (drop_inv_atom1 … H) -H #H destruct
40 | #I #K1 #K2 #V1 #V2 #HK12 #HV12 #IHK12 #I0 #K0 #V0 #i #H elim (drop_inv_O1_pair1 … H) * -H
41   [ -IHK12 #H1 #H2 destruct /3 width=5 by drop_pair, ex3_2_intro/
42   | -HK12 -HV12 #Hi #HK10 elim (IHK12 … HK10) -IHK12 -HK10
43     /3 width=5 by drop_drop_lt, ex3_2_intro/
44   ]
45 ]
46 qed-.
47
48 lemma lpx_sn_deliftable_dropable: ∀R. d_deliftable_sn R → dropable_sn (lpx_sn R).
49 #R #HR #L1 #K1 #s #l #m #H elim H -L1 -K1 -l -m
50 [ #l #m #Hm #X #H >(lpx_sn_inv_atom1 … H) -H
51   /4 width=3 by drop_atom, lpx_sn_atom, ex2_intro/
52 | #I #K1 #V1 #X #H elim (lpx_sn_inv_pair1 … H) -H
53   #L2 #V2 #HL12 #HV12 #H destruct
54   /3 width=5 by drop_pair, lpx_sn_pair, ex2_intro/
55 | #I #L1 #K1 #V1 #m #_ #IHLK1 #X #H elim (lpx_sn_inv_pair1 … H) -H
56   #L2 #V2 #HL12 #HV12 #H destruct
57   elim (IHLK1 … HL12) -L1 /3 width=3 by drop_drop, ex2_intro/
58 | #I #L1 #K1 #V1 #W1 #l #m #HLK1 #HWV1 #IHLK1 #X #H
59   elim (lpx_sn_inv_pair1 … H) -H #L2 #V2 #HL12 #HV12 #H destruct
60   elim (HR … HV12 … HLK1 … HWV1) -V1
61   elim (IHLK1 … HL12) -L1 /3 width=5 by drop_skip, lpx_sn_pair, ex2_intro/
62 ]
63 qed-.
64
65 lemma lpx_sn_liftable_dedropable: ∀R. (∀L. reflexive ? (R L)) →
66                                   d_liftable R → dedropable_sn (lpx_sn R).
67 #R #H1R #H2R #L1 #K1 #s #l #m #H elim H -L1 -K1 -l -m
68 [ #l #m #Hm #X #H >(lpx_sn_inv_atom1 … H) -H
69   /4 width=4 by drop_atom, lpx_sn_atom, ex3_intro/
70 | #I #K1 #V1 #X #H elim (lpx_sn_inv_pair1 … H) -H
71   #K2 #V2 #HK12 #HV12 #H destruct
72   lapply (lpx_sn_fwd_length … HK12)
73   #H @(ex3_intro … (K2.ⓑ{I}V2)) (**) (* explicit constructor *)
74   /3 width=1 by lpx_sn_pair, monotonic_le_plus_l/
75   @lreq_O2 normalize //
76 | #I #L1 #K1 #V1 #m #_ #IHLK1 #K2 #HK12 elim (IHLK1 … HK12) -K1
77   /3 width=5 by drop_drop, lreq_pair, lpx_sn_pair, ex3_intro/
78 | #I #L1 #K1 #V1 #W1 #l #m #HLK1 #HWV1 #IHLK1 #X #H
79   elim (lpx_sn_inv_pair1 … H) -H #K2 #W2 #HK12 #HW12 #H destruct
80   elim (lift_total W2 l m) #V2 #HWV2
81   lapply (H2R … HW12 … HLK1 … HWV1 … HWV2) -W1
82   elim (IHLK1 … HK12) -K1
83   /3 width=6 by drop_skip, lreq_succ, lpx_sn_pair, ex3_intro/
84 ]
85 qed-.
86
87 fact lpx_sn_dropable_aux: ∀R,L2,K2,s,l,m. ⬇[s, l, m] L2 ≡ K2 → ∀L1. lpx_sn R L1 L2 →
88                           l = 0 → ∃∃K1. ⬇[s, 0, m] L1 ≡ K1 & lpx_sn R K1 K2.
89 #R #L2 #K2 #s #l #m #H elim H -L2 -K2 -l -m
90 [ #l #m #Hm #X #H >(lpx_sn_inv_atom2 … H) -H 
91   /4 width=3 by drop_atom, lpx_sn_atom, ex2_intro/
92 | #I #K2 #V2 #X #H elim (lpx_sn_inv_pair2 … H) -H
93   #K1 #V1 #HK12 #HV12 #H destruct
94   /3 width=5 by drop_pair, lpx_sn_pair, ex2_intro/
95 | #I #L2 #K2 #V2 #m #_ #IHLK2 #X #H #_ elim (lpx_sn_inv_pair2 … H) -H
96   #L1 #V1 #HL12 #HV12 #H destruct
97   elim (IHLK2 … HL12) -L2 /3 width=3 by drop_drop, ex2_intro/
98 | #I #L2 #K2 #V2 #W2 #l #m #_ #_ #_ #L1 #_
99   <plus_n_Sm #H destruct
100 ]
101 qed-.
102
103 lemma lpx_sn_dropable: ∀R. dropable_dx (lpx_sn R).
104 /2 width=5 by lpx_sn_dropable_aux/ qed-.