]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/equivalence/lsubss_ssta.ma
milestone in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / equivalence / lsubss_ssta.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_2/static/ssta_ssta.ma".
16 include "basic_2/equivalence/cpcs_cpcs.ma".
17 include "basic_2/equivalence/lsubss_ldrop.ma".
18
19 (* LOCAL ENVIRONMENT REFINEMENT FOR STRATIFIED STATIC TYPE ASSIGNMENT *******)
20
21 (* Properties on stratified native type assignment **************************)
22
23 lemma lsubss_ssta_trans: ∀h,g,L2,T,U2,l. ⦃h, L2⦄ ⊢ T •[g] ⦃l, U2⦄ →
24                          ∀L1. h ⊢ L1 •⊑[g] L2 →
25                          ∃∃U1. ⦃h, L1⦄ ⊢ T •[g] ⦃l, U1⦄ & L1 ⊢ U1 ⬌* U2.
26 #h #g #L2 #T #U #l #H elim H -L2 -T -U -l
27 [ /3 width=3/
28 | #L2 #K2 #V2 #W2 #U2 #i #l #HLK2 #_ #HWU2 #IHVW2 #L1 #HL12
29   elim (lsubss_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
30   elim (lsubss_inv_pair2 … H) -H * #K1 [ | -HWU2 -IHVW2 -HLK1 ]
31   [ #HK12 #H destruct
32     elim (IHVW2 … HK12) -K2 #T2 #HVT2 #HTW2
33     lapply (ldrop_fwd_ldrop2 … HLK1) #H
34     elim (lift_total T2 0 (i+1)) /3 width=11/
35   | #W1 #V1 #W2 #l0 #_ #_ #_ #_ #_ #H destruct
36   ]
37 | #L2 #K2 #W2 #V2 #U2 #i #l #HLK2 #HWV2 #HWU2 #IHWV2 #L1 #HL12
38   elim (lsubss_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
39   elim (lsubss_inv_pair2 … H) -H * #K1 [ -HWV2 | -IHWV2 ]
40   [ #HK12 #H destruct
41     elim (IHWV2 … HK12) -K2 /3 width=6/
42   | #W1 #V1 #T2 #l0 #HVW1 #HWT2 #HW12 #_ #H #_ destruct
43     elim (ssta_mono … HWV2 … HWT2) -HWV2 -HWT2 #H1 #H2 destruct
44     lapply (ldrop_fwd_ldrop2 … HLK1) #H
45     elim (lift_total W1 0 (i+1)) /3 width=11/
46   ]
47 | #a #I #L2 #V2 #T2 #U2 #l #_ #IHTU2 #L1 #HL12
48   elim (IHTU2 (L1.ⓑ{I}V2) …) [2: /2 width=1/ ] -L2 /3 width=3/
49 | #L2 #V2 #T2 #U2 #l #_ #IHTU2 #L1 #HL12
50   elim (IHTU2 … HL12) -L2 /3 width=5/
51 | #L2 #W2 #T2 #U2 #l #_ #IHTU2 #L1 #HL12
52   elim (IHTU2 … HL12) -L2 /3 width=3/
53 ]
54 qed-.