]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/unfold/lstas_aaa.ma
universary milestone in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / unfold / lstas_aaa.ma
index 7497f44b6f5146880897970e7502c377960c0207..405f62ba2dd033420a8d04cc80901b9ff2439c33 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/static/sta_aaa.ma".
-include "basic_2/unfold/lstas.ma".
+include "basic_2/static/aaa_lift.ma".
+include "basic_2/unfold/lstas_lstas.ma".
 
 (* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
 
 (* Properties on atomic arity assignment for terms **************************)
 
-lemma lstas_aaa_conf: ∀h,G,L,l. Conf3 … (aaa G L) (lstas h G L l).
-/3 width=6 by sta_aaa_conf, lstar_Conf3/ qed-.
+lemma aaa_lstas: ∀h,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀d.
+                 ∃∃U. ⦃G, L⦄ ⊢ T •*[h, d] U & ⦃G, L⦄ ⊢ U ⁝ A.
+#h #G #L #T #A #H elim H -G -L -T -A
+[ /2 width=3 by ex2_intro/
+| * #G #L #K #V #B #i #HLK #HV #IHV #d
+  [ elim (IHV d) -IHV #W
+    elim (lift_total W 0 (i+1))
+    lapply (drop_fwd_drop2 … HLK)
+    /3 width=10 by lstas_ldef, aaa_lift, ex2_intro/
+  | @(nat_ind_plus … d) -d
+    [ elim (IHV 0) -IHV /3 width=7 by lstas_zero, aaa_lref, ex2_intro/
+    | #d #_ elim (IHV d) -IHV #W
+      elim (lift_total W 0 (i+1))
+      lapply (drop_fwd_drop2 … HLK)
+      /3 width=10 by lstas_succ, aaa_lift, ex2_intro/
+    ]
+  ]
+| #a #G #L #V #T #B #A #HV #_ #_ #IHT #d elim (IHT d) -IHT
+  /3 width=7 by lstas_bind, aaa_abbr, ex2_intro/
+| #a #G #L #V #T #B #A #HV #_ #_ #IHT #d elim (IHT d) -IHT
+  /3 width=6 by lstas_bind, aaa_abst, ex2_intro/
+| #G #L #V #T #B #A #HV #_ #_ #IHT #d elim (IHT d) -IHT
+  /3 width=6 by lstas_appl, aaa_appl, ex2_intro/
+| #G #L #W #T #A #HW #_ #_ #IHT #d elim (IHT d) -IHT
+  /3 width=3 by lstas_cast, aaa_cast, ex2_intro/
+]
+qed-.
+
+lemma lstas_aaa_conf: ∀h,G,L,d. Conf3 … (aaa G L) (lstas h d G L).
+#h #G #L #d #A #T #HT #U #HTU
+elim (aaa_lstas h … HT d) -HT #X #HTX
+lapply (lstas_mono … HTX … HTU) -T //
+qed-.