--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/static/da_sta.ma".
+include "basic_2/static/sta_aaa.ma".
+
+(* DEGREE ASSIGNMENT FOR TERMS **********************************************)
+
+(* Properties on atomic arity assignment for terms **************************)
+
+lemma aaa_da: ∀h,g,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∃l. ⦃G, L⦄ ⊢ T ▪[h, g] l.
+#h #g #G #L #T #A #H elim (aaa_sta h … H) -A /2 width=2 by sta_da/
+qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/static/da_sta.ma".
-include "basic_2/static/sta_aaa.ma".
-
-(* DEGREE ASSIGNMENT FOR TERMS **********************************************)
-
-(* Properties on atomic arity assignment for terms **************************)
-
-lemma aaa_da: ∀h,g,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∃l. ⦃G, L⦄ ⊢ T ▪[h, g] l.
-#h #g #G #L #T #A #H elim (aaa_sta h … H) -A /2 width=2 by sta_da/
-qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/static/sta.ma".
+include "basic_2/static/da_da.ma".
+
+(* Properties on static type assignment for terms ***************************)
+
+lemma da_sta_conf: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U →
+ ∀l. ⦃G, L⦄ ⊢ T ▪[h, g] l → ⦃G, L⦄ ⊢ U ▪[h, g] l-1.
+#h #g #G #L #T #U #H elim H -G -L -T -U
+[ #G #L #k #l #H
+ lapply (da_inv_sort … H) -H /3 width=1 by da_sort, deg_next/
+| #G #L #K #V #U #W #i #HLK #_ #HWU #IHVW #l #H
+ elim (da_inv_lref … H) -H * #K0 #V0 [| #l0] #HLK0 #HV0
+ lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct
+ lapply (drop_fwd_drop2 … HLK) -HLK /3 width=8 by da_lift/
+| #G #L #K #W #V #U #i #HLK #_ #HWU #IHWV #l #H
+ elim (da_inv_lref … H) -H * #K0 #V0 [| #l0] #HLK0 #HV0 [| #H0 ]
+ lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct
+ lapply (drop_fwd_drop2 … HLK) -HLK /3 width=8 by da_lift/
+| #a #I #G #L #V #T #U #_ #IHTU #l #H
+ lapply (da_inv_bind … H) -H /3 width=1 by da_bind/
+| #G #L #V #T #U #_ #IHTU #l #H
+ lapply (da_inv_flat … H) -H /3 width=1 by da_flat/
+| #G #L #W #T #U #_ #IHTU #l #H
+ lapply (da_inv_flat … H) -H /2 width=1 by/
+]
+qed-.
+
+lemma sta_da: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U →
+ ∃l. ⦃G, L⦄ ⊢ T ▪[h, g] l.
+#h #g #G #L #T #U #H elim H -G -L -T -U
+[ #G #L #k elim (deg_total h g k) /3 width=2 by da_sort, ex_intro/
+| #G #L #K #V #W #W0 #i #HLK #_ #_ * /3 width=5 by da_ldef, ex_intro/
+| #G #L #K #W #V #W0 #i #HLK #_ #_ * /3 width=5 by da_ldec, ex_intro/
+| #a #I #G #L #V #T #U #_ * /3 width=2 by da_bind, ex_intro/
+| #G #L #V #T #U #_ * /3 width=2 by da_flat, ex_intro/
+| #G #L #W #T #U #_ * /3 width=2 by da_flat, ex_intro/
+]
+qed-.
+
+lemma sta_da_ge: ∀h,G,L,T,U,l0. ⦃G, L⦄ ⊢ T •[h] U →
+ ∃∃g,l. ⦃G, L⦄ ⊢ T ▪[h, g] l & l0 ≤ l.
+#h #G #L #T #U #l0 #H elim H -G -L -T -U
+[ /3 width=4 by da_sort, ex2_2_intro/
+| #G #L #K #V #W #W0 #i #HLK #_ #_ * /3 width=5 by da_ldef, ex2_2_intro/
+| #G #L #K #W #V #W0 #i #HLK #_ #_ * /4 width=5 by da_ldec, lt_to_le, le_S_S, ex2_2_intro/
+| #a #I #G #L #V #T #U #_ * /3 width=4 by da_bind, ex2_2_intro/
+| #G #L #V #T #U #_ * /3 width=4 by da_flat, ex2_2_intro/
+| #G #L #W #T #U #_ * /3 width=4 by da_flat, ex2_2_intro/
+]
+qed-.
+
+(* Inversion lrmmas on static type assignment for terms *********************)
+
+lemma da_inv_sta: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l →
+ ∃U. ⦃G, L⦄ ⊢ T •[h] U.
+#h #g #G #L #T #l #H elim H -G -L -T -l
+[ /2 width=2/
+| #G #L #K #V #i #l #HLK #_ * #W #HVW
+ elim (lift_total W 0 (i+1)) /3 width=7 by sta_ldef, ex_intro/
+| #G #L #K #W #i #l #HLK #_ * #V #HWV
+ elim (lift_total W 0 (i+1)) /3 width=7 by sta_ldec, ex_intro/
+| #a #I #G #L #V #T #l #_ * /3 width=2 by sta_bind, ex_intro/
+| * #G #L #V #T #l #_ * /3 width=2 by sta_appl, sta_cast, ex_intro/
+]
+qed-.
+
+lemma sta_inv_refl_pos: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l+1 → ⦃G, L⦄ ⊢ T •[h] T → ⊥.
+#h #g #G #L #T #l #H1T #HTT
+lapply (da_sta_conf … HTT … H1T) -HTT <minus_plus_m_m #H2T
+lapply (da_mono … H2T … H1T) -h -G -L -T #H
+elim (plus_xySz_x_false 0 l 0) //
+qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/static/sta.ma".
-include "basic_2/static/da_da.ma".
-
-(* Properties on static type assignment for terms ***************************)
-
-lemma da_sta_conf: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U →
- ∀l. ⦃G, L⦄ ⊢ T ▪[h, g] l → ⦃G, L⦄ ⊢ U ▪[h, g] l-1.
-#h #g #G #L #T #U #H elim H -G -L -T -U
-[ #G #L #k #l #H
- lapply (da_inv_sort … H) -H /3 width=1 by da_sort, deg_next/
-| #G #L #K #V #U #W #i #HLK #_ #HWU #IHVW #l #H
- elim (da_inv_lref … H) -H * #K0 #V0 [| #l0] #HLK0 #HV0
- lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct
- lapply (drop_fwd_drop2 … HLK) -HLK /3 width=8 by da_lift/
-| #G #L #K #W #V #U #i #HLK #_ #HWU #IHWV #l #H
- elim (da_inv_lref … H) -H * #K0 #V0 [| #l0] #HLK0 #HV0 [| #H0 ]
- lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct
- lapply (drop_fwd_drop2 … HLK) -HLK /3 width=8 by da_lift/
-| #a #I #G #L #V #T #U #_ #IHTU #l #H
- lapply (da_inv_bind … H) -H /3 width=1 by da_bind/
-| #G #L #V #T #U #_ #IHTU #l #H
- lapply (da_inv_flat … H) -H /3 width=1 by da_flat/
-| #G #L #W #T #U #_ #IHTU #l #H
- lapply (da_inv_flat … H) -H /2 width=1 by/
-]
-qed-.
-
-lemma sta_da: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U →
- ∃l. ⦃G, L⦄ ⊢ T ▪[h, g] l.
-#h #g #G #L #T #U #H elim H -G -L -T -U
-[ #G #L #k elim (deg_total h g k) /3 width=2 by da_sort, ex_intro/
-| #G #L #K #V #W #W0 #i #HLK #_ #_ * /3 width=5 by da_ldef, ex_intro/
-| #G #L #K #W #V #W0 #i #HLK #_ #_ * /3 width=5 by da_ldec, ex_intro/
-| #a #I #G #L #V #T #U #_ * /3 width=2 by da_bind, ex_intro/
-| #G #L #V #T #U #_ * /3 width=2 by da_flat, ex_intro/
-| #G #L #W #T #U #_ * /3 width=2 by da_flat, ex_intro/
-]
-qed-.
-
-lemma sta_da_ge: ∀h,G,L,T,U,l0. ⦃G, L⦄ ⊢ T •[h] U →
- ∃∃g,l. ⦃G, L⦄ ⊢ T ▪[h, g] l & l0 ≤ l.
-#h #G #L #T #U #l0 #H elim H -G -L -T -U
-[ /3 width=4 by da_sort, ex2_2_intro/
-| #G #L #K #V #W #W0 #i #HLK #_ #_ * /3 width=5 by da_ldef, ex2_2_intro/
-| #G #L #K #W #V #W0 #i #HLK #_ #_ * /4 width=5 by da_ldec, lt_to_le, le_S_S, ex2_2_intro/
-| #a #I #G #L #V #T #U #_ * /3 width=4 by da_bind, ex2_2_intro/
-| #G #L #V #T #U #_ * /3 width=4 by da_flat, ex2_2_intro/
-| #G #L #W #T #U #_ * /3 width=4 by da_flat, ex2_2_intro/
-]
-qed-.
-
-(* Inversion lrmmas on static type assignment for terms *********************)
-
-lemma da_inv_sta: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l →
- ∃U. ⦃G, L⦄ ⊢ T •[h] U.
-#h #g #G #L #T #l #H elim H -G -L -T -l
-[ /2 width=2/
-| #G #L #K #V #i #l #HLK #_ * #W #HVW
- elim (lift_total W 0 (i+1)) /3 width=7 by sta_ldef, ex_intro/
-| #G #L #K #W #i #l #HLK #_ * #V #HWV
- elim (lift_total W 0 (i+1)) /3 width=7 by sta_ldec, ex_intro/
-| #a #I #G #L #V #T #l #_ * /3 width=2 by sta_bind, ex_intro/
-| * #G #L #V #T #l #_ * /3 width=2 by sta_appl, sta_cast, ex_intro/
-]
-qed-.
-
-lemma sta_inv_refl_pos: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l+1 → ⦃G, L⦄ ⊢ T •[h] T → ⊥.
-#h #g #G #L #T #l #H1T #HTT
-lapply (da_sta_conf … HTT … H1T) -HTT <minus_plus_m_m #H2T
-lapply (da_mono … H2T … H1T) -h -G -L -T #H
-elim (plus_xySz_x_false 0 l 0) //
-qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/statictypestar_6.ma".
+include "basic_2/static/sta.ma".
+
+(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
+
+definition lstas: ∀h. genv → lenv → nat → relation term ≝
+ λh,G,L. lstar … (sta h G L).
+
+interpretation "nat-iterated static type assignment (term)"
+ 'StaticTypeStar h G L l T U = (lstas h G L l T U).
+
+(* Basic eliminators ********************************************************)
+
+lemma lstas_ind_sn: ∀h,G,L,U2. ∀R:relation2 nat term.
+ R 0 U2 → (
+ ∀l,T,U1. ⦃G, L⦄ ⊢ T •[h] U1 → ⦃G, L⦄ ⊢ U1 •* [h, l] U2 →
+ R l U1 → R (l+1) T
+ ) →
+ ∀l,T. ⦃G, L⦄ ⊢ T •*[h, l] U2 → R l T.
+/3 width=5 by lstar_ind_l/ qed-.
+
+lemma lstas_ind_dx: ∀h,G,L,T. ∀R:relation2 nat term.
+ R 0 T → (
+ ∀l,U1,U2. ⦃G, L⦄ ⊢ T •* [h, l] U1 → ⦃G, L⦄ ⊢ U1 •[h] U2 →
+ R l U1 → R (l+1) U2
+ ) →
+ ∀l,U. ⦃G, L⦄ ⊢ T •*[h, l] U → R l U.
+/3 width=5 by lstar_ind_r/ qed-.
+
+(* Basic inversion lemmas ***************************************************)
+
+lemma lstas_inv_O: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •*[h, 0] U → T = U.
+/2 width=4 by lstar_inv_O/ qed-.
+
+lemma lstas_inv_SO: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •*[h, 1] U → ⦃G, L⦄ ⊢ T •[h] U.
+/2 width=1 by lstar_inv_step/ qed-.
+
+lemma lstas_inv_step_sn: ∀h,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 •*[h, l+1] T2 →
+ ∃∃T. ⦃G, L⦄ ⊢ T1 •[h] T & ⦃G, L⦄ ⊢ T •*[h, l] T2.
+/2 width=3 by lstar_inv_S/ qed-.
+
+lemma lstas_inv_step_dx: ∀h,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 •*[h, l+1] T2 →
+ ∃∃T. ⦃G, L⦄ ⊢ T1 •*[h, l] T & ⦃G, L⦄ ⊢ T •[h] T2.
+/2 width=3 by lstar_inv_S_dx/ qed-.
+
+lemma lstas_inv_sort1: ∀h,G,L,X,k,l. ⦃G, L⦄ ⊢ ⋆k •*[h, l] X → X = ⋆((next h)^l k).
+#h #G #L #X #k #l #H @(lstas_ind_dx … H) -X -l //
+#l #X #X0 #_ #H #IHX destruct
+lapply (sta_inv_sort1 … H) -H #H destruct
+>iter_SO //
+qed-.
+
+lemma lstas_inv_gref1: ∀h,G,L,X,p,l. ⦃G, L⦄ ⊢ §p •*[h, l+1] X → ⊥.
+#h #G #L #X #p #l #H elim (lstas_inv_step_sn … H) -H
+#U #H #HUX elim (sta_inv_gref1 … H)
+qed-.
+
+lemma lstas_inv_bind1: ∀h,a,I,G,L,V,T,X,l. ⦃G, L⦄ ⊢ ⓑ{a,I}V.T •*[h, l] X →
+ ∃∃U. ⦃G, L.ⓑ{I}V⦄ ⊢ T •*[h, l] U & X = ⓑ{a,I}V.U.
+#h #a #I #G #L #V #T #X #l #H @(lstas_ind_dx … H) -X -l /2 width=3 by ex2_intro/
+#l #X #X0 #_ #HX0 * #U #HTU #H destruct
+elim (sta_inv_bind1 … HX0) -HX0 #U0 #HU0 #H destruct /3 width=3 by lstar_dx, ex2_intro/
+qed-.
+
+lemma lstas_inv_appl1: ∀h,G,L,V,T,X,l. ⦃G, L⦄ ⊢ ⓐV.T •*[h, l] X →
+ ∃∃U. ⦃G, L⦄ ⊢ T •*[h, l] U & X = ⓐV.U.
+#h #G #L #V #T #X #l #H @(lstas_ind_dx … H) -X -l /2 width=3 by ex2_intro/
+#l #X #X0 #_ #HX0 * #U #HTU #H destruct
+elim (sta_inv_appl1 … HX0) -HX0 #U0 #HU0 #H destruct /3 width=3 by lstar_dx, ex2_intro/
+qed-.
+
+lemma lstas_inv_cast1: ∀h,G,L,W,T,U,l. ⦃G, L⦄ ⊢ ⓝW.T •*[h, l+1] U → ⦃G, L⦄ ⊢ T •*[h, l+1] U.
+#h #G #L #W #T #X #l #H elim (lstas_inv_step_sn … H) -H
+#U #H #HUX lapply (sta_inv_cast1 … H) -H /2 width=3 by lstar_S/
+qed-.
+
+(* Basic properties *********************************************************)
+
+lemma lstas_refl: ∀h,G,L. reflexive … (lstas h G L 0).
+// qed.
+
+lemma sta_lstas: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ⦃G, L⦄ ⊢ T •*[h, 1] U.
+/2 width=1 by lstar_step/ qed.
+
+lemma lstas_step_sn: ∀h,G,L,T1,U1,U2,l. ⦃G, L⦄ ⊢ T1 •[h] U1 → ⦃G, L⦄ ⊢ U1 •*[h, l] U2 →
+ ⦃G, L⦄ ⊢ T1 •*[h, l+1] U2.
+/2 width=3 by lstar_S/ qed.
+
+lemma lstas_step_dx: ∀h,G,L,T1,T2,U2,l. ⦃G, L⦄ ⊢ T1 •*[h, l] T2 → ⦃G, L⦄ ⊢ T2 •[h] U2 →
+ ⦃G, L⦄ ⊢ T1 •*[h, l+1] U2.
+/2 width=3 by lstar_dx/ qed.
+
+lemma lstas_split: ∀h,G,L. inv_ltransitive … (lstas h G L).
+/2 width=1 by lstar_inv_ltransitive/ qed-.
+
+lemma lstas_sort: ∀h,G,L,l,k. ⦃G, L⦄ ⊢ ⋆k •*[h, l] ⋆((next h)^l k).
+#h #G #L #l @(nat_ind_plus … l) -l //
+#l #IHl #k >iter_SO /2 width=3 by sta_sort, lstas_step_dx/
+qed.
+
+lemma lstas_bind: ∀h,I,G,L,V,T,U,l. ⦃G, L.ⓑ{I}V⦄ ⊢ T •*[h, l] U →
+ ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V.T •*[h, l] ⓑ{a,I}V.U.
+#h #I #G #L #V #T #U #l #H @(lstas_ind_dx … H) -U -l /3 width=3 by sta_bind, lstar_O, lstas_step_dx/
+qed.
+
+lemma lstas_appl: ∀h,G,L,T,U,l. ⦃G, L⦄ ⊢ T •*[h, l] U →
+ ∀V.⦃G, L⦄ ⊢ ⓐV.T •*[h, l] ⓐV.U.
+#h #G #L #T #U #l #H @(lstas_ind_dx … H) -U -l /3 width=3 by sta_appl, lstar_O, lstas_step_dx/
+qed.
+
+lemma lstas_cast: ∀h,G,L,T,U,l. ⦃G, L⦄ ⊢ T •*[h, l+1] U →
+ ∀W. ⦃G, L⦄ ⊢ ⓝW.T •*[h, l+1] U.
+#h #G #L #T #U #l #H elim (lstas_inv_step_sn … H) -H /3 width=3 by sta_cast, lstas_step_sn/
+qed.
+
+(* Basic_1: removed theorems 7:
+ sty1_abbr sty1_appl sty1_bind sty1_cast2
+ sty1_correct sty1_lift sty1_trans
+*)
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/statictypestar_6.ma".
-include "basic_2/static/sta.ma".
-
-(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
-
-definition lstas: ∀h. genv → lenv → nat → relation term ≝
- λh,G,L. lstar … (sta h G L).
-
-interpretation "nat-iterated static type assignment (term)"
- 'StaticTypeStar h G L l T U = (lstas h G L l T U).
-
-(* Basic eliminators ********************************************************)
-
-lemma lstas_ind_sn: ∀h,G,L,U2. ∀R:relation2 nat term.
- R 0 U2 → (
- ∀l,T,U1. ⦃G, L⦄ ⊢ T •[h] U1 → ⦃G, L⦄ ⊢ U1 •* [h, l] U2 →
- R l U1 → R (l+1) T
- ) →
- ∀l,T. ⦃G, L⦄ ⊢ T •*[h, l] U2 → R l T.
-/3 width=5 by lstar_ind_l/ qed-.
-
-lemma lstas_ind_dx: ∀h,G,L,T. ∀R:relation2 nat term.
- R 0 T → (
- ∀l,U1,U2. ⦃G, L⦄ ⊢ T •* [h, l] U1 → ⦃G, L⦄ ⊢ U1 •[h] U2 →
- R l U1 → R (l+1) U2
- ) →
- ∀l,U. ⦃G, L⦄ ⊢ T •*[h, l] U → R l U.
-/3 width=5 by lstar_ind_r/ qed-.
-
-(* Basic inversion lemmas ***************************************************)
-
-lemma lstas_inv_O: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •*[h, 0] U → T = U.
-/2 width=4 by lstar_inv_O/ qed-.
-
-lemma lstas_inv_SO: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •*[h, 1] U → ⦃G, L⦄ ⊢ T •[h] U.
-/2 width=1 by lstar_inv_step/ qed-.
-
-lemma lstas_inv_step_sn: ∀h,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 •*[h, l+1] T2 →
- ∃∃T. ⦃G, L⦄ ⊢ T1 •[h] T & ⦃G, L⦄ ⊢ T •*[h, l] T2.
-/2 width=3 by lstar_inv_S/ qed-.
-
-lemma lstas_inv_step_dx: ∀h,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 •*[h, l+1] T2 →
- ∃∃T. ⦃G, L⦄ ⊢ T1 •*[h, l] T & ⦃G, L⦄ ⊢ T •[h] T2.
-/2 width=3 by lstar_inv_S_dx/ qed-.
-
-lemma lstas_inv_sort1: ∀h,G,L,X,k,l. ⦃G, L⦄ ⊢ ⋆k •*[h, l] X → X = ⋆((next h)^l k).
-#h #G #L #X #k #l #H @(lstas_ind_dx … H) -X -l //
-#l #X #X0 #_ #H #IHX destruct
-lapply (sta_inv_sort1 … H) -H #H destruct
->iter_SO //
-qed-.
-
-lemma lstas_inv_gref1: ∀h,G,L,X,p,l. ⦃G, L⦄ ⊢ §p •*[h, l+1] X → ⊥.
-#h #G #L #X #p #l #H elim (lstas_inv_step_sn … H) -H
-#U #H #HUX elim (sta_inv_gref1 … H)
-qed-.
-
-lemma lstas_inv_bind1: ∀h,a,I,G,L,V,T,X,l. ⦃G, L⦄ ⊢ ⓑ{a,I}V.T •*[h, l] X →
- ∃∃U. ⦃G, L.ⓑ{I}V⦄ ⊢ T •*[h, l] U & X = ⓑ{a,I}V.U.
-#h #a #I #G #L #V #T #X #l #H @(lstas_ind_dx … H) -X -l /2 width=3 by ex2_intro/
-#l #X #X0 #_ #HX0 * #U #HTU #H destruct
-elim (sta_inv_bind1 … HX0) -HX0 #U0 #HU0 #H destruct /3 width=3 by lstar_dx, ex2_intro/
-qed-.
-
-lemma lstas_inv_appl1: ∀h,G,L,V,T,X,l. ⦃G, L⦄ ⊢ ⓐV.T •*[h, l] X →
- ∃∃U. ⦃G, L⦄ ⊢ T •*[h, l] U & X = ⓐV.U.
-#h #G #L #V #T #X #l #H @(lstas_ind_dx … H) -X -l /2 width=3 by ex2_intro/
-#l #X #X0 #_ #HX0 * #U #HTU #H destruct
-elim (sta_inv_appl1 … HX0) -HX0 #U0 #HU0 #H destruct /3 width=3 by lstar_dx, ex2_intro/
-qed-.
-
-lemma lstas_inv_cast1: ∀h,G,L,W,T,U,l. ⦃G, L⦄ ⊢ ⓝW.T •*[h, l+1] U → ⦃G, L⦄ ⊢ T •*[h, l+1] U.
-#h #G #L #W #T #X #l #H elim (lstas_inv_step_sn … H) -H
-#U #H #HUX lapply (sta_inv_cast1 … H) -H /2 width=3 by lstar_S/
-qed-.
-
-(* Basic properties *********************************************************)
-
-lemma lstas_refl: ∀h,G,L. reflexive … (lstas h G L 0).
-// qed.
-
-lemma sta_lstas: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ⦃G, L⦄ ⊢ T •*[h, 1] U.
-/2 width=1 by lstar_step/ qed.
-
-lemma lstas_step_sn: ∀h,G,L,T1,U1,U2,l. ⦃G, L⦄ ⊢ T1 •[h] U1 → ⦃G, L⦄ ⊢ U1 •*[h, l] U2 →
- ⦃G, L⦄ ⊢ T1 •*[h, l+1] U2.
-/2 width=3 by lstar_S/ qed.
-
-lemma lstas_step_dx: ∀h,G,L,T1,T2,U2,l. ⦃G, L⦄ ⊢ T1 •*[h, l] T2 → ⦃G, L⦄ ⊢ T2 •[h] U2 →
- ⦃G, L⦄ ⊢ T1 •*[h, l+1] U2.
-/2 width=3 by lstar_dx/ qed.
-
-lemma lstas_split: ∀h,G,L. inv_ltransitive … (lstas h G L).
-/2 width=1 by lstar_inv_ltransitive/ qed-.
-
-lemma lstas_sort: ∀h,G,L,l,k. ⦃G, L⦄ ⊢ ⋆k •*[h, l] ⋆((next h)^l k).
-#h #G #L #l @(nat_ind_plus … l) -l //
-#l #IHl #k >iter_SO /2 width=3 by sta_sort, lstas_step_dx/
-qed.
-
-lemma lstas_bind: ∀h,I,G,L,V,T,U,l. ⦃G, L.ⓑ{I}V⦄ ⊢ T •*[h, l] U →
- ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V.T •*[h, l] ⓑ{a,I}V.U.
-#h #I #G #L #V #T #U #l #H @(lstas_ind_dx … H) -U -l /3 width=3 by sta_bind, lstar_O, lstas_step_dx/
-qed.
-
-lemma lstas_appl: ∀h,G,L,T,U,l. ⦃G, L⦄ ⊢ T •*[h, l] U →
- ∀V.⦃G, L⦄ ⊢ ⓐV.T •*[h, l] ⓐV.U.
-#h #G #L #T #U #l #H @(lstas_ind_dx … H) -U -l /3 width=3 by sta_appl, lstar_O, lstas_step_dx/
-qed.
-
-lemma lstas_cast: ∀h,G,L,T,U,l. ⦃G, L⦄ ⊢ T •*[h, l+1] U →
- ∀W. ⦃G, L⦄ ⊢ ⓝW.T •*[h, l+1] U.
-#h #G #L #T #U #l #H elim (lstas_inv_step_sn … H) -H /3 width=3 by sta_cast, lstas_step_sn/
-qed.
-
-(* Basic_1: removed theorems 7:
- sty1_abbr sty1_appl sty1_bind sty1_cast2
- sty1_correct sty1_lift sty1_trans
-*)
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/static/sta_aaa.ma".
+include "basic_2/unfold/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-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/static/sta_aaa.ma".
-include "basic_2/unfold/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-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/statictypestaralt_6.ma".
+include "basic_2/unfold/lstas_lift.ma".
+
+(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
+
+(* alternative definition of lstas *)
+inductive lstasa (h): genv → relation4 lenv nat term term ≝
+| lstasa_O : ∀G,L,T. lstasa h G L 0 T T
+| lstasa_sort: ∀G,L,l,k. lstasa h G L l (⋆k) (⋆((next h)^l k))
+| lstasa_ldef: ∀G,L,K,V,W,U,i,l. ⇩[i] L ≡ K.ⓓV → lstasa h G K (l+1) V W →
+ ⇧[0, i+1] W ≡ U → lstasa h G L (l+1) (#i) U
+| lstasa_ldec: ∀G,L,K,W,V,V0,U,i,l. ⇩[i] L ≡ K.ⓛW → ⦃G, K⦄ ⊢ W •[h] V0 →
+ lstasa h G K l W V → ⇧[0, i+1] V ≡ U → lstasa h G L (l+1) (#i) U
+| lstasa_bind: ∀a,I,G,L,V,T,U,l. lstasa h G (L.ⓑ{I}V) l T U →
+ lstasa h G L l (ⓑ{a,I}V.T) (ⓑ{a,I}V.U)
+| lstasa_appl: ∀G,L,V,T,U,l. lstasa h G L l T U → lstasa h G L l (ⓐV.T) (ⓐV.U)
+| lstasa_cast: ∀G,L,W,T,U,l. lstasa h G L (l+1) T U → lstasa h G L (l+1) (ⓝW.T) U
+.
+
+interpretation "nat-iterated static type assignment (term) alternative"
+ 'StaticTypeStarAlt h G L l T U = (lstasa h G L l T U).
+
+(* Base properties **********************************************************)
+
+lemma sta_lstasa: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ⦃G, L⦄ ⊢ T ••*[h, 1] U.
+#h #G #L #T #U #H elim H -G -L -T -U
+/2 width=8 by lstasa_O, lstasa_sort, lstasa_ldef, lstasa_ldec, lstasa_bind, lstasa_appl, lstasa_cast/
+qed.
+
+lemma lstasa_step_dx: ∀h,G,L,T1,T,l. ⦃G, L⦄ ⊢ T1 ••*[h, l] T →
+ ∀T2. ⦃G, L⦄ ⊢ T •[h] T2 → ⦃G, L⦄ ⊢ T1 ••*[h, l+1] T2.
+#h #G #L #T1 #T #l #H elim H -G -L -T1 -T -l
+[ /2 width=1 by sta_lstasa/
+| #G #L #l #k #X #H >(sta_inv_sort1 … H) -X >commutative_plus //
+| #G #L #K #V #W #U #i #l #HLK #_ #HWU #IHVW #U2 #HU2
+ lapply (drop_fwd_drop2 … HLK) #H
+ elim (sta_inv_lift1 … HU2 … H … HWU) -H -U /3 width=6 by lstasa_ldef/
+| #G #L #K #W #V #V0 #U #i #l #HLK #HWl0 #_ #HVU #IHWV #U2 #HU2
+ lapply (drop_fwd_drop2 … HLK) #H
+ elim (sta_inv_lift1 … HU2 … H … HVU) -H -U /3 width=8 by lstasa_ldec/
+| #a #I #G #L #V #T1 #U1 #l #_ #IHTU1 #X #H
+ elim (sta_inv_bind1 … H) -H #U #HU1 #H destruct /3 width=1 by lstasa_bind/
+| #G #L #V #T1 #U1 #l #_ #IHTU1 #X #H
+ elim (sta_inv_appl1 … H) -H #U #HU1 #H destruct /3 width=1 by lstasa_appl/
+| /3 width=1 by lstasa_cast/
+]
+qed.
+
+(* Main properties **********************************************************)
+
+theorem lstas_lstasa: ∀h,G,L,T,U,l. ⦃G, L⦄ ⊢ T •*[h, l] U → ⦃G, L⦄ ⊢ T ••*[h, l] U.
+#h #G #L #T #U #l #H @(lstas_ind_dx … H) -U -l /2 width=3 by lstasa_step_dx, lstasa_O/
+qed.
+
+(* Main inversion lemmas ****************************************************)
+
+theorem lstasa_inv_lstas: ∀h,G,L,T,U,l. ⦃G, L⦄ ⊢ T ••*[h, l] U → ⦃G, L⦄ ⊢ T •*[h, l] U.
+#h #G #L #T #U #l #H elim H -G -L -T -U -l
+/2 width=8 by lstas_inv_SO, lstas_ldec, lstas_ldef, lstas_cast, lstas_appl, lstas_bind/
+qed-.
+
+(* Advanced eliminators *****************************************************)
+
+lemma lstas_ind_alt: ∀h. ∀R:genv→relation4 lenv nat term term.
+ (∀G,L,T. R G L O T T) →
+ (∀G,L,l,k. R G L l (⋆k) (⋆((next h)^l k))) → (
+ ∀G,L,K,V,W,U,i,l.
+ ⇩[i] L ≡ K.ⓓV → ⦃G, K⦄ ⊢ V •*[h, l+1] W → ⇧[O, i+1] W ≡ U →
+ R G K (l+1) V W → R G L (l+1) (#i) U
+ ) → (
+ ∀G,L,K,W,V,V0,U,i,l.
+ ⇩[i] L ≡ K.ⓛW → ⦃G, K⦄ ⊢ W •[h] V0 →
+ ⦃G, K⦄ ⊢ W •*[h, l]V → ⇧[O, i+1] V ≡ U →
+ R G K l W V → R G L (l+1) (#i) U
+ ) → (
+ ∀a,I,G,L,V,T,U,l. ⦃G, L.ⓑ{I}V⦄ ⊢ T •*[h, l] U →
+ R G (L.ⓑ{I}V) l T U → R G L l (ⓑ{a,I}V.T) (ⓑ{a,I}V.U)
+ ) → (
+ ∀G,L,V,T,U,l. ⦃G, L⦄ ⊢ T •*[h, l] U →
+ R G L l T U → R G L l (ⓐV.T) (ⓐV.U)
+ ) → (
+ ∀G,L,W,T,U,l. ⦃G, L⦄⊢ T •*[h, l+1] U →
+ R G L (l+1) T U → R G L (l+1) (ⓝW.T) U
+ ) →
+ ∀G,L,l,T,U. ⦃G, L⦄ ⊢ T •*[h, l] U → R G L l T U.
+#h #R #IH1 #IH2 #IH3 #IH4 #IH5 #IH6 #IH7 #G #L #l #T #U #H
+elim (lstas_lstasa … H) /3 width=10 by lstasa_inv_lstas/
+qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/statictypestaralt_6.ma".
-include "basic_2/unfold/lstas_lift.ma".
-
-(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
-
-(* alternative definition of lstas *)
-inductive lstasa (h): genv → relation4 lenv nat term term ≝
-| lstasa_O : ∀G,L,T. lstasa h G L 0 T T
-| lstasa_sort: ∀G,L,l,k. lstasa h G L l (⋆k) (⋆((next h)^l k))
-| lstasa_ldef: ∀G,L,K,V,W,U,i,l. ⇩[i] L ≡ K.ⓓV → lstasa h G K (l+1) V W →
- ⇧[0, i+1] W ≡ U → lstasa h G L (l+1) (#i) U
-| lstasa_ldec: ∀G,L,K,W,V,V0,U,i,l. ⇩[i] L ≡ K.ⓛW → ⦃G, K⦄ ⊢ W •[h] V0 →
- lstasa h G K l W V → ⇧[0, i+1] V ≡ U → lstasa h G L (l+1) (#i) U
-| lstasa_bind: ∀a,I,G,L,V,T,U,l. lstasa h G (L.ⓑ{I}V) l T U →
- lstasa h G L l (ⓑ{a,I}V.T) (ⓑ{a,I}V.U)
-| lstasa_appl: ∀G,L,V,T,U,l. lstasa h G L l T U → lstasa h G L l (ⓐV.T) (ⓐV.U)
-| lstasa_cast: ∀G,L,W,T,U,l. lstasa h G L (l+1) T U → lstasa h G L (l+1) (ⓝW.T) U
-.
-
-interpretation "nat-iterated static type assignment (term) alternative"
- 'StaticTypeStarAlt h G L l T U = (lstasa h G L l T U).
-
-(* Base properties **********************************************************)
-
-lemma sta_lstasa: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ⦃G, L⦄ ⊢ T ••*[h, 1] U.
-#h #G #L #T #U #H elim H -G -L -T -U
-/2 width=8 by lstasa_O, lstasa_sort, lstasa_ldef, lstasa_ldec, lstasa_bind, lstasa_appl, lstasa_cast/
-qed.
-
-lemma lstasa_step_dx: ∀h,G,L,T1,T,l. ⦃G, L⦄ ⊢ T1 ••*[h, l] T →
- ∀T2. ⦃G, L⦄ ⊢ T •[h] T2 → ⦃G, L⦄ ⊢ T1 ••*[h, l+1] T2.
-#h #G #L #T1 #T #l #H elim H -G -L -T1 -T -l
-[ /2 width=1 by sta_lstasa/
-| #G #L #l #k #X #H >(sta_inv_sort1 … H) -X >commutative_plus //
-| #G #L #K #V #W #U #i #l #HLK #_ #HWU #IHVW #U2 #HU2
- lapply (drop_fwd_drop2 … HLK) #H
- elim (sta_inv_lift1 … HU2 … H … HWU) -H -U /3 width=6 by lstasa_ldef/
-| #G #L #K #W #V #V0 #U #i #l #HLK #HWl0 #_ #HVU #IHWV #U2 #HU2
- lapply (drop_fwd_drop2 … HLK) #H
- elim (sta_inv_lift1 … HU2 … H … HVU) -H -U /3 width=8 by lstasa_ldec/
-| #a #I #G #L #V #T1 #U1 #l #_ #IHTU1 #X #H
- elim (sta_inv_bind1 … H) -H #U #HU1 #H destruct /3 width=1 by lstasa_bind/
-| #G #L #V #T1 #U1 #l #_ #IHTU1 #X #H
- elim (sta_inv_appl1 … H) -H #U #HU1 #H destruct /3 width=1 by lstasa_appl/
-| /3 width=1 by lstasa_cast/
-]
-qed.
-
-(* Main properties **********************************************************)
-
-theorem lstas_lstasa: ∀h,G,L,T,U,l. ⦃G, L⦄ ⊢ T •*[h, l] U → ⦃G, L⦄ ⊢ T ••*[h, l] U.
-#h #G #L #T #U #l #H @(lstas_ind_dx … H) -U -l /2 width=3 by lstasa_step_dx, lstasa_O/
-qed.
-
-(* Main inversion lemmas ****************************************************)
-
-theorem lstasa_inv_lstas: ∀h,G,L,T,U,l. ⦃G, L⦄ ⊢ T ••*[h, l] U → ⦃G, L⦄ ⊢ T •*[h, l] U.
-#h #G #L #T #U #l #H elim H -G -L -T -U -l
-/2 width=8 by lstas_inv_SO, lstas_ldec, lstas_ldef, lstas_cast, lstas_appl, lstas_bind/
-qed-.
-
-(* Advanced eliminators *****************************************************)
-
-lemma lstas_ind_alt: ∀h. ∀R:genv→relation4 lenv nat term term.
- (∀G,L,T. R G L O T T) →
- (∀G,L,l,k. R G L l (⋆k) (⋆((next h)^l k))) → (
- ∀G,L,K,V,W,U,i,l.
- ⇩[i] L ≡ K.ⓓV → ⦃G, K⦄ ⊢ V •*[h, l+1] W → ⇧[O, i+1] W ≡ U →
- R G K (l+1) V W → R G L (l+1) (#i) U
- ) → (
- ∀G,L,K,W,V,V0,U,i,l.
- ⇩[i] L ≡ K.ⓛW → ⦃G, K⦄ ⊢ W •[h] V0 →
- ⦃G, K⦄ ⊢ W •*[h, l]V → ⇧[O, i+1] V ≡ U →
- R G K l W V → R G L (l+1) (#i) U
- ) → (
- ∀a,I,G,L,V,T,U,l. ⦃G, L.ⓑ{I}V⦄ ⊢ T •*[h, l] U →
- R G (L.ⓑ{I}V) l T U → R G L l (ⓑ{a,I}V.T) (ⓑ{a,I}V.U)
- ) → (
- ∀G,L,V,T,U,l. ⦃G, L⦄ ⊢ T •*[h, l] U →
- R G L l T U → R G L l (ⓐV.T) (ⓐV.U)
- ) → (
- ∀G,L,W,T,U,l. ⦃G, L⦄⊢ T •*[h, l+1] U →
- R G L (l+1) T U → R G L (l+1) (ⓝW.T) U
- ) →
- ∀G,L,l,T,U. ⦃G, L⦄ ⊢ T •*[h, l] U → R G L l T U.
-#h #R #IH1 #IH2 #IH3 #IH4 #IH5 #IH6 #IH7 #G #L #l #T #U #H
-elim (lstas_lstasa … H) /3 width=10 by lstasa_inv_lstas/
-qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/unfold/lstas.ma".
+include "basic_2/static/da_sta.ma".
+
+(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
+
+(* Properties on degree assignment for terms ********************************)
+
+lemma lstas_da_conf: ∀h,g,G,L,T,U,l1. ⦃G, L⦄ ⊢ T •*[h, l1] U →
+ ∀l2. ⦃G, L⦄ ⊢ T ▪[h, g] l2 → ⦃G, L⦄ ⊢ U ▪[h, g] l2-l1.
+#h #g #G #L #T #U #l1 #H @(lstas_ind_dx … H) -U -l1 //
+#l1 #U #U0 #_ #HU0 #IHTU #l2 #HT
+<minus_plus /3 width=3 by da_sta_conf/
+qed-.
+
+(* Advanced inversion lemmas ************************************************)
+
+lemma lstas_inv_refl_pos: ∀h,G,L,T,l. ⦃G, L⦄ ⊢ T •*[h, l+1] T → ⊥.
+#h #G #L #T #l #H elim (lstas_inv_step_sn … H)
+#U #HTU #_ elim (sta_da_ge … (l+1) HTU) -U
+#g #l0 #HT #Hl0 lapply (lstas_da_conf … H … HT) -H
+#H0T lapply (da_mono … HT … H0T) -h -G -L -T
+#H elim (discr_x_minus_xy … H) -H
+[ #H destruct /2 width=3 by le_plus_xSy_O_false/
+| -Hl0 <plus_n_Sm #H destruct
+]
+qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/unfold/lstas.ma".
-include "basic_2/static/da_sta.ma".
-
-(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
-
-(* Properties on degree assignment for terms ********************************)
-
-lemma lstas_da_conf: ∀h,g,G,L,T,U,l1. ⦃G, L⦄ ⊢ T •*[h, l1] U →
- ∀l2. ⦃G, L⦄ ⊢ T ▪[h, g] l2 → ⦃G, L⦄ ⊢ U ▪[h, g] l2-l1.
-#h #g #G #L #T #U #l1 #H @(lstas_ind_dx … H) -U -l1 //
-#l1 #U #U0 #_ #HU0 #IHTU #l2 #HT
-<minus_plus /3 width=3 by da_sta_conf/
-qed-.
-
-(* Advanced inversion lemmas ************************************************)
-
-lemma lstas_inv_refl_pos: ∀h,G,L,T,l. ⦃G, L⦄ ⊢ T •*[h, l+1] T → ⊥.
-#h #G #L #T #l #H elim (lstas_inv_step_sn … H)
-#U #HTU #_ elim (sta_da_ge … (l+1) HTU) -U
-#g #l0 #HT #Hl0 lapply (lstas_da_conf … H … HT) -H
-#H0T lapply (da_mono … HT … H0T) -h -G -L -T
-#H elim (discr_x_minus_xy … H) -H
-[ #H destruct /2 width=3 by le_plus_xSy_O_false/
-| -Hl0 <plus_n_Sm #H destruct
-]
-qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/static/sta_lift.ma".
+include "basic_2/unfold/lstas.ma".
+
+(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
+
+(* Properties on relocation *************************************************)
+
+lemma lstas_lift: ∀h,G,l. l_liftable (llstar … (sta h G) l).
+/3 width=10 by l_liftable_llstar, sta_lift/ qed.
+
+(* Inversion lemmas on relocation *******************************************)
+
+lemma lstas_inv_lift1: ∀h,G,l. l_deliftable_sn (llstar … (sta h G) l).
+/3 width=6 by l_deliftable_sn_llstar, sta_inv_lift1/ qed-.
+
+(* Advanced inversion lemmas ************************************************)
+
+lemma lstas_inv_lref1: ∀h,G,L,U,i,l. ⦃G, L⦄ ⊢ #i •*[h, l+1] U →
+ (∃∃K,V,W. ⇩[i] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •*[h, l+1] W &
+ ⇧[0, i+1] W ≡ U
+ ) ∨
+ (∃∃K,W,V,V0. ⇩[i] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W •[h] V0 &
+ ⦃G, K⦄ ⊢ W •*[h, l] V & ⇧[0, i+1] V ≡ U
+ ).
+#h #G #L #U #i #l #H elim (lstas_inv_step_sn … H) -H
+#X #H #HXU elim (sta_inv_lref1 … H) -H
+* #K #V #W #HLK #HVW #HWX
+lapply (drop_fwd_drop2 … HLK) #H0LK
+elim (lstas_inv_lift1 … HXU … H0LK … HWX) -H0LK -X
+/4 width=8 by lstas_step_sn, ex4_4_intro, ex3_3_intro, or_introl, or_intror/
+qed-.
+
+(* Advanced forward lemmas **************************************************)
+
+lemma lstas_fwd_correct: ∀h,G,L,T1,U1. ⦃G, L⦄ ⊢ T1 •[h] U1 →
+ ∀T2,l. ⦃G, L⦄ ⊢ T1 •*[h, l] T2 →
+ ∃U2. ⦃G, L⦄ ⊢ T2 •[h] U2.
+#h #G #L #T1 #U1 #HTU1 #T2 #l #H @(lstas_ind_dx … H) -l -T2 /2 width=3 by ex_intro/ -HTU1
+#l #T #T2 #_ #HT2 #_ -T1 -U1 -l
+elim (sta_fwd_correct … HT2) -T /2 width=2 by ex_intro/
+qed-.
+
+(* Advanced properties ******************************************************)
+
+lemma lstas_total: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U →
+ ∀l. ∃U0. ⦃G, L⦄ ⊢ T •*[h, l] U0.
+#h #G #L #T #U #HTU #l @(nat_ind_plus … l) -l /2 width=2 by ex_intro/
+#l * #U0 #HTU0 elim (lstas_fwd_correct … HTU … HTU0) -U
+/3 width=4 by lstas_step_dx, ex_intro/
+qed-.
+
+lemma lstas_ldef: ∀h,G,L,K,V,i. ⇩[i] L ≡ K.ⓓV →
+ ∀W,l. ⦃G, K⦄ ⊢ V •*[h, l+1] W →
+ ∀U. ⇧[0, i+1] W ≡ U → ⦃G, L⦄ ⊢ #i •*[h, l+1] U.
+#h #G #L #K #V #i #HLK #W #l #HVW #U #HWU
+lapply (drop_fwd_drop2 … HLK)
+elim (lstas_inv_step_sn … HVW) -HVW #W0
+elim (lift_total W0 0 (i+1)) /3 width=12 by lstas_step_sn, sta_ldef, lstas_lift/
+qed.
+
+lemma lstas_ldec: ∀h,G,L,K,W,i. ⇩[i] L ≡ K.ⓛW → ∀V0. ⦃G, K⦄ ⊢ W •[h] V0 →
+ ∀V,l. ⦃G, K⦄ ⊢ W •*[h, l] V →
+ ∀U. ⇧[0, i+1] V ≡ U → ⦃G, L⦄ ⊢ #i •*[h, l+1] U.
+#h #G #L #K #W #i #HLK #V0 #HWV0 #V #l #HWV #U #HVU
+lapply (drop_fwd_drop2 … HLK) #H
+elim (lift_total W 0 (i+1)) /3 width=12 by lstas_step_sn, sta_ldec, lstas_lift/
+qed.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/static/sta_lift.ma".
-include "basic_2/unfold/lstas.ma".
-
-(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
-
-(* Properties on relocation *************************************************)
-
-lemma lstas_lift: ∀h,G,l. l_liftable (llstar … (sta h G) l).
-/3 width=10 by l_liftable_llstar, sta_lift/ qed.
-
-(* Inversion lemmas on relocation *******************************************)
-
-lemma lstas_inv_lift1: ∀h,G,l. l_deliftable_sn (llstar … (sta h G) l).
-/3 width=6 by l_deliftable_sn_llstar, sta_inv_lift1/ qed-.
-
-(* Advanced inversion lemmas ************************************************)
-
-lemma lstas_inv_lref1: ∀h,G,L,U,i,l. ⦃G, L⦄ ⊢ #i •*[h, l+1] U →
- (∃∃K,V,W. ⇩[i] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •*[h, l+1] W &
- ⇧[0, i+1] W ≡ U
- ) ∨
- (∃∃K,W,V,V0. ⇩[i] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W •[h] V0 &
- ⦃G, K⦄ ⊢ W •*[h, l] V & ⇧[0, i+1] V ≡ U
- ).
-#h #G #L #U #i #l #H elim (lstas_inv_step_sn … H) -H
-#X #H #HXU elim (sta_inv_lref1 … H) -H
-* #K #V #W #HLK #HVW #HWX
-lapply (drop_fwd_drop2 … HLK) #H0LK
-elim (lstas_inv_lift1 … HXU … H0LK … HWX) -H0LK -X
-/4 width=8 by lstas_step_sn, ex4_4_intro, ex3_3_intro, or_introl, or_intror/
-qed-.
-
-(* Advanced forward lemmas **************************************************)
-
-lemma lstas_fwd_correct: ∀h,G,L,T1,U1. ⦃G, L⦄ ⊢ T1 •[h] U1 →
- ∀T2,l. ⦃G, L⦄ ⊢ T1 •*[h, l] T2 →
- ∃U2. ⦃G, L⦄ ⊢ T2 •[h] U2.
-#h #G #L #T1 #U1 #HTU1 #T2 #l #H @(lstas_ind_dx … H) -l -T2 /2 width=3 by ex_intro/ -HTU1
-#l #T #T2 #_ #HT2 #_ -T1 -U1 -l
-elim (sta_fwd_correct … HT2) -T /2 width=2 by ex_intro/
-qed-.
-
-(* Advanced properties ******************************************************)
-
-lemma lstas_total: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U →
- ∀l. ∃U0. ⦃G, L⦄ ⊢ T •*[h, l] U0.
-#h #G #L #T #U #HTU #l @(nat_ind_plus … l) -l /2 width=2 by ex_intro/
-#l * #U0 #HTU0 elim (lstas_fwd_correct … HTU … HTU0) -U
-/3 width=4 by lstas_step_dx, ex_intro/
-qed-.
-
-lemma lstas_ldef: ∀h,G,L,K,V,i. ⇩[i] L ≡ K.ⓓV →
- ∀W,l. ⦃G, K⦄ ⊢ V •*[h, l+1] W →
- ∀U. ⇧[0, i+1] W ≡ U → ⦃G, L⦄ ⊢ #i •*[h, l+1] U.
-#h #G #L #K #V #i #HLK #W #l #HVW #U #HWU
-lapply (drop_fwd_drop2 … HLK)
-elim (lstas_inv_step_sn … HVW) -HVW #W0
-elim (lift_total W0 0 (i+1)) /3 width=12 by lstas_step_sn, sta_ldef, lstas_lift/
-qed.
-
-lemma lstas_ldec: ∀h,G,L,K,W,i. ⇩[i] L ≡ K.ⓛW → ∀V0. ⦃G, K⦄ ⊢ W •[h] V0 →
- ∀V,l. ⦃G, K⦄ ⊢ W •*[h, l] V →
- ∀U. ⇧[0, i+1] V ≡ U → ⦃G, L⦄ ⊢ #i •*[h, l+1] U.
-#h #G #L #K #W #i #HLK #V0 #HWV0 #V #l #HWV #U #HVU
-lapply (drop_fwd_drop2 … HLK) #H
-elim (lift_total W 0 (i+1)) /3 width=12 by lstas_step_sn, sta_ldec, lstas_lift/
-qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/static/sta_sta.ma".
+include "basic_2/unfold/lstas_lift.ma".
+
+(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
+
+(* Main properties **********************************************************)
+
+theorem lstas_trans: ∀h,G,L. ltransitive … (lstas h G L).
+/2 width=3 by lstar_ltransitive/ qed-.
+
+theorem lstas_mono: ∀h,G,L,l. singlevalued … (lstas h G L l).
+/3 width=7 by sta_mono, lstar_singlevalued/ qed-.
+
+theorem lstas_conf_le: ∀h,G,L,T,U1,l1. ⦃G, L⦄ ⊢ T •*[h, l1] U1 →
+ ∀U2,l2. l1 ≤ l2 → ⦃G, L⦄ ⊢ T •*[h, l2] U2 →
+ ⦃G, L⦄ ⊢ U1 •*[h, l2-l1] U2.
+#h #G #L #T #U1 #l1 #HTU1 #U2 #l2 #Hl12
+>(plus_minus_m_m … Hl12) in ⊢ (%→?); -Hl12 >commutative_plus #H
+elim (lstas_split … H) -H #U #HTU
+>(lstas_mono … HTU … HTU1) -T //
+qed-.
+
+(* Advanced properties ******************************************************)
+
+lemma lstas_sta_conf_pos: ∀h,G,L,T,U1. ⦃G, L⦄ ⊢ T •[h] U1 →
+ ∀U2,l. ⦃G, L⦄ ⊢ T •*[h, l+1] U2 → ⦃G, L⦄ ⊢ U1 •*[h, l] U2.
+#h #G #L #T #U1 #HTU1 #U2 #l #HTU2
+lapply (lstas_conf_le … T U1 1 … HTU2) -HTU2 /2 width=1 by sta_lstas/
+qed-.
+
+lemma lstas_strip_pos: ∀h,G,L,T1,U1. ⦃G, L⦄ ⊢ T1 •[h] U1 →
+ ∀T2,l. ⦃G, L⦄ ⊢ T1 •*[h, l+1] T2 →
+ ∃∃U2. ⦃G, L⦄ ⊢ T2 •[h] U2 & ⦃G, L⦄ ⊢ U1 •*[h, l+1] U2.
+#h #G #L #T1 #U1 #HTU1 #T2 #l #HT12
+elim (lstas_fwd_correct … HTU1 … HT12)
+lapply (lstas_sta_conf_pos … HTU1 … HT12) -T1 /3 width=5 by lstas_step_dx, ex2_intro/
+qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/static/sta_sta.ma".
-include "basic_2/unfold/lstas_lift.ma".
-
-(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
-
-(* Main properties **********************************************************)
-
-theorem lstas_trans: ∀h,G,L. ltransitive … (lstas h G L).
-/2 width=3 by lstar_ltransitive/ qed-.
-
-theorem lstas_mono: ∀h,G,L,l. singlevalued … (lstas h G L l).
-/3 width=7 by sta_mono, lstar_singlevalued/ qed-.
-
-theorem lstas_conf_le: ∀h,G,L,T,U1,l1. ⦃G, L⦄ ⊢ T •*[h, l1] U1 →
- ∀U2,l2. l1 ≤ l2 → ⦃G, L⦄ ⊢ T •*[h, l2] U2 →
- ⦃G, L⦄ ⊢ U1 •*[h, l2-l1] U2.
-#h #G #L #T #U1 #l1 #HTU1 #U2 #l2 #Hl12
->(plus_minus_m_m … Hl12) in ⊢ (%→?); -Hl12 >commutative_plus #H
-elim (lstas_split … H) -H #U #HTU
->(lstas_mono … HTU … HTU1) -T //
-qed-.
-
-(* Advanced properties ******************************************************)
-
-lemma lstas_sta_conf_pos: ∀h,G,L,T,U1. ⦃G, L⦄ ⊢ T •[h] U1 →
- ∀U2,l. ⦃G, L⦄ ⊢ T •*[h, l+1] U2 → ⦃G, L⦄ ⊢ U1 •*[h, l] U2.
-#h #G #L #T #U1 #HTU1 #U2 #l #HTU2
-lapply (lstas_conf_le … T U1 1 … HTU2) -HTU2 /2 width=1 by sta_lstas/
-qed-.
-
-lemma lstas_strip_pos: ∀h,G,L,T1,U1. ⦃G, L⦄ ⊢ T1 •[h] U1 →
- ∀T2,l. ⦃G, L⦄ ⊢ T1 •*[h, l+1] T2 →
- ∃∃U2. ⦃G, L⦄ ⊢ T2 •[h] U2 & ⦃G, L⦄ ⊢ U1 •*[h, l+1] U2.
-#h #G #L #T1 #U1 #HTU1 #T2 #l #HT12
-elim (lstas_fwd_correct … HTU1 … HT12)
-lapply (lstas_sta_conf_pos … HTU1 … HT12) -T1 /3 width=5 by lstas_step_dx, ex2_intro/
-qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/statictype_5.ma".
+include "basic_2/grammar/genv.ma".
+include "basic_2/substitution/drop.ma".
+include "basic_2/static/sh.ma".
+
+(* STATIC TYPE ASSIGNMENT ON TERMS ******************************************)
+
+(* activate genv *)
+inductive sta (h:sh): relation4 genv lenv term term ≝
+| sta_sort: ∀G,L,k. sta h G L (⋆k) (⋆(next h k))
+| sta_ldef: ∀G,L,K,V,W,U,i. ⇩[i] L ≡ K.ⓓV → sta h G K V W →
+ ⇧[0, i + 1] W ≡ U → sta h G L (#i) U
+| sta_ldec: ∀G,L,K,W,V,U,i. ⇩[i] L ≡ K.ⓛW → sta h G K W V →
+ ⇧[0, i + 1] W ≡ U → sta h G L (#i) U
+| sta_bind: ∀a,I,G,L,V,T,U. sta h G (L.ⓑ{I}V) T U →
+ sta h G L (ⓑ{a,I}V.T) (ⓑ{a,I}V.U)
+| sta_appl: ∀G,L,V,T,U. sta h G L T U → sta h G L (ⓐV.T) (ⓐV.U)
+| sta_cast: ∀G,L,W,T,U. sta h G L T U → sta h G L (ⓝW.T) U
+.
+
+interpretation "static type assignment (term)"
+ 'StaticType h G L T U = (sta h G L T U).
+
+(* Basic inversion lemmas ************************************************)
+
+fact sta_inv_sort1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀k0. T = ⋆k0 →
+ U = ⋆(next h k0).
+#h #G #L #T #U * -G -L -T -U
+[ #G #L #k #k0 #H destruct //
+| #G #L #K #V #W #U #i #_ #_ #_ #k0 #H destruct
+| #G #L #K #W #V #U #i #_ #_ #_ #k0 #H destruct
+| #a #I #G #L #V #T #U #_ #k0 #H destruct
+| #G #L #V #T #U #_ #k0 #H destruct
+| #G #L #W #T #U #_ #k0 #H destruct
+qed-.
+
+(* Basic_1: was: sty0_gen_sort *)
+lemma sta_inv_sort1: ∀h,G,L,U,k. ⦃G, L⦄ ⊢ ⋆k •[h] U → U = ⋆(next h k).
+/2 width=5 by sta_inv_sort1_aux/ qed-.
+
+fact sta_inv_lref1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀j. T = #j →
+ (∃∃K,V,W. ⇩[j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •[h] W &
+ ⇧[0, j+1] W ≡ U
+ ) ∨
+ (∃∃K,W,V. ⇩[j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W •[h] V &
+ ⇧[0, j+1] W ≡ U
+ ).
+#h #G #L #T #U * -G -L -T -U
+[ #G #L #k #j #H destruct
+| #G #L #K #V #W #U #i #HLK #HVW #HWU #j #H destruct /3 width=6 by or_introl, ex3_3_intro/
+| #G #L #K #W #V #U #i #HLK #HWV #HWU #j #H destruct /3 width=6 by or_intror, ex3_3_intro/
+| #a #I #G #L #V #T #U #_ #j #H destruct
+| #G #L #V #T #U #_ #j #H destruct
+| #G #L #W #T #U #_ #j #H destruct
+]
+qed-.
+
+(* Basic_1: was sty0_gen_lref *)
+lemma sta_inv_lref1: ∀h,G,L,U,i. ⦃G, L⦄ ⊢ #i •[h] U →
+ (∃∃K,V,W. ⇩[i] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •[h] W &
+ ⇧[0, i+1] W ≡ U
+ ) ∨
+ (∃∃K,W,V. ⇩[i] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W •[h] V &
+ ⇧[0, i+1] W ≡ U
+ ).
+/2 width=3 by sta_inv_lref1_aux/ qed-.
+
+fact sta_inv_gref1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀p0. T = §p0 → ⊥.
+#h #G #L #T #U * -G -L -T -U
+[ #G #L #k #p0 #H destruct
+| #G #L #K #V #W #U #i #_ #_ #_ #p0 #H destruct
+| #G #L #K #W #V #U #i #_ #_ #_ #p0 #H destruct
+| #a #I #G #L #V #T #U #_ #p0 #H destruct
+| #G #L #V #T #U #_ #p0 #H destruct
+| #G #L #W #T #U #_ #p0 #H destruct
+qed-.
+
+lemma sta_inv_gref1: ∀h,G,L,U,p. ⦃G, L⦄ ⊢ §p •[h] U → ⊥.
+/2 width=8 by sta_inv_gref1_aux/ qed-.
+
+fact sta_inv_bind1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀b,J,X,Y. T = ⓑ{b,J}Y.X →
+ ∃∃Z. ⦃G, L.ⓑ{J}Y⦄ ⊢ X •[h] Z & U = ⓑ{b,J}Y.Z.
+#h #G #L #T #U * -G -L -T -U
+[ #G #L #k #b #J #X #Y #H destruct
+| #G #L #K #V #W #U #i #_ #_ #_ #b #J #X #Y #H destruct
+| #G #L #K #W #V #U #i #_ #_ #_ #b #J #X #Y #H destruct
+| #a #I #G #L #V #T #U #HTU #b #J #X #Y #H destruct /2 width=3 by ex2_intro/
+| #G #L #V #T #U #_ #b #J #X #Y #H destruct
+| #G #L #W #T #U #_ #b #J #X #Y #H destruct
+]
+qed-.
+
+(* Basic_1: was: sty0_gen_bind *)
+lemma sta_inv_bind1: ∀h,b,J,G,L,Y,X,U. ⦃G, L⦄ ⊢ ⓑ{b,J}Y.X •[h] U →
+ ∃∃Z. ⦃G, L.ⓑ{J}Y⦄ ⊢ X •[h] Z & U = ⓑ{b,J}Y.Z.
+/2 width=3 by sta_inv_bind1_aux/ qed-.
+
+fact sta_inv_appl1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀X,Y. T = ⓐY.X →
+ ∃∃Z. ⦃G, L⦄ ⊢ X •[h] Z & U = ⓐY.Z.
+#h #G #L #T #U * -G -L -T -U
+[ #G #L #k #X #Y #H destruct
+| #G #L #K #V #W #U #i #_ #_ #_ #X #Y #H destruct
+| #G #L #K #W #V #U #i #_ #_ #_ #X #Y #H destruct
+| #a #I #G #L #V #T #U #_ #X #Y #H destruct
+| #G #L #V #T #U #HTU #X #Y #H destruct /2 width=3 by ex2_intro/
+| #G #L #W #T #U #_ #X #Y #H destruct
+]
+qed-.
+
+(* Basic_1: was: sty0_gen_appl *)
+lemma sta_inv_appl1: ∀h,G,L,Y,X,U. ⦃G, L⦄ ⊢ ⓐY.X •[h] U →
+ ∃∃Z. ⦃G, L⦄ ⊢ X •[h] Z & U = ⓐY.Z.
+/2 width=3 by sta_inv_appl1_aux/ qed-.
+
+fact sta_inv_cast1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀X,Y. T = ⓝY.X →
+ ⦃G, L⦄ ⊢ X •[h] U.
+#h #G #L #T #U * -G -L -T -U
+[ #G #L #k #X #Y #H destruct
+| #G #L #K #V #W #U #i #_ #_ #_ #X #Y #H destruct
+| #G #L #K #W #V #U #i #_ #_ #_ #X #Y #H destruct
+| #a #I #G #L #V #T #U #_ #X #Y #H destruct
+| #G #L #V #T #U #_ #X #Y #H destruct
+| #G #L #W #T #U #HTU #X #Y #H destruct //
+]
+qed-.
+
+(* Basic_1: was: sty0_gen_cast *)
+lemma sta_inv_cast1: ∀h,G,L,X,Y,U. ⦃G, L⦄ ⊢ ⓝY.X •[h] U → ⦃G, L⦄ ⊢ X •[h] U.
+/2 width=4 by sta_inv_cast1_aux/ qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/statictype_5.ma".
-include "basic_2/grammar/genv.ma".
-include "basic_2/substitution/drop.ma".
-include "basic_2/static/sh.ma".
-
-(* STATIC TYPE ASSIGNMENT ON TERMS ******************************************)
-
-(* activate genv *)
-inductive sta (h:sh): relation4 genv lenv term term ≝
-| sta_sort: ∀G,L,k. sta h G L (⋆k) (⋆(next h k))
-| sta_ldef: ∀G,L,K,V,W,U,i. ⇩[i] L ≡ K.ⓓV → sta h G K V W →
- ⇧[0, i + 1] W ≡ U → sta h G L (#i) U
-| sta_ldec: ∀G,L,K,W,V,U,i. ⇩[i] L ≡ K.ⓛW → sta h G K W V →
- ⇧[0, i + 1] W ≡ U → sta h G L (#i) U
-| sta_bind: ∀a,I,G,L,V,T,U. sta h G (L.ⓑ{I}V) T U →
- sta h G L (ⓑ{a,I}V.T) (ⓑ{a,I}V.U)
-| sta_appl: ∀G,L,V,T,U. sta h G L T U → sta h G L (ⓐV.T) (ⓐV.U)
-| sta_cast: ∀G,L,W,T,U. sta h G L T U → sta h G L (ⓝW.T) U
-.
-
-interpretation "static type assignment (term)"
- 'StaticType h G L T U = (sta h G L T U).
-
-(* Basic inversion lemmas ************************************************)
-
-fact sta_inv_sort1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀k0. T = ⋆k0 →
- U = ⋆(next h k0).
-#h #G #L #T #U * -G -L -T -U
-[ #G #L #k #k0 #H destruct //
-| #G #L #K #V #W #U #i #_ #_ #_ #k0 #H destruct
-| #G #L #K #W #V #U #i #_ #_ #_ #k0 #H destruct
-| #a #I #G #L #V #T #U #_ #k0 #H destruct
-| #G #L #V #T #U #_ #k0 #H destruct
-| #G #L #W #T #U #_ #k0 #H destruct
-qed-.
-
-(* Basic_1: was: sty0_gen_sort *)
-lemma sta_inv_sort1: ∀h,G,L,U,k. ⦃G, L⦄ ⊢ ⋆k •[h] U → U = ⋆(next h k).
-/2 width=5 by sta_inv_sort1_aux/ qed-.
-
-fact sta_inv_lref1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀j. T = #j →
- (∃∃K,V,W. ⇩[j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •[h] W &
- ⇧[0, j+1] W ≡ U
- ) ∨
- (∃∃K,W,V. ⇩[j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W •[h] V &
- ⇧[0, j+1] W ≡ U
- ).
-#h #G #L #T #U * -G -L -T -U
-[ #G #L #k #j #H destruct
-| #G #L #K #V #W #U #i #HLK #HVW #HWU #j #H destruct /3 width=6 by or_introl, ex3_3_intro/
-| #G #L #K #W #V #U #i #HLK #HWV #HWU #j #H destruct /3 width=6 by or_intror, ex3_3_intro/
-| #a #I #G #L #V #T #U #_ #j #H destruct
-| #G #L #V #T #U #_ #j #H destruct
-| #G #L #W #T #U #_ #j #H destruct
-]
-qed-.
-
-(* Basic_1: was sty0_gen_lref *)
-lemma sta_inv_lref1: ∀h,G,L,U,i. ⦃G, L⦄ ⊢ #i •[h] U →
- (∃∃K,V,W. ⇩[i] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •[h] W &
- ⇧[0, i+1] W ≡ U
- ) ∨
- (∃∃K,W,V. ⇩[i] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W •[h] V &
- ⇧[0, i+1] W ≡ U
- ).
-/2 width=3 by sta_inv_lref1_aux/ qed-.
-
-fact sta_inv_gref1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀p0. T = §p0 → ⊥.
-#h #G #L #T #U * -G -L -T -U
-[ #G #L #k #p0 #H destruct
-| #G #L #K #V #W #U #i #_ #_ #_ #p0 #H destruct
-| #G #L #K #W #V #U #i #_ #_ #_ #p0 #H destruct
-| #a #I #G #L #V #T #U #_ #p0 #H destruct
-| #G #L #V #T #U #_ #p0 #H destruct
-| #G #L #W #T #U #_ #p0 #H destruct
-qed-.
-
-lemma sta_inv_gref1: ∀h,G,L,U,p. ⦃G, L⦄ ⊢ §p •[h] U → ⊥.
-/2 width=8 by sta_inv_gref1_aux/ qed-.
-
-fact sta_inv_bind1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀b,J,X,Y. T = ⓑ{b,J}Y.X →
- ∃∃Z. ⦃G, L.ⓑ{J}Y⦄ ⊢ X •[h] Z & U = ⓑ{b,J}Y.Z.
-#h #G #L #T #U * -G -L -T -U
-[ #G #L #k #b #J #X #Y #H destruct
-| #G #L #K #V #W #U #i #_ #_ #_ #b #J #X #Y #H destruct
-| #G #L #K #W #V #U #i #_ #_ #_ #b #J #X #Y #H destruct
-| #a #I #G #L #V #T #U #HTU #b #J #X #Y #H destruct /2 width=3 by ex2_intro/
-| #G #L #V #T #U #_ #b #J #X #Y #H destruct
-| #G #L #W #T #U #_ #b #J #X #Y #H destruct
-]
-qed-.
-
-(* Basic_1: was: sty0_gen_bind *)
-lemma sta_inv_bind1: ∀h,b,J,G,L,Y,X,U. ⦃G, L⦄ ⊢ ⓑ{b,J}Y.X •[h] U →
- ∃∃Z. ⦃G, L.ⓑ{J}Y⦄ ⊢ X •[h] Z & U = ⓑ{b,J}Y.Z.
-/2 width=3 by sta_inv_bind1_aux/ qed-.
-
-fact sta_inv_appl1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀X,Y. T = ⓐY.X →
- ∃∃Z. ⦃G, L⦄ ⊢ X •[h] Z & U = ⓐY.Z.
-#h #G #L #T #U * -G -L -T -U
-[ #G #L #k #X #Y #H destruct
-| #G #L #K #V #W #U #i #_ #_ #_ #X #Y #H destruct
-| #G #L #K #W #V #U #i #_ #_ #_ #X #Y #H destruct
-| #a #I #G #L #V #T #U #_ #X #Y #H destruct
-| #G #L #V #T #U #HTU #X #Y #H destruct /2 width=3 by ex2_intro/
-| #G #L #W #T #U #_ #X #Y #H destruct
-]
-qed-.
-
-(* Basic_1: was: sty0_gen_appl *)
-lemma sta_inv_appl1: ∀h,G,L,Y,X,U. ⦃G, L⦄ ⊢ ⓐY.X •[h] U →
- ∃∃Z. ⦃G, L⦄ ⊢ X •[h] Z & U = ⓐY.Z.
-/2 width=3 by sta_inv_appl1_aux/ qed-.
-
-fact sta_inv_cast1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀X,Y. T = ⓝY.X →
- ⦃G, L⦄ ⊢ X •[h] U.
-#h #G #L #T #U * -G -L -T -U
-[ #G #L #k #X #Y #H destruct
-| #G #L #K #V #W #U #i #_ #_ #_ #X #Y #H destruct
-| #G #L #K #W #V #U #i #_ #_ #_ #X #Y #H destruct
-| #a #I #G #L #V #T #U #_ #X #Y #H destruct
-| #G #L #V #T #U #_ #X #Y #H destruct
-| #G #L #W #T #U #HTU #X #Y #H destruct //
-]
-qed-.
-
-(* Basic_1: was: sty0_gen_cast *)
-lemma sta_inv_cast1: ∀h,G,L,X,Y,U. ⦃G, L⦄ ⊢ ⓝY.X •[h] U → ⦃G, L⦄ ⊢ X •[h] U.
-/2 width=4 by sta_inv_cast1_aux/ qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/static/sta.ma".
+include "basic_2/static/aaa_lift.ma".
+
+(* STATIC TYPE ASSIGNMENT FOR TERMS *****************************************)
+
+(* Properties on atomic arity assignment for terms **************************)
+
+lemma aaa_sta: ∀h,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∃U. ⦃G, L⦄ ⊢ T •[h] U.
+#h #G #L #T #A #H elim H -G -L -T -A
+[ /2 width=2 by sta_sort, ex_intro/
+| * #G #L #K [ #V | #W ] #B #i #HLK #_ * [ #W | #V ] #HVW
+ elim (lift_total W 0 (i+1)) /3 width=7 by sta_ldef, sta_ldec, ex_intro/
+| #a #G #L #V #T #B #A #_ #_ #_ * /3 width=2 by sta_bind, ex_intro/
+| #a #G #L #V #T #B #A #_ #_ #_ * /3 width=2 by sta_bind, ex_intro/
+| #G #L #V #T #B #A #_ #_ #_ * /3 width=2 by sta_appl, ex_intro/
+| #G #L #W #T #A #_ #_ #_ * /3 width=2 by sta_cast, ex_intro/
+]
+qed-.
+
+lemma sta_aaa_conf: ∀h,G,L. Conf3 … (aaa G L) (sta h G L).
+#h #G #L #T #A #H elim H -G -L -T -A
+[ #G #L #k #U #H
+ lapply (sta_inv_sort1 … H) -H #H destruct //
+| #I #G #L #K #V #B #i #HLK #HV #IHV #U #H
+ elim (sta_inv_lref1 … H) -H * #K0 #V0 #W0 #HLK0 #HVW0 #HU
+ lapply (drop_mono … HLK0 … HLK) -HLK0 #H0 destruct
+ lapply (drop_fwd_drop2 … HLK) -HLK #HLK
+ @(aaa_lift … HLK … HU) -HU -L /2 width=2 by/
+| #a #G #L #V #T #B #A #HV #_ #_ #IHT #X #H
+ elim (sta_inv_bind1 … H) -H #U #HTU #H destruct /3 width=2 by aaa_abbr/
+| #a #G #L #V #T #B #A #HV #_ #_ #IHT #X #H
+ elim (sta_inv_bind1 … H) -H #U #HTU #H destruct /3 width=2 by aaa_abst/
+| #G #L #V #T #B #A #HV #_ #_ #IHT #X #H
+ elim (sta_inv_appl1 … H) -H #U #HTU #H destruct /3 width=3 by aaa_appl/
+| #G #L #V #T #A #_ #_ #IHV #IHT #X #H
+ lapply (sta_inv_cast1 … H) -H /2 width=2 by/
+]
+qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/static/sta.ma".
-include "basic_2/static/aaa_lift.ma".
-
-(* STATIC TYPE ASSIGNMENT FOR TERMS *****************************************)
-
-(* Properties on atomic arity assignment for terms **************************)
-
-lemma aaa_sta: ∀h,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∃U. ⦃G, L⦄ ⊢ T •[h] U.
-#h #G #L #T #A #H elim H -G -L -T -A
-[ /2 width=2 by sta_sort, ex_intro/
-| * #G #L #K [ #V | #W ] #B #i #HLK #_ * [ #W | #V ] #HVW
- elim (lift_total W 0 (i+1)) /3 width=7 by sta_ldef, sta_ldec, ex_intro/
-| #a #G #L #V #T #B #A #_ #_ #_ * /3 width=2 by sta_bind, ex_intro/
-| #a #G #L #V #T #B #A #_ #_ #_ * /3 width=2 by sta_bind, ex_intro/
-| #G #L #V #T #B #A #_ #_ #_ * /3 width=2 by sta_appl, ex_intro/
-| #G #L #W #T #A #_ #_ #_ * /3 width=2 by sta_cast, ex_intro/
-]
-qed-.
-
-lemma sta_aaa_conf: ∀h,G,L. Conf3 … (aaa G L) (sta h G L).
-#h #G #L #T #A #H elim H -G -L -T -A
-[ #G #L #k #U #H
- lapply (sta_inv_sort1 … H) -H #H destruct //
-| #I #G #L #K #V #B #i #HLK #HV #IHV #U #H
- elim (sta_inv_lref1 … H) -H * #K0 #V0 #W0 #HLK0 #HVW0 #HU
- lapply (drop_mono … HLK0 … HLK) -HLK0 #H0 destruct
- lapply (drop_fwd_drop2 … HLK) -HLK #HLK
- @(aaa_lift … HLK … HU) -HU -L /2 width=2 by/
-| #a #G #L #V #T #B #A #HV #_ #_ #IHT #X #H
- elim (sta_inv_bind1 … H) -H #U #HTU #H destruct /3 width=2 by aaa_abbr/
-| #a #G #L #V #T #B #A #HV #_ #_ #IHT #X #H
- elim (sta_inv_bind1 … H) -H #U #HTU #H destruct /3 width=2 by aaa_abst/
-| #G #L #V #T #B #A #HV #_ #_ #IHT #X #H
- elim (sta_inv_appl1 … H) -H #U #HTU #H destruct /3 width=3 by aaa_appl/
-| #G #L #V #T #A #_ #_ #IHV #IHT #X #H
- lapply (sta_inv_cast1 … H) -H /2 width=2 by/
-]
-qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/substitution/drop_drop.ma".
+include "basic_2/static/sta.ma".
+
+(* STATIC TYPE ASSIGNMENT ON TERMS ******************************************)
+
+(* Properties on relocation *************************************************)
+
+(* Basic_1: was: sty0_lift *)
+lemma sta_lift: ∀h,G. l_liftable (sta h G).
+#h #G #L1 #T1 #U1 #H elim H -G -L1 -T1 -U1
+[ #G #L1 #k #L2 #s #d #e #HL21 #X1 #H1 #X2 #H2
+ >(lift_inv_sort1 … H1) -X1
+ >(lift_inv_sort1 … H2) -X2 //
+| #G #L1 #K1 #V1 #W1 #W #i #HLK1 #_ #HW1 #IHVW1 #L2 #s #d #e #HL21 #X #H #U2 #HWU2
+ elim (lift_inv_lref1 … H) * #Hid #H destruct
+ [ elim (lift_trans_ge … HW1 … HWU2) -W // #W2 #HW12 #HWU2
+ elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K2 #V2 #HK21 #HV12 #H destruct
+ /3 width=9 by sta_ldef/
+ | lapply (lift_trans_be … HW1 … HWU2 ? ?) -W /2 width=1 by le_S/ #HW1U2
+ lapply (drop_trans_ge … HL21 … HLK1 ?) -L1 /3 width=9 by sta_ldef, drop_inv_gen/
+ ]
+| #G #L1 #K1 #W1 #V1 #W #i #HLK1 #_ #HW1 #IHWV1 #L2 #s #d #e #HL21 #X #H #U2 #HWU2
+ elim (lift_inv_lref1 … H) * #Hid #H destruct
+ [ elim (lift_trans_ge … HW1 … HWU2) -W // <minus_plus #W #HW1 #HWU2
+ elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H
+ elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K2 #W2 #HK21 #HW12 #H destruct
+ lapply (lift_mono … HW1 … HW12) -HW1 #H destruct
+ elim (lift_total V1 (d-i-1) e) /3 width=9 by sta_ldec/
+ | lapply (lift_trans_be … HW1 … HWU2 ? ?) -W /2 width=1 by le_S/ #HW1U2
+ lapply (drop_trans_ge … HL21 … HLK1 ?) -L1 /3 width=9 by sta_ldec, drop_inv_gen/
+ ]
+| #a #I #G #L1 #V1 #T1 #U1 #_ #IHTU1 #L2 #s #d #e #HL21 #X1 #H1 #X2 #H2
+ elim (lift_inv_bind1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H destruct
+ elim (lift_inv_bind1 … H2) -H2 #X #U2 #H1 #HU12 #H2 destruct
+ lapply (lift_mono … H1 … HV12) -H1 #H destruct /4 width=6 by sta_bind, drop_skip/
+| #G #L1 #V1 #T1 #U1 #_ #IHTU1 #L2 #s #d #e #HL21 #X1 #H1 #X2 #H2
+ elim (lift_inv_flat1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H destruct
+ elim (lift_inv_flat1 … H2) -H2 #X #U2 #H1 #HU12 #H2 destruct
+ lapply (lift_mono … H1 … HV12) -H1 #H destruct /4 width=6 by sta_appl/
+| #G #L1 #W1 #T1 #U1 #_ #IHTU1 #L2 #s #d #e #HL21 #X #H #U2 #HU12
+ elim (lift_inv_flat1 … H) -H #W2 #T2 #_ #HT12 #H destruct /3 width=6 by sta_cast/
+]
+qed.
+
+(* Note: apparently this was missing in basic_1 *)
+lemma sta_inv_lift1: ∀h,G. l_deliftable_sn (sta h G).
+#h #G #L2 #T2 #U2 #H elim H -G -L2 -T2 -U2
+[ #G #L2 #k #L1 #s #d #e #_ #X #H
+ >(lift_inv_sort2 … H) -X /2 width=3 by sta_sort, lift_sort, ex2_intro/
+| #G #L2 #K2 #V2 #W2 #W #i #HLK2 #HVW2 #HW2 #IHVW2 #L1 #s #d #e #HL21 #X #H
+ elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HVW2 | -IHVW2 ]
+ [ elim (drop_conf_lt … HL21 … HLK2) -L2 // #K1 #V1 #HLK1 #HK21 #HV12
+ elim (IHVW2 … HK21 … HV12) -K2 -V2 #W1 #HW12 #HVW1
+ elim (lift_trans_le … HW12 … HW2) -W2 // >minus_plus <plus_minus_m_m /3 width=8 by sta_ldef, ex2_intro/
+ | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
+ elim (le_inv_plus_l … Hid) -Hid #Hdie #ei
+ elim (lift_split … HW2 d (i-e+1)) -HW2 /2 width=1 by le_S_S, le_S/
+ #W0 #HW20 <le_plus_minus_comm // >minus_minus_m_m /3 width=8 by sta_ldef, le_S, ex2_intro/
+ ]
+| #G #L2 #K2 #W2 #V2 #W #i #HLK2 #HWV2 #HW2 #IHWV2 #L1 #s #d #e #HL21 #X #H
+ elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HWV2 | -IHWV2 ]
+ [ elim (drop_conf_lt … HL21 … HLK2) -L2 // #K1 #W1 #HLK1 #HK21 #HW12
+ elim (IHWV2 … HK21 … HW12) -K2 #V1 #_ #HWV1
+ elim (lift_trans_le … HW12 … HW2) -W2 // >minus_plus <plus_minus_m_m /3 width=8 by sta_ldec, ex2_intro/
+ | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
+ elim (le_inv_plus_l … Hid) -Hid #Hdie #ei
+ elim (lift_split … HW2 d (i-e+1)) -HW2 /2 width=1 by le_S_S, le_S/
+ #W0 #HW20 <le_plus_minus_comm // >minus_minus_m_m /3 width=8 by sta_ldec, le_S, ex2_intro/
+ ]
+| #a #I #G #L2 #V2 #T2 #U2 #_ #IHTU2 #L1 #s #d #e #HL21 #X #H
+ elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
+ elim (IHTU2 (L1.ⓑ{I}V1) … HT12) -IHTU2 -HT12 /3 width=5 by sta_bind, drop_skip, lift_bind, ex2_intro/
+| #G #L2 #V2 #T2 #U2 #_ #IHTU2 #L1 #s #d #e #HL21 #X #H
+ elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
+ elim (IHTU2 … HL21 … HT12) -L2 -HT12 /3 width=5 by sta_appl, lift_flat, ex2_intro/
+| #G #L2 #W2 #T2 #U2 #_ #IHTU2 #L1 #s #d #e #HL21 #X #H
+ elim (lift_inv_flat2 … H) -H #W1 #T1 #_ #HT12 #H destruct
+ elim (IHTU2 … HL21 … HT12) -L2 -HT12 /3 width=3 by sta_cast, ex2_intro/
+]
+qed-.
+
+(* Advanced forward lemmas **************************************************)
+
+(* Basic_1: was: sty0_correct *)
+lemma sta_fwd_correct: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∃T0. ⦃G, L⦄ ⊢ U •[h] T0.
+#h #G #L #T #U #H elim H -G -L -T -U
+[ /2 width=2/
+| #G #L #K #V #W #W0 #i #HLK #_ #HW0 * #V0 #HWV0
+ lapply (drop_fwd_drop2 … HLK) -HLK #HLK
+ elim (lift_total V0 0 (i+1)) /3 width=11 by ex_intro, sta_lift/
+| #G #L #K #W #V #V0 #i #HLK #HWV #HWV0 #_
+ lapply (drop_fwd_drop2 … HLK) -HLK #HLK
+ elim (lift_total V 0 (i+1)) /3 width=11 by ex_intro, sta_lift/
+| #a #I #G #L #V #T #U #_ * /3 width=2 by sta_bind, ex_intro/
+| #G #L #V #T #U #_ * #T0 #HUT0 /3 width=2 by sta_appl, ex_intro/
+| #G #L #W #T #U #_ * /2 width=2 by ex_intro/
+]
+qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/substitution/drop_drop.ma".
-include "basic_2/static/sta.ma".
-
-(* STATIC TYPE ASSIGNMENT ON TERMS ******************************************)
-
-(* Properties on relocation *************************************************)
-
-(* Basic_1: was: sty0_lift *)
-lemma sta_lift: ∀h,G. l_liftable (sta h G).
-#h #G #L1 #T1 #U1 #H elim H -G -L1 -T1 -U1
-[ #G #L1 #k #L2 #s #d #e #HL21 #X1 #H1 #X2 #H2
- >(lift_inv_sort1 … H1) -X1
- >(lift_inv_sort1 … H2) -X2 //
-| #G #L1 #K1 #V1 #W1 #W #i #HLK1 #_ #HW1 #IHVW1 #L2 #s #d #e #HL21 #X #H #U2 #HWU2
- elim (lift_inv_lref1 … H) * #Hid #H destruct
- [ elim (lift_trans_ge … HW1 … HWU2) -W // #W2 #HW12 #HWU2
- elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H
- elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K2 #V2 #HK21 #HV12 #H destruct
- /3 width=9 by sta_ldef/
- | lapply (lift_trans_be … HW1 … HWU2 ? ?) -W /2 width=1 by le_S/ #HW1U2
- lapply (drop_trans_ge … HL21 … HLK1 ?) -L1 /3 width=9 by sta_ldef, drop_inv_gen/
- ]
-| #G #L1 #K1 #W1 #V1 #W #i #HLK1 #_ #HW1 #IHWV1 #L2 #s #d #e #HL21 #X #H #U2 #HWU2
- elim (lift_inv_lref1 … H) * #Hid #H destruct
- [ elim (lift_trans_ge … HW1 … HWU2) -W // <minus_plus #W #HW1 #HWU2
- elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H
- elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K2 #W2 #HK21 #HW12 #H destruct
- lapply (lift_mono … HW1 … HW12) -HW1 #H destruct
- elim (lift_total V1 (d-i-1) e) /3 width=9 by sta_ldec/
- | lapply (lift_trans_be … HW1 … HWU2 ? ?) -W /2 width=1 by le_S/ #HW1U2
- lapply (drop_trans_ge … HL21 … HLK1 ?) -L1 /3 width=9 by sta_ldec, drop_inv_gen/
- ]
-| #a #I #G #L1 #V1 #T1 #U1 #_ #IHTU1 #L2 #s #d #e #HL21 #X1 #H1 #X2 #H2
- elim (lift_inv_bind1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H destruct
- elim (lift_inv_bind1 … H2) -H2 #X #U2 #H1 #HU12 #H2 destruct
- lapply (lift_mono … H1 … HV12) -H1 #H destruct /4 width=6 by sta_bind, drop_skip/
-| #G #L1 #V1 #T1 #U1 #_ #IHTU1 #L2 #s #d #e #HL21 #X1 #H1 #X2 #H2
- elim (lift_inv_flat1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H destruct
- elim (lift_inv_flat1 … H2) -H2 #X #U2 #H1 #HU12 #H2 destruct
- lapply (lift_mono … H1 … HV12) -H1 #H destruct /4 width=6 by sta_appl/
-| #G #L1 #W1 #T1 #U1 #_ #IHTU1 #L2 #s #d #e #HL21 #X #H #U2 #HU12
- elim (lift_inv_flat1 … H) -H #W2 #T2 #_ #HT12 #H destruct /3 width=6 by sta_cast/
-]
-qed.
-
-(* Note: apparently this was missing in basic_1 *)
-lemma sta_inv_lift1: ∀h,G. l_deliftable_sn (sta h G).
-#h #G #L2 #T2 #U2 #H elim H -G -L2 -T2 -U2
-[ #G #L2 #k #L1 #s #d #e #_ #X #H
- >(lift_inv_sort2 … H) -X /2 width=3 by sta_sort, lift_sort, ex2_intro/
-| #G #L2 #K2 #V2 #W2 #W #i #HLK2 #HVW2 #HW2 #IHVW2 #L1 #s #d #e #HL21 #X #H
- elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HVW2 | -IHVW2 ]
- [ elim (drop_conf_lt … HL21 … HLK2) -L2 // #K1 #V1 #HLK1 #HK21 #HV12
- elim (IHVW2 … HK21 … HV12) -K2 -V2 #W1 #HW12 #HVW1
- elim (lift_trans_le … HW12 … HW2) -W2 // >minus_plus <plus_minus_m_m /3 width=8 by sta_ldef, ex2_intro/
- | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
- elim (le_inv_plus_l … Hid) -Hid #Hdie #ei
- elim (lift_split … HW2 d (i-e+1)) -HW2 /2 width=1 by le_S_S, le_S/
- #W0 #HW20 <le_plus_minus_comm // >minus_minus_m_m /3 width=8 by sta_ldef, le_S, ex2_intro/
- ]
-| #G #L2 #K2 #W2 #V2 #W #i #HLK2 #HWV2 #HW2 #IHWV2 #L1 #s #d #e #HL21 #X #H
- elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HWV2 | -IHWV2 ]
- [ elim (drop_conf_lt … HL21 … HLK2) -L2 // #K1 #W1 #HLK1 #HK21 #HW12
- elim (IHWV2 … HK21 … HW12) -K2 #V1 #_ #HWV1
- elim (lift_trans_le … HW12 … HW2) -W2 // >minus_plus <plus_minus_m_m /3 width=8 by sta_ldec, ex2_intro/
- | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
- elim (le_inv_plus_l … Hid) -Hid #Hdie #ei
- elim (lift_split … HW2 d (i-e+1)) -HW2 /2 width=1 by le_S_S, le_S/
- #W0 #HW20 <le_plus_minus_comm // >minus_minus_m_m /3 width=8 by sta_ldec, le_S, ex2_intro/
- ]
-| #a #I #G #L2 #V2 #T2 #U2 #_ #IHTU2 #L1 #s #d #e #HL21 #X #H
- elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
- elim (IHTU2 (L1.ⓑ{I}V1) … HT12) -IHTU2 -HT12 /3 width=5 by sta_bind, drop_skip, lift_bind, ex2_intro/
-| #G #L2 #V2 #T2 #U2 #_ #IHTU2 #L1 #s #d #e #HL21 #X #H
- elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
- elim (IHTU2 … HL21 … HT12) -L2 -HT12 /3 width=5 by sta_appl, lift_flat, ex2_intro/
-| #G #L2 #W2 #T2 #U2 #_ #IHTU2 #L1 #s #d #e #HL21 #X #H
- elim (lift_inv_flat2 … H) -H #W1 #T1 #_ #HT12 #H destruct
- elim (IHTU2 … HL21 … HT12) -L2 -HT12 /3 width=3 by sta_cast, ex2_intro/
-]
-qed-.
-
-(* Advanced forward lemmas **************************************************)
-
-(* Basic_1: was: sty0_correct *)
-lemma sta_fwd_correct: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∃T0. ⦃G, L⦄ ⊢ U •[h] T0.
-#h #G #L #T #U #H elim H -G -L -T -U
-[ /2 width=2/
-| #G #L #K #V #W #W0 #i #HLK #_ #HW0 * #V0 #HWV0
- lapply (drop_fwd_drop2 … HLK) -HLK #HLK
- elim (lift_total V0 0 (i+1)) /3 width=11 by ex_intro, sta_lift/
-| #G #L #K #W #V #V0 #i #HLK #HWV #HWV0 #_
- lapply (drop_fwd_drop2 … HLK) -HLK #HLK
- elim (lift_total V 0 (i+1)) /3 width=11 by ex_intro, sta_lift/
-| #a #I #G #L #V #T #U #_ * /3 width=2 by sta_bind, ex_intro/
-| #G #L #V #T #U #_ * #T0 #HUT0 /3 width=2 by sta_appl, ex_intro/
-| #G #L #W #T #U #_ * /2 width=2 by ex_intro/
-]
-qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/multiple/llpx_sn_drop.ma".
+include "basic_2/static/sta.ma".
+
+(* STRATIFIED STATIC TYPE ASSIGNMENT FOR TERMS ******************************)
+
+(* Properties on lazy sn pointwise extensions *******************************)
+
+lemma sta_llpx_sn_conf: ∀R. (∀L. reflexive … (R L)) → l_liftable R →
+ ∀h,G. s_r_confluent1 … (sta h G) (llpx_sn R 0).
+#R #H1R #H2R #h #G #Ls #T1 #T2 #H elim H -G -Ls -T1 -T2
+[ /3 width=4 by llpx_sn_fwd_length, llpx_sn_sort/
+| #G #Ls #Ks #V1s #W2s #V2s #i #HLKs #_ #HVW2s #IHV12s #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
+ #Kd #V1d #HLKd #HV1s #HV1sd
+ lapply (drop_fwd_drop2 … HLKs) -HLKs #HLKs
+ lapply (drop_fwd_drop2 … HLKd) -HLKd #HLKd
+ @(llpx_sn_lift_le … HLKs HLKd … HVW2s) -HLKs -HLKd -HVW2s /2 width=1 by/ (**) (* full auto too slow *)
+| #G #Ls #Ks #V1s #W1s #V2s #i #HLKs #_ #HV12s #IHVW1s #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
+ #Kd #V1d #HLKd #HV1s #HV1sd
+ lapply (drop_fwd_drop2 … HLKs) -HLKs #HLKs
+ lapply (drop_fwd_drop2 … HLKd) -HLKd #HLKd
+ @(llpx_sn_lift_le … HLKs HLKd … HV12s) -HLKs -HLKd -HV12s /2 width=1 by/ (**) (* full auto too slow *)
+| #a #I #G #Ls #V #T1 #T2 #_ #IHT12 #Ld #H elim (llpx_sn_inv_bind_O … H) -H
+ /4 width=5 by llpx_sn_bind_repl_SO, llpx_sn_bind/
+| #G #Ls #V #T1 #T2 #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
+ /3 width=1 by llpx_sn_flat/
+| #G #Ls #V #T1 #T2 #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
+ /3 width=1 by llpx_sn_flat/
+]
+qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/multiple/llpx_sn_drop.ma".
-include "basic_2/static/sta.ma".
-
-(* STRATIFIED STATIC TYPE ASSIGNMENT FOR TERMS ******************************)
-
-(* Properties on lazy sn pointwise extensions *******************************)
-
-lemma sta_llpx_sn_conf: ∀R. (∀L. reflexive … (R L)) → l_liftable R →
- ∀h,G. s_r_confluent1 … (sta h G) (llpx_sn R 0).
-#R #H1R #H2R #h #G #Ls #T1 #T2 #H elim H -G -Ls -T1 -T2
-[ /3 width=4 by llpx_sn_fwd_length, llpx_sn_sort/
-| #G #Ls #Ks #V1s #W2s #V2s #i #HLKs #_ #HVW2s #IHV12s #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
- #Kd #V1d #HLKd #HV1s #HV1sd
- lapply (drop_fwd_drop2 … HLKs) -HLKs #HLKs
- lapply (drop_fwd_drop2 … HLKd) -HLKd #HLKd
- @(llpx_sn_lift_le … HLKs HLKd … HVW2s) -HLKs -HLKd -HVW2s /2 width=1 by/ (**) (* full auto too slow *)
-| #G #Ls #Ks #V1s #W1s #V2s #i #HLKs #_ #HV12s #IHVW1s #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
- #Kd #V1d #HLKd #HV1s #HV1sd
- lapply (drop_fwd_drop2 … HLKs) -HLKs #HLKs
- lapply (drop_fwd_drop2 … HLKd) -HLKd #HLKd
- @(llpx_sn_lift_le … HLKs HLKd … HV12s) -HLKs -HLKd -HV12s /2 width=1 by/ (**) (* full auto too slow *)
-| #a #I #G #Ls #V #T1 #T2 #_ #IHT12 #Ld #H elim (llpx_sn_inv_bind_O … H) -H
- /4 width=5 by llpx_sn_bind_repl_SO, llpx_sn_bind/
-| #G #Ls #V #T1 #T2 #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
- /3 width=1 by llpx_sn_flat/
-| #G #Ls #V #T1 #T2 #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
- /3 width=1 by llpx_sn_flat/
-]
-qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/substitution/drop_drop.ma".
+include "basic_2/static/sta.ma".
+
+(* STATIC TYPE ASSIGNMENT ON TERMS ******************************************)
+
+(* Main properties **********************************************************)
+
+(* Note: apparently this was missing in basic_1 *)
+theorem sta_mono: ∀h,G,L. singlevalued … (sta h G L).
+#h #G #L #T #U1 #H elim H -G -L -T -U1
+[ #G #L #k #X #H >(sta_inv_sort1 … H) -X //
+| #G #L #K #V #W #U1 #i #HLK #_ #HWU1 #IHVW #U2 #H
+ elim (sta_inv_lref1 … H) -H * #K0 #V0 #W0 #HLK0 #HVW0 #HW0U2
+ lapply (drop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct
+ lapply (IHVW … HVW0) -IHVW -HVW0 #H destruct
+ >(lift_mono … HWU1 … HW0U2) -W0 -U1 //
+| #G #L #K #W #V #U1 #i #HLK #_ #HWU1 #IHWV #U2 #H
+ elim (sta_inv_lref1 … H) -H * #K0 #W0 #V0 #HLK0 #HWV0 #HV0U2
+ lapply (drop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct
+ lapply (IHWV … HWV0) -IHWV -HWV0 #H destruct
+ >(lift_mono … HWU1 … HV0U2) -W -U1 //
+| #a #I #G #L #V #T #U1 #_ #IHTU1 #X #H
+ elim (sta_inv_bind1 … H) -H #U2 #HTU2 #H destruct /3 width=1 by eq_f/
+| #G #L #V #T #U1 #_ #IHTU1 #X #H
+ elim (sta_inv_appl1 … H) -H #U2 #HTU2 #H destruct /3 width=1 by eq_f/
+| #G #L #W #T #U1 #_ #IHTU1 #U2 #H
+ lapply (sta_inv_cast1 … H) -H /2 width=1 by/
+]
+qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/substitution/drop_drop.ma".
-include "basic_2/static/sta.ma".
-
-(* STATIC TYPE ASSIGNMENT ON TERMS ******************************************)
-
-(* Main properties **********************************************************)
-
-(* Note: apparently this was missing in basic_1 *)
-theorem sta_mono: ∀h,G,L. singlevalued … (sta h G L).
-#h #G #L #T #U1 #H elim H -G -L -T -U1
-[ #G #L #k #X #H >(sta_inv_sort1 … H) -X //
-| #G #L #K #V #W #U1 #i #HLK #_ #HWU1 #IHVW #U2 #H
- elim (sta_inv_lref1 … H) -H * #K0 #V0 #W0 #HLK0 #HVW0 #HW0U2
- lapply (drop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct
- lapply (IHVW … HVW0) -IHVW -HVW0 #H destruct
- >(lift_mono … HWU1 … HW0U2) -W0 -U1 //
-| #G #L #K #W #V #U1 #i #HLK #_ #HWU1 #IHWV #U2 #H
- elim (sta_inv_lref1 … H) -H * #K0 #W0 #V0 #HLK0 #HWV0 #HV0U2
- lapply (drop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct
- lapply (IHWV … HWV0) -IHWV -HWV0 #H destruct
- >(lift_mono … HWU1 … HV0U2) -W -U1 //
-| #a #I #G #L #V #T #U1 #_ #IHTU1 #X #H
- elim (sta_inv_bind1 … H) -H #U2 #HTU2 #H destruct /3 width=1 by eq_f/
-| #G #L #V #T #U1 #_ #IHTU1 #X #H
- elim (sta_inv_appl1 … H) -H #U2 #HTU2 #H destruct /3 width=1 by eq_f/
-| #G #L #W #T #U1 #_ #IHTU1 #U2 #H
- lapply (sta_inv_cast1 … H) -H /2 width=1 by/
-]
-qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
+
+notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 • break [ term 46 h ] break term 46 T2 )"
+ non associative with precedence 45
+ for @{ 'StaticType $h $G $L $T1 $T2 }.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-
-notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 • break [ term 46 h ] break term 46 T2 )"
- non associative with precedence 45
- for @{ 'StaticType $h $G $L $T1 $T2 }.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
+
+notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 • • * break [ term 46 h , break term 46 l ] break term 46 T2 )"
+ non associative with precedence 45
+ for @{ 'StaticTypeStarAlt $h $G $L $l $T1 $T2 }.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-
-notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 • • * break [ term 46 h , break term 46 l ] break term 46 T2 )"
- non associative with precedence 45
- for @{ 'StaticTypeStarAlt $h $G $L $l $T1 $T2 }.
include "basic_2/notation/relations/pred_4.ma".
include "basic_2/grammar/genv.ma".
include "basic_2/static/lsubr.ma".
+include "basic_2/unfold/lstas.ma".
(* CONTEXT-SENSITIVE PARALLEL REDUCTION FOR TERMS ***************************)
]
qed-.
+fact lstas_cpr_aux: ∀h,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 •*[h, l] T2 →
+ l = 0 → ⦃G, L⦄ ⊢ T1 ➡ T2.
+#h #G #L #T1 #T2 #l #H elim H -G -L -T1 -T2 -l
+/3 width=1 by cpr_eps, cpr_flat, cpr_bind/
+[ #G #L #l #k #H0 destruct normalize //
+| #G #L #K #V1 #V2 #W2 #i #l #HLK #_ #HVW2 #IHV12 #H destruct
+ /3 width=6 by cpr_delta/
+| #G #L #K #V1 #V2 #W2 #i #l #_ #_ #_ #_ <plus_n_Sm #H destruct
+]
+qed-.
+
+lemma lstas_cpr: ∀h,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 •*[h, 0] T2 → ⦃G, L⦄ ⊢ T1 ➡ T2.
+/2 width=4 by lstas_cpr_aux/ qed.
+
(* Basic inversion lemmas ***************************************************)
fact cpr_inv_atom1_aux: ∀G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ∀I. T1 = ⓪{I} →
include "basic_2/substitution/drop_drop.ma".
include "basic_2/multiple/fqus_alt.ma".
-include "basic_2/static/sta.ma".
include "basic_2/static/da.ma".
include "basic_2/reduction/cpx.ma".
(* Advanced properties ******************************************************)
-lemma sta_cpx: ∀h,g,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 •[h] T2 →
- ⦃G, L⦄ ⊢ T1 ▪[h, g] l+1 → ⦃G, L⦄ ⊢ T1 ➡[h, g] T2.
-#h #g #G #L #T1 #T2 #l #H elim H -G -L -T1 -T2
-[ /3 width=4 by cpx_st, da_inv_sort/
-| #G #L #K #V1 #V2 #W2 #i #HLK #_ #HVW2 #IHV12 #H
+fact sta_cpx_aux: ∀h,g,G,L,T1,T2,l2,l1. ⦃G, L⦄ ⊢ T1 •*[h, l2] T2 → l2 = 1 →
+ ⦃G, L⦄ ⊢ T1 ▪[h, g] l1+1 → ⦃G, L⦄ ⊢ T1 ➡[h, g] T2.
+#h #g #G #L #T1 #T2 #l2 #l1 #H elim H -G -L -T1 -T2 -l2
+[ #G #L #l2 #k #H0 destruct normalize
+ /3 width=4 by cpx_st, da_inv_sort/
+| #G #L #K #V1 #V2 #W2 #i #l2 #HLK #_ #HVW2 #IHV12 #H0 #H destruct
elim (da_inv_lref … H) -H * #K0 #V0 [| #l0 ] #HLK0
lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct /3 width=7 by cpx_delta/
-| #G #L #K #W1 #W2 #V1 #i #HLK #_ #HWV1 #IHW12 #H
- elim (da_inv_lref … H) -H * #K0 #W0 [| #l1 ] #HLK0
- lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct /3 width=7 by cpx_delta/
+| #G #L #K #V1 #V2 #i #_ #_ #_ #H destruct
+| #G #L #K #V1 #V2 #W2 #i #l2 #HLK #HV12 #HVW2 #_ #H0 #H
+ lapply (discr_plus_xy_y … H0) -H0 #H0 destruct
+ elim (da_inv_lref … H) -H * #K0 #V0 [| #l1 ] #HLK0
+ lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct
+ /4 width=7 by cpx_delta, cpr_cpx, lstas_cpr/
| /4 width=2 by cpx_bind, da_inv_bind/
| /4 width=3 by cpx_flat, da_inv_flat/
| /4 width=3 by cpx_eps, da_inv_flat/
]
-qed.
+qed-.
+
+lemma sta_cpx: ∀h,g,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 •*[h, 1] T2 →
+ ⦃G, L⦄ ⊢ T1 ▪[h, g] l+1 → ⦃G, L⦄ ⊢ T1 ➡[h, g] T2.
+/2 width=3 by sta_cpx_aux/ qed.
(* Relocation properties ****************************************************)
qed-.
lemma fqu_sta_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 •[h] U2 →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 •*[h, 1] U2 →
∀l. ⦃G2, L2⦄ ⊢ T2 ▪[h, g] l+1 →
∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, g] U1 & ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, U2⦄.
/3 width=5 by fqu_cpx_trans, sta_cpx/ qed-.
qed-.
lemma fquq_sta_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 •[h] U2 →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 •*[h, 1] U2 →
∀l. ⦃G2, L2⦄ ⊢ T2 ▪[h, g] l+1 →
∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, g] U1 & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
/3 width=5 by fquq_cpx_trans, sta_cpx/ qed-.
(* Advanced properties ******************************************************)
lemma sta_fpb: ∀h,g,G,L,T1,T2,l.
- ⦃G, L⦄ ⊢ T1 ▪[h, g] l+1 → ⦃G, L⦄ ⊢ T1 •[h] T2 →
+ ⦃G, L⦄ ⊢ T1 ▪[h, g] l+1 → ⦃G, L⦄ ⊢ T1 •*[h, 1] T2 →
⦃G, L, T1⦄ ≽[h, g] ⦃G, L, T2⦄.
/3 width=4 by fpb_cpx, sta_cpx/ qed.
(* Advanced inversion lemmas ************************************************)
-lemma zero_eq_plus: ∀x,y. 0 = x + y → 0 = x ∧ 0 = y.
-* /2 width=1 by conj/ #x #y normalize #H destruct
-qed-.
-
lemma lstas_split_aux: ∀h,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 •*[h, l] T2 → ∀l1,l2. l = l1 + l2 →
∃∃T. ⦃G, L⦄ ⊢ T1 •*[h, l1] T & ⦃G, L⦄ ⊢ T •*[h, l2] T2.
#h #G #L #T1 #T2 #l #H elim H -G -L -T1 -T2 -l
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/multiple/llpx_sn_drop.ma".
+include "basic_2/unfold/lstas.ma".
+
+(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
+
+(* Properties on lazy sn pointwise extensions *******************************)
+
+lemma lstas_llpx_sn_conf: ∀R. (∀L. reflexive … (R L)) → l_liftable R →
+ ∀h,G,l. s_r_confluent1 … (lstas h l G) (llpx_sn R 0).
+#R #H1R #H2R #h #G #l #Ls #T1 #T2 #H elim H -G -Ls -T1 -T2 -l
+[ /3 width=4 by llpx_sn_fwd_length, llpx_sn_sort/
+| #G #Ls #Ks #V1s #V2s #W2s #i #l #HLKs #_ #HVW2s #IHV12s #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
+ #Kd #V1d #HLKd #HV1s #HV1sd
+ lapply (drop_fwd_drop2 … HLKs) -HLKs #HLKs
+ lapply (drop_fwd_drop2 … HLKd) -HLKd #HLKd
+ @(llpx_sn_lift_le … HLKs HLKd … HVW2s) -HLKs -HLKd -HVW2s /2 width=1 by/ (**) (* full auto too slow *)
+| //
+| #G #Ls #Ks #V1s #V2s #W2s #i #l #HLKs #_ #HVW2s #IHV12s #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
+ #Kd #V1d #HLKd #HV1s #HV1sd
+ lapply (drop_fwd_drop2 … HLKs) -HLKs #HLKs
+ lapply (drop_fwd_drop2 … HLKd) -HLKd #HLKd
+ @(llpx_sn_lift_le … HLKs HLKd … HVW2s) -HLKs -HLKd -HVW2s /2 width=1 by/ (**) (* full auto too slow *)
+| #a #I #G #Ls #V #T1 #T2 #l #_ #IHT12 #Ld #H elim (llpx_sn_inv_bind_O … H) -H
+ /4 width=5 by llpx_sn_bind_repl_SO, llpx_sn_bind/
+| #G #Ls #V #T1 #T2 #l #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
+ /3 width=1 by llpx_sn_flat/
+| #G #Ls #V #T1 #T2 #l #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
+ /3 width=1 by llpx_sn_flat/
+]
+qed-.
(* Inversion & forward lemmas ***********************************************)
+lemma discr_plus_xy_y: ∀x,y. x + y = y → x = 0.
+// qed-.
+
lemma lt_plus_SO_to_le: ∀x,y. x < y + 1 → x ≤ y.
/2 width=1 by monotonic_pred/ qed-.
#H destruct
qed-.
+lemma zero_eq_plus: ∀x,y. 0 = x + y → 0 = x ∧ 0 = y.
+* /2 width=1 by conj/ #x #y normalize #H destruct
+qed-.
+
(* Iterators ****************************************************************)
(* Note: see also: lib/arithemetics/bigops.ma *)