# elim #######################################################################
elim:
- @echo " MATITADEP"
- $(H)grep "elim (.*?)" $(MAS)
+ @echo " GREP elim"
+ $(H)grep "elim (.*?)" $(MAS) || true
# orig #######################################################################
<section>Summary of the Specification</section>
<body>Here is a numerical acount of the specification's contents
and its timeline.
- Nodes are counted according to the "intrinsinc complexity measure"
- [F. Guidi: "Procedural Representation of CIC Proof Terms"
+ Nodes are counted according to the "intrinsinc complexity measure"
+ [F. Guidi: "Procedural Representation of CIC Proof Terms"
Journal of Automated Reasoning 44(1-2), Springer (February 2010),
- pp. 53-78].
+ pp. 53-78].
</body>
<table name="apps_2_sum"/>
<news date="2012 February 24.">
include "basic_2/notation/relations/ineint_5.ma".
include "basic_2/grammar/aarity.ma".
-include "basic_2/substitution/gr2_gr2.ma".
+include "basic_2/substitution/gr2_gr2.ma".
include "basic_2/substitution/lifts_lift_vector.ma".
include "basic_2/substitution/ldrops_ldrop.ma".
include "basic_2/computation/acp.ma".
(**************************************************************************)
include "basic_2/notation/relations/dpredstar_6.ma".
-include "basic_2/unfold/sstas.ma".
+include "basic_2/unfold/lsstas.ma".
include "basic_2/computation/cprs.ma".
(* DECOMPOSED EXTENDED PARALLEL COMPUTATION ON TERMS ************************)
definition cpds: ∀h. sd h → relation4 genv lenv term term ≝
λh,g,G,L,T1,T2.
- ∃∃T. ⦃G, L⦄ ⊢ T1 •*[h, g] T & ⦃G, L⦄ ⊢ T ➡* T2.
+ ∃∃T,l1,l2. l2 ≤ l1 & ⦃G, L⦄ ⊢ T1 ▪[h, g] l1 & ⦃G, L⦄ ⊢ T1 •*[h, g, l2] T & ⦃G, L⦄ ⊢ T ➡* T2.
interpretation "decomposed extended parallel computation (term)"
'DPRedStar h g G L T1 T2 = (cpds h g G L T1 T2).
(* Basic properties *********************************************************)
-lemma cpds_refl: ∀h,g,G,L. reflexive … (cpds h g G L).
-/2 width=3/ qed.
+lemma ssta_cprs_cpds: ∀h,g,G,L,T1,T,T2,l. ⦃G, L⦄ ⊢ T1 ▪[h, g] l+1 → ⦃G, L⦄ ⊢ T1 •[h, g] T →
+ ⦃G, L⦄ ⊢ T ➡* T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2.
+/3 width=7/ qed.
+
+lemma lsstas_cpds: ∀h,g,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 ▪[h, g] l → ⦃G, L⦄ ⊢ T1 •*[h, g, l] T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2.
+/2 width=7/ qed.
-lemma sstas_cpds: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 •*[h, g] T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2.
-/2 width=3/ qed.
+lemma cprs_cpds: ∀h,g,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 ▪[h, g] l → ⦃G, L⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2.
+/2 width=7/ qed.
-lemma cprs_cpds: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2.
-/2 width=3/ qed.
+lemma cpds_refl: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l → ⦃G, L⦄ ⊢ T •*➡*[h, g] T.
+/2 width=2/ qed.
lemma cpds_strap1: ∀h,g,G,L,T1,T,T2.
⦃G, L⦄ ⊢ T1 •*➡*[h, g] T → ⦃G, L⦄ ⊢ T ➡ T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2.
-#h #g #G #L #T1 #T #T2 * /3 width=5/
-qed.
-
-lemma cpds_strap2: ∀h,g,G,L,T1,T,T2,l.
- ⦃G, L⦄ ⊢ T1 •[h, g] ⦃l+1, T⦄ → ⦃G, L⦄ ⊢ T •*➡*[h, g] T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2.
-#h #g #G #L #T1 #T #T2 #l #HT1 * /3 width=4/
+#h #g #G #L #T1 #T #T2 * /3 width=9/
qed.
-
-lemma ssta_cprs_cpds: ∀h,g,G,L,T1,T,T2,l. ⦃G, L⦄ ⊢ T1 •[h, g] ⦃l+1, T⦄ →
- ⦃G, L⦄ ⊢ T ➡* T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2.
-/3 width=3/ qed.
(* *)
(**************************************************************************)
-include "basic_2/unfold/sstas_aaa.ma".
+include "basic_2/unfold/lsstas_aaa.ma".
include "basic_2/computation/cpxs_aaa.ma".
include "basic_2/computation/cpds.ma".
(* Properties on atomic arity assignment for terms **************************)
-lemma cpds_aaa: ∀h,g,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀U. ⦃G, L⦄ ⊢ T •*➡*[h, g] U → ⦃G, L⦄ ⊢ U ⁝ A.
-#h #g #G #L #T #A #HT #U * /3 width=5 by sstas_aaa, aaa_cprs_conf/
+lemma aaa_cpds_conf: ∀h,g,G,L. Conf3 … (aaa G L) (cpds h g G L).
+#h #g #G #L #A #T #HT #U * /3 width=6 by aaa_lsstas_conf, aaa_cprs_conf/
qed.
(* *)
(**************************************************************************)
-include "basic_2/unfold/sstas_sstas.ma".
+include "basic_2/unfold/lsstas_lsstas.ma".
include "basic_2/computation/lprs_cprs.ma".
include "basic_2/computation/cpxs_cpxs.ma".
include "basic_2/computation/cpds.ma".
(* Advanced properties ******************************************************)
+lemma cpds_strap2: ∀h,g,G,L,T1,T,T2,l. ⦃G, L⦄ ⊢ T1 ▪[h, g] l+1 →
+ ⦃G, L⦄ ⊢ T1 •[h, g] T → ⦃G, L⦄ ⊢ T •*➡*[h, g] T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2.
+#h #g #G #L #T1 #T #T2 #l #Hl #HT1 *
+#T0 #l0 #l1 #Hl10 #HT #HT0 #HT02
+lapply (ssta_da_conf … HT1 … Hl) <minus_plus_m_m #H0T
+lapply (da_mono … H0T … HT) -HT -H0T #H destruct /3 width=7/
+qed.
+
lemma cpds_cprs_trans: ∀h,g,G,L,T1,T,T2.
⦃G, L⦄ ⊢ T1 •*➡*[h, g] T → ⦃G, L⦄ ⊢ T ➡* T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2.
-#h #g #G #L #T1 #T #T2 * #T0 #HT10 #HT0 #HT2
-lapply (cprs_trans … HT0 … HT2) -T /2 width=3/
+#h #g #G #L #T1 #T #T2 * #T0 #l1 #l2 #Hl12 #Hl1 #HT10 #HT0 #HT2
+lapply (cprs_trans … HT0 … HT2) -T /2 width=7/
qed-.
-lemma sstas_cpds_trans: ∀h,g,G,L,T1,T,T2.
- ⦃G, L⦄ ⊢ T1 •*[h, g] T → ⦃G, L⦄ ⊢ T •*➡*[h, g] T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2.
-#h #g #G #L #T1 #T #T2 #HT1 * #T0 #HT0 #HT02
-lapply (sstas_trans … HT1 … HT0) -T /2 width=3/
+lemma lsstas_cpds_trans: ∀h,g,G,L,T1,T,T2,l1,l2.
+ l2 ≤ l1 → ⦃G, L⦄ ⊢ T1 ▪[h, g] l1 →
+ ⦃G, L⦄ ⊢ T1 •*[h, g, l2] T → ⦃G, L⦄ ⊢ T •*➡*[h, g] T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2.
+#h #g #G #L #T1 #T #T2 #l1 #l2 #Hl21 #Hl1 #HT1 * #T0 #l3 #l4 #Hl43 #Hl3 #HT0 #HT02
+lapply (lsstas_da_conf … HT1 … Hl1) #H0T
+lapply (da_mono … H0T … Hl3) -H0T -Hl3 #H destruct
+lapply (le_minus_to_plus_r … Hl21 Hl43) -Hl21 -Hl43
+lapply (lsstas_trans … HT1 … HT0) -T /2 width=7/
qed-.
(* Advanced inversion lemmas ************************************************)
lemma cpds_inv_abst1: ∀h,g,a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{a}V1.T1 •*➡*[h, g] U2 →
∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡* V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 •*➡*[h, g] T2 &
U2 = ⓛ{a}V2.T2.
-#h #g #a #G #L #V1 #T1 #U2 * #X #H1 #H2
-elim (sstas_inv_bind1 … H1) -H1 #U #HTU1 #H destruct
-elim (cprs_inv_abst1 … H2) -H2 #V2 #T2 #HV12 #HUT2 #H destruct /3 width=5/
+#h #g #a #G #L #V1 #T1 #U2 * #X #l1 #l2 #Hl21 #Hl1 #H1 #H2
+lapply (da_inv_bind … Hl1) -Hl1 #Hl1
+elim (lsstas_inv_bind1 … H1) -H1 #U #HTU1 #H destruct
+elim (cprs_inv_abst1 … H2) -H2 #V2 #T2 #HV12 #HUT2 #H destruct /3 width=7/
qed-.
lemma cpds_inv_abbr_abst: ∀h,g,a1,a2,G,L,V1,W2,T1,T2. ⦃G, L⦄ ⊢ ⓓ{a1}V1.T1 •*➡*[h, g] ⓛ{a2}W2.T2 →
∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 •*➡*[h, g] T & ⇧[0, 1] ⓛ{a2}W2.T2 ≡ T & a1 = true.
-#h #g #a1 #a2 #G #L #V1 #W2 #T1 #T2 * #X #H1 #H2
-elim (sstas_inv_bind1 … H1) -H1 #U1 #HTU1 #H destruct
+#h #g #a1 #a2 #G #L #V1 #W2 #T1 #T2 * #X #l1 #l2 #Hl21 #Hl1 #H1 #H2
+lapply (da_inv_bind … Hl1) -Hl1 #Hl1
+elim (lsstas_inv_bind1 … H1) -H1 #U1 #HTU1 #H destruct
elim (cprs_inv_abbr1 … H2) -H2 *
[ #V2 #U2 #HV12 #HU12 #H destruct
-| /3 width=3/
+| /3 width=7/
]
qed-.
(* Advanced forward lemmas **************************************************)
lemma cpds_fwd_cpxs: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2.
-#h #g #G #L #T1 #T2 * /3 width=3 by cpxs_trans, sstas_cpxs, cprs_cpxs/
+#h #g #G #L #T1 #T2 * /3 width=5 by cpxs_trans, lsstas_cpxs, cprs_cpxs/
qed-.
(* *)
(**************************************************************************)
-include "basic_2/unfold/sstas_lift.ma".
+include "basic_2/unfold/lsstas_lift.ma".
include "basic_2/computation/cprs_lift.ma".
include "basic_2/computation/cpds.ma".
(* Relocation properties ****************************************************)
lemma cpds_lift: ∀h,g,G. l_liftable (cpds h g G).
-#h #g #G #K #T1 #T2 * #T #HT1 #HT2 #L #d #e
+#h #g #G #K #T1 #T2 * #T #l1 #l2 #Hl12 #Hl1 #HT1 #HT2 #L #d #e
elim (lift_total T d e)
-/3 width=11 by cprs_lift, sstas_lift, ex2_intro/
+/3 width=15 by cprs_lift, da_lift, lsstas_lift, ex4_3_intro/
qed.
lemma cpds_inv_lift1: ∀h,g,G. l_deliftable_sn (cpds h g G).
-#h #g #G #L #U1 #U2 * #U #HU1 #HU2 #K #d #e #HLK #T1 #HTU1
-elim (sstas_inv_lift1 … HU1 … HLK … HTU1) -U1 #T #HT1 #HTU
-elim (cprs_inv_lift1 … HU2 … HLK … HTU) -U -L /3 width=5/
+#h #g #G #L #U1 #U2 * #U #l1 #l2 #Hl12 #Hl1 #HU1 #HU2 #K #d #e #HLK #T1 #HTU1
+lapply (da_inv_lift … Hl1 … HLK … HTU1) -Hl1 #Hl1
+elim (lsstas_inv_lift1 … HU1 … HLK … HTU1) -U1 #T #HTU #HT1
+elim (cprs_inv_lift1 … HU2 … HLK … HTU) -U -L /3 width=9/
qed-.
(**************************************************************************)
include "basic_2/notation/relations/predstar_6.ma".
-include "basic_2/unfold/sstas.ma".
include "basic_2/reduction/cnx.ma".
include "basic_2/computation/cprs.ma".
/3 width=5 by lsubr_cpx_trans, TC_lsub_trans/
qed-.
-lemma sstas_cpxs: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 •* [h, g] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2.
-#h #g #G #L #T1 #T2 #H @(sstas_ind … H) -T2 //
-/3 width=4 by cpxs_strap1, ssta_cpx/
-qed.
-
lemma cprs_cpxs: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2.
#h #g #G #L #T1 #T2 #H @(cprs_ind … H) -T2 // /3 width=3/
qed.
(**************************************************************************)
include "basic_2/substitution/fsups_fsups.ma".
+include "basic_2/unfold/lsstas_lift.ma".
include "basic_2/reduction/cpx_lift.ma".
include "basic_2/computation/cpxs.ma".
(* Advanced properties ******************************************************)
+lemma lsstas_cpxs: ∀h,g,G,L,T1,T2,l1. ⦃G, L⦄ ⊢ T1 •* [h, g, l1] T2 →
+ ∀l2. ⦃G, L⦄ ⊢ T1 ▪ [h, g] l2 → l1 ≤ l2 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2.
+#h #g #G #L #T1 #T2 #l1 #H @(lsstas_ind_dx … H) -T2 -l1 //
+#l1 #T #T2 #HT1 #HT2 #IHT1 #l2 #Hl2 #Hl12
+lapply (lsstas_da_conf … HT1 … Hl2) -HT1
+>(plus_minus_m_m (l2-l1) 1 ?) [2: /2 width=1/ ]
+/4 width=5 by cpxs_strap1, ssta_cpx, lt_to_le/
+qed.
+
lemma cpxs_delta: ∀h,g,I,G,L,K,V,V2,i.
⇩[0, i] L ≡ K.ⓑ{I}V → ⦃G, K⦄ ⊢ V ➡*[h, g] V2 →
∀W2. ⇧[0, i + 1] V2 ≡ W2 → ⦃G, L⦄ ⊢ #i ➡*[h, g] W2.
elim (IHTU2 … HT2) -T2 /3 width=3/
qed-.
+lemma fsupq_lsstas_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊃⸮ ⦃G2, L2, T2⦄ →
+ ∀U2,l1. ⦃G2, L2⦄ ⊢ T2 •*[h, g, l1] U2 →
+ ∀l2. ⦃G2, L2⦄ ⊢ T2 ▪ [h, g] l2 → l1 ≤ l2 →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] U1 & ⦃G1, L1, U1⦄ ⊃* ⦃G2, L2, U2⦄.
+/3 width=5 by fsupq_cpxs_trans, lsstas_cpxs/ qed-.
+
lemma fsups_cpxs_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊃* ⦃G2, L2, T2⦄ →
∀U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, g] U2 →
∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] U1 & ⦃G1, L1, U1⦄ ⊃* ⦃G2, L2, U2⦄.
lapply (fsups_trans … HT2 … HTU2) -G -L -T2 /2 width=3/
qed-.
-lemma fsup_ssta_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊃ ⦃G2, L2, T2⦄ →
- ∀U2,l. ⦃G2, L2⦄ ⊢ T2 •[h, g] ⦃l+1, U2⦄ →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, g] U1 & ⦃G1, L1, U1⦄ ⊃⸮ ⦃G2, L2, U2⦄.
-/3 width=4 by fsup_cpx_trans, ssta_cpx/ qed-.
+lemma fsups_lsstas_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊃* ⦃G2, L2, T2⦄ →
+ ∀U2,l1. ⦃G2, L2⦄ ⊢ T2 •*[h, g, l1] U2 →
+ ∀l2. ⦃G2, L2⦄ ⊢ T2 ▪ [h, g] l2 → l1 ≤ l2 →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] U1 & ⦃G1, L1, U1⦄ ⊃* ⦃G2, L2, U2⦄.
+/3 width=7 by fsups_cpxs_trans, lsstas_cpxs/ qed-.
| #b #W0 #T0 #HT0 #HU
elim (cpxs_inv_abst1 … HT0) -HT0 #W1 #T1 #HW1 #HT1 #H destruct
lapply (lsubr_cpxs_trans … HT1 (L.ⓓⓝW.V) ?) -HT1 /2 width=1/ #HT1
- @or_intror @(cpxs_trans … HU) -U /3 width=1/ (**) (* explicit constructor *)
+ @or_intror @(cpxs_trans … HU) -U /3 width=1/ (**) (* explicit constructor *)
| #b #V1 #V2 #V0 #T1 #_ #_ #HT1 #_
elim (cpxs_inv_abst1 … HT1) -HT1 #W2 #T2 #_ #_ #H destruct
]
@or_intror @(cpxs_trans … HU) -U (**) (* explicit constructor *)
@(cpxs_trans … (+ⓓV.ⓐV2.ⓛ{b}W2.T2)) [ /3 width=1/ ] -T
@(cpxs_strap2 … (ⓐV1.ⓛ{b}W.T0)) [2: /2 width=1/ ]
- /4 width=7 by cpx_zeta, lift_bind, lift_flat/ (**) (* auto too slow without trace *)
+ /4 width=7 by cpx_zeta, lift_bind, lift_flat/ (**) (* auto too slow without trace *)
]
| #b #V3 #V4 #V0 #T0 #HV13 #HV34 #HT0 #HU
@or_intror @(cpxs_trans … HU) -U (**) (* explicit constructor *)
@(cpxs_strap1 … (ⓐV0.ⓓ{b}V2.T0)) [ /2 width=1/ ] -V -Vs -T /2 width=3/
| @or3_intro2 -T (**) (* explicit constructor *)
@(cpxs_trans … HU) -U
- @(cpxs_strap1 … (ⓐV0.ⓓ{b}V2.T0)) [ /2 width=1/ ] -V -Vs -W /2 width=3/
+ @(cpxs_strap1 … (ⓐV0.ⓓ{b}V2.T0)) [ /2 width=1/ ] -V -Vs -W /2 width=3/
]
]
qed-.
elim (cpx_inv_abst1 … H1) -H1
#W0 #T0 #HLW0 #HLT0 #H destruct
elim (eq_false_inv_tpair_sn … H2) -H2
-[ -IHT #H lapply (csn_cpx_trans … HLT0) // -HT
+[ -IHT #H lapply (csn_cpx_trans … HLT0) // -HT
#HT0 lapply (csn_lpx_conf … (L.ⓛW0) … HT0) -HT0 /2 width=1/ /3 width=1/
| -IHW -HLW0 -HT * #H destruct /3 width=1/
]
| lsubsv_atom: lsubsv h g G (⋆) (⋆)
| lsubsv_pair: ∀I,L1,L2,V. lsubsv h g G L1 L2 →
lsubsv h g G (L1.ⓑ{I}V) (L2.ⓑ{I}V)
-| lsubsv_abbr: ∀L1,L2,W,V,W1,V2,l. ⦃G, L1⦄ ⊢ ⓝW.V ¡[h, g] → ⦃G, L2⦄ ⊢ W ¡[h, g] →
- ⦃G, L1⦄ ⊢ V •[h, g] ⦃l+1, W1⦄ → ⦃G, L2⦄ ⊢ W •[h, g] ⦃l, V2⦄ →
+| lsubsv_abbr: ∀L1,L2,W,V,l. ⦃G, L1⦄ ⊢ W ¡[h, g] → ⦃G, L1⦄ ⊢ V ¡[h, g] →
+ scast h g l G L1 V W → ⦃G, L2⦄ ⊢ W ¡[h, g] →
+ ⦃G, L1⦄ ⊢ V ▪[h, g] l+1 → ⦃G, L2⦄ ⊢ W ▪[h, g] l →
lsubsv h g G L1 L2 → lsubsv h g G (L1.ⓓⓝW.V) (L2.ⓛW)
.
#h #g #G #L1 #L2 * -L1 -L2
[ //
| #I #L1 #L2 #V #_ #H destruct
-| #L1 #L2 #W #V #V1 #V2 #l #_ #_ #_ #_ #_ #H destruct
+| #L1 #L2 #W #V #l #_ #_ #_ #_ #_ #_ #_ #H destruct
]
qed-.
fact lsubsv_inv_pair1_aux: ∀h,g,G,L1,L2. G ⊢ L1 ¡⊑[h, g] L2 →
∀I,K1,X. L1 = K1.ⓑ{I}X →
(∃∃K2. G ⊢ K1 ¡⊑[h, g] K2 & L2 = K2.ⓑ{I}X) ∨
- ∃∃K2,W,V,W1,V2,l. ⦃G, K1⦄ ⊢ X ¡[h, g] & ⦃G, K2⦄ ⊢ W ¡[h, g] &
- ⦃G, K1⦄ ⊢ V •[h, g] ⦃l+1, W1⦄ & ⦃G, K2⦄ ⊢ W •[h, g] ⦃l, V2⦄ &
- G ⊢ K1 ¡⊑[h, g] K2 &
- I = Abbr & L2 = K2.ⓛW & X = ⓝW.V.
+ ∃∃K2,W,V,l. ⦃G, K1⦄ ⊢ W ¡[h, g] & ⦃G, K1⦄ ⊢ V ¡[h, g] &
+ scast h g l G K1 V W & ⦃G, K2⦄ ⊢ W ¡[h, g] &
+ ⦃G, K1⦄ ⊢ V ▪[h, g] l+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] l &
+ G ⊢ K1 ¡⊑[h, g] K2 &
+ I = Abbr & L2 = K2.ⓛW & X = ⓝW.V.
#h #g #G #L1 #L2 * -L1 -L2
[ #J #K1 #X #H destruct
| #I #L1 #L2 #V #HL12 #J #K1 #X #H destruct /3 width=3/
-| #L1 #L2 #W #V #W1 #V2 #l #HV #HW #HW1 #HV2 #HL12 #J #K1 #X #H destruct /3 width=12/
+| #L1 #L2 #W #V #l #H1W #HV #HVW #H2W #H1l #H2l #HL12 #J #K1 #X #H destruct /3 width=13/
]
qed-.
lemma lsubsv_inv_pair1: ∀h,g,I,G,K1,L2,X. G ⊢ K1.ⓑ{I}X ¡⊑[h, g] L2 →
(∃∃K2. G ⊢ K1 ¡⊑[h, g] K2 & L2 = K2.ⓑ{I}X) ∨
- ∃∃K2,W,V,W1,V2,l. ⦃G, K1⦄ ⊢ X ¡[h, g] & ⦃G, K2⦄ ⊢ W ¡[h, g] &
- ⦃G, K1⦄ ⊢ V •[h, g] ⦃l+1, W1⦄ & ⦃G, K2⦄ ⊢ W •[h, g] ⦃l, V2⦄ &
- G ⊢ K1 ¡⊑[h, g] K2 &
- I = Abbr & L2 = K2.ⓛW & X = ⓝW.V.
+ ∃∃K2,W,V,l. ⦃G, K1⦄ ⊢ W ¡[h, g] & ⦃G, K1⦄ ⊢ V ¡[h, g] &
+ scast h g l G K1 V W & ⦃G, K2⦄ ⊢ W ¡[h, g] &
+ ⦃G, K1⦄ ⊢ V ▪[h, g] l+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] l &
+ G ⊢ K1 ¡⊑[h, g] K2 &
+ I = Abbr & L2 = K2.ⓛW & X = ⓝW.V.
/2 width=3 by lsubsv_inv_pair1_aux/ qed-.
fact lsubsv_inv_atom2_aux: ∀h,g,G,L1,L2. G ⊢ L1 ¡⊑[h, g] L2 → L2 = ⋆ → L1 = ⋆.
#h #g #G #L1 #L2 * -L1 -L2
[ //
| #I #L1 #L2 #V #_ #H destruct
-| #L1 #L2 #W #V #V1 #V2 #l #_ #_ #_ #_ #_ #H destruct
+| #L1 #L2 #W #V #l #_ #_ #_ #_ #_ #_ #_ #H destruct
]
qed-.
fact lsubsv_inv_pair2_aux: ∀h,g,G,L1,L2. G ⊢ L1 ¡⊑[h, g] L2 →
∀I,K2,W. L2 = K2.ⓑ{I}W →
(∃∃K1. G ⊢ K1 ¡⊑[h, g] K2 & L1 = K1.ⓑ{I}W) ∨
- ∃∃K1,V,W1,V2,l. ⦃G, K1⦄ ⊢ ⓝW.V ¡[h, g] & ⦃G, K2⦄ ⊢ W ¡[h, g] &
- ⦃G, K1⦄ ⊢ V •[h, g] ⦃l+1, W1⦄ & ⦃G, K2⦄ ⊢ W •[h, g] ⦃l, V2⦄ &
- G ⊢ K1 ¡⊑[h, g] K2 & I = Abst & L1 = K1. ⓓⓝW.V.
+ ∃∃K1,V,l. ⦃G, K1⦄ ⊢ W ¡[h, g] & ⦃G, K1⦄ ⊢ V ¡[h, g] &
+ scast h g l G K1 V W & ⦃G, K2⦄ ⊢ W ¡[h, g] &
+ ⦃G, K1⦄ ⊢ V ▪[h, g] l+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] l &
+ G ⊢ K1 ¡⊑[h, g] K2 & I = Abst & L1 = K1. ⓓⓝW.V.
#h #g #G #L1 #L2 * -L1 -L2
[ #J #K2 #U #H destruct
| #I #L1 #L2 #V #HL12 #J #K2 #U #H destruct /3 width=3/
-| #L1 #L2 #W #V #W1 #V2 #l #HV #HW #HW1 #HV2 #HL12 #J #K2 #U #H destruct /3 width=10/
+| #L1 #L2 #W #V #l #H1W #HV #HVW #H2W #H1l #H2l #HL12 #J #K2 #U #H destruct /3 width=10/
]
qed-.
lemma lsubsv_inv_pair2: ∀h,g,I,G,L1,K2,W. G ⊢ L1 ¡⊑[h, g] K2.ⓑ{I}W →
(∃∃K1. G ⊢ K1 ¡⊑[h, g] K2 & L1 = K1.ⓑ{I}W) ∨
- ∃∃K1,V,W1,V2,l. ⦃G, K1⦄ ⊢ ⓝW.V ¡[h, g] & ⦃G, K2⦄ ⊢ W ¡[h, g] &
- ⦃G, K1⦄ ⊢ V •[h, g] ⦃l+1, W1⦄ & ⦃G, K2⦄ ⊢ W •[h, g] ⦃l, V2⦄ &
- G ⊢ K1 ¡⊑[h, g] K2 & I = Abst & L1 = K1. ⓓⓝW.V.
+ ∃∃K1,V,l. ⦃G, K1⦄ ⊢ W ¡[h, g] & ⦃G, K1⦄ ⊢ V ¡[h, g] &
+ scast h g l G K1 V W & ⦃G, K2⦄ ⊢ W ¡[h, g] &
+ ⦃G, K1⦄ ⊢ V ▪[h, g] l+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] l &
+ G ⊢ K1 ¡⊑[h, g] K2 & I = Abst & L1 = K1. ⓓⓝW.V.
/2 width=3 by lsubsv_inv_pair2_aux/ qed-.
(* Basic_forward lemmas *****************************************************)
(* *)
(**************************************************************************)
-include "basic_2/dynamic/snv_cpcs.ma".
-include "basic_2/dynamic/lsubsv_ssta.ma".
+include "basic_2/dynamic/lsubsv_lsstas.ma".
(* LOCAL ENVIRONMENT REFINEMENT FOR STRATIFIED NATIVE VALIDITY **************)
-(* Properties for the preservation results **********************************)
+(* Properties on decomposed extended parallel computation on terms **********)
-fact lsubsv_sstas_aux: ∀h,g,G0,L0,T0.
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_ssta_cpr_lpr h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_ssta h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_lsubsv h g G1 L1 T1) →
- ∀G,L2,T. ⦃G0, L0, T0⦄ >[h, g] ⦃G, L2, T⦄ → ⦃G, L2⦄ ⊢ T ¡[h, g] →
- ∀L1. G ⊢ L1 ¡⊑[h, g] L2 → ∀U2. ⦃G, L2⦄ ⊢ T •*[h, g] U2 →
- ∃∃U1. ⦃G, L1⦄ ⊢ T •*[h, g] U1 & ⦃G, L1⦄ ⊢ U1 ⬌* U2.
-#h #g #G0 #L0 #T0 #IH4 #IH3 #IH2 #IH1 #G #L2 #T #HLT0 #HT #L1 #HL12 #U2 #H @(sstas_ind … H) -U2 [ /2 width=3/ ]
-#U2 #W #l #HTU2 #HU2W * #U1 #HTU1 #HU12
-lapply (IH1 … HT … HL12) // #H
-lapply (snv_sstas_aux … IH2 … HTU1) // /3 width=5 by ygt_yprs_trans, lsubsv_yprs/ -H #HU1
-lapply (snv_sstas_aux … IH2 … HTU2) // #H
-lapply (IH1 … H … HL12) [ /3 width=5 by ygt_yprs_trans, sstas_yprs/ ] -H #HU2
-elim (snv_fwd_ssta … HU1) #l1 #W1 #HUW1
-elim (lsubsv_ssta_trans … HU2W … HL12) -HU2W #W2 #HUW2 #HW2
-elim (ssta_cpcs_lpr_aux … IH4 IH3 … HU1 HU2 … HUW1 … HUW2 … HU12 L1) -HU1 -HU2 -HUW2 -HU12 //
-[2,3: /4 width=5 by ygt_yprs_trans, sstas_yprs, lsubsv_yprs/ ] -G0 -L2 -L0 -T0 -U2 #H #HW12 destruct
-lapply (cpcs_trans … HW12 … HW2) -W2 /3 width=4/
-qed-.
-
-fact lsubsv_cpds_aux: ∀h,g,G0,L0,T0.
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_ssta_cpr_lpr h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_ssta h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_lsubsv h g G1 L1 T1) →
- ∀G,L2,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G,L2, T1⦄ → ⦃G, L2⦄ ⊢ T1 ¡[h, g] →
- ∀L1. G ⊢ L1 ¡⊑[h, g] L2 → ∀T2. ⦃G, L2⦄ ⊢ T1 •*➡*[h, g] T2 →
- ∃∃T. ⦃G, L1⦄ ⊢ T1 •*➡*[h, g] T & ⦃G, L1⦄ ⊢ T2 ➡* T.
-#h #g #G0 #L0 #T0 #IH4 #IH3 #IH2 #IH1 #G #L2 #T1 #HLT0 #HT1 #L1 #HL12 #T2 * #T #HT1T #HTT2
-lapply (lsubsv_cprs_trans … HL12 … HTT2) -HTT2 #HTT2
-elim (lsubsv_sstas_aux … IH4 IH3 IH2 IH1 … HLT0 … HL12 … HT1T) // -G0 -L2 -L0 -T0 #T0 #HT10 #HT0
-lapply (cpcs_cprs_strap1 … HT0 … HTT2) -T #HT02
-elim (cpcs_inv_cprs … HT02) -HT02 /3 width=3/
+lemma lsubsv_cpds_trans: ∀h,g,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 •*➡*[h, g] T2 →
+ ∀L1. G ⊢ L1 ¡⊑[h, g] L2 →
+ ∃∃T. ⦃G, L1⦄ ⊢ T1 •*➡*[h, g] T & ⦃G, L1⦄ ⊢ T2 ➡* T.
+#h #g #G #L2 #T1 #T2 * #T #l1 #l2 #Hl21 #Hl1 #HT1 #HT2 #L1 #HL12
+lapply (lsubsv_cprs_trans … HL12 … HT2) -HT2 #HT2
+elim (lsubsv_lsstas_trans … HT1 … Hl1 … HL12) // #T0 #HT10 #HT0
+lapply (lsubsv_fwd_lsubd … HL12) -HL12 #HL12
+lapply (lsubd_da_trans … Hl1 … HL12) -L2 #Hl1
+lapply (cpcs_cprs_strap1 … HT0 … HT2) -T #HT02
+elim (cpcs_inv_cprs … HT02) -HT02 /3 width=7/
qed-.
<(ldrop_inv_O2 … H) in HL12; -H /3 width=3/
| elim (IHL12 … HLK1) -L1 /3 width=3/
]
-| #L1 #L2 #W #V #W1 #V2 #l #HV #HW #HW1 #HV2 #_ #IHL12 #K1 #e #H
+| #L1 #L2 #W #V #l #H1W #HV #HVW #H2W #H1l #H2l #_ #IHL12 #K1 #e #H
elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK1
[ destruct
elim (IHL12 L1 0) -IHL12 // #X #HL12 #H
- <(ldrop_inv_O2 … H) in HL12; -H /3 width=6/
+ <(ldrop_inv_O2 … H) in HL12; -H /3 width=4/
| elim (IHL12 … HLK1) -L1 /3 width=3/
]
]
<(ldrop_inv_O2 … H) in HL12; -H /3 width=3/
| elim (IHL12 … HLK2) -L2 /3 width=3/
]
-| #L1 #L2 #W #V #W1 #V2 #l #HV #HW #HW1 #HV2 #_ #IHL12 #K2 #e #H
+| #L1 #L2 #W #V #l #H1W #HV #HVW #H2W #H1l #H2l #_ #IHL12 #K2 #e #H
elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK2
[ destruct
elim (IHL12 L2 0) -IHL12 // #X #HL12 #H
- <(ldrop_inv_O2 … H) in HL12; -H /3 width=6/
+ <(ldrop_inv_O2 … H) in HL12; -H /3 width=4/
| elim (IHL12 … HLK2) -L2 /3 width=3/
]
]
--- /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/lsubd_da.ma".
+include "basic_2/unfold/lsstas_alt.ma".
+include "basic_2/equivalence/cpcs_cpcs.ma".
+include "basic_2/dynamic/lsubsv_ldrop.ma".
+include "basic_2/dynamic/lsubsv_lsubd.ma".
+
+(* LOCAL ENVIRONMENT REFINEMENT FOR STRATIFIED NATIVE VALIDITY **************)
+
+(* Properties on nat-iterated stratified static type assignment *************)
+
+lemma lsubsv_lsstas_trans: ∀h,g,G,L2,T,U2,l1. ⦃G, L2⦄ ⊢ T •*[h, g, l1] U2 →
+ ∀l2. l1 ≤ l2 → ⦃G, L2⦄ ⊢ T ▪[h, g] l2 →
+ ∀L1. G ⊢ L1 ¡⊑[h, g] L2 →
+ ∃∃U1. ⦃G, L1⦄ ⊢ T •*[h, g, l1] U1 & ⦃G, L1⦄ ⊢ U1 ⬌* U2.
+#h #g #G #L2 #T #U #l1 #H @(lsstas_ind_alt … H) -G -L2 -T -U -l1
+[1,2: /2 width=3/
+| #G #L2 #K2 #X #Y #U #i #l1 #HLK2 #_ #HYU #IHXY #l2 #Hl12 #Hl2 #L1 #HL12
+ elim (da_inv_lref … Hl2) -Hl2 * #K0 #V0 [| #l0 ] #HK0 #HV0
+ lapply (ldrop_mono … HK0 … HLK2) -HK0 #H destruct
+ elim (lsubsv_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
+ elim (lsubsv_inv_pair2 … H) -H * #K1 [ | -HYU -IHXY -HLK1 ]
+ [ #HK12 #H destruct
+ elim (IHXY … Hl12 HV0 … HK12) -K2 -l2 #T #HXT #HTY
+ lapply (ldrop_fwd_ldrop2 … HLK1) #H
+ elim (lift_total T 0 (i+1))
+ /3 width=11 by lsstas_ldef, cpcs_lift, ex2_intro/
+ | #V #l0 #_ #_ #_ #_ #_ #_ #_ #H destruct
+ ]
+| #G #L2 #K2 #X #Y #U #i #l1 #l #HLK2 #_ #_ #HYU #IHXY #l2 #Hl12 #Hl2 #L1 #HL12 -l
+ elim (da_inv_lref … Hl2) -Hl2 * #K0 #V0 [| #l0 ] #HK0 #HV0 [| #H1 ]
+ lapply (ldrop_mono … HK0 … HLK2) -HK0 #H2 destruct
+ lapply (le_plus_to_le_r … Hl12) -Hl12 #Hl12
+ elim (lsubsv_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
+ elim (lsubsv_inv_pair2 … H) -H * #K1 [| ]
+ [ #HK12 #H destruct
+ lapply (lsubsv_fwd_lsubd … HK12) #H
+ lapply (lsubd_da_trans … HV0 … H) -H
+ elim (IHXY … Hl12 HV0 … HK12) -K2 -Hl12 #Y0
+ lapply (ldrop_fwd_ldrop2 … HLK1)
+ elim (lift_total Y0 0 (i+1))
+ /3 width=11 by lsstas_ldec, cpcs_lift, ex2_intro/
+ | #V #l #_ #_ #HVX #_ #HV #HX #HK12 #_ #H destruct
+ lapply (da_mono … HX … HV0) -HX #H destruct
+ elim (IHXY … Hl12 HV0 … HK12) -K2 #Y0 #HXY0 #HY0
+ elim (da_ssta … HV) -HV #W #HVW
+ elim (lsstas_total … HVW (l1+1)) -W #W #HVW
+ lapply (HVX … Hl12 HVW HXY0) -HVX -Hl12 -HXY0 #HWY0
+ lapply (cpcs_trans … HWY0 … HY0) -Y0
+ lapply (ldrop_fwd_ldrop2 … HLK1)
+ elim (lift_total W 0 (i+1))
+ /4 width=11 by lsstas_ldef, lsstas_cast, cpcs_lift, ex2_intro/
+ ]
+| #a #I #G #L2 #V2 #T2 #U2 #l1 #_ #IHTU2 #l2 #Hl12 #Hl2 #L1 #HL12
+ lapply (da_inv_bind … Hl2) -Hl2 #Hl2
+ elim (IHTU2 … Hl2 (L1.ⓑ{I}V2) …) // [2: /2 width=1/ ] -L2
+ /3 width=3 by lsstas_bind, cpcs_bind_dx, ex2_intro/
+| #G #L2 #V2 #T2 #U2 #l1 #_ #IHTU2 #l2 #Hl12 #Hl2 #L1 #HL12
+ lapply (da_inv_flat … Hl2) -Hl2 #Hl2
+ elim (IHTU2 … Hl2 … HL12) -L2 //
+ /3 width=5 by lsstas_appl, cpcs_flat, ex2_intro/
+| #G #L2 #W2 #T2 #U2 #l1 #_ #IHTU2 #l2 #Hl12 #Hl2 #L1 #HL12
+ lapply (da_inv_flat … Hl2) -Hl2 #Hl2
+ elim (IHTU2 … Hl2 … HL12) -L2 //
+ /3 width=3 by lsstas_cast, ex2_intro/
+]
+qed-.
+
+lemma lsubsv_ssta_trans: ∀h,g,G,L2,T,U2. ⦃G, L2⦄ ⊢ T •[h, g] U2 →
+ ∀l. ⦃G, L2⦄ ⊢ T ▪[h, g] l+1 →
+ ∀L1. G ⊢ L1 ¡⊑[h, g] L2 →
+ ∃∃U1. ⦃G, L1⦄ ⊢ T •[h, g] U1 & ⦃G, L1⦄ ⊢ U1 ⬌* U2.
+#h #g #G #L2 #T #U2 #H #l #HTl #L1 #HL12
+elim ( lsubsv_lsstas_trans … U2 1 … HTl … HL12)
+/3 width=3 by lsstas_inv_SO, ssta_lsstas, ex2_intro/
+qed-.
lemma lsubsv_fwd_lsuba: ∀h,g,G,L1,L2. G ⊢ L1 ¡⊑[h, g] L2 → G ⊢ L1 ⁝⊑ L2.
#h #g #G #L1 #L2 #H elim H -L1 -L2 // /2 width=1/
-#L1 #L2 #W #V #W1 #V2 #l #HV #HW #_ #_ #_ #IHL12 -W1 -V2
+#L1 #L2 #W #V #l #H1W #HV #HVW #H2W #H1l #_ #_ #IHL12
+lapply (snv_scast … HV H1W HVW H1l) -HV -H1W -HVW -H1l #HV
elim (snv_fwd_aaa … HV) -HV #A #HV
-elim (snv_fwd_aaa … HW) -HW #B #HW
+elim (snv_fwd_aaa … H2W) -H2W #B #HW
elim (aaa_inv_cast … HV) #HWA #_
lapply (lsuba_aaa_trans … HW … IHL12) #HWB
lapply (aaa_mono … HWB … HWA) -HWB -HWA #H destruct /2 width=3/
--- /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/lsubd.ma".
+include "basic_2/dynamic/lsubsv.ma".
+
+(* LOCAL ENVIRONMENT REFINEMENT FOR STRATIFIED NATIVE VALIDITY **************)
+
+(* Forward lemmas on lenv refinement for degree assignment ******************)
+
+lemma lsubsv_fwd_lsubd: ∀h,g,G,L1,L2. G ⊢ L1 ¡⊑[h, g] L2 → G ⊢ L1 ▪⊑[h, g] L2.
+#h #g #G #L1 #L2 #H elim H -L1 -L2 // /2 width=1/ /2 width=3/
+qed-.
(**************************************************************************)
include "basic_2/computation/cpds_cpds.ma".
+include "basic_2/dynamic/snv_aaa.ma".
include "basic_2/dynamic/lsubsv_cpds.ma".
include "basic_2/dynamic/lsubsv_cpcs.ma".
(* Properties concerning stratified native validity *************************)
-fact snv_lsubsv_aux: ∀h,g,G0,L0,T0.
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_ssta_cpr_lpr h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_ssta h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_lsubsv h g G1 L1 T1) →
- ∀G1,L1,T1. G0 = G1 → L0 = L1 → T0 = T1 → IH_snv_lsubsv h g G1 L1 T1.
-#h #g #G0 #L0 #T0 #IH4 #IH3 #IH2 #IH1 #G #L2 * * [||||*] //
-[ #i #HG0 #HL0 #HT0 #H #L1 #HL12 destruct -IH4 -IH3 -IH2
- elim (snv_inv_lref … H) -H #I2 #K2 #W2 #HLK2 #HW2
- elim (lsubsv_ldrop_O1_trans … HL12 … HLK2) -HL12 #X #H #HLK1
+lemma lsubsv_snv_trans: ∀h,g,G,L2,T. ⦃G, L2⦄ ⊢ T ¡[h, g] →
+ ∀L1. G ⊢ L1 ¡⊑[h, g] L2 → ⦃G, L1⦄ ⊢ T ¡[h, g].
+#h #g #G #L2 #T #H elim H -G -L2 -T //
+[ #I #G #L2 #K2 #V #i #HLK2 #_ #IHV #L1 #HL12
+ elim (lsubsv_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
elim (lsubsv_inv_pair2 … H) -H * #K1
- [ #HK12 #H destruct
- /5 width=8 by snv_lref, fsupp_ygt, fsupp_lref/ (**) (* auto too slow without trace *)
- | #V #W1 #V2 #l #HV #_ #_ #_ #_ #_ #H destruct /2 width=5/
+ [ #HK12 #H destruct /3 width=5 by snv_lref/
+ | #W #l #H1V #H1W #HWV #_ #HWl #_ #_ #H1 #H2 destruct -IHV
+ /3 width=10 by snv_scast, snv_lref/
]
-| #p #HG0 #HL0 #HT0 #H #L1 #HL12 destruct -IH4 -IH3 -IH2 -IH1
- elim (snv_inv_gref … H)
-| #a #I #V #T #HG0 #HL0 #HT0 #H #L1 #HL12 destruct -IH4 -IH3 -IH2
- elim (snv_inv_bind … H) -H /4 width=4/
-| #V #T #HG0 #HL0 #HT0 #H #L1 #HL12 destruct
- elim (snv_inv_appl … H) -H #a #W #W0 #U #l #HV #HT #HVW #HW0 #HTU
+| #a #I #G #L2 #V #T #_ #_ #IHV #IHT #L1 #HL12 destruct
+ /4 width=1 by snv_bind, lsubsv_pair/
+| #a #G #L2 #V #W #W0 #T #U #l #_ #_ #HVl #HVW #HW0 #HTU #IHV #IHT #L1 #HL12
lapply (lsubsv_cprs_trans … HL12 … HW0) -HW0 #HW0
- elim (lsubsv_ssta_trans … HVW … HL12) -HVW #W1 #HVW1 #HW1
+ elim (lsubsv_ssta_trans … HVW … HVl … HL12) -HVW #W1 #HVW1 #HW1
lapply (cpcs_cprs_strap1 … HW1 … HW0) -W #HW10
- elim (lsubsv_cpds_aux … IH4 IH3 IH2 IH1 … HL12 … HTU) -IH4 -IH3 -IH2 -HTU // /2 width=1/ #X #HTU #H
- elim (cprs_inv_abst1 … H) -H #W #U2 #HW0 #HU2 #H destruct
+ lapply (lsubd_da_trans … HVl L1 ?) -HVl /2 width=1 by lsubsv_fwd_lsubd/ #HVl
+ elim (lsubsv_cpds_trans … HTU … HL12) -HTU #X #HTU #H
+ elim (cprs_inv_abst1 … H) -H #W #U2 #HW0 #_ #H destruct
lapply (cpcs_cprs_strap1 … HW10 … HW0) -W0 #H
elim (cpcs_inv_cprs … H) -H #W0 #HW10 #HW0
- lapply (cpds_cprs_trans … (ⓛ{a}W0.U2) HTU ?) [ /2 width=1/ ] -HTU -HW0
- /4 width=8 by snv_appl, fsupp_ygt/ (**) (* auto too slow without trace *)
-| #W #T #HG0 #HL0 #HT0 #H #L1 #HL12 destruct -IH4 -IH3 -IH2
- elim (snv_inv_cast … H) -H #U #l #HW #HT #HTU #HUW
+ /4 width=11 by snv_appl, cpds_cprs_trans, cprs_bind/
+| #G #L2 #W #T #U #l #_ #_ #HTl #HTU #HUW #IHW #IHT #L1 #HL12
lapply (lsubsv_cpcs_trans … HL12 … HUW) -HUW #HUW
- elim (lsubsv_ssta_trans … HTU … HL12) -HTU #U0 #HTU0 #HU0
- lapply (cpcs_trans … HU0 … HUW) -U /4 width=4 by snv_cast, fsupp_ygt/ (**) (* auto too slow without trace *)
+ elim (lsubsv_ssta_trans … HTU … HTl … HL12) -HTU #U0 #HTU0 #HU0
+ lapply (lsubd_da_trans … HTl L1 ?) -HTl
+ /4 width=5 by lsubsv_fwd_lsubd, snv_cast, cpcs_trans/
]
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/ssta_ssta.ma".
-include "basic_2/equivalence/cpcs_cpcs.ma".
-include "basic_2/dynamic/lsubsv_ldrop.ma".
-
-(* LOCAL ENVIRONMENT REFINEMENT FOR STRATIFIED NATIVE VALIDITY **************)
-
-(* Properties on stratified static type assignment **************************)
-
-lemma lsubsv_ssta_trans: ∀h,g,G,L2,T,U2,l. ⦃G, L2⦄ ⊢ T •[h, g] ⦃l, U2⦄ →
- ∀L1. G ⊢ L1 ¡⊑[h, g] L2 →
- ∃∃U1. ⦃G, L1⦄ ⊢ T •[h, g] ⦃l, U1⦄ & ⦃G, L1⦄ ⊢ U1 ⬌* U2.
-#h #g #G #L2 #T #U #l #H elim H -G -L2 -T -U -l
-[ /3 width=3/
-| #G #L2 #K2 #X #Y #U #i #l #HLK2 #_ #HYU #IHXY #L1 #HL12
- elim (lsubsv_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
- elim (lsubsv_inv_pair2 … H) -H * #K1 [ | -HYU -IHXY -HLK1 ]
- [ #HK12 #H destruct
- elim (IHXY … HK12) -K2 #T #HXT #HTY
- lapply (ldrop_fwd_ldrop2 … HLK1) #H
- elim (lift_total T 0 (i+1)) /3 width=11/
- | #V #W1 #V2 #l0 #_ #_ #_ #_ #_ #H destruct
- ]
-| #G #L2 #K2 #Y #X #U #i #l #HLK2 #HYX #HYU #IHYX #L1 #HL12
- elim (lsubsv_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
- elim (lsubsv_inv_pair2 … H) -H * #K1 [ -HYX | -IHYX ]
- [ #HK12 #H destruct
- elim (IHYX … HK12) -K2 /3 width=6/
- | #V #W1 #V2 #l0 #HYV #_ #HV #HY #_ #_ #H destruct
- elim (snv_inv_cast … HYV) -HYV #W #l1 #_ #_ #HVW #HWY
- elim (ssta_mono … HVW … HV) -HV #Hl10 #_
- elim (ssta_mono … HYX … HY) -HYX -HY #H #_ destruct
- lapply (ldrop_fwd_ldrop2 … HLK1) #H
- elim (lift_total W 0 (i+1))
- /4 width=11 by cpcs_lift, ex2_intro, ssta_ldef, ssta_cast/
- ]
-| #a #I #G #L2 #V2 #T2 #U2 #l #_ #IHTU2 #L1 #HL12
- elim (IHTU2 (L1.ⓑ{I}V2) …) [2: /2 width=1/ ] -L2 /3 width=3/
-| #G #L2 #V2 #T2 #U2 #l #_ #IHTU2 #L1 #HL12
- elim (IHTU2 … HL12) -L2 /3 width=5/
-| #G #L2 #W2 #T2 #U2 #l #_ #IHTU2 #L1 #HL12
- elim (IHTU2 … HL12) -L2 /3 width=3/
-]
-qed-.
(* STRATIFIED NATIVE VALIDITY FOR TERMS *************************************)
+definition scast: ∀h. sd h → nat → relation4 genv lenv term term ≝
+ λh,g,l,G,L,V,W. ∀V0,W0,l0.
+ l0 ≤ l → ⦃G, L⦄ ⊢ V •*[h, g, l0+1] V0 → ⦃G, L⦄ ⊢ W •*[h, g, l0] W0 → ⦃G, L⦄ ⊢ V0 ⬌* W0.
+
(* activate genv *)
inductive snv (h:sh) (g:sd h): relation3 genv lenv term ≝
| snv_sort: ∀G,L,k. snv h g G L (⋆k)
| snv_lref: ∀I,G,L,K,V,i. ⇩[0, i] L ≡ K.ⓑ{I}V → snv h g G K V → snv h g G L (#i)
| snv_bind: ∀a,I,G,L,V,T. snv h g G L V → snv h g G (L.ⓑ{I}V) T → snv h g G L (ⓑ{a,I}V.T)
| snv_appl: ∀a,G,L,V,W,W0,T,U,l. snv h g G L V → snv h g G L T →
- â¦\83G, Lâ¦\84 â\8a¢ V â\80¢[h, g] â¦\83l+1, Wâ¦\84 → ⦃G, L⦄ ⊢ W ➡* W0 →
+ â¦\83G, Lâ¦\84 â\8a¢ V â\96ª[h, g] l+1 â\86\92 â¦\83G, Lâ¦\84 â\8a¢ V â\80¢[h, g] W → ⦃G, L⦄ ⊢ W ➡* W0 →
⦃G, L⦄ ⊢ T •*➡*[h, g] ⓛ{a}W0.U → snv h g G L (ⓐV.T)
| snv_cast: ∀G,L,W,T,U,l. snv h g G L W → snv h g G L T →
- â¦\83G, Lâ¦\84 â\8a¢ T â\80¢[h, g] â¦\83l+1, Uâ¦\84 → ⦃G, L⦄ ⊢ U ⬌* W → snv h g G L (ⓝW.T)
+ â¦\83G, Lâ¦\84 â\8a¢ T â\96ª[h, g] l+1 â\86\92 â¦\83G, Lâ¦\84 â\8a¢ T â\80¢[h, g] U → ⦃G, L⦄ ⊢ U ⬌* W → snv h g G L (ⓝW.T)
.
interpretation "stratified native validity (term)"
[ #G #L #k #i #H destruct
| #I #G #L #K #V #i0 #HLK #HV #i #H destruct /2 width=5/
| #a #I #G #L #V #T #_ #_ #i #H destruct
-| #a #G #L #V #W #W0 #T #U #l #_ #_ #_ #_ #_ #i #H destruct
-| #G #L #W #T #U #l #_ #_ #_ #_ #i #H destruct
+| #a #G #L #V #W #W0 #T #U #l #_ #_ #_ #_ #_ #_ #i #H destruct
+| #G #L #W #T #U #l #_ #_ #_ #_ #_ #i #H destruct
]
qed-.
[ #G #L #k #p #H destruct
| #I #G #L #K #V #i #_ #_ #p #H destruct
| #a #I #G #L #V #T #_ #_ #p #H destruct
-| #a #G #L #V #W #W0 #T #U #l #_ #_ #_ #_ #_ #p #H destruct
-| #G #L #W #T #U #l #_ #_ #_ #_ #p #H destruct
+| #a #G #L #V #W #W0 #T #U #l #_ #_ #_ #_ #_ #_ #p #H destruct
+| #G #L #W #T #U #l #_ #_ #_ #_ #_ #p #H destruct
]
qed-.
[ #G #L #k #a #I #V #T #H destruct
| #I0 #G #L #K #V0 #i #_ #_ #a #I #V #T #H destruct
| #b #I0 #G #L #V0 #T0 #HV0 #HT0 #a #I #V #T #H destruct /2 width=1/
-| #b #G #L #V0 #W0 #W00 #T0 #U0 #l #_ #_ #_ #_ #_ #a #I #V #T #H destruct
-| #G #L #W0 #T0 #U0 #l #_ #_ #_ #_ #a #I #V #T #H destruct
+| #b #G #L #V0 #W0 #W00 #T0 #U0 #l #_ #_ #_ #_#_ #_ #a #I #V #T #H destruct
+| #G #L #W0 #T0 #U0 #l #_ #_ #_ #_ #_ #a #I #V #T #H destruct
]
qed-.
fact snv_inv_appl_aux: ∀h,g,G,L,X. ⦃G, L⦄ ⊢ X ¡[h, g] → ∀V,T. X = ⓐV.T →
∃∃a,W,W0,U,l. ⦃G, L⦄ ⊢ V ¡[h, g] & ⦃G, L⦄ ⊢ T ¡[h, g] &
- â¦\83G, Lâ¦\84 â\8a¢ V â\80¢[h, g] â¦\83l+1, Wâ¦\84 & ⦃G, L⦄ ⊢ W ➡* W0 &
+ â¦\83G, Lâ¦\84 â\8a¢ V â\96ª[h, g] l+1 & â¦\83G, Lâ¦\84 â\8a¢ V â\80¢[h, g] W & ⦃G, L⦄ ⊢ W ➡* W0 &
⦃G, L⦄ ⊢ T •*➡*[h, g] ⓛ{a}W0.U.
#h #g #G #L #X * -L -X
[ #G #L #k #V #T #H destruct
| #I #G #L #K #V0 #i #_ #_ #V #T #H destruct
| #a #I #G #L #V0 #T0 #_ #_ #V #T #H destruct
-| #a #G #L #V0 #W0 #W00 #T0 #U0 #l #HV0 #HT0 #HVW0 #HW00 #HTU0 #V #T #H destruct /2 width=8/
-| #G #L #W0 #T0 #U0 #l #_ #_ #_ #_ #V #T #H destruct
+| #a #G #L #V0 #W0 #W00 #T0 #U0 #l #HV0 #HT0 #Hl #HVW0 #HW00 #HTU0 #V #T #H destruct /2 width=8/
+| #G #L #W0 #T0 #U0 #l #_ #_ #_ #_ #_ #V #T #H destruct
]
qed-.
lemma snv_inv_appl: ∀h,g,G,L,V,T. ⦃G, L⦄ ⊢ ⓐV.T ¡[h, g] →
∃∃a,W,W0,U,l. ⦃G, L⦄ ⊢ V ¡[h, g] & ⦃G, L⦄ ⊢ T ¡[h, g] &
- â¦\83G, Lâ¦\84 â\8a¢ V â\80¢[h, g] â¦\83l+1, Wâ¦\84 & ⦃G, L⦄ ⊢ W ➡* W0 &
+ â¦\83G, Lâ¦\84 â\8a¢ V â\96ª[h, g] l+1 & â¦\83G, Lâ¦\84 â\8a¢ V â\80¢[h, g] W & ⦃G, L⦄ ⊢ W ➡* W0 &
⦃G, L⦄ ⊢ T •*➡*[h, g] ⓛ{a}W0.U.
/2 width=3 by snv_inv_appl_aux/ qed-.
fact snv_inv_cast_aux: ∀h,g,G,L,X. ⦃G, L⦄ ⊢ X ¡[h, g] → ∀W,T. X = ⓝW.T →
∃∃U,l. ⦃G, L⦄ ⊢ W ¡[h, g] & ⦃G, L⦄ ⊢ T ¡[h, g] &
- â¦\83G, Lâ¦\84 â\8a¢ T â\80¢[h, g] â¦\83l+1, Uâ¦\84 & ⦃G, L⦄ ⊢ U ⬌* W.
+ â¦\83G, Lâ¦\84 â\8a¢ T â\96ª[h, g] l+1 & â¦\83G, Lâ¦\84 â\8a¢ T â\80¢[h, g] U & ⦃G, L⦄ ⊢ U ⬌* W.
#h #g #G #L #X * -G -L -X
[ #G #L #k #W #T #H destruct
| #I #G #L #K #V #i #_ #_ #W #T #H destruct
| #a #I #G #L #V #T0 #_ #_ #W #T #H destruct
-| #a #G #L #V #W0 #W00 #T0 #U #l #_ #_ #_ #_ #_ #W #T #H destruct
-| #G #L #W0 #T0 #U0 #l #HW0 #HT0 #HTU0 #HUW0 #W #T #H destruct /2 width=4/
+| #a #G #L #V #W0 #W00 #T0 #U #l #_ #_ #_ #_ #_ #_ #W #T #H destruct
+| #G #L #W0 #T0 #U0 #l #HW0 #HT0 #Hl #HTU0 #HUW0 #W #T #H destruct /2 width=4/
]
qed-.
lemma snv_inv_cast: ∀h,g,G,L,W,T. ⦃G, L⦄ ⊢ ⓝW.T ¡[h, g] →
∃∃U,l. ⦃G, L⦄ ⊢ W ¡[h, g] & ⦃G, L⦄ ⊢ T ¡[h, g] &
- â¦\83G, Lâ¦\84 â\8a¢ T â\80¢[h, g] â¦\83l+1, Uâ¦\84 & ⦃G, L⦄ ⊢ U ⬌* W.
+ â¦\83G, Lâ¦\84 â\8a¢ T â\96ª[h, g] l+1 & â¦\83G, Lâ¦\84 â\8a¢ T â\80¢[h, g] U & ⦃G, L⦄ ⊢ U ⬌* W.
/2 width=3 by snv_inv_cast_aux/ qed-.
-
-(* Basic forward lemmas *****************************************************)
-
-lemma snv_fwd_ssta: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, g] → ∃∃l,U. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄.
-#h #g #G #L #T #H elim H -G -L -T
-[ #G #L #k elim (deg_total h g k) /3 width=3/
-| * #G #L #K #V #i #HLK #_ * #l0 #W #HVW
- [ elim (lift_total W 0 (i+1)) /3 width=8/
- | elim (lift_total V 0 (i+1)) /3 width=8/
- ]
-| #a #I #G #L #V #T #_ #_ #_ * /3 width=3/
-| #a #G #L #V #W #W1 #T0 #T1 #l #_ #_ #_ #_ #_ #_ * /3 width=3/
-| #G #L #W #T #U #l #_ #_ #HTU #_ #_ #_ /3 width=3/ (**) (* auto fails without the last #_ *)
-]
-qed-.
[ /2 width=2/
| #I #G #L #K #V #i #HLK #_ * /3 width=6/
| #a * #G #L #V #T #_ #_ * #B #HV * #A #HA /3 width=2/
-| #a #G #L #V #W #W0 #T #U #l #_ #_ #HVW #HW0 #HTU * #B #HV * #X #HT
- lapply (cpds_aaa h g … HV W0 ?) [ -HTU /3 width=4/ ] -W #HW0 (**) (* auto fail without -HTU *)
- lapply (cpds_aaa … HT … HTU) -HTU #H
+| #a #G #L #V #W #W0 #T #U #l #_ #_ #Hl #HVW #HW0 #HTU * #B #HV * #X #HT
+ lapply (aaa_cpds_conf h g … HV W0 ?) [ -HTU /3 width=4/ ] -W #HW0 (**) (* auto fail without -HTU *)
+ lapply (aaa_cpds_conf … HT … HTU) -HTU #H
elim (aaa_inv_abst … H) -H #B0 #A #H1 #HU #H2 destruct
lapply (aaa_mono … H1 … HW0) -W0 #H destruct /3 width=4/
-| #G #L #W #T #U #l #_ #_ #HTU #HUW * #B #HW * #A #HT
- lapply (aaa_cpcs_mono … HUW A … HW) -HUW /2 width=7/ -HTU #H destruct /3 width=3/
+| #G #L #W #T #U #l #_ #_ #_ #HTU #HUW * #B #HW * #A #HT
+ lapply (aaa_ssta_conf … HT … HTU) -HTU #H
+ lapply (aaa_cpcs_mono … HUW … H … HW) -HUW -H #H destruct /3 width=3/
]
qed-.
lemma snv_fwd_csn: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, g] → ⦃G, L⦄ ⊢ ⬊*[h, g] T.
#h #g #G #L #T #H elim (snv_fwd_aaa … H) -H /2 width=2/
qed-.
+
+(* Advanced forward lemmas **************************************************)
+
+lemma snv_fwd_da: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, g] → ∃l. ⦃G, L⦄ ⊢ T ▪[h, g] l.
+#h #g #G #L #T #H elim (snv_fwd_aaa … H) -H /2 width=2 by aaa_fwd_da/
+qed-.
+
+lemma snv_fwd_ssta: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, g] → ∃U. ⦃G, L⦄ ⊢ T •[h, g] U.
+#h #g #G #L #T #H elim (snv_fwd_aaa … H) -H /2 width=2 by aaa_fwd_ssta/
+qed-.
+
+lemma snv_lsstas_fwd_correct: ∀h,g,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 ¡[h, g] → ⦃G, L⦄ ⊢ T1 •* [h, g, l] T2 →
+ ∃U2. ⦃G, L⦄ ⊢ T2 •[h, g] U2.
+#h #g #G #L #T1 #T2 #l #HT1 #HT12
+elim (snv_fwd_ssta … HT1) -HT1 /2 width=5 by lsstas_fwd_correct/
+qed-.
+
+(* Advanced properties ******************************************************)
+
+lemma snv_scast: ∀h,g,G,L,V,W,l. ⦃G, L⦄ ⊢ V ¡[h, g] → ⦃G, L⦄ ⊢ W ¡[h, g] →
+ scast h g l G L V W → ⦃G, L⦄ ⊢V ▪[h, g] l+1 → ⦃G, L⦄ ⊢ ⓝW.V ¡[h, g].
+#h #g #G #L #V #W #l #HV #HW #H #Hl
+elim (snv_fwd_ssta … HV) /4 width=6 by snv_cast, ssta_lsstas/
+qed-.
(* *)
(**************************************************************************)
-include "basic_2/unfold/sstas_sstas.ma".
-include "basic_2/equivalence/cpcs_cpcs.ma".
-include "basic_2/dynamic/snv_sstas.ma".
+include "basic_2/unfold/lsstas_lsstas.ma".
+include "basic_2/equivalence/cpes_cpds.ma".
+include "basic_2/dynamic/yprs_lift.ma".
include "basic_2/dynamic/ygt.ma".
(* STRATIFIED NATIVE VALIDITY FOR TERMS *************************************)
λh,g,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
∀T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L2⦄ ⊢ T2 ¡[h, g].
-definition IH_ssta_cpr_lpr: ∀h:sh. sd h → relation3 genv lenv term ≝
- λh,g,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀U1,l. ⦃G, L1⦄ ⊢ T1 •[h, g] ⦃l, U1⦄ →
- ∀T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
- ∃∃U2. ⦃G, L2⦄ ⊢ T2 •[h, g] ⦃l, U2⦄ & ⦃G, L2⦄ ⊢ U1 ⬌* U2.
+definition IH_da_cpr_lpr: ∀h:sh. sd h → relation3 genv lenv term ≝
+ λh,g,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
+ ∀l. ⦃G, L1⦄ ⊢ T1 ▪[h, g] l →
+ ∀T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
+ ⦃G, L2⦄ ⊢ T2 ▪[h, g] l.
-definition IH_snv_ssta: ∀h:sh. sd h → relation3 genv lenv term ≝
- λh,g,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, g] →
- ∀U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l+1, U⦄ → ⦃G, L⦄ ⊢ U ¡[h, g].
+definition IH_lsstas_cpr_lpr: ∀h:sh. sd h → relation3 genv lenv term ≝
+ λh,g,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
+ ∀l1,l2. l2 ≤ l1 → ⦃G, L1⦄ ⊢ T1 ▪[h, g] l1 →
+ ∀U1. ⦃G, L1⦄ ⊢ T1 •*[h, g, l2] U1 →
+ ∀T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
+ ∃∃U2. ⦃G, L2⦄ ⊢ T2 •*[h, g, l2] U2 & ⦃G, L2⦄ ⊢ U1 ⬌* U2.
-definition IH_snv_lsubsv: ∀h:sh. sd h → relation3 genv lenv term ≝
- λh,g,G,L2,T. ⦃G, L2⦄ ⊢ T ¡[h, g] →
- ∀L1. G ⊢ L1 ¡⊑[h, g] L2 → ⦃G, L1⦄ ⊢ T ¡[h, g].
+definition IH_snv_lsstas: ∀h:sh. sd h → relation3 genv lenv term ≝
+ λh,g,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, g] →
+ ∀l1,l2. l2 ≤ l1 → ⦃G, L⦄ ⊢ T ▪[h, g] l1 →
+ ∀U. ⦃G, L⦄ ⊢ T •*[h, g, l2] U → ⦃G, L⦄ ⊢ U ¡[h, g].
(* Properties for the preservation results **********************************)
∀G,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
∀T2. ⦃G, L1⦄ ⊢ T1 ➡* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L2⦄ ⊢ T2 ¡[h, g].
#h #g #G0 #L0 #T0 #IH #G #L1 #T1 #HLT0 #HT1 #T2 #H
-elim H -T2 [ /2 width=6/ ] -HT1
-/4 width=6 by ygt_yprs_trans, cprs_yprs/
+@(cprs_ind … H) -T2 /4 width=6 by ygt_yprs_trans, cprs_yprs/
qed-.
-fact ssta_cprs_lpr_aux: ∀h,g,G0,L0,T0.
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_ssta_cpr_lpr h g G1 L1 T1) →
- ∀G,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀U1,l. ⦃G, L1⦄ ⊢ T1 •[h, g] ⦃l, U1⦄ →
- ∀T2. ⦃G, L1⦄ ⊢ T1 ➡* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
- ∃∃U2. ⦃G, L2⦄ ⊢ T2 •[h, g] ⦃l, U2⦄ & ⦃G, L2⦄ ⊢ U1 ⬌* U2.
-#h #g #G0 #L0 #T0 #IH2 #IH1 #G #L1 #T1 #H01 #HT1 #U1 #l #HTU1 #T2 #H
-elim H -T2 [ /2 width=7/ ]
-#T #T2 #HT1T #HTT2 #IHT1 #L2 #HL12
-elim (IHT1 L1) // -IHT1 #U #HTU #HU1
-elim (IH1 … HTU … HTT2 … HL12) -IH1 -HTU -HTT2
-[2: /3 width=10 by snv_cprs_lpr_aux/
-|3: /5 width=6 by ygt_yprs_trans, cprs_yprs/
-] -G0 -L0 -T0 -T1 -T #U2 #HTU2 #HU2
-lapply (lpr_cpcs_conf … HL12 … HU1) -L1 #HU1
-lapply (cpcs_trans … HU1 … HU2) -U /2 width=3/
+fact da_cprs_lpr_aux: ∀h,g,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h g G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h g G1 L1 T1) →
+ ∀G,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
+ ∀l. ⦃G, L1⦄ ⊢ T1 ▪[h, g] l →
+ ∀T2. ⦃G, L1⦄ ⊢ T1 ➡* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L2⦄ ⊢ T2 ▪[h, g] l.
+#h #g #G0 #L0 #T0 #IH2 #IH1 #G #L1 #T1 #HLT0 #HT1 #l #Hl #T2 #H
+@(cprs_ind … H) -T2 /4 width=10 by snv_cprs_lpr_aux, ygt_yprs_trans, cprs_yprs/
qed-.
-fact ssta_cpcs_lpr_aux: ∀h,g,G0,L0,T0.
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_ssta_cpr_lpr h g G1 L1 T1) →
- ∀G,L1,T1,T2. ⦃G0, L0, T0⦄ >[h, g] ⦃G, L1, T1⦄ → ⦃G0, L0, T0⦄ >[h, g] ⦃G, L1, T2⦄ →
- ⦃G, L1⦄ ⊢ T1 ¡[h, g] → ⦃G, L1⦄ ⊢ T2 ¡[h, g] →
- ∀U1,l1. ⦃G, L1⦄ ⊢ T1 •[h, g] ⦃l1, U1⦄ →
- ∀U2,l2. ⦃G, L1⦄ ⊢ T2 •[h, g] ⦃l2, U2⦄ →
- ⦃G, L1⦄ ⊢ T1 ⬌* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
- l1 = l2 ∧ ⦃G, L2⦄ ⊢ U1 ⬌* U2.
-#h #g #G0 #L0 #T0 #IH2 #IH1 #G #L1 #T1 #T2 #H01 #H02 #HT1 #HT2 #U1 #l1 #HTU1 #U2 #l2 #HTU2 #H #L2 #HL12
-elim (cpcs_inv_cprs … H) -H #T #H1 #H2
-elim (ssta_cprs_lpr_aux … H01 HT1 … HTU1 … H1 … HL12) -T1 /2 width=1/ #W1 #H1 #HUW1
-elim (ssta_cprs_lpr_aux … H02 HT2 … HTU2 … H2 … HL12) -T2 /2 width=1/ #W2 #H2 #HUW2 -L0 -T0
-elim (ssta_mono … H1 … H2) -h -T #H1 #H2 destruct
-lapply (cpcs_canc_dx … HUW1 … HUW2) -W2 /2 width=1/
+fact da_cpcs_aux: ∀h,g,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h g G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h g G1 L1 T1) →
+ ∀G,L,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G, L, T1⦄ → ⦃G, L⦄ ⊢ T1 ¡[h, g] →
+ ∀T2. ⦃G0, L0, T0⦄ >[h, g] ⦃G, L, T2⦄ → ⦃G, L⦄ ⊢ T2 ¡[h, g] →
+ ∀l1. ⦃G, L⦄ ⊢ T1 ▪[h, g] l1 → ∀l2. ⦃G, L⦄ ⊢ T2 ▪[h, g] l2 →
+ ⦃G, L⦄ ⊢ T1 ⬌* T2 → l1 = l2.
+#h #g #G0 #L0 #T0 #IH2 #IH1 #G #L #T1 #HLT01 #HT1 #T2 #HLT02 #HT2 #l1 #Hl1 #l2 #Hl2 #H
+elim (cpcs_inv_cprs … H) -H /4 width=18 by da_cprs_lpr_aux, da_mono/
qed-.
-fact snv_sstas_aux: ∀h,g,G0,L0,T0.
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_ssta h g G1 L1 T1) →
- ∀G,L,T. ⦃G0, L0, T0⦄ >[h, g] ⦃G, L, T⦄ → ⦃G, L⦄ ⊢ T ¡[h, g] →
- ∀U. ⦃G, L⦄ ⊢ T •*[h, g] U → ⦃G, L⦄ ⊢ U ¡[h, g].
-#h #g #G0 #L0 #T0 #IH #G #L #T #H01 #HT #U #H
-@(sstas_ind … H) -U // -HT /4 width=5 by ygt_yprs_trans, sstas_yprs/
+fact ssta_cpr_lpr_aux: ∀h,g,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_lsstas_cpr_lpr h g G1 L1 T1) →
+ ∀G,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
+ ∀l. ⦃G, L1⦄ ⊢ T1 ▪[h, g] l+1 →
+ ∀U1. ⦃G, L1⦄ ⊢ T1 •[h, g] U1 →
+ ∀T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
+ ∃∃U2. ⦃G, L2⦄ ⊢ T2 •[h, g] U2 & ⦃G, L2⦄ ⊢ U1 ⬌* U2.
+#h #g #G0 #L0 #T0 #IH #G #L1 #T1 #H01 #HT1 #l #Hl #U1 #HTU1 #T2 #HT12 #L2 #HL12
+elim (IH … H01 … 1 … Hl U1 … HT12 … HL12)
+/3 width=3 by lsstas_inv_SO, ssta_lsstas, ex2_intro/
+qed-.
+
+fact lsstas_cprs_lpr_aux: ∀h,g,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h g G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h g G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_lsstas_cpr_lpr h g G1 L1 T1) →
+ ∀G,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
+ ∀l1,l2. l2 ≤ l1 → ⦃G, L1⦄ ⊢ T1 ▪[h, g] l1 →
+ ∀U1. ⦃G, L1⦄ ⊢ T1 •*[h, g, l2] U1 →
+ ∀T2. ⦃G, L1⦄ ⊢ T1 ➡* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
+ ∃∃U2. ⦃G, L2⦄ ⊢ T2 •*[h, g, l2] U2 & ⦃G, L2⦄ ⊢ U1 ⬌* U2.
+#h #g #G0 #L0 #T0 #IH3 #IH2 #IH1 #G #L1 #T1 #H01 #HT1 #l1 #l2 #Hl21 #Hl1 #U1 #HTU1 #T2 #H
+@(cprs_ind … H) -T2 [ /2 width=10/ ]
+#T #T2 #HT1T #HTT2 #IHT1 #L2 #HL12
+elim (IHT1 L1) // -IHT1 #U #HTU #HU1
+elim (IH1 … Hl21 … HTU … HTT2 … HL12) -IH1 -HTU -HTT2
+[2: /3 width=12 by da_cprs_lpr_aux/
+|3: /3 width=10 by snv_cprs_lpr_aux/
+|4: /3 width=5 by ygt_yprs_trans, cprs_yprs/
+] -G0 -L0 -T0 -T1 -T -l1 #U2 #HTU2 #HU2
+/4 width=5 by lpr_cpcs_conf, cpcs_trans, ex2_intro/
qed-.
-fact snv_sstas_lpr_aux: ∀h,g,G0,L0,T0.
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_ssta h g G1 L1 T1) →
- ∀G,L1,T. ⦃G0, L0, T0⦄ >[h, g] ⦃G, L1, T⦄ → ⦃G, L1⦄ ⊢ T ¡[h, g] →
- ∀U. ⦃G, L1⦄ ⊢ T •*[h, g] U → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L2⦄ ⊢ U ¡[h, g].
-/4 width=8 by snv_sstas_aux, ygt_yprs_trans, sstas_yprs/
+fact lsstas_cpcs_lpr_aux: ∀h,g,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h g G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h g G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_lsstas_cpr_lpr h g G1 L1 T1) →
+ ∀G,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
+ ∀l,l1. l ≤ l1 → ⦃G, L1⦄ ⊢ T1 ▪[h, g] l1 → ∀U1. ⦃G, L1⦄ ⊢ T1 •*[h, g, l] U1 →
+ ∀T2. ⦃G0, L0, T0⦄ >[h, g] ⦃G, L1, T2⦄ → ⦃G, L1⦄ ⊢ T2 ¡[h, g] →
+ ∀l2. l ≤ l2 → ⦃G, L1⦄ ⊢ T2 ▪[h, g] l2 → ∀U2. ⦃G, L1⦄ ⊢ T2 •*[h, g, l] U2 →
+ ⦃G, L1⦄ ⊢ T1 ⬌* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L2⦄ ⊢ U1 ⬌* U2.
+#h #g #G0 #L0 #T0 #IH3 #IH2 #IH1 #G #L1 #T1 #H01 #HT1 #l #l1 #Hl1 #HTl1 #U1 #HTU1 #T2 #H02 #HT2 #l2 #Hl2 #HTl2 #U2 #HTU2 #H #L2 #HL12
+elim (cpcs_inv_cprs … H) -H #T #H1 #H2
+elim (lsstas_cprs_lpr_aux … H01 HT1 … Hl1 HTl1 … HTU1 … H1 … HL12) -T1 /2 width=1/ #W1 #H1 #HUW1
+elim (lsstas_cprs_lpr_aux … H02 HT2 … Hl2 HTl2 … HTU2 … H2 … HL12) -T2 /2 width=1/ #W2 #H2 #HUW2 -L0 -T0
+lapply (lsstas_mono … H1 … H2) -h -T -l #H destruct /2 width=3 by cpcs_canc_dx/
qed-.
-fact sstas_cprs_lpr_aux: ∀h,g,G0,L0,T0.
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_ssta h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_ssta_cpr_lpr h g G1 L1 T1) →
- ∀G,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀U1. ⦃G, L1⦄ ⊢ T1 •*[h, g] U1 → ∀T2. ⦃G, L1⦄ ⊢ T1 ➡* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
- ∃∃U2. ⦃G, L2⦄ ⊢ T2 •*[h, g] U2 & ⦃G, L2⦄ ⊢ U1 ⬌* U2.
-#h #g #G0 #L0 #T0 #IH3 #IH2 #IH1 #G #L1 #T1 #H01 #HT1 #U1 #H
-@(sstas_ind … H) -U1 [ /3 width=5 by lpr_cprs_conf, ex2_intro/ ]
-#U1 #W1 #l1 #HTU1 #HUW1 #IHTU1 #T2 #HT12 #L2 #HL12
-elim (IHTU1 … HT12 … HL12) -IHTU1 #U2 #HTU2 #HU12
-lapply (snv_cprs_lpr_aux … IH2 … HT1 … HT12 … HL12) // #HT2
-elim (snv_sstas_fwd_correct … HTU2) // #W2 #l2 #HUW2
-elim (IH1 … HUW1 U1 … HL12) -HUW1 //
-[2: /3 width=8 by snv_sstas_aux/
-|3: /3 width=5 by ygt_yprs_trans, sstas_yprs/
-] #W #HU1W #HW1
-elim (ssta_cpcs_lpr_aux … IH2 IH1 … HU1W … HUW2 … HU12 L2) // -IH1 -HU1W -HU12
-[2: /4 width=9 by snv_sstas_aux, ygt_yprs_trans, cprs_lpr_yprs/
-|3: /3 width=11 by snv_sstas_lpr_aux/
-|4,5: /4 width=5 by ygt_yprs_trans, cprs_lpr_yprs, sstas_yprs/
-] -G0 -L0 -T0 -L1 -T1 -HT2 #H #HW12 destruct
-lapply (cpcs_trans … HW1 … HW12) -W /3 width=4/
+fact snv_ssta_aux: ∀h,g,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_lsstas h g G1 L1 T1) →
+ ∀G,L,T. ⦃G0, L0, T0⦄ >[h, g] ⦃G, L, T⦄ → ⦃G, L⦄ ⊢ T ¡[h, g] →
+ ∀l. ⦃G, L⦄ ⊢ T ▪[h, g] l+1 →
+ ∀U. ⦃G, L⦄ ⊢ T •[h, g] U → ⦃G, L⦄ ⊢ U ¡[h, g].
+/3 width=8 by lsstas_inv_SO, ssta_lsstas/ qed-.
+
+fact lsstas_cpds_aux: ∀h,g,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_lsstas h g G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h g G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h g G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_lsstas_cpr_lpr h g G1 L1 T1) →
+ ∀G,L,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G, L, T1⦄ → ⦃G, L⦄ ⊢ T1 ¡[h, g] →
+ ∀l1,l2. l2 ≤ l1 → ⦃G, L⦄ ⊢ T1 ▪[h, g] l1 →
+ ∀U1. ⦃G, L⦄ ⊢ T1 •*[h, g, l2] U1 → ∀T2. ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2 →
+ ∃∃U2,l. l ≤ l2 & ⦃G, L⦄ ⊢ T2 •*[h, g, l] U2 & ⦃G, L⦄ ⊢ U1 •*⬌*[h, g] U2.
+#h #g #G0 #L0 #T0 #IH4 #IH3 #IH2 #IH1 #G #L #T1 #H01 #HT1 #l1 #l2 #Hl21 #Hl1 #U1 #HTU1 #T2 * #T #l0 #l #Hl0 #H #HT1T #HTT2
+lapply (da_mono … H … Hl1) -H #H destruct
+lapply (lsstas_da_conf … HTU1 … Hl1) #Hl12
+elim (le_or_ge l2 l) #Hl2
+[ lapply (lsstas_conf_le … HTU1 … HT1T) -HT1T // #HU1T
+ /5 width=11 by cpds_cpes_dx, monotonic_le_minus_l, ex3_2_intro, ex4_3_intro/
+| lapply (lsstas_da_conf … HT1T … Hl1) #Hl1l
+ lapply (lsstas_conf_le … HT1T … HTU1) -HTU1 // #HTU1
+ elim (lsstas_cprs_lpr_aux … IH3 IH2 IH1 … Hl1l … HTU1 … HTT2 L)
+ /3 width=8 by ygt_yprs_trans, lsstas_yprs, monotonic_le_minus_l/ -T #U2 #HTU2 #HU12
+ /3 width=5 by cpcs_cpes, ex3_2_intro/
+]
qed-.
-fact cpds_cprs_lpr_aux: ∀h,g,G0,L0,T0.
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_ssta h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_ssta_cpr_lpr h g G1 L1 T1) →
- ∀G,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀U1. ⦃G, L1⦄ ⊢ T1 •*➡*[h, g] U1 →
- ∀T2. ⦃G, L1⦄ ⊢ T1 ➡* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
- ∃∃U2. ⦃G, L2⦄ ⊢ T2 •*➡*[h, g] U2 & ⦃G, L2⦄ ⊢ U1 ➡* U2.
-#h #g #G0 #L0 #T0 #IH3 #IH2 #IH1 #G #L1 #T1 #H01 #HT1 #U1 * #W1 #HTW1 #HWU1 #T2 #HT12 #L2 #HL12
-elim (sstas_cprs_lpr_aux … IH3 IH2 IH1 … H01 … HTW1 … HT12 … HL12) // -L0 -T0 -T1 #W2 #HTW2 #HW12
+fact cpds_cpr_lpr_aux: ∀h,g,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h g G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_lsstas_cpr_lpr h g G1 L1 T1) →
+ ∀G,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
+ ∀U1. ⦃G, L1⦄ ⊢ T1 •*➡*[h, g] U1 →
+ ∀T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
+ ∃∃U2. ⦃G, L2⦄ ⊢ T2 •*➡*[h, g] U2 & ⦃G, L2⦄ ⊢ U1 ➡* U2.
+#h #g #G0 #L0 #T0 #IH2 #IH1 #G #L1 #T1 #H01 #HT1 #U1 * #W1 #l1 #l2 #Hl21 #Hl1 #HTW1 #HWU1 #T2 #HT12 #L2 #HL12
+elim (IH1 … H01 … HTW1 … HT12 … HL12) -IH1 // #W2 #HTW2 #HW12
+lapply (IH2 … H01 … Hl1 … HT12 … HL12) -L0 -T0 // -T1 #Hl1
lapply (lpr_cprs_conf … HL12 … HWU1) -L1 #HWU1
lapply (cpcs_canc_sn … HW12 HWU1) -W1 #H
-elim (cpcs_inv_cprs … H) -H /3 width=3/
-qed-.
-
-fact ssta_cpds_aux: ∀h,g,G0,L0,T0.
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_ssta_cpr_lpr h g G1 L1 T1) →
- ∀G,L,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G, L, T1⦄ → ⦃G, L⦄ ⊢ T1 ¡[h, g] →
- ∀l,U1. ⦃G, L⦄ ⊢ T1 •[h, g] ⦃l+1, U1⦄ → ∀T2. ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2 →
- ∃∃U,U2. ⦃G, L⦄ ⊢ U1 •*[h, g] U & ⦃G, L⦄ ⊢ T2 •*[h, g] U2 & ⦃G, L⦄ ⊢ U ⬌* U2.
-#h #g #G0 #L0 #T0 #IH2 #IH1 #G #L #T1 #H01 #HT1 #l #U1 #HTU1 #T2 * #T #HT1T #HTT2
-elim (sstas_strip … HT1T … HTU1) #HU1T destruct [ -HT1T | -L0 -T0 -T1 ]
-[ elim (ssta_cprs_lpr_aux … IH2 IH1 … HTU1 … HTT2 L) // -G0 -L0 -T0 -T /3 width=5/
-| @(ex3_2_intro …T2 HU1T) // /2 width=1/
-]
+elim (cpcs_inv_cprs … H) -H /3 width=7 by ex4_3_intro, 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/lsubd_da.ma".
+include "basic_2/computation/cpds_cpds.ma".
+include "basic_2/dynamic/snv_aaa.ma".
+include "basic_2/dynamic/snv_cpcs.ma".
+
+(* STRATIFIED NATIVE VALIDITY FOR TERMS *************************************)
+
+(* Properties on degree assignment for terms ********************************)
+
+fact da_cpr_lpr_aux: ∀h,g,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_lsstas h g G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h g G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h g G1 L1 T1) →
+ ∀G1,L1,T1. G0 = G1 → L0 = L1 → T0 = T1 → IH_da_cpr_lpr h g G1 L1 T1.
+#h #g #G0 #L0 #T0 #IH3 #IH2 #IH1 #G1 #L1 * * [|||| * ]
+[ #k #_ #_ #_ #_ #l #H2 #X3 #H3 #L2 #_ -IH3 -IH2 -IH1
+ lapply (da_inv_sort … H2) -H2
+ lapply (cpr_inv_sort1 … H3) -H3 #H destruct /2 width=1/
+| #i #HG0 #HL0 #HT0 #H1 #l #H2 #X3 #H3 #L2 #HL12 destruct -IH3 -IH2
+ elim (snv_inv_lref … H1) -H1 #I0 #K0 #X0 #H #HX0
+ elim (da_inv_lref … H2) -H2 * #K1 [ #V1 | #W1 #l1 ] #HLK1 [ #HV1 | #HW1 #H ] destruct
+ lapply (ldrop_mono … H … HLK1) -H #H destruct
+ elim (cpr_inv_lref1 … H3) -H3
+ [1,3: #H destruct
+ lapply (fsupp_lref … G1 … HLK1)
+ elim (lpr_ldrop_conf … HLK1 … HL12) -HLK1 -HL12 #X #H #HLK2
+ elim (lpr_inv_pair1 … H) -H #K2 #V2 #HK12 #HV12 #H destruct
+ /4 width=10 by da_ldef, da_ldec, fsupp_ygt/
+ |2,4: * #K0 #V0 #W0 #H #HVW0 #HW0
+ lapply (ldrop_mono … H … HLK1) -H #H destruct
+ lapply (fsupp_lref … G1 … HLK1)
+ elim (lpr_ldrop_conf … HLK1 … HL12) -HLK1 -HL12 #X #H #HLK2
+ elim (lpr_inv_pair1 … H) -H #K2 #V2 #HK12 #_ #H destruct
+ lapply (ldrop_fwd_ldrop2 … HLK2) -V2
+ /4 width=7 by da_lift, fsupp_ygt/
+ ]
+| #p #_ #_ #HT0 #H1 destruct -IH3 -IH2 -IH1
+ elim (snv_inv_gref … H1)
+| #a #I #V1 #T1 #HG0 #HL0 #HT0 #H1 #l #H2 #X3 #H3 #L2 #HL12 destruct -IH2
+ elim (snv_inv_bind … H1) -H1 #_ #HT1
+ lapply (da_inv_bind … H2) -H2
+ elim (cpr_inv_bind1 … H3) -H3 *
+ [ #V2 #T2 #HV12 #HT12 #H destruct
+ /4 width=9 by da_bind, fsupp_ygt, lpr_pair/
+ | #T2 #HT12 #HT2 #H1 #H2 destruct
+ /4 width=11 by da_inv_lift, fsupp_ygt, lpr_pair, ldrop_ldrop/
+ ]
+| #V1 #T1 #HG0 #HL0 #HT0 #H1 #l #H2 #X3 #H3 #L2 #HL12 destruct
+ elim (snv_inv_appl … H1) -H1 #b0 #W1 #W0 #T0 #l0 #HV1 #HT1 #Hl0 #HVW1 #HW10 #HT10
+ lapply (da_inv_flat … H2) -H2 #Hl
+ elim (cpr_inv_appl1 … H3) -H3 *
+ [ #V2 #T2 #HV12 #HT12 #H destruct -IH3 -IH2 /4 width=7 by da_flat, fsupp_ygt/
+ | #b #V2 #W #W2 #U1 #U2 #HV12 #HW2 #HU12 #H1 #H2 destruct
+ elim (snv_inv_bind … HT1) -HT1 #HW #HU1
+ lapply (da_inv_bind … Hl) -Hl #Hl
+ elim (cpds_inv_abst1 … HT10) -HT10 #W3 #U3 #HW3 #_ #H destruct -U3
+ lapply (cprs_div … HW3 … HW10) -W3 #HWW1
+ lapply (ssta_da_conf … HVW1 … Hl0) <minus_plus_m_m #H
+ elim (snv_fwd_da … HW) #l1 #Hl1
+ lapply (IH3 … HV1 … 1 … Hl0 W1 ?) /2 width=2 by fsupp_ygt, ssta_lsstas/ #HW1
+ lapply (da_cpcs_aux … IH2 IH1 … Hl1 … H … HWW1) -H
+ /3 width=5 by ygt_yprs_trans, fsupp_ygt, ssta_yprs/ #H destruct
+ lapply (IH1 … HV1 … Hl0 … HV12 … HL12) -HV1 -Hl0 -HV12 [ /2 by fsupp_ygt/ ] #Hl0
+ lapply (IH1 … Hl1 … HW2 … HL12) -Hl1 // /2 width=1 by fsupp_ygt/ -HW
+ lapply (IH1 … HU1 … Hl … HU12 (L2.ⓛW2) ?) -IH1 -HU1 -Hl -HU12 [1,2: /2 by fsupp_ygt, lpr_pair/ ] -HL12 -HW2
+ /4 width=6 by da_bind, lsubd_da_trans, lsubd_abbr/
+ | #b #V #V2 #W #W2 #U1 #U2 #HV1 #HV2 #HW2 #HU12 #H1 #H2 destruct -IH3 -IH2 -V -W0 -T0 -l0 -HV1 -HVW1
+ elim (snv_inv_bind … HT1) -HT1 #_
+ lapply (da_inv_bind … Hl) -Hl
+ /5 width=9 by da_bind, da_flat, fsupp_ygt, lpr_pair/
+ ]
+| #W1 #T1 #HG0 #HL0 #HT0 #H1 #l #H2 #X3 #H3 #L2 #HL12 destruct -IH3 -IH2
+ elim (snv_inv_cast … H1) -H1 #U1 #l0 #HW1 #HT1 #Hl0 #HTU1 #HUW1
+ lapply (da_inv_flat … H2) -H2 #Hl
+ elim (cpr_inv_cast1 … H3) -H3
+ [ * #W2 #T2 #HW12 #HT12 #H destruct /4 width=7 by da_flat, fsupp_ygt/
+ | /3 width=7 by fsupp_ygt/
+ ]
+]
+qed-.
| #a #I #G #K #V #T #_ #_ #IHV #IHT #L #d #e #HLK #X #H
elim (lift_inv_bind1 … H) -H #W #U #HVW #HTU #H destruct
/4 width=4/
-| #a #G #K #V #V0 #V1 #T #T1 #l #_ #_ #HV0 #HV01 #HT1 #IHV #IHT #L #d #e #HLK #X #H
+| #a #G #K #V #V0 #V1 #T #T1 #l #_ #_ #Hl #HV0 #HV01 #HT1 #IHV #IHT #L #d #e #HLK #X #H
elim (lift_inv_flat1 … H) -H #W #U #HVW #HTU #H destruct
elim (lift_total V0 d e) #W0 #HVW0
elim (lift_total V1 d e) #W1 #HVW1
elim (lift_total T1 (d+1) e) #U1 #HTU1
@(snv_appl … a … W0 … W1 … U1 l)
- [ /2 width=4/ | /2 width=4/ | /2 width=9/ | /2 width=9/ ]
+ [1,2: /2 width=4/ | /2 width=7/ |4,5: /2 width=9/ ]
@(cpds_lift … HT1 … HLK … HTU) /2 width=1/
-| #G #K #V0 #T #V #l #_ #_ #HTV #HV0 #IHV0 #IHT #L #d #e #HLK #X #H
+| #G #K #V0 #T #V #l #_ #_ #Hl #HTV #HV0 #IHV0 #IHT #L #d #e #HLK #X #H
elim (lift_inv_flat1 … H) -H #W0 #U #HVW0 #HTU #H destruct
elim (lift_total V d e) #W #HVW
- @(snv_cast … W l) [ /2 width=4/ | /2 width=4/ | /2 width=9/ | /2 width=9/ ]
+ @(snv_cast … W l) [1,2: /2 width=4/ | /2 width=7/ |4,5: /2 width=9/ ]
]
qed.
]
| #a #I #G #L #W #U #_ #_ #IHW #IHU #K #d #e #HLK #X #H
elim (lift_inv_bind2 … H) -H #V #T #HVW #HTU #H destruct /4 width=4/
-| #a #G #L #W #W0 #W1 #U #U1 #l #_ #_ #HW0 #HW01 #HU1 #IHW #IHU #K #d #e #HLK #X #H
+| #a #G #L #W #W0 #W1 #U #U1 #l #_ #_ #Hl #HW0 #HW01 #HU1 #IHW #IHU #K #d #e #HLK #X #H
elim (lift_inv_flat2 … H) -H #V #T #HVW #HTU #H destruct
- elim (ssta_inv_lift1 … HW0 … HLK … HVW) -HW0 #V0 #HV0 #HVW0
+ lapply (da_inv_lift … Hl … HLK … HVW) -Hl #Hl
+ elim (ssta_inv_lift1 … HW0 … HLK … HVW) -HW0 #V0 #HVW0 #HV0
elim (cprs_inv_lift1 … HW01 … HLK … HVW0) -W0 #V1 #HVW1 #HV01
elim (cpds_inv_lift1 … HU1 … HLK … HTU) -HU1 #X #H #HTU
elim (lift_inv_bind2 … H) -H #Y #T1 #HY #HTU1 #H destruct
lapply (lift_inj … HY … HVW1) -HY #H destruct /3 width=8/
-| #G #L #W0 #U #W #l #_ #_ #HUW #HW0 #IHW0 #IHU #K #d #e #HLK #X #H
+| #G #L #W0 #U #W #l #_ #_ #Hl #HUW #HW0 #IHW0 #IHU #K #d #e #HLK #X #H
elim (lift_inv_flat2 … H) -H #V0 #T #HVW0 #HTU #H destruct
- elim (ssta_inv_lift1 … HUW … HLK … HTU) -HUW #V #HTV #HVW
+ lapply (da_inv_lift … Hl … HLK … HTU) -Hl #Hl
+ elim (ssta_inv_lift1 … HUW … HLK … HTU) -HUW #V #HVW #HTV
lapply (cpcs_inv_lift G … HLK … HVW … HVW0 ?) // -W /3 width=4/
]
qed-.
(* *)
(**************************************************************************)
-include "basic_2/computation/cpds_cpds.ma".
include "basic_2/dynamic/snv_lift.ma".
include "basic_2/dynamic/snv_cpcs.ma".
+include "basic_2/dynamic/lsubsv_snv.ma".
+include "basic_2/dynamic/yprs_yprs.ma".
(* STRATIFIED NATIVE VALIDITY FOR TERMS *************************************)
(* Properties on context-free parallel reduction for local environments *****)
fact snv_cpr_lpr_aux: ∀h,g,G0,L0,T0.
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_lsubsv h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_ssta_cpr_lpr h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_ssta h g G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_lsstas h g G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_lsstas_cpr_lpr h g G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h g G1 L1 T1) →
(∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h g G1 L1 T1) →
∀G1,L1,T1. G0 = G1 → L0 = L1 → T0 = T1 → IH_snv_cpr_lpr h g G1 L1 T1.
-#h #g #G0 #L0 #T0 #IH4 #IH3 #IH2 #IH1 #G1 #L1 * * [||||*]
+#h #g #G0 #L0 #T0 #IH4 #IH3 #IH2 #IH1 #G1 #L1 * * [|||| * ]
[ #k #HG0 #HL0 #HT0 #H1 #X #H2 #L2 #_ destruct -IH4 -IH3 -IH2 -IH1 -H1
>(cpr_inv_sort1 … H2) -X //
| #i #HG0 #HL0 #HT0 #H1 #X #H2 #L2 #HL12 destruct -IH4 -IH3 -IH2
elim (lpr_inv_pair1 … H) -H #K2 #V2 #HK12 #HV12 #H destruct
lapply (fsupp_lref … G1 … HLK1) #HKL
elim (cpr_inv_lref1 … H2) -H2
- [ #H destruct -HLK1
- lapply (IH1 … HV12 … HK12) -IH1 -HV12 -HK12 // -HV1 [ /2 width=1/ ] -HKL /2 width=5/
+ [ #H destruct -HLK1 /4 width=10 by fsupp_ygt, snv_lref/
| * #K0 #V0 #W0 #H #HVW0 #W0 -HV12
lapply (ldrop_mono … H … HLK1) -HLK1 -H #H destruct
- lapply (ldrop_fwd_ldrop2 … HLK2) -HLK2 #HLK2
- lapply (IH1 … HVW0 … HK12) -IH1 -HVW0 -HK12 // -HV1 [ /2 width=1/ ] -HKL /3 width=7/
+ lapply (ldrop_fwd_ldrop2 … HLK2) -HLK2 /4 width=7 by fsupp_ygt, snv_lift/
]
| #p #HG0 #HL0 #HT0 #H1 #X #H2 #L2 #HL12 destruct -IH4 -IH3 -IH2 -IH1
elim (snv_inv_gref … H1)
| #a #I #V1 #T1 #HG0 #HL0 #HT0 #H1 #X #H2 #L2 #HL12 destruct -IH4 -IH3 -IH2
elim (snv_inv_bind … H1) -H1 #HV1 #HT1
elim (cpr_inv_bind1 … H2) -H2 *
- [ #V2 #T2 #HV12 #HT12 #H destruct
- lapply (IH1 … HV12 … HL12) // -HV1 [ /2 width=1/ ] #HV2
- lapply (IH1 … HT12 (L2.ⓑ{I}V2) ?) // -IH1 /2 width=1/
+ [ #V2 #T2 #HV12 #HT12 #H destruct /4 width=8 by fsupp_ygt, snv_bind, lpr_pair/
| #T2 #HT12 #HXT2 #H1 #H2 destruct -HV1
- lapply (IH1 … HT1 … HT12 (L2.ⓓV1) ?) -IH1 [1,2: /2 width=1/ ] -L1 -T1 #HT2
- lapply (snv_inv_lift … HT2 L2 … HXT2) -T2 // /2 width=1/
+ /4 width=10 by fsupp_ygt, snv_inv_lift, lpr_pair, ldrop_ldrop/
]
| #V1 #T1 #HG0 #HL0 #HT0 #H1 #X #H2 #L2 #HL12 destruct
- elim (snv_inv_appl … H1) -H1 #a #W10 #W1 #U1 #l #HV1 #HT1 #HVW1 #HW10 #HTU1
+ elim (snv_inv_appl … H1) -H1 #a #W10 #W1 #U1 #l0 #HV1 #HT1 #Hl0 #HVW1 #HW10 #HTU1
elim (cpr_inv_appl1 … H2) -H2 *
[ #V2 #T2 #HV12 #HT12 #H destruct -IH4
- lapply (IH1 … HV1 … HV12 … HL12) [ /2 width=1/ ] #HV2
- lapply (IH1 … HT1 … HT12 … HL12) [ /2 width=1/ ] #HT2
- elim (IH3 … HVW1 … HV12 … HL12) -HVW1 -HV12 // -HV1 [2: /2 width=1/ ] #W2 #HVW2 #HW12
- elim (cpds_cprs_lpr_aux … IH2 IH1 IH3 … HTU1 … T2 … HL12) // [2,3: /2 width=1/ ] -IH2 -IH1 -IH3 -HT1 -HT12 -HTU1 #X #HTU2 #H
+ lapply (IH1 … HV12 … HL12) /2 width=1 by fsupp_ygt/ #HV2
+ lapply (IH1 … HT12 … HL12) /2 width=1 by fsupp_ygt/ #HT2
+ lapply (IH2 … Hl0 … HV12 … HL12) /2 width=1 by fsupp_ygt/ #H2l0
+ elim (ssta_cpr_lpr_aux … IH3 … Hl0 … HVW1 … HV12 … HL12) -Hl0 -HVW1 -HV12 /2 width=1 by fsupp_ygt/ -HV1 #W2 #HVW2 #HW12
+ elim (cpds_cpr_lpr_aux … IH2 IH3 … HTU1 … HT12 … HL12) /2 width=1 by fsupp_ygt/ -HT12 -HTU1 #X #HTU2 #H
elim (cprs_inv_abst1 … H) -H #W20 #U2 #HW120 #_ #H destruct
lapply (lpr_cprs_conf … HL12 … HW10) -L1 #HW10
lapply (cpcs_cprs_strap1 … HW10 … HW120) -W1 #HW120
lapply (cpcs_canc_sn … HW12 HW120) -W10 #HW20
elim (cpcs_inv_cprs … HW20) -HW20 #W0 #HW20 #HW200
- lapply (cpds_cprs_trans … (ⓛ{a}W0.U2) HTU2 ?) [ /2 width=1/ ] -HW200 -HTU2 /2 width=8/
+ lapply (cpds_cprs_trans … (ⓛ{a}W0.U2) HTU2 ?)
+ /2 width=7 by snv_appl, cprs_bind/
| #b #V2 #W20 #W2 #T20 #T2 #HV12 #HW202 #HT202 #H1 #H2 destruct
elim (snv_inv_bind … HT1) -HT1 #HW20 #HT20
elim (cpds_inv_abst1 … HTU1) -HTU1 #W30 #T30 #HW230 #_ #H destruct -T30
lapply (cprs_div … HW10 … HW230) -W30 #HW120
+ lapply (snv_ssta_aux … IH4 … Hl0 … HVW1) /2 width=1 by fsupp_ygt/ #HW10
+ lapply (ssta_da_conf … HVW1 … Hl0) <minus_plus_m_m #HlW10
+ elim (snv_fwd_da … HW20) #l #Hl
+ lapply (da_cpcs_aux … IH1 IH2 … HlW10 … Hl … HW120) // -HlW10
+ /3 width=5 by ygt_yprs_trans, fsupp_ygt, ssta_yprs/ #H destruct
+ lapply (IH2 … Hl0 … HV12 … HL12) /2 width=1 by fsupp_ygt/ #HlV2
+ lapply (IH2 … Hl … HW202 … HL12) /2 width=1 by fsupp_ygt/ #HlW2
+ elim (ssta_cpr_lpr_aux … IH3 … Hl0 … HVW1 … HV12 … HL12) /2 width=1 by fsupp_ygt/ #W3 #HV2W3 #HW103
+ lapply (ssta_da_conf … HV2W3 … HlV2) <minus_plus_m_m #HlW3
lapply (cpcs_cpr_strap1 … HW120 … HW202) -HW120 #HW102
lapply (lpr_cpcs_conf … HL12 … HW102) -HW102 #HW102
- elim (IH3 … HVW1 … HV12 … HL12) // [2: /2 width=1/ ] -HVW1 #W3 #HV2W3 #HW103
lapply (cpcs_canc_sn … HW103 … HW102) -W10 #HW32
- lapply (IH1 … HV12 … HL12) // [ /2 width=1/ ] -HV1 #HV2
- lapply (IH1 … HW202 … HL12) // [ /2 width=1/ ] -HW20 #HW2
- lapply (IH1 … HT20 … HT202 … (L2.ⓛW2) ?) [1,2: /2 width=1/ ] -HT20 #HT2
- lapply (IH2 … HV2W3) //
- [ @(ygt_yprs_trans … G1 G1 … L1 L1 … V1) (**) (* auto /4 width=5/ is a bit slow even with trace *)
- [ /2 width=1 by fsupp_ygt/
- | /3 width=1 by cprs_lpr_yprs, cpr_cprs/
- ]
- ] #HW3
- elim (snv_fwd_ssta … HW2) #l0 #U2 #HWU2
- elim (ssta_fwd_correct … HV2W3) <minus_plus_m_m #U3 #HWU3
- elim (ssta_cpcs_lpr_aux … IH1 IH3 … HWU3 … HWU2 … HW32 … L2) // -IH3
- [2: /4 width=5 by ygt_yprs_trans, fsupp_ygt, cprs_lpr_yprs, cpr_cprs/
- |3: @(ygt_yprs_trans … G1 G1 L1 L2 … V2) (**) (* auto not tried *)
- [ @(ygt_yprs_trans … G1 G1 L1 L1 … V1)
- [ /2 width=1 by fsupp_ygt/
- | /3 width=1 by cprs_lpr_yprs, cpr_cprs/
- ]
- | /3 width=2 by ypr_ssta, ypr_yprs/
- ]
- ] #H #_ destruct -IH2 -U3
- lapply (IH4 … HT2 (L2.ⓓⓝW2.V2) ?)
- [ /3 width=5/
- | @(ygt_yprs_trans … G1 G1 … (L1.ⓛW20) … T2) (**) (* auto /5 width=5/ is too slow even with trace timeout=35 *)
- [ /4 width=5 by ygt_yprs_trans, fsupp_ygt, ypr_yprs, ypr_cpr/
- | /4 width=1 by ypr_yprs, ypr_lpr, lpr_pair/
- ]
- ] -L1 -V1 -T20 -U2 /3 width=4/
+ lapply (IH1 … HV12 … HL12) /2 width=1 by fsupp_ygt/ -HV1 #HV2
+ lapply (IH1 … HW202 … HL12) /2 width=1 by fsupp_ygt/ -HW20 #HW2
+ lapply (IH1 … HT20 … HT202 … (L2.ⓛW2) ?) /2 width=1 by fsupp_ygt, lpr_pair/ -HT20 #HT2
+ lapply (snv_ssta_aux … IH4 … HlV2 … HV2W3)
+ /3 width=5 by ygt_yprs_trans, fsupp_ygt, cpr_lpr_yprs/ #HW3
+ lapply (lsubsv_snv_trans … HT2 (L2.ⓓⓝW2.V2) ?) -HT2 /3 width=3 by snv_bind, snv_cast/
+ @(lsubsv_abbr … l) /3 width=7 by fsupp_ygt/ #W #W0 #l0 #Hl0 #HV2W #HW20
+ lapply (lsstas_ssta_conf_pos … HV2W3 … HV2W) -HV2W #HW3W
+ @(lsstas_cpcs_lpr_aux … IH1 IH2 IH3 … HlW3 … HW3W … HlW2 … HW20 … HW32) //
+ [ /3 width=9 by ygt_yprs_trans, fsupp_ygt, cpr_lpr_ssta_yprs/
+ | /3 width=5 by ygt_yprs_trans, fsupp_ygt, cpr_lpr_yprs/
+ ]
| #b #V0 #V2 #W0 #W2 #T0 #T2 #HV10 #HV02 #HW02 #HT02 #H1 #H2 destruct -IH4
elim (snv_inv_bind … HT1) -HT1 #HW0 #HT0
elim (cpds_inv_abbr_abst … HTU1) -HTU1 #X #HTU0 #HX #H destruct
elim (lift_inv_bind1 … HX) -HX #W3 #U3 #HW13 #_ #H destruct
lapply (lpr_cprs_conf … HL12 … HW10) -HW10 #HW10
- elim (cpds_cprs_lpr_aux … IH2 IH1 IH3 … HTU0 T2 … (L2.ⓓW2)) // [2,3,4: /2 width=1/ ] -IH2 -HTU0 #X #HTU2 #H
+ elim (cpds_cpr_lpr_aux … IH2 IH3 … HTU0 … HT02 (L2.ⓓW2)) /2 width=1 by fsupp_ygt, lpr_pair/ -HTU0 #X #HTU2 #H
elim (cprs_inv_abst1 … H) -H #W #U2 #HW1 #_ #H destruct -U3
- elim (IH3 … HVW1 … HV10 … HL12) // /2 width=1/ -IH3 -HVW1 #X #H1 #H2
+ elim (ssta_cpr_lpr_aux … IH3 … HVW1 … HV10 … HL12) /2 width=2 by fsupp_ygt/ -IH3 -HVW1 #X #H1 #H2
lapply (cpcs_canc_sn … H2 HW10) -W10 #H2
elim (lift_total X 0 1) #W20 #H3
- lapply (ssta_lift … H1 (L2.ⓓW2) … HV02 … H3) /2 width=1/ -H1 #HVW20
- lapply (cpcs_lift … (L2.ⓓW2) … H3 … HW13 H2) /2 width=1/ -HW13 -H3 -H2 #HW320
+ lapply (ssta_lift … H1 (L2.ⓓW2) … HV02 … H3) /2 width=1 by ldrop_ldrop/ -H1 #HVW20
+ lapply (cpcs_lift … (L2.ⓓW2) … H3 … HW13 H2) /2 width=1 by ldrop_ldrop/ -HW13 -H3 -H2 #HW320
lapply (cpcs_cprs_strap1 … HW320 … HW1) -W3 #HW20
elim (cpcs_inv_cprs … HW20) -HW20 #W3 #HW203 #HW3
- lapply (cpds_cprs_trans … (ⓛ{a}W3.U2) HTU2 ?) [ /2 width=1/ ] -HW3 -HTU2 #HTU2
- lapply (IH1 … HW02 … HL12) // [ /2 width=1/ ] -HW0 #HW2
- lapply (IH1 … HV10 … HL12) // [ /2 width=1/ ] -HV1 -HV10 #HV0
- lapply (IH1 … HT02 (L2.ⓓW2) ?) // [1,2: /2 width=1/ ] -L1 #HT2
- lapply (snv_lift … HV0 (L2.ⓓW2) … HV02) /2 width=1/ -HV0 -HV02 /3 width=8/
+ lapply (cpds_cprs_trans … (ⓛ{a}W3.U2) HTU2 ?) /2 width=1 by cprs_bind/ -HW3 -HTU2 #HTU2
+ lapply (IH2 … Hl0 … HV10 … HL12) /2 width=1 by fsupp_ygt/ -IH2 -Hl0 #Hl0
+ lapply (da_lift … Hl0 (L2.ⓓW2) … HV02) /2 width=1 by ldrop_ldrop/ -Hl0 #Hl0
+ lapply (IH1 … HW02 … HL12) /2 width=1 by fsupp_ygt/ -HW0 #HW2
+ lapply (IH1 … HV10 … HL12) /2 width=1 by fsupp_ygt/ -HV1 -HV10 #HV0
+ lapply (IH1 … HT02 (L2.ⓓW2) ?) /2 width=1 by fsupp_ygt, lpr_pair/ -L1 #HT2
+ lapply (snv_lift … HV0 (L2.ⓓW2) … HV02) /3 width=7 by snv_bind, snv_appl, ldrop_ldrop/
]
-| #W1 #T1 #HG0 #HL0 #HT0 #H1 #X #H2 #L2 #HL12 destruct -IH4 -IH2
- elim (snv_inv_cast … H1) -H1 #U1 #l #HW1 #HT1 #HTU1 #HUW1
+| #W1 #T1 #HG0 #HL0 #HT0 #H1 #X #H2 #L2 #HL12 destruct -IH4
+ elim (snv_inv_cast … H1) -H1 #U1 #l0 #HW1 #HT1 #Hl0 #HTU1 #HUW1
elim (cpr_inv_cast1 … H2) -H2
[ * #W2 #T2 #HW12 #HT12 #H destruct
- lapply (cpcs_cprs_strap1 … HUW1 W2 ?) [ /2 width=1/ ] -HUW1 #H1
- lapply (IH1 … HW12 … HL12) // /2 width=1/ -HW1 -HW12 #HW2
- lapply (IH1 … HT12 … HL12) // /2 width=1/ -IH1 #HT2
- elim (IH3 … HTU1 … HT12 … HL12) // /2 width=1/ -IH3 -HT1 -HT12 -HTU1 #U2 #HTU2 #HU12
- lapply (lpr_cpcs_conf … HL12 … H1) -L1 #H1
- lapply (cpcs_canc_sn … HU12 H1) -U1 /2 width=4/
- | #H -IH3 -HW1 -HTU1 -HUW1
- lapply (IH1 … H … HL12) // /2 width=1/
+ lapply (cpcs_cprs_strap1 … HUW1 W2 ?) /2 width=1 by cpr_cprs/ -HUW1 #H1
+ lapply (IH1 … HW12 … HL12) /2 width=1 by fsupp_ygt/ -HW1 -HW12 #HW2
+ lapply (IH1 … HT12 … HL12) /2 width=1 by fsupp_ygt/ -IH1 #HT2
+ elim (ssta_cpr_lpr_aux … IH3 … Hl0 … HTU1 … HT12 … HL12) /2 width=2 by fsupp_ygt/ -IH3 -HTU1 #U2 #HTU2 #HU12
+ lapply (IH2 … Hl0 … HT12 … HL12) /2 width=1 by fsupp_ygt/ -IH2 -HT1 -HT12 -Hl0 #Hl0
+ /4 width=7 by snv_cast, lpr_cpcs_conf, cpcs_canc_sn/
+ | #H -IH3 -IH2 -HW1 -HTU1 -HUW1
+ lapply (IH1 … H … HL12) /2 width=1 by fsupp_ygt/
]
]
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/dynamic/snv_lift.ma".
+include "basic_2/dynamic/snv_cpcs.ma".
+
+(* STRATIFIED NATIVE VALIDITY FOR TERMS *************************************)
+
+(* Properties on nat-iterated stratified static type assignment for terms ***)
+
+fact snv_lsstas_aux: ∀h,g,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h g G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h g G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_lsstas_cpr_lpr h g G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_lsstas h g G1 L1 T1) →
+ ∀G1,L1,T1. G0 = G1 → L0 = L1 → T0 = T1 → IH_snv_lsstas h g G1 L1 T1.
+#h #g #G0 #L0 #T0 #IH4 #IH3 #IH2 #IH1 #G1 #L1 * * [|||| * ]
+[ #k #HG0 #HL0 #HT0 #_ #l1 #l2 #Hl21 #Hl1 #X #H2 destruct -IH4 -IH3 -IH2 -IH1
+ >(lsstas_inv_sort1 … H2) -X //
+| #i #HG0 #HL0 #HT0 #H1 #l1 #l2 @(nat_ind_plus … l2) -l2 [ #_ | #l2 #_ #Hl21 ] #Hl1 #X #H2 destruct -IH4 -IH3 -IH2
+ [ lapply (lsstas_inv_O … H2) -H2 #H destruct // ]
+ elim (snv_inv_lref … H1) -H1 #I0 #K0 #X0 #HLK0 #HX0
+ elim (da_inv_lref … Hl1) -Hl1 * #K1 [ #V1 | #W1 #l ] #HLK1 [ #Hl1 | #Hl #H ]
+ lapply (ldrop_mono … HLK0 … HLK1) -HLK0 #H0 destruct
+ elim (lsstas_inv_lref1 … H2) -H2 * #K0 #Y0 #X0 [2,4: #l0 ] #HLK0 [1,2: #HYl0 ] #HYX0 #HX0
+ lapply (ldrop_mono … HLK0 … HLK1) -HLK0 #H destruct
+ [ lapply (le_plus_to_le_r … Hl21) -Hl21 #Hl21 ]
+ lapply (fsupp_lref … G1 … HLK1) #H
+ lapply (ldrop_fwd_ldrop2 … HLK1) -HLK1 /4 width=8 by fsupp_ygt, snv_lift/
+| #p #HG0 #HL0 #HT0 #H1 #l1 #l2 #Hl21 #Hl1 #X #H2 destruct -IH4 -IH3 -IH2 -IH1
+ elim (snv_inv_gref … H1)
+| #a #I #V1 #T1 #HG0 #HL0 #HT0 #H1 #l1 #l2 #Hl21 #Hl1 #X #H2 destruct -IH4 -IH3 -IH2
+ elim (snv_inv_bind … H1) -H1 #HV1 #HT1
+ lapply (da_inv_bind … Hl1) -Hl1 #Hl1
+ elim (lsstas_inv_bind1 … H2) -H2 #U1 #HTU1 #H destruct /4 width=8 by fsupp_ygt, snv_bind/
+| #V1 #T1 #HG0 #HL0 #HT0 #H1 #l1 #l2 #Hl21 #Hl1 #X #H2 destruct
+ elim (snv_inv_appl … H1) -H1 #a #W1 #W0 #T0 #l0 #HV1 #HT1 #Hl0 #HVW1 #HW10 #HT10
+ lapply (da_inv_flat … Hl1) -Hl1 #Hl1
+ elim (lsstas_inv_appl1 … H2) -H2 #U1 #HTU1 #H destruct
+ lapply (IH1 … HT1 … Hl1 … HTU1) /2 width=1 by fsupp_ygt/ #HU1
+ elim (lsstas_cpds_aux … IH1 IH4 IH3 IH2 … Hl1 … HTU1 … HT10) -IH4 -IH3 -IH2 -IH1 /2 width=1 by fsupp_ygt/ -T1 -l1 #X #l #_ #H #HU10 -l2
+ elim (lsstas_inv_bind1 … H) -H #U0 #_ #H destruct -T0 -l
+ elim (cpes_inv_abst2 … HU10) -HU10 #W2 #U2 #HU12 #HU02
+ elim (cprs_inv_abst … HU02) -HU02 #HW02 #_
+ /3 width=7 by snv_appl, cprs_trans/
+| #W1 #T1 #HG0 #HL0 #HT0 #H1 #l1 #l2 @(nat_ind_plus … l2) -l2 [ #_ | #l2 #_ #Hl21 ] #Hl1 #X #H2 destruct -IH4 -IH3 -IH2
+ [ lapply (lsstas_inv_O … H2) -H2 #H destruct // ]
+ elim (snv_inv_cast … H1) -H1
+ lapply (da_inv_flat … Hl1) -Hl1
+ lapply (lsstas_inv_cast1 … H2) -H2 /3 width=8 by fsupp_ygt/
+]
+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/computation/cpds_cpds.ma".
+include "basic_2/dynamic/snv_aaa.ma".
+include "basic_2/dynamic/snv_cpcs.ma".
+include "basic_2/dynamic/lsubsv_lsstas.ma".
+include "basic_2/dynamic/yprs_yprs.ma".
+
+(* STRATIFIED NATIVE VALIDITY FOR TERMS *************************************)
+
+(* Properties on sn parallel reduction for local environments ***************)
+
+fact lsstas_cpr_lpr_aux: ∀h,g,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_lsstas h g G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h g G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h g G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_lsstas_cpr_lpr h g G1 L1 T1) →
+ ∀G1,L1,T1. G0 = G1 → L0 = L1 → T0 = T1 → IH_lsstas_cpr_lpr h g G1 L1 T1.
+#h #g #G0 #L0 #T0 #IH4 #IH3 #IH2 #IH1 #G1 #L1 * * [|||| * ]
+[ #k #_ #_ #_ #_ #l1 #l2 #_ #_ #X2 #H2 #X3 #H3 #L2 #_ -IH4 -IH3 -IH2 -IH1
+ >(lsstas_inv_sort1 … H2) -X2
+ >(cpr_inv_sort1 … H3) -X3 /2 width=3 by cpr_atom, ex2_intro/
+| #i #HG0 #HL0 #HT0 #H1 #l1 #l2 @(nat_ind_plus … l2) -l2 [ #_ | #l2 #_ #Hl21 ] #Hl1 #X2 #H2 #X3 #H3 #L2 #HL12 destruct -IH4 -IH3
+ [ lapply (lsstas_inv_O … H2) -H2 #H destruct -IH1 -H1 -l1 /4 width=5 by lpr_cpcs_conf, cpr_cpcs_dx, ex2_intro/ ]
+ elim (snv_inv_lref … H1) -H1 #I0 #K0 #X0 #HK0 #HX0
+ elim (da_inv_lref … Hl1) -Hl1 * #K1 [ #V1 | #W1 #l0 ] #HLK1 [ #HVl1 | #HWl1 #H destruct ]
+ lapply (ldrop_mono … HK0 … HLK1) -HK0 #H destruct
+ elim (lsstas_inv_lref1 … H2) -H2 * #K0 #V0 #W0 [2,4: #l ] #HK0 [1,2: #Hl ] #HW0 #HX2
+ lapply (ldrop_mono … HK0 … HLK1) -HK0 #H destruct
+ [ lapply (da_mono … Hl … HWl1) -Hl #H destruct
+ lapply (le_plus_to_le_r … Hl21) -Hl21 #Hl21
+ ]
+ lapply (fsupp_lref … G1 … HLK1) #HKV1
+ elim (lpr_ldrop_conf … HLK1 … HL12) -HL12 #X #H #HLK2
+ elim (lpr_inv_pair1 … H) -H #K2 [ #W2 | #V2 ] #HK12 [ #HW12 | #HV12 ] #H destruct
+ lapply (ldrop_fwd_ldrop2 … HLK2) #H2
+ elim (cpr_inv_lref1 … H3) -H3
+ [1,3: #H destruct -HLK1
+ |2,4: * #K0 #V0 #X0 #H #HVX0 #HX0
+ lapply (ldrop_mono … H … HLK1) -H -HLK1 #H destruct
+ ]
+ [ elim (IH1 … HWl1 … HW0 … HW12 … HK12) -IH1 -HW0 /2 width=1 by fsupp_ygt/ #V2 #HWV2 #HV2
+ elim (lift_total V2 0 (i+1))
+ /6 width=11 by fsupp_ygt, cpcs_lift, lsstas_ldec, ex2_intro/
+ | elim (IH1 … HVl1 … HW0 … HV12 … HK12) -IH1 -HVl1 -HW0 -HV12 -HK12 -IH2 /2 width=1 by fsupp_ygt/ #W2 #HVW2 #HW02
+ elim (lift_total W2 0 (i+1))
+ /4 width=11 by cpcs_lift, lsstas_ldef, ex2_intro/
+ | elim (IH1 … HVl1 … HW0 … HVX0 … HK12) -IH1 -HVl1 -HW0 -HVX0 -HK12 -IH2 -V2 /2 width=1 by fsupp_ygt/ -l1 #W2 #HXW2 #HW02
+ elim (lift_total W2 0 (i+1))
+ /3 width=11 by cpcs_lift, lsstas_lift, ex2_intro/
+ ]
+| #p #_ #_ #HT0 #H1 destruct -IH4 -IH3 -IH2 -IH1
+ elim (snv_inv_gref … H1)
+| #a #I #V1 #T1 #HG0 #HL0 #HT0 #H1 #l1 #l2 #Hl21 #Hl1 #X2 #H2 #X3 #H3 #L2 #HL12 destruct -IH4 -IH3 -IH2
+ elim (snv_inv_bind … H1) -H1 #_ #HT1
+ lapply (da_inv_bind … Hl1) -Hl1 #Hl1
+ elim (lsstas_inv_bind1 … H2) -H2 #U1 #HTU1 #H destruct
+ elim (cpr_inv_bind1 … H3) -H3 *
+ [ #V2 #T2 #HV12 #HT12 #H destruct
+ elim (IH1 … Hl1 … HTU1 … HT12 (L2.ⓑ{I}V2)) -IH1 -Hl1 -HTU1 -HT12 /2 width=1 by fsupp_ygt, lpr_pair/ -T1
+ /4 width=5 by cpcs_bind2, lpr_cpr_conf, lsstas_bind, ex2_intro/
+ | #T3 #HT13 #HXT3 #H1 #H2 destruct
+ elim (IH1 … Hl1 … HTU1 … HT13 (L2.ⓓV1)) -IH1 -Hl1 -HTU1 -HT13 /2 width=1 by fsupp_ygt, lpr_pair/ -T1 -HL12 #U3 #HTU3 #HU13
+ elim (lsstas_inv_lift1 … HTU3 L2 … HXT3) -T3
+ /5 width=8 by cpcs_cpr_strap1, cpcs_bind1, cpr_zeta, ldrop_ldrop, ex2_intro/
+ ]
+| #V1 #T1 #HG0 #HL0 #HT0 #H1 #l1 #l2 #Hl21 #Hl1 #X2 #H2 #X3 #H3 #L2 #HL12 destruct
+ elim (snv_inv_appl … H1) -H1 #a #W1 #W10 #U10 #l0 #HV1 #HT1 #Hl0 #HVW1 #HW10 #HTU10
+ lapply (da_inv_flat … Hl1) -Hl1 #Hl1
+ elim (lsstas_inv_appl1 … H2) -H2 #U1 #HTU1 #H destruct
+ elim (cpr_inv_appl1 … H3) -H3 *
+ [ #V2 #T2 #HV12 #HT12 #H destruct -a -l0 -W1 -W10 -U10 -HV1 -IH4 -IH3 -IH2
+ elim (IH1 … Hl1 … HTU1 … HT12 … HL12) -IH1 -Hl1 -HTU1
+ /4 width=5 by fsupp_ygt, cpcs_flat, lpr_cpr_conf, lsstas_appl, ex2_intro/
+ | #b #V2 #W2 #W3 #T2 #T3 #HV12 #HW23 #HT23 #H1 #H2 destruct
+ elim (snv_inv_bind … HT1) -HT1 #HW2 #HT2
+ lapply (da_inv_bind … Hl1) -Hl1 #Hl1
+ elim (lsstas_inv_bind1 … HTU1) -HTU1 #U2 #HTU2 #H destruct
+ elim (cpds_inv_abst1 … HTU10) -HTU10 #W0 #U0 #HW20 #_ #H destruct
+ lapply (cprs_div … HW10 … HW20) -W0 #HW12
+ lapply (ssta_da_conf … HVW1 … Hl0) <minus_plus_m_m #H
+ elim (snv_fwd_da … HW2) #l #Hl
+ lapply (IH4 … HV1 … 1 … Hl0 W1 ?) /2 width=1 by fsupp_ygt, ssta_lsstas/ #HW1
+ lapply (da_cpcs_aux … IH3 IH2 … H … Hl … HW12) // -H
+ /3 width=5 by ygt_yprs_trans, fsupp_ygt, ssta_yprs/ #H destruct
+ lapply (IH3 … HV12 … HL12) /2 width=1 by fsupp_ygt/ #HV2
+ lapply (IH2 … Hl0 … HV12 … HL12) /2 width=1 by fsupp_ygt/ #HV2l
+ elim (IH1 … 1 … Hl0 … W1 … HV12 … HL12) /2 width=1 by fsupp_ygt, ssta_lsstas/ -HVW1 #W4 #H #HW14
+ lapply (lsstas_inv_SO … H) #HV2W4
+ lapply (ssta_da_conf … HV2W4 … HV2l) <minus_plus_m_m #HW4l
+ lapply (IH4 … HV2 … HV2l … H) -H /3 width=5 by ygt_yprs_trans, fsupp_ygt, cpr_lpr_yprs/ #HW4
+ lapply (IH3 … HW23 … HL12) /2 width=1 by fsupp_ygt/ #HW3
+ lapply (IH2 … Hl … HW23 … HL12) /2 width=1 by fsupp_ygt/ #HW3l
+ elim (IH1 … Hl1 … HTU2 … HT23 (L2.ⓛW3)) -HTU2 /2 width=1 by fsupp_ygt, lpr_pair/ #U3 #HTU3 #HU23
+ lapply (cpcs_cpr_strap1 … HW12 … HW23) #H
+ lapply (lpr_cpcs_conf … HL12 … H) -H #H
+ lapply (cpcs_canc_sn … HW14 H) -H #HW43
+ elim (lsubsv_lsstas_trans … HTU3 … Hl21 … (L2.ⓓⓝW3.V2)) -HTU3
+ [ #U4 #HT3U4 #HU43 -HW12 -HW3 -HW3l -W4 -IH2 -IH3 -IH4
+ @(ex2_intro … (ⓓ{b}ⓝW3.V2.U4)) /2 width=1 by lsstas_bind/ -HT3U4
+ @(cpcs_canc_dx … (ⓓ{b}ⓝW3.V2.U3)) /2 width=1 by cpcs_bind_dx/ -HU43
+ @(cpcs_cpr_strap1 … (ⓐV2.ⓛ{b}W3.U3)) /2 width=1 by cpr_beta/
+ /4 width=3 by cpcs_flat, cpcs_bind2, lpr_cpr_conf/
+ | -U3
+ @(lsubsv_abbr … l) /3 width=7 by fsupp_ygt/
+ #W #W0 #l0 #Hl0 #HV2W #HW30
+ lapply (lsstas_ssta_conf_pos … HV2W4 … HV2W) -HV2W #HW4W
+ @(lsstas_cpcs_lpr_aux … IH3 IH2 IH1 … Hl0 … HW4W … Hl0 … HW30 … HW43) //
+ [ /3 width=9 by ygt_yprs_trans, fsupp_ygt, cpr_lpr_ssta_yprs/
+ | /3 width=5 by ygt_yprs_trans, fsupp_ygt, cpr_lpr_yprs/
+ ]
+ | -IH1 -IH3 -IH4 /3 width=9 by fsupp_ygt, lpr_pair/
+ ]
+ | #b #V0 #V2 #W0 #W2 #T0 #T2 #HV10 #HV02 #HW02 #HT02 #H1 #H2 destruct -a -l0 -W1 -W10 -HV1 -IH4 -IH3 -IH2
+ elim (snv_inv_bind … HT1) -HT1 #_ #HT0
+ lapply (da_inv_bind … Hl1) -Hl1 #Hl1
+ elim (lsstas_inv_bind1 … HTU1) -HTU1 #U0 #HTU0 #H destruct
+ elim (IH1 … Hl1 … HTU0 … HT02 (L2.ⓓW2)) -IH1 -Hl1 -HTU0 /2 width=1 by fsupp_ygt, lpr_pair/ -T0 #U2 #HTU2 #HU02
+ lapply (lpr_cpr_conf … HL12 … HV10) -HV10 #HV10
+ lapply (lpr_cpr_conf … HL12 … HW02) -L1 #HW02
+ lapply (cpcs_bind2 b … HW02 … HU02) -HW02 -HU02 #HU02
+ lapply (cpcs_flat … HV10 … HU02 Appl) -HV10 -HU02 #HU02
+ lapply (cpcs_cpr_strap1 … HU02 (ⓓ{b}W2.ⓐV2.U2) ?)
+ /4 width=3 by lsstas_appl, lsstas_bind, cpr_theta, ex2_intro/
+ ]
+| #W1 #T1 #HG0 #HL0 #HT0 #H1 #l1 #l2 @(nat_ind_plus … l2) -l2 [ #_ | #l2 #_ #Hl21 ] #Hl1 #X2 #H2 #X3 #H3 #L2 #HL12 destruct -IH4 -IH3 -IH2
+ [ lapply (lsstas_inv_O … H2) -H2 #H destruct -IH1 -H1 -l1 /4 width=5 by lpr_cpcs_conf, cpr_cpcs_dx, ex2_intro/ ]
+ elim (snv_inv_cast … H1) -H1 #U1 #l #_ #HT1 #_ #_ #_ -U1 -l
+ lapply (da_inv_flat … Hl1) -Hl1 #Hl1
+ lapply (lsstas_inv_cast1 … H2) -H2 #HTU1
+ elim (cpr_inv_cast1 … H3) -H3
+ [ * #U2 #T2 #_ #HT12 #H destruct
+ elim (IH1 … Hl1 … HTU1 … HT12 … HL12) -IH1 -Hl1 -HTU1 -HL12
+ /3 width=3 by fsupp_ygt, lsstas_cast, ex2_intro/
+ | #HT1X3
+ elim (IH1 … Hl1 … HTU1 … HT1X3 … HL12) -IH1 -Hl1 -HTU1 -HL12
+ /2 width=3 by fsupp_ygt, 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/dynamic/snv_lift.ma".
-include "basic_2/dynamic/snv_cpcs.ma".
-
-(* STRATIFIED NATIVE VALIDITY FOR TERMS *************************************)
-
-(* Properties on stratified static type assignment for terms ****************)
-
-fact snv_ssta_aux: ∀h,g,G0,L0,T0.
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_ssta_cpr_lpr h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_ssta h g G1 L1 T1) →
- ∀G1,L1,T1. G0 = G1 → L0 = L1 → T0 = T1 → IH_snv_ssta h g G1 L1 T1.
-#h #g #G0 #L0 #T0 #IH3 #IH2 #IH1 #G1 #L1 * * [||||*]
-[ #k #HG0 #HL0 #HT0 #_ #X #l #H2 destruct -IH3 -IH2 -IH1
- elim (ssta_inv_sort1 … H2) -H2 #_ #H destruct //
-| #i #HG0 #HL0 #HT0 #H1 #X #l #H2 destruct -IH3 -IH2
- elim (snv_inv_lref … H1) -H1 #I #K1 #V1 #HLK1 #HV1
- elim (ssta_inv_lref1 … H2) -H2 * #K0 #V0 #W1 [| #l ] #H #HVW1 #HX [| #_ ]
- lapply (ldrop_mono … H … HLK1) -H #H destruct
- lapply (fsupp_lref … G1 … HLK1) #H
- lapply (ldrop_fwd_ldrop2 … HLK1) -HLK1 /4 width=7/
-| #p #HG0 #HL0 #HT0 #H1 #X #l #H2 destruct -IH3 -IH2 -IH1
- elim (snv_inv_gref … H1)
-| #a #I #V1 #T1 #HG0 #HL0 #HT0 #H1 #X #l #H2 destruct -IH3 -IH2
- elim (snv_inv_bind … H1) -H1 #HV1 #HT1
- elim (ssta_inv_bind1 … H2) -H2 #U1 #HTU1 #H destruct /4 width=5/
-| #V1 #T1 #HG0 #HL0 #HT0 #H1 #X #l #H2 destruct
- elim (snv_inv_appl … H1) -H1 #a #W1 #W0 #T0 #l0 #HV1 #HT1 #HVW1 #HW10 #HT10
- elim (ssta_inv_appl1 … H2) -H2 #U1 #HTU1 #H destruct
- lapply (IH1 … HT1 … HTU1) -IH1 /2 width=1/ #HU1
- elim (ssta_cpds_aux … IH3 IH2 … HTU1 … HT10) -IH3 -IH2 // /2 width=2/ -T1 #U #X #HU1U #H #HU0
- elim (sstas_inv_bind1 … H) -H #U0 #HTU0 #H destruct
- elim (cpcs_inv_abst2 … HU0) -HU0 #W2 #U2 #HU2 #HU02
- elim (cprs_inv_abst … HU02) -HU02 #HW02 #_
- lapply (cprs_trans … HW10 … HW02) -W0 /3 width=10 by snv_appl, ex2_intro/ (**) (* auto is too slow without trace *)
-| #W1 #T1 #HG0 #HL0 #HT0 #H1 #X #l #H2 destruct -IH3 -IH2
- elim (snv_inv_cast … H1) -H1 #U1 #l0 #HW1 #HT1 #HTU1 #HUW1
- lapply (ssta_inv_cast1 … H2) -H2 /3 width=5/
-]
-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/computation/cpds_cpds.ma".
-include "basic_2/dynamic/snv_cpcs.ma".
-include "basic_2/dynamic/lsubsv_ssta.ma".
-
-(* STRATIFIED NATIVE VALIDITY FOR TERMS *************************************)
-
-(* Properties on sn parallel reduction for local environments ***************)
-
-fact ssta_cpr_lpr_aux: ∀h,g,G0,L0,T0.
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_ssta h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >[h, g] ⦃G1, L1, T1⦄ → IH_ssta_cpr_lpr h g G1 L1 T1) →
- ∀G1,L1,T1. G0 = G1 → L0 = L1 → T0 = T1 → IH_ssta_cpr_lpr h g G1 L1 T1.
-#h #g #G0 #L0 #T0 #IH3 #IH2 #IH1 #G1 #L1 * * [|||| *]
-[ #k #_ #_ #_ #_ #X2 #l #H2 #X3 #H3 #L2 #HL12 -IH3 -IH2 -IH1
- elim (ssta_inv_sort1 … H2) -H2 #Hkl #H destruct
- >(cpr_inv_sort1 … H3) -X3 /4 width=6/
-| #i #HG0 #HL0 #HT0 #H1 #X2 #l #H2 #X3 #H3 #L2 #HL12 destruct -IH3 -IH2
- elim (snv_inv_lref … H1) -H1 #I0 #K0 #V0 #H #HV1
- elim (ssta_inv_lref1 … H2) -H2 * #K1
- #V1 #W1 [| #l0 ] #HLK1 #HVW1 #HWU1 [| #H destruct ]
- lapply (ldrop_mono … H … HLK1) -H #H destruct
- lapply (fsupp_lref … G1 … HLK1) #HKV1
- elim (lpr_ldrop_conf … HLK1 … HL12) -HL12 #X #H #HLK2
- elim (lpr_inv_pair1 … H) -H #K2 #V2 #HK12 #HV12 #H destruct
- lapply (ldrop_fwd_ldrop2 … HLK2) #H2
- elim (cpr_inv_lref1 … H3) -H3
- [1,3: #H destruct -HLK1
- |2,4: * #K0 #V0 #W0 #H #HVW0 #HW0
- lapply (ldrop_mono … H … HLK1) -H -HLK1 #H destruct
- ]
- [ elim (IH1 … HVW1 … HV12 … HK12) -IH1 -HVW1 // [2: /2 width=1/ ] -K1 -V1 #W2 #HVW2 #HW12
- elim (lift_total W2 0 (i+1)) #U2 #HWU2
- lapply (cpcs_lift … H2 … HWU1 … HWU2 HW12) -H2 -W1 /3 width=6/
- | elim (IH1 … HVW1 … HV12 … HK12) -IH1 -HVW1 // [2: /2 width=1/ ] #W2 #HVW2 #_
- elim (lift_total V2 0 (i+1)) #U2 #HVU2
- lapply (lpr_cpr_conf … HK12 … HV12) -HV12 #HV12
- lapply (cpcs_lift … H2 … HWU1 … HVU2 HV12) -H2 -V1 /3 width=6/
- | elim (IH1 … HVW1 … HVW0 … HK12) -IH1 -HVW1 // [2: /2 width=1/ ] -K1 -V1 #W2 #HVW2 #HW12
- elim (lift_total W2 0 (i+1)) #U2 #HWU2
- lapply (ssta_lift … HVW2 … H2 … HW0 … HWU2) -HVW2 -HW0
- lapply (cpcs_lift … H2 … HWU1 … HWU2 HW12) -H2 -W1 /3 width=6/
- ]
-| #p #_ #_ #HT0 #H1 destruct -IH3 -IH2 -IH1
- elim (snv_inv_gref … H1)
-| #a #I #V1 #T1 #HG0 #HL0 #HT0 #H1 #X2 #l #H2 #X3 #H3 #L2 #HL12 destruct -IH3 -IH2
- elim (snv_inv_bind … H1) -H1 #_ #HT1
- elim (ssta_inv_bind1 … H2) -H2 #U1 #HTU1 #H destruct
- elim (cpr_inv_bind1 … H3) -H3 *
- [ #V2 #T2 #HV12 #HT12 #H destruct
- elim (IH1 … HTU1 … HT12 (L2.ⓑ{I}V2)) // [2,3: /2 width=1/ ] -T1
- lapply (lpr_cpr_conf … HL12 … HV12) -L1 /3 width=5/
- | #T2 #HT12 #HT2 #H1 #H2 destruct
- elim (IH1 … HTU1 … HT12 (L2.ⓓV1)) -IH1 -HTU1 // [2,3: /2 width=1/ ] -T1 -HL12 #U2 #HTU2 #HU12
- elim (ssta_inv_lift1 … HTU2 … HT2) -T2 [3: /2 width=1/ |2: skip ] #U #HXU #HU2
- lapply (cpcs_bind1 true … V1 V1 … HU12) // -HU12 #HU12
- lapply (cpcs_cpr_strap1 … HU12 U ?) -HU12 /2 width=3/
- ]
-| #V1 #T1 #HG0 #HL0 #HT0 #H1 #X2 #l #H2 #X3 #H3 #L2 #HL12 destruct
- elim (snv_inv_appl … H1) -H1 #a #W1 #W10 #U10 #l0 #HV1 #HT1 #HVW1 #HW10 #HTU10
- elim (ssta_inv_appl1 … H2) -H2 #U1 #HTU1 #H destruct
- elim (cpr_inv_appl1 … H3) -H3 *
- [ #V2 #T2 #HV12 #HT12 #H destruct -a -l0 -W1 -W10 -U10 -HV1 -IH3 -IH2
- elim (IH1 … HTU1 … HT12 … HL12) -IH1 -HTU1 // [2: /2 width=1/ ] -T1 #U2 #HTU2 #HU12
- lapply (lpr_cpr_conf … HL12 … HV12) -L1 /3 width=5/
- | #b #V2 #W2 #W20 #T2 #T20 #HV12 #HW20 #HT20 #H1 #H2 destruct
- elim (snv_inv_bind … HT1) -HT1 #HW2 #HT2
- elim (ssta_inv_bind1 … HTU1) -HTU1 #U2 #HTU2 #H destruct
- elim (cpds_inv_abst1 … HTU10) -HTU10 #W0 #U0 #HW0 #_ #H destruct
- lapply (cprs_div … HW10 … HW0) -W0 #HW12
- elim (ssta_fwd_correct … HVW1) <minus_plus_m_m #X1 #HWX1
- elim (snv_fwd_ssta … HW2) #l1 #V22 #HWV22
- lapply (IH3 … HVW1) -IH3 // [ /2 width=1/ ] #HW1
- elim (ssta_cpcs_lpr_aux … IH2 IH1 … HWX1 … HWV22 … L1) -HWX1 //
- [2: /2 width=1/
- |3: /4 width=5 by ygt_yprs_trans, fsupp_ygt, sstas_yprs, ssta_sstas/
- ] #H #_ destruct -X1
- lapply (IH2 … HV1 … HV12 … HL12) [ /2 width=1/ ] #HV2
- lapply (IH2 … HW2 … HW20 … HL12) [ /2 width=1/ ] -IH2 #H2W20
- elim (IH1 … HVW1 … HV12 … HL12) -HVW1 // -HV1 [2: /2 width=1/ ] #W12 #HVW12 #HW112
- elim (IH1 … HWV22 … HW20 … HL12) -HWV22 // -HW2 [2: /2 width=1/ ] #V20 #HWV20 #_
- elim (IH1 … HTU2 … HT20 (L2.ⓛW20)) -IH1 -HTU2 -HT20 // [2,3: /2 width=1/ ] -HT2 #U20 #HTU20 #HU20
- lapply (lpr_cpcs_conf … HL12 … HW12) -HW12 #HW12
- lapply (lpr_cpr_conf … HL12 … HW20) -HW20 #HW20
- lapply (lpr_cpr_conf … HL12 … HV12) -L1 #HV12
- lapply (cpcs_canc_sn … HW12 HW112) -W1 #HW12
- lapply (cpcs_canc_sn … HW12 HW20) -HW12 #HW12
- elim (lsubsv_ssta_trans … HTU20 (L2.ⓓⓝW20.V2)) -HTU20
- [ #U #HTU20 #HUU20 -HVW12 -HWV20 -HW12
- @(ex2_intro … (ⓓ{b}ⓝW20.V2.U)) [ /3 width=1/ ] -HTU20
- @(cpcs_canc_dx … (ⓓ{b}ⓝW20.V2.U20)) [2: /2 width=1/ ] -HUU20
- @(cpcs_cpr_strap1 … (ⓐV2.ⓛ{b}W20.U20)) [2: /2 width=1/ ]
- /3 by cpcs_bind2, cpcs_flat/
- | -HU20 -HW20 -HV12 /3 width=5 by lsubsv_abbr, snv_cast/
- ]
- | #b #V0 #V2 #W0 #W2 #T0 #T2 #HV10 #HV02 #HW02 #HT02 #H1 #H2 destruct -a -l0 -W1 -W10 -HV1 -IH3 -IH2
- elim (ssta_inv_bind1 … HTU1) -HTU1 #U0 #HTU0 #H destruct
- elim (snv_inv_bind … HT1) -HT1 #_ #HT0
- elim (IH1 … HTU0 … HT02 (L2.ⓓW2)) -IH1 -HTU0 // [2,3: /2 width=1/ ] -T0 #U2 #HTU2 #HU02
- lapply (lpr_cpr_conf … HL12 … HV10) -HV10 #HV10
- lapply (lpr_cpr_conf … HL12 … HW02) -L1 #HW02
- lapply (cpcs_bind2 b … HW02 … HU02) -HW02 -HU02 #HU02
- lapply (cpcs_flat … HV10 … HU02 Appl) -HV10 -HU02 #HU02
- lapply (cpcs_cpr_strap1 … HU02 (ⓓ{b}W2.ⓐV2.U2) ?) [ /2 width=3/ ] -V0 /4 width=3/
- ]
-| #U0 #T1 #HG0 #HL0 #HT0 #H1 #X2 #l #H2 #X3 #H3 #L2 #HL12 destruct -IH3 -IH2
- elim (snv_inv_cast … H1) -H1 #T0 #l0 #_ #HT1 #HT10 #_
- lapply (ssta_inv_cast1 … H2) -H2 #HTU1
- elim (ssta_mono … HT10 … HTU1) -HT10 #H1 #H2 destruct
- elim (cpr_inv_cast1 … H3) -H3
- [ * #U2 #T2 #_ #HT12 #H destruct
- elim (IH1 … HTU1 … HT12 … HL12) -IH1 -HTU1 -HL12 // [2: /2 width=1/ ] -T1 -U0 /3 width=3/
- | #HT1X3
- elim (IH1 … HTU1 … HT1X3 … HL12) -IH1 -HTU1 -HL12 // [2: /2 width=1/ ] -U0 -T1 /2 width=3/
- ]
-]
-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/sstas_lift.ma".
-include "basic_2/dynamic/snv.ma".
-
-(* STRATIFIED NATIVE VALIDITY FOR TERMS *************************************)
-
-(* Forward_lemmas on iterated stratified static type assignment for terms ***)
-
-lemma snv_sstas_fwd_correct: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ¡[h, g] → ⦃G, L⦄ ⊢ T1 •* [h, g] T2 →
- ∃∃U2,l. ⦃G, L⦄ ⊢ T2 •[h, g] ⦃l, U2⦄.
-#h #g #G #L #T1 #T2 #HT1 #HT12
-elim (snv_fwd_ssta … HT1) -HT1 /2 width=5 by sstas_fwd_correct/
-qed-.
]
qed.
-lemma sstas_ygt: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 •*[h, g] T2 → (T1 = T2 → ⊥) →
- ⦃G, L, T1⦄ >[h, g] ⦃G, L, T2⦄.
-#h #g #G #L #T1 #T2 #H @(sstas_ind … H) -T2
-[ #H elim H //
-| #T #T2 #l #_ #HT2 #IHT1 #HT12 -HT12
- elim (term_eq_dec T1 T) #H destruct
- [ -IHT1 /3 width=2/
- | lapply (IHT1 … H) -IHT1 -H #HT1
- @(ygt_strap1 … HT1) -HT1 /2 width=2/
- ]
-]
-qed.
-
lemma lsubsv_ygt: ∀h,g,G,L1,L2,T. G ⊢ L2 ¡⊑[h, g] L1 → (L1 = L2 → ⊥) →
⦃G, L1, T⦄ >[h, g] ⦃G, L2, T⦄.
/4 width=1/ 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/lsstas_lift.ma".
+include "basic_2/dynamic/ygt.ma".
+
+(* "BIG TREE" PARALLEL COMPUTATION FOR CLOSURES *****************************)
+
+(* Advanced properties ******************************************************)
+
+lemma lsstas_ygt: ∀h,g,G,L,T1,T2,l2. ⦃G, L⦄ ⊢ T1 •*[h, g, l2] T2 → (T1 = T2 → ⊥) →
+ ∀l1. l2 ≤ l1 → ⦃G, L⦄ ⊢ T1 ▪[h, g] l1 → ⦃G, L, T1⦄ >[h, g] ⦃G, L, T2⦄.
+#h #g #G #L #T1 #T2 #l2 #H @(lsstas_ind_dx … H) -l2 -T2
+[ #H elim H //
+| #l2 #T #T2 #HT1 #HT2 #IHT1 #HT12 #l1 #Hl21
+ elim (term_eq_dec T1 T) #H destruct [ -IHT1 |]
+ [ elim (le_inv_plus_l … Hl21) -Hl21 #_ #Hl1
+ >(plus_minus_m_m … Hl1) -Hl1 /3 width=5 by ysc_ssta, ygt_inj/
+ | #Hl1 >commutative_plus in Hl21; #Hl21
+ elim (le_inv_plus_l … Hl21) -Hl21 #Hl12 #Hl21
+ lapply (lsstas_da_conf … HT1 … Hl1) -HT1
+ >(plus_minus_m_m … Hl12) -Hl12
+ /4 width=5 by ypr_ssta, ygt_strap1/
+ ]
+]
+qed.
| ypr_fsup : ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊃ ⦃G2, L2, T2⦄ → ypr h g G1 L1 T1 G2 L2 T2
| ypr_lpr : ∀L2. ⦃G1, L1⦄ ⊢ ➡ L2 → ypr h g G1 L1 T1 G1 L2 T1
| ypr_cpr : ∀T2. ⦃G1, L1⦄ ⊢ T1 ➡ T2 → ypr h g G1 L1 T1 G1 L1 T2
-| ypr_ssta : â\88\80T2,l. â¦\83G1, L1â¦\84 â\8a¢ T1 â\80¢[h, g] â¦\83l+1, T2â¦\84 → ypr h g G1 L1 T1 G1 L1 T2
+| ypr_ssta : â\88\80T2,l. â¦\83G1, L1â¦\84 â\8a¢ T1 â\96ª[h, g] l+1 â\86\92 â¦\83G1, L1â¦\84 â\8a¢ T1 â\80¢[h, g] T2 → ypr h g G1 L1 T1 G1 L1 T2
| ypr_lsubsv: ∀L2. G1 ⊢ L2 ¡⊑[h, g] L1 → ypr h g G1 L1 T1 G1 L2 T1
.
#h #g #G #L1 #L2 #T #H @(lprs_ind … H) -L2 // /3 width=5 by ypr_lpr, yprs_strap1/
qed.
-lemma sstas_yprs: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 •*[h, g] T2 →
- ⦃G, L, T1⦄ ≥[h, g] ⦃G, L, T2⦄.
-#h #g #G #L #T1 #T2 #H @(sstas_ind … H) -T2 // /3 width=5 by ypr_ssta, yprs_strap1/
-qed.
-
lemma lsubsv_yprs: ∀h,g,G,L1,L2,T. G ⊢ L2 ¡⊑[h, g] L1 → ⦃G, L1, T⦄ ≥[h, g] ⦃G, L2, T⦄.
/3 width=1/ qed.
-lemma cprs_lpr_yprs: ∀h,g,G,L1,L2,T1,T2. ⦃G, L1⦄ ⊢ T1 ➡* T2 → ⦃G, L1⦄ ⊢ ➡ L2 →
- ⦃G, L1, T1⦄ ≥[h, g] ⦃G, L2, T2⦄.
-/3 width=5 by yprs_strap1, ypr_lpr, cprs_yprs/
-qed.
+lemma cpr_lpr_yprs: ∀h,g,G,L1,L2,T1,T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ⦃G, L1⦄ ⊢ ➡ L2 →
+ ⦃G, L1, T1⦄ ≥[h, g] ⦃G, L2, T2⦄.
+/4 width=5 by yprs_strap1, ypr_lpr, ypr_cpr/ qed.
+
+lemma ssta_yprs: ∀h,g,G,L,T,U,l.
+ ⦃G, L⦄ ⊢ T ▪[h, g] l+1 → ⦃G, L⦄ ⊢ T •[h, g] U →
+ ⦃G, L, T⦄ ≥[h, g] ⦃G, L, U⦄.
+/3 width=2 by ypr_yprs, ypr_ssta/ 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/lsstas_lift.ma".
+include "basic_2/dynamic/yprs.ma".
+
+(* "BIG TREE" PARALLEL COMPUTATION FOR CLOSURES *****************************)
+
+(* Advanced properties ******************************************************)
+
+lemma lsstas_yprs: ∀h,g,G,L,T1,T2,l2. ⦃G, L⦄ ⊢ T1 •*[h, g, l2] T2 →
+ ∀l1. l2 ≤ l1 → ⦃G, L⦄ ⊢ T1 ▪[h, g] l1 → ⦃G, L, T1⦄ ≥[h, g] ⦃G, L, T2⦄.
+#h #g #G #L #T1 #T2 #l2 #H @(lsstas_ind_dx … H) -l2 -T2 //
+#l2 #T #T2 #HT1 #HT2 #IHT1 #l1 >commutative_plus #Hl21 #Hl1
+elim (le_inv_plus_l … Hl21) -Hl21 #Hl12 #Hl21
+lapply (lsstas_da_conf … HT1 … Hl1) -HT1
+>(plus_minus_m_m … Hl12) -Hl12
+/3 width=5 by ypr_ssta, yprs_strap1/
+qed.
(* "BIG TREE" PARALLEL COMPUTATION FOR CLOSURES *****************************)
+(* Main properties **********************************************************)
+
theorem yprs_trans: ∀h,g. tri_transitive … (yprs h g).
/2 width=5/ qed-.
+
+(* Advanced properties ******************************************************)
+
+lemma cpr_lpr_ssta_yprs: ∀h,g,G,L1,L2,T1,T2,U2,l2.
+ ⦃G, L1⦄ ⊢ T1 ➡ T2 → ⦃G, L1⦄ ⊢ ➡ L2 →
+ ⦃G, L2⦄ ⊢ T2 ▪[h, g] l2+1 → ⦃G, L2⦄ ⊢ T2 •[h, g] U2 →
+ ⦃G, L1, T1⦄ ≥[h, g] ⦃G, L2, U2⦄.
+/3 width=5 by yprs_trans, cpr_lpr_yprs, ssta_yprs/ qed.
inductive ysc (h) (g) (G1) (L1) (T1): relation3 genv lenv term ≝
| ysc_fsup : ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊃ ⦃G2, L2, T2⦄ → ysc h g G1 L1 T1 G2 L2 T2
| ysc_cpr : ∀T2. ⦃G1, L1⦄ ⊢ T1 ➡ T2 → (T1 = T2 → ⊥) → ysc h g G1 L1 T1 G1 L1 T2
-| ysc_ssta : â\88\80T2,l. â¦\83G1, L1â¦\84 â\8a¢ T1 â\80¢[h, g] â¦\83l+1, T2â¦\84 → ysc h g G1 L1 T1 G1 L1 T2
+| ysc_ssta : â\88\80T2,l. â¦\83G1, L1â¦\84 â\8a¢ T1 â\96ª[h, g] l+1 â\86\92 â¦\83G1, L1â¦\84 â\8a¢ T1 â\80¢[h, g] T2 → ysc h g G1 L1 T1 G1 L1 T2
| ysc_lsubsv: ∀L2. G1 ⊢ L2 ¡⊑[h, g] L1 → (L1 = L2 → ⊥) → ysc h g G1 L1 T1 G1 L2 T1
.
(* Inversion lemmas on "big tree" parallel reduction for closures ***********)
lemma ypr_inv_ysc: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≽[h, g] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄ ∨
+ ⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄ ∨
∧∧ G1 = G2 & ⦃G1, L1⦄ ⊢ ➡ L2 & T1 = T2.
#h #g #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2 /3 width=1/ /3 width=2/
[ #T2 #HT12 elim (term_eq_dec T1 T2) #H destruct /3 width=1/ /4 width=1/
--- /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/dpconvstar_6.ma".
+include "basic_2/unfold/lsstas.ma".
+include "basic_2/equivalence/cpcs.ma".
+
+(* DECOMPOSED EXTENDED PARALLEL EQUIVALENCE FOR TERMS ***********************)
+
+definition cpes: ∀h. sd h → relation4 genv lenv term term ≝
+ λh,g,G,L,T1,T2.
+ ∃∃T,l1,l2. l2 ≤ l1 & ⦃G, L⦄ ⊢ T1 ▪[h, g] l1 & ⦃G, L⦄ ⊢ T1 •*[h, g, l2] T & ⦃G, L⦄ ⊢ T ⬌* T2.
+
+interpretation "decomposed extended parallel equivalence (term)"
+ 'DPConvStar h g G L T1 T2 = (cpes h g G L T1 T2).
+
+(* Basic properties *********************************************************)
+
+lemma ssta_cpcs_cpes: ∀h,g,G,L,T1,T,T2,l. ⦃G, L⦄ ⊢ T1 ▪[h, g] l+1 → ⦃G, L⦄ ⊢ T1 •[h, g] T →
+ ⦃G, L⦄ ⊢ T ⬌* T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, g] T2.
+/3 width=7/ qed.
+
+lemma lsstas_cpes: ∀h,g,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 ▪[h, g] l → ⦃G, L⦄ ⊢ T1 •*[h, g, l] T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, g] T2.
+/2 width=7/ qed.
+
+lemma cpcs_cpes: ∀h,g,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 ▪[h, g] l → ⦃G, L⦄ ⊢ T1 ⬌* T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, g] T2.
+/2 width=7/ qed.
+
+lemma cpes_refl: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l → ⦃G, L⦄ ⊢ T •*⬌*[h, g] T.
+/2 width=2/ qed.
+
+lemma cpes_strap1: ∀h,g,G,L,T1,T,T2.
+ ⦃G, L⦄ ⊢ T1 •*⬌*[h, g] T → ⦃G, L⦄ ⊢ T ⬌ T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, g] T2.
+#h #g #G #L #T1 #T #T2 * /3 width=9/
+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/computation/cpds.ma".
+include "basic_2/equivalence/cpcs_cpcs.ma".
+include "basic_2/equivalence/cpes.ma".
+
+(* DECOMPOSED EXTENDED PARALLEL EQUIVALENCE FOR TERMS ***********************)
+
+(* Advanced properties ******************************************************)
+
+lemma cpds_cpes_dx: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, g] T2.
+#h #g #G #L #T1 #T2 * /3 width=7 by cpcs_cprs_dx, ex4_3_intro/
+qed.
+
+(* Advanced inversion lemmas ************************************************)
+
+lemma cpes_inv_abst2: ∀h,g,a,G,L,W1,T1,T. ⦃G, L⦄ ⊢ T •*⬌*[h, g] ⓛ{a}W1.T1 →
+ ∃∃W2,T2. ⦃G, L⦄ ⊢ T •*➡*[h, g] ⓛ{a}W2.T2 & ⦃G, L⦄ ⊢ ⓛ{a}W1.T1 ➡* ⓛ{a}W2.T2.
+#h #g #a #G #L #W1 #T1 #T * #T0 #l #l0 #Hl0 #Hl #HT0 #H
+elim (cpcs_inv_abst2 … H) -H /3 width=7 by ex4_3_intro, ex2_2_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/ssta_ltpss.ma".
+include "basic_2/unwind/sstas.ma".
+
+(* ITERATED STRATIFIED STATIC TYPE ASSIGNMENTON TERMS ***********************)
+
+(* Properties about parallel unfold *****************************************)
+
+lemma sstas_ltpss_tpss_conf: ∀h,g,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 •*[g] U1 →
+ ∀L2,d,e. L1 ▶* [d, e] L2 →
+ ∀T2. L2 ⊢ T1 ▶* [d, e] T2 →
+ ∃∃U2. ⦃h, L2⦄ ⊢ T2 •*[g] U2 &
+ L2 ⊢ U1 ▶* [d, e] U2.
+#h #g #L1 #T1 #U1 #H @(sstas_ind_alt … H) -T1
+[ #T1 #HUT1 #L2 #d #e #HL12 #U2 #HU12
+ elim (ssta_ltpss_tpss_conf … HUT1 … HL12 … HU12) -HUT1 -HL12 /3 width=3/
+| #T0 #U0 #l0 #HTU0 #_ #IHU01 #L2 #d #e #HL12 #T #HT0
+ elim (ssta_ltpss_tpss_conf … HTU0 … HL12 … HT0) -HTU0 -HT0 #U #HTU #HU0
+ elim (IHU01 … HL12 … HU0) -IHU01 -HL12 -U0 /3 width=4/
+]
+qed.
+
+lemma sstas_ltpss_tps_conf: ∀h,g,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 •*[g] U1 →
+ ∀L2,d,e. L1 ▶* [d, e] L2 →
+ ∀T2. L2 ⊢ T1 ▶ [d, e] T2 →
+ ∃∃U2. ⦃h, L2⦄ ⊢ T2 •*[g] U2 & L2 ⊢ U1 ▶* [d, e] U2.
+/3 width=5/ qed.
+
+lemma sstas_ltpss_conf: ∀h,g,L1,T,U1. ⦃h, L1⦄ ⊢ T •*[g] U1 →
+ ∀L2,d,e. L1 ▶* [d, e] L2 →
+ ∃∃U2. ⦃h, L2⦄ ⊢ T •*[g] U2 & L2 ⊢ U1 ▶* [d, e] U2.
+/2 width=5/ qed.
+
+lemma sstas_tpss_conf: ∀h,g,L,T1,U1. ⦃h, L⦄ ⊢ T1 •*[g] U1 →
+ ∀T2,d,e. L ⊢ T1 ▶* [d, e] T2 →
+ ∃∃U2. ⦃h, L⦄ ⊢ T2 •*[g] U2 & L ⊢ U1 ▶* [d, e] U2.
+/2 width=5/ qed.
+
+lemma sstas_tps_conf: ∀h,g,L,T1,U1. ⦃h, L⦄ ⊢ T1 •*[g] U1 →
+ ∀T2,d,e. L ⊢ T1 ▶ [d, e] T2 →
+ ∃∃U2. ⦃h, L⦄ ⊢ T2 •*[g] U2 & L ⊢ U1 ▶* [d, e] U2.
+/2 width=5/ 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/tpss_tpss.ma".
+include "basic_2/unfold/ltpss_tpss.ma".
+include "basic_2/static/sta_lift.ma".
+
+(* STATIC TYPE ASSIGNMENT ON TERMS ******************************************)
+
+(* Properties about parallel unfold *****************************************)
+
+lemma sta_ltpss_tpss_conf: ∀h,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 • U1 →
+ ∀L2,d,e. L1 ▶* [d, e] L2 →
+ ∀T2. L2 ⊢ T1 ▶* [d, e] T2 →
+ ∃∃U2. ⦃h, L2⦄ ⊢ T2 • U2 & L2 ⊢ U1 ▶* [d, e] U2.
+#h #L1 #T1 #U #H elim H -L1 -T1 -U
+[ #L1 #k1 #L2 #d #e #_ #T2 #H
+ >(tpss_inv_sort1 … H) -H /2 width=3/
+| #L1 #K1 #V1 #W1 #U1 #i #HLK1 #HVW1 #HWU1 #IHVW1 #L2 #d #e #HL12 #T2 #H
+ elim (tpss_inv_lref1 … H) -H [ | -HVW1 ]
+ [ #H destruct
+ elim (lt_or_ge i d) #Hdi [ -HVW1 | ]
+ [ elim (ltpss_ldrop_conf_le … HL12 … HLK1 ?) -L1 /2 width=2/ #X #H #HLK2
+ elim (ltpss_inv_tpss11 … H ?) -H /2 width=1/ #K2 #V2 #HK12 #HV12 #H destruct
+ elim (IHVW1 … HK12 … HV12) -IHVW1 -HK12 -HV12 #W2 #HVW2 #HW12
+ lapply (ldrop_fwd_ldrop2 … HLK2) #H
+ elim (lift_total W2 0 (i+1)) #U2 #HWU2
+ lapply (tpss_lift_ge … HW12 … H … HWU1 … HWU2) // -HW12 -H -HWU1
+ >minus_plus <plus_minus_m_m // -Hdi /3 width=6/
+ | elim (lt_or_ge i (d + e)) #Hide [ -HVW1 | -Hdi -IHVW1 ]
+ [ elim (ltpss_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HLK2
+ elim (ltpss_inv_tpss21 … H ?) -H /2 width=1/ #K2 #V2 #HK12 #HV12 #H destruct
+ elim (IHVW1 … HK12 … HV12) -IHVW1 -HK12 -HV12 #W2 #HVW2 #HW12
+ lapply (ldrop_fwd_ldrop2 … HLK2) #H
+ elim (lift_total W2 0 (i+1)) #U2 #HWU2
+ lapply (tpss_lift_ge … HW12 … H … HWU1 … HWU2) // -HW12 -H -HWU1 >minus_plus #H
+ lapply (tpss_weak … H d e ? ?) [1,2: normalize [ >commutative_plus <plus_minus_m_m // | /2 width=1/ ]] -Hdi -Hide /3 width=6/
+ | lapply (ltpss_ldrop_conf_ge … HL12 … HLK1 ?) -L1 // -Hide /3 width=6/
+ ]
+ ]
+ | * #K2 #V2 #W2 #Hdi #Hide #HLK2 #HVW2 #HWT2
+ elim (ltpss_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HL2K0
+ elim (ltpss_inv_tpss21 … H ?) -H /2 width=1/ #K0 #V0 #HK12 #HV12 #H destruct
+ lapply (ldrop_mono … HL2K0 … HLK2) -HL2K0 #H destruct
+ lapply (ldrop_fwd_ldrop2 … HLK2) -HLK2 #HLK2
+ lapply (tpss_trans_eq … HV12 HVW2) -V2 #HV1W2
+ elim (IHVW1 … HK12 … HV1W2) -IHVW1 -HK12 -HV1W2 #V2 #HWV2 #HW1V2
+ elim (lift_total V2 0 (i+1)) #U2 #HVU2
+ lapply (sta_lift … HWV2 … HLK2 … HWT2 … HVU2) -HWV2 -HWT2 #HTU2
+ lapply (tpss_lift_ge … HW1V2 … HLK2 … HWU1 … HVU2) // -HW1V2 -HLK2 -HWU1 -HVU2 >minus_plus #H
+ lapply (tpss_weak … H d e ? ?) [1,2: normalize [ >commutative_plus <plus_minus_m_m // | /2 width=1/ ]] -Hdi -Hide /2 width=3/
+ ]
+| #L1 #K1 #W1 #V1 #U1 #i #HLK1 #HWV1 #HWU1 #IHWV1 #L2 #d #e #HL12 #T2 #H
+ elim (tpss_inv_lref1 … H) -H [ | -HWV1 -HWU1 -IHWV1 ]
+ [ #H destruct
+ elim (lt_or_ge i d) #Hdi [ -HWV1 ]
+ [ elim (ltpss_ldrop_conf_le … HL12 … HLK1 ?) -L1 /2 width=2/ #X #H #HLK2
+ elim (ltpss_inv_tpss11 … H ?) -H /2 width=1/ #K2 #W2 #HK12 #HW12 #H destruct
+ elim (IHWV1 … HK12 … HW12) -IHWV1 -HK12 #V2 #HWV2 #_
+ lapply (ldrop_fwd_ldrop2 … HLK2) #HLK
+ elim (lift_total W2 0 (i+1)) #U2 #HWU2
+ lapply (tpss_lift_ge … HW12 … HLK … HWU1 … HWU2) // -HW12 -HLK -HWU1
+ >minus_plus <plus_minus_m_m // -Hdi /3 width=6/
+ | elim (lt_or_ge i (d + e)) #Hide [ -HWV1 | -IHWV1 -Hdi ]
+ [ elim (ltpss_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HLK2
+ elim (ltpss_inv_tpss21 … H ?) -H /2 width=1/ #K2 #W2 #HK12 #HW12 #H destruct
+ elim (IHWV1 … HK12 … HW12) -IHWV1 -HK12 #V2 #HWV2 #_
+ lapply (ldrop_fwd_ldrop2 … HLK2) #HLK
+ elim (lift_total W2 0 (i+1)) #U2 #HWU2
+ lapply (tpss_lift_ge … HW12 … HLK … HWU1 … HWU2) // -HW12 -HLK -HWU1 >minus_plus #H
+ lapply (tpss_weak … H d e ? ?) [1,2: normalize [ >commutative_plus <plus_minus_m_m // | /2 width=1/ ]] -Hdi -Hide /3 width=6/
+ | lapply (ltpss_ldrop_conf_ge … HL12 … HLK1 ?) -L1 // -Hide /3 width=6/
+ ]
+ ]
+ | * #K2 #V2 #W2 #Hdi #Hide #HLK2 #_ #_
+ elim (ltpss_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HL2K0
+ elim (ltpss_inv_tpss21 … H ?) -H /2 width=1/ -Hdi -Hide #K0 #V0 #_ #_ #H destruct
+ lapply (ldrop_mono … HL2K0 … HLK2) -HL2K0 -HLK2 #H destruct
+ ]
+| #I #L1 #V1 #T1 #U1 #_ #IHTU1 #L2 #d #e #HL12 #X #H
+ elim (tpss_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
+ elim (IHTU1 … HT12) -IHTU1 -HT12 /2 width=1/ -HL12 /3 width=5/
+| #L1 #V1 #T1 #U1 #_ #IHTU1 #L2 #d #e #HL12 #X #H
+ elim (tpss_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
+ elim (IHTU1 … HT12) -IHTU1 -HT12 // -HL12 /3 width=5/
+| #L1 #W1 #T1 #U1 #_ #IHTU1 #L2 #d #e #HL12 #X #H
+ elim (tpss_inv_flat1 … H) -H #W2 #T2 #HW12 #HT12 #H destruct
+ elim (IHTU1 … HT12) -IHTU1 -HT12 // -HL12 /3 width=3/
+]
+qed.
+
+lemma sta_ltpss_tps_conf: ∀h,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 • U1 →
+ ∀L2,d,e. L1 ▶* [d, e] L2 →
+ ∀T2. L2 ⊢ T1 ▶ [d, e] T2 →
+ ∃∃U2. ⦃h, L2⦄ ⊢ T2 • U2 & L2 ⊢ U1 ▶* [d, e] U2.
+/3 width=5/ qed.
+
+lemma sta_ltpss_conf: ∀h,L1,T,U1. ⦃h, L1⦄ ⊢ T • U1 →
+ ∀L2,d,e. L1 ▶* [d, e] L2 →
+ ∃∃U2. ⦃h, L2⦄ ⊢ T • U2 & L2 ⊢ U1 ▶* [d, e] U2.
+/2 width=5/ qed.
+
+lemma sta_tpss_conf: ∀h,L,T1,U1. ⦃h, L⦄ ⊢ T1 • U1 →
+ ∀T2,d,e. L ⊢ T1 ▶* [d, e] T2 →
+ ∃∃U2. ⦃h, L⦄ ⊢ T2 • U2 & L ⊢ U1 ▶* [d, e] U2.
+/2 width=5/ qed.
+
+lemma sta_tps_conf: ∀h,L,T1,U1. ⦃h, L⦄ ⊢ T1 • U1 →
+ ∀T2,d,e. L ⊢ T1 ▶ [d, e] T2 →
+ ∃∃U2. ⦃h, L⦄ ⊢ T2 • U2 & L ⊢ U1 ▶* [d, e] U2.
+/2 width=5/ qed.
--- /dev/null
+(* 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/
+| #G #L #K #W #i #l #HLK #_ * #V #HWV
+ elim (lift_total W 0 (i+1)) /3 width=7/
+| #a #I #G #L #V #T #l #_ * /3 width=2/
+| * #G #L #V #T #l #_ * /3 width=2/
+]
+qed-.
+
+(* Properties on static type assignment for terms ***************************)
+
+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/
+| #G #L #K #V #W #W0 #i #HLK #_ #_ * /3 width=5/
+| #G #L #K #W #V #W0 #i #HLK #_ #_ * /3 width=5/
+| #a #I #G #L #V #T #U #_ * /3 width=2/
+| #G #L #V #T #U #_ * /3 width=2/
+| #G #L #W #T #U #_ * /3 width=2/
+]
+qed-.
--- /dev/null
+(* Advanced inversion lemmas ************************************************)
+
+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 (sta_da_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
+(* Advanced properties ******************************************************)
+
+lemma sta_da_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/
+| #G #L #K #V #W #W0 #i #HLK #_ #HW0 #IHVW #l #H
+ elim (da_inv_lref … H) -H * #K0 #V0 [2: #l0] #HLK0 #HV0 [ #H0 ]
+ lapply (ldrop_mono … HLK0 … HLK) -HLK0 #H destruct
+ lapply (ldrop_fwd_ldrop2 … HLK) -HLK /3 width=7/
+| #G #L #K #W #V #W0 #i #HLK #HWV #HW0 #_ #l #H
+ elim (da_inv_lref … H) -H * #K0 #V0 [2: #l0] #HLK0 #HV0 [ #H0 ]
+ lapply (ldrop_mono … HLK0 … HLK) -HLK0 #H destruct
+ lapply (ldrop_fwd_ldrop2 … HLK) -HLK /3 width=7/
+| #a #I #G #L #V #T #U #_ #IHTU #l #H
+ lapply (da_inv_bind … H) -H /3 width=1/
+| #G #L #V #T #U #_ #IHTU #l #H
+ lapply (da_inv_flat … H) -H /3 width=1/
+| #G #L #W #T #U #_ #IHTU #l #H
+ lapply (da_inv_flat … H) -H /2 width=1/
+]
+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_7.ma".
+include "basic_2/grammar/genv.ma".
+include "basic_2/relocation/ldrop.ma".
+include "basic_2/static/sd.ma".
+
+(* STRATIFIED STATIC TYPE ASSIGNMENT ON TERMS *******************************)
+
+(* activate genv *)
+inductive ssta (h:sh) (g:sd h): nat → relation4 genv lenv term term ≝
+| ssta_sort: ∀G,L,k,l. deg h g k l → ssta h g l G L (⋆k) (⋆(next h k))
+| ssta_ldef: ∀G,L,K,V,W,U,i,l. ⇩[0, i] L ≡ K. ⓓV → ssta h g l G K V W →
+ ⇧[0, i + 1] W ≡ U → ssta h g l G L (#i) U
+| ssta_ldec: ∀G,L,K,W,V,U,i,l. ⇩[0, i] L ≡ K. ⓛW → ssta h g l G K W V →
+ ⇧[0, i + 1] W ≡ U → ssta h g (l+1) G L (#i) U
+| ssta_bind: ∀a,I,G,L,V,T,U,l. ssta h g l G (L. ⓑ{I} V) T U →
+ ssta h g l G L (ⓑ{a,I}V.T) (ⓑ{a,I}V.U)
+| ssta_appl: ∀G,L,V,T,U,l. ssta h g l G L T U →
+ ssta h g l G L (ⓐV.T) (ⓐV.U)
+| ssta_cast: ∀G,L,W,T,U,l. ssta h g l G L T U → ssta h g l G L (ⓝW.T) U
+.
+
+interpretation "stratified static type assignment (term)"
+ 'StaticType h g G L T U l = (ssta h g l G L T U).
+
+definition ssta_step: ∀h. sd h → relation4 genv lenv term term ≝
+ λh,g,G,L,T,U. ∃l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l+1, U⦄.
+
+(* Basic inversion lemmas ************************************************)
+
+fact ssta_inv_sort1_aux: ∀h,g,G,L,T,U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ → ∀k0. T = ⋆k0 →
+ deg h g k0 l ∧ U = ⋆(next h k0).
+#h #g #G #L #T #U #l * -G -L -T -U -l
+[ #G #L #k #l #Hkl #k0 #H destruct /2 width=1/
+| #G #L #K #V #W #U #i #l #_ #_ #_ #k0 #H destruct
+| #G #L #K #W #V #U #i #l #_ #_ #_ #k0 #H destruct
+| #a #I #G #L #V #T #U #l #_ #k0 #H destruct
+| #G #L #V #T #U #l #_ #k0 #H destruct
+| #G #L #W #T #U #l #_ #k0 #H destruct
+qed-.
+
+lemma ssta_inv_sort1: ∀h,g,G,L,U,k,l. ⦃G, L⦄ ⊢ ⋆k •[h, g] ⦃l, U⦄ →
+ deg h g k l ∧ U = ⋆(next h k).
+/2 width=5 by ssta_inv_sort1_aux/ qed-.
+
+fact ssta_inv_lref1_aux: ∀h,g,G,L,T,U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ → ∀j. T = #j →
+ (∃∃K,V,W. ⇩[0, j] L ≡ K. ⓓV & ⦃G, K⦄ ⊢ V •[h, g] ⦃l, W⦄ &
+ ⇧[0, j + 1] W ≡ U
+ ) ∨
+ (∃∃K,W,V,l0. ⇩[0, j] L ≡ K. ⓛW & ⦃G, K⦄ ⊢ W •[h, g] ⦃l0, V⦄ &
+ ⇧[0, j + 1] W ≡ U & l = l0 + 1
+ ).
+#h #g #G #L #T #U #l * -G -L -T -U -l
+[ #G #L #k #l #_ #j #H destruct
+| #G #L #K #V #W #U #i #l #HLK #HVW #HWU #j #H destruct /3 width=6/
+| #G #L #K #W #V #U #i #l #HLK #HWV #HWU #j #H destruct /3 width=8/
+| #a #I #G #L #V #T #U #l #_ #j #H destruct
+| #G #L #V #T #U #l #_ #j #H destruct
+| #G #L #W #T #U #l #_ #j #H destruct
+]
+qed-.
+
+lemma ssta_inv_lref1: ∀h,g,G,L,U,i,l. ⦃G, L⦄ ⊢ #i •[h, g] ⦃l, U⦄ →
+ (∃∃K,V,W. ⇩[0, i] L ≡ K. ⓓV & ⦃G, K⦄ ⊢ V •[h, g] ⦃l, W⦄ &
+ ⇧[0, i + 1] W ≡ U
+ ) ∨
+ (∃∃K,W,V,l0. ⇩[0, i] L ≡ K. ⓛW & ⦃G, K⦄ ⊢ W •[h, g] ⦃l0, V⦄ &
+ ⇧[0, i + 1] W ≡ U & l = l0 + 1
+ ).
+/2 width=3 by ssta_inv_lref1_aux/ qed-.
+
+fact ssta_inv_gref1_aux: ∀h,g,G,L,T,U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ → ∀p0. T = §p0 → ⊥.
+#h #g #G #L #T #U #l * -G -L -T -U -l
+[ #G #L #k #l #_ #p0 #H destruct
+| #G #L #K #V #W #U #i #l #_ #_ #_ #p0 #H destruct
+| #G #L #K #W #V #U #i #l #_ #_ #_ #p0 #H destruct
+| #a #I #G #L #V #T #U #l #_ #p0 #H destruct
+| #G #L #V #T #U #l #_ #p0 #H destruct
+| #G #L #W #T #U #l #_ #p0 #H destruct
+qed-.
+
+lemma ssta_inv_gref1: ∀h,g,G,L,U,p,l. ⦃G, L⦄ ⊢ §p •[h, g] ⦃l, U⦄ → ⊥.
+/2 width=10 by ssta_inv_gref1_aux/ qed-.
+
+fact ssta_inv_bind1_aux: ∀h,g,G,L,T,U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ →
+ ∀a,I,X,Y. T = ⓑ{a,I}Y.X →
+ ∃∃Z. ⦃G, L.ⓑ{I}Y⦄ ⊢ X •[h, g] ⦃l, Z⦄ & U = ⓑ{a,I}Y.Z.
+#h #g #G #L #T #U #l * -G -L -T -U -l
+[ #G #L #k #l #_ #a #I #X #Y #H destruct
+| #G #L #K #V #W #U #i #l #_ #_ #_ #a #I #X #Y #H destruct
+| #G #L #K #W #V #U #i #l #_ #_ #_ #a #I #X #Y #H destruct
+| #b #J #G #L #V #T #U #l #HTU #a #I #X #Y #H destruct /2 width=3/
+| #G #L #V #T #U #l #_ #a #I #X #Y #H destruct
+| #G #L #W #T #U #l #_ #a #I #X #Y #H destruct
+]
+qed-.
+
+lemma ssta_inv_bind1: ∀h,g,a,I,G,L,Y,X,U,l. ⦃G, L⦄ ⊢ ⓑ{a,I}Y.X •[h, g] ⦃l, U⦄ →
+ ∃∃Z. ⦃G, L.ⓑ{I}Y⦄ ⊢ X •[h, g] ⦃l, Z⦄ & U = ⓑ{a,I}Y.Z.
+/2 width=3 by ssta_inv_bind1_aux/ qed-.
+
+fact ssta_inv_appl1_aux: ∀h,g,G,L,T,U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ → ∀X,Y. T = ⓐY.X →
+ ∃∃Z. ⦃G, L⦄ ⊢ X •[h, g] ⦃l, Z⦄ & U = ⓐY.Z.
+#h #g #G #L #T #U #l * -G -L -T -U -l
+[ #G #L #k #l #_ #X #Y #H destruct
+| #G #L #K #V #W #U #i #l #_ #_ #_ #X #Y #H destruct
+| #G #L #K #W #V #U #i #l #_ #_ #_ #X #Y #H destruct
+| #a #I #G #L #V #T #U #l #_ #X #Y #H destruct
+| #G #L #V #T #U #l #HTU #X #Y #H destruct /2 width=3/
+| #G #L #W #T #U #l #_ #X #Y #H destruct
+]
+qed-.
+
+lemma ssta_inv_appl1: ∀h,g,G,L,Y,X,U,l. ⦃G, L⦄ ⊢ ⓐY.X •[h, g] ⦃l, U⦄ →
+ ∃∃Z. ⦃G, L⦄ ⊢ X •[h, g] ⦃l, Z⦄ & U = ⓐY.Z.
+/2 width=3 by ssta_inv_appl1_aux/ qed-.
+
+fact ssta_inv_cast1_aux: ∀h,g,G,L,T,U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ →
+ ∀X,Y. T = ⓝY.X → ⦃G, L⦄ ⊢ X •[h, g] ⦃l, U⦄.
+#h #g #G #L #T #U #l * -G -L -T -U -l
+[ #G #L #k #l #_ #X #Y #H destruct
+| #G #L #K #V #W #U #l #i #_ #_ #_ #X #Y #H destruct
+| #G #L #K #W #V #U #l #i #_ #_ #_ #X #Y #H destruct
+| #a #I #G #L #V #T #U #l #_ #X #Y #H destruct
+| #G #L #V #T #U #l #_ #X #Y #H destruct
+| #G #L #W #T #U #l #HTU #X #Y #H destruct //
+]
+qed-.
+
+lemma ssta_inv_cast1: ∀h,g,G,L,X,Y,U,l. ⦃G, L⦄ ⊢ ⓝY.X •[h, g] ⦃l, U⦄ →
+ ⦃G, L⦄ ⊢ X •[h, g] ⦃l, U⦄.
+/2 width=4 by ssta_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/aaa_lift.ma".
+include "basic_2/static/ssta.ma".
+
+(* STRATIFIED STATIC TYPE ASSIGNMENT ON TERMS *******************************)
+
+(* Properties on atomic arity assignment for terms **************************)
+
+lemma ssta_aaa: ∀h,g,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ → ⦃G, L⦄ ⊢ U ⁝ A.
+#h #g #G #L #T #A #H elim H -G -L -T -A
+[ #G #L #k #U #l #H
+ elim (ssta_inv_sort1 … H) -H #_ #H destruct //
+| #I #G #L #K #V #B #i #HLK #HV #IHV #U #l #H
+ elim (ssta_inv_lref1 … H) -H * #K0 #V0 #W0 [2: #l0 ] #HLK0 #HVW0 #HU [ #H ]
+ lapply (ldrop_mono … HLK0 … HLK) -HLK0 #H0 destruct
+ lapply (ldrop_fwd_ldrop2 … HLK) -HLK #HLK
+ @(aaa_lift … HLK … HU) -HU -L // -HV /2 width=2/
+| #a #G #L #V #T #B #A #HV #_ #_ #IHT #X #l #H
+ elim (ssta_inv_bind1 … H) -H #U #HTU #H destruct /3 width=2/
+| #a #G #L #V #T #B #A #HV #_ #_ #IHT #X #l #H
+ elim (ssta_inv_bind1 … H) -H #U #HTU #H destruct /3 width=2/
+| #G #L #V #T #B #A #HV #_ #_ #IHT #X #l #H
+ elim (ssta_inv_appl1 … H) -H #U #HTU #H destruct /3 width=3/
+| #G #L #V #T #A #_ #_ #IHV #IHT #X #l #H
+ lapply (ssta_inv_cast1 … H) -H /2 width=2/
+]
+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/relocation/ldrop_ldrop.ma".
+include "basic_2/static/ssta.ma".
+
+(* STRATIFIED STATIC TYPE ASSIGNMENT ON TERMS *******************************)
+
+(* Properties on relocation *************************************************)
+
+lemma ssta_lift: ∀h,g,G,L1,T1,U1,l. ⦃G, L1⦄ ⊢ T1 •[h, g] ⦃l, U1⦄ →
+ ∀L2,d,e. ⇩[d, e] L2 ≡ L1 → ∀T2. ⇧[d, e] T1 ≡ T2 →
+ ∀U2. ⇧[d, e] U1 ≡ U2 → ⦃G, L2⦄ ⊢ T2 •[h, g] ⦃l, U2⦄.
+#h #g #G #L1 #T1 #U1 #l #H elim H -G -L1 -T1 -U1 -l
+[ #G #L1 #k #l #Hkl #L2 #d #e #HL21 #X1 #H1 #X2 #H2
+ >(lift_inv_sort1 … H1) -X1
+ >(lift_inv_sort1 … H2) -X2 /2 width=1/
+| #G #L1 #K1 #V1 #W1 #W #i #l #HLK1 #_ #HW1 #IHVW1 #L2 #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 (ldrop_trans_le … HL21 … HLK1 ?) -L1 /2 width=2/ #X #HLK2 #H
+ elim (ldrop_inv_skip2 … H ?) -H /2 width=1/ -Hid #K2 #V2 #HK21 #HV12 #H destruct
+ /3 width=8/
+ | lapply (lift_trans_be … HW1 … HWU2 ? ?) -W // /2 width=1/ #HW1U2
+ lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 // -Hid /3 width=8/
+ ]
+| #G #L1 #K1 #W1 #V1 #W #i #l #HLK1 #_ #HW1 #IHWV1 #L2 #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 (ldrop_trans_le … HL21 … HLK1 ?) -L1 /2 width=2/ #X #HLK2 #H
+ elim (ldrop_inv_skip2 … H ?) -H /2 width=1/ -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=8/
+ | lapply (lift_trans_be … HW1 … HWU2 ? ?) -W // /2 width=1/ #HW1U2
+ lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 // -Hid /3 width=8/
+ ]
+| #a #I #G #L1 #V1 #T1 #U1 #l #_ #IHTU1 #L2 #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=5/
+| #G #L1 #V1 #T1 #U1 #l #_ #IHTU1 #L2 #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=5/
+| #G #L1 #W1 #T1 #U1 #l #_ #IHTU1 #L2 #d #e #HL21 #X #H #U2 #HU12
+ elim (lift_inv_flat1 … H) -H #W2 #T2 #HW12 #HT12 #H destruct /3 width=5/
+]
+qed.
+
+lemma ssta_inv_lift1: ∀h,g,G,L2,T2,U2,l. ⦃G, L2⦄ ⊢ T2 •[h, g] ⦃l, U2⦄ →
+ ∀L1,d,e. ⇩[d, e] L2 ≡ L1 → ∀T1. ⇧[d, e] T1 ≡ T2 →
+ ∃∃U1. ⦃G, L1⦄ ⊢ T1 •[h, g] ⦃l, U1⦄ & ⇧[d, e] U1 ≡ U2.
+#h #g #G #L2 #T2 #U2 #l #H elim H -G -L2 -T2 -U2 -l
+[ #G #L2 #k #l #Hkl #L1 #d #e #_ #X #H
+ >(lift_inv_sort2 … H) -X /3 width=3/
+| #G #L2 #K2 #V2 #W2 #W #i #l #HLK2 #HVW2 #HW2 #IHVW2 #L1 #d #e #HL21 #X #H
+ elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HVW2 | -IHVW2 ]
+ [ elim (ldrop_conf_lt … HL21 … HLK2 ?) -L2 // #K1 #V1 #HLK1 #HK21 #HV12
+ elim (IHVW2 … HK21 … HV12) -K2 -V2 #W1 #HVW1 #HW12
+ elim (lift_trans_le … HW12 … HW2 ?) -W2 // >minus_plus <plus_minus_m_m // -Hid /3 width=6/
+ | lapply (ldrop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
+ elim (le_inv_plus_l … Hid) -Hid #Hdie #ei
+ elim (lift_split … HW2 d (i-e+1) ? ? ?) -HW2 // [3: /2 width=1/ ]
+ [ #W0 #HW20 <le_plus_minus_comm // >minus_minus_m_m /2 width=1/ /3 width=6/
+ | <le_plus_minus_comm //
+ ]
+ ]
+| #G #L2 #K2 #W2 #V2 #W #i #l #HLK2 #HWV2 #HW2 #IHWV2 #L1 #d #e #HL21 #X #H
+ elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HWV2 | -IHWV2 ]
+ [ elim (ldrop_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 // -Hid /3 width=6/
+ | lapply (ldrop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
+ elim (le_inv_plus_l … Hid) -Hid #Hdie #ei
+ elim (lift_split … HW2 d (i-e+1) ? ? ?) -HW2 // [3: /2 width=1/ ]
+ [ #W0 #HW20 <le_plus_minus_comm // >minus_minus_m_m /2 width=1/ /3 width=6/
+ | <le_plus_minus_comm //
+ ]
+ ]
+| #a #I #G #L2 #V2 #T2 #U2 #l #_ #IHTU2 #L1 #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 /2 width=1/ -HL21 /3 width=5/
+| #G #L2 #V2 #T2 #U2 #l #_ #IHTU2 #L1 #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/
+| #G #L2 #W2 #T2 #U2 #l #_ #IHTU2 #L1 #d #e #HL21 #X #H
+ elim (lift_inv_flat2 … H) -H #W1 #T1 #HW12 #HT12 #H destruct
+ elim (IHTU2 … HL21 … HT12) -L2 -HT12 /3 width=3/
+]
+qed-.
+
+(* Advanced forvard lemmas **************************************************)
+
+lemma ssta_fwd_correct: ∀h,g,G,L,T,U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ →
+ ∃T0. ⦃G, L⦄ ⊢ U •[h, g] ⦃l-1, T0⦄.
+#h #g #G #L #T #U #l #H elim H -G -L -T -U -l
+[ /4 width=2/
+| #G #L #K #V #W #W0 #i #l #HLK #_ #HW0 * #V0 #HWV0
+ lapply (ldrop_fwd_ldrop2 … HLK) -HLK #HLK
+ elim (lift_total V0 0 (i+1)) /3 width=10/
+| #G #L #K #W #V #V0 #i #l #HLK #HWV #HWV0 #_
+ lapply (ldrop_fwd_ldrop2 … HLK) -HLK #HLK
+ elim (lift_total V 0 (i+1)) /3 width=10/
+| #a #I #G #L #V #T #U #l #_ * /3 width=2/
+| #G #L #V #T #U #l #_ * #T0 #HUT0 /3 width=2/
+| #G #L #W #T #U #l #_ * /2 width=2/
+]
+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/ssta_lift.ma".
+
+(* STRATIFIED STATIC TYPE ASSIGNMENT ON TERMS *******************************)
+
+(* Main properties **********************************************************)
+
+theorem ssta_mono: ∀h,g,G,L,T,U1,l1. ⦃G, L⦄ ⊢ T •[h, g] ⦃l1, U1⦄ →
+ ∀U2,l2. ⦃G, L⦄ ⊢ T •[h, g] ⦃l2, U2⦄ → l1 = l2 ∧ U1 = U2.
+#h #g #G #L #T #U1 #l1 #H elim H -G -L -T -U1 -l1
+[ #G #L #k #l #Hkl #X #l2 #H
+ elim (ssta_inv_sort1 … H) -H #Hkl2 #H destruct
+ >(deg_mono … Hkl2 … Hkl) -g -L -l2 /2 width=1/
+| #G #L #K #V #W #U1 #i #l1 #HLK #_ #HWU1 #IHVW #U2 #l2 #H
+ elim (ssta_inv_lref1 … H) -H * #K0 #V0 #W0 [2: #l0] #HLK0 #HVW0 #HW0U2
+ lapply (ldrop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct
+ lapply (IHVW … HVW0) -IHVW -HVW0 * #H1 #H2 destruct
+ >(lift_mono … HWU1 … HW0U2) -W0 -U1 /2 width=1/
+| #G #L #K #W #V #U1 #i #l1 #HLK #_ #HWU1 #IHWV #U2 #l2 #H
+ elim (ssta_inv_lref1 … H) -H * #K0 #W0 #V0 [2: #l0 ] #HLK0 #HWV0 #HV0U2
+ lapply (ldrop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct
+ lapply (IHWV … HWV0) -IHWV -HWV0 * #H1 #H2 destruct
+ >(lift_mono … HWU1 … HV0U2) -W -U1 /2 width=1/
+| #a #I #G #L #V #T #U1 #l1 #_ #IHTU1 #X #l2 #H
+ elim (ssta_inv_bind1 … H) -H #U2 #HTU2 #H destruct
+ elim (IHTU1 … HTU2) -T /3 width=1/
+| #G #L #V #T #U1 #l1 #_ #IHTU1 #X #l2 #H
+ elim (ssta_inv_appl1 … H) -H #U2 #HTU2 #H destruct
+ elim (IHTU1 … HTU2) -T /3 width=1/
+| #G #L #W1 #T #U1 #l1 #_ #IHTU1 #U2 #l2 #H
+ lapply (ssta_inv_cast1 … H) -H #HTU2
+ elim (IHTU1 … HTU2) -T /2 width=1/
+]
+qed-.
+
+(* Advanced inversion lemmas ************************************************)
+
+lemma ssta_inv_refl_pos: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l+1, T⦄ → ⊥.
+#h #g #G #L #T #l #HTT
+elim (ssta_fwd_correct … HTT) <minus_plus_m_m #U #HTU
+elim (ssta_mono … HTU … HTT) -h -L #H #_ -T -U
+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 *)
+(* *)
+(**************************************************************************)
+
+(* 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 g ] break ⦃ term 46 l , break term 46 T2 ⦄ )"
+ non associative with precedence 45
+ for @{ 'StaticType $h $g $G $L $T1 $T2 $l }.
(* *)
(**************************************************************************)
+include "basic_2/notation/relations/statictypestar_6.ma".
include "basic_2/static/ssta.ma".
-(* ITERATED STRATIFIED STATIC TYPE ASSIGNMENTON TERMS ***********************)
+(* ITERATED STRATIFIED STATIC TYPE ASSIGNMENT FOR TERMS *********************)
-inductive sstas (h:sh) (g:sd h) (L:lenv): relation term ≝
-| sstas_refl: ∀T,U. ⦃h, L⦄ ⊢ T •[g, 0] U → sstas h g L T T
-| sstas_step: ∀T,U1,U2,l. ⦃h, L⦄ ⊢ T •[g, l+1] U1 → sstas h g L U1 U2 →
- sstas h g L T U2.
+definition sstas: ∀h. sd h → relation4 genv lenv term term ≝
+ λh,g,G,L. star … (ssta_step h g G L).
-interpretation "stratified unwind (term)"
- 'StaticTypeStar h g L T U = (sstas h g L T U).
+interpretation "iterated stratified static type assignment (term)"
+ 'StaticTypeStar h g G L T U = (sstas h g G L T U).
(* Basic eliminators ********************************************************)
-fact sstas_ind_alt_aux: ∀h,g,L,U2. ∀R:predicate term.
- (∀T. ⦃h, L⦄ ⊢ U2 •[g , 0] T → R U2) →
- (∀T,U1,l. ⦃h, L⦄ ⊢ T •[g, l + 1] U1 →
- ⦃h, L⦄ ⊢ U1 •* [g] U2 → R U1 → R T
- ) →
- ∀T,U. ⦃h, L⦄ ⊢ T •*[g] U → U = U2 → R T.
-#h #g #L #U2 #R #H1 #H2 #T #U #H elim H -H -T -U /2 width=2/ /3 width=5/
+lemma sstas_ind: ∀h,g,G,L,T. ∀R:predicate term.
+ R T → (
+ ∀U1,U2,l. ⦃G, L⦄ ⊢ T •* [h, g] U1 → ⦃G, L⦄ ⊢ U1 •[h, g] ⦃l+1, U2⦄ →
+ R U1 → R U2
+ ) →
+ ∀U. ⦃G, L⦄ ⊢ T •*[h, g] U → R U.
+#h #g #G #L #T #R #IH1 #IH2 #U #H elim H -U //
+#U1 #U2 #H * /2 width=5/
qed-.
-lemma sstas_ind_alt: ∀h,g,L,U2. ∀R:predicate term.
- (∀T. ⦃h, L⦄ ⊢ U2 •[g , 0] T → R U2) →
- (∀T,U1,l. ⦃h, L⦄ ⊢ T •[g, l + 1] U1 →
- ⦃h, L⦄ ⊢ U1 •* [g] U2 → R U1 → R T
- ) →
- ∀T. ⦃h, L⦄ ⊢ T •*[g] U2 → R T.
-/3 width=9 by sstas_ind_alt_aux/ qed-.
-
-(* Basic inversion lemmas ***************************************************)
-
-fact sstas_inv_sort1_aux: ∀h,g,L,T,U. ⦃h, L⦄ ⊢ T •*[g] U → ∀k. T = ⋆k →
- ∀l. deg h g k l → U = ⋆((next h)^l k).
-#h #g #L #T #U #H @(sstas_ind_alt … H) -T
-[ #U0 #HU0 #k #H #l #Hkl destruct
- elim (ssta_inv_sort1 … HU0) -L #HkO #_ -U0
- >(deg_mono … Hkl HkO) -g -l //
-| #T0 #U0 #l0 #HTU0 #_ #IHU0 #k #H #l #Hkl destruct
- elim (ssta_inv_sort1 … HTU0) -L #HkS #H destruct
- lapply (deg_mono … Hkl HkS) -Hkl #H destruct
- >(IHU0 (next h k) ? l0) -IHU0 // /2 width=1/ >iter_SO >iter_n_Sm //
-]
-qed.
+lemma sstas_ind_dx: ∀h,g,G,L,U2. ∀R:predicate term.
+ R U2 → (
+ ∀T,U1,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l+1, U1⦄ → ⦃G, L⦄ ⊢ U1 •* [h, g] U2 →
+ R U1 → R T
+ ) →
+ ∀T. ⦃G, L⦄ ⊢ T •*[h, g] U2 → R T.
+#h #g #G #L #U2 #R #IH1 #IH2 #T #H @(star_ind_l … T H) -T //
+#T #T0 * /2 width=5/
+qed-.
-lemma sstas_inv_sort1: ∀h,g,L,U,k. ⦃h, L⦄ ⊢ ⋆k •*[g] U → ∀l. deg h g k l →
- U = ⋆((next h)^l k).
-/2 width=6/ qed-.
+(* Basic properties *********************************************************)
-fact sstas_inv_bind1_aux: ∀h,g,L,T,U. ⦃h, L⦄ ⊢ T •*[g] U →
- ∀J,X,Y. T = ⓑ{J}Y.X →
- ∃∃Z. ⦃h, L.ⓑ{J}Y⦄ ⊢ X •*[g] Z & U = ⓑ{J}Y.Z.
-#h #g #L #T #U #H @(sstas_ind_alt … H) -T
-[ #U0 #HU0 #J #X #Y #H destruct
- elim (ssta_inv_bind1 … HU0) -HU0 #X0 #HX0 #H destruct /3 width=3/
-| #T0 #U0 #l #HTU0 #_ #IHU0 #J #X #Y #H destruct
- elim (ssta_inv_bind1 … HTU0) -HTU0 #X0 #HX0 #H destruct
- elim (IHU0 J X0 Y ?) -IHU0 // #X1 #HX01 #H destruct /3 width=4/
-]
-qed.
+lemma sstas_refl: ∀h,g,G,L. reflexive … (sstas h g G L).
+// qed.
-lemma sstas_inv_bind1: ∀h,g,J,L,Y,X,U. ⦃h, L⦄ ⊢ ⓑ{J}Y.X •*[g] U →
- ∃∃Z. ⦃h, L.ⓑ{J}Y⦄ ⊢ X •*[g] Z & U = ⓑ{J}Y.Z.
-/2 width=3/ qed-.
+lemma ssta_sstas: ∀h,g,G,L,T,U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l+1, U⦄ → ⦃G, L⦄ ⊢ T •*[h, g] U.
+/3 width=2 by R_to_star, ex_intro/ qed. (**) (* auto fails without trace *)
-fact sstas_inv_appl1_aux: ∀h,g,L,T,U. ⦃h, L⦄ ⊢ T •*[g] U → ∀X,Y. T = ⓐY.X →
- ∃∃Z. ⦃h, L⦄ ⊢ X •*[g] Z & U = ⓐY.Z.
-#h #g #L #T #U #H @(sstas_ind_alt … H) -T
-[ #U0 #HU0 #X #Y #H destruct
- elim (ssta_inv_appl1 … HU0) -HU0 #X0 #HX0 #H destruct /3 width=3/
-| #T0 #U0 #l #HTU0 #_ #IHU0 #X #Y #H destruct
- elim (ssta_inv_appl1 … HTU0) -HTU0 #X0 #HX0 #H destruct
- elim (IHU0 X0 Y ?) -IHU0 // #X1 #HX01 #H destruct /3 width=4/
-]
+lemma sstas_strap1: ∀h,g,G,L,T1,T2,U2,l. ⦃G, L⦄ ⊢ T1 •*[h, g] T2 → ⦃G, L⦄ ⊢ T2 •[h, g] ⦃l+1, U2⦄ →
+ ⦃G, L⦄ ⊢ T1 •*[h, g] U2.
+/3 width=4 by sstep, ex_intro/ (**) (* auto fails without trace *)
qed.
-lemma sstas_inv_appl1: ∀h,g,L,Y,X,U. ⦃h, L⦄ ⊢ ⓐY.X •*[g] U →
- ∃∃Z. ⦃h, L⦄ ⊢ X •*[g] Z & U = ⓐY.Z.
-/2 width=3/ qed-.
+lemma sstas_strap2: ∀h,g,G,L,T1,U1,U2,l. ⦃G, L⦄ ⊢ T1 •[h, g] ⦃l+1, U1⦄ → ⦃G, L⦄ ⊢ U1 •*[h, g] U2 →
+ ⦃G, L⦄ ⊢ T1 •*[h, g] U2.
+/3 width=3 by star_compl, ex_intro/ (**) (* auto fails without trace *)
+qed.
-(* Basic forward lemmas *****************************************************)
+(* Basic inversion lemmas ***************************************************)
-lemma sstas_fwd_correct: ∀h,g,L,T,U. ⦃h, L⦄ ⊢ T •*[g] U →
- ∃∃W. ⦃h, L⦄ ⊢ U •[g, 0] W & ⦃h, L⦄ ⊢ U •*[g] U.
-#h #g #L #T #U #H @(sstas_ind_alt … H) -T /2 width=1/ /3 width=2/
+lemma sstas_inv_bind1: ∀h,g,a,I,G,L,Y,X,U. ⦃G, L⦄ ⊢ ⓑ{a,I}Y.X •*[h, g] U →
+ ∃∃Z. ⦃G, L.ⓑ{I}Y⦄ ⊢ X •*[h, g] Z & U = ⓑ{a,I}Y.Z.
+#h #g #a #I #G #L #Y #X #U #H @(sstas_ind … H) -U /2 width=3/
+#T #U #l #_ #HTU * #Z #HXZ #H destruct
+elim (ssta_inv_bind1 … HTU) -HTU #Z0 #HZ0 #H destruct /3 width=4/
qed-.
-(* Basic_1: removed theorems 7:
- sty1_bind sty1_abbr sty1_appl sty1_cast2
- sty1_lift sty1_correct sty1_trans
-*)
+lemma sstas_inv_appl1: ∀h,g,G,L,Y,X,U. ⦃G, L⦄ ⊢ ⓐY.X •*[h, g] U →
+ ∃∃Z. ⦃G, L⦄ ⊢ X •*[h, g] Z & U = ⓐY.Z.
+#h #g #G #L #Y #X #U #H @(sstas_ind … H) -U /2 width=3/
+#T #U #l #_ #HTU * #Z #HXZ #H destruct
+elim (ssta_inv_appl1 … HTU) -HTU #Z0 #HZ0 #H destruct /3 width=4/
+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/ssta_aaa.ma".
+include "basic_2/unfold/sstas.ma".
+
+(* ITERATED STRATIFIED STATIC TYPE ASSIGNMENT FOR TERMS *********************)
+
+(* Properties on atomic arity assignment for terms **************************)
+
+lemma sstas_aaa: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •*[h, g] U →
+ ∀A. ⦃G, L⦄ ⊢ T ⁝ A → ⦃G, L⦄ ⊢ U ⁝ A.
+#h #g #G #L #T #U #H @(sstas_ind_dx … H) -T // /3 width=6/
+qed.
(**************************************************************************)
include "basic_2/static/ssta_lift.ma".
-include "basic_2/unwind/sstas.ma".
+include "basic_2/unfold/sstas.ma".
-(* ITERATED STRATIFIED STATIC TYPE ASSIGNMENTON TERMS ***********************)
+(* ITERATED STRATIFIED STATIC TYPE ASSIGNMENT FOR TERMS *********************)
-(* Advanced properties ******************************************************)
+(* Advanced forward lemmas **************************************************)
-lemma sstas_total_S: ∀h,g,L,l,T,U. ⦃h, L⦄ ⊢ T•[g, l + 1]U →
- ∃∃W. ⦃h, L⦄ ⊢ T •*[g] W & ⦃h, L⦄ ⊢ U •*[g] W.
-#h #g #L #l @(nat_ind_plus … l) -l
-[ #T #U #HTU
- elim (ssta_fwd_correct … HTU) /4 width=4/
-| #l #IHl #T #U #HTU
- elim (ssta_fwd_correct … HTU) <minus_plus_m_m #V #HUV
- elim (IHl … HUV) -IHl -HUV /3 width=4/
-]
+lemma sstas_fwd_correct: ∀h,g,G,L,T1,U1,l1. ⦃G, L⦄ ⊢ T1 •[h, g] ⦃l1, U1⦄ →
+ ∀T2. ⦃G, L⦄ ⊢ T1 •*[h, g] T2 →
+ ∃∃U2,l2. ⦃G, L⦄ ⊢ T2 •[h, g] ⦃l2, U2⦄.
+#h #g #G #L #T1 #U1 #l1 #HTU1 #T2 #H @(sstas_ind … H) -T2 [ /2 width=3/ ] -HTU1
+#T #T2 #l #_ #HT2 * #U #l0 #_ -l0
+elim (ssta_fwd_correct … HT2) -T /2 width=3/
qed-.
(* Properties on relocation *************************************************)
-lemma sstas_lift: ∀h,g,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 •*[g] U1 →
+lemma sstas_lift: ∀h,g,G,L1,T1,U1. ⦃G, L1⦄ ⊢ T1 •*[h, g] U1 →
∀L2,d,e. ⇩[d, e] L2 ≡ L1 → ∀T2. ⇧[d, e] T1 ≡ T2 →
- ∀U2. ⇧[d, e] U1 ≡ U2 → ⦃h, L2⦄ ⊢ T2 •*[g] U2.
-#h #g #L1 #T1 #U1 #H @(sstas_ind_alt … H) -T1
-[ #T1 #HUT1 #L2 #d #e #HL21 #X #HX #U2 #HU12
- >(lift_mono … HX … HU12) -X
- elim (lift_total T1 d e) /3 width=10/
+ ∀U2. ⇧[d, e] U1 ≡ U2 → ⦃G, L2⦄ ⊢ T2 •*[h, g] U2.
+#h #g #G #L1 #T1 #U1 #H @(sstas_ind_dx … H) -T1
+[ #L2 #d #e #HL21 #X #HX #U2 #HU12
+ >(lift_mono … HX … HU12) -X //
| #T0 #U0 #l0 #HTU0 #_ #IHU01 #L2 #d #e #HL21 #T2 #HT02 #U2 #HU12
elim (lift_total U0 d e) /3 width=10/
]
qed.
-lemma sstas_inv_lift1: ∀h,g,L2,T2,U2. ⦃h, L2⦄ ⊢ T2 •*[g] U2 →
+(* Inversion lemmas on relocation *******************************************)
+
+lemma sstas_inv_lift1: ∀h,g,G,L2,T2,U2. ⦃G, L2⦄ ⊢ T2 •*[h, g] U2 →
∀L1,d,e. ⇩[d, e] L2 ≡ L1 → ∀T1. ⇧[d, e] T1 ≡ T2 →
- ∃∃U1. ⦃h, L1⦄ ⊢ T1 •*[g] U1 & ⇧[d, e] U1 ≡ U2.
-#h #g #L2 #T2 #U2 #H @(sstas_ind_alt … H) -T2
-[ #T2 #HUT2 #L1 #d #e #HL21 #U1 #HU12
- elim (ssta_inv_lift1 … HUT2 … HL21 … HU12) -HUT2 -HL21 /3 width=3/
-| #T0 #U0 #l0 #HTU0 #_ #IHU01 #L1 #d #e #HL21 #U1 #HU12
- elim (ssta_inv_lift1 … HTU0 … HL21 … HU12) -HTU0 -HU12 #U #HU1 #HU0
- elim (IHU01 … HL21 … HU0) -IHU01 -HL21 -U0 /3 width=4/
-]
+ ∃∃U1. ⦃G, L1⦄ ⊢ T1 •*[h, g] U1 & ⇧[d, e] U1 ≡ U2.
+#h #g #G #L2 #T2 #U2 #H @(sstas_ind_dx … H) -T2 /2 width=3/
+#T0 #U0 #l0 #HTU0 #_ #IHU01 #L1 #d #e #HL21 #U1 #HU12
+elim (ssta_inv_lift1 … HTU0 … HL21 … HU12) -HTU0 -HU12 #U #HU1 #HU0
+elim (IHU01 … HL21 … HU0) -IHU01 -HL21 -U0 /3 width=4/
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/ssta_ltpss.ma".
-include "basic_2/unwind/sstas.ma".
-
-(* ITERATED STRATIFIED STATIC TYPE ASSIGNMENTON TERMS ***********************)
-
-(* Properties about parallel unfold *****************************************)
-
-lemma sstas_ltpss_tpss_conf: ∀h,g,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 •*[g] U1 →
- ∀L2,d,e. L1 ▶* [d, e] L2 →
- ∀T2. L2 ⊢ T1 ▶* [d, e] T2 →
- ∃∃U2. ⦃h, L2⦄ ⊢ T2 •*[g] U2 &
- L2 ⊢ U1 ▶* [d, e] U2.
-#h #g #L1 #T1 #U1 #H @(sstas_ind_alt … H) -T1
-[ #T1 #HUT1 #L2 #d #e #HL12 #U2 #HU12
- elim (ssta_ltpss_tpss_conf … HUT1 … HL12 … HU12) -HUT1 -HL12 /3 width=3/
-| #T0 #U0 #l0 #HTU0 #_ #IHU01 #L2 #d #e #HL12 #T #HT0
- elim (ssta_ltpss_tpss_conf … HTU0 … HL12 … HT0) -HTU0 -HT0 #U #HTU #HU0
- elim (IHU01 … HL12 … HU0) -IHU01 -HL12 -U0 /3 width=4/
-]
-qed.
-
-lemma sstas_ltpss_tps_conf: ∀h,g,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 •*[g] U1 →
- ∀L2,d,e. L1 ▶* [d, e] L2 →
- ∀T2. L2 ⊢ T1 ▶ [d, e] T2 →
- ∃∃U2. ⦃h, L2⦄ ⊢ T2 •*[g] U2 & L2 ⊢ U1 ▶* [d, e] U2.
-/3 width=5/ qed.
-
-lemma sstas_ltpss_conf: ∀h,g,L1,T,U1. ⦃h, L1⦄ ⊢ T •*[g] U1 →
- ∀L2,d,e. L1 ▶* [d, e] L2 →
- ∃∃U2. ⦃h, L2⦄ ⊢ T •*[g] U2 & L2 ⊢ U1 ▶* [d, e] U2.
-/2 width=5/ qed.
-
-lemma sstas_tpss_conf: ∀h,g,L,T1,U1. ⦃h, L⦄ ⊢ T1 •*[g] U1 →
- ∀T2,d,e. L ⊢ T1 ▶* [d, e] T2 →
- ∃∃U2. ⦃h, L⦄ ⊢ T2 •*[g] U2 & L ⊢ U1 ▶* [d, e] U2.
-/2 width=5/ qed.
-
-lemma sstas_tps_conf: ∀h,g,L,T1,U1. ⦃h, L⦄ ⊢ T1 •*[g] U1 →
- ∀T2,d,e. L ⊢ T1 ▶ [d, e] T2 →
- ∃∃U2. ⦃h, L⦄ ⊢ T2 •*[g] U2 & L ⊢ U1 ▶* [d, e] U2.
-/2 width=5/ qed.
(* *)
(**************************************************************************)
-include "basic_2/unfold/delift_lift.ma".
include "basic_2/static/ssta_ssta.ma".
-include "basic_2/unwind/sstas_lift.ma".
+include "basic_2/unfold/sstas.ma".
-(* ITERATED STRATIFIED STATIC TYPE ASSIGNMENTON TERMS ***********************)
+(* ITERATED STRATIFIED STATIC TYPE ASSIGNMENT FOR TERMS *********************)
(* Advanced inversion lemmas ************************************************)
-lemma sstas_inv_O: ∀h,g,L,T,U. ⦃h, L⦄ ⊢ T •*[g] U →
- ∀T0. ⦃h, L⦄ ⊢ T •[g , 0] T0 → U = T.
-#h #g #L #T #U #H @(sstas_ind_alt … H) -T //
+lemma sstas_inv_O: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •*[h, g] U →
+ ∀T0. ⦃G, L⦄ ⊢ T •[h, g] ⦃0, T0⦄ → U = T.
+#h #g #G #L #T #U #H @(sstas_ind_dx … H) -T //
#T0 #U0 #l0 #HTU0 #_ #_ #T1 #HT01
elim (ssta_mono … HTU0 … HT01) <plus_n_Sm #H destruct
qed-.
-lemma sstas_inv_S: ∀h,g,L,T,U. ⦃h, L⦄ ⊢ T •*[g] U →
- ∀T0,l. ⦃h, L⦄ ⊢ T •[g , l+1] T0 → ⦃h, L⦄ ⊢ T0 •*[g] U.
-#h #g #L #T #U #H @(sstas_ind_alt … H) -T
-[ #U0 #HU0 #T #l #HUT
- elim (ssta_mono … HUT … HU0) <plus_n_Sm #H destruct
-| #T0 #U0 #l0 #HTU0 #HU0 #_ #T #l #HT0
- elim (ssta_mono … HT0 … HTU0) -T0 #_ #H destruct -l0 //
-]
-qed-.
-
-(* Main properties **********************************************************)
+(* Advanced properties ******************************************************)
-theorem sstas_mono: ∀h,g,L,T,U1. ⦃h, L⦄ ⊢ T •*[g] U1 →
- ∀U2. ⦃h, L⦄ ⊢ T •*[g] U2 → U1 = U2.
-#h #g #L #T #U1 #H @(sstas_ind_alt … H) -T
-[ #T1 #HUT1 #U2 #HU12
- >(sstas_inv_O … HU12 … HUT1) -h -L -T1 -U2 //
-| #T0 #U0 #l0 #HTU0 #_ #IHU01 #U2 #HU12
- lapply (sstas_inv_S … HU12 … HTU0) -T0 -l0 /2 width=1/
-]
+lemma sstas_strip: ∀h,g,G,L,T,U1. ⦃G, L⦄ ⊢ T •*[h, g] U1 →
+ ∀U2,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U2⦄ →
+ T = U1 ∨ ⦃G, L⦄ ⊢ U2 •*[h, g] U1.
+#h #g #G #L #T #U1 #H1 @(sstas_ind_dx … H1) -T /2 width=1/
+#T #U #l0 #HTU #HU1 #_ #U2 #l #H2
+elim (ssta_mono … H2 … HTU) -H2 -HTU #H1 #H2 destruct /2 width=1/
qed-.
-(* More advancd inversion lemmas ********************************************)
+(* Main properties **********************************************************)
-fact sstas_inv_lref1_aux: ∀h,g,L,T,U. ⦃h, L⦄ ⊢ T •*[g] U → ∀j. T = #j →
- ∃∃I,K,V,W. ⇩[0, j] L ≡ K. ⓑ{I}V & ⦃h, K⦄ ⊢ V •*[g] W &
- L ⊢ ▼*[0, j + 1] U ≡ W.
-#h #g #L #T #U #H @(sstas_ind_alt … H) -T
-[ #T #HUT #j #H destruct
- elim (ssta_inv_lref1 … HUT) -HUT * #K #V #W [2: #l] #HLK #HVW #HVT
- [ <plus_n_Sm #H destruct
- | /3 width=12/
- ]
-| #T0 #U0 #l0 #HTU0 #HU0 #_ #j #H destruct
- elim (ssta_inv_lref1 … HTU0) -HTU0 * #K #V #W [2: #l] #HLK #HVW #HVU0
- [ #_ -HVW
- lapply (ldrop_fwd_ldrop2 … HLK) #H
- elim (sstas_inv_lift1 … HU0 … H … HVU0) -HU0 -H -HVU0 /3 width=7/
- | elim (sstas_total_S … HVW) -HVW #T #HVT #HWT
- lapply (ldrop_fwd_ldrop2 … HLK) #H
- elim (sstas_inv_lift1 … HU0 … H … HVU0) -HU0 -H -HVU0 #X #HWX
- >(sstas_mono … HWX … HWT) -X -W /3 width=7/
- ]
+theorem sstas_trans: ∀h,g,G,L,T1,U. ⦃G, L⦄ ⊢ T1 •*[h, g] U →
+ ∀T2. ⦃G, L⦄ ⊢ U •*[h, g] T2 → ⦃G, L⦄ ⊢ T1 •*[h, g] T2.
+/2 width=3/ qed-.
+
+theorem sstas_conf: ∀h,g,G,L,T,U1. ⦃G, L⦄ ⊢ T •*[h, g] U1 →
+ ∀U2. ⦃G, L⦄ ⊢ T •*[h, g] U2 →
+ ⦃G, L⦄ ⊢ U1 •*[h, g] U2 ∨ ⦃G, L⦄ ⊢ U2 •*[h, g] U1.
+#h #g #G #L #T #U1 #H1 @(sstas_ind_dx … H1) -T /2 width=1/
+#T #U #l #HTU #HU1 #IHU1 #U2 #H2
+elim (sstas_strip … H2 … HTU) #H destruct
+[ -H2 -IHU1 /3 width=4/
+| -T /2 width=1/
]
qed-.
(* *)
(**************************************************************************)
-include "basic_2/substitution/ldrop.ma".
+include "basic_2/notation/relations/statictype_5.ma".
+include "basic_2/grammar/genv.ma".
+include "basic_2/relocation/ldrop.ma".
include "basic_2/static/sh.ma".
(* STATIC TYPE ASSIGNMENT ON TERMS ******************************************)
-inductive sta (h:sh): lenv → relation term ≝
-| sta_sort: ∀L,k. sta h L (⋆k) (⋆(next h k))
-| sta_ldef: ∀L,K,V,W,U,i. ⇩[0, i] L ≡ K. ⓓV → sta h K V W →
- ⇧[0, i + 1] W ≡ U → sta h L (#i) U
-| sta_ldec: ∀L,K,W,V,U,i. ⇩[0, i] L ≡ K. ⓛW → sta h K W V →
- ⇧[0, i + 1] W ≡ U → sta h L (#i) U
-| sta_bind: ∀I,L,V,T,U. sta h (L. ⓑ{I} V) T U →
- sta h L (ⓑ{I}V.T) (ⓑ{I}V.U)
-| sta_appl: ∀L,V,T,U. sta h L T U →
- sta h L (ⓐV.T) (ⓐV.U)
-| sta_cast: ∀L,W,T,U. sta h L T U → sta h L (ⓝW. T) U
+(* 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. ⇩[0, 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. ⇩[0, 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 L T U = (sta h L T U).
+ 'StaticType h G L T U = (sta h G L T U).
(* Basic inversion lemmas ************************************************)
-fact sta_inv_sort1_aux: ∀h,L,T,U. ⦃h, L⦄ ⊢ T • U → ∀k0. T = ⋆k0 →
+fact sta_inv_sort1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀k0. T = ⋆k0 →
U = ⋆(next h k0).
-#h #L #T #U * -L -T -U
-[ #L #k #k0 #H destruct //
-| #L #K #V #W #U #i #_ #_ #_ #k0 #H destruct
-| #L #K #W #V #U #i #_ #_ #_ #k0 #H destruct
-| #I #L #V #T #U #_ #k0 #H destruct
-| #L #V #T #U #_ #k0 #H destruct
-| #L #W #T #U #_ #k0 #H destruct
-qed.
+#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,L,U,k. ⦃h, L⦄ ⊢ ⋆k • U → U = ⋆(next h k).
-/2 width=4/ qed-.
+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,L,T,U. ⦃h, L⦄ ⊢ T • U → ∀j. T = #j →
- (∃∃K,V,W. ⇩[0, j] L ≡ K. ⓓV & ⦃h, K⦄ ⊢ V • W &
+fact sta_inv_lref1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀j. T = #j →
+ (∃∃K,V,W. ⇩[0, j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •[h] W &
⇧[0, j + 1] W ≡ U
) ∨
- (∃∃K,W,V. ⇩[0, j] L ≡ K. ⓛW & ⦃h, K⦄ ⊢ W • V &
+ (∃∃K,W,V. ⇩[0, j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W •[h] V &
⇧[0, j + 1] W ≡ U
).
-#h #L #T #U * -L -T -U
-[ #L #k #j #H destruct
-| #L #K #V #W #U #i #HLK #HVW #HWU #j #H destruct /3 width=6/
-| #L #K #W #V #U #i #HLK #HWV #HWU #j #H destruct /3 width=6/
-| #I #L #V #T #U #_ #j #H destruct
-| #L #V #T #U #_ #j #H destruct
-| #L #W #T #U #_ #j #H destruct
+#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/
+| #G #L #K #W #V #U #i #HLK #HWV #HWU #j #H destruct /3 width=6/
+| #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.
+qed-.
(* Basic_1: was sty0_gen_lref *)
-lemma sta_inv_lref1: ∀h,L,U,i. ⦃h, L⦄ ⊢ #i • U →
- (∃∃K,V,W. ⇩[0, i] L ≡ K. ⓓV & ⦃h, K⦄ ⊢ V • W &
+lemma sta_inv_lref1: ∀h,G,L,U,i. ⦃G, L⦄ ⊢ #i •[h] U →
+ (∃∃K,V,W. ⇩[0, i] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •[h] W &
⇧[0, i + 1] W ≡ U
) ∨
- (∃∃K,W,V. ⇩[0, i] L ≡ K. ⓛW & ⦃h, K⦄ ⊢ W • V &
+ (∃∃K,W,V. ⇩[0, i] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W •[h] V &
⇧[0, i + 1] W ≡ U
).
-/2 width=3/ qed-.
-
-fact sta_inv_bind1_aux: ∀h,L,T,U. ⦃h, L⦄ ⊢ T • U → ∀J,X,Y. T = ⓑ{J}Y.X →
- ∃∃Z. ⦃h, L.ⓑ{J}Y⦄ ⊢ X • Z & U = ⓑ{J}Y.Z.
-#h #L #T #U * -L -T -U
-[ #L #k #J #X #Y #H destruct
-| #L #K #V #W #U #i #_ #_ #_ #J #X #Y #H destruct
-| #L #K #W #V #U #i #_ #_ #_ #J #X #Y #H destruct
-| #I #L #V #T #U #HTU #J #X #Y #H destruct /2 width=3/
-| #L #V #T #U #_ #J #X #Y #H destruct
-| #L #W #T #U #_ #J #X #Y #H destruct
+/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/
+| #G #L #V #T #U #_ #b #J #X #Y #H destruct
+| #G #L #W #T #U #_ #b #J #X #Y #H destruct
]
-qed.
+qed-.
(* Basic_1: was: sty0_gen_bind *)
-lemma sta_inv_bind1: ∀h,J,L,Y,X,U. ⦃h, L⦄ ⊢ ⓑ{J}Y.X • U →
- ∃∃Z. ⦃h, L.ⓑ{J}Y⦄ ⊢ X • Z & U = ⓑ{J}Y.Z.
-/2 width=3/ qed-.
-
-fact sta_inv_appl1_aux: ∀h,L,T,U. ⦃h, L⦄ ⊢ T • U → ∀X,Y. T = ⓐY.X →
- ∃∃Z. ⦃h, L⦄ ⊢ X • Z & U = ⓐY.Z.
-#h #L #T #U * -L -T -U
-[ #L #k #X #Y #H destruct
-| #L #K #V #W #U #i #_ #_ #_ #X #Y #H destruct
-| #L #K #W #V #U #i #_ #_ #_ #X #Y #H destruct
-| #I #L #V #T #U #_ #X #Y #H destruct
-| #L #V #T #U #HTU #X #Y #H destruct /2 width=3/
-| #L #W #T #U #_ #X #Y #H destruct
+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/
+| #G #L #W #T #U #_ #X #Y #H destruct
]
-qed.
+qed-.
(* Basic_1: was: sty0_gen_appl *)
-lemma sta_inv_appl1: ∀h,L,Y,X,U. ⦃h, L⦄ ⊢ ⓐY.X • U →
- ∃∃Z. ⦃h, L⦄ ⊢ X • Z & U = ⓐY.Z.
-/2 width=3/ qed-.
-
-fact sta_inv_cast1_aux: ∀h,L,T,U. ⦃h, L⦄ ⊢ T • U → ∀X,Y. T = ⓝY.X →
- ⦃h, L⦄ ⊢ X • U.
-#h #L #T #U * -L -T -U
-[ #L #k #X #Y #H destruct
-| #L #K #V #W #U #i #_ #_ #_ #X #Y #H destruct
-| #L #K #W #V #U #i #_ #_ #_ #X #Y #H destruct
-| #I #L #V #T #U #_ #X #Y #H destruct
-| #L #V #T #U #_ #X #Y #H destruct
-| #L #W #T #U #HTU #X #Y #H destruct //
+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.
+qed-.
(* Basic_1: was: sty0_gen_cast *)
-lemma sta_inv_cast1: ∀h,L,X,Y,U. ⦃h, L⦄ ⊢ ⓝY.X • U → ⦃h, L⦄ ⊢ X • U.
-/2 width=4/ qed-.
+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-.
+
+(* 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/
+| #G #L #K #W #i #l #HLK #_ * #V #HWV
+ elim (lift_total W 0 (i+1)) /3 width=7/
+| #a #I #G #L #V #T #l #_ * /3 width=2/
+| * #G #L #V #T #l #_ * /3 width=2/
+]
+qed-.
+
+(* Properties on static type assignment for terms ***************************)
+
+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/
+| #G #L #K #V #W #W0 #i #HLK #_ #_ * /3 width=5/
+| #G #L #K #W #V #W0 #i #HLK #_ #_ * /3 width=5/
+| #a #I #G #L #V #T #U #_ * /3 width=2/
+| #G #L #V #T #U #_ * /3 width=2/
+| #G #L #W #T #U #_ * /3 width=2/
+]
+qed-.
--- /dev/null
+(* Forward lemmas on stratified static type assignment for terms ************)
+
+lemma aaa_fwd_ssta: ∀h,g,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∃U. ⦃G, L⦄ ⊢ T •[h, g] U.
+#h #G #L #T #A #H elim H -G -L -T -A
+[ /2 width=2/
+| * #G #L #K [ #V | #W ] #B #i #HLK #_ * [ #W | #V ] #HVW
+ elim (lift_total W 0 (i+1)) /3 width=7/
+| #a #G #L #V #T #B #A #_ #_ #_ * /3 width=2/
+| #a #G #L #V #T #B #A #_ #_ #_ * /3 width=2/
+| #G #L #V #T #B #A #_ #_ #_ * /3 width=2/
+| #G #L #W #T #A #_ #_ #_ * /3 width=2/
+]
+qed-.
(* *)
(**************************************************************************)
-include "basic_2/substitution/ldrop_ldrop.ma".
+include "basic_2/relocation/ldrop_ldrop.ma".
include "basic_2/static/sta.ma".
(* STATIC TYPE ASSIGNMENT ON TERMS ******************************************)
(* Properties on relocation *************************************************)
(* Basic_1: was: sty0_lift *)
-lemma sta_lift: ∀h,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 • U1 → ∀L2,d,e. ⇩[d, e] L2 ≡ L1 →
- ∀T2. ⇧[d, e] T1 ≡ T2 → ∀U2. ⇧[d, e] U1 ≡ U2 → ⦃h, L2⦄ ⊢ T2 • U2.
-#h #L1 #T1 #U1 #H elim H -L1 -T1 -U1
-[ #L1 #k #L2 #d #e #HL21 #X1 #H1 #X2 #H2
+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 #d #e #HL21 #X1 #H1 #X2 #H2
>(lift_inv_sort1 … H1) -X1
>(lift_inv_sort1 … H2) -X2 //
-| #L1 #K1 #V1 #W1 #W #i #HLK1 #_ #HW1 #IHVW1 #L2 #d #e #HL21 #X #H #U2 #HWU2
+| #G #L1 #K1 #V1 #W1 #W #i #HLK1 #_ #HW1 #IHVW1 #L2 #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 (ldrop_trans_le … HL21 … HLK1 ?) -L1 /2 width=2/ #X #HLK2 #H
- elim (ldrop_inv_skip2 … H ?) -H /2 width=1/ -Hid #K2 #V2 #HK21 #HV12 #H destruct
+ [ elim (lift_trans_ge … HW1 … HWU2) -W // #W2 #HW12 #HWU2
+ elim (ldrop_trans_le … HL21 … HLK1) -L1 /2 width=2/ #X #HLK2 #H
+ elim (ldrop_inv_skip2 … H) -H /2 width=1/ -Hid #K2 #V2 #HK21 #HV12 #H destruct
/3 width=8/
| lapply (lift_trans_be … HW1 … HWU2 ? ?) -W // /2 width=1/ #HW1U2
lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 // -Hid /3 width=8/
]
-| #L1 #K1 #W1 #V1 #W #i #HLK1 #_ #HW1 #IHWV1 #L2 #d #e #HL21 #X #H #U2 #HWU2
+| #G #L1 #K1 #W1 #V1 #W #i #HLK1 #_ #HW1 #IHWV1 #L2 #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 (ldrop_trans_le … HL21 … HLK1 ?) -L1 /2 width=2/ #X #HLK2 #H
- elim (ldrop_inv_skip2 … H ?) -H /2 width=1/ -Hid #K2 #W2 #HK21 #HW12 #H destruct
+ [ elim (lift_trans_ge … HW1 … HWU2) -W // <minus_plus #W #HW1 #HWU2
+ elim (ldrop_trans_le … HL21 … HLK1) -L1 /2 width=2/ #X #HLK2 #H
+ elim (ldrop_inv_skip2 … H) -H /2 width=1/ -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=8/
| lapply (lift_trans_be … HW1 … HWU2 ? ?) -W // /2 width=1/ #HW1U2
lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 // -Hid /3 width=8/
]
-| #I #L1 #V1 #T1 #U1 #_ #IHTU1 #L2 #d #e #HL21 #X1 #H1 #X2 #H2
+| #a #I #G #L1 #V1 #T1 #U1 #_ #IHTU1 #L2 #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=5/
-| #L1 #V1 #T1 #U1 #_ #IHTU1 #L2 #d #e #HL21 #X1 #H1 #X2 #H2
+| #G #L1 #V1 #T1 #U1 #_ #IHTU1 #L2 #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=5/
-| #L1 #W1 #T1 #U1 #_ #IHTU1 #L2 #d #e #HL21 #X #H #U2 #HU12
+| #G #L1 #W1 #T1 #U1 #_ #IHTU1 #L2 #d #e #HL21 #X #H #U2 #HU12
elim (lift_inv_flat1 … H) -H #W2 #T2 #_ #HT12 #H destruct /3 width=5/
]
qed.
(* Note: apparently this was missing in basic_1 *)
-lemma sta_inv_lift1: ∀h,L2,T2,U2. ⦃h, L2⦄ ⊢ T2 • U2 → ∀L1,d,e. ⇩[d, e] L2 ≡ L1 →
- ∀T1. ⇧[d, e] T1 ≡ T2 →
- ∃∃U1. ⦃h, L1⦄ ⊢ T1 • U1 & ⇧[d, e] U1 ≡ U2.
-#h #L2 #T2 #U2 #H elim H -L2 -T2 -U2
-[ #L2 #k #L1 #d #e #_ #X #H
+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 #d #e #_ #X #H
>(lift_inv_sort2 … H) -X /2 width=3/
-| #L2 #K2 #V2 #W2 #W #i #HLK2 #HVW2 #HW2 #IHVW2 #L1 #d #e #HL21 #X #H
+| #G #L2 #K2 #V2 #W2 #W #i #HLK2 #HVW2 #HW2 #IHVW2 #L1 #d #e #HL21 #X #H
elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HVW2 | -IHVW2 ]
- [ elim (ldrop_conf_lt … HL21 … HLK2 ?) -L2 // #K1 #V1 #HLK1 #HK21 #HV12
- elim (IHVW2 … HK21 … HV12) -K2 -V2 #W1 #HVW1 #HW12
- elim (lift_trans_le … HW12 … HW2 ?) -W2 // >minus_plus <plus_minus_m_m // -Hid /3 width=6/
+ [ elim (ldrop_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 // -Hid /3 width=8/
| lapply (ldrop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
elim (le_inv_plus_l … Hid) -Hid #Hdie #ei
- elim (lift_split … HW2 d (i-e+1) ? ? ?) -HW2 // [3: /2 width=1/ ]
- [ #W0 #HW20 <le_plus_minus_comm // >minus_minus_m_m /2 width=1/ /3 width=6/
- | <le_plus_minus_comm // /2 width=1/
+ elim (lift_split … HW2 d (i-e+1)) -HW2 // [3: /2 width=1/ ]
+ [ #W0 #HW20 <le_plus_minus_comm // >minus_minus_m_m /2 width=1/ /3 width=8/
+ | <le_plus_minus_comm //
]
]
-| #L2 #K2 #W2 #V2 #W #i #HLK2 #HWV2 #HW2 #IHWV2 #L1 #d #e #HL21 #X #H
+| #G #L2 #K2 #W2 #V2 #W #i #HLK2 #HWV2 #HW2 #IHWV2 #L1 #d #e #HL21 #X #H
elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HWV2 | -IHWV2 ]
- [ elim (ldrop_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 // -Hid /3 width=6/
+ [ elim (ldrop_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 // -Hid /3 width=8/
| lapply (ldrop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
elim (le_inv_plus_l … Hid) -Hid #Hdie #ei
- elim (lift_split … HW2 d (i-e+1) ? ? ?) -HW2 // [3: /2 width=1/ ]
- [ #W0 #HW20 <le_plus_minus_comm // >minus_minus_m_m /2 width=1/ /3 width=6/
- | <le_plus_minus_comm // /2 width=1/
+ elim (lift_split … HW2 d (i-e+1)) -HW2 // [3: /2 width=1/ ]
+ [ #W0 #HW20 <le_plus_minus_comm // >minus_minus_m_m /2 width=1/ /3 width=8/
+ | <le_plus_minus_comm //
]
]
-| #I #L2 #V2 #T2 #U2 #_ #IHTU2 #L1 #d #e #HL21 #X #H
+| #a #I #G #L2 #V2 #T2 #U2 #_ #IHTU2 #L1 #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 /2 width=1/ -HL21 /3 width=5/
-| #L2 #V2 #T2 #U2 #_ #IHTU2 #L1 #d #e #HL21 #X #H
+| #G #L2 #V2 #T2 #U2 #_ #IHTU2 #L1 #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/
-| #L2 #W2 #T2 #U2 #_ #IHTU2 #L1 #d #e #HL21 #X #H
+| #G #L2 #W2 #T2 #U2 #_ #IHTU2 #L1 #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/
]
-qed.
+qed-.
(* Advanced forvard lemmas **************************************************)
(* Basic_1: was: sty0_correct *)
-lemma sta_fwd_correct: ∀h,L,T,U. ⦃h, L⦄ ⊢ T • U → ∃T0. ⦃h, L⦄ ⊢ U • T0.
-#h #L #T #U #H elim H -L -T -U
+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/
-| #L #K #V #W #W0 #i #HLK #_ #HW0 * #V0 #HWV0
+| #G #L #K #V #W #W0 #i #HLK #_ #HW0 * #V0 #HWV0
lapply (ldrop_fwd_ldrop2 … HLK) -HLK #HLK
elim (lift_total V0 0 (i+1)) /3 width=10/
-| #L #K #W #V #V0 #i #HLK #HWV #HWV0 #_
+| #G #L #K #W #V #V0 #i #HLK #HWV #HWV0 #_
lapply (ldrop_fwd_ldrop2 … HLK) -HLK #HLK
elim (lift_total V 0 (i+1)) /3 width=10/
-| #I #L #V #T #U #_ * /3 width=2/
-| #L #V #T #U #_ * #T0 #HUT0 /3 width=2/
-| #L #W #T #U #_ * /2 width=2/
+| #a #I #G #L #V #T #U #_ * /3 width=2/
+| #G #L #V #T #U #_ * #T0 #HUT0 /3 width=2/
+| #G #L #W #T #U #_ * /2 width=2/
]
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/tpss_tpss.ma".
-include "basic_2/unfold/ltpss_tpss.ma".
-include "basic_2/static/sta_lift.ma".
-
-(* STATIC TYPE ASSIGNMENT ON TERMS ******************************************)
-
-(* Properties about parallel unfold *****************************************)
-
-lemma sta_ltpss_tpss_conf: ∀h,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 • U1 →
- ∀L2,d,e. L1 ▶* [d, e] L2 →
- ∀T2. L2 ⊢ T1 ▶* [d, e] T2 →
- ∃∃U2. ⦃h, L2⦄ ⊢ T2 • U2 & L2 ⊢ U1 ▶* [d, e] U2.
-#h #L1 #T1 #U #H elim H -L1 -T1 -U
-[ #L1 #k1 #L2 #d #e #_ #T2 #H
- >(tpss_inv_sort1 … H) -H /2 width=3/
-| #L1 #K1 #V1 #W1 #U1 #i #HLK1 #HVW1 #HWU1 #IHVW1 #L2 #d #e #HL12 #T2 #H
- elim (tpss_inv_lref1 … H) -H [ | -HVW1 ]
- [ #H destruct
- elim (lt_or_ge i d) #Hdi [ -HVW1 | ]
- [ elim (ltpss_ldrop_conf_le … HL12 … HLK1 ?) -L1 /2 width=2/ #X #H #HLK2
- elim (ltpss_inv_tpss11 … H ?) -H /2 width=1/ #K2 #V2 #HK12 #HV12 #H destruct
- elim (IHVW1 … HK12 … HV12) -IHVW1 -HK12 -HV12 #W2 #HVW2 #HW12
- lapply (ldrop_fwd_ldrop2 … HLK2) #H
- elim (lift_total W2 0 (i+1)) #U2 #HWU2
- lapply (tpss_lift_ge … HW12 … H … HWU1 … HWU2) // -HW12 -H -HWU1
- >minus_plus <plus_minus_m_m // -Hdi /3 width=6/
- | elim (lt_or_ge i (d + e)) #Hide [ -HVW1 | -Hdi -IHVW1 ]
- [ elim (ltpss_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HLK2
- elim (ltpss_inv_tpss21 … H ?) -H /2 width=1/ #K2 #V2 #HK12 #HV12 #H destruct
- elim (IHVW1 … HK12 … HV12) -IHVW1 -HK12 -HV12 #W2 #HVW2 #HW12
- lapply (ldrop_fwd_ldrop2 … HLK2) #H
- elim (lift_total W2 0 (i+1)) #U2 #HWU2
- lapply (tpss_lift_ge … HW12 … H … HWU1 … HWU2) // -HW12 -H -HWU1 >minus_plus #H
- lapply (tpss_weak … H d e ? ?) [1,2: normalize [ >commutative_plus <plus_minus_m_m // | /2 width=1/ ]] -Hdi -Hide /3 width=6/
- | lapply (ltpss_ldrop_conf_ge … HL12 … HLK1 ?) -L1 // -Hide /3 width=6/
- ]
- ]
- | * #K2 #V2 #W2 #Hdi #Hide #HLK2 #HVW2 #HWT2
- elim (ltpss_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HL2K0
- elim (ltpss_inv_tpss21 … H ?) -H /2 width=1/ #K0 #V0 #HK12 #HV12 #H destruct
- lapply (ldrop_mono … HL2K0 … HLK2) -HL2K0 #H destruct
- lapply (ldrop_fwd_ldrop2 … HLK2) -HLK2 #HLK2
- lapply (tpss_trans_eq … HV12 HVW2) -V2 #HV1W2
- elim (IHVW1 … HK12 … HV1W2) -IHVW1 -HK12 -HV1W2 #V2 #HWV2 #HW1V2
- elim (lift_total V2 0 (i+1)) #U2 #HVU2
- lapply (sta_lift … HWV2 … HLK2 … HWT2 … HVU2) -HWV2 -HWT2 #HTU2
- lapply (tpss_lift_ge … HW1V2 … HLK2 … HWU1 … HVU2) // -HW1V2 -HLK2 -HWU1 -HVU2 >minus_plus #H
- lapply (tpss_weak … H d e ? ?) [1,2: normalize [ >commutative_plus <plus_minus_m_m // | /2 width=1/ ]] -Hdi -Hide /2 width=3/
- ]
-| #L1 #K1 #W1 #V1 #U1 #i #HLK1 #HWV1 #HWU1 #IHWV1 #L2 #d #e #HL12 #T2 #H
- elim (tpss_inv_lref1 … H) -H [ | -HWV1 -HWU1 -IHWV1 ]
- [ #H destruct
- elim (lt_or_ge i d) #Hdi [ -HWV1 ]
- [ elim (ltpss_ldrop_conf_le … HL12 … HLK1 ?) -L1 /2 width=2/ #X #H #HLK2
- elim (ltpss_inv_tpss11 … H ?) -H /2 width=1/ #K2 #W2 #HK12 #HW12 #H destruct
- elim (IHWV1 … HK12 … HW12) -IHWV1 -HK12 #V2 #HWV2 #_
- lapply (ldrop_fwd_ldrop2 … HLK2) #HLK
- elim (lift_total W2 0 (i+1)) #U2 #HWU2
- lapply (tpss_lift_ge … HW12 … HLK … HWU1 … HWU2) // -HW12 -HLK -HWU1
- >minus_plus <plus_minus_m_m // -Hdi /3 width=6/
- | elim (lt_or_ge i (d + e)) #Hide [ -HWV1 | -IHWV1 -Hdi ]
- [ elim (ltpss_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HLK2
- elim (ltpss_inv_tpss21 … H ?) -H /2 width=1/ #K2 #W2 #HK12 #HW12 #H destruct
- elim (IHWV1 … HK12 … HW12) -IHWV1 -HK12 #V2 #HWV2 #_
- lapply (ldrop_fwd_ldrop2 … HLK2) #HLK
- elim (lift_total W2 0 (i+1)) #U2 #HWU2
- lapply (tpss_lift_ge … HW12 … HLK … HWU1 … HWU2) // -HW12 -HLK -HWU1 >minus_plus #H
- lapply (tpss_weak … H d e ? ?) [1,2: normalize [ >commutative_plus <plus_minus_m_m // | /2 width=1/ ]] -Hdi -Hide /3 width=6/
- | lapply (ltpss_ldrop_conf_ge … HL12 … HLK1 ?) -L1 // -Hide /3 width=6/
- ]
- ]
- | * #K2 #V2 #W2 #Hdi #Hide #HLK2 #_ #_
- elim (ltpss_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HL2K0
- elim (ltpss_inv_tpss21 … H ?) -H /2 width=1/ -Hdi -Hide #K0 #V0 #_ #_ #H destruct
- lapply (ldrop_mono … HL2K0 … HLK2) -HL2K0 -HLK2 #H destruct
- ]
-| #I #L1 #V1 #T1 #U1 #_ #IHTU1 #L2 #d #e #HL12 #X #H
- elim (tpss_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
- elim (IHTU1 … HT12) -IHTU1 -HT12 /2 width=1/ -HL12 /3 width=5/
-| #L1 #V1 #T1 #U1 #_ #IHTU1 #L2 #d #e #HL12 #X #H
- elim (tpss_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
- elim (IHTU1 … HT12) -IHTU1 -HT12 // -HL12 /3 width=5/
-| #L1 #W1 #T1 #U1 #_ #IHTU1 #L2 #d #e #HL12 #X #H
- elim (tpss_inv_flat1 … H) -H #W2 #T2 #HW12 #HT12 #H destruct
- elim (IHTU1 … HT12) -IHTU1 -HT12 // -HL12 /3 width=3/
-]
-qed.
-
-lemma sta_ltpss_tps_conf: ∀h,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 • U1 →
- ∀L2,d,e. L1 ▶* [d, e] L2 →
- ∀T2. L2 ⊢ T1 ▶ [d, e] T2 →
- ∃∃U2. ⦃h, L2⦄ ⊢ T2 • U2 & L2 ⊢ U1 ▶* [d, e] U2.
-/3 width=5/ qed.
-
-lemma sta_ltpss_conf: ∀h,L1,T,U1. ⦃h, L1⦄ ⊢ T • U1 →
- ∀L2,d,e. L1 ▶* [d, e] L2 →
- ∃∃U2. ⦃h, L2⦄ ⊢ T • U2 & L2 ⊢ U1 ▶* [d, e] U2.
-/2 width=5/ qed.
-
-lemma sta_tpss_conf: ∀h,L,T1,U1. ⦃h, L⦄ ⊢ T1 • U1 →
- ∀T2,d,e. L ⊢ T1 ▶* [d, e] T2 →
- ∃∃U2. ⦃h, L⦄ ⊢ T2 • U2 & L ⊢ U1 ▶* [d, e] U2.
-/2 width=5/ qed.
-
-lemma sta_tps_conf: ∀h,L,T1,U1. ⦃h, L⦄ ⊢ T1 • U1 →
- ∀T2,d,e. L ⊢ T1 ▶ [d, e] T2 →
- ∃∃U2. ⦃h, L⦄ ⊢ T2 • U2 & L ⊢ U1 ▶* [d, e] U2.
-/2 width=5/ qed.
(* *)
(**************************************************************************)
-include "basic_2/substitution/ldrop_ldrop.ma".
+include "basic_2/relocation/ldrop_ldrop.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,L,T,U1. ⦃h, L⦄ ⊢ T • U1 →
- ∀U2. ⦃h, L⦄ ⊢ T • U2 → U1 = U2.
-#h #L #T #U1 #H elim H -L -T -U1
-[ #L #k #X #H >(sta_inv_sort1 … H) -X //
-| #L #K #V #W #U1 #i #HLK #_ #HWU1 #IHVW #U2 #H
+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 (ldrop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct
lapply (IHVW … HVW0) -IHVW -HVW0 #H destruct
>(lift_mono … HWU1 … HW0U2) -W0 -U1 //
-| #L #K #W #V #U1 #i #HLK #_ #HWU1 #IHWV #U2 #H
+| #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 (ldrop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct
lapply (IHWV … HWV0) -IHWV -HWV0 #H destruct
>(lift_mono … HWU1 … HV0U2) -W -U1 //
-| #I #L #V #T #U1 #_ #IHTU1 #X #H
+| #a #I #G #L #V #T #U1 #_ #IHTU1 #X #H
elim (sta_inv_bind1 … H) -H #U2 #HTU2 #H destruct /3 width=1/
-| #L #V #T #U1 #_ #IHTU1 #X #H
+| #G #L #V #T #U1 #_ #IHTU1 #X #H
elim (sta_inv_appl1 … H) -H #U2 #HTU2 #H destruct /3 width=1/
-| #L #W #T #U1 #_ #IHTU1 #U2 #H
+| #G #L #W #T #U1 #_ #IHTU1 #U2 #H
lapply (sta_inv_cast1 … H) -H /2 width=1/
]
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 }.
| #L1 #H1 #V1 #IH * normalize
[ #T1 #T2 #_ <plus_n_Sm #H destruct
| #L2 #I2 #V2 #T1 #T2 #H1 #H2
- elim (IH … H1 ?) -IH -H1 /2 width=1/ -H2 #H1 #H2 destruct /2 width=1/
+ elim (IH … H1) -IH -H1 /2 width=1/ -H2 #H1 #H2 destruct /2 width=1/
]
]
qed-.
[ #L1 #L2 #_ <plus_n_Sm #H destruct
| #K2 #I2 #V2 #L1 #L2 #H1 #H2
elim (destruct_lpair_lpair … H1) -H1 #H1 #H3 #H4 destruct (**) (* destruct lemma needed *)
- elim (IH … H1 ?) -IH -H1 // -H2 /2 width=1/
+ elim (IH … H1) -IH -H1 // -H2 /2 width=1/
]
]
qed-.
elim (plus_xySz_x_false … (sym_eq … H))
| #K2 #I2 #V2 #L1 #L2 #H1 #H2
elim (destruct_lpair_lpair … H1) -H1 #H1 #H3 #H4 destruct (**) (* destruct lemma needed *)
- elim (IH … H1 ?) -IH -H1 // -H2 /2 width=1/
+ elim (IH … H1) -IH -H1 // -H2 /2 width=1/
]
]
qed-.
lemma append_inv_refl_dx: ∀L,K. L @@ K = L → K = ⋆.
#L #K #H
-elim (append_inj_dx … (⋆) … H ?) //
+elim (append_inj_dx … (⋆) … H) //
qed-.
lemma append_inv_pair_dx: ∀I,L,K,V. L @@ K = L.ⓑ{I}V → K = ⋆.ⓑ{I}V.
#I #L #K #V #H
-elim (append_inj_dx … (⋆.ⓑ{I}V) … H ?) //
+elim (append_inj_dx … (⋆.ⓑ{I}V) … H) //
qed-.
lemma length_inv_pos_dx_append: ∀d,L. |L| = d + 1 →
@(ex2_3_intro … (⋆)) // (**) (* explicit constructor *)
| #d #IHd #L #H elim (length_inv_pos_sn … H) -H #I #K #V #H1 #H2 destruct
>H1 in IHd; -H1 #IHd
- elim (IHd K ?) -IHd // #J #L #W #H1 #H2 destruct
+ elim (IHd K) -IHd // #J #L #W #H1 #H2 destruct
@(ex2_3_intro … (L.ⓑ{I}V)) // (**) (* explicit constructor *)
>append_length /2 width=1/
]
c: conversion
d: decomposed extended reduction
+e: decomposed extended conversion
q: restricted reduction
r: reduction
s: substitution
--- /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 T ▪ break [ term 46 h , break term 46 g ] break term 46 l )"
+ non associative with precedence 45
+ for @{ 'Degree $h $g $G $L $T $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 *)
+(* *)
+(**************************************************************************)
+
+(* 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 g ] break term 46 T2 )"
+ non associative with precedence 45
+ for @{ 'DPConvStar $h $g $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( G ⊢ break term 46 L1 ▪ ⊑ break [ term 46 h, break term 46 g ] break term 46 L2 )"
+ non associative with precedence 45
+ for @{ 'LRSubEqD $h $g $G $L1 $L2 }.
--- /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 g ] break term 46 T2 )"
+ non associative with precedence 45
+ for @{ 'StaticType $h $g $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 g ] break ⦃ term 46 l , break term 46 T2 ⦄ )"
- non associative with precedence 45
- for @{ 'StaticType $h $g $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 *)
-(* *)
-(**************************************************************************)
-
-(* 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 g ] break term 46 T2 )"
- non associative with precedence 45
- for @{ 'StaticTypeStar $h $g $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 g , break term 46 l ] break term 46 T2 )"
+ non associative with precedence 45
+ for @{ 'StaticTypeStar $h $g $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 g , break term 46 l ] break term 46 T2 )"
+ non associative with precedence 45
+ for @{ 'StaticTypeStarAlt $h $g $G $L $l $T1 $T2 }.
(* Basic forward lemmas *****************************************************)
-lemma cix_inv_cir: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ 𝐈[h, g]⦃T⦄ → ⦃G, L⦄ ⊢ 𝐈⦃T⦄.
+lemma cix_inv_cir: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ 𝐈[h, g]⦃T⦄ → ⦃G, L⦄ ⊢ 𝐈⦃T⦄.
/3 width=1/ qed-.
(* Basic properties *********************************************************)
lemma cnx_sort: ∀h,g,G,L,k. deg h g k 0 → ⦃G, L⦄ ⊢ 𝐍[h, g]⦃⋆k⦄.
#h #g #G #L #k #Hk #X #H elim (cpx_inv_sort1 … H) -H // * #l #Hkl #_
-lapply (deg_mono … Hkl Hk) -h -L <plus_n_Sm #H destruct
+lapply (deg_mono … Hkl Hk) -h -L <plus_n_Sm #H destruct
qed.
lemma cnx_sort_iter: ∀h,g,G,L,k,l. deg h g k l → ⦃G, L⦄ ⊢ 𝐍[h, g]⦃⋆((next h)^l k)⦄.
#h #g #G #L #T #H elim H -L -T
[ #L #k #l #Hkl #H
lapply (cnx_inv_sort … H) -H #H
- lapply (deg_mono … H Hkl) -h -L -k <plus_n_Sm #H destruct
+ lapply (deg_mono … H Hkl) -h -L -k <plus_n_Sm #H destruct
| #I #L #K #V #i #HLK #H
elim (cnx_inv_delta … HLK H)
| #L #V #T #_ #IHV #H
#I #G #L #V1 #T1 #T #H #b
elim (cpr_inv_bind1 … H) -H *
[ #V2 #T2 #HV12 #HT12 #H destruct /3 width=4/
-| #T2 #_ #_ #H destruct
+| #T2 #_ #_ #H destruct
]
qed-.
elim (IHV12 … HLK … HV01) -V1 #V3 #HV32 #HV03
elim (IHT12 (K.ⓛW0) … HT01) -T1 /2 width=1/ #T3 #HT32 #HT03
elim (IHW12 … HLK … HW01) -W1 #W3 #HW32 #HW03
- @ex2_intro [2: /3 width=2/ | skip |3: /2 width=1/ ] (**) (* /4 width=6/ is slow *)
+ @ex2_intro [2: /3 width=2/ | skip |3: /2 width=1/ ] (**) (* /4 width=6/ is slow *)
| #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #K #d #e #HLK #X #HX
elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct
elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct
(**************************************************************************)
include "basic_2/notation/relations/pred_6.ma".
-include "basic_2/static/ssta.ma".
+include "basic_2/static/sd.ma".
include "basic_2/reduction/cpr.ma".
(* CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION FOR TERMS ******************)
#h #g #G #L #T1 #T2 #H elim H -L -T1 -T2 // /2 width=1/ /2 width=3/ /2 width=7/
qed.
-fact ssta_cpx_aux: ∀h,g,G,L,T1,T2,l0. ⦃G, L⦄ ⊢ T1 •[h, g] ⦃l0, T2⦄ →
- ∀l. l0 = l+1 → ⦃G, L⦄ ⊢ T1 ➡[h, g] T2.
-#h #g #G #L #T1 #T2 #l0 #H elim H -L -T1 -T2 -l0 /2 width=2/ /2 width=7/ /3 width=2/ /3 width=7/
-qed-.
-
-lemma ssta_cpx: ∀h,g,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 •[h, g] ⦃l+1, T2⦄ → ⦃G, L⦄ ⊢ T1 ➡[h, g] T2.
-/2 width=4 by ssta_cpx_aux/ qed.
-
lemma cpx_pair_sn: ∀h,g,I,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 →
∀T. ⦃G, L⦄ ⊢ ②{I}V1.T ➡[h, g] ②{I}V2.T.
#h #g * /2 width=1/ qed.
#h #g #I #G #L #V1 #T1 #T #H #b
elim (cpx_inv_bind1 … H) -H *
[ #V2 #T2 #HV12 #HT12 #H destruct /3 width=4/
-| #T2 #_ #_ #H destruct
+| #T2 #_ #_ #H destruct
]
qed-.
include "basic_2/relocation/ldrop_ldrop.ma".
include "basic_2/relocation/fsupq.ma".
+include "basic_2/static/ssta.ma".
include "basic_2/reduction/cpx.ma".
(* CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION FOR TERMS ******************)
+(* Advanced properties ******************************************************)
+
+lemma ssta_cpx: ∀h,g,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 •[h, g] 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
+[ #G #L #k #H lapply (da_inv_sort … H) -H /2 width=2/
+| #G #L #K #V #U #W #i #HLK #_ #HWU #IHVW #H
+ elim (da_inv_lref … H) -H * #K0 #V0 [| #l0 ] #HLK0
+ lapply (ldrop_mono … HLK0 … HLK) -HLK0 #H destruct /3 width=7/
+| #G #L #K #W #U #l0 #i #HLK #_ #HWU #H
+ elim (da_inv_lref … H) -H * #K0 #W0 [| #l1 ] #HLK0
+ lapply (ldrop_mono … HLK0 … HLK) -HLK0 #H destruct /2 width=7/
+| #a #I #G #L #V #T #U #_ #IHTU #H lapply (da_inv_bind … H) -H /3 width=1/
+| #G #L #V #T #U #_ #IHTU #H lapply (da_inv_flat … H) -H /3 width=1/
+| #G #L #W #T #U #_ #IHTU #H lapply (da_inv_flat … H) -H /3 width=1/
+]
+qed.
+
(* Relocation properties ****************************************************)
lemma cpx_lift: ∀h,g,G. l_liftable (cpx h g G).
elim (IHV12 … HLK … HV01) -V1 #V3 #HV32 #HV03
elim (IHT12 (K.ⓛW0) … HT01) -T1 /2 width=1/ #T3 #HT32 #HT03
elim (IHW12 … HLK … HW01) -W1 #W3 #HW32 #HW03
- @ex2_intro [2: /3 width=2/ | skip |3: /2 width=1/ ] (**) (* /4 width=6/ is slow *)
+ @ex2_intro [2: /3 width=2/ | skip |3: /2 width=1/ ] (**) (* /4 width=6/ is slow *)
| #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #K #d #e #HLK #X #HX
elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct
elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct
qed-.
lemma fsupq_ssta_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊃⸮ ⦃G2, L2, T2⦄ →
- ∀U2,l. ⦃G2, L2⦄ ⊢ T2 •[h, g] ⦃l+1, U2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 •[h, g] U2 →
+ ∀l. ⦃G2, L2⦄ ⊢ T2 ▪[h, g] l+1 →
∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, g] U1 & ⦃G1, L1, U1⦄ ⊃⸮ ⦃G2, L2, U2⦄.
-/3 width=4 by fsupq_cpx_trans, ssta_cpx/ qed-.
+/3 width=5 by fsupq_cpx_trans, ssta_cpx/ qed-.
lemma fsup_cpx_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊃ ⦃G2, L2, T2⦄ →
∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, g] U2 →
/3 width=3 by fsupq_cpx_trans, fsup_fsupq/ qed-.
lemma fsup_ssta_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊃ ⦃G2, L2, T2⦄ →
- ∀U2,l. ⦃G2, L2⦄ ⊢ T2 •[h, g] ⦃l+1, U2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 •[h, g] U2 →
+ ∀l. ⦃G2, L2⦄ ⊢ T2 ▪[h, g] l+1 →
∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, g] U1 & ⦃G1, L1, U1⦄ ⊃⸮ ⦃G2, L2, U2⦄.
-/3 width=4 by fsupq_ssta_trans, fsup_fsupq/ qed-.
+/3 width=5 by fsupq_ssta_trans, fsup_fsupq/ qed-.
(* SN PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS *****************************)
-definition lpr: relation3 genv lenv lenv ≝ λG. lpx_sn (cpr G).
+definition lpr: relation3 genv lenv lenv ≝ λG. lpx_sn (cpr G).
interpretation "parallel reduction (local environment, sn variant)"
'PRedSn G L1 L2 = (lpr G L1 L2).
#T2 #HT02 #X2 #HXT2 #L1 #HL01 #L2 #HL02
elim (IH … HT01 … HT02 (L1.ⓓV0) … (L2.ⓓV0)) -IH -HT01 -HT02 // /2 width=1/ -L0 -T0 #T #HT1 #HT2
elim (cpr_inv_lift1 … HT1 L1 … HXT1) -T1 /2 width=1/ #T1 #HT1 #HXT1
-elim (cpr_inv_lift1 … HT2 L2 … HXT2) -T2 /2 width=1/ #T2 #HT2 #HXT2
+elim (cpr_inv_lift1 … HT2 L2 … HXT2) -T2 /2 width=1/ #T2 #HT2 #HXT2
lapply (lift_inj … HT2 … HT1) -T #H destruct /2 width=3/
qed-.
elim (IH … HT01 … HT02 (L1.ⓓW2) … (L2.ⓓW2)) /2 width=1/ -L0 -V0 -W0 -T0 #T #HT1 #HT2
elim (cpr_inv_lift1 … HT1 L1 … HXT1) -HXT1 /2 width=1/ #Y #HYT #HXY
@(ex2_intro … (ⓐV.Y)) /2 width=1/ /3 width=5/ (**) (* auto /4 width=9/ is too slow *)
-]
+]
qed-.
fact cpr_conf_lpr_beta_beta:
(* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************)
definition lpx: ∀h. sd h → relation3 genv lenv lenv ≝
- λh,g,G. lpx_sn (cpx h g G).
+ λh,g,G. lpx_sn (cpx h g G).
interpretation "extended parallel reduction (local environment, sn variant)"
'PRedSn h g G L1 L2 = (lpx h g G L1 L2).
| #L #K #I #V #e #_ #IHLK #L2 #H
lapply (ldrop_inv_ldrop1 … H ?) -H // /2 width=1/
| #L #K1 #I #T #V1 #d #e #_ #HVT1 #IHLK1 #X #H
- elim (ldrop_inv_skip1 … H ?) -H // <minus_plus_m_m #K2 #V2 #HLK2 #HVT2 #H destruct
+ elim (ldrop_inv_skip1 … H) -H // <minus_plus_m_m #K2 #V2 #HLK2 #HVT2 #H destruct
>(lift_inj … HVT1 … HVT2) -HVT1 -HVT2
>(IHLK1 … HLK2) -IHLK1 -HLK2 //
]
lapply (ldrop_inv_O1_pair1 … H) -H * * #He2 #HL20
[ -IHLK0 -He21 destruct <minus_n_O /3 width=3/
| -HLK0 <minus_le_minus_minus_comm //
- elim (IHLK0 … HL20 ? ?) -L0 // /2 width=1/ /2 width=3/
+ elim (IHLK0 … HL20) -L0 // /2 width=1/ /2 width=3/
]
| #L0 #K0 #I #V0 #V1 #d1 #e1 >plus_plus_comm_23 #_ #_ #IHLK0 #L2 #e2 #H #Hd1e2 #He2de1
elim (le_inv_plus_l … Hd1e2) #_ #He2
lapply (le_n_O_to_eq … H) -H #H destruct /2 width=3/
| #L1 #L2 #I #V #e #_ #IHL12 #e2 #L #HL2 #H
lapply (le_n_O_to_eq … H) -H #H destruct
- elim (IHL12 … HL2 ?) -IHL12 -HL2 // #L0 #H #HL0
+ elim (IHL12 … HL2) -IHL12 -HL2 // #L0 #H #HL0
lapply (ldrop_inv_O2 … H) -H #H destruct /3 width=5/
| #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #IHL12 #e2 #L #H #He2d
elim (ldrop_inv_O1_pair1 … H) -H *
elim (lpx_sn_inv_pair2 … H) -H #K1 #V1 #HK12 #HV12 #H destruct /3 width=5/
| #L2 #K2 #I #V2 #e #_ #IHLK2 #X #H #_
elim (lpx_sn_inv_pair2 … H) -H #L1 #V1 #HL12 #HV12 #H destruct
- elim (IHLK2 … HL12 ?) -L2 // /3 width=3/
+ elim (IHLK2 … HL12) -L2 // /3 width=3/
| #L2 #K2 #I #V2 #W2 #d #e #_ #_ #_ #L1 #_
>commutative_plus normalize #H destruct
]
* [ #a ] #I #T2 #V1 #U1 #d #e #H
[ elim (lift_inv_bind1 … H) -H /2 width=4/
| elim (lift_inv_flat1 … H) -H /2 width=4/
-]
+]
qed-.
lemma lift_fwd_pair2: ∀I,T1,V2,U2,d,e. ⇧[d,e] T1 ≡ ②{I}V2.U2 →
* [ #a ] #I #T1 #V2 #U2 #d #e #H
[ elim (lift_inv_bind2 … H) -H /2 width=4/
| elim (lift_inv_flat2 … H) -H /2 width=4/
-]
+]
qed-.
lemma lift_fwd_tw: ∀d,e,T1,T2. ⇧[d, e] T1 ≡ T2 → ♯{T1} = ♯{T2}.
lapply (lift_inv_gref2 … Hk) -Hk #Hk destruct /3 width=3/
| #a #I #W1 #W #U1 #U #d1 #e1 #_ #_ #IHW #IHU #d2 #e2 #T2 #H #Hd12
lapply (lift_inv_bind2 … H) -H * #W2 #U2 #HW2 #HU2 #H destruct
- elim (IHW … HW2 ?) // -IHW -HW2 #W0 #HW2 #HW1
- >plus_plus_comm_23 in HU2; #HU2 elim (IHU … HU2 ?) /2 width=1/ /3 width=5/
+ elim (IHW … HW2) // -IHW -HW2 #W0 #HW2 #HW1
+ >plus_plus_comm_23 in HU2; #HU2 elim (IHU … HU2) /2 width=1/ /3 width=5/
| #I #W1 #W #U1 #U #d1 #e1 #_ #_ #IHW #IHU #d2 #e2 #T2 #H #Hd12
lapply (lift_inv_flat2 … H) -H * #W2 #U2 #HW2 #HU2 #H destruct
- elim (IHW … HW2 ?) // -IHW -HW2 #W0 #HW2 #HW1
- elim (IHU … HU2 ?) // /3 width=5/
+ elim (IHW … HW2) // -IHW -HW2 #W0 #HW2 #HW1
+ elim (IHU … HU2) // /3 width=5/
]
qed.
>(lift_inv_lref2_lt … H) -H [ /3 width=3/ | /2 width=3/ ]
| #i #d1 #e1 #Hid1 #e #e2 #T2 #H #He1 #He1e2
elim (lt_or_ge (i+e1) (d1+e+e2)) #Hie1d1e2
- [ elim (lift_inv_lref2_be … H ? ?) -H // /2 width=1/
+ [ elim (lift_inv_lref2_be … H) -H // /2 width=1/
| >(lift_inv_lref2_ge … H ?) -H //
lapply (le_plus_to_minus … Hie1d1e2) #Hd1e21i
elim (le_inv_plus_l … Hie1d1e2) -Hie1d1e2 #Hd1e12 #He2ie1
| #p #d1 #e1 #e #e2 #T2 #H >(lift_inv_gref2 … H) -H /2 width=3/
| #a #I #V1 #V #T1 #T #d1 #e1 #_ #_ #IHV1 #IHT1 #e #e2 #X #H #He1 #He1e2
elim (lift_inv_bind2 … H) -H #V2 #T2 #HV2 #HT2 #H destruct
- elim (IHV1 … HV2 ? ?) -V // >plus_plus_comm_23 in HT2; #HT2
- elim (IHT1 … HT2 ? ?) -T // -He1 -He1e2 /3 width=5/
+ elim (IHV1 … HV2) -V // >plus_plus_comm_23 in HT2; #HT2
+ elim (IHT1 … HT2) -T // -He1 -He1e2 /3 width=5/
| #I #V1 #V #T1 #T #d1 #e1 #_ #_ #IHV1 #IHT1 #e #e2 #X #H #He1 #He1e2
elim (lift_inv_flat2 … H) -H #V2 #T2 #HV2 #HT2 #H destruct
- elim (IHV1 … HV2 ? ?) -V //
- elim (IHT1 … HT2 ? ?) -T // -He1 -He1e2 /3 width=5/
+ elim (IHV1 … HV2) -V //
+ elim (IHT1 … HT2) -T // -He1 -He1e2 /3 width=5/
]
qed.
>(lift_inv_gref1 … HX) -HX /2 width=3/
| #a #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd21
elim (lift_inv_bind1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct
- elim (IHV12 … HV20 ?) -IHV12 -HV20 //
- elim (IHT12 … HT20 ?) -IHT12 -HT20 /2 width=1/ /3 width=5/
+ elim (IHV12 … HV20) -IHV12 -HV20 //
+ elim (IHT12 … HT20) -IHT12 -HT20 /2 width=1/ /3 width=5/
| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd21
elim (lift_inv_flat1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct
- elim (IHV12 … HV20 ?) -IHV12 -HV20 //
- elim (IHT12 … HT20 ?) -IHT12 -HT20 // /3 width=5/
+ elim (IHV12 … HV20) -IHV12 -HV20 //
+ elim (IHT12 … HT20) -IHT12 -HT20 // /3 width=5/
]
qed.
>(lift_inv_gref1 … HX) -HX /2 width=3/
| #a #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hded
elim (lift_inv_bind1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct
- elim (IHV12 … HV20 ?) -IHV12 -HV20 //
- elim (IHT12 … HT20 ?) -IHT12 -HT20 /2 width=1/ #T
+ elim (IHV12 … HV20) -IHV12 -HV20 //
+ elim (IHT12 … HT20) -IHT12 -HT20 /2 width=1/ #T
<plus_minus /2 width=2/ /3 width=5/
| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hded
elim (lift_inv_flat1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct
- elim (IHV12 … HV20 ?) -IHV12 -HV20 //
- elim (IHT12 … HT20 ?) -IHT12 -HT20 // /3 width=5/
+ elim (IHV12 … HV20) -IHV12 -HV20 //
+ elim (IHT12 … HT20) -IHT12 -HT20 // /3 width=5/
]
qed.
∃∃T0. ⇧[0, e2] T1 ≡ T0 & ⇧[d1 + e2, e1] T2 ≡ T0.
#T #T1 #d1 #e1 #HT1 #T2 #e2 #HT2
elim (lift_total T1 0 e2) #T0 #HT10
-elim (lift_trans_le … HT1 … HT10 ?) -HT1 // #X #HTX #HT20
+elim (lift_trans_le … HT1 … HT10) -HT1 // #X #HTX #HT20
lapply (lift_mono … HTX … HT2) -T #H destruct /2 width=3/
qed.
lemma lift_conf_be: ∀T,T1,d,e1. ⇧[d, e1] T ≡ T1 → ∀T2,e2. ⇧[d, e2] T ≡ T2 →
e1 ≤ e2 → ⇧[d + e1, e2 - e1] T1 ≡ T2.
#T #T1 #d #e1 #HT1 #T2 #e2 #HT2 #He12
-elim (lift_split … HT2 (d+e1) e1 ? ? ?) -HT2 // #X #H
+elim (lift_split … HT2 (d+e1) e1) -HT2 // #X #H
>(lift_mono … H … HT1) -T //
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.ma".
+include "basic_2/static/aaa.ma".
+
+(* ATONIC ARITY ASSIGNMENT FOR TERMS ****************************************)
+
+(* Forward lemmas on degree assignment for terms ****************************)
+
+lemma aaa_fwd_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 H -G -L -T -A
+[ #G #L #k elim (deg_total … g k) /3 width=2/
+| * #G #L #K [ #V | #W ] #B #i #HLK #_ * /3 width=5/
+| #a #G #L #V #T #B #A #_ #_ #_ * /3 width=2/
+| #a #G #L #V #T #B #A #_ #_ #_ * /3 width=2/
+| #G #L #V #T #B #A #_ #_ #_ * /3 width=2/
+| #G #L #W #T #A #_ #_ #_ * /3 width=2/
+]
+qed-.
>(lift_inv_sort1 … H) -H //
| #I #G #L1 #K1 #V1 #B #i #HLK1 #_ #IHB #L2 #d #e #HL21 #T2 #H
elim (lift_inv_lref1 … H) -H * #Hid #H destruct
- [ elim (ldrop_trans_le … HL21 … HLK1 ?) -L1 /2 width=2/ #X #HLK2 #H
- elim (ldrop_inv_skip2 … H ?) -H /2 width=1/ -Hid #K2 #V2 #HK21 #HV12 #H destruct
+ [ elim (ldrop_trans_le … HL21 … HLK1) -L1 /2 width=2/ #X #HLK2 #H
+ elim (ldrop_inv_skip2 … H) -H /2 width=1/ -Hid #K2 #V2 #HK21 #HV12 #H destruct
/3 width=8/
| lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 // -Hid /3 width=8/
]
>(lift_inv_sort2 … H) -H //
| #I #G #L2 #K2 #V2 #B #i #HLK2 #_ #IHB #L1 #d #e #HL21 #T1 #H
elim (lift_inv_lref2 … H) -H * #Hid #H destruct
- [ elim (ldrop_conf_lt … HL21 … HLK2 ?) -L2 // -Hid /3 width=8/
+ [ elim (ldrop_conf_lt … HL21 … HLK2) -L2 // -Hid /3 width=8/
| lapply (ldrop_conf_ge … HL21 … HLK2 ?) -L2 // -Hid /3 width=8/
]
| #a #G #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #d #e #HL21 #X #H
--- /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/ssta.ma".
+include "basic_2/static/aaa_lift.ma".
+include "basic_2/static/aaa_da.ma".
+
+(* ATONIC ARITY ASSIGNMENT FOR TERMS ****************************************)
+
+(* Properties on stratified static type assignment for terms ****************)
+
+lemma aaa_ssta_conf: ∀h,g,G,L. Conf3 … (aaa G L) (ssta h g G L).
+#h #g #G #L #T #A #H elim H -G -L -T -A
+[ #G #L #k #U #H
+ lapply (ssta_inv_sort1 … H) -H #H destruct //
+| #I #G #L #K #V #B #i #HLK #HV #IHV #U #H
+ elim (ssta_inv_lref1 … H) -H * #K0 #V0 #W0 #HLK0 #HVW0 #HU
+ lapply (ldrop_mono … HLK0 … HLK) -HLK0 #H0 destruct
+ lapply (ldrop_fwd_ldrop2 … HLK) -HLK #HLK
+ @(aaa_lift … HLK … HU) -HU -L // -HV /2 width=2/
+| #a #G #L #V #T #B #A #HV #_ #_ #IHT #X #H
+ elim (ssta_inv_bind1 … H) -H #U #HTU #H destruct /3 width=2/
+| #a #G #L #V #T #B #A #HV #_ #_ #IHT #X #H
+ elim (ssta_inv_bind1 … H) -H #U #HTU #H destruct /3 width=2/
+| #G #L #V #T #B #A #HV #_ #_ #IHT #X #H
+ elim (ssta_inv_appl1 … H) -H #U #HTU #H destruct /3 width=3/
+| #G #L #V #T #A #_ #_ #IHV #IHT #X #H
+ lapply (ssta_inv_cast1 … H) -H /2 width=2/
+]
+qed-.
+
+(* Forward lemmas on stratified static type assignment for terms ************)
+
+lemma aaa_fwd_ssta: ∀h,g,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∃U. ⦃G, L⦄ ⊢ T •[h, g] U.
+#h #g #G #L #T #A #H elim (aaa_fwd_da … H) -H /2 width=3 by da_ssta/
+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/degree_6.ma".
+include "basic_2/grammar/genv.ma".
+include "basic_2/relocation/ldrop.ma".
+include "basic_2/static/sd.ma".
+
+(* DEGREE ASSIGNMENT FOR TERMS **********************************************)
+
+(* activate genv *)
+inductive da (h:sh) (g:sd h): relation4 genv lenv term nat ≝
+| da_sort: ∀G,L,k,l. deg h g k l → da h g G L (⋆k) l
+| da_ldef: ∀G,L,K,V,i,l. ⇩[0, i] L ≡ K.ⓓV → da h g G K V l → da h g G L (#i) l
+| da_ldec: ∀G,L,K,W,i,l. ⇩[0, i] L ≡ K.ⓛW → da h g G K W l → da h g G L (#i) (l+1)
+| da_bind: ∀a,I,G,L,V,T,l. da h g G (L.ⓑ{I}V) T l → da h g G L (ⓑ{a,I}V.T) l
+| da_flat: ∀I,G,L,V,T,l. da h g G L T l → da h g G L (ⓕ{I}V.T) l
+.
+
+interpretation "degree assignment (term)"
+ 'Degree h g G L T l = (da h g G L T l).
+
+(* Basic inversion lemmas ***************************************************)
+
+fact da_inv_sort_aux: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l →
+ ∀k0. T = ⋆k0 → deg h g k0 l.
+#h #g #G #L #T #l * -G -L -T -l
+[ #G #L #k #l #Hkl #k0 #H destruct //
+| #G #L #K #V #i #l #_ #_ #k0 #H destruct
+| #G #L #K #W #i #l #_ #_ #k0 #H destruct
+| #a #I #G #L #V #T #l #_ #k0 #H destruct
+| #I #G #L #V #T #l #_ #k0 #H destruct
+]
+qed-.
+
+lemma da_inv_sort: ∀h,g,G,L,k,l. ⦃G, L⦄ ⊢ ⋆k ▪[h, g] l → deg h g k l.
+/2 width=5 by da_inv_sort_aux/ qed-.
+
+fact da_inv_lref_aux: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l → ∀j. T = #j →
+ (∃∃K,V. ⇩[0, j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V ▪[h, g] l) ∨
+ (∃∃K,W,l0. ⇩[0, j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W ▪[h, g] l0 &
+ l = l0 + 1
+ ).
+#h #g #G #L #T #l * -G -L -T -l
+[ #G #L #k #l #_ #j #H destruct
+| #G #L #K #V #i #l #HLK #HV #j #H destruct /3 width=4/
+| #G #L #K #W #i #l #HLK #HW #j #H destruct /3 width=6/
+| #a #I #G #L #V #T #l #_ #j #H destruct
+| #I #G #L #V #T #l #_ #j #H destruct
+]
+qed-.
+
+lemma da_inv_lref: ∀h,g,G,L,j,l. ⦃G, L⦄ ⊢ #j ▪[h, g] l →
+ (∃∃K,V. ⇩[0, j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V ▪[h, g] l) ∨
+ (∃∃K,W,l0. ⇩[0, j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W ▪[h, g] l0 & l = l0+1).
+/2 width=3 by da_inv_lref_aux/ qed-.
+
+fact da_inv_gref_aux: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l → ∀p0. T = §p0 → ⊥.
+#h #g #G #L #T #l * -G -L -T -l
+[ #G #L #k #l #_ #p0 #H destruct
+| #G #L #K #V #i #l #_ #_ #p0 #H destruct
+| #G #L #K #W #i #l #_ #_ #p0 #H destruct
+| #a #I #G #L #V #T #l #_ #p0 #H destruct
+| #I #G #L #V #T #l #_ #p0 #H destruct
+]
+qed-.
+
+lemma da_inv_gref: ∀h,g,G,L,p,l. ⦃G, L⦄ ⊢ §p ▪[h, g] l → ⊥.
+/2 width=9 by da_inv_gref_aux/ qed-.
+
+fact da_inv_bind_aux: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l →
+ ∀b,J,X,Y. T = ⓑ{b,J}Y.X → ⦃G, L.ⓑ{J}Y⦄ ⊢ X ▪[h, g] l.
+#h #g #G #L #T #l * -G -L -T -l
+[ #G #L #k #l #_ #b #J #X #Y #H destruct
+| #G #L #K #V #i #l #_ #_ #b #J #X #Y #H destruct
+| #G #L #K #W #i #l #_ #_ #b #J #X #Y #H destruct
+| #a #I #G #L #V #T #l #HT #b #J #X #Y #H destruct //
+| #I #G #L #V #T #l #_ #b #J #X #Y #H destruct
+]
+qed-.
+
+lemma da_inv_bind: ∀h,g,b,J,G,L,Y,X,l. ⦃G, L⦄ ⊢ ⓑ{b,J}Y.X ▪[h, g] l → ⦃G, L.ⓑ{J}Y⦄ ⊢ X ▪[h, g] l.
+/2 width=4 by da_inv_bind_aux/ qed-.
+
+fact da_inv_flat_aux: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l →
+ ∀J,X,Y. T = ⓕ{J}Y.X → ⦃G, L⦄ ⊢ X ▪[h, g] l.
+#h #g #G #L #T #l * -G -L -T -l
+[ #G #L #k #l #_ #J #X #Y #H destruct
+| #G #L #K #V #i #l #_ #_ #J #X #Y #H destruct
+| #G #L #K #W #i #l #_ #_ #J #X #Y #H destruct
+| #a #I #G #L #V #T #l #_ #J #X #Y #H destruct
+| #I #G #L #V #T #l #HT #J #X #Y #H destruct //
+]
+qed-.
+
+lemma da_inv_flat: ∀h,g,J,G,L,Y,X,l. ⦃G, L⦄ ⊢ ⓕ{J}Y.X ▪[h, g] l → ⦃G, L⦄ ⊢ X ▪[h, g] l.
+/2 width=5 by da_inv_flat_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/da_lift.ma".
+
+(* DEGREE ASSIGNMENT FOR TERMS **********************************************)
+
+(* Main properties **********************************************************)
+
+theorem da_mono: ∀h,g,G,L,T,l1. ⦃G, L⦄ ⊢ T ▪[h, g] l1 →
+ ∀l2. ⦃G, L⦄ ⊢ T ▪[h, g] l2 → l1 = l2.
+#h #g #G #L #T #l1 #H elim H -G -L -T -l1
+[ #G #L #k #l1 #Hkl1 #l2 #H
+ lapply (da_inv_sort … H) -G -L #Hkl2
+ >(deg_mono … Hkl2 … Hkl1) -h -k -l2 //
+| #G #L #K #V #i #l1 #HLK #_ #IHV #l2 #H
+ elim (da_inv_lref … H) -H * #K0 #V0 [| #l0 ] #HLK0 #HV0 [| #Hl0 ]
+ lapply (ldrop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct /2 width=1/
+| #G #L #K #W #i #l1 #HLK #_ #IHW #l2 #H
+ elim (da_inv_lref … H) -H * #K0 #W0 [| #l0 ] #HLK0 #HW0 [| #Hl0 ]
+ lapply (ldrop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct /3 width=1/
+| #a #I #G #L #V #T #l1 #_ #IHT #l2 #H
+ lapply (da_inv_bind … H) -H /2 width=1/
+| #I #G #L #V #T #l1 #_ #IHT #l2 #H
+ lapply (da_inv_flat … H) -H /2 width=1/
+]
+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/relocation/ldrop_ldrop.ma".
+include "basic_2/static/da.ma".
+
+(* DEGREE ASSIGNMENT FOR TERMS **********************************************)
+
+(* Properties on relocation *************************************************)
+
+lemma da_lift: ∀h,g,G,L1,T1,l. ⦃G, L1⦄ ⊢ T1 ▪[h, g] l →
+ ∀L2,d,e. ⇩[d, e] L2 ≡ L1 → ∀T2. ⇧[d, e] T1 ≡ T2 →
+ ⦃G, L2⦄ ⊢ T2 ▪[h, g] l.
+#h #g #G #L1 #T1 #l #H elim H -G -L1 -T1 -l
+[ #G #L1 #k #l #Hkl #L2 #d #e #_ #X #H
+ >(lift_inv_sort1 … H) -X /2 width=1/
+| #G #L1 #K1 #V1 #i #l #HLK1 #_ #IHV1 #L2 #d #e #HL21 #X #H
+ elim (lift_inv_lref1 … H) * #Hid #H destruct
+ [ elim (ldrop_trans_le … HL21 … HLK1) -L1 /2 width=2/ #X #HLK2 #H
+ elim (ldrop_inv_skip2 … H) -H /2 width=1/ -Hid #K2 #V2 #HK21 #HV12 #H destruct
+ /3 width=7/
+ | lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 // -Hid /3 width=7/
+ ]
+| #G #L1 #K1 #W1 #i #l #HLK1 #_ #IHW1 #L2 #d #e #HL21 #X #H
+ elim (lift_inv_lref1 … H) * #Hid #H destruct
+ [ elim (ldrop_trans_le … HL21 … HLK1) -L1 /2 width=2/ #X #HLK2 #H
+ elim (ldrop_inv_skip2 … H) -H /2 width=1/ -Hid #K2 #W2 #HK21 #HW12 #H destruct
+ /3 width=7/
+ | lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 // -Hid /3 width=7/
+ ]
+| #a #I #G #L1 #V1 #T1 #l #_ #IHT1 #L2 #d #e #HL21 #X #H
+ elim (lift_inv_bind1 … H) -H #V2 #T2 #HV12 #HU12 #H destruct /4 width=4/
+| #I #G #L1 #V1 #T1 #l #_ #IHT1 #L2 #d #e #HL21 #X #H
+ elim (lift_inv_flat1 … H) -H #V2 #T2 #HV12 #HU12 #H destruct /3 width=4/
+]
+qed.
+
+(* Inversion lemmas on relocation *******************************************)
+
+lemma da_inv_lift: ∀h,g,G,L2,T2,l. ⦃G, L2⦄ ⊢ T2 ▪[h, g] l →
+ ∀L1,d,e. ⇩[d, e] L2 ≡ L1 → ∀T1. ⇧[d, e] T1 ≡ T2 →
+ ⦃G, L1⦄ ⊢ T1 ▪[h, g] l.
+#h #g #G #L2 #T2 #l #H elim H -G -L2 -T2 -l
+[ #G #L2 #k #l #Hkl #L1 #d #e #_ #X #H
+ >(lift_inv_sort2 … H) -X /2 width=1/
+| #G #L2 #K2 #V2 #i #l #HLK2 #HV2 #IHV2 #L1 #d #e #HL21 #X #H
+ elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HV2 | -IHV2 ]
+ [ elim (ldrop_conf_lt … HL21 … HLK2) -L2 // /3 width=7/
+ | lapply (ldrop_conf_ge … HL21 … HLK2 ?) -L2 // /2 width=4/
+ ]
+| #G #L2 #K2 #W2 #i #l #HLK2 #HW2 #IHW2 #L1 #d #e #HL21 #X #H
+ elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HW2 | -IHW2 ]
+ [ elim (ldrop_conf_lt … HL21 … HLK2) -L2 // /3 width=7/
+ | lapply (ldrop_conf_ge … HL21 … HLK2 ?) -L2 // /2 width=4/
+ ]
+| #a #I #G #L2 #V2 #T2 #l #_ #IHT2 #L1 #d #e #HL21 #X #H
+ elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct /4 width=4/
+| #I #G #L2 #V2 #T2 #l #_ #IHT2 #L1 #d #e #HL21 #X #H
+ elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct /3 width=4/
+]
+qed-.
| #I #L1 #L2 #V #_ #IHL12 #K1 #e #H
elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK1
[ destruct
- elim (IHL12 L1 0 ?) -IHL12 // #X #HL12 #H
+ elim (IHL12 L1 0) -IHL12 // #X #HL12 #H
<(ldrop_inv_O2 … H) in HL12; -H /3 width=3/
| elim (IHL12 … HLK1) -L1 /3 width=3/
]
--- /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/lrsubeqd_5.ma".
+include "basic_2/substitution/lsubr.ma".
+include "basic_2/static/da.ma".
+
+(* LOCAL ENVIRONMENT REFINEMENT FOR DEGREE ASSIGNMENT ***********************)
+
+inductive lsubd (h) (g) (G): relation lenv ≝
+| lsubd_atom: lsubd h g G (⋆) (⋆)
+| lsubd_pair: ∀I,L1,L2,V. lsubd h g G L1 L2 →
+ lsubd h g G (L1.ⓑ{I}V) (L2.ⓑ{I}V)
+| lsubd_abbr: ∀L1,L2,W,V,l. ⦃G, L1⦄ ⊢ V ▪[h, g] l+1 → ⦃G, L2⦄ ⊢ W ▪[h, g] l →
+ lsubd h g G L1 L2 → lsubd h g G (L1.ⓓⓝW.V) (L2.ⓛW)
+.
+
+interpretation
+ "local environment refinement (degree assignment)"
+ 'LRSubEqD h g G L1 L2 = (lsubd h g G L1 L2).
+
+(* Basic_forward lemmas *****************************************************)
+
+lemma lsubd_fwd_lsubr: ∀h,g,G,L1,L2. G ⊢ L1 ▪⊑[h, g] L2 → L1 ⊑ L2.
+#h #g #G #L1 #L2 #H elim H -L1 -L2 // /2 width=1/
+qed-.
+
+(* Basic inversion lemmas ***************************************************)
+
+fact lsubd_inv_atom1_aux: ∀h,g,G,L1,L2. G ⊢ L1 ▪⊑[h, g] L2 → L1 = ⋆ → L2 = ⋆.
+#h #g #G #L1 #L2 * -L1 -L2
+[ //
+| #I #L1 #L2 #V #_ #H destruct
+| #L1 #L2 #W #V #l #_ #_ #_ #H destruct
+]
+qed-.
+
+lemma lsubd_inv_atom1: ∀h,g,G,L2. G ⊢ ⋆ ▪⊑[h, g] L2 → L2 = ⋆.
+/2 width=6 by lsubd_inv_atom1_aux/ qed-.
+
+fact lsubd_inv_pair1_aux: ∀h,g,G,L1,L2. G ⊢ L1 ▪⊑[h, g] L2 →
+ ∀I,K1,X. L1 = K1.ⓑ{I}X →
+ (∃∃K2. G ⊢ K1 ▪⊑[h, g] K2 & L2 = K2.ⓑ{I}X) ∨
+ ∃∃K2,W,V,l. ⦃G, K1⦄ ⊢ V ▪[h, g] l+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] l &
+ G ⊢ K1 ▪⊑[h, g] K2 &
+ I = Abbr & L2 = K2.ⓛW & X = ⓝW.V.
+#h #g #G #L1 #L2 * -L1 -L2
+[ #J #K1 #X #H destruct
+| #I #L1 #L2 #V #HL12 #J #K1 #X #H destruct /3 width=3/
+| #L1 #L2 #W #V #l #HV #HW #HL12 #J #K1 #X #H destruct /3 width=9/
+]
+qed-.
+
+lemma lsubd_inv_pair1: ∀h,g,I,G,K1,L2,X. G ⊢ K1.ⓑ{I}X ▪⊑[h, g] L2 →
+ (∃∃K2. G ⊢ K1 ▪⊑[h, g] K2 & L2 = K2.ⓑ{I}X) ∨
+ ∃∃K2,W,V,l. ⦃G, K1⦄ ⊢ V ▪[h, g] l+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] l &
+ G ⊢ K1 ▪⊑[h, g] K2 &
+ I = Abbr & L2 = K2.ⓛW & X = ⓝW.V.
+/2 width=3 by lsubd_inv_pair1_aux/ qed-.
+
+fact lsubd_inv_atom2_aux: ∀h,g,G,L1,L2. G ⊢ L1 ▪⊑[h, g] L2 → L2 = ⋆ → L1 = ⋆.
+#h #g #G #L1 #L2 * -L1 -L2
+[ //
+| #I #L1 #L2 #V #_ #H destruct
+| #L1 #L2 #W #V #l #_ #_ #_ #H destruct
+]
+qed-.
+
+lemma lsubd_inv_atom2: ∀h,g,G,L1. G ⊢ L1 ▪⊑[h, g] ⋆ → L1 = ⋆.
+/2 width=6 by lsubd_inv_atom2_aux/ qed-.
+
+fact lsubd_inv_pair2_aux: ∀h,g,G,L1,L2. G ⊢ L1 ▪⊑[h, g] L2 →
+ ∀I,K2,W. L2 = K2.ⓑ{I}W →
+ (∃∃K1. G ⊢ K1 ▪⊑[h, g] K2 & L1 = K1.ⓑ{I}W) ∨
+ ∃∃K1,V,l. ⦃G, K1⦄ ⊢ V ▪[h, g] l+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] l &
+ G ⊢ K1 ▪⊑[h, g] K2 & I = Abst & L1 = K1. ⓓⓝW.V.
+#h #g #G #L1 #L2 * -L1 -L2
+[ #J #K2 #U #H destruct
+| #I #L1 #L2 #V #HL12 #J #K2 #U #H destruct /3 width=3/
+| #L1 #L2 #W #V #l #HV #HW #HL12 #J #K2 #U #H destruct /3 width=7/
+]
+qed-.
+
+lemma lsubd_inv_pair2: ∀h,g,I,G,L1,K2,W. G ⊢ L1 ▪⊑[h, g] K2.ⓑ{I}W →
+ (∃∃K1. G ⊢ K1 ▪⊑[h, g] K2 & L1 = K1.ⓑ{I}W) ∨
+ ∃∃K1,V,l. ⦃G, K1⦄ ⊢ V ▪[h, g] l+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] l &
+ G ⊢ K1 ▪⊑[h, g] K2 & I = Abst & L1 = K1. ⓓⓝW.V.
+/2 width=3 by lsubd_inv_pair2_aux/ qed-.
+
+(* Basic properties *********************************************************)
+
+lemma lsubd_refl: ∀h,g,G,L. G ⊢ L ▪⊑[h, g] L.
+#h #g #G #L elim L -L // /2 width=1/
+qed.
+
+(* Note: the constant 0 cannot be generalized *)
+lemma lsubd_ldrop_O1_conf: ∀h,g,G,L1,L2. G ⊢ L1 ▪⊑[h, g] L2 →
+ ∀K1,e. ⇩[0, e] L1 ≡ K1 →
+ ∃∃K2. G ⊢ K1 ▪⊑[h, g] K2 & ⇩[0, e] L2 ≡ K2.
+#h #g #G #L1 #L2 #H elim H -L1 -L2
+[ /2 width=3/
+| #I #L1 #L2 #V #_ #IHL12 #K1 #e #H
+ elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK1
+ [ destruct
+ elim (IHL12 L1 0) -IHL12 // #X #HL12 #H
+ <(ldrop_inv_O2 … H) in HL12; -H /3 width=3/
+ | elim (IHL12 … HLK1) -L1 /3 width=3/
+ ]
+| #L1 #L2 #W #V #l #HV #HW #_ #IHL12 #K1 #e #H
+ elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK1
+ [ destruct
+ elim (IHL12 L1 0) -IHL12 // #X #HL12 #H
+ <(ldrop_inv_O2 … H) in HL12; -H /3 width=3/
+ | elim (IHL12 … HLK1) -L1 /3 width=3/
+ ]
+]
+qed-.
+
+(* Note: the constant 0 cannot be generalized *)
+lemma lsubd_ldrop_O1_trans: ∀h,g,G,L1,L2. G ⊢ L1 ▪⊑[h, g] L2 →
+ ∀K2,e. ⇩[0, e] L2 ≡ K2 →
+ ∃∃K1. G ⊢ K1 ▪⊑[h, g] K2 & ⇩[0, e] L1 ≡ K1.
+#h #g #G #L1 #L2 #H elim H -L1 -L2
+[ /2 width=3/
+| #I #L1 #L2 #V #_ #IHL12 #K2 #e #H
+ elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK2
+ [ destruct
+ elim (IHL12 L2 0) -IHL12 // #X #HL12 #H
+ <(ldrop_inv_O2 … H) in HL12; -H /3 width=3/
+ | elim (IHL12 … HLK2) -L2 /3 width=3/
+ ]
+| #L1 #L2 #W #V #l #HV #HW #_ #IHL12 #K2 #e #H
+ elim (ldrop_inv_O1_pair1 … H) -H * #He #HLK2
+ [ destruct
+ elim (IHL12 L2 0) -IHL12 // #X #HL12 #H
+ <(ldrop_inv_O2 … H) in HL12; -H /3 width=3/
+ | elim (IHL12 … HLK2) -L2 /3 width=3/
+ ]
+]
+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_da.ma".
+include "basic_2/static/lsubd.ma".
+
+(* LOCAL ENVIRONMENT REFINEMENT FOR DEGREE ASSIGNMENT ***********************)
+
+(* Properties on degree assignment ******************************************)
+
+lemma lsubd_da_trans: ∀h,g,G,L2,T,l. ⦃G, L2⦄ ⊢ T ▪[h, g] l →
+ ∀L1. G ⊢ L1 ▪⊑[h, g] L2 → ⦃G, L1⦄ ⊢ T ▪[h, g] l.
+#h #g #G #L2 #T #l #H elim H -G -L2 -T -l
+[ /2 width=1/
+| #G #L2 #K2 #V #i #l #HLK2 #_ #IHV #L1 #HL12
+ elim (lsubd_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
+ elim (lsubd_inv_pair2 … H) -H * #K1 [ | -IHV -HLK1 ]
+ [ #HK12 #H destruct /3 width=4/
+ | #W #l0 #_ #_ #_ #H destruct
+ ]
+| #G #L2 #K2 #W #i #l #HLK2 #HW #IHW #L1 #HL12
+ elim (lsubd_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
+ elim (lsubd_inv_pair2 … H) -H * #K1 [ -HW | -IHW ]
+ [ #HK12 #H destruct /3 width=4/
+ | #V #l0 #HV #H0W #_ #_ #H destruct
+ lapply (da_mono … H0W … HW) -H0W -HW #H destruct /3 width=7/
+ ]
+| /4 width=1/
+| /3 width=1/
+]
+qed-.
+
+lemma lsubd_da_conf: ∀h,g,G,L1,T,l. ⦃G, L1⦄ ⊢ T ▪[h, g] l →
+ ∀L2. G ⊢ L1 ▪⊑[h, g] L2 → ⦃G, L2⦄ ⊢ T ▪[h, g] l.
+#h #g #G #L1 #T #l #H elim H -G -L1 -T -l
+[ /2 width=1/
+| #G #L1 #K1 #V #i #l #HLK1 #HV #IHV #L2 #HL12
+ elim (lsubd_ldrop_O1_conf … HL12 … HLK1) -L1 #X #H #HLK2
+ elim (lsubd_inv_pair1 … H) -H * #K2 [ -HV | -IHV ]
+ [ #HK12 #H destruct /3 width=4/
+ | #W0 #V0 #l0 #HV0 #HW0 #_ #_ #H1 #H2 destruct
+ lapply (da_inv_flat … HV) -HV #H0V0
+ lapply (da_mono … H0V0 … HV0) -H0V0 -HV0 #H destruct /2 width=4/
+ ]
+| #G #L1 #K1 #W #i #l #HLK1 #HW #IHW #L2 #HL12
+ elim (lsubd_ldrop_O1_conf … HL12 … HLK1) -L1 #X #H #HLK2
+ elim (lsubd_inv_pair1 … H) -H * #K2 [ -HW | -IHW ]
+ [ #HK12 #H destruct /3 width=4/
+ | #W0 #V0 #l0 #HV0 #HW0 #_ #H destruct
+ ]
+| /4 width=1/
+| /3 width=1/
+]
+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/lsubd_da.ma".
+
+(* LOCAL ENVIRONMENT REFINEMENT FOR DEGREE ASSIGNMENT ***********************)
+
+(* Main properties **********************************************************)
+
+theorem lsubd_trans: ∀h,g,G. Transitive … (lsubd h g G).
+#h #g #G #L1 #L #H elim H -L1 -L
+[ #X #H >(lsubd_inv_atom1 … H) -H //
+| #I #L1 #L #Y #HL1 #IHL1 #X #H
+ elim (lsubd_inv_pair1 … H) -H * #L2
+ [ #HL2 #H destruct /3 width=1/
+ | #W #V #l #HV #HW #HL2 #H1 #H2 #H3 destruct
+ /3 width=3 by lsubd_abbr, lsubd_da_trans/
+ ]
+| #L1 #L #W #V #l #HV #HW #HL1 #IHL1 #X #H
+ elim (lsubd_inv_pair1 … H) -H * #L2
+ [ #HL2 #H destruct /3 width=5 by lsubd_abbr, lsubd_da_conf/
+ | #W0 #V0 #l0 #_ #_ #_ #H destruct
+ ]
+]
+qed-.
| #K0 #l1 #l2 * [ #l01 ] #H1 * [1,3: #l02 ] #H2 //
[ < H2 in H1; -H2 #H
lapply (nexts_inj … H) -H #H destruct //
- | elim (H1 ?) /2 width=2/
- | elim (H2 ?) /2 width=2/
+ | elim H1 /2 width=2/
+ | elim H2 /2 width=2/
]
| #k0 #l0 *
[ #l #H destruct elim l -l normalize /2 width=1/
(* *)
(**************************************************************************)
-include "basic_2/notation/relations/statictype_7.ma".
-include "basic_2/grammar/genv.ma".
-include "basic_2/relocation/ldrop.ma".
-include "basic_2/static/sd.ma".
+include "basic_2/notation/relations/statictype_6.ma".
+include "basic_2/static/da.ma".
-(* STRATIFIED STATIC TYPE ASSIGNMENT ON TERMS *******************************)
+(* STRATIFIED STATIC TYPE ASSIGNMENT FOR TERMS ******************************)
(* activate genv *)
-inductive ssta (h:sh) (g:sd h): nat → relation4 genv lenv term term ≝
-| ssta_sort: ∀G,L,k,l. deg h g k l → ssta h g l G L (⋆k) (⋆(next h k))
-| ssta_ldef: ∀G,L,K,V,W,U,i,l. ⇩[0, i] L ≡ K. ⓓV → ssta h g l G K V W →
- ⇧[0, i + 1] W ≡ U → ssta h g l G L (#i) U
-| ssta_ldec: ∀G,L,K,W,V,U,i,l. ⇩[0, i] L ≡ K. ⓛW → ssta h g l G K W V →
- ⇧[0, i + 1] W ≡ U → ssta h g (l+1) G L (#i) U
-| ssta_bind: ∀a,I,G,L,V,T,U,l. ssta h g l G (L. ⓑ{I} V) T U →
- ssta h g l G L (ⓑ{a,I}V.T) (ⓑ{a,I}V.U)
-| ssta_appl: ∀G,L,V,T,U,l. ssta h g l G L T U →
- ssta h g l G L (ⓐV.T) (ⓐV.U)
-| ssta_cast: ∀G,L,W,T,U,l. ssta h g l G L T U → ssta h g l G L (ⓝW.T) U
+inductive ssta (h) (g): relation4 genv lenv term term ≝
+| ssta_sort: ∀G,L,k. ssta h g G L (⋆k) (⋆(next h k))
+| ssta_ldef: ∀G,L,K,V,U,W,i. ⇩[0, i] L ≡ K.ⓓV → ssta h g G K V W →
+ ⇧[0, i + 1] W ≡ U → ssta h g G L (#i) U
+| ssta_ldec: ∀G,L,K,W,U,l,i. ⇩[0, i] L ≡ K.ⓛW → ⦃G, K⦄ ⊢ W ▪[h, g] l →
+ ⇧[0, i + 1] W ≡ U → ssta h g G L (#i) U
+| ssta_bind: ∀a,I,G,L,V,T,U. ssta h g G (L.ⓑ{I}V) T U →
+ ssta h g G L (ⓑ{a,I}V.T) (ⓑ{a,I}V.U)
+| ssta_appl: ∀G,L,V,T,U. ssta h g G L T U → ssta h g G L (ⓐV.T) (ⓐV.U)
+| ssta_cast: ∀G,L,W,T,U. ssta h g G L T U → ssta h g G L (ⓝW.T) U
.
interpretation "stratified static type assignment (term)"
- 'StaticType h g G L T U l = (ssta h g l G L T U).
-
-definition ssta_step: ∀h. sd h → relation4 genv lenv term term ≝
- λh,g,G,L,T,U. ∃l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l+1, U⦄.
+ 'StaticType h g G L T U = (ssta h g G L T U).
(* Basic inversion lemmas ************************************************)
-fact ssta_inv_sort1_aux: ∀h,g,G,L,T,U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ → ∀k0. T = ⋆k0 →
- deg h g k0 l ∧ U = ⋆(next h k0).
-#h #g #G #L #T #U #l * -G -L -T -U -l
-[ #G #L #k #l #Hkl #k0 #H destruct /2 width=1/
-| #G #L #K #V #W #U #i #l #_ #_ #_ #k0 #H destruct
-| #G #L #K #W #V #U #i #l #_ #_ #_ #k0 #H destruct
-| #a #I #G #L #V #T #U #l #_ #k0 #H destruct
-| #G #L #V #T #U #l #_ #k0 #H destruct
-| #G #L #W #T #U #l #_ #k0 #H destruct
+fact ssta_inv_sort1_aux: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h, g] U → ∀k0. T = ⋆k0 →
+ U = ⋆(next h k0).
+#h #g #G #L #T #U * -G -L -T -U
+[ #G #L #k #k0 #H destruct //
+| #G #L #K #V #U #W #i #_ #_ #_ #k0 #H destruct
+| #G #L #K #W #U #l #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 just: sty0_gen_sort *)
-lemma ssta_inv_sort1: ∀h,g,G,L,U,k,l. ⦃G, L⦄ ⊢ ⋆k •[h, g] ⦃l, U⦄ →
- deg h g k l ∧ U = ⋆(next h k).
-/2 width=5 by ssta_inv_sort1_aux/ qed-.
+lemma ssta_inv_sort1: ∀h,g,G,L,U,k. ⦃G, L⦄ ⊢ ⋆k •[h, g] U → U = ⋆(next h k).
+/2 width=6 by ssta_inv_sort1_aux/ qed-.
-fact ssta_inv_lref1_aux: ∀h,g,G,L,T,U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ → ∀j. T = #j →
- (∃∃K,V,W. ⇩[0, j] L ≡ K. ⓓV & ⦃G, K⦄ ⊢ V •[h, g] ⦃l, W⦄ &
+fact ssta_inv_lref1_aux: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h, g] U → ∀j. T = #j →
+ (∃∃K,V,W. ⇩[0, j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •[h, g] W &
⇧[0, j + 1] W ≡ U
) ∨
- (∃∃K,W,V,l0. ⇩[0, j] L ≡ K. ⓛW & ⦃G, K⦄ ⊢ W •[h, g] ⦃l0, V⦄ &
- ⇧[0, j + 1] W ≡ U & l = l0 + 1
+ (∃∃K,W,l. ⇩[0, j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W ▪[h, g] l &
+ ⇧[0, j + 1] W ≡ U
).
-#h #g #G #L #T #U #l * -G -L -T -U -l
-[ #G #L #k #l #_ #j #H destruct
-| #G #L #K #V #W #U #i #l #HLK #HVW #HWU #j #H destruct /3 width=6/
-| #G #L #K #W #V #U #i #l #HLK #HWV #HWU #j #H destruct /3 width=8/
-| #a #I #G #L #V #T #U #l #_ #j #H destruct
-| #G #L #V #T #U #l #_ #j #H destruct
-| #G #L #W #T #U #l #_ #j #H destruct
+#h #g #G #L #T #U * -G -L -T -U
+[ #G #L #k #j #H destruct
+| #G #L #K #V #U #W #i #HLK #HVW #HWU #j #H destruct /3 width=6/
+| #G #L #K #W #U #l #i #HLK #HWl #HWU #j #H destruct /3 width=6/
+| #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 just: sty0_gen_lref *)
-lemma ssta_inv_lref1: ∀h,g,G,L,U,i,l. ⦃G, L⦄ ⊢ #i •[h, g] ⦃l, U⦄ →
- (∃∃K,V,W. ⇩[0, i] L ≡ K. ⓓV & ⦃G, K⦄ ⊢ V •[h, g] ⦃l, W⦄ &
+lemma ssta_inv_lref1: ∀h,g,G,L,U,i. ⦃G, L⦄ ⊢ #i •[h, g] U →
+ (∃∃K,V,W. ⇩[0, i] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •[h, g] W &
⇧[0, i + 1] W ≡ U
) ∨
- (∃∃K,W,V,l0. ⇩[0, i] L ≡ K. ⓛW & ⦃G, K⦄ ⊢ W •[h, g] ⦃l0, V⦄ &
- ⇧[0, i + 1] W ≡ U & l = l0 + 1
+ (∃∃K,W,l. ⇩[0, i] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W ▪[h, g] l &
+ ⇧[0, i + 1] W ≡ U
).
/2 width=3 by ssta_inv_lref1_aux/ qed-.
-fact ssta_inv_gref1_aux: ∀h,g,G,L,T,U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ → ∀p0. T = §p0 → ⊥.
-#h #g #G #L #T #U #l * -G -L -T -U -l
-[ #G #L #k #l #_ #p0 #H destruct
-| #G #L #K #V #W #U #i #l #_ #_ #_ #p0 #H destruct
-| #G #L #K #W #V #U #i #l #_ #_ #_ #p0 #H destruct
-| #a #I #G #L #V #T #U #l #_ #p0 #H destruct
-| #G #L #V #T #U #l #_ #p0 #H destruct
-| #G #L #W #T #U #l #_ #p0 #H destruct
+fact ssta_inv_gref1_aux: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h, g] U → ∀p0. T = §p0 → ⊥.
+#h #g #G #L #T #U * -G -L -T -U
+[ #G #L #k #p0 #H destruct
+| #G #L #K #V #U #W #i #_ #_ #_ #p0 #H destruct
+| #G #L #K #W #U #l #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 ssta_inv_gref1: ∀h,g,G,L,U,p,l. ⦃G, L⦄ ⊢ §p •[h, g] ⦃l, U⦄ → ⊥.
-/2 width=10 by ssta_inv_gref1_aux/ qed-.
-
-fact ssta_inv_bind1_aux: ∀h,g,G,L,T,U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ →
- ∀a,I,X,Y. T = ⓑ{a,I}Y.X →
- ∃∃Z. ⦃G, L.ⓑ{I}Y⦄ ⊢ X •[h, g] ⦃l, Z⦄ & U = ⓑ{a,I}Y.Z.
-#h #g #G #L #T #U #l * -G -L -T -U -l
-[ #G #L #k #l #_ #a #I #X #Y #H destruct
-| #G #L #K #V #W #U #i #l #_ #_ #_ #a #I #X #Y #H destruct
-| #G #L #K #W #V #U #i #l #_ #_ #_ #a #I #X #Y #H destruct
-| #b #J #G #L #V #T #U #l #HTU #a #I #X #Y #H destruct /2 width=3/
-| #G #L #V #T #U #l #_ #a #I #X #Y #H destruct
-| #G #L #W #T #U #l #_ #a #I #X #Y #H destruct
+lemma ssta_inv_gref1: ∀h,g,G,L,U,p. ⦃G, L⦄ ⊢ §p •[h, g] U → ⊥.
+/2 width=9 by ssta_inv_gref1_aux/ qed-.
+
+fact ssta_inv_bind1_aux: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h, g] U →
+ ∀b,J,X,Y. T = ⓑ{b,J}Y.X →
+ ∃∃Z. ⦃G, L.ⓑ{J}Y⦄ ⊢ X •[h, g] Z & U = ⓑ{b,J}Y.Z.
+#h #g #G #L #T #U * -G -L -T -U
+[ #G #L #k #b #J #X #Y #H destruct
+| #G #L #K #V #U #W #i #_ #_ #_ #b #J #X #Y #H destruct
+| #G #L #K #W #U #l #i #_ #_ #_ #b #J #X #Y #H destruct
+| #a #I #G #L #V #T #U #HTU #b #J #X #Y #H destruct /2 width=3/
+| #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 just: sty0_gen_bind *)
-lemma ssta_inv_bind1: ∀h,g,a,I,G,L,Y,X,U,l. ⦃G, L⦄ ⊢ ⓑ{a,I}Y.X •[h, g] ⦃l, U⦄ →
- ∃∃Z. ⦃G, L.ⓑ{I}Y⦄ ⊢ X •[h, g] ⦃l, Z⦄ & U = ⓑ{a,I}Y.Z.
+lemma ssta_inv_bind1: ∀h,g,b,J,G,L,Y,X,U. ⦃G, L⦄ ⊢ ⓑ{b,J}Y.X •[h, g] U →
+ ∃∃Z. ⦃G, L.ⓑ{J}Y⦄ ⊢ X •[h, g] Z & U = ⓑ{b,J}Y.Z.
/2 width=3 by ssta_inv_bind1_aux/ qed-.
-fact ssta_inv_appl1_aux: ∀h,g,G,L,T,U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ → ∀X,Y. T = ⓐY.X →
- ∃∃Z. ⦃G, L⦄ ⊢ X •[h, g] ⦃l, Z⦄ & U = ⓐY.Z.
-#h #g #G #L #T #U #l * -G -L -T -U -l
-[ #G #L #k #l #_ #X #Y #H destruct
-| #G #L #K #V #W #U #i #l #_ #_ #_ #X #Y #H destruct
-| #G #L #K #W #V #U #i #l #_ #_ #_ #X #Y #H destruct
-| #a #I #G #L #V #T #U #l #_ #X #Y #H destruct
-| #G #L #V #T #U #l #HTU #X #Y #H destruct /2 width=3/
-| #G #L #W #T #U #l #_ #X #Y #H destruct
+fact ssta_inv_appl1_aux: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h, g] U → ∀X,Y. T = ⓐY.X →
+ ∃∃Z. ⦃G, L⦄ ⊢ X •[h, g] Z & U = ⓐY.Z.
+#h #g #G #L #T #U * -G -L -T -U
+[ #G #L #k #X #Y #H destruct
+| #G #L #K #V #U #W #i #_ #_ #_ #X #Y #H destruct
+| #G #L #K #W #U #l #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/
+| #G #L #W #T #U #_ #X #Y #H destruct
]
qed-.
-(* Basic_1: was just: sty0_gen_appl *)
-lemma ssta_inv_appl1: ∀h,g,G,L,Y,X,U,l. ⦃G, L⦄ ⊢ ⓐY.X •[h, g] ⦃l, U⦄ →
- ∃∃Z. ⦃G, L⦄ ⊢ X •[h, g] ⦃l, Z⦄ & U = ⓐY.Z.
+lemma ssta_inv_appl1: ∀h,g,G,L,Y,X,U. ⦃G, L⦄ ⊢ ⓐY.X •[h, g] U →
+ ∃∃Z. ⦃G, L⦄ ⊢ X •[h, g] Z & U = ⓐY.Z.
/2 width=3 by ssta_inv_appl1_aux/ qed-.
-fact ssta_inv_cast1_aux: ∀h,g,G,L,T,U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ →
- â\88\80X,Y. T = â\93\9dY.X â\86\92 â¦\83G, Lâ¦\84 â\8a¢ X â\80¢[h, g] â¦\83l, Uâ¦\84.
-#h #g #G #L #T #U #l * -G -L -T -U -l
-[ #G #L #k #l #_ #X #Y #H destruct
-| #G #L #K #V #W #U #l #i #_ #_ #_ #X #Y #H destruct
-| #G #L #K #W #V #U #l #i #_ #_ #_ #X #Y #H destruct
-| #a #I #G #L #V #T #U #l #_ #X #Y #H destruct
-| #G #L #V #T #U #l #_ #X #Y #H destruct
-| #G #L #W #T #U #l #HTU #X #Y #H destruct //
+fact ssta_inv_cast1_aux: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h, g] U → ∀X,Y. T = ⓝY.X →
+ â¦\83G, Lâ¦\84 â\8a¢ X â\80¢[h, g] U.
+#h #g #G #L #T #U * -G -L -T -U
+[ #G #L #k #X #Y #H destruct
+| #G #L #K #V #U #W #i #_ #_ #_ #X #Y #H destruct
+| #G #L #K #W #U #l #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 just: sty0_gen_cast *)
-lemma ssta_inv_cast1: ∀h,g,G,L,X,Y,U,l. ⦃G, L⦄ ⊢ ⓝY.X •[h, g] ⦃l, U⦄ →
- ⦃G, L⦄ ⊢ X •[h, g] ⦃l, U⦄.
+lemma ssta_inv_cast1: ∀h,g,G,L,X,Y,U. ⦃G, L⦄ ⊢ ⓝY.X •[h, g] U → ⦃G, L⦄ ⊢ X •[h, g] U.
/2 width=4 by ssta_inv_cast1_aux/ qed-.
+
+(* Inversion lemmas on degree assignment for terms **************************)
+
+lemma ssta_inv_da: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h, g] 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/
+| #G #L #K #V #U #W #i #HLK #_ #_ * /3 width=5/
+| #G #L #K #W #U #l #i #HLK #HWl #_ /3 width=5/
+| #a #I #G #L #V #T #U #_ * /3 width=2/
+| #G #L #V #T #U #_ * /3 width=2/
+| #G #L #W #T #U #_ * /3 width=2/
+]
+qed-.
+
+(* Properties on degree assignment for terms ********************************)
+
+lemma da_ssta: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l →
+ ∃U. ⦃G, L⦄ ⊢ T •[h, g] 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/
+| #G #L #K #W #i #l #HLK #HW #_
+ elim (lift_total W 0 (i+1)) /3 width=7/
+| #a #I #G #L #V #T #l #_ * /3 width=2/
+| * #G #L #V #T #l #_ * /3 width=2/
+]
+qed-.
+
+(* Basic_1: removed theorems 7:
+ sty0_gen_sort sty0_gen_lref sty0_gen_bind sty0_gen_appl sty0_gen_cast
+ sty0_lift sty0_correct
+*)
+++ /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/aaa_lift.ma".
-include "basic_2/static/ssta.ma".
-
-(* STRATIFIED STATIC TYPE ASSIGNMENT ON TERMS *******************************)
-
-(* Properties on atomic arity assignment for terms **************************)
-
-lemma ssta_aaa: ∀h,g,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ → ⦃G, L⦄ ⊢ U ⁝ A.
-#h #g #G #L #T #A #H elim H -G -L -T -A
-[ #G #L #k #U #l #H
- elim (ssta_inv_sort1 … H) -H #_ #H destruct //
-| #I #G #L #K #V #B #i #HLK #HV #IHV #U #l #H
- elim (ssta_inv_lref1 … H) -H * #K0 #V0 #W0 [2: #l0 ] #HLK0 #HVW0 #HU [ #H ]
- lapply (ldrop_mono … HLK0 … HLK) -HLK0 #H0 destruct
- lapply (ldrop_fwd_ldrop2 … HLK) -HLK #HLK
- @(aaa_lift … HLK … HU) -HU -L // -HV /2 width=2/
-| #a #G #L #V #T #B #A #HV #_ #_ #IHT #X #l #H
- elim (ssta_inv_bind1 … H) -H #U #HTU #H destruct /3 width=2/
-| #a #G #L #V #T #B #A #HV #_ #_ #IHT #X #l #H
- elim (ssta_inv_bind1 … H) -H #U #HTU #H destruct /3 width=2/
-| #G #L #V #T #B #A #HV #_ #_ #IHT #X #l #H
- elim (ssta_inv_appl1 … H) -H #U #HTU #H destruct /3 width=3/
-| #G #L #V #T #A #_ #_ #IHV #IHT #X #l #H
- lapply (ssta_inv_cast1 … H) -H /2 width=2/
-]
-qed.
(* *)
(**************************************************************************)
-include "basic_2/relocation/ldrop_ldrop.ma".
+include "basic_2/static/da_lift.ma".
include "basic_2/static/ssta.ma".
-(* STRATIFIED STATIC TYPE ASSIGNMENT ON TERMS *******************************)
+(* STRATIFIED STATIC TYPE ASSIGNMENT FOR TERMS ******************************)
(* Properties on relocation *************************************************)
-(* Basic_1: was just: sty0_lift *)
-lemma ssta_lift: ∀h,g,G,L1,T1,U1,l. ⦃G, L1⦄ ⊢ T1 •[h, g] ⦃l, U1⦄ →
- ∀L2,d,e. ⇩[d, e] L2 ≡ L1 → ∀T2. ⇧[d, e] T1 ≡ T2 →
- ∀U2. ⇧[d, e] U1 ≡ U2 → ⦃G, L2⦄ ⊢ T2 •[h, g] ⦃l, U2⦄.
-#h #g #G #L1 #T1 #U1 #l #H elim H -G -L1 -T1 -U1 -l
-[ #G #L1 #k #l #Hkl #L2 #d #e #HL21 #X1 #H1 #X2 #H2
+lemma ssta_lift: ∀h,g,G. l_liftable (ssta h g G).
+#h #g #G #L1 #T1 #U1 #H elim H -G -L1 -T1 -U1
+[ #G #L1 #k #L2 #d #e #HL21 #X1 #H1 #X2 #H2
>(lift_inv_sort1 … H1) -X1
- >(lift_inv_sort1 … H2) -X2 /2 width=1/
-| #G #L1 #K1 #V1 #W1 #W #i #l #HLK1 #_ #HW1 #IHVW1 #L2 #d #e #HL21 #X #H #U2 #HWU2
+ >(lift_inv_sort1 … H2) -X2 //
+| #G #L1 #K1 #V1 #U1 #W1 #i #HLK1 #_ #HWU1 #IHVW1 #L2 #d #e #HL21 #X #H #U2 #HU12
elim (lift_inv_lref1 … H) * #Hid #H destruct
- [ elim (lift_trans_ge … HW1 … HWU2 ?) -W // #W2 #HW12 #HWU2
- elim (ldrop_trans_le … HL21 … HLK1 ?) -L1 /2 width=2/ #X #HLK2 #H
- elim (ldrop_inv_skip2 … H ?) -H /2 width=1/ -Hid #K2 #V2 #HK21 #HV12 #H destruct
+ [ elim (lift_trans_ge … HWU1 … HU12) -U1 // #W2 #HW12 #HWU2
+ elim (ldrop_trans_le … HL21 … HLK1) -L1 /2 width=2/ #X #HLK2 #H
+ elim (ldrop_inv_skip2 … H) -H /2 width=1/ -Hid #K2 #V2 #HK21 #HV12 #H destruct
/3 width=8/
- | lapply (lift_trans_be … HW1 … HWU2 ? ?) -W // /2 width=1/ #HW1U2
+ | lapply (lift_trans_be … HWU1 … HU12 ? ?) -U1 // /2 width=1/ #HW1U2
lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 // -Hid /3 width=8/
]
-| #G #L1 #K1 #W1 #V1 #W #i #l #HLK1 #_ #HW1 #IHWV1 #L2 #d #e #HL21 #X #H #U2 #HWU2
+| #G #L1 #K1 #W1 #U1 #l #i #HLK1 #HW1l #HWU1 #L2 #d #e #HL21 #X #H #U2 #HU12
elim (lift_inv_lref1 … H) * #Hid #H destruct
- [ elim (lift_trans_ge … HW1 … HWU2 ?) -W // <minus_plus #W #HW1 #HWU2
- elim (ldrop_trans_le … HL21 … HLK1 ?) -L1 /2 width=2/ #X #HLK2 #H
- elim (ldrop_inv_skip2 … H ?) -H /2 width=1/ -Hid #K2 #W2 #HK21 #HW12 #H destruct
+ [ elim (lift_trans_ge … HWU1 … HU12) -U1 // <minus_plus #W2 #HW12 #HWU2
+ elim (ldrop_trans_le … HL21 … HLK1) -L1 /2 width=2/ #X #HLK2 #H
+ elim (ldrop_inv_skip2 … H) -H /2 width=1/ -Hid #K2 #W #HK21 #HW1 #H destruct
lapply (lift_mono … HW1 … HW12) -HW1 #H destruct
- elim (lift_total V1 (d-i-1) e) /3 width=8/
- | lapply (lift_trans_be … HW1 … HWU2 ? ?) -W // /2 width=1/ #HW1U2
+ /3 width=10 by da_lift, ssta_ldec/
+ | lapply (lift_trans_be … HWU1 … HU12 ? ?) -U1 // /2 width=1/ #HW1U2
lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 // -Hid /3 width=8/
]
-| #a #I #G #L1 #V1 #T1 #U1 #l #_ #IHTU1 #L2 #d #e #HL21 #X1 #H1 #X2 #H2
+| #a #I #G #L1 #V1 #T1 #U1 #_ #IHTU1 #L2 #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=5/
-| #G #L1 #V1 #T1 #U1 #l #_ #IHTU1 #L2 #d #e #HL21 #X1 #H1 #X2 #H2
+| #G #L1 #V1 #T1 #U1 #_ #IHTU1 #L2 #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=5/
-| #G #L1 #W1 #T1 #U1 #l #_ #IHTU1 #L2 #d #e #HL21 #X #H #U2 #HU12
- elim (lift_inv_flat1 … H) -H #W2 #T2 #HW12 #HT12 #H destruct /3 width=5/
+| #G #L1 #W1 #T1 #U1 #_ #IHTU1 #L2 #d #e #HL21 #X #H #U2 #HU12
+ elim (lift_inv_flat1 … H) -H #W2 #T2 #_ #HT12 #H destruct /3 width=5/
]
qed.
-(* Note: apparently this was missing in basic_1 *)
-lemma ssta_inv_lift1: ∀h,g,G,L2,T2,U2,l. ⦃G, L2⦄ ⊢ T2 •[h, g] ⦃l, U2⦄ →
- ∀L1,d,e. ⇩[d, e] L2 ≡ L1 → ∀T1. ⇧[d, e] T1 ≡ T2 →
- ∃∃U1. ⦃G, L1⦄ ⊢ T1 •[h, g] ⦃l, U1⦄ & ⇧[d, e] U1 ≡ U2.
-#h #g #G #L2 #T2 #U2 #l #H elim H -G -L2 -T2 -U2 -l
-[ #G #L2 #k #l #Hkl #L1 #d #e #_ #X #H
- >(lift_inv_sort2 … H) -X /3 width=3/
-| #G #L2 #K2 #V2 #W2 #W #i #l #HLK2 #HVW2 #HW2 #IHVW2 #L1 #d #e #HL21 #X #H
+(* Inversion lemmas on relocation *******************************************)
+
+lemma ssta_inv_lift1: ∀h,g,G. l_deliftable_sn (ssta h g G).
+#h #g #G #L2 #T2 #U2 #H elim H -G -L2 -T2 -U2
+[ #G #L2 #k #L1 #d #e #_ #X #H
+ >(lift_inv_sort2 … H) -X /2 width=3/
+| #G #L2 #K2 #V2 #U2 #W2 #i #HLK2 #HVW2 #HWU2 #IHVW2 #L1 #d #e #HL21 #X #H
elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HVW2 | -IHVW2 ]
- [ elim (ldrop_conf_lt … HL21 … HLK2 ?) -L2 // #K1 #V1 #HLK1 #HK21 #HV12
- elim (IHVW2 … HK21 … HV12) -K2 -V2 #W1 #HVW1 #HW12
- elim (lift_trans_le … HW12 … HW2 ?) -W2 // >minus_plus <plus_minus_m_m // -Hid /3 width=6/
+ [ elim (ldrop_conf_lt … HL21 … HLK2) -L2 // #K1 #V1 #HLK1 #HK21 #HV12
+ elim (IHVW2 … HK21 … HV12) -K2 -V2 #W1 #HW12 #HVW1
+ elim (lift_trans_le … HW12 … HWU2) -W2 // >minus_plus <plus_minus_m_m // -Hid /3 width=8/
| lapply (ldrop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
elim (le_inv_plus_l … Hid) -Hid #Hdie #ei
- elim (lift_split … HW2 d (i-e+1) ? ? ?) -HW2 // [3: /2 width=1/ ]
- [ #W0 #HW20 <le_plus_minus_comm // >minus_minus_m_m /2 width=1/ /3 width=6/
+ elim (lift_split … HWU2 d (i-e+1)) -HWU2 // [3: /2 width=1/ ]
+ [ #W0 #HW20 <le_plus_minus_comm // >minus_minus_m_m /2 width=1/ /3 width=8/
| <le_plus_minus_comm //
]
]
-| #G #L2 #K2 #W2 #V2 #W #i #l #HLK2 #HWV2 #HW2 #IHWV2 #L1 #d #e #HL21 #X #H
- elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HWV2 | -IHWV2 ]
- [ elim (ldrop_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 // -Hid /3 width=6/
+| #G #L2 #K2 #W2 #U2 #l #i #HLK2 #HW2l #HWU2 #L1 #d #e #HL21 #X #H
+ elim (lift_inv_lref2 … H) * #Hid #H destruct
+ [ elim (ldrop_conf_lt … HL21 … HLK2) -L2 // #K1 #W1 #HLK1 #HK21 #HW12
+ lapply (da_inv_lift … HW2l … HK21 … HW12) -K2
+ elim (lift_trans_le … HW12 … HWU2) -W2 // >minus_plus <plus_minus_m_m // -Hid /3 width=8/
| lapply (ldrop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
elim (le_inv_plus_l … Hid) -Hid #Hdie #ei
- elim (lift_split … HW2 d (i-e+1) ? ? ?) -HW2 // [3: /2 width=1/ ]
- [ #W0 #HW20 <le_plus_minus_comm // >minus_minus_m_m /2 width=1/ /3 width=6/
+ elim (lift_split … HWU2 d (i-e+1)) -HWU2 // [3: /2 width=1/ ]
+ [ #W0 #HW20 <le_plus_minus_comm // >minus_minus_m_m /2 width=1/ /3 width=8/
| <le_plus_minus_comm //
]
]
-| #a #I #G #L2 #V2 #T2 #U2 #l #_ #IHTU2 #L1 #d #e #HL21 #X #H
+| #a #I #G #L2 #V2 #T2 #U2 #_ #IHTU2 #L1 #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 /2 width=1/ -HL21 /3 width=5/
-| #G #L2 #V2 #T2 #U2 #l #_ #IHTU2 #L1 #d #e #HL21 #X #H
+| #G #L2 #V2 #T2 #U2 #_ #IHTU2 #L1 #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/
-| #G #L2 #W2 #T2 #U2 #l #_ #IHTU2 #L1 #d #e #HL21 #X #H
- elim (lift_inv_flat2 … H) -H #W1 #T1 #HW12 #HT12 #H destruct
+| #G #L2 #W2 #T2 #U2 #_ #IHTU2 #L1 #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/
]
qed-.
+(* Advanced properties ******************************************************)
+
+lemma ssta_da_conf: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h, g] 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/
+| #G #L #K #V #U #W #i #HLK #_ #HWU #IHVW #l #H
+ elim (da_inv_lref … H) -H * #K0 #V0 [| #l0] #HLK0 #HV0
+ lapply (ldrop_mono … HLK0 … HLK) -HLK0 #H destruct
+ lapply (ldrop_fwd_ldrop2 … HLK) -HLK /3 width=7/
+| #G #L #K #W #U #l0 #i #HLK #_ #HWU #l #H -l0
+ elim (da_inv_lref … H) -H * #K0 #V0 [| #l1] #HLK0 #HV0 [| #H0 ]
+ lapply (ldrop_mono … HLK0 … HLK) -HLK0 #H destruct
+ lapply (ldrop_fwd_ldrop2 … HLK) -HLK /3 width=7/
+| #a #I #G #L #V #T #U #_ #IHTU #l #H
+ lapply (da_inv_bind … H) -H /3 width=1/
+| #G #L #V #T #U #_ #IHTU #l #H
+ lapply (da_inv_flat … H) -H /3 width=1/
+| #G #L #W #T #U #_ #IHTU #l #H
+ lapply (da_inv_flat … H) -H /2 width=1/
+]
+qed-.
+
(* Advanced forvard lemmas **************************************************)
-(* Basic_1: was just: sty0_correct *)
-lemma ssta_fwd_correct: ∀h,g,G,L,T,U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ →
- ∃T0. ⦃G, L⦄ ⊢ U •[h, g] ⦃l-1, T0⦄.
-#h #g #G #L #T #U #l #H elim H -G -L -T -U -l
-[ /4 width=2/
-| #G #L #K #V #W #W0 #i #l #HLK #_ #HW0 * #V0 #HWV0
+lemma ssta_fwd_correct: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h, g] U →
+ ∃T0. ⦃G, L⦄ ⊢ U •[h, g] T0.
+#h #g #G #L #T #U #H elim H -G -L -T -U
+[ /2 width=2/
+| #G #L #K #V #U #W #i #HLK #_ #HWU * #T #HWT
lapply (ldrop_fwd_ldrop2 … HLK) -HLK #HLK
- elim (lift_total V0 0 (i+1)) /3 width=10/
-| #G #L #K #W #V #V0 #i #l #HLK #HWV #HWV0 #_
+ elim (lift_total T 0 (i+1)) /3 width=10/
+| #G #L #K #W #U #l #i #HLK #HWl #HWU
+ elim (da_ssta … HWl) -HWl #T #HWT
lapply (ldrop_fwd_ldrop2 … HLK) -HLK #HLK
- elim (lift_total V 0 (i+1)) /3 width=10/
-| #a #I #G #L #V #T #U #l #_ * /3 width=2/
-| #G #L #V #T #U #l #_ * #T0 #HUT0 /3 width=2/
-| #G #L #W #T #U #l #_ * /2 width=2/
+ elim (lift_total T 0 (i+1)) /3 width=10/
+| #a #I #G #L #V #T #U #_ * /3 width=2/
+| #G #L #V #T #U #_ * #T0 #HUT0 /3 width=2/
+| #G #L #W #T #U #_ * /2 width=2/
]
qed-.
(* *)
(**************************************************************************)
+include "basic_2/static/da_da.ma".
include "basic_2/static/ssta_lift.ma".
(* STRATIFIED STATIC TYPE ASSIGNMENT ON TERMS *******************************)
+(* Advanced inversion lemmas ************************************************)
+
+lemma ssta_inv_refl_pos: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l+1 → ⦃G, L⦄ ⊢ T •[h, g] T → ⊥.
+#h #g #G #L #T #l #H1T #HTT
+lapply (ssta_da_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-.
+
(* Main properties **********************************************************)
-(* Note: apparently this was missing in basic_1 *)
-theorem ssta_mono: ∀h,g,G,L,T,U1,l1. ⦃G, L⦄ ⊢ T •[h, g] ⦃l1, U1⦄ →
- ∀U2,l2. ⦃G, L⦄ ⊢ T •[h, g] ⦃l2, U2⦄ → l1 = l2 ∧ U1 = U2.
-#h #g #G #L #T #U1 #l1 #H elim H -G -L -T -U1 -l1
-[ #G #L #k #l #Hkl #X #l2 #H
- elim (ssta_inv_sort1 … H) -H #Hkl2 #H destruct
- >(deg_mono … Hkl2 … Hkl) -g -L -l2 /2 width=1/
-| #G #L #K #V #W #U1 #i #l1 #HLK #_ #HWU1 #IHVW #U2 #l2 #H
- elim (ssta_inv_lref1 … H) -H * #K0 #V0 #W0 [2: #l0] #HLK0 #HVW0 #HW0U2
+theorem ssta_mono: ∀h,g,G,L. singlevalued … (ssta h g G L).
+#h #g #G #L #T #U1 #H elim H -G -L -T -U1
+[ #G #L #k #X #H >(ssta_inv_sort1 … H) -X //
+| #G #L #K #V #U1 #W #i #HLK #_ #HWU1 #IHVW #U2 #H
+ elim (ssta_inv_lref1 … H) -H * #K0 #V0 #W0 #HLK0 #HVW0 #HW0U2
lapply (ldrop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct
- lapply (IHVW … HVW0) -IHVW -HVW0 * #H1 #H2 destruct
- >(lift_mono … HWU1 … HW0U2) -W0 -U1 /2 width=1/
-| #G #L #K #W #V #U1 #i #l1 #HLK #_ #HWU1 #IHWV #U2 #l2 #H
- elim (ssta_inv_lref1 … H) -H * #K0 #W0 #V0 [2: #l0 ] #HLK0 #HWV0 #HV0U2
+ lapply (IHVW … HVW0) -IHVW -HVW0 #H destruct
+ >(lift_mono … HWU1 … HW0U2) -W0 -U1 //
+| #G #L #K #W #U1 #l #i #HLK #HWl #HWU1 #U2 #H
+ elim (ssta_inv_lref1 … H) -H * #K0 #W0 #l0 #HLK0 #HWl0 #HW0U2
lapply (ldrop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct
- lapply (IHWV … HWV0) -IHWV -HWV0 * #H1 #H2 destruct
- >(lift_mono … HWU1 … HV0U2) -W -U1 /2 width=1/
-| #a #I #G #L #V #T #U1 #l1 #_ #IHTU1 #X #l2 #H
- elim (ssta_inv_bind1 … H) -H #U2 #HTU2 #H destruct
- elim (IHTU1 … HTU2) -T /3 width=1/
-| #G #L #V #T #U1 #l1 #_ #IHTU1 #X #l2 #H
- elim (ssta_inv_appl1 … H) -H #U2 #HTU2 #H destruct
- elim (IHTU1 … HTU2) -T /3 width=1/
-| #G #L #W1 #T #U1 #l1 #_ #IHTU1 #U2 #l2 #H
- lapply (ssta_inv_cast1 … H) -H #HTU2
- elim (IHTU1 … HTU2) -T /2 width=1/
+ lapply (da_mono … HWl0 … HWl) -HWl0 #H destruct
+ >(lift_mono … HWU1 … HW0U2) -W -U1 //
+| #a #I #G #L #V #T #U1 #_ #IHTU1 #X #H
+ elim (ssta_inv_bind1 … H) -H #U2 #HTU2 #H destruct /3 width=1/
+| #G #L #V #T #U1 #_ #IHTU1 #X #H
+ elim (ssta_inv_appl1 … H) -H #U2 #HTU2 #H destruct /3 width=1/
+| #G #L #W #T #U1 #_ #IHTU1 #U2 #H
+ lapply (ssta_inv_cast1 … H) -H /2 width=1/
]
qed-.
-
-(* Advanced inversion lemmas ************************************************)
-
-lemma ssta_inv_refl_pos: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l+1, T⦄ → ⊥.
-#h #g #G #L #T #l #HTT
-elim (ssta_fwd_correct … HTT) <minus_plus_m_m #U #HTU
-elim (ssta_mono … HTU … HTT) -h -L #H #_ -T -U
-elim (plus_xySz_x_false 0 l 0 ?) //
-qed-.
>(ldrops_inv_nil … H) -L1 /2 width=7/
| #des #des2 #d #e #i #Hid #_ #IHdes2 #L1 #K2 #V2 #H
elim (ldrops_inv_cons … H) -H #L #HL1 #H
- elim (ldrop_inv_skip2 … H ?) -H /2 width=1/ #K #V >minus_plus #HK2 #HV2 #H destruct
+ elim (ldrop_inv_skip2 … H) -H /2 width=1/ #K #V >minus_plus #HK2 #HV2 #H destruct
elim (IHdes2 … HL1) -IHdes2 -HL1 #K1 #V1 #des1 #Hdes1 #HK1 #HV1 #X destruct
@(ex4_3_intro … K1 V1 … ) // [3,4: /2 width=7/ | skip ]
normalize >plus_minus // @minuss_lt // /2 width=1/ (**) (* explicit constructors, /3 width=1/ is a bit slow *)
[ /2 width=7/
| #L1 #L3 #L #des3 #d #e #_ #HL3 #IHL13 #L2 #i #HL2
elim (lt_or_ge i d) #Hid
- [ elim (ldrop_trans_le … HL3 … HL2 ?) -L /2 width=2/ #L #HL3 #HL2
+ [ elim (ldrop_trans_le … HL3 … HL2) -L /2 width=2/ #L #HL3 #HL2
elim (IHL13 … HL3) -L3 /3 width=7/
| lapply (ldrop_trans_ge … HL3 … HL2 ?) -L // #HL32
elim (IHL13 … HL32) -L3 /3 width=7/
[ /2 width=3/
| #T1 #T3 #T #des #d #e #HT13 #_ #IHT13 #T2 #HT2
elim (IHT13 … HT2) -T #T #HT3 #HT2
- elim (lift_trans_le … HT13 … HT3 ?) -T3 // /3 width=5/
+ elim (lift_trans_le … HT13 … HT3) -T3 // /3 width=5/
]
qed-.
>(lifts_inv_nil … H) -T1 /2 width=3/
| #d #e #des #IHdes #i #i0 #H1 #des0 #H2 #T1 #T0 #HT10 #T2 #HT02
elim (at_inv_cons … H1) -H1 * #Hid #Hi0
- [ elim (minuss_inv_cons1_lt … H2 ?) -H2 [2: /2 width=1/ ] #des1 #Hdes1 <minus_le_minus_minus_comm // <minus_plus_m_m #H
+ [ elim (minuss_inv_cons1_lt … H2) -H2 [2: /2 width=1/ ] #des1 #Hdes1 <minus_le_minus_minus_comm // <minus_plus_m_m #H
elim (pluss_inv_cons2 … H) -H #des2 #H1 #H2 destruct
elim (lifts_inv_cons … HT10) -HT10 #T >minus_plus #HT1 #HT0
elim (IHdes … Hi0 … Hdes1 … HT0 … HT02) -IHdes -Hi0 -Hdes1 -T0 #T0 #HT0 #HT02
- elim (lift_trans_le … HT1 … HT0 ?) -T /2 width=1/ #T #HT1 <plus_minus_m_m /2 width=1/ /3 width=5/
+ elim (lift_trans_le … HT1 … HT0) -T /2 width=1/ #T #HT1 <plus_minus_m_m /2 width=1/ /3 width=5/
| >commutative_plus in Hi0; #Hi0
lapply (minuss_inv_cons1_ge … H2 ?) -H2 [ /2 width=1/ ] <associative_plus #Hdes0
elim (IHdes … Hi0 … Hdes0 … HT10 … HT02) -IHdes -Hi0 -Hdes0 -T0 #T0 #HT0 #HT02
- elim (lift_split … HT0 d (i+1) ? ? ?) -HT0 /2 width=1/ /3 width=5/
+ elim (lift_split … HT0 d (i+1)) -HT0 /2 width=1/ /3 width=5/
]
]
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_7.ma".
+include "basic_2/static/ssta.ma".
+
+(* NAT-ITERATED STRATIFIED STATIC TYPE ASSIGNMENT FOR TERMS *****************)
+
+definition lsstas: ∀h. sd h → genv → lenv → nat → relation term ≝
+ λh,g,G,L. lstar … (ssta h g G L).
+
+interpretation "nat-iterated stratified static type assignment (term)"
+ 'StaticTypeStar h g G L l T U = (lsstas h g G L l T U).
+
+(* Basic eliminators ********************************************************)
+
+lemma lsstas_ind_sn: ∀h,g,G,L,U2. ∀R:relation2 nat term.
+ R 0 U2 → (
+ ∀l,T,U1. ⦃G, L⦄ ⊢ T •[h, g] U1 → ⦃G, L⦄ ⊢ U1 •* [h, g, l] U2 →
+ R l U1 → R (l+1) T
+ ) →
+ ∀l,T. ⦃G, L⦄ ⊢ T •*[h, g, l] U2 → R l T.
+/3 width=5 by lstar_ind_l/ qed-.
+
+lemma lsstas_ind_dx: ∀h,g,G,L,T. ∀R:relation2 nat term.
+ R 0 T → (
+ ∀l,U1,U2. ⦃G, L⦄ ⊢ T •* [h, g, l] U1 → ⦃G, L⦄ ⊢ U1 •[h, g] U2 →
+ R l U1 → R (l+1) U2
+ ) →
+ ∀l,U. ⦃G, L⦄ ⊢ T •*[h, g, l] U → R l U.
+/3 width=5 by lstar_ind_r/ qed-.
+
+(* Basic inversion lemmas ***************************************************)
+
+lemma lsstas_inv_O: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •*[h, g, 0] U → T = U.
+/2 width=4 by lstar_inv_O/ qed-.
+
+lemma lsstas_inv_SO: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •*[h, g, 1] U → ⦃G, L⦄ ⊢ T •[h, g] U.
+/2 width=1 by lstar_inv_step/ qed-.
+
+lemma lsstas_inv_step_sn: ∀h,g,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 •*[h, g, l+1] T2 →
+ ∃∃T. ⦃G, L⦄ ⊢ T1 •[h, g] T & ⦃G, L⦄ ⊢ T •*[h, g, l] T2.
+/2 width=3 by lstar_inv_S/ qed-.
+
+lemma lsstas_inv_step_dx: ∀h,g,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 •*[h, g, l+1] T2 →
+ ∃∃T. ⦃G, L⦄ ⊢ T1 •*[h, g, l] T & ⦃G, L⦄ ⊢ T •[h, g] T2.
+/2 width=3 by lstar_inv_S_dx/ qed-.
+
+lemma lsstas_inv_sort1: ∀h,g,G,L,X,k,l. ⦃G, L⦄ ⊢ ⋆k •*[h, g, l] X → X = ⋆((next h)^l k).
+#h #g #G #L #X #k #l #H @(lsstas_ind_dx … H) -X -l //
+#l #X #X0 #_ #H #IHX destruct
+lapply (ssta_inv_sort1 … H) -H #H destruct
+>iter_SO //
+qed-.
+
+lemma lsstas_inv_gref1: ∀h,g,G,L,X,p,l. ⦃G, L⦄ ⊢ §p •*[h, g, l+1] X → ⊥.
+#h #g #G #L #X #p #l #H elim (lsstas_inv_step_sn … H) -H
+#U #H #HUX elim (ssta_inv_gref1 … H)
+qed-.
+
+lemma lsstas_inv_bind1: ∀h,g,a,I,G,L,V,T,X,l. ⦃G, L⦄ ⊢ ⓑ{a,I}V.T •*[h, g, l] X →
+ ∃∃U. ⦃G, L.ⓑ{I}V⦄ ⊢ T •*[h, g, l] U & X = ⓑ{a,I}V.U.
+#h #g #a #I #G #L #V #T #X #l #H @(lsstas_ind_dx … H) -X -l [ /2 width=3/ ]
+#l #X #X0 #_ #HX0 * #U #HTU #H destruct
+elim (ssta_inv_bind1 … HX0) -HX0 #U0 #HU0 #H destruct /3 width=3/
+qed-.
+
+lemma lsstas_inv_appl1: ∀h,g,G,L,V,T,X,l. ⦃G, L⦄ ⊢ ⓐV.T •*[h, g, l] X →
+ ∃∃U. ⦃G, L⦄ ⊢ T •*[h, g, l] U & X = ⓐV.U.
+#h #g #G #L #V #T #X #l #H @(lsstas_ind_dx … H) -X -l [ /2 width=3/ ]
+#l #X #X0 #_ #HX0 * #U #HTU #H destruct
+elim (ssta_inv_appl1 … HX0) -HX0 #U0 #HU0 #H destruct /3 width=3/
+qed-.
+
+lemma lsstas_inv_cast1: ∀h,g,G,L,W,T,U,l. ⦃G, L⦄ ⊢ ⓝW.T •*[h, g, l+1] U → ⦃G, L⦄ ⊢ T •*[h, g, l+1] U.
+#h #g #G #L #W #T #X #l #H elim (lsstas_inv_step_sn … H) -H
+#U #H #HUX lapply (ssta_inv_cast1 … H) -H /2 width=3/
+qed-.
+
+(* Basic properties *********************************************************)
+
+lemma lsstas_refl: ∀h,g,G,L. reflexive … (lsstas h g G L 0).
+// qed.
+
+lemma ssta_lsstas: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h, g] U → ⦃G, L⦄ ⊢ T •*[h, g, 1] U.
+/2 width=1/ qed.
+
+lemma lsstas_step_sn: ∀h,g,G,L,T1,U1,U2,l. ⦃G, L⦄ ⊢ T1 •[h, g] U1 → ⦃G, L⦄ ⊢ U1 •*[h, g, l] U2 →
+ ⦃G, L⦄ ⊢ T1 •*[h, g, l+1] U2.
+/2 width=3/ qed.
+
+lemma lsstas_step_dx: ∀h,g,G,L,T1,T2,U2,l. ⦃G, L⦄ ⊢ T1 •*[h, g, l] T2 → ⦃G, L⦄ ⊢ T2 •[h, g] U2 →
+ ⦃G, L⦄ ⊢ T1 •*[h, g, l+1] U2.
+/2 width=3/ qed.
+
+lemma lsstas_split: ∀h,g,G,L. inv_ltransitive … (lsstas h g G L).
+/2 width=1 by lstar_inv_ltransitive/ qed-.
+
+lemma lsstas_sort: ∀h,g,G,L,l,k. ⦃G, L⦄ ⊢ ⋆k •*[h, g, l] ⋆((next h)^l k).
+#h #g #G #L #l @(nat_ind_plus … l) -l //
+#l #IHl #k >iter_SO /2 width=3/
+qed.
+
+lemma lsstas_bind: ∀h,g,I,G,L,V,T,U,l. ⦃G, L.ⓑ{I}V⦄ ⊢ T •*[h, g, l] U →
+ ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V.T •*[h, g, l] ⓑ{a,I}V.U.
+#h #g #I #G #L #V #T #U #l #H @(lsstas_ind_dx … H) -U -l // /3 width=3/
+qed.
+
+lemma lsstas_appl: ∀h,g,G,L,T,U,l. ⦃G, L⦄ ⊢ T •*[h, g, l] U →
+ ∀V.⦃G, L⦄ ⊢ ⓐV.T •*[h, g, l] ⓐV.U.
+#h #g #G #L #T #U #l #H @(lsstas_ind_dx … H) -U -l // /3 width=3/
+qed.
+
+lemma lsstas_cast: ∀h,g,G,L,T,U,l. ⦃G, L⦄ ⊢ T •*[h, g, l+1] U →
+ ∀W. ⦃G, L⦄ ⊢ ⓝW.T •*[h, g, l+1] U.
+#h #g #G #L #T #U #l #H elim (lsstas_inv_step_sn … H) -H /3 width=3/
+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/aaa_ssta.ma".
+include "basic_2/unfold/lsstas.ma".
+
+(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
+
+(* Properties on atomic arity assignment for terms **************************)
+
+lemma aaa_lsstas_conf: ∀h,g,G,L,l. Conf3 … (aaa G L) (lsstas h g G L l).
+/3 width=6 by aaa_ssta_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_7.ma".
+include "basic_2/unfold/lsstas_lift.ma".
+
+(* NAT-ITERATED STRATIFIED STATIC TYPE ASSIGNMENT FOR TERMS *****************)
+
+(* alternative definition of lsstas *)
+inductive lsstasa (h) (g): genv → relation4 lenv nat term term ≝
+| lsstasa_O : ∀G,L,T. lsstasa h g G L 0 T T
+| lsstasa_sort: ∀G,L,l,k. lsstasa h g G L l (⋆k) (⋆((next h)^l k))
+| lsstasa_ldef: ∀G,L,K,V,W,U,i,l. ⇩[0, i] L ≡ K.ⓓV → lsstasa h g G K (l+1) V W →
+ ⇧[0, i+1] W ≡ U → lsstasa h g G L (l+1) (#i) U
+| lsstasa_ldec: ∀G,L,K,W,V,U,i,l,l0. ⇩[0, i] L ≡ K.ⓛW → ⦃G, K⦄ ⊢ W ▪[h, g] l0 →
+ lsstasa h g G K l W V → ⇧[0, i+1] V ≡ U → lsstasa h g G L (l+1) (#i) U
+| lsstasa_bind: ∀a,I,G,L,V,T,U,l. lsstasa h g G (L.ⓑ{I}V) l T U →
+ lsstasa h g G L l (ⓑ{a,I}V.T) (ⓑ{a,I}V.U)
+| lsstasa_appl: ∀G,L,V,T,U,l. lsstasa h g G L l T U → lsstasa h g G L l (ⓐV.T) (ⓐV.U)
+| lsstasa_cast: ∀G,L,W,T,U,l. lsstasa h g G L (l+1) T U → lsstasa h g G L (l+1) (ⓝW.T) U
+.
+
+interpretation "nat-iterated stratified static type assignment (term) alternative"
+ 'StaticTypeStarAlt h g G L l T U = (lsstasa h g G L l T U).
+
+(* Base properties **********************************************************)
+
+lemma ssta_lsstasa: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h, g] U → ⦃G, L⦄ ⊢ T ••*[h, g, 1] U.
+#h #g #G #L #T #U #H elim H -G -L -T -U
+// /2 width=1/ /2 width=6/ /2 width=8/
+qed.
+
+lemma lsstasa_step_dx: ∀h,g,G,L,T1,T,l. ⦃G, L⦄ ⊢ T1 ••*[h, g, l] T →
+ ∀T2. ⦃G, L⦄ ⊢ T •[h, g] T2 → ⦃G, L⦄ ⊢ T1 ••*[h, g, l+1] T2.
+#h #g #G #L #T1 #T #l #H elim H -G -L -T1 -T -l
+[ /2 width=1/
+| #G #L #l #k #X #H >(ssta_inv_sort1 … H) -X >commutative_plus //
+| #G #L #K #V #W #U #i #l #HLK #_ #HWU #IHVW #U2 #HU2
+ lapply (ldrop_fwd_ldrop2 … HLK) #H
+ elim (ssta_inv_lift1 … HU2 … H … HWU) -H -U /3 width=6/
+| #G #L #K #W #V #U #i #l #l0 #HLK #HWl0 #_ #HVU #IHWV #U2 #HU2
+ lapply (ldrop_fwd_ldrop2 … HLK) #H
+ elim (ssta_inv_lift1 … HU2 … H … HVU) -H -U /3 width=8/
+| #a #I #G #L #V #T1 #U1 #l #_ #IHTU1 #X #H
+ elim (ssta_inv_bind1 … H) -H #U #HU1 #H destruct /3 width=1/
+| #G #L #V #T1 #U1 #l #_ #IHTU1 #X #H
+ elim (ssta_inv_appl1 … H) -H #U #HU1 #H destruct /3 width=1/
+| /3 width=1/
+]
+qed.
+
+(* Main properties **********************************************************)
+
+theorem lsstas_lsstasa: ∀h,g,G,L,T,U,l. ⦃G, L⦄ ⊢ T •*[h, g, l] U → ⦃G, L⦄ ⊢ T ••*[h, g, l] U.
+#h #g #G #L #T #U #l #H @(lsstas_ind_dx … H) -U -l // /2 width=3/
+qed.
+
+(* Main inversion lemmas ****************************************************)
+
+theorem lsstasa_inv_lsstas: ∀h,g,G,L,T,U,l. ⦃G, L⦄ ⊢ T ••*[h, g, l] U → ⦃G, L⦄ ⊢ T •*[h, g, l] U.
+#h #g #G #L #T #U #l #H elim H -G -L -T -U -l
+// /2 width=1/ /2 width=6/ /3 width=8 by lsstas_ldec, lsstas_inv_SO/
+qed-.
+
+(* Advanced eliminators *****************************************************)
+
+lemma lsstas_ind_alt: ∀h,g. ∀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.
+ ⇩[O, i] L ≡ K.ⓓV → ⦃G, K⦄ ⊢ V •*[h, g, 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,U,i,l,l0.
+ ⇩[O, i] L ≡ K.ⓛW → ⦃G, K⦄ ⊢ W ▪[h, g] l0 →
+ ⦃G, K⦄ ⊢ W •*[h, g, 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, g, 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, g, 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, g, 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, g, l] U → R G L l T U.
+#h #g #R #IH1 #IH2 #IH3 #IH4 #IH5 #IH6 #IH7 #G #L #l #T #U #H
+elim (lsstas_lsstasa … H) /3 width=10 by lsstasa_inv_lsstas/
+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/ssta_lift.ma".
+include "basic_2/unfold/lsstas.ma".
+
+(* NAT-ITERATED STRATIFIED STATIC TYPE ASSIGNMENT FOR TERMS *****************)
+
+(* Properties on relocation *************************************************)
+
+lemma lsstas_lift: ∀h,g,G,l. l_liftable (llstar … (ssta h g G) l).
+/2 width=1/ qed.
+
+(* Inversion lemmas on relocation *******************************************)
+
+lemma lsstas_inv_lift1: ∀h,g,G,l. l_deliftable_sn (llstar … (ssta h g G) l).
+/3 width=5 by l_deliftable_sn_llstar, ssta_inv_lift1/ qed-.
+
+(* Advanced inversion lemmas ************************************************)
+
+lemma lsstas_inv_lref1: ∀h,g,G,L,U,i,l. ⦃G, L⦄ ⊢ #i •*[h, g, l+1] U →
+ (∃∃K,V,W. ⇩[0, i] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •*[h, g, l+1] W &
+ ⇧[0, i + 1] W ≡ U
+ ) ∨
+ (∃∃K,W,V,l0. ⇩[0, i] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W ▪[h, g] l0 &
+ ⦃G, K⦄ ⊢ W •*[h, g, l] V & ⇧[0, i + 1] V ≡ U
+ ).
+#h #g #G #L #U #i #l #H elim (lsstas_inv_step_sn … H) -H
+#X #H #HXU elim (ssta_inv_lref1 … H) -H
+* #K [ #V #W | #W #l0 ] #HLK [ #HVW | #HWl0 ] #HWX
+lapply (ldrop_fwd_ldrop2 … HLK) #H0LK
+elim (lsstas_inv_lift1 … HXU … H0LK … HWX) -H0LK -X /3 width=8/ /4 width=6/
+qed-.
+
+(* Advanced forward lemmas **************************************************)
+
+lemma lsstas_fwd_correct: ∀h,g,G,L,T1,U1. ⦃G, L⦄ ⊢ T1 •[h, g] U1 →
+ ∀T2,l. ⦃G, L⦄ ⊢ T1 •*[h, g, l] T2 →
+ ∃U2. ⦃G, L⦄ ⊢ T2 •[h, g] U2.
+#h #g #G #L #T1 #U1 #HTU1 #T2 #l #H @(lsstas_ind_dx … H) -l -T2 [ /2 width=3/ ] -HTU1
+#l #T #T2 #_ #HT2 #_ -T1 -U1 -l
+elim (ssta_fwd_correct … HT2) -T /2 width=2/
+qed-.
+
+(* Advanced properties ******************************************************)
+
+lemma lsstas_total: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h, g] U →
+ ∀l. ∃U0. ⦃G, L⦄ ⊢ T •*[h, g, l] U0.
+#h #g #G #L #T #U #HTU #l @(nat_ind_plus … l) -l [ /2 width=2/ ]
+#l * #U0 #HTU0
+elim (lsstas_fwd_correct … HTU … HTU0) -U /3 width=4/
+qed-.
+
+lemma lsstas_ldef: ∀h,g,G,L,K,V,i. ⇩[0, i] L ≡ K.ⓓV →
+ ∀W,l. ⦃G, K⦄ ⊢ V •*[h, g, l+1] W →
+ ∀U. ⇧[0, i+1] W ≡ U → ⦃G, L⦄ ⊢ #i •*[h, g, l+1] U.
+#h #g #G #L #K #V #i #HLK #W #l #HVW #U #HWU
+lapply (ldrop_fwd_ldrop2 … HLK)
+elim (lsstas_inv_step_sn … HVW) -HVW #W0
+elim (lift_total W0 0 (i+1)) /3 width=11/
+qed.
+
+lemma lsstas_ldec: ∀h,g,G,L,K,W,i. ⇩[0, i] L ≡ K.ⓛW → ∀l0. ⦃G, K⦄ ⊢ W ▪[h, g] l0 →
+ ∀V,l. ⦃G, K⦄ ⊢ W •*[h, g, l] V →
+ ∀U. ⇧[0, i+1] V ≡ U → ⦃G, L⦄ ⊢ #i •*[h, g, l+1] U.
+#h #g #G #L #K #W #i #HLK #T #HWT #V #l #HWV #U #HVU
+lapply (ldrop_fwd_ldrop2 … HLK) #H
+elim (lift_total W 0 (i+1)) /3 width=11/
+qed.
+
+(* Properties on degree assignment for terms ********************************)
+
+lemma lsstas_da_conf: ∀h,g,G,L,T,U,l1. ⦃G, L⦄ ⊢ T •*[h, g, l1] U →
+ ∀l2. ⦃G, L⦄ ⊢ T ▪[h, g] l2 → ⦃G, L⦄ ⊢ U ▪[h, g] l2-l1.
+#h #g #G #L #T #U #l1 #H @(lsstas_ind_dx … H) -U -l1 //
+#l1 #U #U0 #_ #HU0 #IHTU #l2 #HT
+<minus_plus /3 width=3 by ssta_da_conf/
+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/ssta_ssta.ma".
+include "basic_2/unfold/lsstas_lift.ma".
+
+(* NAT-ITERATED STRATIFIED STATIC TYPE ASSIGNMENT FOR TERMS *****************)
+
+(* Main properties **********************************************************)
+
+theorem lsstas_trans: ∀h,g,G,L. ltransitive … (lsstas h g G L).
+/2 width=3 by lstar_ltransitive/ qed-.
+
+theorem lsstas_mono: ∀h,g,G,L,l. singlevalued … (lsstas h g G L l).
+/3 width=7 by ssta_mono, lstar_singlevalued/ qed-.
+
+theorem lsstas_conf_le: ∀h,g,G,L,T,U1,l1. ⦃G, L⦄ ⊢ T •*[h, g, l1] U1 →
+ ∀U2,l2. l1 ≤ l2 → ⦃G, L⦄ ⊢ T •*[h, g, l2] U2 →
+ ⦃G, L⦄ ⊢ U1 •*[h, g, l2 - l1] U2.
+#h #g #G #L #T #U1 #l1 #HTU1 #U2 #l2 #Hl12
+>(plus_minus_m_m … Hl12) in ⊢ (%→?); -Hl12 >commutative_plus #H
+elim (lsstas_split … H) -H #U #HTU
+>(lsstas_mono … HTU … HTU1) -T //
+qed-.
+
+(* Advanced properties ******************************************************)
+
+lemma lsstas_ssta_conf_pos: ∀h,g,G,L,T,U1. ⦃G, L⦄ ⊢ T •[h, g] U1 →
+ ∀U2,l. ⦃G, L⦄ ⊢ T •*[h, g, l+1] U2 → ⦃G, L⦄ ⊢ U1 •*[h, g, l] U2.
+#h #g #G #L #T #U1 #HTU1 #U2 #l #HTU2
+lapply (lsstas_conf_le … T U1 1 … HTU2) -HTU2 // /2 width=1/
+qed-.
+
+lemma lsstas_strip_pos: ∀h,g,G,L,T1,U1. ⦃G, L⦄ ⊢ T1 •[h, g] U1 →
+ ∀T2,l. ⦃G, L⦄ ⊢ T1 •*[h, g, l+1] T2 →
+ ∃∃U2. ⦃G, L⦄ ⊢ T2 •[h, g] U2 & ⦃G, L⦄ ⊢ U1 •*[h, g, l+1] U2.
+#h #g #G #L #T1 #U1 #HTU1 #T2 #l #HT12
+elim (lsstas_fwd_correct … HTU1 … HT12)
+lapply (lsstas_ssta_conf_pos … HTU1 … HT12) -T1 /3 width=5/
+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/ssta.ma".
-
-(* ITERATED STRATIFIED STATIC TYPE ASSIGNMENT FOR TERMS *********************)
-
-(* Note: includes: stass_refl *)
-definition sstas: ∀h. sd h → relation4 genv lenv term term ≝
- λh,g,G,L. star … (ssta_step h g G L).
-
-interpretation "iterated stratified static type assignment (term)"
- 'StaticTypeStar h g G L T U = (sstas h g G L T U).
-
-(* Basic eliminators ********************************************************)
-
-lemma sstas_ind: ∀h,g,G,L,T. ∀R:predicate term.
- R T → (
- ∀U1,U2,l. ⦃G, L⦄ ⊢ T •* [h, g] U1 → ⦃G, L⦄ ⊢ U1 •[h, g] ⦃l+1, U2⦄ →
- R U1 → R U2
- ) →
- ∀U. ⦃G, L⦄ ⊢ T •*[h, g] U → R U.
-#h #g #G #L #T #R #IH1 #IH2 #U #H elim H -U //
-#U1 #U2 #H * /2 width=5/
-qed-.
-
-lemma sstas_ind_dx: ∀h,g,G,L,U2. ∀R:predicate term.
- R U2 → (
- ∀T,U1,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l+1, U1⦄ → ⦃G, L⦄ ⊢ U1 •* [h, g] U2 →
- R U1 → R T
- ) →
- ∀T. ⦃G, L⦄ ⊢ T •*[h, g] U2 → R T.
-#h #g #G #L #U2 #R #IH1 #IH2 #T #H @(star_ind_l … T H) -T //
-#T #T0 * /2 width=5/
-qed-.
-
-(* Basic properties *********************************************************)
-
-lemma sstas_refl: ∀h,g,G,L. reflexive … (sstas h g G L).
-// qed.
-
-lemma ssta_sstas: ∀h,g,G,L,T,U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l+1, U⦄ → ⦃G, L⦄ ⊢ T •*[h, g] U.
-/3 width=2 by R_to_star, ex_intro/ qed. (**) (* auto fails without trace *)
-
-lemma sstas_strap1: ∀h,g,G,L,T1,T2,U2,l. ⦃G, L⦄ ⊢ T1 •*[h, g] T2 → ⦃G, L⦄ ⊢ T2 •[h, g] ⦃l+1, U2⦄ →
- ⦃G, L⦄ ⊢ T1 •*[h, g] U2.
-/3 width=4 by sstep, ex_intro/ (**) (* auto fails without trace *)
-qed.
-
-lemma sstas_strap2: ∀h,g,G,L,T1,U1,U2,l. ⦃G, L⦄ ⊢ T1 •[h, g] ⦃l+1, U1⦄ → ⦃G, L⦄ ⊢ U1 •*[h, g] U2 →
- ⦃G, L⦄ ⊢ T1 •*[h, g] U2.
-/3 width=3 by star_compl, ex_intro/ (**) (* auto fails without trace *)
-qed.
-
-(* Basic inversion lemmas ***************************************************)
-
-lemma sstas_inv_bind1: ∀h,g,a,I,G,L,Y,X,U. ⦃G, L⦄ ⊢ ⓑ{a,I}Y.X •*[h, g] U →
- ∃∃Z. ⦃G, L.ⓑ{I}Y⦄ ⊢ X •*[h, g] Z & U = ⓑ{a,I}Y.Z.
-#h #g #a #I #G #L #Y #X #U #H @(sstas_ind … H) -U /2 width=3/
-#T #U #l #_ #HTU * #Z #HXZ #H destruct
-elim (ssta_inv_bind1 … HTU) -HTU #Z0 #HZ0 #H destruct /3 width=4/
-qed-.
-
-lemma sstas_inv_appl1: ∀h,g,G,L,Y,X,U. ⦃G, L⦄ ⊢ ⓐY.X •*[h, g] U →
- ∃∃Z. ⦃G, L⦄ ⊢ X •*[h, g] Z & U = ⓐY.Z.
-#h #g #G #L #Y #X #U #H @(sstas_ind … H) -U /2 width=3/
-#T #U #l #_ #HTU * #Z #HXZ #H destruct
-elim (ssta_inv_appl1 … HTU) -HTU #Z0 #HZ0 #H destruct /3 width=4/
-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/ssta_aaa.ma".
-include "basic_2/unfold/sstas.ma".
-
-(* ITERATED STRATIFIED STATIC TYPE ASSIGNMENT FOR TERMS *********************)
-
-(* Properties on atomic arity assignment for terms **************************)
-
-lemma sstas_aaa: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •*[h, g] U →
- ∀A. ⦃G, L⦄ ⊢ T ⁝ A → ⦃G, L⦄ ⊢ U ⁝ A.
-#h #g #G #L #T #U #H @(sstas_ind_dx … H) -T // /3 width=6/
-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/ssta_lift.ma".
-include "basic_2/unfold/sstas.ma".
-
-(* ITERATED STRATIFIED STATIC TYPE ASSIGNMENT FOR TERMS *********************)
-
-(* Advanced forward lemmas **************************************************)
-
-lemma sstas_fwd_correct: ∀h,g,G,L,T1,U1,l1. ⦃G, L⦄ ⊢ T1 •[h, g] ⦃l1, U1⦄ →
- ∀T2. ⦃G, L⦄ ⊢ T1 •*[h, g] T2 →
- ∃∃U2,l2. ⦃G, L⦄ ⊢ T2 •[h, g] ⦃l2, U2⦄.
-#h #g #G #L #T1 #U1 #l1 #HTU1 #T2 #H @(sstas_ind … H) -T2 [ /2 width=3/ ] -HTU1
-#T #T2 #l #_ #HT2 * #U #l0 #_ -l0
-elim (ssta_fwd_correct … HT2) -T /2 width=3/
-qed-.
-
-(* Properties on relocation *************************************************)
-
-lemma sstas_lift: ∀h,g,G,L1,T1,U1. ⦃G, L1⦄ ⊢ T1 •*[h, g] U1 →
- ∀L2,d,e. ⇩[d, e] L2 ≡ L1 → ∀T2. ⇧[d, e] T1 ≡ T2 →
- ∀U2. ⇧[d, e] U1 ≡ U2 → ⦃G, L2⦄ ⊢ T2 •*[h, g] U2.
-#h #g #G #L1 #T1 #U1 #H @(sstas_ind_dx … H) -T1
-[ #L2 #d #e #HL21 #X #HX #U2 #HU12
- >(lift_mono … HX … HU12) -X //
-| #T0 #U0 #l0 #HTU0 #_ #IHU01 #L2 #d #e #HL21 #T2 #HT02 #U2 #HU12
- elim (lift_total U0 d e) /3 width=10/
-]
-qed.
-
-(* Inversion lemmas on relocation *******************************************)
-
-lemma sstas_inv_lift1: ∀h,g,G,L2,T2,U2. ⦃G, L2⦄ ⊢ T2 •*[h, g] U2 →
- ∀L1,d,e. ⇩[d, e] L2 ≡ L1 → ∀T1. ⇧[d, e] T1 ≡ T2 →
- ∃∃U1. ⦃G, L1⦄ ⊢ T1 •*[h, g] U1 & ⇧[d, e] U1 ≡ U2.
-#h #g #G #L2 #T2 #U2 #H @(sstas_ind_dx … H) -T2 /2 width=3/
-#T0 #U0 #l0 #HTU0 #_ #IHU01 #L1 #d #e #HL21 #U1 #HU12
-elim (ssta_inv_lift1 … HTU0 … HL21 … HU12) -HTU0 -HU12 #U #HU1 #HU0
-elim (IHU01 … HL21 … HU0) -IHU01 -HL21 -U0 /3 width=4/
-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/ssta_ssta.ma".
-include "basic_2/unfold/sstas.ma".
-
-(* ITERATED STRATIFIED STATIC TYPE ASSIGNMENT FOR TERMS *********************)
-
-(* Advanced inversion lemmas ************************************************)
-
-lemma sstas_inv_O: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •*[h, g] U →
- ∀T0. ⦃G, L⦄ ⊢ T •[h, g] ⦃0, T0⦄ → U = T.
-#h #g #G #L #T #U #H @(sstas_ind_dx … H) -T //
-#T0 #U0 #l0 #HTU0 #_ #_ #T1 #HT01
-elim (ssta_mono … HTU0 … HT01) <plus_n_Sm #H destruct
-qed-.
-
-(* Advanced properties ******************************************************)
-
-lemma sstas_strip: ∀h,g,G,L,T,U1. ⦃G, L⦄ ⊢ T •*[h, g] U1 →
- ∀U2,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U2⦄ →
- T = U1 ∨ ⦃G, L⦄ ⊢ U2 •*[h, g] U1.
-#h #g #G #L #T #U1 #H1 @(sstas_ind_dx … H1) -T /2 width=1/
-#T #U #l0 #HTU #HU1 #_ #U2 #l #H2
-elim (ssta_mono … H2 … HTU) -H2 -HTU #H1 #H2 destruct /2 width=1/
-qed-.
-
-(* Main properties **********************************************************)
-
-theorem sstas_trans: ∀h,g,G,L,T1,U. ⦃G, L⦄ ⊢ T1 •*[h, g] U →
- ∀T2. ⦃G, L⦄ ⊢ U •*[h, g] T2 → ⦃G, L⦄ ⊢ T1 •*[h, g] T2.
-/2 width=3/ qed-.
-
-theorem sstas_conf: ∀h,g,G,L,T,U1. ⦃G, L⦄ ⊢ T •*[h, g] U1 →
- ∀U2. ⦃G, L⦄ ⊢ T •*[h, g] U2 →
- ⦃G, L⦄ ⊢ U1 •*[h, g] U2 ∨ ⦃G, L⦄ ⊢ U2 •*[h, g] U1.
-#h #g #G #L #T #U1 #H1 @(sstas_ind_dx … H1) -T /2 width=1/
-#T #U #l #HTU #HU1 #IHU1 #U2 #H2
-elim (sstas_strip … H2 … HTU) #H destruct
-[ -H2 -IHU1 /3 width=4/
-| -T /2 width=1/
-]
-qed-.
<section>Summary of the Specification</section>
<body>Here is a numerical acount of the specification's contents
and its timeline.
- Nodes are counted according to the "intrinsinc complexity measure"
- [F. Guidi: "Procedural Representation of CIC Proof Terms"
+ Nodes are counted according to the "intrinsinc complexity measure"
+ [F. Guidi: "Procedural Representation of CIC Proof Terms"
Journal of Automated Reasoning 44(1-2), Springer (February 2010),
- pp. 53-78].
+ pp. 53-78].
</body>
<table name="basic_2_sum"/>
<news date="In progress.">
</news>
<news date="2013 August 7.">
Passive support for global environments.
- </news>
+ </news>
<news date="2013 July 27.">
Reaxiomatized β-reductum as in extended β-reduction
</news>
for simply typed terms.
</news>
<news date="2013 April 16.">
- Reaxiomatized substitution and reduction
+ Reaxiomatized substitution and reduction
commute with respect to subclosure
(anniversary milestone).
</news>
]
*)
[ { "\"big tree\" parallel computation" * } {
- [ "yprs ( ? ⊢ ⦃?,?,?⦄ ≥[?,?] ⦃?,?,?⦄ )" "yprs_yprs" "ygt ( ⦃?,?,?⦄ >[?,?] ⦃?,?,?⦄ )" "ygt_ygt" * ]
+ [ "ygt ( ⦃?,?,?⦄ >[?,?] ⦃?,?,?⦄ )" "ygt_lift" + "ygt_ygt" * ]
+ [ "yprs ( ? ⊢ ⦃?,?,?⦄ ≥[?,?] ⦃?,?,?⦄ )" "yprs_lift" + "yprs_yprs" * ]
}
]
[ { "\"big tree\" parallel reduction" * } {
}
]
[ { "local env. ref. for stratified native validity" * } {
- [ "lsubsv ( ? ⊢ ? ¡⊑[?,?] ? )" "lsubsv_ldrop" + "lsubsv_lsuba" + "lsubsv_ssta" + "lsubsv_dxprs" + "lsubsv_cpcs" + "lsubsv_snv" * ]
+ [ "lsubsv ( ? ⊢ ? ¡⊑[?,?] ? )" "lsubsv_ldrop" + "lsubsv_lsubd" + "lsubsv_lsuba" + "lsubsv_lsstas" + "lsubsv_cpds" + "lsubsv_cpcs" + "lsubsv_snv" * ]
}
]
[ { "stratified native validity" * } {
- [ "snv ( ⦃?,?⦄ ⊢ ? ¡[?,?] )" "snv_lift" + "snv_aaa" + "snv_ssta" + "snv_sstas" + "snv_ssta_lpr" + "snv_lpr" + "snv_cpcs" * ]
+ [ "snv ( ⦃?,?⦄ ⊢ ? ¡[?,?] )" "snv_lift" + "snv_da_lpr" + "snv_aaa" + "snv_lsstas" + "snv_lsstas_lpr" + "snv_lpr" + "snv_cpcs" * ]
}
]
}
]
class "blue"
[ { "equivalence" * } {
+ [ { "decomposed extended equivalence" * } {
+ [ "cpes ( ⦃?,?⦄ ⊢ ? •*⬌*[?,?] ? )" "cpes_cpds" * ]
+ }
+ ]
[ { "context-sensitive equivalence" * } {
[ "cpcs ( ⦃?,?⦄ ⊢ ? ⬌* ? )" "cpcs_aaa" + "cpcs_cprs" + "cpcs_cpcs" * ]
}
}
]
[ { "decomposed extended computation" * } {
- [ "dxprs ( ⦃?,?⦄ ⊢ ? •*➡*[?,?] ? )" "dxprs_lift" + "dxprs_aaa" + "dxprs_dxprs" * ]
+ [ "cpds ( ⦃?,?⦄ ⊢ ? •*➡*[?,?] ? )" "cpds_lift" + "cpds_aaa" + "cpds_cpds" * ]
}
]
[ { "context-sensitive extended computation" * } {
[ "unfold ( ⦃?,?⦄ ⊢ ? ⧫* ? )" * ]
}
]
- [ { "iterated stratified static type assignment" * } {
- [ "sstas ( ⦃?,?⦄ ⊢ ? •*[?] ? )" "sstas_lift" + "sstas_aaa" + "sstas_sstas" * ]
+ [ { "iterated static type assignment" * } {
+ [ "lsstas ( ⦃?,?⦄ ⊢ ? •*[?,?,?] ? )" "lsstas_alt ( ⦃?,?⦄ ⊢ ? ••*[?,?,?] ? )" "lsstas_lift" + "lsstas_aaa" + "lsstas_lsstas" * ]
}
]
}
]
class "grass"
[ { "static typing" * } {
- [ { "stratified static type assignment" * } {
- [ "ssta ( ⦃?,?⦄ ⊢ ? •[?,?] ⦃?,?⦄ )" "ssta_lift" + "ssta_aaa" + "ssta_ssta" * ]
- }
- ]
[ { "local env. ref. for atomic arity assignment" * } {
[ "lsuba ( ? ⊢ ? ⁝⊑ ? )" "lsuba_ldrop" + "lsuba_aaa" + "lsuba_lsuba" * ]
}
]
[ { "atomic arity assignment" * } {
- [ "aaa ( ⦃?,?⦄ ⊢ ? ⁝ ? )" "aaa_lift" + "aaa_lifts" + "aaa_aaa" * ]
+ [ "aaa ( ⦃?,?⦄ ⊢ ? ⁝ ? )" "aaa_lift" + "aaa_lifts" + "aaa_da" + "aaa_ssta" + "aaa_aaa" * ]
+ }
+ ]
+ [ { "stratified static type assignment" * } {
+ [ "ssta ( ⦃?,?⦄ ⊢ ? •[?,?] ? )" "ssta_lift" + "ssta_ssta" * ]
+ }
+ ]
+ [ { "local env. ref. for degree assignment" * } {
+ [ "lsubd ( ? ⊢ ? ▪⊑ ? )" "lsubd_da" + "lsubd_lsubd" * ]
+ }
+ ]
+ [ { "degree assignment" * } {
+ [ "da ( ⦃?,?⦄ ⊢ ? ▪[?,?] ? )" "da_lift" + "da_da" * ]
}
]
[ { "parameters" * } {
definition NF_dec: ∀A. relation A → relation A → Prop ≝
λA,R,S. ∀a1. NF A R S a1 ∨
- ∃∃a2. R … a1 a2 & (S a2 a1 → ⊥).
+ ∃∃a2. R … a1 a2 & (S a2 a1 → ⊥).
inductive SN (A) (R,S:relation A): predicate A ≝
| SN_intro: ∀a1. (∀a2. R a1 a2 → (S a2 a1 → ⊥) → SN A R S a2) → SN A R S a1
<key name="ex">6 6</key>
<key name="ex">6 7</key>
<key name="ex">7 4</key>
- <key name="ex">7 5</key>
<key name="ex">7 7</key>
<key name="ex">8 5</key>
- <key name="ex">8 6</key>
+ <key name="ex">9 3</key>
+ <key name="ex">10 4</key>
<key name="or">3</key>
<key name="or">4</key>
<key name="or">5</key>
interpretation "multiple existental quantifier (7, 4)" 'Ex P0 P1 P2 P3 P4 P5 P6 = (ex7_4 ? ? ? ? P0 P1 P2 P3 P4 P5 P6).
-(* multiple existental quantifier (7, 5) *)
-
-inductive ex7_5 (A0,A1,A2,A3,A4:Type[0]) (P0,P1,P2,P3,P4,P5,P6:A0→A1→A2→A3→A4→Prop) : Prop ≝
- | ex7_5_intro: ∀x0,x1,x2,x3,x4. P0 x0 x1 x2 x3 x4 → P1 x0 x1 x2 x3 x4 → P2 x0 x1 x2 x3 x4 → P3 x0 x1 x2 x3 x4 → P4 x0 x1 x2 x3 x4 → P5 x0 x1 x2 x3 x4 → P6 x0 x1 x2 x3 x4 → ex7_5 ? ? ? ? ? ? ? ? ? ? ? ?
-.
-
-interpretation "multiple existental quantifier (7, 5)" 'Ex P0 P1 P2 P3 P4 P5 P6 = (ex7_5 ? ? ? ? ? P0 P1 P2 P3 P4 P5 P6).
-
(* multiple existental quantifier (7, 7) *)
inductive ex7_7 (A0,A1,A2,A3,A4,A5,A6:Type[0]) (P0,P1,P2,P3,P4,P5,P6:A0→A1→A2→A3→A4→A5→A6→Prop) : Prop ≝
interpretation "multiple existental quantifier (8, 5)" 'Ex P0 P1 P2 P3 P4 P5 P6 P7 = (ex8_5 ? ? ? ? ? P0 P1 P2 P3 P4 P5 P6 P7).
-(* multiple existental quantifier (8, 6) *)
+(* multiple existental quantifier (9, 3) *)
+
+inductive ex9_3 (A0,A1,A2:Type[0]) (P0,P1,P2,P3,P4,P5,P6,P7,P8:A0→A1→A2→Prop) : Prop ≝
+ | ex9_3_intro: ∀x0,x1,x2. P0 x0 x1 x2 → P1 x0 x1 x2 → P2 x0 x1 x2 → P3 x0 x1 x2 → P4 x0 x1 x2 → P5 x0 x1 x2 → P6 x0 x1 x2 → P7 x0 x1 x2 → P8 x0 x1 x2 → ex9_3 ? ? ? ? ? ? ? ? ? ? ? ?
+.
+
+interpretation "multiple existental quantifier (9, 3)" 'Ex P0 P1 P2 P3 P4 P5 P6 P7 P8 = (ex9_3 ? ? ? P0 P1 P2 P3 P4 P5 P6 P7 P8).
+
+(* multiple existental quantifier (10, 4) *)
-inductive ex8_6 (A0,A1,A2,A3,A4,A5:Type[0]) (P0,P1,P2,P3,P4,P5,P6,P7:A0→A1→A2→A3→A4→A5→Prop) : Prop ≝
- | ex8_6_intro: ∀x0,x1,x2,x3,x4,x5. P0 x0 x1 x2 x3 x4 x5 → P1 x0 x1 x2 x3 x4 x5 → P2 x0 x1 x2 x3 x4 x5 → P3 x0 x1 x2 x3 x4 x5 → P4 x0 x1 x2 x3 x4 x5 → P5 x0 x1 x2 x3 x4 x5 → P6 x0 x1 x2 x3 x4 x5 → P7 x0 x1 x2 x3 x4 x5 → ex8_6 ? ? ? ? ? ? ? ? ? ? ? ? ? ?
+inductive ex10_4 (A0,A1,A2,A3:Type[0]) (P0,P1,P2,P3,P4,P5,P6,P7,P8,P9:A0→A1→A2→A3→Prop) : Prop ≝
+ | ex10_4_intro: ∀x0,x1,x2,x3. P0 x0 x1 x2 x3 → P1 x0 x1 x2 x3 → P2 x0 x1 x2 x3 → P3 x0 x1 x2 x3 → P4 x0 x1 x2 x3 → P5 x0 x1 x2 x3 → P6 x0 x1 x2 x3 → P7 x0 x1 x2 x3 → P8 x0 x1 x2 x3 → P9 x0 x1 x2 x3 → ex10_4 ? ? ? ? ? ? ? ? ? ? ? ? ? ?
.
-interpretation "multiple existental quantifier (8, 6)" 'Ex P0 P1 P2 P3 P4 P5 P6 P7 = (ex8_6 ? ? ? ? ? ? P0 P1 P2 P3 P4 P5 P6 P7).
+interpretation "multiple existental quantifier (10, 4)" 'Ex P0 P1 P2 P3 P4 P5 P6 P7 P8 P9 = (ex10_4 ? ? ? ? P0 P1 P2 P3 P4 P5 P6 P7 P8 P9).
(* multiple disjunction connective (3) *)
non associative with precedence 20
for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P4) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P5) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P6) }.
-(* multiple existental quantifier (7, 5) *)
-
-notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5 break & term 19 P6)"
- non associative with precedence 20
- for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.$P3) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.$P4) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.$P5) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.$P6) }.
-
-notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5 break & term 19 P6)"
- non associative with precedence 20
- for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P4) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P5) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P6) }.
-
(* multiple existental quantifier (7, 7) *)
notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 , ident x6 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5 break & term 19 P6)"
non associative with precedence 20
for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P4) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P5) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P6) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P7) }.
-(* multiple existental quantifier (8, 6) *)
+(* multiple existental quantifier (9, 3) *)
+
+notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5 break & term 19 P6 break & term 19 P7 break & term 19 P8)"
+ non associative with precedence 20
+ for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P3) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P4) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P5) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P6) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P7) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P8) }.
+
+notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5 break & term 19 P6 break & term 19 P7 break & term 19 P8)"
+ non associative with precedence 20
+ for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P4) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P5) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P6) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P7) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P8) }.
+
+(* multiple existental quantifier (10, 4) *)
-notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5 break & term 19 P6 break & term 19 P7)"
+notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5 break & term 19 P6 break & term 19 P7 break & term 19 P8 break & term 19 P9)"
non associative with precedence 20
- for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P3) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P4) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P5) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P6) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P7) }.
+ for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P3) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P4) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P5) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P6) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P7) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P8) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P9) }.
-notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5 break & term 19 P6 break & term 19 P7)"
+notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5 break & term 19 P6 break & term 19 P7 break & term 19 P8 break & term 19 P9)"
non associative with precedence 20
- for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P4) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P5) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P6) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P7) }.
+ for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P4) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P5) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P6) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P7) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P8) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P9) }.
(* multiple disjunction connective (3) *)