]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2A/etc/cpr/ssta_ltpss_dx.etc
milestone update in ground_2 and basic_2A
[helm.git] / matita / matita / contribs / lambdadelta / basic_2A / etc / cpr / ssta_ltpss_dx.etc
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_2/unfold/tpss_tpss.ma".
16 include "basic_2/unfold/ltpss_dx_tpss.ma".
17 include "basic_2/static/ssta_lift.ma".
18
19 (* STRATIFIED STATIC TYPE ASSIGNMENT ON TERMS *******************************)
20
21 (* Properties about dx parallel unfold **************************************)
22
23 (* Note: apparently this was missing in basic_1 *)
24 lemma ssta_ltpss_dx_tpss_conf: ∀h,g,L1,T1,U1,l. ⦃h, L1⦄ ⊢ T1 •[g] ⦃l, U1⦄ →
25                                ∀L2,d,e. L1 ▶* [d, e] L2 →
26                                ∀T2. L2 ⊢ T1 ▶* [d, e] T2 →
27                                ∃∃U2. ⦃h, L2⦄ ⊢ T2 •[g] ⦃l, U2⦄ &
28                                      L2 ⊢ U1 ▶* [d, e] U2.
29 #h #g #L1 #T1 #U #l #H elim H -L1 -T1 -U -l
30 [ #L1 #k1 #l1 #Hkl1 #L2 #d #e #_ #T2 #H
31   >(tpss_inv_sort1 … H) -H /3 width=3/
32 | #L1 #K1 #V1 #W1 #U1 #i #l #HLK1 #HVW1 #HWU1 #IHVW1 #L2 #d #e #HL12 #T2 #H
33   elim (tpss_inv_lref1 … H) -H [ | -HVW1 ]
34   [ #H destruct
35     elim (lt_or_ge i d) #Hdi [ -HVW1 | ]
36     [ elim (ltpss_dx_ldrop_conf_le … HL12 … HLK1 ?) -L1 /2 width=2/ #X #H #HLK2
37       elim (ltpss_dx_inv_tpss11 … H ?) -H /2 width=1/ #K2 #V2 #HK12 #HV12 #H destruct
38       elim (IHVW1 … HK12 … HV12) -IHVW1 -HK12 -HV12 #W2 #HVW2 #HW12
39       lapply (ldrop_fwd_ldrop2 … HLK2) #H
40       elim (lift_total W2 0 (i+1)) #U2 #HWU2
41       lapply (tpss_lift_ge … HW12 … H … HWU1 … HWU2) // -HW12 -H -HWU1
42       >minus_plus <plus_minus_m_m // -Hdi /3 width=6/
43     | elim (lt_or_ge i (d + e)) #Hide [ -HVW1 | -Hdi -IHVW1 ]
44       [ elim (ltpss_dx_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HLK2
45         elim (ltpss_dx_inv_tpss21 … H ?) -H /2 width=1/ #K2 #V2 #HK12 #HV12 #H destruct
46         elim (IHVW1 … HK12 … HV12) -IHVW1 -HK12 -HV12 #W2 #HVW2 #HW12
47         lapply (ldrop_fwd_ldrop2 … HLK2) #H
48         elim (lift_total W2 0 (i+1)) #U2 #HWU2
49         lapply (tpss_lift_ge … HW12 … H … HWU1 … HWU2) // -HW12 -H -HWU1 >minus_plus #H
50         lapply (tpss_weak … H d e ? ?) [1,2: normalize [ >commutative_plus <plus_minus_m_m // | /2 width=1/ ]] -Hdi -Hide /3 width=6/
51       | lapply (ltpss_dx_ldrop_conf_ge … HL12 … HLK1 ?) -L1 // -Hide /3 width=6/
52       ]
53     ]
54   | * #K2 #V2 #W2 #Hdi #Hide #HLK2 #HVW2 #HWT2
55     elim (ltpss_dx_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HL2K0
56     elim (ltpss_dx_inv_tpss21 … H ?) -H /2 width=1/ #K0 #V0 #HK12 #HV12 #H destruct
57     lapply (ldrop_mono … HL2K0 … HLK2) -HL2K0 #H destruct
58     lapply (ldrop_fwd_ldrop2 … HLK2) -HLK2 #HLK2
59     lapply (tpss_trans_eq … HV12 HVW2) -V2 #HV1W2
60     elim (IHVW1 … HK12 … HV1W2) -IHVW1 -HK12 -HV1W2 #V2 #HWV2 #HW1V2
61     elim (lift_total V2 0 (i+1)) #U2 #HVU2
62     lapply (ssta_lift … HWV2 … HLK2 … HWT2 … HVU2) -HWV2 -HWT2 #HTU2
63     lapply (tpss_lift_ge … HW1V2 … HLK2 … HWU1 … HVU2) // -HW1V2 -HLK2 -HWU1 -HVU2 >minus_plus #H
64     lapply (tpss_weak … H d e ? ?) [1,2: normalize [ >commutative_plus <plus_minus_m_m // | /2 width=1/ ]] -Hdi -Hide /2 width=3/
65   ]
66 | #L1 #K1 #W1 #V1 #U1 #i #l #HLK1 #HWV1 #HWU1 #IHWV1 #L2 #d #e #HL12 #T2 #H
67   elim (tpss_inv_lref1 … H) -H [ | -HWV1 -HWU1 -IHWV1 ]
68   [ #H destruct
69     elim (lt_or_ge i d) #Hdi [ -HWV1 ]
70     [ elim (ltpss_dx_ldrop_conf_le … HL12 … HLK1 ?) -L1 /2 width=2/ #X #H #HLK2
71       elim (ltpss_dx_inv_tpss11 … H ?) -H /2 width=1/ #K2 #W2 #HK12 #HW12 #H destruct
72       elim (IHWV1 … HK12 … HW12) -IHWV1 -HK12 #V2 #HWV2 #_
73       lapply (ldrop_fwd_ldrop2 … HLK2) #HLK
74       elim (lift_total W2 0 (i+1)) #U2 #HWU2
75       lapply (tpss_lift_ge … HW12 … HLK … HWU1 … HWU2) // -HW12 -HLK -HWU1
76       >minus_plus <plus_minus_m_m // -Hdi /3 width=6/
77     | elim (lt_or_ge i (d + e)) #Hide [ -HWV1 | -IHWV1 -Hdi ]
78       [ elim (ltpss_dx_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HLK2
79         elim (ltpss_dx_inv_tpss21 … H ?) -H /2 width=1/ #K2 #W2 #HK12 #HW12 #H destruct
80         elim (IHWV1 … HK12 … HW12) -IHWV1 -HK12 #V2 #HWV2 #_
81         lapply (ldrop_fwd_ldrop2 … HLK2) #HLK
82         elim (lift_total W2 0 (i+1)) #U2 #HWU2
83         lapply (tpss_lift_ge … HW12 … HLK … HWU1 … HWU2) // -HW12 -HLK -HWU1 >minus_plus #H
84         lapply (tpss_weak … H d e ? ?) [1,2: normalize [ >commutative_plus <plus_minus_m_m // | /2 width=1/ ]] -Hdi -Hide /3 width=6/
85       | lapply (ltpss_dx_ldrop_conf_ge … HL12 … HLK1 ?) -L1 // -Hide /3 width=6/
86       ]
87     ]
88   | * #K2 #V2 #W2 #Hdi #Hide #HLK2 #_ #_
89     elim (ltpss_dx_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HL2K0
90     elim (ltpss_dx_inv_tpss21 … H ?) -H /2 width=1/ -Hdi -Hide #K0 #V0 #_ #_ #H destruct
91     lapply (ldrop_mono … HL2K0 … HLK2) -HL2K0 -HLK2 #H destruct
92   ]
93 | #a #I #L1 #V1 #T1 #U1 #l #_ #IHTU1 #L2 #d #e #HL12 #X #H
94   elim (tpss_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
95   elim (IHTU1 … HT12) -IHTU1 -HT12 /2 width=1/ -HL12 /3 width=5/
96 | #L1 #V1 #T1 #U1 #l #_ #IHTU1 #L2 #d #e #HL12 #X #H
97   elim (tpss_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
98   elim (IHTU1 … HT12) -IHTU1 -HT12 // -HL12 /3 width=5/
99 | #L1 #W1 #T1 #U1 #l #_ #IHTU1 #L2 #d #e #HL12 #X #H
100   elim (tpss_inv_flat1 … H) -H #W2 #T2 #HW12 #HT12 #H destruct
101   elim (IHTU1 … HT12) -IHTU1 -HT12 // -HL12 /3 width=3/
102 ]
103 qed.
104
105 lemma ssta_ltpss_dx_tps_conf: ∀h,g,L1,T1,U1,l. ⦃h, L1⦄ ⊢ T1 •[g] ⦃l, U1⦄ →
106                               ∀L2,d,e. L1 ▶* [d, e] L2 →
107                               ∀T2. L2 ⊢ T1 ▶ [d, e] T2 →
108                               ∃∃U2. ⦃h, L2⦄ ⊢ T2 •[g] ⦃l, U2⦄ &
109                                     L2 ⊢ U1 ▶* [d, e] U2.
110 /3 width=5/ qed.
111
112 lemma ssta_ltpss_dx_conf: ∀h,g,L1,T,U1,l. ⦃h, L1⦄ ⊢ T •[g] ⦃l, U1⦄ →
113                           ∀L2,d,e. L1 ▶* [d, e] L2 →
114                           ∃∃U2. ⦃h, L2⦄ ⊢ T •[g] ⦃l, U2⦄ & L2 ⊢ U1 ▶* [d, e] U2.
115 /2 width=5/ qed.
116
117 lemma ssta_tpss_conf: ∀h,g,L,T1,U1,l. ⦃h, L⦄ ⊢ T1 •[g] ⦃l, U1⦄ →
118                       ∀T2,d,e. L ⊢ T1 ▶* [d, e] T2 →
119                       ∃∃U2. ⦃h, L⦄ ⊢ T2 •[g] ⦃l, U2⦄ & L ⊢ U1 ▶* [d, e] U2.
120 /2 width=5/ qed.
121
122 lemma ssta_tps_conf: ∀h,g,L,T1,U1,l. ⦃h, L⦄ ⊢ T1 •[g] ⦃l, U1⦄ →
123                      ∀T2,d,e. L ⊢ T1 ▶ [d, e] T2 →
124                      ∃∃U2. ⦃h, L⦄ ⊢ T2 •[g] ⦃l, U2⦄ & L ⊢ U1 ▶* [d, e] U2.
125 /2 width=5/ qed.