]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2A/unfold/lstas_lstas.ma
update in lambdadelta
[helm.git] / matita / matita / contribs / lambdadelta / basic_2A / unfold / lstas_lstas.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/unfold/lstas_lift.ma".
16
17 (* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
18
19 (* Main properties **********************************************************)
20
21 theorem lstas_trans: ∀h,G,L,T1,T,d1. ⦃G, L⦄ ⊢ T1 •*[h, d1] T →
22                      ∀T2,d2. ⦃G, L⦄ ⊢ T •*[h, d2] T2 → ⦃G, L⦄ ⊢ T1 •*[h, d1+d2] T2.
23 #h #G #L #T1 #T #d1 #H elim H -G -L -T1 -T -d1
24 [ #G #L #d1 #k #X #d2 #H >(lstas_inv_sort1 … H) -X
25   <iter_plus /2 width=1 by lstas_sort/
26 | #G #L #K #V1 #V #U #i #d1 #HLK #_ #HVU #IHV1 #U2 #d2 #HU2
27   lapply (drop_fwd_drop2 … HLK) #H0
28   elim (lstas_inv_lift1 … HU2 … H0 … HVU)
29   /3 width=6 by lstas_ldef/
30 | //
31 | #G #L #K #W1 #W #U #i #d1 #HLK #_ #HWU #IHW1 #U2 #d2 #HU2
32   lapply (drop_fwd_drop2 … HLK) #H0
33   elim (lstas_inv_lift1 … HU2 … H0 … HWU)
34   /3 width=6 by lstas_succ/
35 | #a #I #G #L #V #T1 #T #d1 #_ #IHT1 #X #d2 #H
36   elim (lstas_inv_bind1 … H) -H #T2 #HT2 #H destruct
37   /3 width=1 by lstas_bind/
38 | #G #L #V #T1 #T #d1 #_ #IHT1 #X #d2 #H
39   elim (lstas_inv_appl1 … H) -H #T2 #HT2 #H destruct
40   /3 width=1 by lstas_appl/
41 | /3 width=1 by lstas_cast/
42 ]
43 qed-.
44
45 (* Note: apparently this was missing in basic_1 *)
46 theorem lstas_mono: ∀h,G,L,d. singlevalued … (lstas h d G L).
47 #h #G #L #d #T #T1 #H elim H -G -L -T -T1 -d
48 [ #G #L #d #k #X #H >(lstas_inv_sort1 … H) -X //
49 | #G #L #K #V #V1 #U1 #i #d #HLK #_ #HVU1 #IHV1 #X #H
50   elim (lstas_inv_lref1 … H) -H *
51   #K0 #V0 #W0 [3: #d0 ] #HLK0
52   lapply (drop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct
53   #HVW0 #HX lapply (IHV1 … HVW0) -IHV1 -HVW0 #H destruct
54   /2 width=5 by lift_mono/
55 | #G #L #K #W #W1 #i #HLK #_ #_ #X #H
56   elim (lstas_inv_lref1_O … H) -H *
57   #K0 #V0 #W0 #HLK0
58   lapply (drop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct //
59 | #G #L #K #W #W1 #U1 #i #d #HLK #_ #HWU1 #IHWV #X #H
60   elim (lstas_inv_lref1_S … H) -H * #K0 #W0 #V0 #HLK0
61   lapply (drop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct
62   #HW0 #HX lapply (IHWV … HW0) -IHWV -HW0 #H destruct
63   /2 width=5 by lift_mono/
64 | #a #I #G #L #V #T #U1 #d #_ #IHTU1 #X #H
65   elim (lstas_inv_bind1 … H) -H #U2 #HTU2 #H destruct /3 width=1 by eq_f/
66 | #G #L #V #T #U1 #d #_ #IHTU1 #X #H
67   elim (lstas_inv_appl1 … H) -H #U2 #HTU2 #H destruct /3 width=1 by eq_f/
68 | #G #L #W #T #U1 #d #_ #IHTU1 #U2 #H
69   lapply (lstas_inv_cast1 … H) -H /2 width=1 by/
70 ]
71 qed-.
72
73 (* Advanced inversion lemmas ************************************************)
74
75 (* Basic_1: was just: sty0_correct *)
76 lemma lstas_correct: ∀h,G,L,T1,T,d1. ⦃G, L⦄ ⊢ T1 •*[h, d1] T →
77                      ∀d2. ∃T2. ⦃G, L⦄ ⊢ T •*[h, d2] T2.
78 #h #G #L #T1 #T #d1 #H elim H -G -L -T1 -T -d1
79 [ /2 width=2 by lstas_sort, ex_intro/
80 | #G #L #K #V1 #V #U #i #d #HLK #_ #HVU #IHV1 #d2
81   elim (IHV1 d2) -IHV1 #V2
82   elim (lift_total V2 0 (i+1))
83   lapply (drop_fwd_drop2 … HLK) -HLK
84   /3 width=11 by ex_intro, lstas_lift/
85 | #G #L #K #W1 #W #i #HLK #HW1 #IHW1 #d2
86   @(nat_ind_plus … d2) -d2 /3 width=5 by lstas_zero, ex_intro/
87   #d2 #_ elim (IHW1 d2) -IHW1 #W2 #HW2
88   lapply (lstas_trans … HW1 … HW2) -W
89   elim (lift_total W2 0 (i+1))
90   /3 width=7 by lstas_succ, ex_intro/
91 | #G #L #K #W1 #W #U #i #d #HLK #_ #HWU #IHW1 #d2
92   elim (IHW1 d2) -IHW1 #W2
93   elim (lift_total W2 0 (i+1))
94   lapply (drop_fwd_drop2 … HLK) -HLK
95   /3 width=11 by ex_intro, lstas_lift/
96 | #a #I #G #L #V #T #U #d #_ #IHTU #d2
97   elim (IHTU d2) -IHTU /3 width=2 by lstas_bind, ex_intro/
98 | #G #L #V #T #U #d #_ #IHTU #d2
99   elim (IHTU d2) -IHTU /3 width=2 by lstas_appl, ex_intro/
100 | #G #L #W #T #U #d #_ #IHTU #d2
101   elim (IHTU d2) -IHTU /2 width=2 by ex_intro/
102 ]
103 qed-.
104
105 (* more main properties *****************************************************)
106
107 theorem lstas_conf_le: ∀h,G,L,T,U1,d1. ⦃G, L⦄ ⊢ T •*[h, d1] U1 →
108                        ∀U2,d2. d1 ≤ d2 → ⦃G, L⦄ ⊢ T •*[h, d2] U2 →
109                        ⦃G, L⦄ ⊢ U1 •*[h, d2-d1] U2.
110 #h #G #L #T #U1 #d1 #HTU1 #U2 #d2 #Hd12
111 >(plus_minus_m_m … Hd12) in ⊢ (%→?); -Hd12 >commutative_plus #H
112 elim (lstas_split … H) -H #U #HTU
113 >(lstas_mono … HTU … HTU1) -T //
114 qed-.
115
116 theorem lstas_conf: ∀h,G,L,T0,T1,d1. ⦃G, L⦄ ⊢ T0 •*[h, d1] T1 →
117                     ∀T2,d2. ⦃G, L⦄ ⊢ T0 •*[h, d2] T2 →
118                     ∃∃T. ⦃G, L⦄ ⊢ T1 •*[h, d2] T & ⦃G, L⦄ ⊢ T2 •*[h, d1] T.
119 #h #G #L #T0 #T1 #d1 #HT01 #T2 #d2 #HT02
120 elim (lstas_lstas … HT01 (d1+d2)) #T #HT0
121 lapply (lstas_conf_le … HT01 … HT0) // -HT01 <minus_plus_m_m_commutative
122 lapply (lstas_conf_le … HT02 … HT0) // -T0 <minus_plus_m_m
123 /2 width=3 by ex2_intro/
124 qed-.