(* LOCAL ENVIRONMENT REFINEMENT FOR STRATIFIED NATIVE VALIDITY **************)
(* Note: this is not transitive *)
-inductive lsubsv (h) (g) (G): relation lenv ≝
-| 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_beta: ∀L1,L2,W,V,d1. ⦃G, L1⦄ ⊢ ⓝW.V ¡[h, g, d1] → ⦃G, L2⦄ ⊢ W ¡[h, g] →
- ⦃G, L1⦄ ⊢ V ▪[h, g] d1+1 → ⦃G, L2⦄ ⊢ W ▪[h, g] d1 →
- lsubsv h g G L1 L2 → lsubsv h g G (L1.ⓓⓝW.V) (L2.ⓛW)
+inductive lsubsv (h) (o) (G): relation lenv ≝
+| lsubsv_atom: lsubsv h o G (⋆) (⋆)
+| lsubsv_pair: ∀I,L1,L2,V. lsubsv h o G L1 L2 →
+ lsubsv h o G (L1.ⓑ{I}V) (L2.ⓑ{I}V)
+| lsubsv_beta: ∀L1,L2,W,V,d1. ⦃G, L1⦄ ⊢ ⓝW.V ¡[h, o, d1] → ⦃G, L2⦄ ⊢ W ¡[h, o] →
+ ⦃G, L1⦄ ⊢ V ▪[h, o] d1+1 → ⦃G, L2⦄ ⊢ W ▪[h, o] d1 →
+ lsubsv h o G L1 L2 → lsubsv h o G (L1.ⓓⓝW.V) (L2.ⓛW)
.
interpretation
"local environment refinement (stratified native validity)"
- 'LRSubEqV h g G L1 L2 = (lsubsv h g G L1 L2).
+ 'LRSubEqV h o G L1 L2 = (lsubsv h o G L1 L2).
(* Basic inversion lemmas ***************************************************)
-fact lsubsv_inv_atom1_aux: ∀h,g,G,L1,L2. G ⊢ L1 ⫃¡[h, g] L2 → L1 = ⋆ → L2 = ⋆.
-#h #g #G #L1 #L2 * -L1 -L2
+fact lsubsv_inv_atom1_aux: ∀h,o,G,L1,L2. G ⊢ L1 ⫃¡[h, o] L2 → L1 = ⋆ → L2 = ⋆.
+#h #o #G #L1 #L2 * -L1 -L2
[ //
| #I #L1 #L2 #V #_ #H destruct
| #L1 #L2 #W #V #d1 #_ #_ #_ #_ #_ #H destruct
]
qed-.
-lemma lsubsv_inv_atom1: ∀h,g,G,L2. G ⊢ ⋆ ⫃¡[h, g] L2 → L2 = ⋆.
+lemma lsubsv_inv_atom1: ∀h,o,G,L2. G ⊢ ⋆ ⫃¡[h, o] L2 → L2 = ⋆.
/2 width=6 by lsubsv_inv_atom1_aux/ qed-.
-fact lsubsv_inv_pair1_aux: ∀h,g,G,L1,L2. G ⊢ L1 ⫃¡[h, g] L2 →
+fact lsubsv_inv_pair1_aux: ∀h,o,G,L1,L2. G ⊢ L1 ⫃¡[h, o] L2 →
∀I,K1,X. L1 = K1.ⓑ{I}X →
- (∃∃K2. G ⊢ K1 ⫃¡[h, g] K2 & L2 = K2.ⓑ{I}X) ∨
- ∃∃K2,W,V,d1. ⦃G, K1⦄ ⊢ ⓝW.V ¡[h, g, d1] & ⦃G, K2⦄ ⊢ W ¡[h, g] &
- ⦃G, K1⦄ ⊢ V ▪[h, g] d1+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] d1 &
- G ⊢ K1 ⫃¡[h, g] K2 &
+ (∃∃K2. G ⊢ K1 ⫃¡[h, o] K2 & L2 = K2.ⓑ{I}X) ∨
+ ∃∃K2,W,V,d1. ⦃G, K1⦄ ⊢ ⓝW.V ¡[h, o, d1] & ⦃G, K2⦄ ⊢ W ¡[h, o] &
+ ⦃G, K1⦄ ⊢ V ▪[h, o] d1+1 & ⦃G, K2⦄ ⊢ W ▪[h, o] d1 &
+ G ⊢ K1 ⫃¡[h, o] K2 &
I = Abbr & L2 = K2.ⓛW & X = ⓝW.V.
-#h #g #G #L1 #L2 * -L1 -L2
+#h #o #G #L1 #L2 * -L1 -L2
[ #J #K1 #X #H destruct
| #I #L1 #L2 #V #HL12 #J #K1 #X #H destruct /3 width=3 by ex2_intro, or_introl/
| #L1 #L2 #W #V #d1 #HWV #HW #HVd1 #HWd1 #HL12 #J #K1 #X #H destruct /3 width=11 by or_intror, ex8_4_intro/
]
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,d1. ⦃G, K1⦄ ⊢ ⓝW.V ¡[h, g, d1] & ⦃G, K2⦄ ⊢ W ¡[h, g] &
- ⦃G, K1⦄ ⊢ V ▪[h, g] d1+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] d1 &
- G ⊢ K1 ⫃¡[h, g] K2 &
+lemma lsubsv_inv_pair1: ∀h,o,I,G,K1,L2,X. G ⊢ K1.ⓑ{I}X ⫃¡[h, o] L2 →
+ (∃∃K2. G ⊢ K1 ⫃¡[h, o] K2 & L2 = K2.ⓑ{I}X) ∨
+ ∃∃K2,W,V,d1. ⦃G, K1⦄ ⊢ ⓝW.V ¡[h, o, d1] & ⦃G, K2⦄ ⊢ W ¡[h, o] &
+ ⦃G, K1⦄ ⊢ V ▪[h, o] d1+1 & ⦃G, K2⦄ ⊢ W ▪[h, o] d1 &
+ G ⊢ K1 ⫃¡[h, o] 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
+fact lsubsv_inv_atom2_aux: ∀h,o,G,L1,L2. G ⊢ L1 ⫃¡[h, o] L2 → L2 = ⋆ → L1 = ⋆.
+#h #o #G #L1 #L2 * -L1 -L2
[ //
| #I #L1 #L2 #V #_ #H destruct
| #L1 #L2 #W #V #d1 #_ #_ #_ #_ #_ #H destruct
]
qed-.
-lemma lsubsv_inv_atom2: ∀h,g,G,L1. G ⊢ L1 ⫃¡[h, g] ⋆ → L1 = ⋆.
+lemma lsubsv_inv_atom2: ∀h,o,G,L1. G ⊢ L1 ⫃¡[h, o] ⋆ → L1 = ⋆.
/2 width=6 by lsubsv_inv_atom2_aux/ qed-.
-fact lsubsv_inv_pair2_aux: ∀h,g,G,L1,L2. G ⊢ L1 ⫃¡[h, g] L2 →
+fact lsubsv_inv_pair2_aux: ∀h,o,G,L1,L2. G ⊢ L1 ⫃¡[h, o] L2 →
∀I,K2,W. L2 = K2.ⓑ{I}W →
- (∃∃K1. G ⊢ K1 ⫃¡[h, g] K2 & L1 = K1.ⓑ{I}W) ∨
- ∃∃K1,V,d1. ⦃G, K1⦄ ⊢ ⓝW.V ¡[h, g, d1] & ⦃G, K2⦄ ⊢ W ¡[h, g] &
- ⦃G, K1⦄ ⊢ V ▪[h, g] d1+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] d1 &
- G ⊢ K1 ⫃¡[h, g] K2 & I = Abst & L1 = K1.ⓓⓝW.V.
-#h #g #G #L1 #L2 * -L1 -L2
+ (∃∃K1. G ⊢ K1 ⫃¡[h, o] K2 & L1 = K1.ⓑ{I}W) ∨
+ ∃∃K1,V,d1. ⦃G, K1⦄ ⊢ ⓝW.V ¡[h, o, d1] & ⦃G, K2⦄ ⊢ W ¡[h, o] &
+ ⦃G, K1⦄ ⊢ V ▪[h, o] d1+1 & ⦃G, K2⦄ ⊢ W ▪[h, o] d1 &
+ G ⊢ K1 ⫃¡[h, o] K2 & I = Abst & L1 = K1.ⓓⓝW.V.
+#h #o #G #L1 #L2 * -L1 -L2
[ #J #K2 #U #H destruct
| #I #L1 #L2 #V #HL12 #J #K2 #U #H destruct /3 width=3 by ex2_intro, or_introl/
| #L1 #L2 #W #V #d1 #HWV #HW #HVd1 #HWd1 #HL12 #J #K2 #U #H destruct /3 width=8 by or_intror, ex7_3_intro/
]
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,d1. ⦃G, K1⦄ ⊢ ⓝW.V ¡[h, g, d1] & ⦃G, K2⦄ ⊢ W ¡[h, g] &
- ⦃G, K1⦄ ⊢ V ▪[h, g] d1+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] d1 &
- G ⊢ K1 ⫃¡[h, g] K2 & I = Abst & L1 = K1.ⓓⓝW.V.
+lemma lsubsv_inv_pair2: ∀h,o,I,G,L1,K2,W. G ⊢ L1 ⫃¡[h, o] K2.ⓑ{I}W →
+ (∃∃K1. G ⊢ K1 ⫃¡[h, o] K2 & L1 = K1.ⓑ{I}W) ∨
+ ∃∃K1,V,d1. ⦃G, K1⦄ ⊢ ⓝW.V ¡[h, o, d1] & ⦃G, K2⦄ ⊢ W ¡[h, o] &
+ ⦃G, K1⦄ ⊢ V ▪[h, o] d1+1 & ⦃G, K2⦄ ⊢ W ▪[h, o] d1 &
+ G ⊢ K1 ⫃¡[h, o] K2 & I = Abst & L1 = K1.ⓓⓝW.V.
/2 width=3 by lsubsv_inv_pair2_aux/ qed-.
(* Basic forward lemmas *****************************************************)
-lemma lsubsv_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 by lsubr_pair, lsubr_beta/
+lemma lsubsv_fwd_lsubr: ∀h,o,G,L1,L2. G ⊢ L1 ⫃¡[h, o] L2 → L1 ⫃ L2.
+#h #o #G #L1 #L2 #H elim H -L1 -L2 /2 width=1 by lsubr_pair, lsubr_beta/
qed-.
(* Basic properties *********************************************************)
-lemma lsubsv_refl: ∀h,g,G,L. G ⊢ L ⫃¡[h, g] L.
-#h #g #G #L elim L -L /2 width=1 by lsubsv_pair/
+lemma lsubsv_refl: ∀h,o,G,L. G ⊢ L ⫃¡[h, o] L.
+#h #o #G #L elim L -L /2 width=1 by lsubsv_pair/
qed.
-lemma lsubsv_cprs_trans: ∀h,g,G,L1,L2. G ⊢ L1 ⫃¡[h, g] L2 →
+lemma lsubsv_cprs_trans: ∀h,o,G,L1,L2. G ⊢ L1 ⫃¡[h, o] L2 →
∀T1,T2. ⦃G, L2⦄ ⊢ T1 ➡* T2 → ⦃G, L1⦄ ⊢ T1 ➡* T2.
/3 width=6 by lsubsv_fwd_lsubr, lsubr_cprs_trans/
qed-.
(* Note: the constant 0 cannot be generalized *)
-lemma lsubsv_drop_O1_conf: ∀h,g,G,L1,L2. G ⊢ L1 ⫃¡[h, g] L2 →
- ∀K1,s,m. ⬇[s, 0, m] L1 ≡ K1 →
- ∃∃K2. G ⊢ K1 ⫃¡[h, g] K2 & ⬇[s, 0, m] L2 ≡ K2.
-#h #g #G #L1 #L2 #H elim H -L1 -L2
+lemma lsubsv_drop_O1_conf: ∀h,o,G,L1,L2. G ⊢ L1 ⫃¡[h, o] L2 →
+ ∀K1,c,k. ⬇[c, 0, k] L1 ≡ K1 →
+ ∃∃K2. G ⊢ K1 ⫃¡[h, o] K2 & ⬇[c, 0, k] L2 ≡ K2.
+#h #o #G #L1 #L2 #H elim H -L1 -L2
[ /2 width=3 by ex2_intro/
-| #I #L1 #L2 #V #_ #IHL12 #K1 #s #m #H
+| #I #L1 #L2 #V #_ #IHL12 #K1 #c #k #H
elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK1
[ destruct
- elim (IHL12 L1 s 0) -IHL12 // #X #HL12 #H
+ elim (IHL12 L1 c 0) -IHL12 // #X #HL12 #H
<(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsubsv_pair, drop_pair, ex2_intro/
| elim (IHL12 … HLK1) -L1 /3 width=3 by drop_drop_lt, ex2_intro/
]
-| #L1 #L2 #W #V #d1 #HWV #HW #HVd1 #HWd1 #_ #IHL12 #K1 #s #m #H
+| #L1 #L2 #W #V #d1 #HWV #HW #HVd1 #HWd1 #_ #IHL12 #K1 #c #k #H
elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK1
[ destruct
- elim (IHL12 L1 s 0) -IHL12 // #X #HL12 #H
+ elim (IHL12 L1 c 0) -IHL12 // #X #HL12 #H
<(drop_inv_O2 … H) in HL12; -H /3 width=4 by lsubsv_beta, drop_pair, ex2_intro/
| elim (IHL12 … HLK1) -L1 /3 width=3 by drop_drop_lt, ex2_intro/
]
qed-.
(* Note: the constant 0 cannot be generalized *)
-lemma lsubsv_drop_O1_trans: ∀h,g,G,L1,L2. G ⊢ L1 ⫃¡[h, g] L2 →
- ∀K2,s, m. ⬇[s, 0, m] L2 ≡ K2 →
- ∃∃K1. G ⊢ K1 ⫃¡[h, g] K2 & ⬇[s, 0, m] L1 ≡ K1.
-#h #g #G #L1 #L2 #H elim H -L1 -L2
+lemma lsubsv_drop_O1_trans: ∀h,o,G,L1,L2. G ⊢ L1 ⫃¡[h, o] L2 →
+ ∀K2,c, k. ⬇[c, 0, k] L2 ≡ K2 →
+ ∃∃K1. G ⊢ K1 ⫃¡[h, o] K2 & ⬇[c, 0, k] L1 ≡ K1.
+#h #o #G #L1 #L2 #H elim H -L1 -L2
[ /2 width=3 by ex2_intro/
-| #I #L1 #L2 #V #_ #IHL12 #K2 #s #m #H
+| #I #L1 #L2 #V #_ #IHL12 #K2 #c #k #H
elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK2
[ destruct
- elim (IHL12 L2 s 0) -IHL12 // #X #HL12 #H
+ elim (IHL12 L2 c 0) -IHL12 // #X #HL12 #H
<(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsubsv_pair, drop_pair, ex2_intro/
| elim (IHL12 … HLK2) -L2 /3 width=3 by drop_drop_lt, ex2_intro/
]
-| #L1 #L2 #W #V #d1 #HWV #HW #HVd1 #HWd1 #_ #IHL12 #K2 #s #m #H
+| #L1 #L2 #W #V #d1 #HWV #HW #HVd1 #HWd1 #_ #IHL12 #K2 #c #k #H
elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK2
[ destruct
- elim (IHL12 L2 s 0) -IHL12 // #X #HL12 #H
+ elim (IHL12 L2 c 0) -IHL12 // #X #HL12 #H
<(drop_inv_O2 … H) in HL12; -H /3 width=4 by lsubsv_beta, drop_pair, ex2_intro/
| elim (IHL12 … HLK2) -L2 /3 width=3 by drop_drop_lt, ex2_intro/
]
(* Properties on context-sensitive parallel equivalence for terms ***********)
-lemma lsubsv_cpcs_trans: ∀h,g,G,L1,L2. G ⊢ L1 ⫃¡[h, g] L2 →
+lemma lsubsv_cpcs_trans: ∀h,o,G,L1,L2. G ⊢ L1 ⫃¡[h, o] L2 →
∀T1,T2. ⦃G, L2⦄ ⊢ T1 ⬌* T2 → ⦃G, L1⦄ ⊢ T1 ⬌* T2.
/3 width=6 by lsubsv_fwd_lsubr, lsubr_cpcs_trans/
qed-.
(* Properties on nat-iterated static type assignment ************************)
-lemma lsubsv_lstas_trans: ∀h,g,G,L2,T,U2,d2. ⦃G, L2⦄ ⊢ T •*[h, d2] U2 →
- ∀d1. d2 ≤ d1 → ⦃G, L2⦄ ⊢ T ▪[h, g] d1 →
- ∀L1. G ⊢ L1 ⫃¡[h, g] L2 →
+lemma lsubsv_lstas_trans: ∀h,o,G,L2,T,U2,d2. ⦃G, L2⦄ ⊢ T •*[h, d2] U2 →
+ ∀d1. d2 ≤ d1 → ⦃G, L2⦄ ⊢ T ▪[h, o] d1 →
+ ∀L1. G ⊢ L1 ⫃¡[h, o] L2 →
∃∃U1. ⦃G, L1⦄ ⊢ T •*[h, d2] U1 & ⦃G, L1⦄ ⊢ U1 ⬌* U2.
-#h #g #G #L2 #T #U #d2 #H elim H -G -L2 -T -U -d2
+#h #o #G #L2 #T #U #d2 #H elim H -G -L2 -T -U -d2
[ /2 width=3 by ex2_intro/
| #G #L2 #K2 #V #W #U #i #d2 #HLK2 #_ #HWU #IHVW #d1 #Hd21 #Hd1 #L1 #HL12
elim (da_inv_lref … Hd1) -Hd1 * #K0 #V0 [| #d0 ] #HK0 #HV0
]
qed-.
-lemma lsubsv_sta_trans: ∀h,g,G,L2,T,U2. ⦃G, L2⦄ ⊢ T •*[h, 1] U2 →
- ∀d. ⦃G, L2⦄ ⊢ T ▪[h, g] d+1 →
- ∀L1. G ⊢ L1 ⫃¡[h, g] L2 →
+lemma lsubsv_sta_trans: ∀h,o,G,L2,T,U2. ⦃G, L2⦄ ⊢ T •*[h, 1] U2 →
+ ∀d. ⦃G, L2⦄ ⊢ T ▪[h, o] d+1 →
+ ∀L1. G ⊢ L1 ⫃¡[h, o] L2 →
∃∃U1. ⦃G, L1⦄ ⊢ T •*[h, 1] U1 & ⦃G, L1⦄ ⊢ U1 ⬌* U2.
/2 width=7 by lsubsv_lstas_trans/ qed-.
(* Forward lemmas on lenv refinement for atomic arity assignment ************)
-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 by lsuba_pair/
+lemma lsubsv_fwd_lsuba: ∀h,o,G,L1,L2. G ⊢ L1 ⫃¡[h, o] L2 → G ⊢ L1 ⫃⁝ L2.
+#h #o #G #L1 #L2 #H elim H -L1 -L2 /2 width=1 by lsuba_pair/
#L1 #L2 #W #V #d1 #H #_ #_ #_ #_ #IHL12
elim (shnv_inv_cast … H) -H #HW #HV #H
lapply (H 0 ?) // -d1 #HWV
(* 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=3 by lsubd_pair, lsubd_beta/
+lemma lsubsv_fwd_lsubd: ∀h,o,G,L1,L2. G ⊢ L1 ⫃¡[h, o] L2 → G ⊢ L1 ⫃▪[h, o] L2.
+#h #o #G #L1 #L2 #H elim H -L1 -L2 /2 width=3 by lsubd_pair, lsubd_beta/
qed-.
(* Properties on decomposed extended parallel computation on terms **********)
-lemma lsubsv_scpds_trans: ∀h,g,G,L2,T1,T2,d. ⦃G, L2⦄ ⊢ T1 •*➡*[h, g, d] T2 →
- ∀L1. G ⊢ L1 ⫃¡[h, g] L2 →
- ∃∃T. ⦃G, L1⦄ ⊢ T1 •*➡*[h, g, d] T & ⦃G, L1⦄ ⊢ T2 ➡* T.
-#h #g #G #L2 #T1 #T2 #d2 * #T #d1 #Hd21 #Hd1 #HT1 #HT2 #L1 #HL12
+lemma lsubsv_scpds_trans: ∀h,o,G,L2,T1,T2,d. ⦃G, L2⦄ ⊢ T1 •*➡*[h, o, d] T2 →
+ ∀L1. G ⊢ L1 ⫃¡[h, o] L2 →
+ ∃∃T. ⦃G, L1⦄ ⊢ T1 •*➡*[h, o, d] T & ⦃G, L1⦄ ⊢ T2 ➡* T.
+#h #o #G #L2 #T1 #T2 #d2 * #T #d1 #Hd21 #Hd1 #HT1 #HT2 #L1 #HL12
lapply (lsubsv_cprs_trans … HL12 … HT2) -HT2 #HT2
elim (lsubsv_lstas_trans … HT1 … Hd1 … HL12) // #T0 #HT10 #HT0
lapply (cpcs_cprs_strap1 … HT0 … HT2) -T #HT02
(* Properties concerning stratified native validity *************************)
-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 //
+lemma lsubsv_snv_trans: ∀h,o,G,L2,T. ⦃G, L2⦄ ⊢ T ¡[h, o] →
+ ∀L1. G ⊢ L1 ⫃¡[h, o] L2 → ⦃G, L1⦄ ⊢ T ¡[h, o].
+#h #o #G #L2 #T #H elim H -G -L2 -T //
[ #I #G #L2 #K2 #V #i #HLK2 #_ #IHV #L1 #HL12
elim (lsubsv_drop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1
elim (lsubsv_inv_pair2 … H) -H * #K1
(* STRATIFIED HIGHER NATIVE VALIDITY FOR TERMS ******************************)
-inductive shnv (h) (g) (d1) (G) (L): predicate term ≝
-| shnv_cast: ∀U,T. ⦃G, L⦄ ⊢ U ¡[h, g] → ⦃G, L⦄ ⊢ T ¡[h, g] →
- (∀d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ U •*⬌*[h, g, d2, d2+1] T) →
- shnv h g d1 G L (ⓝU.T)
+inductive shnv (h) (o) (d1) (G) (L): predicate term ≝
+| shnv_cast: ∀U,T. ⦃G, L⦄ ⊢ U ¡[h, o] → ⦃G, L⦄ ⊢ T ¡[h, o] →
+ (∀d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ U •*⬌*[h, o, d2, d2+1] T) →
+ shnv h o d1 G L (ⓝU.T)
.
interpretation "stratified higher native validity (term)"
- 'NativeValid h g d G L T = (shnv h g d G L T).
+ 'NativeValid h o d G L T = (shnv h o d G L T).
(* Basic inversion lemmas ***************************************************)
-fact shnv_inv_cast_aux: ∀h,g,G,L,X,d1. ⦃G, L⦄ ⊢ X ¡[h, g, d1] → ∀U,T. X = ⓝU.T →
- ∧∧ ⦃G, L⦄ ⊢ U ¡[h, g] & ⦃G, L⦄ ⊢ T ¡[h, g]
- & (∀d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ U •*⬌*[h, g, d2, d2+1] T).
-#h #g #G #L #X #d1 * -X
+fact shnv_inv_cast_aux: ∀h,o,G,L,X,d1. ⦃G, L⦄ ⊢ X ¡[h, o, d1] → ∀U,T. X = ⓝU.T →
+ ∧∧ ⦃G, L⦄ ⊢ U ¡[h, o] & ⦃G, L⦄ ⊢ T ¡[h, o]
+ & (∀d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ U •*⬌*[h, o, d2, d2+1] T).
+#h #o #G #L #X #d1 * -X
#U #T #HU #HT #HUT #U1 #T1 #H destruct /3 width=1 by and3_intro/
qed-.
-lemma shnv_inv_cast: ∀h,g,G,L,U,T,d1. ⦃G, L⦄ ⊢ ⓝU.T ¡[h, g, d1] →
- ∧∧ ⦃G, L⦄ ⊢ U ¡[h, g] & ⦃G, L⦄ ⊢ T ¡[h, g]
- & (∀d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ U •*⬌*[h, g, d2, d2+1] T).
+lemma shnv_inv_cast: ∀h,o,G,L,U,T,d1. ⦃G, L⦄ ⊢ ⓝU.T ¡[h, o, d1] →
+ ∧∧ ⦃G, L⦄ ⊢ U ¡[h, o] & ⦃G, L⦄ ⊢ T ¡[h, o]
+ & (∀d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ U •*⬌*[h, o, d2, d2+1] T).
/2 width=3 by shnv_inv_cast_aux/ qed-.
-lemma shnv_inv_snv: ∀h,g,G,L,T,d. ⦃G, L⦄ ⊢ T ¡[h, g, d] → ⦃G, L⦄ ⊢ T ¡[h, g].
-#h #g #G #L #T #d * -T
+lemma shnv_inv_snv: ∀h,o,G,L,T,d. ⦃G, L⦄ ⊢ T ¡[h, o, d] → ⦃G, L⦄ ⊢ T ¡[h, o].
+#h #o #G #L #T #d * -T
#U #T #HU #HT #HUT elim (HUT 0) -HUT /2 width=3 by snv_cast/
qed-.
(* Basic properties *********************************************************)
-lemma snv_shnv_cast: ∀h,g,G,L,U,T. ⦃G, L⦄ ⊢ ⓝU.T ¡[h, g] → ⦃G, L⦄ ⊢ ⓝU.T ¡[h, g, 0].
-#h #g #G #L #U #T #H elim (snv_inv_cast … H) -H
+lemma snv_shnv_cast: ∀h,o,G,L,U,T. ⦃G, L⦄ ⊢ ⓝU.T ¡[h, o] → ⦃G, L⦄ ⊢ ⓝU.T ¡[h, o, 0].
+#h #o #G #L #U #T #H elim (snv_inv_cast … H) -H
#U0 #HU #HT #HU0 #HTU0 @shnv_cast // -HU -HT
#d #H <(le_n_O_to_eq … H) -d /2 width=3 by scpds_div/
qed.
(* STRATIFIED NATIVE VALIDITY FOR TERMS *************************************)
(* activate genv *)
-inductive snv (h) (g): relation3 genv lenv term ≝
-| snv_sort: ∀G,L,k. snv h g G L (⋆k)
-| snv_lref: ∀I,G,L,K,V,i. ⬇[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,W0,T,U0,d. snv h g G L V → snv h g G L T →
- ⦃G, L⦄ ⊢ V •*➡*[h, g, 1] W0 → ⦃G, L⦄ ⊢ T •*➡*[h, g, d] ⓛ{a}W0.U0 → snv h g G L (ⓐV.T)
-| snv_cast: ∀G,L,U,T,U0. snv h g G L U → snv h g G L T →
- ⦃G, L⦄ ⊢ U •*➡*[h, g, 0] U0 → ⦃G, L⦄ ⊢ T •*➡*[h, g, 1] U0 → snv h g G L (ⓝU.T)
+inductive snv (h) (o): relation3 genv lenv term ≝
+| snv_sort: ∀G,L,s. snv h o G L (⋆s)
+| snv_lref: ∀I,G,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → snv h o G K V → snv h o G L (#i)
+| snv_bind: ∀a,I,G,L,V,T. snv h o G L V → snv h o G (L.ⓑ{I}V) T → snv h o G L (ⓑ{a,I}V.T)
+| snv_appl: ∀a,G,L,V,W0,T,U0,d. snv h o G L V → snv h o G L T →
+ ⦃G, L⦄ ⊢ V •*➡*[h, o, 1] W0 → ⦃G, L⦄ ⊢ T •*➡*[h, o, d] ⓛ{a}W0.U0 → snv h o G L (ⓐV.T)
+| snv_cast: ∀G,L,U,T,U0. snv h o G L U → snv h o G L T →
+ ⦃G, L⦄ ⊢ U •*➡*[h, o, 0] U0 → ⦃G, L⦄ ⊢ T •*➡*[h, o, 1] U0 → snv h o G L (ⓝU.T)
.
interpretation "stratified native validity (term)"
- 'NativeValid h g G L T = (snv h g G L T).
+ 'NativeValid h o G L T = (snv h o G L T).
(* Basic inversion lemmas ***************************************************)
-fact snv_inv_lref_aux: ∀h,g,G,L,X. ⦃G, L⦄ ⊢ X ¡[h, g] → ∀i. X = #i →
- ∃∃I,K,V. ⬇[i] L ≡ K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ¡[h, g].
-#h #g #G #L #X * -G -L -X
-[ #G #L #k #i #H destruct
+fact snv_inv_lref_aux: ∀h,o,G,L,X. ⦃G, L⦄ ⊢ X ¡[h, o] → ∀i. X = #i →
+ ∃∃I,K,V. ⬇[i] L ≡ K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ¡[h, o].
+#h #o #G #L #X * -G -L -X
+[ #G #L #s #i #H destruct
| #I #G #L #K #V #i0 #HLK #HV #i #H destruct /2 width=5 by ex2_3_intro/
| #a #I #G #L #V #T #_ #_ #i #H destruct
| #a #G #L #V #W0 #T #U0 #d #_ #_ #_ #_ #i #H destruct
]
qed-.
-lemma snv_inv_lref: ∀h,g,G,L,i. ⦃G, L⦄ ⊢ #i ¡[h, g] →
- ∃∃I,K,V. ⬇[i] L ≡ K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ¡[h, g].
+lemma snv_inv_lref: ∀h,o,G,L,i. ⦃G, L⦄ ⊢ #i ¡[h, o] →
+ ∃∃I,K,V. ⬇[i] L ≡ K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ¡[h, o].
/2 width=3 by snv_inv_lref_aux/ qed-.
-fact snv_inv_gref_aux: ∀h,g,G,L,X. ⦃G, L⦄ ⊢ X ¡[h, g] → ∀p. X = §p → ⊥.
-#h #g #G #L #X * -G -L -X
-[ #G #L #k #p #H destruct
+fact snv_inv_gref_aux: ∀h,o,G,L,X. ⦃G, L⦄ ⊢ X ¡[h, o] → ∀p. X = §p → ⊥.
+#h #o #G #L #X * -G -L -X
+[ #G #L #s #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 #W0 #T #U0 #d #_ #_ #_ #_ #p #H destruct
]
qed-.
-lemma snv_inv_gref: ∀h,g,G,L,p. ⦃G, L⦄ ⊢ §p ¡[h, g] → ⊥.
+lemma snv_inv_gref: ∀h,o,G,L,p. ⦃G, L⦄ ⊢ §p ¡[h, o] → ⊥.
/2 width=8 by snv_inv_gref_aux/ qed-.
-fact snv_inv_bind_aux: ∀h,g,G,L,X. ⦃G, L⦄ ⊢ X ¡[h, g] → ∀a,I,V,T. X = ⓑ{a,I}V.T →
- ⦃G, L⦄ ⊢ V ¡[h, g] ∧ ⦃G, L.ⓑ{I}V⦄ ⊢ T ¡[h, g].
-#h #g #G #L #X * -G -L -X
-[ #G #L #k #b #Z #X1 #X2 #H destruct
+fact snv_inv_bind_aux: ∀h,o,G,L,X. ⦃G, L⦄ ⊢ X ¡[h, o] → ∀a,I,V,T. X = ⓑ{a,I}V.T →
+ ⦃G, L⦄ ⊢ V ¡[h, o] ∧ ⦃G, L.ⓑ{I}V⦄ ⊢ T ¡[h, o].
+#h #o #G #L #X * -G -L -X
+[ #G #L #s #b #Z #X1 #X2 #H destruct
| #I #G #L #K #V #i #_ #_ #b #Z #X1 #X2 #H destruct
| #a #I #G #L #V #T #HV #HT #b #Z #X1 #X2 #H destruct /2 width=1 by conj/
| #a #G #L #V #W0 #T #U0 #d #_ #_ #_ #_ #b #Z #X1 #X2 #H destruct
]
qed-.
-lemma snv_inv_bind: ∀h,g,a,I,G,L,V,T. ⦃G, L⦄ ⊢ ⓑ{a,I}V.T ¡[h, g] →
- ⦃G, L⦄ ⊢ V ¡[h, g] ∧ ⦃G, L.ⓑ{I}V⦄ ⊢ T ¡[h, g].
+lemma snv_inv_bind: ∀h,o,a,I,G,L,V,T. ⦃G, L⦄ ⊢ ⓑ{a,I}V.T ¡[h, o] →
+ ⦃G, L⦄ ⊢ V ¡[h, o] ∧ ⦃G, L.ⓑ{I}V⦄ ⊢ T ¡[h, o].
/2 width=4 by snv_inv_bind_aux/ qed-.
-fact snv_inv_appl_aux: ∀h,g,G,L,X. ⦃G, L⦄ ⊢ X ¡[h, g] → ∀V,T. X = ⓐV.T →
- ∃∃a,W0,U0,d. ⦃G, L⦄ ⊢ V ¡[h, g] & ⦃G, L⦄ ⊢ T ¡[h, g] &
- ⦃G, L⦄ ⊢ V •*➡*[h, g, 1] W0 & ⦃G, L⦄ ⊢ T •*➡*[h, g, d] ⓛ{a}W0.U0.
-#h #g #G #L #X * -L -X
-[ #G #L #k #X1 #X2 #H destruct
+fact snv_inv_appl_aux: ∀h,o,G,L,X. ⦃G, L⦄ ⊢ X ¡[h, o] → ∀V,T. X = ⓐV.T →
+ ∃∃a,W0,U0,d. ⦃G, L⦄ ⊢ V ¡[h, o] & ⦃G, L⦄ ⊢ T ¡[h, o] &
+ ⦃G, L⦄ ⊢ V •*➡*[h, o, 1] W0 & ⦃G, L⦄ ⊢ T •*➡*[h, o, d] ⓛ{a}W0.U0.
+#h #o #G #L #X * -L -X
+[ #G #L #s #X1 #X2 #H destruct
| #I #G #L #K #V #i #_ #_ #X1 #X2 #H destruct
| #a #I #G #L #V #T #_ #_ #X1 #X2 #H destruct
| #a #G #L #V #W0 #T #U0 #d #HV #HT #HVW0 #HTU0 #X1 #X2 #H destruct /2 width=6 by ex4_4_intro/
]
qed-.
-lemma snv_inv_appl: ∀h,g,G,L,V,T. ⦃G, L⦄ ⊢ ⓐV.T ¡[h, g] →
- ∃∃a,W0,U0,d. ⦃G, L⦄ ⊢ V ¡[h, g] & ⦃G, L⦄ ⊢ T ¡[h, g] &
- ⦃G, L⦄ ⊢ V •*➡*[h, g, 1] W0 & ⦃G, L⦄ ⊢ T •*➡*[h, g, d] ⓛ{a}W0.U0.
+lemma snv_inv_appl: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ⓐV.T ¡[h, o] →
+ ∃∃a,W0,U0,d. ⦃G, L⦄ ⊢ V ¡[h, o] & ⦃G, L⦄ ⊢ T ¡[h, o] &
+ ⦃G, L⦄ ⊢ V •*➡*[h, o, 1] W0 & ⦃G, L⦄ ⊢ T •*➡*[h, o, d] ⓛ{a}W0.U0.
/2 width=3 by snv_inv_appl_aux/ qed-.
-fact snv_inv_cast_aux: ∀h,g,G,L,X. ⦃G, L⦄ ⊢ X ¡[h, g] → ∀U,T. X = ⓝU.T →
- ∃∃U0. ⦃G, L⦄ ⊢ U ¡[h, g] & ⦃G, L⦄ ⊢ T ¡[h, g] &
- ⦃G, L⦄ ⊢ U •*➡*[h, g, 0] U0 & ⦃G, L⦄ ⊢ T •*➡*[h, g, 1] U0.
-#h #g #G #L #X * -G -L -X
-[ #G #L #k #X1 #X2 #H destruct
+fact snv_inv_cast_aux: ∀h,o,G,L,X. ⦃G, L⦄ ⊢ X ¡[h, o] → ∀U,T. X = ⓝU.T →
+ ∃∃U0. ⦃G, L⦄ ⊢ U ¡[h, o] & ⦃G, L⦄ ⊢ T ¡[h, o] &
+ ⦃G, L⦄ ⊢ U •*➡*[h, o, 0] U0 & ⦃G, L⦄ ⊢ T •*➡*[h, o, 1] U0.
+#h #o #G #L #X * -G -L -X
+[ #G #L #s #X1 #X2 #H destruct
| #I #G #L #K #V #i #_ #_ #X1 #X2 #H destruct
| #a #I #G #L #V #T #_ #_ #X1 #X2 #H destruct
| #a #G #L #V #W0 #T #U0 #d #_ #_ #_ #_ #X1 #X2 #H destruct
]
qed-.
-lemma snv_inv_cast: ∀h,g,G,L,U,T. ⦃G, L⦄ ⊢ ⓝU.T ¡[h, g] →
- ∃∃U0. ⦃G, L⦄ ⊢ U ¡[h, g] & ⦃G, L⦄ ⊢ T ¡[h, g] &
- ⦃G, L⦄ ⊢ U •*➡*[h, g, 0] U0 & ⦃G, L⦄ ⊢ T •*➡*[h, g, 1] U0.
+lemma snv_inv_cast: ∀h,o,G,L,U,T. ⦃G, L⦄ ⊢ ⓝU.T ¡[h, o] →
+ ∃∃U0. ⦃G, L⦄ ⊢ U ¡[h, o] & ⦃G, L⦄ ⊢ T ¡[h, o] &
+ ⦃G, L⦄ ⊢ U •*➡*[h, o, 0] U0 & ⦃G, L⦄ ⊢ T •*➡*[h, o, 1] U0.
/2 width=3 by snv_inv_cast_aux/ qed-.
(* Forward lemmas on atomic arity assignment for terms **********************)
-lemma snv_fwd_aaa: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, g] → ∃A. ⦃G, L⦄ ⊢ T ⁝ A.
-#h #g #G #L #T #H elim H -G -L -T
+lemma snv_fwd_aaa: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, o] → ∃A. ⦃G, L⦄ ⊢ T ⁝ A.
+#h #o #G #L #T #H elim H -G -L -T
[ /2 width=2 by aaa_sort, ex_intro/
| #I #G #L #K #V #i #HLK #_ * /3 width=6 by aaa_lref, ex_intro/
| #a * #G #L #V #T #_ #_ * #B #HV * #A #HA /3 width=2 by aaa_abbr, aaa_abst, ex_intro/
(* Advanced forward lemmas **************************************************)
-lemma snv_fwd_da: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, g] → ∃d. ⦃G, L⦄ ⊢ T ▪[h, g] d.
-#h #g #G #L #T #H elim (snv_fwd_aaa … H) -H /2 width=2 by aaa_da/
+lemma snv_fwd_da: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, o] → ∃d. ⦃G, L⦄ ⊢ T ▪[h, o] d.
+#h #o #G #L #T #H elim (snv_fwd_aaa … H) -H /2 width=2 by aaa_da/
qed-.
-lemma snv_fwd_lstas: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, g] →
+lemma snv_fwd_lstas: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, o] →
∀d. ∃U. ⦃G, L⦄ ⊢ T •*[h, d] U.
-#h #g #G #L #T #H #d elim (snv_fwd_aaa … H) -H
+#h #o #G #L #T #H #d elim (snv_fwd_aaa … H) -H
#A #HT elim (aaa_lstas h … HT d) -HT /2 width=2 by ex_intro/
qed-.
(* 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_lstas 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 #_ #_ #_ #_ #d #H2 #X3 #H3 #L2 #_ -IH3 -IH2 -IH1
+fact da_cpr_lpr_aux: ∀h,o,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_snv_lstas h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h o G1 L1 T1) →
+ ∀G1,L1,T1. G0 = G1 → L0 = L1 → T0 = T1 → IH_da_cpr_lpr h o G1 L1 T1.
+#h #o #G0 #L0 #T0 #IH3 #IH2 #IH1 #G1 #L1 * * [|||| * ]
+[ #s #_ #_ #_ #_ #d #H2 #X3 #H3 #L2 #_ -IH3 -IH2 -IH1
lapply (da_inv_sort … H2) -H2
lapply (cpr_inv_sort1 … H3) -H3 #H destruct /2 width=1 by da_sort/
| #i #HG0 #HL0 #HT0 #H1 #d #H2 #X3 #H3 #L2 #HL12 destruct -IH3 -IH2
elim (snv_fwd_da … HW) #d0 #Hd0
lapply (cprs_scpds_div … HW3 … Hd0 … 1 HVW1) -W3 #H
elim (da_scpes_aux … IH3 IH2 IH1 … Hd0 … Hd1 … H) -IH3 -IH2 -H /2 width=1 by fqup_fpbg/ #_ #H1
- <minus_n_O #H destruct >(plus_minus_m_m d1 1) in Hd1; // -H1 #Hd1
+ <minus_n_O #H destruct >(plus_minus_k_k d1 1) in Hd1; // -H1 #Hd1
lapply (IH1 … HV1 … Hd1 … HV12 … HL12) -HV1 -Hd1 -HV12 [ /2 by fqup_fpbg/ ]
lapply (IH1 … Hd0 … HW2 … HL12) -Hd0 /2 width=1 by fqup_fpbg/ -HW
lapply (IH1 … HU … Hd … HU2 (L2.ⓛW2) ?) -IH1 -HU -Hd -HU2 [1,2: /2 by fqup_fpbg, lpr_pair/ ] -HL12 -HW2
(* forward lemmas on "qrst" strongly normalizing closures *********************)
-lemma snv_fwd_fsb: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, g] → ⦥[h, g] ⦃G, L, T⦄.
-#h #g #G #L #T #H elim (snv_fwd_aaa … H) -H /2 width=2 by aaa_fsb/
+lemma snv_fwd_fsb: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, o] → ⦥[h, o] ⦃G, L, T⦄.
+#h #o #G #L #T #H elim (snv_fwd_aaa … H) -H /2 width=2 by aaa_fsb/
qed-.
(* Relocation properties ****************************************************)
-lemma snv_lift: ∀h,g,G,K,T. ⦃G, K⦄ ⊢ T ¡[h, g] → ∀L,s,l,m. ⬇[s, l, m] L ≡ K →
- ∀U. ⬆[l, m] T ≡ U → ⦃G, L⦄ ⊢ U ¡[h, g].
-#h #g #G #K #T #H elim H -G -K -T
-[ #G #K #k #L #s #l #m #_ #X #H
- >(lift_inv_sort1 … H) -X -K -l -m //
-| #I #G #K #K0 #V #i #HK0 #_ #IHV #L #s #l #m #HLK #X #H
+lemma snv_lift: ∀h,o,G,K,T. ⦃G, K⦄ ⊢ T ¡[h, o] → ∀L,c,l,k. ⬇[c, l, k] L ≡ K →
+ ∀U. ⬆[l, k] T ≡ U → ⦃G, L⦄ ⊢ U ¡[h, o].
+#h #o #G #K #T #H elim H -G -K -T
+[ #G #K #s #L #c #l #k #_ #X #H
+ >(lift_inv_sort1 … H) -X -K -l -k //
+| #I #G #K #K0 #V #i #HK0 #_ #IHV #L #c #l #k #HLK #X #H
elim (lift_inv_lref1 … H) * #Hil #H destruct
[ elim (drop_trans_le … HLK … HK0) -K /2 width=2 by ylt_fwd_le/ #X #HL0 #H
elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ -Hil #L0 #W #HLK0 #HVW #H destruct
| lapply (drop_trans_ge … HLK … HK0 ?) -K
/3 width=9 by snv_lref, drop_inv_gen/
]
-| #a #I #G #K #V #T #_ #_ #IHV #IHT #L #s #l #m #HLK #X #H
+| #a #I #G #K #V #T #_ #_ #IHV #IHT #L #c #l #k #HLK #X #H
elim (lift_inv_bind1 … H) -H #W #U #HVW #HTU #H destruct
/4 width=5 by snv_bind, drop_skip/
-| #a #G #K #V #W0 #T #U0 #d #_ #_ #HVW0 #HTU0 #IHV #IHT #L #s #l #m #HLK #X #H
+| #a #G #K #V #W0 #T #U0 #d #_ #_ #HVW0 #HTU0 #IHV #IHT #L #c #l #k #HLK #X #H
elim (lift_inv_flat1 … H) -H #W #U #HVW #HTU #H destruct
- elim (lift_total W0 l m)
- elim (lift_total U0 (l+1) m)
+ elim (lift_total W0 l k)
+ elim (lift_total U0 (l+1) k)
/4 width=17 by snv_appl, scpds_lift, lift_bind/
-| #G #K #V #T #U0 #_ #_ #HVU0 #HTU0 #IHV #IHT #L #s #l #m #HLK #X #H
+| #G #K #V #T #U0 #_ #_ #HVU0 #HTU0 #IHV #IHT #L #c #l #k #HLK #X #H
elim (lift_inv_flat1 … H) -H #W #U #HVW #HTU #H destruct
- elim (lift_total U0 l m)
+ elim (lift_total U0 l k)
/3 width=12 by snv_cast, scpds_lift/
]
qed.
-lemma snv_inv_lift: ∀h,g,G,L,U. ⦃G, L⦄ ⊢ U ¡[h, g] → ∀K,s,l,m. ⬇[s, l, m] L ≡ K →
- ∀T. ⬆[l, m] T ≡ U → ⦃G, K⦄ ⊢ T ¡[h, g].
-#h #g #G #L #U #H elim H -G -L -U
-[ #G #L #k #K #s #l #m #_ #X #H
- >(lift_inv_sort2 … H) -X -L -l -m //
-| #I #G #L #L0 #W #i #HL0 #_ #IHW #K #s #l #m #HLK #X #H
+lemma snv_inv_lift: ∀h,o,G,L,U. ⦃G, L⦄ ⊢ U ¡[h, o] → ∀K,c,l,k. ⬇[c, l, k] L ≡ K →
+ ∀T. ⬆[l, k] T ≡ U → ⦃G, K⦄ ⊢ T ¡[h, o].
+#h #o #G #L #U #H elim H -G -L -U
+[ #G #L #s #K #c #l #k #_ #X #H
+ >(lift_inv_sort2 … H) -X -L -l -k //
+| #I #G #L #L0 #W #i #HL0 #_ #IHW #K #c #l #k #HLK #X #H
elim (lift_inv_lref2 … H) * #Hil #H destruct
[ elim (drop_conf_le … HLK … HL0) -L /2 width=2 by ylt_fwd_le/ #X #HK0 #H
elim (drop_inv_skip1 … H) -H /2 width=1 by ylt_to_minus/ -Hil #K0 #V #HLK0 #HVW #H destruct
/3 width=12 by snv_lref/
| lapply (drop_conf_ge … HLK … HL0 ?) -L /3 width=9 by snv_lref/
]
-| #a #I #G #L #W #U #_ #_ #IHW #IHU #K #s #l #m #HLK #X #H
+| #a #I #G #L #W #U #_ #_ #IHW #IHU #K #c #l #k #HLK #X #H
elim (lift_inv_bind2 … H) -H #V #T #HVW #HTU #H destruct
/4 width=5 by snv_bind, drop_skip/
-| #a #G #L #W #W1 #U #U1 #d #_ #_ #HW1 #HU1 #IHW #IHU #K #s #l #m #HLK #X #H
+| #a #G #L #W #W1 #U #U1 #d #_ #_ #HW1 #HU1 #IHW #IHU #K #c #l #k #HLK #X #H
elim (lift_inv_flat2 … H) -H #V #T #HVW #HTU #H destruct
elim (scpds_inv_lift1 … HW1 … HLK … HVW) -HW1 #W0 #HW01 #HVW0
elim (scpds_inv_lift1 … HU1 … HLK … HTU) -HU1 #X #H #HTU0
elim (lift_inv_bind2 … H) -H #Y #U0 #HY #HU01 #H destruct
lapply (lift_inj … HY … HW01) -HY #H destruct
/3 width=6 by snv_appl/
-| #G #L #W #U #U1 #_ #_ #HWU1 #HU1 #IHW #IHU #K #s #l #m #HLK #X #H
+| #G #L #W #U #U1 #_ #_ #HWU1 #HU1 #IHW #IHU #K #c #l #k #HLK #X #H
elim (lift_inv_flat2 … H) -H #V #T #HVW #HTU #H destruct
elim (scpds_inv_lift1 … HWU1 … HLK … HVW) -HWU1 #U0 #HU01 #HVU0
elim (scpds_inv_lift1 … HU1 … HLK … HTU) -HU1 #X #HX #HTU0
(* Properties on subclosure *************************************************)
-lemma snv_fqu_conf: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ⦃G1, L1⦄ ⊢ T1 ¡[h, g] → ⦃G2, L2⦄ ⊢ T2 ¡[h, g].
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
+lemma snv_fqu_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ⦃G1, L1⦄ ⊢ T1 ¡[h, o] → ⦃G2, L2⦄ ⊢ T2 ¡[h, o].
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
[ #I1 #G1 #L1 #V1 #H
elim (snv_inv_lref … H) -H #I2 #L2 #V2 #H #HV2
lapply (drop_inv_O2 … H) -H #H destruct //
]
qed-.
-lemma snv_fquq_conf: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ⦃G1, L1⦄ ⊢ T1 ¡[h, g] → ⦃G2, L2⦄ ⊢ T2 ¡[h, g].
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H elim (fquq_inv_gen … H) -H [|*]
+lemma snv_fquq_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ⦃G1, L1⦄ ⊢ T1 ¡[h, o] → ⦃G2, L2⦄ ⊢ T2 ¡[h, o].
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim (fquq_inv_gen … H) -H [|*]
/2 width=5 by snv_fqu_conf/
qed-.
-lemma snv_fqup_conf: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ⦃G1, L1⦄ ⊢ T1 ¡[h, g] → ⦃G2, L2⦄ ⊢ T2 ¡[h, g].
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
+lemma snv_fqup_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
+ ⦃G1, L1⦄ ⊢ T1 ¡[h, o] → ⦃G2, L2⦄ ⊢ T2 ¡[h, o].
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
/3 width=5 by fqup_strap1, snv_fqu_conf/
qed-.
-lemma snv_fqus_conf: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
- ⦃G1, L1⦄ ⊢ T1 ¡[h, g] → ⦃G2, L2⦄ ⊢ T2 ¡[h, g].
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H elim (fqus_inv_gen … H) -H [|*]
+lemma snv_fqus_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+ ⦃G1, L1⦄ ⊢ T1 ¡[h, o] → ⦃G2, L2⦄ ⊢ T2 ¡[h, o].
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim (fqus_inv_gen … H) -H [|*]
/2 width=5 by snv_fqup_conf/
qed-.
(* 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_lstas h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G1, L1, T1⦄ → IH_lstas_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 * * [|||| * ]
-[ #k #HG0 #HL0 #HT0 #H1 #X #H2 #L2 #_ destruct -IH4 -IH3 -IH2 -IH1 -H1
+fact snv_cpr_lpr_aux: ∀h,o,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_snv_lstas h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_lstas_cpr_lpr h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h o G1 L1 T1) →
+ ∀G1,L1,T1. G0 = G1 → L0 = L1 → T0 = T1 → IH_snv_cpr_lpr h o G1 L1 T1.
+#h #o #G0 #L0 #T0 #IH4 #IH3 #IH2 #IH1 #G1 #L1 * * [|||| * ]
+[ #s #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 (snv_inv_lref … H1) -H1 #I #K1 #V1 #HLK1 #HV1
elim (snv_fwd_da … HW10) #d0 #HW10d
lapply (cprs_scpds_div … HW130 … HW10d … 1 HVW1) -W30 #HVW10
elim (da_scpes_aux … IH4 IH1 IH2 … HW10d … HV1d … HVW10) /2 width=1 by fqup_fpbg/
- #_ #Hd <minus_n_O #H destruct >(plus_minus_m_m d 1) in HV1d; // -Hd #HV1d
+ #_ #Hd <minus_n_O #H destruct >(plus_minus_k_k d 1) in HV1d; // -Hd #HV1d
lapply (scpes_cpr_lpr_aux … IH2 IH3 … HVW10 … HW120 … HV12 … HL12) /2 width=1 by fqup_fpbg/ -HVW10 #HVW20
lapply (IH2 … HV1d … HV12 … HL12) /2 width=1 by fqup_fpbg/ -HV1d #HV2d
lapply (IH2 … HW10d … HW120 … HL12) /2 width=1 by fqup_fpbg/ -HW10d #HW20d
(* Properties on nat-iterated stratified static type assignment for terms ***)
-fact snv_lstas_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_lstas_cpr_lpr h g G1 L1 T1) →
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G1, L1, T1⦄ → IH_snv_lstas h g G1 L1 T1) →
- ∀G1,L1,T1. G0 = G1 → L0 = L1 → T0 = T1 → IH_snv_lstas h g G1 L1 T1.
-#h #g #G0 #L0 #T0 #IH4 #IH3 #IH2 #IH1 #G1 #L1 * * [|||| * ]
-[ #k #HG0 #HL0 #HT0 #_ #d1 #d2 #Hd21 #Hd1 #X #H2 destruct -IH4 -IH3 -IH2 -IH1
+fact snv_lstas_aux: ∀h,o,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_lstas_cpr_lpr h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_snv_lstas h o G1 L1 T1) →
+ ∀G1,L1,T1. G0 = G1 → L0 = L1 → T0 = T1 → IH_snv_lstas h o G1 L1 T1.
+#h #o #G0 #L0 #T0 #IH4 #IH3 #IH2 #IH1 #G1 #L1 * * [|||| * ]
+[ #s #HG0 #HL0 #HT0 #_ #d1 #d2 #Hd21 #Hd1 #X #H2 destruct -IH4 -IH3 -IH2 -IH1
>(lstas_inv_sort1 … H2) -X //
| #i #HG0 #HL0 #HT0 #H1 #d1 #d2 #Hd21 #Hd1 #T #H2 destruct -IH4 -IH3 -IH2
elim (snv_inv_lref … H1) -H1 #I0 #K0 #X0 #HLK0 #HX0
(* Properties on sn parallel reduction for local environments ***************)
-fact lstas_cpr_lpr_aux: ∀h,g,G0,L0,T0.
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G1, L1, T1⦄ → IH_snv_lstas 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_lstas_cpr_lpr h g G1 L1 T1) →
- ∀G1,L1,T1. G0 = G1 → L0 = L1 → T0 = T1 → IH_lstas_cpr_lpr h g G1 L1 T1.
-#h #g #G0 #L0 #T0 #IH4 #IH3 #IH2 #IH1 #G1 #L1 * * [|||| * ]
-[ #k #_ #_ #_ #_ #d1 #d2 #_ #_ #X2 #H2 #X3 #H3 #L2 #_ -IH4 -IH3 -IH2 -IH1
+fact lstas_cpr_lpr_aux: ∀h,o,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_snv_lstas h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_lstas_cpr_lpr h o G1 L1 T1) →
+ ∀G1,L1,T1. G0 = G1 → L0 = L1 → T0 = T1 → IH_lstas_cpr_lpr h o G1 L1 T1.
+#h #o #G0 #L0 #T0 #IH4 #IH3 #IH2 #IH1 #G1 #L1 * * [|||| * ]
+[ #s #_ #_ #_ #_ #d1 #d2 #_ #_ #X2 #H2 #X3 #H3 #L2 #_ -IH4 -IH3 -IH2 -IH1
>(lstas_inv_sort1 … H2) -X2
>(cpr_inv_sort1 … H3) -X3 /2 width=3 by ex2_intro/
| #i #HG0 #HL0 #HT0 #H1 #d1 #d2 #Hd21 #Hd1 #X2 #H2 #X3 #H3 #L2 #HL12 destruct -IH4 -IH3
lapply (cprs_scpds_div … HW20 … HW2d … HVW1) -W0 #H21
elim (snv_fwd_da … HV1) #d #HV1d
elim (da_scpes_aux … IH4 IH3 IH2 … HW2d … HV1d … H21) /2 width=1 by fqup_fpbg/ #_ #H
- <minus_n_O #H0 destruct >(plus_minus_m_m d 1) in HV1d; // -H #HV1d
+ <minus_n_O #H0 destruct >(plus_minus_k_k d 1) in HV1d; // -H #HV1d
lapply (scpes_cpr_lpr_aux … IH2 IH1 … H21 … HW23 … HV12 … HL12) -H21 /2 width=1 by fqup_fpbg/ #H32
lapply (IH3 … HW23 … HL12) /2 width=1 by fqup_fpbg/ #HW3
lapply (IH3 … HV12 … HL12) /2 width=1 by fqup_fpbg/ #HV2
lapply (IH2 … HW2d … HW23 … HL12) /2 width=1 by fqup_fpbg/ -HW2 -HW2d #HW3d
lapply (IH2 … HV1d … HV12 … HL12) /2 width=1 by fqup_fpbg/ -HV1 -HV1d #HV2d
elim (IH1 … Hd1 … HT2U … HT23 (L2.ⓛW3)) -HT2U /2 width=1 by fqup_fpbg, lpr_pair/ #U3 #HTU3 #HU23
- elim (lsubsv_lstas_trans … g … HTU3 … Hd21 … (L2.ⓓⓝW3.V2)) -HTU3
+ elim (lsubsv_lstas_trans … o … HTU3 … Hd21 … (L2.ⓓⓝW3.V2)) -HTU3
[ #U4 #HT3U4 #HU43 -IH1 -IH2 -IH3 -IH4 -d -d1 -HW3 -HV2 -HT2
@(ex2_intro … (ⓓ{b}ⓝW3.V2.U4)) /2 width=1 by lstas_bind/ -HT3U4
@(cpcs_canc_dx … (ⓓ{b}ⓝW3.V2.U3)) /2 width=1 by cpcs_bind_dx/ -HU43
(* Main preservation properties *********************************************)
-lemma snv_preserve: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, g] →
- ∧∧ IH_da_cpr_lpr h g G L T
- & IH_snv_cpr_lpr h g G L T
- & IH_snv_lstas h g G L T
- & IH_lstas_cpr_lpr h g G L T.
-#h #g #G #L #T #HT elim (snv_fwd_aaa … HT) -HT
-#A #HT @(aaa_ind_fpbg h g … HT) -G -L -T -A
+lemma snv_preserve: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, o] →
+ ∧∧ IH_da_cpr_lpr h o G L T
+ & IH_snv_cpr_lpr h o G L T
+ & IH_snv_lstas h o G L T
+ & IH_lstas_cpr_lpr h o G L T.
+#h #o #G #L #T #HT elim (snv_fwd_aaa … HT) -HT
+#A #HT @(aaa_ind_fpbg h o … HT) -G -L -T -A
#G #L #T #A #_ #IH -A @and4_intro
[ letin aux ≝ da_cpr_lpr_aux | letin aux ≝ snv_cpr_lpr_aux
| letin aux ≝ snv_lstas_aux | letin aux ≝ lstas_cpr_lpr_aux
@(aux … G L T) // #G0 #L0 #T0 #H elim (IH … H) -IH -H //
qed-.
-theorem da_cpr_lpr: ∀h,g,G,L,T. IH_da_cpr_lpr h g G L T.
-#h #g #G #L #T #HT elim (snv_preserve … HT) /2 width=1 by/
+theorem da_cpr_lpr: ∀h,o,G,L,T. IH_da_cpr_lpr h o G L T.
+#h #o #G #L #T #HT elim (snv_preserve … HT) /2 width=1 by/
qed-.
-theorem snv_cpr_lpr: ∀h,g,G,L,T. IH_snv_cpr_lpr h g G L T.
-#h #g #G #L #T #HT elim (snv_preserve … HT) /2 width=1 by/
+theorem snv_cpr_lpr: ∀h,o,G,L,T. IH_snv_cpr_lpr h o G L T.
+#h #o #G #L #T #HT elim (snv_preserve … HT) /2 width=1 by/
qed-.
-theorem snv_lstas: ∀h,g,G,L,T. IH_snv_lstas h g G L T.
-#h #g #G #L #T #HT elim (snv_preserve … HT) /2 width=5 by/
+theorem snv_lstas: ∀h,o,G,L,T. IH_snv_lstas h o G L T.
+#h #o #G #L #T #HT elim (snv_preserve … HT) /2 width=5 by/
qed-.
-theorem lstas_cpr_lpr: ∀h,g,G,L,T. IH_lstas_cpr_lpr h g G L T.
-#h #g #G #L #T #HT elim (snv_preserve … HT) /2 width=3 by/
+theorem lstas_cpr_lpr: ∀h,o,G,L,T. IH_lstas_cpr_lpr h o G L T.
+#h #o #G #L #T #HT elim (snv_preserve … HT) /2 width=3 by/
qed-.
(* Advanced preservation properties *****************************************)
-lemma snv_cprs_lpr: ∀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 #G #L1 #T1 #HT1 #T2 #H
+lemma snv_cprs_lpr: ∀h,o,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, o] →
+ ∀T2. ⦃G, L1⦄ ⊢ T1 ➡* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L2⦄ ⊢ T2 ¡[h, o].
+#h #o #G #L1 #T1 #HT1 #T2 #H
@(cprs_ind … H) -T2 /3 width=5 by snv_cpr_lpr/
qed-.
-lemma da_cprs_lpr: ∀h,g,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀d. ⦃G, L1⦄ ⊢ T1 ▪[h, g] d →
- ∀T2. ⦃G, L1⦄ ⊢ T1 ➡* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L2⦄ ⊢ T2 ▪[h, g] d.
-#h #g #G #L1 #T1 #HT1 #d #Hd #T2 #H
+lemma da_cprs_lpr: ∀h,o,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, o] →
+ ∀d. ⦃G, L1⦄ ⊢ T1 ▪[h, o] d →
+ ∀T2. ⦃G, L1⦄ ⊢ T1 ➡* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L2⦄ ⊢ T2 ▪[h, o] d.
+#h #o #G #L1 #T1 #HT1 #d #Hd #T2 #H
@(cprs_ind … H) -T2 /3 width=6 by snv_cprs_lpr, da_cpr_lpr/
qed-.
-lemma lstas_cprs_lpr: ∀h,g,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀d1,d2. d2 ≤ d1 → ⦃G, L1⦄ ⊢ T1 ▪[h, g] d1 →
+lemma lstas_cprs_lpr: ∀h,o,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, o] →
+ ∀d1,d2. d2 ≤ d1 → ⦃G, L1⦄ ⊢ T1 ▪[h, o] d1 →
∀U1. ⦃G, L1⦄ ⊢ T1 •*[h, d2] U1 →
∀T2. ⦃G, L1⦄ ⊢ T1 ➡* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
∃∃U2. ⦃G, L2⦄ ⊢ T2 •*[h, d2] U2 & ⦃G, L2⦄ ⊢ U1 ⬌* U2.
-#h #g #G #L1 #T1 #HT1 #d1 #d2 #Hd21 #Hd1 #U1 #HTU1 #T2 #H
+#h #o #G #L1 #T1 #HT1 #d1 #d2 #Hd21 #Hd1 #U1 #HTU1 #T2 #H
@(cprs_ind … H) -T2 [ /2 width=10 by lstas_cpr_lpr/ ]
#T #T2 #HT1T #HTT2 #IHT1 #L2 #HL12
elim (IHT1 L1) // -IHT1 #U #HTU #HU1
-elim (lstas_cpr_lpr … g … Hd21 … HTU … HTT2 … HL12) -HTU -HTT2
+elim (lstas_cpr_lpr … o … Hd21 … HTU … HTT2 … HL12) -HTU -HTT2
[2,3: /2 width=7 by snv_cprs_lpr, da_cprs_lpr/ ] -T1 -T -d1
/4 width=5 by lpr_cpcs_conf, cpcs_trans, ex2_intro/
qed-.
-lemma lstas_cpcs_lpr: ∀h,g,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀d,d1. d ≤ d1 → ⦃G, L1⦄ ⊢ T1 ▪[h, g] d1 → ∀U1. ⦃G, L1⦄ ⊢ T1 •*[h, d] U1 →
- ∀T2. ⦃G, L1⦄ ⊢ T2 ¡[h, g] →
- ∀d2. d ≤ d2 → ⦃G, L1⦄ ⊢ T2 ▪[h, g] d2 → ∀U2. ⦃G, L1⦄ ⊢ T2 •*[h, d] U2 →
+lemma lstas_cpcs_lpr: ∀h,o,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, o] →
+ ∀d,d1. d ≤ d1 → ⦃G, L1⦄ ⊢ T1 ▪[h, o] d1 → ∀U1. ⦃G, L1⦄ ⊢ T1 •*[h, d] U1 →
+ ∀T2. ⦃G, L1⦄ ⊢ T2 ¡[h, o] →
+ ∀d2. d ≤ d2 → ⦃G, L1⦄ ⊢ T2 ▪[h, o] d2 → ∀U2. ⦃G, L1⦄ ⊢ T2 •*[h, d] U2 →
⦃G, L1⦄ ⊢ T1 ⬌* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L2⦄ ⊢ U1 ⬌* U2.
-#h #g #G #L1 #T1 #HT1 #d #d1 #Hd1 #HTd1 #U1 #HTU1 #T2 #HT2 #d2 #Hd2 #HTd2 #U2 #HTU2 #H #L2 #HL12
+#h #o #G #L1 #T1 #HT1 #d #d1 #Hd1 #HTd1 #U1 #HTU1 #T2 #HT2 #d2 #Hd2 #HTd2 #U2 #HTU2 #H #L2 #HL12
elim (cpcs_inv_cprs … H) -H #T #H1 #H2
elim (lstas_cprs_lpr … HT1 … Hd1 HTd1 … HTU1 … H1 … HL12) -T1 #W1 #H1 #HUW1
elim (lstas_cprs_lpr … HT2 … Hd2 HTd2 … HTU2 … H2 … HL12) -T2 #W2 #H2 #HUW2
(* Inductive premises for the preservation results **************************)
definition IH_snv_cpr_lpr: ∀h:sh. sd h → relation3 genv lenv term ≝
- λ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,o,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, o] →
+ ∀T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L2⦄ ⊢ T2 ¡[h, o].
definition IH_da_cpr_lpr: ∀h:sh. sd h → relation3 genv lenv term ≝
- λh,g,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀d. ⦃G, L1⦄ ⊢ T1 ▪[h, g] d →
+ λh,o,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, o] →
+ ∀d. ⦃G, L1⦄ ⊢ T1 ▪[h, o] d →
∀T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
- ⦃G, L2⦄ ⊢ T2 ▪[h, g] d.
+ ⦃G, L2⦄ ⊢ T2 ▪[h, o] d.
definition IH_lstas_cpr_lpr: ∀h:sh. sd h → relation3 genv lenv term ≝
- λh,g,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀d1,d2. d2 ≤ d1 → ⦃G, L1⦄ ⊢ T1 ▪[h, g] d1 →
+ λh,o,G,L1,T1. ⦃G, L1⦄ ⊢ T1 ¡[h, o] →
+ ∀d1,d2. d2 ≤ d1 → ⦃G, L1⦄ ⊢ T1 ▪[h, o] d1 →
∀U1. ⦃G, L1⦄ ⊢ T1 •*[h, d2] U1 →
∀T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
∃∃U2. ⦃G, L2⦄ ⊢ T2 •*[h, d2] U2 & ⦃G, L2⦄ ⊢ U1 ⬌* U2.
definition IH_snv_lstas: ∀h:sh. sd h → relation3 genv lenv term ≝
- λh,g,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, g] →
- ∀d1,d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ T ▪[h, g] d1 →
- ∀U. ⦃G, L⦄ ⊢ T •*[h, d2] U → ⦃G, L⦄ ⊢ U ¡[h, g].
+ λh,o,G,L,T. ⦃G, L⦄ ⊢ T ¡[h, o] →
+ ∀d1,d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ T ▪[h, o] d1 →
+ ∀U. ⦃G, L⦄ ⊢ T •*[h, d2] U → ⦃G, L⦄ ⊢ U ¡[h, o].
(* Properties for the preservation results **********************************)
-fact snv_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) →
- ∀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
+fact snv_cprs_lpr_aux: ∀h,o,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h o G1 L1 T1) →
+ ∀G,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, o] →
+ ∀T2. ⦃G, L1⦄ ⊢ T1 ➡* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L2⦄ ⊢ T2 ¡[h, o].
+#h #o #G0 #L0 #T0 #IH #G #L1 #T1 #HLT0 #HT1 #T2 #H
@(cprs_ind … H) -T2 /4 width=6 by fpbg_fpbs_trans, cprs_fpbs/
qed-.
-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] →
- ∀d. ⦃G, L1⦄ ⊢ T1 ▪[h, g] d →
- ∀T2. ⦃G, L1⦄ ⊢ T1 ➡* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L2⦄ ⊢ T2 ▪[h, g] d.
-#h #g #G0 #L0 #T0 #IH2 #IH1 #G #L1 #T1 #HLT0 #HT1 #d #Hd #T2 #H
+fact da_cprs_lpr_aux: ∀h,o,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h o G1 L1 T1) →
+ ∀G,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, o] →
+ ∀d. ⦃G, L1⦄ ⊢ T1 ▪[h, o] d →
+ ∀T2. ⦃G, L1⦄ ⊢ T1 ➡* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L2⦄ ⊢ T2 ▪[h, o] d.
+#h #o #G0 #L0 #T0 #IH2 #IH1 #G #L1 #T1 #HLT0 #HT1 #d #Hd #T2 #H
@(cprs_ind … H) -T2 /4 width=10 by snv_cprs_lpr_aux, fpbg_fpbs_trans, cprs_fpbs/
qed-.
-fact da_scpds_lpr_aux: ∀h,g,G0,L0,T0.
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G1, L1, T1⦄ → IH_snv_lstas 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) →
- ∀G,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀d1. ⦃G, L1⦄ ⊢ T1 ▪[h, g] d1 →
- ∀T2,d2. ⦃G, L1⦄ ⊢ T1 •*➡*[h, g, d2] T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
- d2 ≤ d1 ∧ ⦃G, L2⦄ ⊢ T2 ▪[h, g] d1-d2.
-#h #g #G0 #L0 #T0 #IH3 #IH2 #IH1 #G #L1 #T1 #HLT0 #HT1 #d1 #Hd1 #T2 #d2 * #T #d0 #Hd20 #H #HT1 #HT2 #L2 #HL12
+fact da_scpds_lpr_aux: ∀h,o,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_snv_lstas h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h o G1 L1 T1) →
+ ∀G,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, o] →
+ ∀d1. ⦃G, L1⦄ ⊢ T1 ▪[h, o] d1 →
+ ∀T2,d2. ⦃G, L1⦄ ⊢ T1 •*➡*[h, o, d2] T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
+ d2 ≤ d1 ∧ ⦃G, L2⦄ ⊢ T2 ▪[h, o] d1-d2.
+#h #o #G0 #L0 #T0 #IH3 #IH2 #IH1 #G #L1 #T1 #HLT0 #HT1 #d1 #Hd1 #T2 #d2 * #T #d0 #Hd20 #H #HT1 #HT2 #L2 #HL12
lapply (da_mono … H … Hd1) -H #H destruct
lapply (lstas_da_conf … HT1 … Hd1) #Hd12
lapply (da_cprs_lpr_aux … IH2 IH1 … Hd12 … HT2 … HL12) -IH2 -IH1 -HT2 -HL12
/3 width=8 by fpbg_fpbs_trans, lstas_fpbs, conj/
qed-.
-fact da_scpes_aux: ∀h,g,G0,L0,T0.
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G1, L1, T1⦄ → IH_snv_lstas 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) →
- ∀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] →
- ∀d11. ⦃G, L⦄ ⊢ T1 ▪[h, g] d11 → ∀d12. ⦃G, L⦄ ⊢ T2 ▪[h, g] d12 →
- ∀d21,d22. ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d21, d22] T2 →
+fact da_scpes_aux: ∀h,o,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_snv_lstas h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h o G1 L1 T1) →
+ ∀G,L,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G, L, T1⦄ → ⦃G, L⦄ ⊢ T1 ¡[h, o] →
+ ∀T2. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G, L, T2⦄ → ⦃G, L⦄ ⊢ T2 ¡[h, o] →
+ ∀d11. ⦃G, L⦄ ⊢ T1 ▪[h, o] d11 → ∀d12. ⦃G, L⦄ ⊢ T2 ▪[h, o] d12 →
+ ∀d21,d22. ⦃G, L⦄ ⊢ T1 •*⬌*[h, o, d21, d22] T2 →
∧∧ d21 ≤ d11 & d22 ≤ d12 & d11 - d21 = d12 - d22.
-#h #g #G0 #L0 #T0 #IH3 #IH2 #IH1 #G #L #T1 #HLT01 #HT1 #T2 #HLT02 #HT2 #d11 #Hd11 #d12 #Hd12 #d21 #d22 * #T #HT1 #HT2
+#h #o #G0 #L0 #T0 #IH3 #IH2 #IH1 #G #L #T1 #HLT01 #HT1 #T2 #HLT02 #HT2 #d11 #Hd11 #d12 #Hd12 #d21 #d22 * #T #HT1 #HT2
elim (da_scpds_lpr_aux … IH3 IH2 IH1 … Hd11 … HT1 … L) -Hd11 -HT1 //
elim (da_scpds_lpr_aux … IH3 IH2 IH1 … Hd12 … HT2 … L) -Hd12 -HT2 //
/3 width=7 by da_mono, and3_intro/
qed-.
-fact lstas_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_lstas_cpr_lpr h g G1 L1 T1) →
- ∀G,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀d1,d2. d2 ≤ d1 → ⦃G, L1⦄ ⊢ T1 ▪[h, g] d1 →
+fact lstas_cprs_lpr_aux: ∀h,o,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_lstas_cpr_lpr h o G1 L1 T1) →
+ ∀G,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, o] →
+ ∀d1,d2. d2 ≤ d1 → ⦃G, L1⦄ ⊢ T1 ▪[h, o] d1 →
∀U1. ⦃G, L1⦄ ⊢ T1 •*[h, d2] U1 →
∀T2. ⦃G, L1⦄ ⊢ T1 ➡* T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
∃∃U2. ⦃G, L2⦄ ⊢ T2 •*[h, d2] U2 & ⦃G, L2⦄ ⊢ U1 ⬌* U2.
-#h #g #G0 #L0 #T0 #IH3 #IH2 #IH1 #G #L1 #T1 #H01 #HT1 #d1 #d2 #Hd21 #Hd1 #U1 #HTU1 #T2 #H
+#h #o #G0 #L0 #T0 #IH3 #IH2 #IH1 #G #L1 #T1 #H01 #HT1 #d1 #d2 #Hd21 #Hd1 #U1 #HTU1 #T2 #H
@(cprs_ind … H) -T2 [ /2 width=10 by/ ]
#T #T2 #HT1T #HTT2 #IHT1 #L2 #HL12
elim (IHT1 L1) // -IHT1 #U #HTU #HU1
/4 width=5 by lpr_cpcs_conf, cpcs_trans, ex2_intro/
qed-.
-fact scpds_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_lstas_cpr_lpr h g G1 L1 T1) →
- ∀G,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀U1,d. ⦃G, L1⦄ ⊢ T1 •*➡*[h, g, d] U1 →
+fact scpds_cpr_lpr_aux: ∀h,o,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_lstas_cpr_lpr h o G1 L1 T1) →
+ ∀G,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, o] →
+ ∀U1,d. ⦃G, L1⦄ ⊢ T1 •*➡*[h, o, d] U1 →
∀T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
- ∃∃U2. ⦃G, L2⦄ ⊢ T2 •*➡*[h, g, d] U2 & ⦃G, L2⦄ ⊢ U1 ➡* U2.
-#h #g #G0 #L0 #T0 #IH2 #IH1 #G #L1 #T1 #H01 #HT1 #U1 #d2 * #W1 #d1 #Hd21 #HTd1 #HTW1 #HWU1 #T2 #HT12 #L2 #HL12
+ ∃∃U2. ⦃G, L2⦄ ⊢ T2 •*➡*[h, o, d] U2 & ⦃G, L2⦄ ⊢ U1 ➡* U2.
+#h #o #G0 #L0 #T0 #IH2 #IH1 #G #L1 #T1 #H01 #HT1 #U1 #d2 * #W1 #d1 #Hd21 #HTd1 #HTW1 #HWU1 #T2 #HT12 #L2 #HL12
elim (IH1 … H01 … HTW1 … HT12 … HL12) -IH1 // #W2 #HTW2 #HW12
lapply (IH2 … H01 … HTd1 … HT12 … HL12) -L0 -T0 // -T1
lapply (lpr_cprs_conf … HL12 … HWU1) -L1 #HWU1
elim (cpcs_inv_cprs … H) -H /3 width=6 by ex4_2_intro, ex2_intro/
qed-.
-fact scpes_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_lstas_cpr_lpr h g G1 L1 T1) →
- ∀G,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, g] →
- ∀T2. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G, L1, T2⦄ → ⦃G, L1⦄ ⊢ T2 ¡[h, g] →
- ∀d1,d2. ⦃G, L1⦄ ⊢ T1 •*⬌*[h, g, d1, d2] T2 →
+fact scpes_cpr_lpr_aux: ∀h,o,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_lstas_cpr_lpr h o G1 L1 T1) →
+ ∀G,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G, L1, T1⦄ → ⦃G, L1⦄ ⊢ T1 ¡[h, o] →
+ ∀T2. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G, L1, T2⦄ → ⦃G, L1⦄ ⊢ T2 ¡[h, o] →
+ ∀d1,d2. ⦃G, L1⦄ ⊢ T1 •*⬌*[h, o, d1, d2] T2 →
∀U1. ⦃G, L1⦄ ⊢ T1 ➡ U1 → ∀U2. ⦃G, L1⦄ ⊢ T2 ➡ U2 → ∀L2. ⦃G, L1⦄ ⊢ ➡ L2 →
- ⦃G, L2⦄ ⊢ U1 •*⬌*[h, g, d1, d2] U2.
-#h #g #G0 #L0 #T0 #IH2 #IH1 #G #L1 #T1 #H01 #HT1 #T2 #HT02 #HT2 #d1 #d2 * #T0 #HT10 #HT20 #U1 #HTU1 #U2 #HTU2 #L2 #HL12
+ ⦃G, L2⦄ ⊢ U1 •*⬌*[h, o, d1, d2] U2.
+#h #o #G0 #L0 #T0 #IH2 #IH1 #G #L1 #T1 #H01 #HT1 #T2 #HT02 #HT2 #d1 #d2 * #T0 #HT10 #HT20 #U1 #HTU1 #U2 #HTU2 #L2 #HL12
elim (scpds_cpr_lpr_aux … IH2 IH1 … HT10 … HTU1 … HL12) -HT10 -HTU1 // #X1 #HUX1 #H1
elim (scpds_cpr_lpr_aux … IH2 IH1 … HT20 … HTU2 … HL12) -HT20 -HTU2 // #X2 #HUX2 #H2
elim (cprs_conf … H1 … H2) -T0
/3 width=5 by scpds_div, scpds_cprs_trans/
qed-.
-fact lstas_scpds_aux: ∀h,g,G0,L0,T0.
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G1, L1, T1⦄ → IH_snv_lstas 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_lstas_cpr_lpr h g G1 L1 T1) →
- ∀G,L,T. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G, L, T⦄ → ⦃G, L⦄ ⊢ T ¡[h, g] →
- ∀d,d1. d1 ≤ d → ⦃G, L⦄ ⊢ T ▪[h, g] d → ∀T1. ⦃G, L⦄ ⊢ T •*[h, d1] T1 →
- ∀T2,d2. ⦃G, L⦄ ⊢ T •*➡*[h, g, d2] T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d2-d1, d1-d2] T2.
-#h #g #G0 #L0 #T0 #IH4 #IH3 #IH2 #IH1 #G #L #T #H0 #HT #d #d1 #Hd1 #HTd #T1 #HT1 #T2 #d2 * #X #d0 #Hd20 #H #HTX #HXT2
+fact lstas_scpds_aux: ∀h,o,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_snv_lstas h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_lstas_cpr_lpr h o G1 L1 T1) →
+ ∀G,L,T. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G, L, T⦄ → ⦃G, L⦄ ⊢ T ¡[h, o] →
+ ∀d,d1. d1 ≤ d → ⦃G, L⦄ ⊢ T ▪[h, o] d → ∀T1. ⦃G, L⦄ ⊢ T •*[h, d1] T1 →
+ ∀T2,d2. ⦃G, L⦄ ⊢ T •*➡*[h, o, d2] T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, o, d2-d1, d1-d2] T2.
+#h #o #G0 #L0 #T0 #IH4 #IH3 #IH2 #IH1 #G #L #T #H0 #HT #d #d1 #Hd1 #HTd #T1 #HT1 #T2 #d2 * #X #d0 #Hd20 #H #HTX #HXT2
lapply (da_mono … H … HTd) -H #H destruct
lapply (lstas_da_conf … HT1 … HTd) #HTd1
lapply (lstas_da_conf … HTX … HTd) #HXd
]
qed-.
-fact scpes_le_aux: ∀h,g,G0,L0,T0.
- (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, g] ⦃G1, L1, T1⦄ → IH_snv_lstas 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_lstas_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] →
- ∀d11. ⦃G, L⦄ ⊢ T1 ▪[h, g] d11 → ∀d12. ⦃G, L⦄ ⊢ T2 ▪[h, g] d12 →
+fact scpes_le_aux: ∀h,o,G0,L0,T0.
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_snv_lstas h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_snv_cpr_lpr h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_da_cpr_lpr h o G1 L1 T1) →
+ (∀G1,L1,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G1, L1, T1⦄ → IH_lstas_cpr_lpr h o G1 L1 T1) →
+ ∀G,L,T1. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G, L, T1⦄ → ⦃G, L⦄ ⊢ T1 ¡[h, o] →
+ ∀T2. ⦃G0, L0, T0⦄ >≡[h, o] ⦃G, L, T2⦄ → ⦃G, L⦄ ⊢ T2 ¡[h, o] →
+ ∀d11. ⦃G, L⦄ ⊢ T1 ▪[h, o] d11 → ∀d12. ⦃G, L⦄ ⊢ T2 ▪[h, o] d12 →
∀d21,d22,d. d21 + d ≤ d11 → d22 + d ≤ d12 →
- ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d21, d22] T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d21+d, d22+d] T2.
-#h #g #G0 #L0 #T0 #IH4 #IH3 #IH2 #IH1 #G #L #T1 #H01 #HT1 #T2 #H02 #HT2 #d11 #Hd11 #Hd12 #Hd12 #d21 #d22 #d #H1 #H2 * #T0 #HT10 #HT20
+ ⦃G, L⦄ ⊢ T1 •*⬌*[h, o, d21, d22] T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, o, d21+d, d22+d] T2.
+#h #o #G0 #L0 #T0 #IH4 #IH3 #IH2 #IH1 #G #L #T1 #H01 #HT1 #T2 #H02 #HT2 #d11 #Hd11 #Hd12 #Hd12 #d21 #d22 #d #H1 #H2 * #T0 #HT10 #HT20
elim (da_lstas … Hd11 (d21+d)) #X1 #HTX1 #_
elim (da_lstas … Hd12 (d22+d)) #X2 #HTX2 #_
lapply (lstas_scpds_aux … IH4 IH3 IH2 IH1 … Hd11 … HTX1 … HT10) -HT10
-[1,2,3: // | >eq_minus_O [2: // ] <minus_plus_m_m_commutative #HX1 ]
+[1,2,3: // | >eq_minus_O [2: // ] <minus_plus_k_k_commutative #HX1 ]
lapply (lstas_scpds_aux … IH4 IH3 IH2 IH1 … Hd12 … HTX2 … HT20) -HT20
-[1,2,3: // | >eq_minus_O [2: // ] <minus_plus_m_m_commutative #HX2 ]
+[1,2,3: // | >eq_minus_O [2: // ] <minus_plus_k_k_commutative #HX2 ]
lapply (lstas_scpes_trans … Hd11 … HTX1 … HX1) [ // ] -Hd11 -HTX1 -HX1 -H1 #H1
lapply (lstas_scpes_trans … Hd12 … HTX2 … HX2) [ // ] -Hd12 -HTX2 -HX2 -H2 #H2
/2 width=4 by scpes_canc_dx/ (**) (* full auto fails *)
(* Advanced properties ******************************************************)
-lemma snv_cast_scpes: ∀h,g,G,L,U,T. ⦃G, L⦄ ⊢ U ¡[h, g] → ⦃G, L⦄ ⊢ T ¡[h, g] →
- ⦃G, L⦄ ⊢ U •*⬌*[h, g, 0, 1] T → ⦃G, L⦄ ⊢ ⓝU.T ¡[h, g].
-#h #g #G #L #U #T #HU #HT * /2 width=3 by snv_cast/
+lemma snv_cast_scpes: ∀h,o,G,L,U,T. ⦃G, L⦄ ⊢ U ¡[h, o] → ⦃G, L⦄ ⊢ T ¡[h, o] →
+ ⦃G, L⦄ ⊢ U •*⬌*[h, o, 0, 1] T → ⦃G, L⦄ ⊢ ⓝU.T ¡[h, o].
+#h #o #G #L #U #T #HU #HT * /2 width=3 by snv_cast/
qed.
qed-.
(* Basic_1: was: pc3_gen_sort *)
-lemma cpcs_inv_sort: ∀G,L,k1,k2. ⦃G, L⦄ ⊢ ⋆k1 ⬌* ⋆k2 → k1 = k2.
-#G #L #k1 #k2 #H elim (cpcs_inv_cprs … H) -H
+lemma cpcs_inv_sort: ∀G,L,s1,s2. ⦃G, L⦄ ⊢ ⋆s1 ⬌* ⋆s2 → s1 = s2.
+#G #L #s1 #s2 #H elim (cpcs_inv_cprs … H) -H
#T #H1 >(cprs_inv_sort1 … H1) -T #H2
lapply (cprs_inv_sort1 … H2) -L #H destruct //
qed-.
/3 width=1 by cpcs_inv_abst1, cpcs_sym/ qed-.
(* Basic_1: was: pc3_gen_sort_abst *)
-lemma cpcs_inv_sort_abst: ∀a,G,L,W,T,k. ⦃G, L⦄ ⊢ ⋆k ⬌* ⓛ{a}W.T → ⊥.
-#a #G #L #W #T #k #H
+lemma cpcs_inv_sort_abst: ∀a,G,L,W,T,s. ⦃G, L⦄ ⊢ ⋆s ⬌* ⓛ{a}W.T → ⊥.
+#a #G #L #W #T #s #H
elim (cpcs_inv_cprs … H) -H #X #H1
>(cprs_inv_sort1 … H1) -X #H2
elim (cprs_inv_abst1 … H2) -H2 #W0 #T0 #_ #_ #H destruct
qed-.
(* Basic_1: was: pc3_gen_lift *)
-lemma cpcs_inv_lift: ∀G,L,K,s,l,m. ⬇[s, l, m] L ≡ K →
- ∀T1,U1. ⬆[l, m] T1 ≡ U1 → ∀T2,U2. ⬆[l, m] T2 ≡ U2 →
+lemma cpcs_inv_lift: ∀G,L,K,c,l,k. ⬇[c, l, k] L ≡ K →
+ ∀T1,U1. ⬆[l, k] T1 ≡ U1 → ∀T2,U2. ⬆[l, k] T2 ≡ U2 →
⦃G, L⦄ ⊢ U1 ⬌* U2 → ⦃G, K⦄ ⊢ T1 ⬌* T2.
-#G #L #K #s #l #m #HLK #T1 #U1 #HTU1 #T2 #U2 #HTU2 #HU12
+#G #L #K #c #l #k #HLK #T1 #U1 #HTU1 #T2 #U2 #HTU2 #HU12
elim (cpcs_inv_cprs … HU12) -HU12 #U #HU1 #HU2
elim (cprs_inv_lift1 … HU1 … HLK … HTU1) -U1 #T #HTU #HT1
elim (cprs_inv_lift1 … HU2 … HLK … HTU2) -L -U2 #X #HXU
->(lift_inj … HXU … HTU) -X -U -l -m /2 width=3 by cprs_div/
+>(lift_inj … HXU … HTU) -X -U -l -k /2 width=3 by cprs_div/
qed-.
(* Advanced properties ******************************************************)
qed-.
(* Basic_1: was: pc3_lift *)
-lemma cpcs_lift: ∀G,L,K,s,l,m. ⬇[s, l, m] L ≡ K →
- ∀T1,U1. ⬆[l, m] T1 ≡ U1 → ∀T2,U2. ⬆[l, m] T2 ≡ U2 →
+lemma cpcs_lift: ∀G,L,K,c,l,k. ⬇[c, l, k] L ≡ K →
+ ∀T1,U1. ⬆[l, k] T1 ≡ U1 → ∀T2,U2. ⬆[l, k] T2 ≡ U2 →
⦃G, K⦄ ⊢ T1 ⬌* T2 → ⦃G, L⦄ ⊢ U1 ⬌* U2.
-#G #L #K #s #l #m #HLK #T1 #U1 #HTU1 #T2 #U2 #HTU2 #HT12
+#G #L #K #c #l #k #HLK #T1 #U1 #HTU1 #T2 #U2 #HTU2 #HT12
elim (cpcs_inv_cprs … HT12) -HT12 #T #HT1 #HT2
-elim (lift_total T l m) /3 width=12 by cprs_div, cprs_lift/
+elim (lift_total T l k) /3 width=12 by cprs_div, cprs_lift/
qed.
lemma cpcs_strip: ∀G,L,T1,T. ⦃G, L⦄ ⊢ T ⬌* T1 → ∀T2. ⦃G, L⦄ ⊢ T ⬌ T2 →
(* STRATIFIED DECOMPOSED PARALLEL EQUIVALENCE FOR TERMS *********************)
definition scpes: ∀h. sd h → nat → nat → relation4 genv lenv term term ≝
- λh,g,d1,d2,G,L,T1,T2.
- ∃∃T. ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d1] T & ⦃G, L⦄ ⊢ T2 •*➡*[h, g, d2] T.
+ λh,o,d1,d2,G,L,T1,T2.
+ ∃∃T. ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d1] T & ⦃G, L⦄ ⊢ T2 •*➡*[h, o, d2] T.
interpretation "stratified decomposed parallel equivalence (term)"
- 'DPConvStar h g d1 d2 G L T1 T2 = (scpes h g d1 d2 G L T1 T2).
+ 'DPConvStar h o d1 d2 G L T1 T2 = (scpes h o d1 d2 G L T1 T2).
(* Basic properties *********************************************************)
-lemma scpds_div: ∀h,g,G,L,T1,T2,T,d1,d2.
- ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d1] T → ⦃G, L⦄ ⊢ T2 •*➡*[h, g, d2] T →
- ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d1, d2] T2.
+lemma scpds_div: ∀h,o,G,L,T1,T2,T,d1,d2.
+ ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d1] T → ⦃G, L⦄ ⊢ T2 •*➡*[h, o, d2] T →
+ ⦃G, L⦄ ⊢ T1 •*⬌*[h, o, d1, d2] T2.
/2 width=3 by ex2_intro/ qed.
-lemma scpes_sym: ∀h,g,G,L,T1,T2,d1,d2. ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d1, d2] T2 →
- ⦃G, L⦄ ⊢ T2 •*⬌*[h, g, d2, d1] T1.
-#h #g #G #L #T1 #T2 #L1 #d2 * /2 width=3 by scpds_div/
+lemma scpes_sym: ∀h,o,G,L,T1,T2,d1,d2. ⦃G, L⦄ ⊢ T1 •*⬌*[h, o, d1, d2] T2 →
+ ⦃G, L⦄ ⊢ T2 •*⬌*[h, o, d2, d1] T1.
+#h #o #G #L #T1 #T2 #L1 #d2 * /2 width=3 by scpds_div/
qed-.
(* Main inversion lemmas about atomic arity assignment on terms *************)
-theorem scpes_aaa_mono: ∀h,g,G,L,T1,T2,d1,d2. ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d1, d2] T2 →
+theorem scpes_aaa_mono: ∀h,o,G,L,T1,T2,d1,d2. ⦃G, L⦄ ⊢ T1 •*⬌*[h, o, d1, d2] T2 →
∀A1. ⦃G, L⦄ ⊢ T1 ⁝ A1 → ∀A2. ⦃G, L⦄ ⊢ T2 ⁝ A2 →
A1 = A2.
-#h #g #G #L #T1 #T2 #d1 #d2 * #T #HT1 #HT2 #A1 #HA1 #A2 #HA2
+#h #o #G #L #T1 #T2 #d1 #d2 * #T #HT1 #HT2 #A1 #HA1 #A2 #HA2
lapply (scpds_aaa_conf … HA1 … HT1) -T1 #HA1
lapply (scpds_aaa_conf … HA2 … HT2) -T2 #HA2
lapply (aaa_mono … HA1 … HA2) -L -T //
(* Inversion lemmas on parallel equivalence for terms ***********************)
-lemma scpes_inv_lstas_eq: ∀h,g,G,L,T1,T2,d1,d2. ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d1, d2] T2 →
+lemma scpes_inv_lstas_eq: ∀h,o,G,L,T1,T2,d1,d2. ⦃G, L⦄ ⊢ T1 •*⬌*[h, o, d1, d2] T2 →
∀U1. ⦃G, L⦄ ⊢ T1 •*[h, d1] U1 →
∀U2. ⦃G, L⦄ ⊢ T2 •*[h, d2] U2 → ⦃G, L⦄ ⊢ U1 ⬌* U2.
-#h #g #G #L #T1 #T2 #d1 #d2 * #T #HT1 #HT2 #U1 #HTU1 #U2 #HTU2
+#h #o #G #L #T1 #T2 #d1 #d2 * #T #HT1 #HT2 #U1 #HTU1 #U2 #HTU2
/3 width=8 by scpds_inv_lstas_eq, cprs_div/
qed-.
(* Properties on parallel equivalence for terms *****************************)
-lemma cpcs_scpes: ∀h,g,G,L,T1,d11. ⦃G, L⦄ ⊢ T1 ▪[h, g] d11 →
+lemma cpcs_scpes: ∀h,o,G,L,T1,d11. ⦃G, L⦄ ⊢ T1 ▪[h, o] d11 →
∀U1,d12. d12 ≤ d11 → ⦃G, L⦄ ⊢ T1 •*[h, d12] U1 →
- ∀T2,d21. ⦃G, L⦄ ⊢ T2 ▪[h, g] d21 →
+ ∀T2,d21. ⦃G, L⦄ ⊢ T2 ▪[h, o] d21 →
∀U2,d22. d22 ≤ d21 → ⦃G, L⦄ ⊢ T2 •*[h, d22] U2 →
- ⦃G, L⦄ ⊢ U1 ⬌* U2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d12, d22] T2.
-#h #g #G #L #T1 #d11 #HT1 #U1 #d12 #Hd121 #HTU1 #T2 #d21 #HT2 #U2 #d22 #Hd221 #HTU2 #HU12
+ ⦃G, L⦄ ⊢ U1 ⬌* U2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, o, d12, d22] T2.
+#h #o #G #L #T1 #d11 #HT1 #U1 #d12 #Hd121 #HTU1 #T2 #d21 #HT2 #U2 #d22 #Hd221 #HTU2 #HU12
elim (cpcs_inv_cprs … HU12) -HU12 /3 width=6 by scpds_div, ex4_2_intro/
qed.
(* Advanced inversion lemmas ************************************************)
-lemma scpes_inv_abst2: ∀h,g,a,G,L,T1,T2,W2,d1,d2. ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d1, d2] ⓛ{a}W2.T2 →
- ∃∃W,T. ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d1] ⓛ{a}W.T & ⦃G, L⦄ ⊢ W2 ➡* W &
- ⦃G, L.ⓛW2⦄ ⊢ T2 •*➡*[h, g, d2] T.
-#h #g #a #G #L #T1 #T2 #W2 #d1 #d2 * #T0 #HT10 #H
+lemma scpes_inv_abst2: ∀h,o,a,G,L,T1,T2,W2,d1,d2. ⦃G, L⦄ ⊢ T1 •*⬌*[h, o, d1, d2] ⓛ{a}W2.T2 →
+ ∃∃W,T. ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d1] ⓛ{a}W.T & ⦃G, L⦄ ⊢ W2 ➡* W &
+ ⦃G, L.ⓛW2⦄ ⊢ T2 •*➡*[h, o, d2] T.
+#h #o #a #G #L #T1 #T2 #W2 #d1 #d2 * #T0 #HT10 #H
elim (scpds_inv_abst1 … H) -H #W #T #HW2 #HT2 #H destruct /2 width=5 by ex3_2_intro/
qed-.
(* Advanced properties ******************************************************)
-lemma scpes_refl: ∀h,g,G,L,T,d1,d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ T ▪[h, g] d1 →
- ⦃G, L⦄ ⊢ T •*⬌*[h, g, d2, d2] T.
-#h #g #G #L #T #d1 #d2 #Hd21 #Hd1
+lemma scpes_refl: ∀h,o,G,L,T,d1,d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ T ▪[h, o] d1 →
+ ⦃G, L⦄ ⊢ T •*⬌*[h, o, d2, d2] T.
+#h #o #G #L #T #d1 #d2 #Hd21 #Hd1
elim (da_lstas … Hd1 … d2) #U #HTU #_
/3 width=3 by scpds_div, lstas_scpds/
qed.
-lemma lstas_scpes_trans: ∀h,g,G,L,T1,d0,d1. ⦃G, L⦄ ⊢ T1 ▪[h, g] d0 → d1 ≤ d0 →
+lemma lstas_scpes_trans: ∀h,o,G,L,T1,d0,d1. ⦃G, L⦄ ⊢ T1 ▪[h, o] d0 → d1 ≤ d0 →
∀T. ⦃G, L⦄ ⊢ T1 •*[h, d1] T →
- ∀T2,d,d2. ⦃G, L⦄ ⊢ T •*⬌*[h,g,d,d2] T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h,g,d1+d,d2] T2.
-#h #g #G #L #T1 #d0 #d1 #Hd0 #Hd10 #T #HT1 #T2 #d #d2 *
+ ∀T2,d,d2. ⦃G, L⦄ ⊢ T •*⬌*[h,o,d,d2] T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h,o,d1+d,d2] T2.
+#h #o #G #L #T1 #d0 #d1 #Hd0 #Hd10 #T #HT1 #T2 #d #d2 *
/3 width=3 by scpds_div, lstas_scpds_trans/ qed-.
(* Properties on parallel computation for terms *****************************)
-lemma cprs_scpds_div: ∀h,g,G,L,T1,T. ⦃G, L⦄ ⊢ T1 ➡* T →
- ∀d. ⦃G, L⦄ ⊢ T1 ▪[h, g] d →
- ∀T2,d2. ⦃G, L⦄ ⊢ T2 •*➡*[h, g, d2] T →
- ⦃G, L⦄⊢ T1 •*⬌*[h, g, 0, d2] T2.
-#h #g #G #L #T1 #T #HT1 #d #Hd elim (da_lstas … Hd 0)
+lemma cprs_scpds_div: ∀h,o,G,L,T1,T. ⦃G, L⦄ ⊢ T1 ➡* T →
+ ∀d. ⦃G, L⦄ ⊢ T1 ▪[h, o] d →
+ ∀T2,d2. ⦃G, L⦄ ⊢ T2 •*➡*[h, o, d2] T →
+ ⦃G, L⦄⊢ T1 •*⬌*[h, o, 0, d2] T2.
+#h #o #G #L #T1 #T #HT1 #d #Hd elim (da_lstas … Hd 0)
#X1 #HTX1 #_ elim (cprs_strip … HT1 X1) -HT1
/3 width=5 by scpds_strap1, scpds_div, lstas_cpr, ex4_2_intro/
qed.
(* Main properties **********************************************************)
-theorem scpes_trans: ∀h,g,G,L,T1,T,d1,d. ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d1, d] T →
- ∀T2,d2. ⦃G, L⦄ ⊢ T •*⬌*[h, g, d, d2] T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d1, d2] T2.
-#h #g #G #L #T1 #T #d1 #d * #X1 #HT1X1 #HTX1 #T2 #d2 * #X2 #HTX2 #HT2X2
+theorem scpes_trans: ∀h,o,G,L,T1,T,d1,d. ⦃G, L⦄ ⊢ T1 •*⬌*[h, o, d1, d] T →
+ ∀T2,d2. ⦃G, L⦄ ⊢ T •*⬌*[h, o, d, d2] T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, o, d1, d2] T2.
+#h #o #G #L #T1 #T #d1 #d * #X1 #HT1X1 #HTX1 #T2 #d2 * #X2 #HTX2 #HT2X2
elim (scpds_conf_eq … HTX1 … HTX2) -T -d /3 width=5 by scpds_cprs_trans, scpds_div/
qed-.
-theorem scpes_canc_sn: ∀h,g,G,L,T,T1,d,d1. ⦃G, L⦄ ⊢ T •*⬌*[h, g, d, d1] T1 →
- ∀T2,d2. ⦃G, L⦄ ⊢ T •*⬌*[h, g, d, d2] T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d1, d2] T2.
+theorem scpes_canc_sn: ∀h,o,G,L,T,T1,d,d1. ⦃G, L⦄ ⊢ T •*⬌*[h, o, d, d1] T1 →
+ ∀T2,d2. ⦃G, L⦄ ⊢ T •*⬌*[h, o, d, d2] T2 → ⦃G, L⦄ ⊢ T1 •*⬌*[h, o, d1, d2] T2.
/3 width=4 by scpes_trans, scpes_sym/ qed-.
-theorem scpes_canc_dx: ∀h,g,G,L,T1,T,d1,d. ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d1, d] T →
- ∀T2,d2. ⦃G, L⦄ ⊢ T2 •*⬌*[h, g, d2, d] T → ⦃G, L⦄ ⊢ T1 •*⬌*[h, g, d1, d2] T2.
+theorem scpes_canc_dx: ∀h,o,G,L,T1,T,d1,d. ⦃G, L⦄ ⊢ T1 •*⬌*[h, o, d1, d] T →
+ ∀T2,d2. ⦃G, L⦄ ⊢ T2 •*⬌*[h, o, d2, d] T → ⦃G, L⦄ ⊢ T1 •*⬌*[h, o, d1, d2] T2.
/3 width=4 by scpes_trans, scpes_sym/ qed-.
(* activate genv *)
inductive lstas (h): nat → relation4 genv lenv term term ≝
-| lstas_sort: ∀G,L,d,k. lstas h d G L (⋆k) (⋆((next h)^d k))
+| lstas_sort: ∀G,L,d,s. lstas h d G L (⋆s) (⋆((next h)^d s))
| lstas_ldef: ∀G,L,K,V,W,U,i,d. ⬇[i] L ≡ K.ⓓV → lstas h d G K V W →
⬆[0, i+1] W ≡ U → lstas h d G L (#i) U
| lstas_zero: ∀G,L,K,W,V,i. ⬇[i] L ≡ K.ⓛW → lstas h 0 G K W V →
(* Basic inversion lemmas ***************************************************)
-fact lstas_inv_sort1_aux: ∀h,G,L,T,U,d. ⦃G, L⦄ ⊢ T •*[h, d] U → ∀k0. T = ⋆k0 →
- U = ⋆((next h)^d k0).
+fact lstas_inv_sort1_aux: ∀h,G,L,T,U,d. ⦃G, L⦄ ⊢ T •*[h, d] U → ∀s0. T = ⋆s0 →
+ U = ⋆((next h)^d s0).
#h #G #L #T #U #d * -G -L -T -U -d
-[ #G #L #d #k #k0 #H destruct //
-| #G #L #K #V #W #U #i #d #_ #_ #_ #k0 #H destruct
-| #G #L #K #W #V #i #_ #_ #k0 #H destruct
-| #G #L #K #W #V #U #i #d #_ #_ #_ #k0 #H destruct
-| #a #I #G #L #V #T #U #d #_ #k0 #H destruct
-| #G #L #V #T #U #d #_ #k0 #H destruct
-| #G #L #W #T #U #d #_ #k0 #H destruct
+[ #G #L #d #s #s0 #H destruct //
+| #G #L #K #V #W #U #i #d #_ #_ #_ #s0 #H destruct
+| #G #L #K #W #V #i #_ #_ #s0 #H destruct
+| #G #L #K #W #V #U #i #d #_ #_ #_ #s0 #H destruct
+| #a #I #G #L #V #T #U #d #_ #s0 #H destruct
+| #G #L #V #T #U #d #_ #s0 #H destruct
+| #G #L #W #T #U #d #_ #s0 #H destruct
qed-.
(* Basic_1: was just: sty0_gen_sort *)
-lemma lstas_inv_sort1: ∀h,G,L,X,k,d. ⦃G, L⦄ ⊢ ⋆k •*[h, d] X → X = ⋆((next h)^d k).
+lemma lstas_inv_sort1: ∀h,G,L,X,s,d. ⦃G, L⦄ ⊢ ⋆s •*[h, d] X → X = ⋆((next h)^d s).
/2 width=5 by lstas_inv_sort1_aux/
qed-.
⬆[0, j+1] V ≡ U & d = d0+1
).
#h #G #L #T #U #d * -G -L -T -U -d
-[ #G #L #d #k #j #H destruct
+[ #G #L #d #s #j #H destruct
| #G #L #K #V #W #U #i #d #HLK #HVW #HWU #j #H destruct /3 width=6 by or3_intro0, ex3_3_intro/
| #G #L #K #W #V #i #HLK #HWV #j #H destruct /3 width=5 by or3_intro1, ex4_3_intro/
| #G #L #K #W #V #U #i #d #HLK #HWV #HWU #j #H destruct /3 width=8 by or3_intro2, ex4_4_intro/
fact lstas_inv_gref1_aux: ∀h,G,L,T,U,d. ⦃G, L⦄ ⊢ T •*[h, d] U → ∀p0. T = §p0 → ⊥.
#h #G #L #T #U #d * -G -L -T -U -d
-[ #G #L #d #k #p0 #H destruct
+[ #G #L #d #s #p0 #H destruct
| #G #L #K #V #W #U #i #d #_ #_ #_ #p0 #H destruct
| #G #L #K #W #V #i #_ #_ #p0 #H destruct
| #G #L #K #W #V #U #i #d #_ #_ #_ #p0 #H destruct
fact lstas_inv_bind1_aux: ∀h,G,L,T,U,d. ⦃G, L⦄ ⊢ T •*[h, d] U → ∀b,J,X,Y. T = ⓑ{b,J}Y.X →
∃∃Z. ⦃G, L.ⓑ{J}Y⦄ ⊢ X •*[h, d] Z & U = ⓑ{b,J}Y.Z.
#h #G #L #T #U #d * -G -L -T -U -d
-[ #G #L #d #k #b #J #X #Y #H destruct
+[ #G #L #d #s #b #J #X #Y #H destruct
| #G #L #K #V #W #U #i #d #_ #_ #_ #b #J #X #Y #H destruct
| #G #L #K #W #V #i #_ #_ #b #J #X #Y #H destruct
| #G #L #K #W #V #U #i #d #_ #_ #_ #b #J #X #Y #H destruct
fact lstas_inv_appl1_aux: ∀h,G,L,T,U,d. ⦃G, L⦄ ⊢ T •*[h, d] U → ∀X,Y. T = ⓐY.X →
∃∃Z. ⦃G, L⦄ ⊢ X •*[h, d] Z & U = ⓐY.Z.
#h #G #L #T #U #d * -G -L -T -U -d
-[ #G #L #d #k #X #Y #H destruct
+[ #G #L #d #s #X #Y #H destruct
| #G #L #K #V #W #U #i #d #_ #_ #_ #X #Y #H destruct
| #G #L #K #W #V #i #_ #_ #X #Y #H destruct
| #G #L #K #W #V #U #i #d #_ #_ #_ #X #Y #H destruct
fact lstas_inv_cast1_aux: ∀h,G,L,T,U,d. ⦃G, L⦄ ⊢ T •*[h, d] U → ∀X,Y. T = ⓝY.X →
⦃G, L⦄ ⊢ X •*[h, d] U.
#h #G #L #T #U #d * -G -L -T -U -d
-[ #G #L #d #k #X #Y #H destruct
+[ #G #L #d #s #X #Y #H destruct
| #G #L #K #V #W #U #i #d #_ #_ #_ #X #Y #H destruct
| #G #L #K #W #V #i #_ #_ #X #Y #H destruct
| #G #L #K #W #V #U #i #d #_ #_ #_ #X #Y #H destruct
(* Properties on degree assignment for terms ********************************)
-lemma da_lstas: ∀h,g,G,L,T,d1. ⦃G, L⦄ ⊢ T ▪[h, g] d1 → ∀d2.
- ∃∃U. ⦃G, L⦄ ⊢ T •*[h, d2] U & ⦃G, L⦄ ⊢ U ▪[h, g] d1-d2.
-#h #g #G #L #T #d1 #H elim H -G -L -T -d1
+lemma da_lstas: ∀h,o,G,L,T,d1. ⦃G, L⦄ ⊢ T ▪[h, o] d1 → ∀d2.
+ ∃∃U. ⦃G, L⦄ ⊢ T •*[h, d2] U & ⦃G, L⦄ ⊢ U ▪[h, o] d1-d2.
+#h #o #G #L #T #d1 #H elim H -G -L -T -d1
[ /4 width=3 by da_sort, deg_iter, ex2_intro/
| #G #L #K #V #i #d1 #HLK #_ #IHV #d2
elim (IHV d2) -IHV #W
]
qed-.
-lemma lstas_da_conf: ∀h,g,G,L,T,U,d2. ⦃G, L⦄ ⊢ T •*[h, d2] U →
- ∀d1. ⦃G, L⦄ ⊢ T ▪[h, g] d1 → ⦃G, L⦄ ⊢ U ▪[h, g] d1-d2.
-#h #g #G #L #T #U #d2 #HTU #d1 #HT
+lemma lstas_da_conf: ∀h,o,G,L,T,U,d2. ⦃G, L⦄ ⊢ T •*[h, d2] U →
+ ∀d1. ⦃G, L⦄ ⊢ T ▪[h, o] d1 → ⦃G, L⦄ ⊢ U ▪[h, o] d1-d2.
+#h #o #G #L #T #U #d2 #HTU #d1 #HT
elim (da_lstas … HT d2) -HT #X #HTX
lapply (lstas_mono … HTX … HTU) -T //
qed-.
(* inversion lemmas on degree assignment for terms **************************)
-lemma lstas_inv_da: ∀h,g,G,L,T,U,d2. ⦃G, L⦄ ⊢ T •*[h, d2] U →
- ∃∃d1. ⦃G, L⦄ ⊢ T ▪[h, g] d1 & ⦃G, L⦄ ⊢ U ▪[h, g] d1-d2.
-#h #g #G #L #T #U #d2 #H elim H -G -L -T -U -d2
-[ #G #L #d2 #k elim (deg_total h g k) /4 width=3 by da_sort, deg_iter, ex2_intro/
+lemma lstas_inv_da: ∀h,o,G,L,T,U,d2. ⦃G, L⦄ ⊢ T •*[h, d2] U →
+ ∃∃d1. ⦃G, L⦄ ⊢ T ▪[h, o] d1 & ⦃G, L⦄ ⊢ U ▪[h, o] d1-d2.
+#h #o #G #L #T #U #d2 #H elim H -G -L -T -U -d2
+[ #G #L #d2 #s elim (deg_total h o s) /4 width=3 by da_sort, deg_iter, ex2_intro/
| #G #L #K #V #W #U #i #d2 #HLK #_ #HWU *
lapply (drop_fwd_drop2 … HLK) /3 width=10 by da_lift, da_ldef, ex2_intro/
| #G #L #K #W #V #i #HLK #_ * /3 width=6 by da_ldec, ex2_intro/
qed-.
lemma lstas_inv_da_ge: ∀h,G,L,T,U,d2,d. ⦃G, L⦄ ⊢ T •*[h, d2] U →
- ∃∃g,d1. ⦃G, L⦄ ⊢ T ▪[h, g] d1 & ⦃G, L⦄ ⊢ U ▪[h, g] d1-d2 & d ≤ d1.
+ ∃∃o,d1. ⦃G, L⦄ ⊢ T ▪[h, o] d1 & ⦃G, L⦄ ⊢ U ▪[h, o] d1-d2 & d ≤ d1.
#h #G #L #T #U #d2 #d #H elim H -G -L -T -U -d2
[ /4 width=5 by da_sort, deg_iter, ex3_2_intro/
| #G #L #K #V #W #U #i #d2 #HLK #_ #HWU *
lapply (drop_fwd_drop2 … HLK) /3 width=10 by da_lift, da_ldef, ex3_2_intro/
| #G #L #K #W #V #i #HLK #_ *
- #g #d1 #HW #HV #Hd1 /4 width=6 by da_ldec, lt_to_le, le_S_S, ex3_2_intro/
+ #o #d1 #HW #HV #Hd1 /4 width=6 by da_ldec, lt_to_le, le_S_S, ex3_2_intro/
| #G #L #K #W #V #U #i #d2 #HLK #_ #HVU *
lapply (drop_fwd_drop2 … HLK)
/4 width=10 by da_lift, da_ldec, lt_to_le, le_S_S, ex3_2_intro/
lemma lstas_inv_refl_pos: ∀h,G,L,T,d. ⦃G, L⦄ ⊢ T •*[h, d+1] T → ⊥.
#h #G #L #T #d2 #H elim (lstas_inv_da_ge … (d2+1) H) -H
-#g #d1 #HT1 #HT12 #Hd21 lapply (da_mono … HT1 … HT12) -h -G -L -T
+#o #d1 #HT1 #HT12 #Hd21 lapply (da_mono … HT1 … HT12) -h -G -L -T
#H elim (discr_x_minus_xy … H) -H
[ #H destruct /2 width=3 by le_plus_xSy_O_false/
| -d1 <plus_n_Sm #H destruct
(* Basic_1: was just: sty0_lift *)
lemma lstas_lift: ∀h,G,d. d_liftable (lstas h G d).
#h #G #d #L1 #T1 #U1 #H elim H -G -L1 -T1 -U1 -d
-[ #G #L1 #d #k #L2 #s #l #m #HL21 #X1 #H1 #X2 #H2
+[ #G #L1 #d #s #L2 #c #l #k #HL21 #X1 #H1 #X2 #H2
>(lift_inv_sort1 … H1) -X1
>(lift_inv_sort1 … H2) -X2 //
-| #G #L1 #K1 #V1 #W1 #W #i #d #HLK1 #_ #HW1 #IHVW1 #L2 #s #l #m #HL21 #X #H #U2 #HWU2
+| #G #L1 #K1 #V1 #W1 #W #i #d #HLK1 #_ #HW1 #IHVW1 #L2 #c #l #k #HL21 #X #H #U2 #HWU2
elim (lift_inv_lref1 … H) * #Hil #H destruct
[ elim (lift_trans_ge … HW1 … HWU2) -W /2 width=1 by ylt_fwd_le_succ1/ #W2 #HW12 #HWU2
elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by ylt_fwd_le/ #X #HLK2 #H
| lapply (lift_trans_be … HW1 … HWU2 ? ?) -W /2 width=1 by yle_succ_dx/ #HW1U2
lapply (drop_trans_ge … HL21 … HLK1 ?) -L1 /3 width=9 by lstas_ldef, drop_inv_gen/
]
-| #G #L1 #K1 #V1 #W1 #i #HLK1 #_ #IHVW1 #L2 #s #l #m #HL21 #X #H #U2 #HWU2
+| #G #L1 #K1 #V1 #W1 #i #HLK1 #_ #IHVW1 #L2 #c #l #k #HL21 #X #H #U2 #HWU2
>(lift_mono … HWU2 … H) -U2
elim (lift_inv_lref1 … H) * #Hil #H destruct
- [ elim (lift_total W1 (l-i-1) m) #W2 #HW12
+ [ elim (lift_total W1 (l-i-1) k) #W2 #HW12
elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by ylt_fwd_le/ #X #HLK2 #H
elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ -Hil #K2 #V2 #HK21 #HV12 #H destruct
/3 width=10 by lstas_zero/
| lapply (drop_trans_ge … HL21 … HLK1 ?) -L1
/3 width=10 by lstas_zero, drop_inv_gen/
]
-| #G #L1 #K1 #W1 #V1 #W #i #d #HLK1 #_ #HW1 #IHWV1 #L2 #s #l #m #HL21 #X #H #U2 #HWU2
+| #G #L1 #K1 #W1 #V1 #W #i #d #HLK1 #_ #HW1 #IHWV1 #L2 #c #l #k #HL21 #X #H #U2 #HWU2
elim (lift_inv_lref1 … H) * #Hil #H destruct
[ elim (lift_trans_ge … HW1 … HWU2) -W /2 width=1 by ylt_fwd_le_succ1/ #W #HW1 #HWU2
elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by ylt_fwd_le/ #X #HLK2 #H
| lapply (lift_trans_be … HW1 … HWU2 ? ?) -W /2 width=1 by yle_succ_dx/ #HW1U2
lapply (drop_trans_ge … HL21 … HLK1 ?) -L1 /3 width=9 by lstas_succ, drop_inv_gen/
]
-| #a #I #G #L1 #V1 #T1 #U1 #d #_ #IHTU1 #L2 #s #l #m #HL21 #X1 #H1 #X2 #H2
+| #a #I #G #L1 #V1 #T1 #U1 #d #_ #IHTU1 #L2 #c #l #k #HL21 #X1 #H1 #X2 #H2
elim (lift_inv_bind1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H destruct
elim (lift_inv_bind1 … H2) -H2 #X #U2 #H1 #HU12 #H2 destruct
lapply (lift_mono … H1 … HV12) -H1 #H destruct /4 width=6 by lstas_bind, drop_skip/
-| #G #L1 #V1 #T1 #U1 #d #_ #IHTU1 #L2 #s #l #m #HL21 #X1 #H1 #X2 #H2
+| #G #L1 #V1 #T1 #U1 #d #_ #IHTU1 #L2 #c #l #k #HL21 #X1 #H1 #X2 #H2
elim (lift_inv_flat1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H destruct
elim (lift_inv_flat1 … H2) -H2 #X #U2 #H1 #HU12 #H2 destruct
lapply (lift_mono … H1 … HV12) -H1 #H destruct /4 width=6 by lstas_appl/
-| #G #L1 #W1 #T1 #U1 #d #_ #IHTU1 #L2 #s #l #m #HL21 #X #H #U2 #HU12
+| #G #L1 #W1 #T1 #U1 #d #_ #IHTU1 #L2 #c #l #k #HL21 #X #H #U2 #HU12
elim (lift_inv_flat1 … H) -H #W2 #T2 #_ #HT12 #H destruct /3 width=6 by lstas_cast/
]
qed.
(* Note: apparently this was missing in basic_1 *)
lemma lstas_inv_lift1: ∀h,G,d. d_deliftable_sn (lstas h G d).
#h #G #d #L2 #T2 #U2 #H elim H -G -L2 -T2 -U2 -d
-[ #G #L2 #d #k #L1 #s #l #m #_ #X #H
+[ #G #L2 #d #s #L1 #c #l #k #_ #X #H
>(lift_inv_sort2 … H) -X /2 width=3 by lstas_sort, lift_sort, ex2_intro/
-| #G #L2 #K2 #V2 #W2 #W #i #d #HLK2 #HVW2 #HW2 #IHVW2 #L1 #s #l #m #HL21 #X #H
+| #G #L2 #K2 #V2 #W2 #W #i #d #HLK2 #HVW2 #HW2 #IHVW2 #L1 #c #l #k #HL21 #X #H
elim (lift_inv_lref2 … H) * #Hil #H destruct [ -HVW2 | -IHVW2 ]
[ elim (drop_conf_lt … HL21 … HLK2) -L2 // #K1 #V1 #HLK1 #HK21 #HV12
elim (IHVW2 … HK21 … HV12) -K2 -V2 #W1 #HW12 #HVW1
elim (lift_trans_le … HW12 … HW2) -W2 // <yminus_succ2 <yplus_inj >yplus_SO2 >ymax_pre_sn /3 width=8 by lstas_ldef, ylt_fwd_le_succ1, ex2_intro/
| lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
elim (yle_inv_plus_inj2 … Hil) -Hil #Hlim #mi
- elim (lift_split … HW2 l (i-m+1)) -HW2 /2 width=1 by yle_succ_dx, le_S_S/
- #W0 #HW20 <le_plus_minus_comm /2 width=1 by yle_inv_inj/ >minus_minus_m_m /3 width=8 by lstas_ldef, yle_inv_inj, le_S, ex2_intro/
+ elim (lift_split … HW2 l (i-k+1)) -HW2 /2 width=1 by yle_succ_dx, le_S_S/
+ #W0 #HW20 <le_plus_minus_comm /2 width=1 by yle_inv_inj/ >minus_minus_k_k /3 width=8 by lstas_ldef, yle_inv_inj, le_S, ex2_intro/
]
-| #G #L2 #K2 #W2 #V2 #i #HLK2 #HWV2 #IHWV2 #L1 #s #l #m #HL21 #X #H
+| #G #L2 #K2 #W2 #V2 #i #HLK2 #HWV2 #IHWV2 #L1 #c #l #k #HL21 #X #H
elim (lift_inv_lref2 … H) * #Hil #H destruct [ -HWV2 | -IHWV2 ]
[ elim (drop_conf_lt … HL21 … HLK2) -L2 // #K1 #W1 #HLK1 #HK21 #HW12
elim (IHWV2 … HK21 … HW12) -K2
| lapply (drop_conf_ge … HL21 … HLK2 ?) -L2
/3 width=5 by lstas_zero, lift_lref_ge_minus, ex2_intro/
]
-| #G #L2 #K2 #W2 #V2 #W #i #d #HLK2 #HWV2 #HW2 #IHWV2 #L1 #s #l #m #HL21 #X #H
+| #G #L2 #K2 #W2 #V2 #W #i #d #HLK2 #HWV2 #HW2 #IHWV2 #L1 #c #l #k #HL21 #X #H
elim (lift_inv_lref2 … H) * #Hil #H destruct [ -HWV2 | -IHWV2 ]
[ elim (drop_conf_lt … HL21 … HLK2) -L2 // #K1 #W1 #HLK1 #HK21 #HW12
elim (IHWV2 … HK21 … HW12) -K2 #V1 #HV12 #HWV1
elim (lift_trans_le … HV12 … HW2) -W2 // <yminus_succ2 <yplus_inj >yplus_SO2 >ymax_pre_sn /3 width=8 by lstas_succ, ylt_fwd_le_succ1, ex2_intro/
| lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
elim (yle_inv_plus_inj2 … Hil) -Hil #Hlim #mi
- elim (lift_split … HW2 l (i-m+1)) -HW2 /2 width=1 by yle_succ_dx, le_S_S/
- #W0 #HW20 <le_plus_minus_comm /2 width=1 by yle_inv_inj/ >minus_minus_m_m /3 width=8 by lstas_succ, yle_inv_inj, le_S, ex2_intro/
+ elim (lift_split … HW2 l (i-k+1)) -HW2 /2 width=1 by yle_succ_dx, le_S_S/
+ #W0 #HW20 <le_plus_minus_comm /2 width=1 by yle_inv_inj/ >minus_minus_k_k /3 width=8 by lstas_succ, yle_inv_inj, le_S, ex2_intro/
]
-| #a #I #G #L2 #V2 #T2 #U2 #d #_ #IHTU2 #L1 #s #l #m #HL21 #X #H
+| #a #I #G #L2 #V2 #T2 #U2 #d #_ #IHTU2 #L1 #c #l #k #HL21 #X #H
elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
elim (IHTU2 (L1.ⓑ{I}V1) … HT12) -IHTU2 -HT12 /3 width=5 by lstas_bind, drop_skip, lift_bind, ex2_intro/
-| #G #L2 #V2 #T2 #U2 #d #_ #IHTU2 #L1 #s #l #m #HL21 #X #H
+| #G #L2 #V2 #T2 #U2 #d #_ #IHTU2 #L1 #c #l #k #HL21 #X #H
elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
elim (IHTU2 … HL21 … HT12) -L2 -HT12 /3 width=5 by lstas_appl, lift_flat, ex2_intro/
-| #G #L2 #W2 #T2 #U2 #d #_ #IHTU2 #L1 #s #l #m #HL21 #X #H
+| #G #L2 #W2 #T2 #U2 #d #_ #IHTU2 #L1 #c #l #k #HL21 #X #H
elim (lift_inv_flat2 … H) -H #W1 #T1 #_ #HT12 #H destruct
elim (IHTU2 … HL21 … HT12) -L2 -HT12 /3 width=3 by lstas_cast, ex2_intro/
]
lemma lstas_split_aux: ∀h,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 •*[h, d] T2 → ∀d1,d2. d = d1 + d2 →
∃∃T. ⦃G, L⦄ ⊢ T1 •*[h, d1] T & ⦃G, L⦄ ⊢ T •*[h, d2] T2.
#h #G #L #T1 #T2 #d #H elim H -G -L -T1 -T2 -d
-[ #G #L #d #k #d1 #d2 #H destruct
+[ #G #L #d #s #d1 #d2 #H destruct
>commutative_plus >iter_plus /2 width=3 by lstas_sort, ex2_intro/
| #G #L #K #V1 #V2 #U2 #i #d #HLK #_ #VU2 #IHV12 #d1 #d2 #H destruct
elim (IHV12 d1 d2) -IHV12 // #V
(* Properties on lazy sn pointwise extensions *******************************)
lemma lstas_llpx_sn_conf: ∀R. (∀L. reflexive … (R L)) → d_liftable R →
- ∀h,G,d. s_r_confluent1 … (lstas h d G) (llpx_sn R 0).
+ ∀h,G,d. c_r_confluent1 … (lstas h d G) (llpx_sn R 0).
#R #H1R #H2R #h #G #d #Ls #T1 #T2 #H elim H -G -Ls -T1 -T2 -d
[ /3 width=4 by llpx_sn_fwd_length, llpx_sn_sort/
-| #G #Ls #Ks #V1s #V2s #W2s #i #d #HLKs #_ #HVW2s #IHV12s #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
- #Kd #V1l #HLKd #HV1s #HV1sd
+| #G #Ls #Ks #V1c #V2c #W2c #i #d #HLKs #_ #HVW2c #IHV12c #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
+ #Kd #V1l #HLKd #HV1c #HV1sd
lapply (drop_fwd_drop2 … HLKs) -HLKs #HLKs
lapply (drop_fwd_drop2 … HLKd) -HLKd #HLKd
- @(llpx_sn_lift_le … HLKs HLKd … HVW2s) -HLKs -HLKd -HVW2s /2 width=1 by/ (**) (* full auto too slow *)
+ @(llpx_sn_lift_le … HLKs HLKd … HVW2c) -HLKs -HLKd -HVW2c /2 width=1 by/ (**) (* full auto too slow *)
| //
-| #G #Ls #Ks #V1s #V2s #W2s #i #d #HLKs #_ #HVW2s #IHV12s #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
- #Kd #V1l #HLKd #HV1s #HV1sd
+| #G #Ls #Ks #V1c #V2c #W2c #i #d #HLKs #_ #HVW2c #IHV12c #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
+ #Kd #V1l #HLKd #HV1c #HV1sd
lapply (drop_fwd_drop2 … HLKs) -HLKs #HLKs
lapply (drop_fwd_drop2 … HLKd) -HLKd #HLKd
- @(llpx_sn_lift_le … HLKs HLKd … HVW2s) -HLKs -HLKd -HVW2s /2 width=1 by/ (**) (* full auto too slow *)
+ @(llpx_sn_lift_le … HLKs HLKd … HVW2c) -HLKs -HLKd -HVW2c /2 width=1 by/ (**) (* full auto too slow *)
| #a #I #G #Ls #V #T1 #T2 #d #_ #IHT12 #Ld #H elim (llpx_sn_inv_bind_O … H) -H
/4 width=5 by llpx_sn_bind_repl_SO, llpx_sn_bind/
| #G #Ls #V #T1 #T2 #d #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
theorem lstas_trans: ∀h,G,L,T1,T,d1. ⦃G, L⦄ ⊢ T1 •*[h, d1] T →
∀T2,d2. ⦃G, L⦄ ⊢ T •*[h, d2] T2 → ⦃G, L⦄ ⊢ T1 •*[h, d1+d2] T2.
#h #G #L #T1 #T #d1 #H elim H -G -L -T1 -T -d1
-[ #G #L #d1 #k #X #d2 #H >(lstas_inv_sort1 … H) -X
+[ #G #L #d1 #s #X #d2 #H >(lstas_inv_sort1 … H) -X
<iter_plus /2 width=1 by lstas_sort/
| #G #L #K #V1 #V #U #i #d1 #HLK #_ #HVU #IHV1 #U2 #d2 #HU2
lapply (drop_fwd_drop2 … HLK) #H0
(* Note: apparently this was missing in basic_1 *)
theorem lstas_mono: ∀h,G,L,d. singlevalued … (lstas h d G L).
#h #G #L #d #T #T1 #H elim H -G -L -T -T1 -d
-[ #G #L #d #k #X #H >(lstas_inv_sort1 … H) -X //
+[ #G #L #d #s #X #H >(lstas_inv_sort1 … H) -X //
| #G #L #K #V #V1 #U1 #i #d #HLK #_ #HVU1 #IHV1 #X #H
elim (lstas_inv_lref1 … H) -H *
#K0 #V0 #W0 [3: #d0 ] #HLK0
∀U2,d2. d1 ≤ d2 → ⦃G, L⦄ ⊢ T •*[h, d2] U2 →
⦃G, L⦄ ⊢ U1 •*[h, d2-d1] U2.
#h #G #L #T #U1 #d1 #HTU1 #U2 #d2 #Hd12
->(plus_minus_m_m … Hd12) in ⊢ (%→?); -Hd12 >commutative_plus #H
+>(plus_minus_k_k … Hd12) in ⊢ (%→?); -Hd12 >commutative_plus #H
elim (lstas_split … H) -H #U #HTU
>(lstas_mono … HTU … HTU1) -T //
qed-.
∃∃T. ⦃G, L⦄ ⊢ T1 •*[h, d2] T & ⦃G, L⦄ ⊢ T2 •*[h, d1] T.
#h #G #L #T0 #T1 #d1 #HT01 #T2 #d2 #HT02
elim (lstas_lstas … HT01 (d1+d2)) #T #HT0
-lapply (lstas_conf_le … HT01 … HT0) // -HT01 <minus_plus_m_m_commutative
-lapply (lstas_conf_le … HT02 … HT0) // -T0 <minus_plus_m_m
+lapply (lstas_conf_le … HT01 … HT0) // -HT01 <minus_plus_k_k_commutative
+lapply (lstas_conf_le … HT02 … HT0) // -T0 <minus_plus_k_k
/2 width=3 by ex2_intro/
qed-.
(* activate genv *)
inductive unfold: relation4 genv lenv term lenv ≝
-| unfold_sort: ∀G,L,k. unfold G L (⋆k) L
+| unfold_sort: ∀G,L,s. unfold G L (⋆s) L
| unfold_lref: ∀I,G,L1,L2,K1,K2,V,i. ⬇[i] L1 ≡ K1. ⓑ{I}V →
unfold G K1 V K2 → ⬇[Ⓣ, |L2|, i] L2 ≡ K2 →
unfold G L1 (#i) (L1@@L2)
(* Basic properties *********************************************************)
-lemma Delta_lift: ∀W1,W2,l,m. ⬆[l, m] W1 ≡ W2 →
- ⬆[l, m] (Delta W1) ≡ (Delta W2).
+lemma Delta_lift: ∀W1,W2,l,k. ⬆[l, k] W1 ≡ W2 →
+ ⬆[l, k] (Delta W1) ≡ (Delta W2).
/4 width=1 by lift_flat, lift_bind, lift_lref_lt/ qed.
(* Main properties **********************************************************)
(* Reflexivity of proper qrst-computation: the term ApplOmega ***************)
-definition ApplDelta: term → nat → term ≝ λW,k. +ⓛW.ⓐ⋆k.ⓐ#0.#0.
+definition ApplDelta: term → nat → term ≝ λW,s. +ⓛW.ⓐ⋆s.ⓐ#0.#0.
-definition ApplOmega1: term → nat → term ≝ λW,k. ⓐ(ApplDelta W k).(ApplDelta W k).
+definition ApplOmega1: term → nat → term ≝ λW,s. ⓐ(ApplDelta W s).(ApplDelta W s).
-definition ApplOmega2: term → nat → term ≝ λW,k. +ⓓⓝW.(ApplDelta W k).ⓐ⋆k.ⓐ#0.#0.
+definition ApplOmega2: term → nat → term ≝ λW,s. +ⓓⓝW.(ApplDelta W s).ⓐ⋆s.ⓐ#0.#0.
-definition ApplOmega3: term → nat → term ≝ λW,k. ⓐ⋆k.(ApplOmega1 W k).
+definition ApplOmega3: term → nat → term ≝ λW,s. ⓐ⋆s.(ApplOmega1 W s).
(* Basic properties *********************************************************)
-lemma ApplDelta_lift: ∀W1,W2,k,l,m. ⬆[l, m] W1 ≡ W2 →
- ⬆[l, m] (ApplDelta W1 k) ≡ (ApplDelta W2 k).
+lemma ApplDelta_lift: ∀W1,W2,s,l,k. ⬆[l, k] W1 ≡ W2 →
+ ⬆[l, k] (ApplDelta W1 s) ≡ (ApplDelta W2 s).
/5 width=1 by lift_flat, lift_bind, lift_lref_lt/ qed.
-lemma cpr_ApplOmega_12: ∀G,L,W,k. ⦃G, L⦄ ⊢ ApplOmega1 W k ➡ ApplOmega2 W k.
+lemma cpr_ApplOmega_12: ∀G,L,W,s. ⦃G, L⦄ ⊢ ApplOmega1 W s ➡ ApplOmega2 W s.
/2 width=1 by cpr_beta/ qed.
-lemma cpr_ApplOmega_23: ∀G,L,W,k. ⦃G, L⦄ ⊢ ApplOmega2 W k ➡ ApplOmega3 W k.
-#G #L #W1 #k elim (lift_total W1 0 1) #W2 #HW12
-@(cpr_zeta … (ApplOmega3 W2 k)) /4 width=1 by ApplDelta_lift, lift_flat/
-@cpr_flat // @cpr_flat @(cpr_delta … (ApplDelta W1 k) ? 0)
+lemma cpr_ApplOmega_23: ∀G,L,W,s. ⦃G, L⦄ ⊢ ApplOmega2 W s ➡ ApplOmega3 W s.
+#G #L #W1 #s elim (lift_total W1 0 1) #W2 #HW12
+@(cpr_zeta … (ApplOmega3 W2 s)) /4 width=1 by ApplDelta_lift, lift_flat/
+@cpr_flat // @cpr_flat @(cpr_delta … (ApplDelta W1 s) ? 0)
[3,5,8,10: /2 width=2 by ApplDelta_lift/ |4,9: /2 width=1 by cpr_eps/ |*: skip ]
qed.
-lemma cpxs_ApplOmega_13: ∀h,g,G,L,W,k. ⦃G, L⦄ ⊢ ApplOmega1 W k ➡*[h,g] ApplOmega3 W k.
+lemma cpxs_ApplOmega_13: ∀h,o,G,L,W,s. ⦃G, L⦄ ⊢ ApplOmega1 W s ➡*[h,o] ApplOmega3 W s.
/4 width=3 by cpxs_strap1, cpr_cpx/ qed.
-lemma fqup_ApplOmega_13: ∀G,L,W,k. ⦃G, L, ApplOmega3 W k⦄ ⊐+ ⦃G, L, ApplOmega1 W k⦄.
+lemma fqup_ApplOmega_13: ∀G,L,W,s. ⦃G, L, ApplOmega3 W s⦄ ⊐+ ⦃G, L, ApplOmega1 W s⦄.
/2 width=1 by/ qed.
(* Main properties **********************************************************)
-theorem fpbg_refl: ∀h,g,G,L,W,k. ⦃G, L, ApplOmega1 W k⦄ >≡[h,g] ⦃G, L, ApplOmega1 W k⦄.
+theorem fpbg_refl: ∀h,o,G,L,W,s. ⦃G, L, ApplOmega1 W s⦄ >≡[h,o] ⦃G, L, ApplOmega1 W s⦄.
/3 width=5 by fpbs_fpbg_trans, fqup_fpbg, cpxs_fpbs/ qed.
(* Extended validy (basic?_2) vs. restricted validity (basic_1) *************)
(* extended validity of a closure, last arg of snv_appl > 1 *)
-lemma snv_extended: ∀h,g,a,G,L,k. ⦃G, L.ⓛ⋆k.ⓛⓛ{a}⋆k.⋆k.ⓛ#0⦄ ⊢ ⓐ#2.#0 ¡[h, g].
-#h #g #a #G #L #k elim (deg_total h g k)
-#d #Hd @(snv_appl … a … (⋆k) … (⋆k) (0+1+1))
+lemma snv_extended: ∀h,o,a,G,L,s. ⦃G, L.ⓛ⋆s.ⓛⓛ{a}⋆s.⋆s.ⓛ#0⦄ ⊢ ⓐ#2.#0 ¡[h, o].
+#h #o #a #G #L #s elim (deg_total h o s)
+#d #Hd @(snv_appl … a … (⋆s) … (⋆s) (0+1+1))
[ /4 width=5 by snv_lref, drop_drop_lt/
| /4 width=13 by snv_bind, snv_lref/
| /5 width=6 by lstas_scpds, lstas_succ, da_ldec, da_sort, drop_drop_lt/
qed.
(* restricted validity of the η-expanded closure, last arg of snv_appl = 1 **)
-lemma snv_restricted: ∀h,g,a,G,L,k. ⦃G, L.ⓛ⋆k.ⓛⓛ{a}⋆k.⋆k.ⓛⓛ{a}⋆k.ⓐ#0.#1⦄ ⊢ ⓐ#2.#0 ¡[h, g].
-#h #g #a #G #L #k elim (deg_total h g k)
-#d #Hd @(snv_appl … a … (⋆k) … (ⓐ#0.#2) (0+1))
+lemma snv_restricted: ∀h,o,a,G,L,s. ⦃G, L.ⓛ⋆s.ⓛⓛ{a}⋆s.⋆s.ⓛⓛ{a}⋆s.ⓐ#0.#1⦄ ⊢ ⓐ#2.#0 ¡[h, o].
+#h #o #a #G #L #s elim (deg_total h o s)
+#d #Hd @(snv_appl … a … (⋆s) … (ⓐ#0.#2) (0+1))
[ /4 width=5 by snv_lref, drop_drop_lt/
| @snv_lref [4: // |1,2,3: skip ]
@snv_bind //
- @(snv_appl … a … (⋆k) … (⋆k) (0+1))
+ @(snv_appl … a … (⋆s) … (⋆s) (0+1))
[ @snv_lref [4: // |1,2,3: skip ] //
| @snv_lref [4: /2 width=1 by drop_drop_lt/ |1,2,3: skip ] @snv_bind //
| @(lstas_scpds … (d+1)) /3 width=6 by da_sort, da_ldec, lstas_succ/
(* Static type assignment (iterated vs primitive): the declared variable ****)
(* basic_1: we have "L.ⓛⓝ⋆k1.⋆k2⦄ ⊢ #0 • ⓝ⋆k1.⋆k2". *)
-theorem sta_ldec: ∀h,G,L,k1,k2. ⦃G, L.ⓛⓝ⋆k1.⋆k2⦄ ⊢ #0 •*[h, 1] ⋆k2.
+theorem sta_ldec: ∀h,G,L,s1,s2. ⦃G, L.ⓛⓝ⋆s1.⋆s2⦄ ⊢ #0 •*[h, 1] ⋆s2.
/3 width=6 by lstas_sort, lstas_succ, lstas_cast, drop_pair/ qed-.
c: conversion
d: decomposed rt-reduction
e: decomposed rt-conversion
+g: generic rt-transition
q: restricted reduction
r: reduction
s: substitution
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( ⦃ term 46 G, break term 46 L ⦄ ⊢ break term 46 T1 ➡ break [ term 46 o , break term 46 h ] break term 46 T2 )"
+notation "hvbox( ⦃ term 46 G, break term 46 L ⦄ ⊢ break term 46 T1 ➡ break [ term 46 h , break term 46 r ] break term 46 T2 )"
non associative with precedence 45
- for @{ 'PRed $h $o $G $L $T1 $T2 }.
+ for @{ 'PRed $h $r $G $L $T1 $T2 }.
interpretation "uniform slicing (local environment)"
'RDropStar i L1 L2 = (drops true (uni i) L1 L2).
-interpretation "general slicing (local environment)"
+interpretation "generic slicing (local environment)"
'RDropStar c f L1 L2 = (drops c f L1 L2).
definition d_liftable1: relation2 lenv term → predicate bool ≝
--- /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 "ground_2/steps/rtc_shift.ma".
+include "ground_2/steps/rtc_plus.ma".
+include "basic_2/notation/relations/pred_6.ma".
+include "basic_2/grammar/lenv.ma".
+include "basic_2/grammar/genv.ma".
+include "basic_2/relocation/lifts.ma".
+include "basic_2/static/sh.ma".
+
+(* CONTEXT-SENSITIVE GENERIC PARALLEL RT-TRANSITION FOR TERMS ***************)
+
+(* avtivate genv *)
+inductive cpg (h): rtc → relation4 genv lenv term term ≝
+| cpg_atom : ∀I,G,L. cpg h (𝟘𝟘) G L (⓪{I}) (⓪{I})
+| cpg_ess : ∀G,L,s. cpg h (𝟘𝟙) G L (⋆s) (⋆(next h s))
+| cpg_delta: ∀r,G,L,V1,V2,W2. cpg h r G L V1 V2 →
+ ⬆*[1] V2 ≡ W2 → cpg h (↓r) G (L.ⓓV1) (#0) W2
+| cpg_ell : ∀r,G,L,V1,V2,W2. cpg h r G L V1 V2 →
+ ⬆*[1] V2 ≡ W2 → cpg h ((↓r)+𝟘𝟙) G (L.ⓛV1) (#0) W2
+| cpt_lref : ∀r,I,G,L,V,T,U,i. cpg h r G L (#i) T →
+ ⬆*[1] T ≡ U → cpg h r G (L.ⓑ{I}V) (#⫯i) U
+| cpg_bind : ∀rV,rT,p,I,G,L,V1,V2,T1,T2.
+ cpg h rV G L V1 V2 → cpg h rT G (L.ⓑ{I}V1) T1 T2 →
+ cpg h ((↓rV)+rT) G L (ⓑ{p,I}V1.T1) (ⓑ{p,I}V2.T2)
+| cpg_flat : ∀rV,rT,I,G,L,V1,V2,T1,T2.
+ cpg h rV G L V1 V2 → cpg h rT G L T1 T2 →
+ cpg h ((↓rV)+rT) G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2)
+| cpg_zeta : ∀r,G,L,V,T1,T,T2. cpg h r G (L.ⓓV) T1 T →
+ ⬆*[1] T2 ≡ T → cpg h ((↓r)+𝟙𝟘) G L (+ⓓV.T1) T2
+| cpg_eps : ∀r,G,L,V,T1,T2. cpg h r G L T1 T2 → cpg h ((↓r)+𝟙𝟘) G L (ⓝV.T1) T2
+| cpg_ee : ∀r,G,L,V1,V2,T. cpg h r G L V1 V2 → cpg h ((↓r)+𝟘𝟙) G L (ⓝV1.T) V2
+| cpg_beta : ∀rV,rW,rT,p,G,L,V1,V2,W1,W2,T1,T2.
+ cpg h rV G L V1 V2 → cpg h rW G L W1 W2 → cpg h rT G (L.ⓛW1) T1 T2 →
+ cpg h ((↓rV)+(↓rW)+(↓rT)+𝟙𝟘) G L (ⓐV1.ⓛ{p}W1.T1) (ⓓ{p}ⓝW2.V2.T2)
+| cpg_theta: ∀rV,rW,rT,p,G,L,V1,V,V2,W1,W2,T1,T2.
+ cpg h rV G L V1 V → ⬆*[1] V ≡ V2 → cpg h rW G L W1 W2 →
+ cpg h rT G (L.ⓓW1) T1 T2 →
+ cpg h ((↓rV)+(↓rW)+(↓rT)+𝟙𝟘) G L (ⓐV1.ⓓ{p}W1.T1) (ⓓ{p}W2.ⓐV2.T2)
+.
+
+interpretation
+ "context-sensitive generic parallel rt-transition (term)"
+ 'PRed h r G L T1 T2 = (cpg h r G L T1 T2).
+
+(* Basic properties *********************************************************)
+
+(* Note: this is "∀h,g,L. reflexive … (cpg h (𝟘𝟘) L)" *)
+lemma cpg_refl: ∀h,G,T,L. ⦃G, L⦄ ⊢ T ➡[h, 𝟘𝟘] T.
+#h #G #T elim T -T // * /2 width=1 by cpg_bind, cpg_flat/
+qed.
+
+lemma cpg_pair_sn: ∀h,r,I,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, r] V2 →
+ ∀T. ⦃G, L⦄ ⊢ ②{I}V1.T ➡[h, ↓r] ②{I}V2.T.
+#h #r * /2 width=1 by cpg_bind, cpg_flat/
+qed.
+
+(* Basic inversion lemmas ***************************************************)
+
+fact cpg_inv_atom1_aux: ∀h,r,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, r] T2 → ∀J. T1 = ⓪{J} →
+ ∨∨ T2 = ⓪{J}
+ | ∃∃s. J = Sort s & T2 = ⋆(next h s)
+ | ∃∃rV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[h, rV] V2 & ⬆*[1] V2 ≡ T2 &
+ L = K.ⓓV1 & J = LRef 0 & r = ↓rV
+ | ∃∃rV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[h, rV] V2 & ⬆*[1] V2 ≡ T2 &
+ L = K.ⓛV1 & J = LRef 0 & r = (↓rV)+𝟘𝟙
+ | ∃∃I,K,V,T,i. ⦃G, K⦄ ⊢ #i ➡[h, r] T & ⬆*[1] T ≡ T2 &
+ L = K.ⓑ{I}V & J = LRef (⫯i).
+#h #r #G #L #T1 #T2 * -r -G -L -T1 -T2
+[ #I #G #L #J #H destruct /2 width=1 by or5_intro0/
+| #G #L #s #J #H destruct /3 width=3 by or5_intro1, ex2_intro/
+| #r #G #L #V1 #V2 #W2 #HV12 #VW2 #J #H destruct /3 width=8 by or5_intro2, ex5_4_intro/
+| #r #G #L #V1 #V2 #W2 #HV12 #VW2 #J #H destruct /3 width=8 by or5_intro3, ex5_4_intro/
+| #r #I #G #L #V #T #U #i #HT #HTU #J #H destruct /3 width=9 by or5_intro4, ex4_5_intro/
+| #rV #rT #p #I #G #L #V1 #V2 #T1 #T2 #_ #_ #J #H destruct
+| #rV #rT #I #G #L #V1 #V2 #T1 #T2 #_ #_ #J #H destruct
+| #r #G #L #V #T1 #T #T2 #_ #_ #J #H destruct
+| #r #G #L #V #T1 #T2 #_ #J #H destruct
+| #r #G #L #V1 #V2 #T #_ #J #H destruct
+| #rV #rW #rT #p #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #J #H destruct
+| #rV #rW #rT #p #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #J #H destruct
+]
+qed-.
+
+lemma cpg_inv_atom1: ∀h,r,J,G,L,T2. ⦃G, L⦄ ⊢ ⓪{J} ➡[h, r] T2 →
+ ∨∨ T2 = ⓪{J}
+ | ∃∃s. J = Sort s & T2 = ⋆(next h s)
+ | ∃∃rV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[h, rV] V2 & ⬆*[1] V2 ≡ T2 &
+ L = K.ⓓV1 & J = LRef 0 & r = ↓rV
+ | ∃∃rV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[h, rV] V2 & ⬆*[1] V2 ≡ T2 &
+ L = K.ⓛV1 & J = LRef 0 & r = (↓rV)+𝟘𝟙
+ | ∃∃I,K,V,T,i. ⦃G, K⦄ ⊢ #i ➡[h, r] T & ⬆*[1] T ≡ T2 &
+ L = K.ⓑ{I}V & J = LRef (⫯i).
+/2 width=3 by cpg_inv_atom1_aux/ qed-.
+
+lemma cpg_inv_sort1: ∀h,r,G,L,T2,s. ⦃G, L⦄ ⊢ ⋆s ➡[h, r] T2 →
+ T2 = ⋆s ∨ T2 = ⋆(next h s).
+#h #r #G #L #T2 #s #H
+elim (cpg_inv_atom1 … H) -H /2 width=1 by or_introl/ *
+[ #s0 #H destruct /2 width=1 by or_intror/
+|2,3: #rV #K #V1 #V2 #_ #_ #_ #H destruct
+| #I #K #V1 #V2 #i #_ #_ #_ #H destruct
+]
+qed-.
+
+lemma cpg_inv_zero1: ∀h,r,G,L,T2. ⦃G, L⦄ ⊢ #0 ➡[h, r] T2 →
+ ∨∨ T2 = #0
+ | ∃∃rV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[h, rV] V2 & ⬆*[1] V2 ≡ T2 &
+ L = K.ⓓV1 & r = ↓rV
+ | ∃∃rV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[h, rV] V2 & ⬆*[1] V2 ≡ T2 &
+ L = K.ⓛV1 & r = (↓rV)+𝟘𝟙.
+#h #r #G #L #T2 #H
+elim (cpg_inv_atom1 … H) -H /2 width=1 by or3_intro0/ *
+[ #s #H destruct
+|2,3: #rV #K #V1 #V2 #HV12 #HVT2 #H1 #_ #H2 destruct /3 width=8 by or3_intro1, or3_intro2, ex4_4_intro/
+| #I #K #V1 #V2 #i #_ #_ #_ #H destruct
+]
+qed-.
+
+lemma cpg_inv_lref1: ∀h,r,G,L,T2,i. ⦃G, L⦄ ⊢ #⫯i ➡[h, r] T2 →
+ (T2 = #⫯i) ∨
+ ∃∃I,K,V,T. ⦃G, K⦄ ⊢ #i ➡[h, r] T & ⬆*[1] T ≡ T2 & L = K.ⓑ{I}V.
+#h #r #G #L #T2 #i #H
+elim (cpg_inv_atom1 … H) -H /2 width=1 by or_introl/ *
+[ #s #H destruct
+|2,3: #rV #K #V1 #V2 #_ #_ #_ #H destruct
+| #I #K #V1 #V2 #j #HV2 #HVT2 #H1 #H2 destruct /3 width=7 by ex3_4_intro, or_intror/
+]
+qed-.
+
+lemma cpg_inv_gref1: ∀h,r,G,L,T2,l. ⦃G, L⦄ ⊢ §l ➡[h, r] T2 → T2 = §l.
+#h #r #G #L #T2 #l #H
+elim (cpg_inv_atom1 … H) -H // *
+[ #s #H destruct
+|2,3: #rV #K #V1 #V2 #_ #_ #_ #H destruct
+| #I #K #V1 #V2 #i #_ #_ #_ #H destruct
+]
+qed-.
+
+fact cpg_inv_bind1_aux: ∀h,r,G,L,U,U2. ⦃G, L⦄ ⊢ U ➡[h, r] U2 →
+ ∀p,J,V1,U1. U = ⓑ{p,J}V1.U1 → (
+ ∃∃rV,rT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V2 & ⦃G, L.ⓑ{J}V1⦄ ⊢ U1 ➡[h, rT] T2 &
+ U2 = ⓑ{p,J}V2.T2 & r = (↓rV)+rT
+ ) ∨
+ ∃∃rT,T. ⦃G, L.ⓓV1⦄ ⊢ U1 ➡[h, rT] T & ⬆*[1] U2 ≡ T &
+ p = true & J = Abbr & r = (↓rT)+𝟙𝟘.
+#h #r #G #L #U #U2 * -r -G -L -U -U2
+[ #I #G #L #q #J #W #U1 #H destruct
+| #G #L #s #q #J #W #U1 #H destruct
+| #r #G #L #V1 #V2 #W2 #_ #_ #q #J #W #U1 #H destruct
+| #r #G #L #V1 #V2 #W2 #_ #_ #q #J #W #U1 #H destruct
+| #r #I #G #L #V #T #U #i #_ #_ #q #J #W #U1 #H destruct
+| #rv #rT #p #I #G #L #V1 #V2 #T1 #T2 #HV12 #HT12 #q #J #W #U1 #H destruct /3 width=8 by ex4_4_intro, or_introl/
+| #rv #rT #I #G #L #V1 #V2 #T1 #T2 #_ #_ #q #J #W #U1 #H destruct
+| #r #G #L #V #T1 #T #T2 #HT1 #HT2 #q #J #W #U1 #H destruct /3 width=5 by ex5_2_intro, or_intror/
+| #r #G #L #V #T1 #T2 #_ #q #J #W #U1 #H destruct
+| #r #G #L #V1 #V2 #T #_ #q #J #W #U1 #H destruct
+| #rv #rW #rT #p #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #q #J #W #U1 #H destruct
+| #rv #rW #rT #p #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #q #J #W #U1 #H destruct
+]
+qed-.
+
+lemma cpg_inv_bind1: ∀h,r,p,I,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓑ{p,I}V1.T1 ➡[h, r] U2 → (
+ ∃∃rV,rT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V2 & ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ➡[h, rT] T2 &
+ U2 = ⓑ{p,I}V2.T2 & r = (↓rV)+rT
+ ) ∨
+ ∃∃rT,T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[h, rT] T & ⬆*[1] U2 ≡ T &
+ p = true & I = Abbr & r = (↓rT)+𝟙𝟘.
+/2 width=3 by cpg_inv_bind1_aux/ qed-.
+
+lemma cpg_inv_abbr1: ∀h,r,p,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{p}V1.T1 ➡[h, r] U2 → (
+ ∃∃rV,rT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V2 & ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[h, rT] T2 &
+ U2 = ⓓ{p}V2.T2 & r = (↓rV)+rT
+ ) ∨
+ ∃∃rT,T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[h, rT] T & ⬆*[1] U2 ≡ T &
+ p = true & r = (↓rT)+𝟙𝟘.
+#h #r #p #G #L #V1 #T1 #U2 #H elim (cpg_inv_bind1 … H) -H *
+/3 width=8 by ex4_4_intro, ex4_2_intro, or_introl, or_intror/
+qed-.
+
+lemma cpg_inv_abst1: ∀h,r,p,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{p}V1.T1 ➡[h, r] U2 →
+ ∃∃rV,rT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 ➡[h, rT] T2 &
+ U2 = ⓛ{p} V2. T2 & r = (↓rV)+rT.
+#h #r #p #G #L #V1 #T1 #U2 #H elim (cpg_inv_bind1 … H) -H *
+[ /3 width=8 by ex4_4_intro/
+| #r #T #_ #_ #_ #H destruct
+]
+qed-.
+
+fact cpg_inv_flat1_aux: ∀h,r,G,L,U,U2. ⦃G, L⦄ ⊢ U ➡[h, r] U2 →
+ ∀J,V1,U1. U = ⓕ{J}V1.U1 →
+ ∨∨ ∃∃rV,rT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, rT] T2 &
+ U2 = ⓕ{J}V2.T2 & r = (↓rV)+rT
+ | ∃∃rT. ⦃G, L⦄ ⊢ U1 ➡[h, rT] U2 & J = Cast & r = (↓rT)+𝟙𝟘
+ | ∃∃rV. ⦃G, L⦄ ⊢ V1 ➡[h, rV] U2 & J = Cast & r = (↓rV)+𝟘𝟙
+ | ∃∃rV,rW,rT,p,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V2 & ⦃G, L⦄ ⊢ W1 ➡[h, rW] W2 & ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[h, rT] T2 &
+ J = Appl & U1 = ⓛ{p}W1.T1 & U2 = ⓓ{p}ⓝW2.V2.T2 & r = (↓rV)+(↓rW)+(↓rT)+𝟙𝟘
+ | ∃∃rV,rW,rT,p,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V & ⬆*[1] V ≡ V2 & ⦃G, L⦄ ⊢ W1 ➡[h, rW] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[h, rT] T2 &
+ J = Appl & U1 = ⓓ{p}W1.T1 & U2 = ⓓ{p}W2.ⓐV2.T2 & r = (↓rV)+(↓rW)+(↓rT)+𝟙𝟘.
+#h #r #G #L #U #U2 * -r -G -L -U -U2
+[ #I #G #L #J #W #U1 #H destruct
+| #G #L #s #J #W #U1 #H destruct
+| #r #G #L #V1 #V2 #W2 #_ #_ #J #W #U1 #H destruct
+| #r #G #L #V1 #V2 #W2 #_ #_ #J #W #U1 #H destruct
+| #r #I #G #L #V #T #U #i #_ #_ #J #W #U1 #H destruct
+| #rv #rT #p #I #G #L #V1 #V2 #T1 #T2 #_ #_ #J #W #U1 #H destruct
+| #rv #rT #I #G #L #V1 #V2 #T1 #T2 #HV12 #HT12 #J #W #U1 #H destruct /3 width=8 by or5_intro0, ex4_4_intro/
+| #r #G #L #V #T1 #T #T2 #_ #_ #J #W #U1 #H destruct
+| #r #G #L #V #T1 #T2 #HT12 #J #W #U1 #H destruct /3 width=3 by or5_intro1, ex3_intro/
+| #r #G #L #V1 #V2 #T #HV12 #J #W #U1 #H destruct /3 width=3 by or5_intro2, ex3_intro/
+| #rv #rW #rT #p #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HW12 #HT12 #J #W #U1 #H destruct /3 width=15 by or5_intro3, ex7_9_intro/
+| #rv #rW #rT #p #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 #HW12 #HT12 #J #W #U1 #H destruct /3 width=17 by or5_intro4, ex8_10_intro/
+]
+qed-.
+
+lemma cpg_inv_flat1: ∀h,r,I,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓕ{I}V1.U1 ➡[h, r] U2 →
+ ∨∨ ∃∃rV,rT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, rT] T2 &
+ U2 = ⓕ{I}V2.T2 & r = (↓rV)+rT
+ | ∃∃rT. ⦃G, L⦄ ⊢ U1 ➡[h, rT] U2 & I = Cast & r = (↓rT)+𝟙𝟘
+ | ∃∃rV. ⦃G, L⦄ ⊢ V1 ➡[h, rV] U2 & I = Cast & r = (↓rV)+𝟘𝟙
+ | ∃∃rV,rW,rT,p,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V2 & ⦃G, L⦄ ⊢ W1 ➡[h, rW] W2 & ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[h, rT] T2 &
+ I = Appl & U1 = ⓛ{p}W1.T1 & U2 = ⓓ{p}ⓝW2.V2.T2 & r = (↓rV)+(↓rW)+(↓rT)+𝟙𝟘
+ | ∃∃rV,rW,rT,p,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V & ⬆*[1] V ≡ V2 & ⦃G, L⦄ ⊢ W1 ➡[h, rW] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[h, rT] T2 &
+ I = Appl & U1 = ⓓ{p}W1.T1 & U2 = ⓓ{p}W2.ⓐV2.T2 & r = (↓rV)+(↓rW)+(↓rT)+𝟙𝟘.
+/2 width=3 by cpg_inv_flat1_aux/ qed-.
+
+lemma cpg_inv_appl1: ∀h,r,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓐV1.U1 ➡[h, r] U2 →
+ ∨∨ ∃∃rV,rT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, rT] T2 &
+ U2 = ⓐV2.T2 & r = (↓rV)+rT
+ | ∃∃rV,rW,rT,p,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V2 & ⦃G, L⦄ ⊢ W1 ➡[h, rW] W2 & ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[h, rT] T2 &
+ U1 = ⓛ{p}W1.T1 & U2 = ⓓ{p}ⓝW2.V2.T2 & r = (↓rV)+(↓rW)+(↓rT)+𝟙𝟘
+ | ∃∃rV,rW,rT,p,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V & ⬆*[1] V ≡ V2 & ⦃G, L⦄ ⊢ W1 ➡[h, rW] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[h, rT] T2 &
+ U1 = ⓓ{p}W1.T1 & U2 = ⓓ{p}W2.ⓐV2.T2 & r = (↓rV)+(↓rW)+(↓rT)+𝟙𝟘.
+#h #r #G #L #V1 #U1 #U2 #H elim (cpg_inv_flat1 … H) -H *
+[ /3 width=8 by or3_intro0, ex4_4_intro/
+|2,3: #r #_ #H destruct
+| /3 width=15 by or3_intro1, ex6_9_intro/
+| /3 width=17 by or3_intro2, ex7_10_intro/
+]
+qed-.
+
+lemma cpg_inv_cast1: ∀h,r,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓝV1.U1 ➡[h, r] U2 →
+ ∨∨ ∃∃rV,rT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, rT] T2 &
+ U2 = ⓝV2.T2 & r = (↓rV)+rT
+ | ∃∃rT. ⦃G, L⦄ ⊢ U1 ➡[h, rT] U2 & r = (↓rT)+𝟙𝟘
+ | ∃∃rV. ⦃G, L⦄ ⊢ V1 ➡[h, rV] U2 & r = (↓rV)+𝟘𝟙.
+#h #r #G #L #V1 #U1 #U2 #H elim (cpg_inv_flat1 … H) -H *
+[ /3 width=8 by or3_intro0, ex4_4_intro/
+|2,3: /3 width=3 by or3_intro1, or3_intro2, ex2_intro/
+| #rv #rW #rT #p #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #H destruct
+| #rv #rW #rT #p #V #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #H destruct
+]
+qed-.
+
+(* Basic forward lemmas *****************************************************)
+
+lemma cpg_fwd_bind1_minus: ∀h,r,I,G,L,V1,T1,T. ⦃G, L⦄ ⊢ -ⓑ{I}V1.T1 ➡[h, r] T → ∀b.
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ ⓑ{b,I}V1.T1 ➡[h, r] ⓑ{b,I}V2.T2 &
+ T = -ⓑ{I}V2.T2.
+#h #r #I #G #L #V1 #T1 #T #H #b elim (cpg_inv_bind1 … H) -H *
+[ #rV #rT #V2 #T2 #HV12 #HT12 #H1 #H2 destruct /3 width=4 by cpg_bind, ex2_2_intro/
+| #r #T2 #_ #_ #H destruct
+]
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/relocation/drops.ma".
+include "basic_2/rt_transition/cpg.ma".
+
+(* CONTEXT-SENSITIVE GENERIC PARALLEL RT-TRANSITION FOR TERMS ***************)
+
+(* Properties with generic slicing for local environments *******************)
+
+lemma cpg_delift: ∀h,r,I,G,K,V,T1,L,l. ⬇*[i] L ≡ (K.ⓑ{I}V) →
+ ∃∃T2,T. ⦃G, L⦄ ⊢ T1 ➡[h, 𝟘𝟘] T2 & ⬆*[↑1] T ≡ T2.
+#h #r #I #G #K #V #T1 elim T1 -T1
+[ * #i #L #l /2 width=4 by cpg_atom, lift_sort, lift_gref, ex2_2_intro/
+ elim (lt_or_eq_or_gt i l) #Hil [1,3: /4 width=4 by cpg_atom, lift_lref_ge_minus, lift_lref_lt, ylt_inj, yle_inj, ex2_2_intro/ ]
+ destruct
+ elim (lift_total V 0 (i+1)) #W #HVW
+ elim (lift_split … HVW i i) /3 width=7 by cpg_delta, ex2_2_intro/
+| * [ #a ] #I #W1 #U1 #IHW1 #IHU1 #L #l #HLK
+ elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2
+ [ elim (IHU1 (L. ⓑ{I} W1) (l+1)) -IHU1 /3 width=9 by cpg_bind, drop_drop, lift_bind, ex2_2_intro/
+ | elim (IHU1 … HLK) -IHU1 -HLK /3 width=8 by cpg_flat, lift_flat, ex2_2_intro/
+ ]
+]
+qed-.
+
+lemma cpg_inv_lref1: ∀h,r,G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡[h, r] T2 →
+ T2 = #i ∨
+ ∃∃I,K,V,V2. ⬇[i] L ≡ K. ⓑ{I}V & ⦃G, K⦄ ⊢ V ➡[h, r] V2 &
+ ⬆[O, i+1] V2 ≡ T2.
+#h #r #G #L #T2 #i #H
+elim (cpg_inv_atom1 … H) -H /2 width=1 by or_introl/ *
+[ #s #d #_ #_ #H destruct
+| #I #K #V #V2 #j #HLK #HV2 #HVT2 #H destruct /3 width=7 by ex3_4_intro, or_intror/
+]
+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/drops.ma".
+include "basic_2/rt_transition/cpg.ma".
+
+(* CONTEXT-SENSITIVE GENERIC PARALLEL RT-TRANSITION FOR TERMS ***************)
+
+(* Properties with length for local environments ****************************)
+
+lemma cpg_inv_lref1_ge: ∀h,r,G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡[h, r] T2 → |L| ≤ i → T2 = #i.
+#h #r #G #L #T2 #i #H elim (cpg_inv_lref1 … H) -H // *
+#I #K #V1 #V2 #HLK #_ #_ #HL -h -G -V2 lapply (drop_fwd_length_lt2 … HLK) -K -I -V1
+#H elim (lt_refl_false i) /2 width=3 by lt_to_le_to_lt/
+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/lsubr.ma".
+include "basic_2/rt_transition/cpg.ma".
+
+(* CONTEXT-SENSITIVE GENERIC PARALLEL RT-TRANSITION FOR TERMS ***************)
+
+(* Properties with restricted refinement for local environments *************)
+
+lemma lsubr_cpg_trans: ∀h,r,G. lsub_trans … (cpg h r G) lsubr.
+#h #r #G #L1 #T1 #T2 #H elim H -r -G -L1 -T1 -T2
+[ //
+| /2 width=2 by cpg_st/
+| #r #G #L1 #V1 #V2 #W2 #_ #HVW2 #IH #X #H
+ elim (lsubr_inv_abbr2 … H) -H #L2 #HL21 #H destruct
+ /3 width=3 by cpg_delta/
+| #r #G #L1 #V1 #V2 #W2 #_ #HVW2 #IH #X #H
+ elim (lsubr_inv_abst2 … H) -H * #L2 [2: #V ] #HL21 #H destruct
+ /4 width=3 by cpg_delta, cpg_lt, cpg_ct/
+| #r #I1 #G #L1 #V1 #T1 #U1 #i #_ #HTU1 #IH #X #H
+ elim (lsubr_fwd_pair2 … H) -H #I2 #L2 #V2 #HL21 #H destruct
+ /3 width=3 by cpt_lref/
+|6,11: /4 width=1 by cpg_bind, cpg_beta, lsubr_pair/
+|7,9,10: /3 width=1 by cpg_flat, cpg_eps, cpg_ct/
+|8,12: /4 width=3 by cpg_zeta, cpg_theta, lsubr_pair/
+]
+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/drops.ma".
+include "basic_2/rt_transition/cpg.ma".
+
+(* CONTEXT-SENSITIVE GENERIC PARALLEL RT-TRANSITION FOR TERMS ***************)
+
+(* Properties with generic slicing for local environments *******************)
+
+(* Note: the main property of simple terms *)
+lemma cpg_inv_appl1_simple: ∀h,r,G,L,V1,T1,U. ⦃G, L⦄ ⊢ ⓐV1.T1 ➡[h, r] U → 𝐒⦃T1⦄ →
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, r] V2 & ⦃G, L⦄ ⊢ T1 ➡[h, r] T2 &
+ U = ⓐV2.T2.
+#h #r #G #L #V1 #T1 #U #H #HT1
+elim (cpg_inv_appl1 … H) -H *
+[ /2 width=5 by ex3_2_intro/
+| #a #V2 #W1 #W2 #U1 #U2 #_ #_ #_ #H #_ destruct
+ elim (simple_inv_bind … HT1)
+| #a #V #V2 #W1 #W2 #U1 #U2 #_ #_ #_ #_ #H #_ destruct
+ elim (simple_inv_bind … HT1)
+]
+qed-.
.
interpretation
- "local environment refinement (restricted)"
+ "restricted refinement (local environment)"
'LRSubEqC L1 L2 = (lsubr L1 L2).
(* Basic properties *********************************************************)
L2 = ⋆ ∨ ∃∃K2. K1 ⫃ K2 & L2 = K2.ⓛW.
/2 width=3 by lsubr_inv_abst1_aux/ qed-.
-fact lsubr_inv_abbr2_aux: ∀L1,L2. L1 ⫃ L2 → ∀K2,W. L2 = K2.ⓓW →
- ∃∃K1. K1 ⫃ K2 & L1 = K1.ⓓW.
+fact lsubr_inv_abbr2_aux: ∀L1,L2. L1 ⫃ L2 → ∀K2,V. L2 = K2.ⓓV →
+ ∃∃K1. K1 ⫃ K2 & L1 = K1.ⓓV.
#L1 #L2 * -L1 -L2
[ #L #K2 #W #H destruct
| #I #L1 #L2 #V #HL12 #K2 #W #H destruct /2 width=3 by ex2_intro/
]
qed-.
-lemma lsubr_inv_abbr2: ∀L1,K2,W. L1 ⫃ K2.ⓓW →
- ∃∃K1. K1 ⫃ K2 & L1 = K1.ⓓW.
+lemma lsubr_inv_abbr2: ∀L1,K2,V. L1 ⫃ K2.ⓓV →
+ ∃∃K1. K1 ⫃ K2 & L1 = K1.ⓓV.
/2 width=3 by lsubr_inv_abbr2_aux/ qed-.
+
+fact lsubr_inv_abst2_aux: ∀L1,L2. L1 ⫃ L2 → ∀K2,W. L2 = K2.ⓛW →
+ (∃∃K1. K1 ⫃ K2 & L1 = K1.ⓛW) ∨
+ ∃∃K1,V. K1 ⫃ K2 & L1 = K1.ⓓⓝW.V.
+#L1 #L2 * -L1 -L2
+[ #L #K2 #W #H destruct
+| #I #L1 #L2 #V #HL12 #K2 #W #H destruct /3 width=3 by ex2_intro, or_introl/
+| #L1 #L2 #V1 #V2 #HL12 #K2 #W #H destruct /3 width=4 by ex2_2_intro, or_intror/
+]
+qed-.
+
+lemma lsubr_inv_abst2: ∀L1,K2,W. L1 ⫃ K2.ⓛW →
+ (∃∃K1. K1 ⫃ K2 & L1 = K1.ⓛW) ∨
+ ∃∃K1,V. K1 ⫃ K2 & L1 = K1.ⓓⓝW.V.
+/2 width=3 by lsubr_inv_abst2_aux/ qed-.
+
+(* Basic forward lemmas *****************************************************)
+
+lemma lsubr_fwd_pair2: ∀I2,L1,K2,V2. L1 ⫃ K2.ⓑ{I2}V2 →
+ ∃∃I1,K1,V1. K1 ⫃ K2 & L1 = K1.ⓑ{I1}V1.
+* #L1 #K2 #V2 #H
+[ elim (lsubr_inv_abbr2 … H) | elim (lsubr_inv_abst2 … H) * ] -H
+/2 width=5 by ex2_3_intro/
+qed-.
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.λ${ident x6}:$T6.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P4) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P5) }.
+(* multiple existental quantifier (6, 9) *)
+
+notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 , ident x6 , ident x7 , ident x8 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5)"
+ non associative with precedence 20
+ for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.$P3) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.$P4) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.$P5) }.
+
+notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 , ident x6 , ident x7 , ident x8 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5)"
+ 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.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.$P4) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.$P5) }.
+
(* multiple existental quantifier (7, 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)"
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.λ${ident x6}:$T6.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P4) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P5) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P6) }.
+(* multiple existental quantifier (7, 9) *)
+
+notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 , ident x6 , ident x7 , ident x8 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}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.$P3) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.$P4) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.$P5) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.$P6) }.
+
+notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 , ident x6 , ident x7 , ident x8 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.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.$P4) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.$P5) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.$P6) }.
+
+(* multiple existental quantifier (7, 10) *)
+
+notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 , ident x6 , ident x7 , ident x8 , ident x9 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}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.λ${ident x9}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.λ${ident x9}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.λ${ident x9}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.λ${ident x9}.$P3) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.λ${ident x9}.$P4) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.λ${ident x9}.$P5) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.λ${ident x9}.$P6) }.
+
+notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 , ident x6 , ident x7 , ident x8 , ident x9 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.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.λ${ident x9}:$T9.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.λ${ident x9}:$T9.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.λ${ident x9}:$T9.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.λ${ident x9}:$T9.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.λ${ident x9}:$T9.$P4) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.λ${ident x9}:$T9.$P5) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.λ${ident x9}:$T9.$P6) }.
+
(* multiple existental quantifier (8, 4) *)
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)"
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, 10) *)
+
+notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 , ident x6 , ident x7 , ident x8 , ident x9 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)"
+ non associative with precedence 20
+ for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.λ${ident x9}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.λ${ident x9}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.λ${ident x9}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.λ${ident x9}.$P3) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.λ${ident x9}.$P4) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.λ${ident x9}.$P5) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.λ${ident x9}.$P6) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.λ${ident x7}.λ${ident x8}.λ${ident x9}.$P7) }.
+
+notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 , ident x6 , ident x7 , ident x8 , ident x9 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)"
+ 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.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.λ${ident x9}:$T9.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.λ${ident x9}:$T9.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.λ${ident x9}:$T9.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.λ${ident x9}:$T9.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.λ${ident x9}:$T9.$P4) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.λ${ident x9}:$T9.$P5) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.λ${ident x9}:$T9.$P6) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.λ${ident x7}:$T7.λ${ident x8}:$T8.λ${ident x9}:$T9.$P7) }.
+
(* multiple disjunction connective (3) *)
notation "hvbox(∨∨ term 29 P0 break | term 29 P1 break | term 29 P2)"
interpretation "constructor (rtc)"
'Tuple ri rh ti th = (mk_rtc ri rh ti th).
-(* Note: for one structural step *)
-definition OO: rtc ≝ 〈0, 0, 0, 0〉.
-
interpretation "one structural step (rtc)"
- 'ZeroZero = (OO).
-
-(* Note: for one r-step *)
-definition UO: rtc ≝ 〈0, 1, 0, 0〉.
+ 'ZeroZero = (mk_rtc O O O O).
interpretation "one r-step (rtc)"
- 'OneZero = (UO).
-
-(* Note: for one t-step *)
-definition OU: rtc ≝ 〈0, 0, 0, 1〉.
+ 'OneZero = (mk_rtc O (S O) O O).
interpretation "one t-step (rtc)"
- 'ZeroOne = (OU).
+ 'ZeroOne = (mk_rtc O O O (S O)).
interpretation "plus (rtc)"
'plus r1 r2 = (plus r1 r2).
+
+(* Basic properties *********************************************************)
+
+lemma plus_OO_r: ∀r. r = 𝟘𝟘 + r.
+* normalize //
+qed.
+
+lemma plus_r_OO: ∀r. r = r + 𝟘𝟘.
+* normalize //
+qed.
interpretation "shift (rtc)"
'Drop r = (shift r).
+
+(* Basic properties *********************************************************)
+
+lemma shift_refl: ∀ri,ti. 〈ri, 0, ti, 0〉 = ↓〈ri, 0, ti, 0〉.
+normalize //
+qed.
}
]
class "water"
- [ { "rt-transition counter" * } {
+ [ { "generic rt-transition counter" * } {
[ { "" * } {
[ "rtc ( 〈?,?,?,?〉 ) ( 𝟘𝟘 ) ( 𝟙𝟘 ) ( 𝟘𝟙 )" "rtc_shift ( ↓? )" "rtc_plus ( ? + ? )" * ]
}
<key name="ex">6 5</key>
<key name="ex">6 6</key>
<key name="ex">6 7</key>
+ <key name="ex">6 9</key>
<key name="ex">7 3</key>
<key name="ex">7 4</key>
<key name="ex">7 7</key>
+ <key name="ex">7 9</key>
+ <key name="ex">7 10</key>
<key name="ex">8 4</key>
<key name="ex">8 5</key>
+ <key name="ex">8 10</key>
<key name="or">3</key>
<key name="or">4</key>
<key name="or">5</key>
interpretation "multiple existental quantifier (6, 7)" 'Ex P0 P1 P2 P3 P4 P5 = (ex6_7 ? ? ? ? ? ? ? P0 P1 P2 P3 P4 P5).
+(* multiple existental quantifier (6, 9) *)
+
+inductive ex6_9 (A0,A1,A2,A3,A4,A5,A6,A7,A8:Type[0]) (P0,P1,P2,P3,P4,P5:A0→A1→A2→A3→A4→A5→A6→A7→A8→Prop) : Prop ≝
+ | ex6_9_intro: ∀x0,x1,x2,x3,x4,x5,x6,x7,x8. P0 x0 x1 x2 x3 x4 x5 x6 x7 x8 → P1 x0 x1 x2 x3 x4 x5 x6 x7 x8 → P2 x0 x1 x2 x3 x4 x5 x6 x7 x8 → P3 x0 x1 x2 x3 x4 x5 x6 x7 x8 → P4 x0 x1 x2 x3 x4 x5 x6 x7 x8 → P5 x0 x1 x2 x3 x4 x5 x6 x7 x8 → ex6_9 ? ? ? ? ? ? ? ? ? ? ? ? ? ? ?
+.
+
+interpretation "multiple existental quantifier (6, 9)" 'Ex P0 P1 P2 P3 P4 P5 = (ex6_9 ? ? ? ? ? ? ? ? ? P0 P1 P2 P3 P4 P5).
+
(* multiple existental quantifier (7, 3) *)
inductive ex7_3 (A0,A1,A2:Type[0]) (P0,P1,P2,P3,P4,P5,P6:A0→A1→A2→Prop) : Prop ≝
interpretation "multiple existental quantifier (7, 7)" 'Ex P0 P1 P2 P3 P4 P5 P6 = (ex7_7 ? ? ? ? ? ? ? P0 P1 P2 P3 P4 P5 P6).
+(* multiple existental quantifier (7, 9) *)
+
+inductive ex7_9 (A0,A1,A2,A3,A4,A5,A6,A7,A8:Type[0]) (P0,P1,P2,P3,P4,P5,P6:A0→A1→A2→A3→A4→A5→A6→A7→A8→Prop) : Prop ≝
+ | ex7_9_intro: ∀x0,x1,x2,x3,x4,x5,x6,x7,x8. P0 x0 x1 x2 x3 x4 x5 x6 x7 x8 → P1 x0 x1 x2 x3 x4 x5 x6 x7 x8 → P2 x0 x1 x2 x3 x4 x5 x6 x7 x8 → P3 x0 x1 x2 x3 x4 x5 x6 x7 x8 → P4 x0 x1 x2 x3 x4 x5 x6 x7 x8 → P5 x0 x1 x2 x3 x4 x5 x6 x7 x8 → P6 x0 x1 x2 x3 x4 x5 x6 x7 x8 → ex7_9 ? ? ? ? ? ? ? ? ? ? ? ? ? ? ? ?
+.
+
+interpretation "multiple existental quantifier (7, 9)" 'Ex P0 P1 P2 P3 P4 P5 P6 = (ex7_9 ? ? ? ? ? ? ? ? ? P0 P1 P2 P3 P4 P5 P6).
+
+(* multiple existental quantifier (7, 10) *)
+
+inductive ex7_10 (A0,A1,A2,A3,A4,A5,A6,A7,A8,A9:Type[0]) (P0,P1,P2,P3,P4,P5,P6:A0→A1→A2→A3→A4→A5→A6→A7→A8→A9→Prop) : Prop ≝
+ | ex7_10_intro: ∀x0,x1,x2,x3,x4,x5,x6,x7,x8,x9. P0 x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 → P1 x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 → P2 x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 → P3 x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 → P4 x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 → P5 x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 → P6 x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 → ex7_10 ? ? ? ? ? ? ? ? ? ? ? ? ? ? ? ? ?
+.
+
+interpretation "multiple existental quantifier (7, 10)" 'Ex P0 P1 P2 P3 P4 P5 P6 = (ex7_10 ? ? ? ? ? ? ? ? ? ? P0 P1 P2 P3 P4 P5 P6).
+
(* multiple existental quantifier (8, 4) *)
inductive ex8_4 (A0,A1,A2,A3:Type[0]) (P0,P1,P2,P3,P4,P5,P6,P7:A0→A1→A2→A3→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, 10) *)
+
+inductive ex8_10 (A0,A1,A2,A3,A4,A5,A6,A7,A8,A9:Type[0]) (P0,P1,P2,P3,P4,P5,P6,P7:A0→A1→A2→A3→A4→A5→A6→A7→A8→A9→Prop) : Prop ≝
+ | ex8_10_intro: ∀x0,x1,x2,x3,x4,x5,x6,x7,x8,x9. P0 x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 → P1 x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 → P2 x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 → P3 x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 → P4 x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 → P5 x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 → P6 x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 → P7 x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 → ex8_10 ? ? ? ? ? ? ? ? ? ? ? ? ? ? ? ? ? ?
+.
+
+interpretation "multiple existental quantifier (8, 10)" 'Ex P0 P1 P2 P3 P4 P5 P6 P7 = (ex8_10 ? ? ? ? ? ? ? ? ? ? P0 P1 P2 P3 P4 P5 P6 P7).
+
(* multiple disjunction connective (3) *)
inductive or3 (P0,P1,P2:Prop) : Prop ≝