]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/etc/da/da.etc
preservation of stratified vaildity through ordinary reduction and static typing
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / etc / da / da.etc
1 (* Inversion lrmmas on static type assignment for terms *********************)
2
3 lemma da_inv_sta: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l →
4                   ∃U. ⦃G, L⦄ ⊢ T •[h] U.
5 #h #g #G #L #T #l #H elim H -G -L -T -l
6 [ /2 width=2/
7 | #G #L #K #V #i #l #HLK #_ * #W #HVW
8   elim (lift_total W 0 (i+1)) /3 width=7/
9 | #G #L #K #W #i #l #HLK #_ * #V #HWV
10   elim (lift_total W 0 (i+1)) /3 width=7/
11 | #a #I #G #L #V #T #l #_ * /3 width=2/
12 | * #G #L #V #T #l #_ * /3 width=2/
13 ]
14 qed-.
15
16 (* Properties on static type assignment for terms ***************************)
17
18 lemma sta_da: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U →
19               ∃l. ⦃G, L⦄ ⊢ T ▪[h, g] l.
20 #h #g #G #L #T #U #H elim H -G -L -T -U
21 [ #G #L #k elim (deg_total h g k) /3 width=2/
22 | #G #L #K #V #W #W0 #i #HLK #_ #_ * /3 width=5/
23 | #G #L #K #W #V #W0 #i #HLK #_ #_ * /3 width=5/
24 | #a #I #G #L #V #T #U #_ * /3 width=2/
25 | #G #L #V #T #U #_ * /3 width=2/
26 | #G #L #W #T #U #_ * /3 width=2/
27 ]
28 qed-.