]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/unfold/lstas_da.ma
8d1feca9f9e09032abb2970e56812406b55e3281
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / unfold / lstas_da.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/da_da.ma".
16 include "basic_2/unfold/lstas_lstas.ma".
17
18 (* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
19
20 (* Properties on degree assignment for terms ********************************)
21
22 lemma da_lstas: ∀h,g,G,L,T,l1. ⦃G, L⦄ ⊢ T ▪[h, g] l1 → ∀l2.
23                 ∃∃U. ⦃G, L⦄ ⊢ T •*[h, l2] U & ⦃G, L⦄ ⊢ U ▪[h, g] l1-l2.
24 #h #g #G #L #T #l1 #H elim H -G -L -T -l1
25 [ /4 width=3 by da_sort, deg_iter, ex2_intro/
26 | #G #L #K #V #i #l1 #HLK #_ #IHV #l2
27   elim (IHV l2) -IHV #W
28   elim (lift_total W 0 (i+1))
29   lapply (drop_fwd_drop2 … HLK)
30   /3 width=10 by lstas_ldef, da_lift, ex2_intro/
31 | #G #L #K #W #i #l1 #HLK #HW #IHW #l2 @(nat_ind_plus … l2) -l2
32   [ elim (IHW 0) -IHW /3 width=6 by lstas_zero, da_ldec, ex2_intro/
33   | #l #_ elim (IHW l) -IHW #V
34     elim (lift_total V 0 (i+1))
35     lapply (drop_fwd_drop2 … HLK)
36     /3 width=10 by lstas_succ, da_lift, ex2_intro/
37   ]
38 | #a #I #G #L #V #T #l1 #_ #IHT #l2 elim (IHT … l2) -IHT
39   /3 width=6 by lstas_bind, da_bind, ex2_intro/
40 | * #G #L #V #T #l1 #_ #IHT #l2 elim (IHT … l2) -IHT
41   /3 width=5 by lstas_appl, lstas_cast, da_flat, ex2_intro/
42 ]
43 qed-.
44
45 lemma lstas_da_conf: ∀h,g,G,L,T,U,l2. ⦃G, L⦄ ⊢ T •*[h, l2] U →
46                      ∀l1. ⦃G, L⦄ ⊢ T ▪[h, g] l1 → ⦃G, L⦄ ⊢ U ▪[h, g] l1-l2.
47 #h #g #G #L #T #U #l2 #HTU #l1 #HT
48 elim (da_lstas … HT l2) -HT #X #HTX
49 lapply (lstas_mono … HTX … HTU) -T //
50 qed-.
51
52 (* inversion lemmas on degree assignment for terms **************************)
53
54 lemma lstas_inv_da: ∀h,g,G,L,T,U,l2. ⦃G, L⦄ ⊢ T •*[h, l2] U →
55                     ∃∃l1. ⦃G, L⦄ ⊢ T ▪[h, g] l1 & ⦃G, L⦄ ⊢ U ▪[h, g] l1-l2.
56 #h #g #G #L #T #U #l2 #H elim H -G -L -T -U -l2
57 [ #G #L #l2 #k elim (deg_total h g k) /4 width=3 by da_sort, deg_iter, ex2_intro/
58 | #G #L #K #V #W #U #i #l2 #HLK #_ #HWU *
59   lapply (drop_fwd_drop2 … HLK) /3 width=10 by da_lift, da_ldef, ex2_intro/
60 | #G #L #K #W #V #i #HLK #_ * /3 width=6 by da_ldec, ex2_intro/
61 | #G #L #K #W #V #U #i #l2 #HLK #_ #HVU *
62   lapply (drop_fwd_drop2 … HLK) /3 width=10 by da_lift, da_ldec, ex2_intro/
63 | #a #I #G #L #V #T #U #l2 #_ * /3 width=3 by da_bind, ex2_intro/
64 | #G #L #V #T #U #l2 #_ * /3 width=3 by da_flat, ex2_intro/
65 | #G #L #W #T #U #l2 #_ * /3 width=3 by da_flat, ex2_intro/
66 ]
67 qed-.
68
69 lemma lstas_inv_da_ge: ∀h,G,L,T,U,l2,l. ⦃G, L⦄ ⊢ T •*[h, l2] U →
70                        ∃∃g,l1. ⦃G, L⦄ ⊢ T ▪[h, g] l1 & ⦃G, L⦄ ⊢ U ▪[h, g] l1-l2 & l ≤ l1.
71 #h #G #L #T #U #l2 #l #H elim H -G -L -T -U -l2
72 [ /4 width=5 by da_sort, deg_iter, ex3_2_intro/
73 | #G #L #K #V #W #U #i #l2 #HLK #_ #HWU *
74   lapply (drop_fwd_drop2 … HLK) /3 width=10 by da_lift, da_ldef, ex3_2_intro/
75 | #G #L #K #W #V #i #HLK #_ * 
76   #g #l1 #HW #HV #Hl1 /4 width=6 by da_ldec, lt_to_le, le_S_S, ex3_2_intro/
77 | #G #L #K #W #V #U #i #l2 #HLK #_ #HVU *
78   lapply (drop_fwd_drop2 … HLK)
79   /4 width=10 by da_lift, da_ldec, lt_to_le, le_S_S, ex3_2_intro/
80 | #a #I #G #L #V #T #U #l2 #_ * /3 width=5 by da_bind, ex3_2_intro/
81 | #G #L #V #T #U #l2 #_ * /3 width=5 by da_flat, ex3_2_intro/
82 | #G #L #W #T #U #l2 #_ * /3 width=5 by da_flat, ex3_2_intro/
83 ]
84 qed-.
85
86 (* Advanced inversion lemmas ************************************************)
87
88 lemma lstas_inv_refl_pos: ∀h,G,L,T,l. ⦃G, L⦄ ⊢ T •*[h, l+1] T → ⊥.
89 #h #G #L #T #l2 #H elim (lstas_inv_da_ge … (l2+1) H) -H
90 #g #l1 #HT1 #HT12 #Hl21 lapply (da_mono … HT1 … HT12) -h -G -L -T
91 #H elim (discr_x_minus_xy … H) -H
92 [ #H destruct /2 width=3 by le_plus_xSy_O_false/
93 | -l1 <plus_n_Sm #H destruct 
94 ]
95 qed-.