]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2A/unfold/lstas_aaa.ma
update in lambdadelta
[helm.git] / matita / matita / contribs / lambdadelta / basic_2A / unfold / lstas_aaa.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/static/aaa_lift.ma".
16 include "basic_2A/unfold/lstas_lstas.ma".
17
18 (* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
19
20 (* Properties on atomic arity assignment for terms **************************)
21
22 lemma aaa_lstas: ∀h,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀d.
23                  ∃∃U. ⦃G, L⦄ ⊢ T •*[h, d] U & ⦃G, L⦄ ⊢ U ⁝ A.
24 #h #G #L #T #A #H elim H -G -L -T -A
25 [ /2 width=3 by ex2_intro/
26 | * #G #L #K #V #B #i #HLK #HV #IHV #d
27   [ elim (IHV d) -IHV #W
28     elim (lift_total W 0 (i+1))
29     lapply (drop_fwd_drop2 … HLK)
30     /3 width=10 by lstas_ldef, aaa_lift, ex2_intro/
31   | @(nat_ind_plus … d) -d
32     [ elim (IHV 0) -IHV /3 width=7 by lstas_zero, aaa_lref, ex2_intro/
33     | #d #_ elim (IHV d) -IHV #W
34       elim (lift_total W 0 (i+1))
35       lapply (drop_fwd_drop2 … HLK)
36       /3 width=10 by lstas_succ, aaa_lift, ex2_intro/
37     ]
38   ]
39 | #a #G #L #V #T #B #A #HV #_ #_ #IHT #d elim (IHT d) -IHT
40   /3 width=7 by lstas_bind, aaa_abbr, ex2_intro/
41 | #a #G #L #V #T #B #A #HV #_ #_ #IHT #d elim (IHT d) -IHT
42   /3 width=6 by lstas_bind, aaa_abst, ex2_intro/
43 | #G #L #V #T #B #A #HV #_ #_ #IHT #d elim (IHT d) -IHT
44   /3 width=6 by lstas_appl, aaa_appl, ex2_intro/
45 | #G #L #W #T #A #HW #_ #_ #IHT #d elim (IHT d) -IHT
46   /3 width=3 by lstas_cast, aaa_cast, ex2_intro/
47 ]
48 qed-.
49
50 lemma lstas_aaa_conf: ∀h,G,L,d. Conf3 … (aaa G L) (lstas h d G L).
51 #h #G #L #d #A #T #HT #U #HTU
52 elim (aaa_lstas h … HT d) -HT #X #HTX
53 lapply (lstas_mono … HTX … HTU) -T //
54 qed-.