(* Basic properties *********************************************************)
(* Basic_1: was just: nf2_sn3 *)
-lemma csx_cpre: ∀h,g,G,L,T1. ⦃G, L⦄ ⊢ ⬊*[h, g] T1 → ∃T2. ⦃G, L⦄ ⊢ T1 ➡* 𝐍⦃T2⦄.
-#h #g #G #L #T1 #H @(csx_ind … H) -T1
+lemma csx_cpre: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬊*[h, o] T1 → ∃T2. ⦃G, L⦄ ⊢ T1 ➡* 𝐍⦃T2⦄.
+#h #o #G #L #T1 #H @(csx_ind … H) -T1
#T1 #_ #IHT1 elim (cnr_dec G L T1) /3 width=3 by ex_intro, conj/
* #T #H1T1 #H2T1 elim (IHT1 … H2T1) -IHT1 -H2T1 /2 width=2 by cpr_cpx/
#T2 * /4 width=3 by cprs_strap2, ex_intro, conj/
(* Basic inversion lemmas ***************************************************)
(* Basic_1: was: pr3_gen_sort *)
-lemma cprs_inv_sort1: ∀G,L,U2,k. ⦃G, L⦄ ⊢ ⋆k ➡* U2 → U2 = ⋆k.
-#G #L #U2 #k #H @(cprs_ind … H) -U2 //
+lemma cprs_inv_sort1: ∀G,L,U2,s. ⦃G, L⦄ ⊢ ⋆s ➡* U2 → U2 = ⋆s.
+#G #L #U2 #s #H @(cprs_ind … H) -U2 //
#U2 #U #_ #HU2 #IHU2 destruct
>(cpr_inv_sort1 … HU2) -HU2 //
qed-.
(* Basic_1: was just: pr3_pr2_pr2_t *)
(* Basic_1: includes: pr3_pr0_pr2_t *)
-lemma lpr_cpr_trans: ∀G. s_r_transitive … (cpr G) (λ_. lpr G).
+lemma lpr_cpr_trans: ∀G. c_r_transitive … (cpr G) (λ_. lpr G).
#G #L2 #T1 #T2 #HT12 elim HT12 -G -L2 -T1 -T2
[ /2 width=3 by/
| #G #L2 #K2 #V0 #V2 #W2 #i #HLK2 #_ #HVW2 #IHV02 #L1 #HL12
(* Advanced properties ******************************************************)
(* Basic_1: was only: pr3_pr2_pr3_t pr3_wcpr0_t *)
-lemma lpr_cprs_trans: ∀G. s_rs_transitive … (cpr G) (λ_. lpr G).
-#G @s_r_trans_LTC1 /2 width=3 by lpr_cpr_trans/ (**) (* full auto fails *)
+lemma lpr_cprs_trans: ∀G. c_rs_transitive … (cpr G) (λ_. lpr G).
+#G @c_r_trans_LTC1 /2 width=3 by lpr_cpr_trans/ (**) (* full auto fails *)
qed-.
(* Basic_1: was: pr3_strip *)
(* EVALUATION FOR CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION ON TERMS ****)
definition cpxe: ∀h. sd h → relation4 genv lenv term term ≝
- λh,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 ∧ ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T2⦄.
+ λh,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 ∧ ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T2⦄.
interpretation "evaluation for context-sensitive extended parallel reduction (term)"
- 'PRedEval h g G L T1 T2 = (cpxe h g G L T1 T2).
+ 'PRedEval h o G L T1 T2 = (cpxe h o G L T1 T2).
(* Basic properties *********************************************************)
-lemma csx_cpxe: ∀h,g,G,L,T1. ⦃G, L⦄ ⊢ ⬊*[h, g] T1 → ∃T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] 𝐍⦃T2⦄.
-#h #g #G #L #T1 #H @(csx_ind … H) -T1
-#T1 #_ #IHT1 elim (cnx_dec h g G L T1) /3 width=3 by ex_intro, conj/
+lemma csx_cpxe: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬊*[h, o] T1 → ∃T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] 𝐍⦃T2⦄.
+#h #o #G #L #T1 #H @(csx_ind … H) -T1
+#T1 #_ #IHT1 elim (cnx_dec h o G L T1) /3 width=3 by ex_intro, conj/
* #T #H1T1 #H2T1 elim (IHT1 … H1T1 H2T1) -IHT1 -H2T1
#T2 * /4 width=3 by cpxs_strap2, ex_intro, conj/
qed-.
(* CONTEXT-SENSITIVE EXTENDED PARALLEL COMPUTATION ON TERMS *****************)
definition cpxs: ∀h. sd h → relation4 genv lenv term term ≝
- λh,g,G. LTC … (cpx h g G).
+ λh,o,G. LTC … (cpx h o G).
interpretation "extended context-sensitive parallel computation (term)"
- 'PRedStar h g G L T1 T2 = (cpxs h g G L T1 T2).
+ 'PRedStar h o G L T1 T2 = (cpxs h o G L T1 T2).
(* Basic eliminators ********************************************************)
-lemma cpxs_ind: ∀h,g,G,L,T1. ∀R:predicate term. R T1 →
- (∀T,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T → ⦃G, L⦄ ⊢ T ➡[h, g] T2 → R T → R T2) →
- ∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → R T2.
-#h #g #L #G #T1 #R #HT1 #IHT1 #T2 #HT12
+lemma cpxs_ind: ∀h,o,G,L,T1. ∀R:predicate term. R T1 →
+ (∀T,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T → ⦃G, L⦄ ⊢ T ➡[h, o] T2 → R T → R T2) →
+ ∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → R T2.
+#h #o #L #G #T1 #R #HT1 #IHT1 #T2 #HT12
@(TC_star_ind … HT1 IHT1 … HT12) //
qed-.
-lemma cpxs_ind_dx: ∀h,g,G,L,T2. ∀R:predicate term. R T2 →
- (∀T1,T. ⦃G, L⦄ ⊢ T1 ➡[h, g] T → ⦃G, L⦄ ⊢ T ➡*[h, g] T2 → R T → R T1) →
- ∀T1. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → R T1.
-#h #g #G #L #T2 #R #HT2 #IHT2 #T1 #HT12
+lemma cpxs_ind_dx: ∀h,o,G,L,T2. ∀R:predicate term. R T2 →
+ (∀T1,T. ⦃G, L⦄ ⊢ T1 ➡[h, o] T → ⦃G, L⦄ ⊢ T ➡*[h, o] T2 → R T → R T1) →
+ ∀T1. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → R T1.
+#h #o #G #L #T2 #R #HT2 #IHT2 #T1 #HT12
@(TC_star_ind_dx … HT2 IHT2 … HT12) //
qed-.
(* Basic properties *********************************************************)
-lemma cpxs_refl: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ T ➡*[h, g] T.
+lemma cpxs_refl: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ T ➡*[h, o] T.
/2 width=1 by inj/ qed.
-lemma cpx_cpxs: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2.
+lemma cpx_cpxs: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2.
/2 width=1 by inj/ qed.
-lemma cpxs_strap1: ∀h,g,G,L,T1,T. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T →
- ∀T2. ⦃G, L⦄ ⊢ T ➡[h, g] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2.
+lemma cpxs_strap1: ∀h,o,G,L,T1,T. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T →
+ ∀T2. ⦃G, L⦄ ⊢ T ➡[h, o] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2.
normalize /2 width=3 by step/ qed.
-lemma cpxs_strap2: ∀h,g,G,L,T1,T. ⦃G, L⦄ ⊢ T1 ➡[h, g] T →
- ∀T2. ⦃G, L⦄ ⊢ T ➡*[h, g] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2.
+lemma cpxs_strap2: ∀h,o,G,L,T1,T. ⦃G, L⦄ ⊢ T1 ➡[h, o] T →
+ ∀T2. ⦃G, L⦄ ⊢ T ➡*[h, o] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2.
normalize /2 width=3 by TC_strap/ qed.
-lemma lsubr_cpxs_trans: ∀h,g,G. lsub_trans … (cpxs h g G) lsubr.
+lemma lsubr_cpxs_trans: ∀h,o,G. lsub_trans … (cpxs h o G) lsubr.
/3 width=5 by lsubr_cpx_trans, LTC_lsub_trans/
qed-.
-lemma cprs_cpxs: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2.
-#h #g #G #L #T1 #T2 #H @(cprs_ind … H) -T2 /3 width=3 by cpxs_strap1, cpr_cpx/
+lemma cprs_cpxs: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2.
+#h #o #G #L #T1 #T2 #H @(cprs_ind … H) -T2 /3 width=3 by cpxs_strap1, cpr_cpx/
qed.
-lemma cpxs_sort: ∀h,g,G,L,k,d1. deg h g k d1 →
- ∀d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ ⋆k ➡*[h, g] ⋆((next h)^d2 k).
-#h #g #G #L #k #d1 #Hkd1 #d2 @(nat_ind_plus … d2) -d2 /2 width=1 by cpx_cpxs/
+lemma cpxs_sort: ∀h,o,G,L,s,d1. deg h o s d1 →
+ ∀d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ ⋆s ➡*[h, o] ⋆((next h)^d2 s).
+#h #o #G #L #s #d1 #Hkd1 #d2 @(nat_ind_plus … d2) -d2 /2 width=1 by cpx_cpxs/
#d2 #IHd2 #Hd21 >iter_SO
-@(cpxs_strap1 … (⋆(iter d2 ℕ (next h) k)))
+@(cpxs_strap1 … (⋆(iter d2 ℕ (next h) s)))
[ /3 width=3 by lt_to_le/
-| @(cpx_st … (d1-d2-1)) <plus_minus_m_m
+| @(cpx_st … (d1-d2-1)) <plus_minus_k_k
/2 width=1 by deg_iter, monotonic_le_minus_r/
]
qed.
-lemma cpxs_bind_dx: ∀h,g,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 →
- ∀I,T1,T2. ⦃G, L. ⓑ{I}V1⦄ ⊢ T1 ➡*[h, g] T2 →
- ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, g] ⓑ{a,I}V2.T2.
-#h #g #G #L #V1 #V2 #HV12 #I #T1 #T2 #HT12 #a @(cpxs_ind_dx … HT12) -T1
+lemma cpxs_bind_dx: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 →
+ ∀I,T1,T2. ⦃G, L. ⓑ{I}V1⦄ ⊢ T1 ➡*[h, o] T2 →
+ ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, o] ⓑ{a,I}V2.T2.
+#h #o #G #L #V1 #V2 #HV12 #I #T1 #T2 #HT12 #a @(cpxs_ind_dx … HT12) -T1
/3 width=3 by cpxs_strap2, cpx_cpxs, cpx_pair_sn, cpx_bind/
qed.
-lemma cpxs_flat_dx: ∀h,g,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 →
- ∀T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 →
- ∀I. ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ➡*[h, g] ⓕ{I}V2.T2.
-#h #g #G #L #V1 #V2 #HV12 #T1 #T2 #HT12 @(cpxs_ind … HT12) -T2
+lemma cpxs_flat_dx: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 →
+ ∀T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 →
+ ∀I. ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ➡*[h, o] ⓕ{I}V2.T2.
+#h #o #G #L #V1 #V2 #HV12 #T1 #T2 #HT12 @(cpxs_ind … HT12) -T2
/3 width=5 by cpxs_strap1, cpx_cpxs, cpx_pair_sn, cpx_flat/
qed.
-lemma cpxs_flat_sn: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 →
- ∀V1,V2. ⦃G, L⦄ ⊢ V1 ➡*[h, g] V2 →
- ∀I. ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ➡*[h, g] ⓕ{I}V2.T2.
-#h #g #G #L #T1 #T2 #HT12 #V1 #V2 #H @(cpxs_ind … H) -V2
+lemma cpxs_flat_sn: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 →
+ ∀V1,V2. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 →
+ ∀I. ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ➡*[h, o] ⓕ{I}V2.T2.
+#h #o #G #L #T1 #T2 #HT12 #V1 #V2 #H @(cpxs_ind … H) -V2
/3 width=5 by cpxs_strap1, cpx_cpxs, cpx_pair_sn, cpx_flat/
qed.
-lemma cpxs_pair_sn: ∀h,g,I,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡*[h, g] V2 →
- ∀T. ⦃G, L⦄ ⊢ ②{I}V1.T ➡*[h, g] ②{I}V2.T.
-#h #g #I #G #L #V1 #V2 #H @(cpxs_ind … H) -V2
+lemma cpxs_pair_sn: ∀h,o,I,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 →
+ ∀T. ⦃G, L⦄ ⊢ ②{I}V1.T ➡*[h, o] ②{I}V2.T.
+#h #o #I #G #L #V1 #V2 #H @(cpxs_ind … H) -V2
/3 width=3 by cpxs_strap1, cpx_pair_sn/
qed.
-lemma cpxs_zeta: ∀h,g,G,L,V,T1,T,T2. ⬆[0, 1] T2 ≡ T →
- ⦃G, L.ⓓV⦄ ⊢ T1 ➡*[h, g] T → ⦃G, L⦄ ⊢ +ⓓV.T1 ➡*[h, g] T2.
-#h #g #G #L #V #T1 #T #T2 #HT2 #H @(cpxs_ind_dx … H) -T1
+lemma cpxs_zeta: ∀h,o,G,L,V,T1,T,T2. ⬆[0, 1] T2 ≡ T →
+ ⦃G, L.ⓓV⦄ ⊢ T1 ➡*[h, o] T → ⦃G, L⦄ ⊢ +ⓓV.T1 ➡*[h, o] T2.
+#h #o #G #L #V #T1 #T #T2 #HT2 #H @(cpxs_ind_dx … H) -T1
/3 width=3 by cpxs_strap2, cpx_cpxs, cpx_bind, cpx_zeta/
qed.
-lemma cpxs_eps: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 →
- ∀V. ⦃G, L⦄ ⊢ ⓝV.T1 ➡*[h, g] T2.
-#h #g #G #L #T1 #T2 #H @(cpxs_ind … H) -T2
+lemma cpxs_eps: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 →
+ ∀V. ⦃G, L⦄ ⊢ ⓝV.T1 ➡*[h, o] T2.
+#h #o #G #L #T1 #T2 #H @(cpxs_ind … H) -T2
/3 width=3 by cpxs_strap1, cpx_cpxs, cpx_eps/
qed.
-lemma cpxs_ct: ∀h,g,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡*[h, g] V2 →
- ∀T. ⦃G, L⦄ ⊢ ⓝV1.T ➡*[h, g] V2.
-#h #g #G #L #V1 #V2 #H @(cpxs_ind … H) -V2
+lemma cpxs_ct: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 →
+ ∀T. ⦃G, L⦄ ⊢ ⓝV1.T ➡*[h, o] V2.
+#h #o #G #L #V1 #V2 #H @(cpxs_ind … H) -V2
/3 width=3 by cpxs_strap1, cpx_cpxs, cpx_ct/
qed.
-lemma cpxs_beta_dx: ∀h,g,a,G,L,V1,V2,W1,W2,T1,T2.
- ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 → ⦃G, L.ⓛW1⦄ ⊢ T1 ➡*[h, g] T2 → ⦃G, L⦄ ⊢ W1 ➡[h, g] W2 →
- ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡*[h, g] ⓓ{a}ⓝW2.V2.T2.
-#h #g #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 * -T2
+lemma cpxs_beta_dx: ∀h,o,a,G,L,V1,V2,W1,W2,T1,T2.
+ ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 → ⦃G, L.ⓛW1⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 →
+ ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡*[h, o] ⓓ{a}ⓝW2.V2.T2.
+#h #o #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 * -T2
/4 width=7 by cpx_cpxs, cpxs_strap1, cpxs_bind_dx, cpxs_flat_dx, cpx_beta/
qed.
-lemma cpxs_theta_dx: ∀h,g,a,G,L,V1,V,V2,W1,W2,T1,T2.
- ⦃G, L⦄ ⊢ V1 ➡[h, g] V → ⬆[0, 1] V ≡ V2 → ⦃G, L.ⓓW1⦄ ⊢ T1 ➡*[h, g] T2 →
- ⦃G, L⦄ ⊢ W1 ➡[h, g] W2 → ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡*[h, g] ⓓ{a}W2.ⓐV2.T2.
-#h #g #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 * -T2
+lemma cpxs_theta_dx: ∀h,o,a,G,L,V1,V,V2,W1,W2,T1,T2.
+ ⦃G, L⦄ ⊢ V1 ➡[h, o] V → ⬆[0, 1] V ≡ V2 → ⦃G, L.ⓓW1⦄ ⊢ T1 ➡*[h, o] T2 →
+ ⦃G, L⦄ ⊢ W1 ➡[h, o] W2 → ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡*[h, o] ⓓ{a}W2.ⓐV2.T2.
+#h #o #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 * -T2
/4 width=9 by cpx_cpxs, cpxs_strap1, cpxs_bind_dx, cpxs_flat_dx, cpx_theta/
qed.
(* Basic inversion lemmas ***************************************************)
-lemma cpxs_inv_sort1: ∀h,g,G,L,U2,k. ⦃G, L⦄ ⊢ ⋆k ➡*[h, g] U2 →
- ∃∃n,d. deg h g k (n+d) & U2 = ⋆((next h)^n k).
-#h #g #G #L #U2 #k #H @(cpxs_ind … H) -U2
-[ elim (deg_total h g k) #d #Hkd
+lemma cpxs_inv_sort1: ∀h,o,G,L,U2,s. ⦃G, L⦄ ⊢ ⋆s ➡*[h, o] U2 →
+ ∃∃n,d. deg h o s (n+d) & U2 = ⋆((next h)^n s).
+#h #o #G #L #U2 #s #H @(cpxs_ind … H) -U2
+[ elim (deg_total h o s) #d #Hkd
@(ex2_2_intro … 0 … Hkd) -Hkd //
| #U #U2 #_ #HU2 * #n #d #Hknd #H destruct
elim (cpx_inv_sort1 … HU2) -HU2
]
qed-.
-lemma cpxs_inv_cast1: ∀h,g,G,L,W1,T1,U2. ⦃G, L⦄ ⊢ ⓝW1.T1 ➡*[h, g] U2 →
- ∨∨ ∃∃W2,T2. ⦃G, L⦄ ⊢ W1 ➡*[h, g] W2 & ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 & U2 = ⓝW2.T2
- | ⦃G, L⦄ ⊢ T1 ➡*[h, g] U2
- | ⦃G, L⦄ ⊢ W1 ➡*[h, g] U2.
-#h #g #G #L #W1 #T1 #U2 #H @(cpxs_ind … H) -U2 /3 width=5 by or3_intro0, ex3_2_intro/
+lemma cpxs_inv_cast1: ∀h,o,G,L,W1,T1,U2. ⦃G, L⦄ ⊢ ⓝW1.T1 ➡*[h, o] U2 →
+ ∨∨ ∃∃W2,T2. ⦃G, L⦄ ⊢ W1 ➡*[h, o] W2 & ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 & U2 = ⓝW2.T2
+ | ⦃G, L⦄ ⊢ T1 ➡*[h, o] U2
+ | ⦃G, L⦄ ⊢ W1 ➡*[h, o] U2.
+#h #o #G #L #W1 #T1 #U2 #H @(cpxs_ind … H) -U2 /3 width=5 by or3_intro0, ex3_2_intro/
#U2 #U #_ #HU2 * /3 width=3 by cpxs_strap1, or3_intro1, or3_intro2/ *
#W #T #HW1 #HT1 #H destruct
elim (cpx_inv_cast1 … HU2) -HU2 /3 width=3 by cpxs_strap1, or3_intro1, or3_intro2/ *
lapply (cpxs_strap1 … HT1 … HT2) -T /3 width=5 by or3_intro0, ex3_2_intro/
qed-.
-lemma cpxs_inv_cnx1: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T ➡*[h, g] U → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T⦄ → T = U.
-#h #g #G #L #T #U #H @(cpxs_ind_dx … H) -T //
+lemma cpxs_inv_cnx1: ∀h,o,G,L,T,U. ⦃G, L⦄ ⊢ T ➡*[h, o] U → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → T = U.
+#h #o #G #L #T #U #H @(cpxs_ind_dx … H) -T //
#T0 #T #H1T0 #_ #IHT #H2T0
lapply (H2T0 … H1T0) -H1T0 #H destruct /2 width=1 by/
qed-.
-lemma cpxs_neq_inv_step_sn: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → (T1 = T2 → ⊥) →
- ∃∃T. ⦃G, L⦄ ⊢ T1 ➡[h, g] T & T1 = T → ⊥ & ⦃G, L⦄ ⊢ T ➡*[h, g] T2.
-#h #g #G #L #T1 #T2 #H @(cpxs_ind_dx … H) -T1
+lemma cpxs_neq_inv_step_sn: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 = T2 → ⊥) →
+ ∃∃T. ⦃G, L⦄ ⊢ T1 ➡[h, o] T & T1 = T → ⊥ & ⦃G, L⦄ ⊢ T ➡*[h, o] T2.
+#h #o #G #L #T1 #T2 #H @(cpxs_ind_dx … H) -T1
[ #H elim H -H //
| #T1 #T #H1 #H2 #IH2 #H12 elim (eq_term_dec T1 T) #H destruct
[ -H1 -H2 /3 width=1 by/
(* Properties about atomic arity assignment on terms ************************)
-lemma cpxs_aaa_conf: ∀h,g,G,L,T1,A. ⦃G, L⦄ ⊢ T1 ⁝ A →
- ∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → ⦃G, L⦄ ⊢ T2 ⁝ A.
-#h #g #G #L #T1 #A #HT1 #T2 #HT12
+lemma cpxs_aaa_conf: ∀h,o,G,L,T1,A. ⦃G, L⦄ ⊢ T1 ⁝ A →
+ ∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ T2 ⁝ A.
+#h #o #G #L #T1 #A #HT1 #T2 #HT12
@(TC_Conf3 … HT1 ? HT12) -A -T1 -T2 /2 width=5 by cpx_aaa_conf/
qed-.
(* Main properties **********************************************************)
-theorem cpxs_trans: ∀h,g,G,L. Transitive … (cpxs h g G L).
+theorem cpxs_trans: ∀h,o,G,L. Transitive … (cpxs h o G L).
normalize /2 width=3 by trans_TC/ qed-.
-theorem cpxs_bind: ∀h,g,a,I,G,L,V1,V2,T1,T2. ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ➡*[h, g] T2 →
- ⦃G, L⦄ ⊢ V1 ➡*[h, g] V2 →
- ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, g] ⓑ{a,I}V2.T2.
-#h #g #a #I #G #L #V1 #V2 #T1 #T2 #HT12 #H @(cpxs_ind … H) -V2
+theorem cpxs_bind: ∀h,o,a,I,G,L,V1,V2,T1,T2. ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ➡*[h, o] T2 →
+ ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 →
+ ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, o] ⓑ{a,I}V2.T2.
+#h #o #a #I #G #L #V1 #V2 #T1 #T2 #HT12 #H @(cpxs_ind … H) -V2
/3 width=5 by cpxs_trans, cpxs_bind_dx/
qed.
-theorem cpxs_flat: ∀h,g,I,G,L,V1,V2,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 →
- ⦃G, L⦄ ⊢ V1 ➡*[h, g] V2 →
- ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ➡*[h, g] ⓕ{I}V2.T2.
-#h #g #I #G #L #V1 #V2 #T1 #T2 #HT12 #H @(cpxs_ind … H) -V2
+theorem cpxs_flat: ∀h,o,I,G,L,V1,V2,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 →
+ ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 →
+ ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ➡*[h, o] ⓕ{I}V2.T2.
+#h #o #I #G #L #V1 #V2 #T1 #T2 #HT12 #H @(cpxs_ind … H) -V2
/3 width=5 by cpxs_trans, cpxs_flat_dx/
qed.
-theorem cpxs_beta_rc: ∀h,g,a,G,L,V1,V2,W1,W2,T1,T2.
- ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 → ⦃G, L.ⓛW1⦄ ⊢ T1 ➡*[h, g] T2 → ⦃G, L⦄ ⊢ W1 ➡*[h, g] W2 →
- ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡*[h, g] ⓓ{a}ⓝW2.V2.T2.
-#h #g #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HT12 #H @(cpxs_ind … H) -W2
+theorem cpxs_beta_rc: ∀h,o,a,G,L,V1,V2,W1,W2,T1,T2.
+ ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 → ⦃G, L.ⓛW1⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ W1 ➡*[h, o] W2 →
+ ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡*[h, o] ⓓ{a}ⓝW2.V2.T2.
+#h #o #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HT12 #H @(cpxs_ind … H) -W2
/4 width=5 by cpxs_trans, cpxs_beta_dx, cpxs_bind_dx, cpx_pair_sn/
qed.
-theorem cpxs_beta: ∀h,g,a,G,L,V1,V2,W1,W2,T1,T2.
- ⦃G, L.ⓛW1⦄ ⊢ T1 ➡*[h, g] T2 → ⦃G, L⦄ ⊢ W1 ➡*[h, g] W2 → ⦃G, L⦄ ⊢ V1 ➡*[h, g] V2 →
- ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡*[h, g] ⓓ{a}ⓝW2.V2.T2.
-#h #g #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HT12 #HW12 #H @(cpxs_ind … H) -V2
+theorem cpxs_beta: ∀h,o,a,G,L,V1,V2,W1,W2,T1,T2.
+ ⦃G, L.ⓛW1⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ W1 ➡*[h, o] W2 → ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 →
+ ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡*[h, o] ⓓ{a}ⓝW2.V2.T2.
+#h #o #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HT12 #HW12 #H @(cpxs_ind … H) -V2
/4 width=5 by cpxs_trans, cpxs_beta_rc, cpxs_bind_dx, cpx_flat/
qed.
-theorem cpxs_theta_rc: ∀h,g,a,G,L,V1,V,V2,W1,W2,T1,T2.
- ⦃G, L⦄ ⊢ V1 ➡[h, g] V → ⬆[0, 1] V ≡ V2 →
- ⦃G, L.ⓓW1⦄ ⊢ T1 ➡*[h, g] T2 → ⦃G, L⦄ ⊢ W1 ➡*[h, g] W2 →
- ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡*[h, g] ⓓ{a}W2.ⓐV2.T2.
-#h #g #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 #HT12 #H @(cpxs_ind … H) -W2
+theorem cpxs_theta_rc: ∀h,o,a,G,L,V1,V,V2,W1,W2,T1,T2.
+ ⦃G, L⦄ ⊢ V1 ➡[h, o] V → ⬆[0, 1] V ≡ V2 →
+ ⦃G, L.ⓓW1⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ W1 ➡*[h, o] W2 →
+ ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡*[h, o] ⓓ{a}W2.ⓐV2.T2.
+#h #o #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 #HT12 #H @(cpxs_ind … H) -W2
/3 width=5 by cpxs_trans, cpxs_theta_dx, cpxs_bind_dx/
qed.
-theorem cpxs_theta: ∀h,g,a,G,L,V1,V,V2,W1,W2,T1,T2.
- ⬆[0, 1] V ≡ V2 → ⦃G, L⦄ ⊢ W1 ➡*[h, g] W2 →
- ⦃G, L.ⓓW1⦄ ⊢ T1 ➡*[h, g] T2 → ⦃G, L⦄ ⊢ V1 ➡*[h, g] V →
- ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡*[h, g] ⓓ{a}W2.ⓐV2.T2.
-#h #g #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV2 #HW12 #HT12 #H @(TC_ind_dx … V1 H) -V1
+theorem cpxs_theta: ∀h,o,a,G,L,V1,V,V2,W1,W2,T1,T2.
+ ⬆[0, 1] V ≡ V2 → ⦃G, L⦄ ⊢ W1 ➡*[h, o] W2 →
+ ⦃G, L.ⓓW1⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ V1 ➡*[h, o] V →
+ ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡*[h, o] ⓓ{a}W2.ⓐV2.T2.
+#h #o #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV2 #HW12 #HT12 #H @(TC_ind_dx … V1 H) -V1
/3 width=5 by cpxs_trans, cpxs_theta_rc, cpxs_flat_dx/
qed.
(* Advanced inversion lemmas ************************************************)
-lemma cpxs_inv_appl1: ∀h,g,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓐV1.T1 ➡*[h, g] U2 →
- ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡*[h, g] V2 & ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 &
+lemma cpxs_inv_appl1: ∀h,o,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓐV1.T1 ➡*[h, o] U2 →
+ ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 & ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 &
U2 = ⓐV2. T2
- | ∃∃a,W,T. ⦃G, L⦄ ⊢ T1 ➡*[h, g] ⓛ{a}W.T & ⦃G, L⦄ ⊢ ⓓ{a}ⓝW.V1.T ➡*[h, g] U2
- | ∃∃a,V0,V2,V,T. ⦃G, L⦄ ⊢ V1 ➡*[h, g] V0 & ⬆[0,1] V0 ≡ V2 &
- ⦃G, L⦄ ⊢ T1 ➡*[h, g] ⓓ{a}V.T & ⦃G, L⦄ ⊢ ⓓ{a}V.ⓐV2.T ➡*[h, g] U2.
-#h #g #G #L #V1 #T1 #U2 #H @(cpxs_ind … H) -U2 [ /3 width=5 by or3_intro0, ex3_2_intro/ ]
+ | ∃∃a,W,T. ⦃G, L⦄ ⊢ T1 ➡*[h, o] ⓛ{a}W.T & ⦃G, L⦄ ⊢ ⓓ{a}ⓝW.V1.T ➡*[h, o] U2
+ | ∃∃a,V0,V2,V,T. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V0 & ⬆[0,1] V0 ≡ V2 &
+ ⦃G, L⦄ ⊢ T1 ➡*[h, o] ⓓ{a}V.T & ⦃G, L⦄ ⊢ ⓓ{a}V.ⓐV2.T ➡*[h, o] U2.
+#h #o #G #L #V1 #T1 #U2 #H @(cpxs_ind … H) -U2 [ /3 width=5 by or3_intro0, ex3_2_intro/ ]
#U #U2 #_ #HU2 * *
[ #V0 #T0 #HV10 #HT10 #H destruct
elim (cpx_inv_appl1 … HU2) -HU2 *
(* Properties on sn extended parallel reduction for local environments ******)
-lemma lpx_cpx_trans: ∀h,g,G. s_r_transitive … (cpx h g G) (λ_.lpx h g G).
-#h #g #G #L2 #T1 #T2 #HT12 elim HT12 -G -L2 -T1 -T2
+lemma lpx_cpx_trans: ∀h,o,G. c_r_transitive … (cpx h o G) (λ_.lpx h o G).
+#h #o #G #L2 #T1 #T2 #HT12 elim HT12 -G -L2 -T1 -T2
[ /2 width=3 by/
| /3 width=2 by cpx_cpxs, cpx_st/
| #I #G #L2 #K2 #V0 #V2 #W2 #i #HLK2 #_ #HVW2 #IHV02 #L1 #HL12
]
qed-.
-lemma cpx_bind2: ∀h,g,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 →
- ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ➡[h, g] T2 →
- ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, g] ⓑ{a,I}V2.T2.
+lemma cpx_bind2: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 →
+ ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ➡[h, o] T2 →
+ ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, o] ⓑ{a,I}V2.T2.
/4 width=5 by lpx_cpx_trans, cpxs_bind_dx, lpx_pair/ qed.
(* Advanced properties ******************************************************)
-lemma lpx_cpxs_trans: ∀h,g,G. s_rs_transitive … (cpx h g G) (λ_.lpx h g G).
-#h #g #G @s_r_trans_LTC1 /2 width=3 by lpx_cpx_trans/ (**) (* full auto fails *)
+lemma lpx_cpxs_trans: ∀h,o,G. c_rs_transitive … (cpx h o G) (λ_.lpx h o G).
+#h #o #G @c_r_trans_LTC1 /2 width=3 by lpx_cpx_trans/ (**) (* full auto fails *)
qed-.
-lemma cpxs_bind2_dx: ∀h,g,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 →
- ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ➡*[h, g] T2 →
- ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, g] ⓑ{a,I}V2.T2.
+lemma cpxs_bind2_dx: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 →
+ ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ➡*[h, o] T2 →
+ ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, o] ⓑ{a,I}V2.T2.
/4 width=5 by lpx_cpxs_trans, cpxs_bind_dx, lpx_pair/ qed.
(* Properties on supclosure *************************************************)
-lemma fqu_cpxs_trans_neq: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, g] U2 → (T2 = U2 → ⊥) →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, U2⦄.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
+lemma fqu_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 → (T2 = U2 → ⊥) →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
[ #I #G #L #V1 #V2 #HV12 #_ elim (lift_total V2 0 1)
#U2 #HVU2 @(ex3_intro … U2)
[1,3: /3 width=7 by fqu_drop, cpxs_delta, drop_pair, drop_drop/
[1,3: /2 width=4 by fqu_flat_dx, cpxs_flat/
| #H0 destruct /2 width=1 by/
]
-| #G #L #K #T1 #U1 #m #HLK #HTU1 #T2 #HT12 #H elim (lift_total T2 0 (m+1))
+| #G #L #K #T1 #U1 #k #HLK #HTU1 #T2 #HT12 #H elim (lift_total T2 0 (k+1))
#U2 #HTU2 @(ex3_intro … U2)
[1,3: /2 width=10 by cpxs_lift, fqu_drop/
| #H0 destruct /3 width=5 by lift_inj/
]
qed-.
-lemma fquq_cpxs_trans_neq: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, g] U2 → (T2 = U2 → ⊥) →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H12 #U2 #HTU2 #H elim (fquq_inv_gen … H12) -H12
+lemma fquq_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 → (T2 = U2 → ⊥) →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #U2 #HTU2 #H elim (fquq_inv_gen … H12) -H12
[ #H12 elim (fqu_cpxs_trans_neq … H12 … HTU2 H) -T2
/3 width=4 by fqu_fquq, ex3_intro/
| * #HG #HL #HT destruct /3 width=4 by ex3_intro/
]
qed-.
-lemma fqup_cpxs_trans_neq: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, g] U2 → (T2 = U2 → ⊥) →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐+ ⦃G2, L2, U2⦄.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind_dx … H) -G1 -L1 -T1
+lemma fqup_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 → (T2 = U2 → ⊥) →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐+ ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind_dx … H) -G1 -L1 -T1
[ #G1 #L1 #T1 #H12 #U2 #HTU2 #H elim (fqu_cpxs_trans_neq … H12 … HTU2 H) -T2
/3 width=4 by fqu_fqup, ex3_intro/
| #G #G1 #L #L1 #T #T1 #H1 #_ #IH12 #U2 #HTU2 #H elim (IH12 … HTU2 H) -T2
]
qed-.
-lemma fqus_cpxs_trans_neq: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, g] U2 → (T2 = U2 → ⊥) →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐* ⦃G2, L2, U2⦄.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H12 #U2 #HTU2 #H elim (fqus_inv_gen … H12) -H12
+lemma fqus_cpxs_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 → (T2 = U2 → ⊥) →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐* ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #U2 #HTU2 #H elim (fqus_inv_gen … H12) -H12
[ #H12 elim (fqup_cpxs_trans_neq … H12 … HTU2 H) -T2
/3 width=4 by fqup_fqus, ex3_intro/
| * #HG #HL #HT destruct /3 width=4 by ex3_intro/
(* Advanced properties ******************************************************)
-lemma cpxs_delta: ∀h,g,I,G,L,K,V,V2,i.
- ⬇[i] L ≡ K.ⓑ{I}V → ⦃G, K⦄ ⊢ V ➡*[h, g] V2 →
- ∀W2. ⬆[0, i+1] V2 ≡ W2 → ⦃G, L⦄ ⊢ #i ➡*[h, g] W2.
-#h #g #I #G #L #K #V #V2 #i #HLK #H elim H -V2
+lemma cpxs_delta: ∀h,o,I,G,L,K,V,V2,i.
+ ⬇[i] L ≡ K.ⓑ{I}V → ⦃G, K⦄ ⊢ V ➡*[h, o] V2 →
+ ∀W2. ⬆[0, i+1] V2 ≡ W2 → ⦃G, L⦄ ⊢ #i ➡*[h, o] W2.
+#h #o #I #G #L #K #V #V2 #i #HLK #H elim H -V2
[ /3 width=9 by cpx_cpxs, cpx_delta/
| #V1 lapply (drop_fwd_drop2 … HLK) -HLK
elim (lift_total V1 0 (i+1)) /4 width=12 by cpx_lift, cpxs_strap1/
]
qed.
-lemma lstas_cpxs: ∀h,g,G,L,T1,T2,d2. ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 →
- ∀d1. ⦃G, L⦄ ⊢ T1 ▪[h, g] d1 → d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2.
-#h #g #G #L #T1 #T2 #d2 #H elim H -G -L -T1 -T2 -d2 //
+lemma lstas_cpxs: ∀h,o,G,L,T1,T2,d2. ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 →
+ ∀d1. ⦃G, L⦄ ⊢ T1 ▪[h, o] d1 → d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2.
+#h #o #G #L #T1 #T2 #d2 #H elim H -G -L -T1 -T2 -d2 //
[ /3 width=3 by cpxs_sort, da_inv_sort/
| #G #L #K #V1 #V2 #W2 #i #d2 #HLK #_ #HVW2 #IHV12 #d1 #H #Hd21
elim (da_inv_lref … H) -H * #K0 #V0 [| #d0 ] #HLK0
(* Advanced inversion lemmas ************************************************)
-lemma cpxs_inv_lref1: ∀h,g,G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡*[h, g] T2 →
+lemma cpxs_inv_lref1: ∀h,o,G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡*[h, o] T2 →
T2 = #i ∨
- ∃∃I,K,V1,T1. ⬇[i] L ≡ K.ⓑ{I}V1 & ⦃G, K⦄ ⊢ V1 ➡*[h, g] T1 &
+ ∃∃I,K,V1,T1. ⬇[i] L ≡ K.ⓑ{I}V1 & ⦃G, K⦄ ⊢ V1 ➡*[h, o] T1 &
⬆[0, i+1] T1 ≡ T2.
-#h #g #G #L #T2 #i #H @(cpxs_ind … H) -T2 /2 width=1 by or_introl/
+#h #o #G #L #T2 #i #H @(cpxs_ind … H) -T2 /2 width=1 by or_introl/
#T #T2 #_ #HT2 *
[ #H destruct
elim (cpx_inv_lref1 … HT2) -HT2 /2 width=1 by or_introl/
(* Relocation properties ****************************************************)
-lemma cpxs_lift: ∀h,g,G. d_liftable (cpxs h g G).
+lemma cpxs_lift: ∀h,o,G. d_liftable (cpxs h o G).
/3 width=10 by cpx_lift, cpxs_strap1, d_liftable_LTC/ qed.
-lemma cpxs_inv_lift1: ∀h,g,G. d_deliftable_sn (cpxs h g G).
+lemma cpxs_inv_lift1: ∀h,o,G. d_deliftable_sn (cpxs h o G).
/3 width=6 by d_deliftable_sn_LTC, cpx_inv_lift1/
qed-.
(* Properties on supclosure *************************************************)
-lemma fqu_cpxs_trans: ∀h,g,G1,G2,L1,L2,T2,U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, g] U2 →
+lemma fqu_cpxs_trans: ∀h,o,G1,G2,L1,L2,T2,U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 →
∀T1. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] U1 & ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, U2⦄.
-#h #g #G1 #G2 #L1 #L2 #T2 #U2 #H @(cpxs_ind_dx … H) -T2 /2 width=3 by ex2_intro/
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T2 #U2 #H @(cpxs_ind_dx … H) -T2 /2 width=3 by ex2_intro/
#T #T2 #HT2 #_ #IHTU2 #T1 #HT1 elim (fqu_cpx_trans … HT1 … HT2) -T
#T #HT1 #HT2 elim (IHTU2 … HT2) -T2 /3 width=3 by cpxs_strap2, ex2_intro/
qed-.
-lemma fquq_cpxs_trans: ∀h,g,G1,G2,L1,L2,T2,U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, g] U2 →
+lemma fquq_cpxs_trans: ∀h,o,G1,G2,L1,L2,T2,U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 →
∀T1. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] U1 & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
-#h #g #G1 #G2 #L1 #L2 #T2 #U2 #HTU2 #T1 #H elim (fquq_inv_gen … H) -H
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T2 #U2 #HTU2 #T1 #H elim (fquq_inv_gen … H) -H
[ #HT12 elim (fqu_cpxs_trans … HTU2 … HT12) /3 width=3 by fqu_fquq, ex2_intro/
| * #H1 #H2 #H3 destruct /2 width=3 by ex2_intro/
]
qed-.
-lemma fquq_lstas_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+lemma fquq_lstas_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
∀U2,d1. ⦃G2, L2⦄ ⊢ T2 •*[h, d1] U2 →
- ∀d2. ⦃G2, L2⦄ ⊢ T2 ▪[h, g] d2 → d1 ≤ d2 →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] U1 & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
+ ∀d2. ⦃G2, L2⦄ ⊢ T2 ▪[h, o] d2 → d1 ≤ d2 →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄.
/3 width=5 by fquq_cpxs_trans, lstas_cpxs/ qed-.
-lemma fqup_cpxs_trans: ∀h,g,G1,G2,L1,L2,T2,U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, g] U2 →
+lemma fqup_cpxs_trans: ∀h,o,G1,G2,L1,L2,T2,U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 →
∀T1. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] U1 & ⦃G1, L1, U1⦄ ⊐+ ⦃G2, L2, U2⦄.
-#h #g #G1 #G2 #L1 #L2 #T2 #U2 #H @(cpxs_ind_dx … H) -T2 /2 width=3 by ex2_intro/
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & ⦃G1, L1, U1⦄ ⊐+ ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T2 #U2 #H @(cpxs_ind_dx … H) -T2 /2 width=3 by ex2_intro/
#T #T2 #HT2 #_ #IHTU2 #T1 #HT1 elim (fqup_cpx_trans … HT1 … HT2) -T
#U1 #HTU1 #H2 elim (IHTU2 … H2) -T2 /3 width=3 by cpxs_strap2, ex2_intro/
qed-.
-lemma fqus_cpxs_trans: ∀h,g,G1,G2,L1,L2,T2,U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, g] U2 →
+lemma fqus_cpxs_trans: ∀h,o,G1,G2,L1,L2,T2,U2. ⦃G2, L2⦄ ⊢ T2 ➡*[h, o] U2 →
∀T1. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] U1 & ⦃G1, L1, U1⦄ ⊐* ⦃G2, L2, U2⦄.
-#h #g #G1 #G2 #L1 #L2 #T2 #U2 #HTU2 #T1 #H elim (fqus_inv_gen … H) -H
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & ⦃G1, L1, U1⦄ ⊐* ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T2 #U2 #HTU2 #T1 #H elim (fqus_inv_gen … H) -H
[ #HT12 elim (fqup_cpxs_trans … HTU2 … HT12) /3 width=3 by fqup_fqus, ex2_intro/
| * #H1 #H2 #H3 destruct /2 width=3 by ex2_intro/
]
qed-.
-lemma fqus_lstas_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+lemma fqus_lstas_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
∀U2,d1. ⦃G2, L2⦄ ⊢ T2 •*[h, d1] U2 →
- ∀d2. ⦃G2, L2⦄ ⊢ T2 ▪[h, g] d2 → d1 ≤ d2 →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] U1 & ⦃G1, L1, U1⦄ ⊐* ⦃G2, L2, U2⦄.
+ ∀d2. ⦃G2, L2⦄ ⊢ T2 ▪[h, o] d2 → d1 ≤ d2 →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] U1 & ⦃G1, L1, U1⦄ ⊐* ⦃G2, L2, U2⦄.
/3 width=6 by fqus_cpxs_trans, lstas_cpxs/ qed-.
(* Properties on lazy equivalence for local environments ********************)
-lemma lleq_cpxs_trans: ∀h,g,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ➡*[h, g] T2 →
- ∀L1. L1 ≡[T1, 0] L2 → ⦃G, L1⦄ ⊢ T1 ➡*[h, g] T2.
-#h #g #G #L2 #T1 #T2 #H @(cpxs_ind_dx … H) -T1
+lemma lleq_cpxs_trans: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ➡*[h, o] T2 →
+ ∀L1. L1 ≡[T1, 0] L2 → ⦃G, L1⦄ ⊢ T1 ➡*[h, o] T2.
+#h #o #G #L2 #T1 #T2 #H @(cpxs_ind_dx … H) -T1
/4 width=6 by cpx_lleq_conf_dx, lleq_cpx_trans, cpxs_strap2/
qed-.
-lemma cpxs_lleq_conf: ∀h,g,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ➡*[h, g] T2 →
- ∀L1. L2 ≡[T1, 0] L1 → ⦃G, L1⦄ ⊢ T1 ➡*[h, g] T2.
+lemma cpxs_lleq_conf: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ➡*[h, o] T2 →
+ ∀L1. L2 ≡[T1, 0] L1 → ⦃G, L1⦄ ⊢ T1 ➡*[h, o] T2.
/3 width=3 by lleq_cpxs_trans, lleq_sym/ qed-.
-lemma cpxs_lleq_conf_dx: ∀h,g,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ➡*[h, g] T2 →
+lemma cpxs_lleq_conf_dx: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ➡*[h, o] T2 →
∀L1. L1 ≡[T1, 0] L2 → L1 ≡[T2, 0] L2.
-#h #g #G #L2 #T1 #T2 #H @(cpxs_ind … H) -T2 /3 width=6 by cpx_lleq_conf_dx/
+#h #o #G #L2 #T1 #T2 #H @(cpxs_ind … H) -T2 /3 width=6 by cpx_lleq_conf_dx/
qed-.
-lemma cpxs_lleq_conf_sn: ∀h,g,G,L1,T1,T2. ⦃G, L1⦄ ⊢ T1 ➡*[h, g] T2 →
+lemma cpxs_lleq_conf_sn: ∀h,o,G,L1,T1,T2. ⦃G, L1⦄ ⊢ T1 ➡*[h, o] T2 →
∀L2. L1 ≡[T1, 0] L2 → L1 ≡[T2, 0] L2.
/4 width=6 by cpxs_lleq_conf_dx, lleq_sym/ qed-.
(* Properties on equivalence for local environments *************************)
-lemma lreq_cpxs_trans: ∀h,g,G. lsub_trans … (cpxs h g G) (lreq 0 (∞)).
+lemma lreq_cpxs_trans: ∀h,o,G. lsub_trans … (cpxs h o G) (lreq 0 (∞)).
/3 width=5 by lreq_cpx_trans, LTC_lsub_trans/
qed-.
(* Forward lemmas involving same top term structure *************************)
-lemma cpxs_fwd_cnx: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T⦄ → ∀U. ⦃G, L⦄ ⊢ T ➡*[h, g] U → T ≂ U.
-#h #g #G #L #T #HT #U #H
+lemma cpxs_fwd_cnx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → ∀U. ⦃G, L⦄ ⊢ T ➡*[h, o] U → T ≂ U.
+#h #o #G #L #T #HT #U #H
>(cpxs_inv_cnx1 … H HT) -G -L -T //
qed-.
-lemma cpxs_fwd_sort: ∀h,g,G,L,U,k. ⦃G, L⦄ ⊢ ⋆k ➡*[h, g] U →
- ⋆k ≂ U ∨ ⦃G, L⦄ ⊢ ⋆(next h k) ➡*[h, g] U.
-#h #g #G #L #U #k #H
-elim (cpxs_inv_sort1 … H) -H #n #d generalize in match k; -k @(nat_ind_plus … n) -n
-[ #k #_ #H -d destruct /2 width=1 by or_introl/
-| #n #IHn #k >plus_plus_comm_23 #Hnd #H destruct
+lemma cpxs_fwd_sort: ∀h,o,G,L,U,s. ⦃G, L⦄ ⊢ ⋆s ➡*[h, o] U →
+ ⋆s ≂ U ∨ ⦃G, L⦄ ⊢ ⋆(next h s) ➡*[h, o] U.
+#h #o #G #L #U #s #H
+elim (cpxs_inv_sort1 … H) -H #n #d generalize in match s; -s @(nat_ind_plus … n) -n
+[ #s #_ #H -d destruct /2 width=1 by or_introl/
+| #n #IHn #s >plus_plus_comm_23 #Hnd #H destruct
lapply (deg_next_SO … Hnd) -Hnd #Hnd
elim (IHn … Hnd) -IHn
[ #H lapply (tsts_inv_atom1 … H) -H #H >H -H /2 width=1 by or_intror/
qed-.
(* Basic_1: was just: pr3_iso_beta *)
-lemma cpxs_fwd_beta: ∀h,g,a,G,L,V,W,T,U. ⦃G, L⦄ ⊢ ⓐV.ⓛ{a}W.T ➡*[h, g] U →
- ⓐV.ⓛ{a}W.T ≂ U ∨ ⦃G, L⦄ ⊢ ⓓ{a}ⓝW.V.T ➡*[h, g] U.
-#h #g #a #G #L #V #W #T #U #H
+lemma cpxs_fwd_beta: ∀h,o,a,G,L,V,W,T,U. ⦃G, L⦄ ⊢ ⓐV.ⓛ{a}W.T ➡*[h, o] U →
+ ⓐV.ⓛ{a}W.T ≂ U ∨ ⦃G, L⦄ ⊢ ⓓ{a}ⓝW.V.T ➡*[h, o] U.
+#h #o #a #G #L #V #W #T #U #H
elim (cpxs_inv_appl1 … H) -H *
[ #V0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or_introl/
| #b #W0 #T0 #HT0 #HU
qed-.
(* Note: probably this is an inversion lemma *)
-lemma cpxs_fwd_delta: ∀h,g,I,G,L,K,V1,i. ⬇[i] L ≡ K.ⓑ{I}V1 →
+lemma cpxs_fwd_delta: ∀h,o,I,G,L,K,V1,i. ⬇[i] L ≡ K.ⓑ{I}V1 →
∀V2. ⬆[0, i + 1] V1 ≡ V2 →
- ∀U. ⦃G, L⦄ ⊢ #i ➡*[h, g] U →
- #i ≂ U ∨ ⦃G, L⦄ ⊢ V2 ➡*[h, g] U.
-#h #g #I #G #L #K #V1 #i #HLK #V2 #HV12 #U #H
+ ∀U. ⦃G, L⦄ ⊢ #i ➡*[h, o] U →
+ #i ≂ U ∨ ⦃G, L⦄ ⊢ V2 ➡*[h, o] U.
+#h #o #I #G #L #K #V1 #i #HLK #V2 #HV12 #U #H
elim (cpxs_inv_lref1 … H) -H /2 width=1 by or_introl/
* #I0 #K0 #V0 #U0 #HLK0 #HVU0 #HU0
lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct
/4 width=10 by cpxs_lift, drop_fwd_drop2, or_intror/
qed-.
-lemma cpxs_fwd_theta: ∀h,g,a,G,L,V1,V,T,U. ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}V.T ➡*[h, g] U →
+lemma cpxs_fwd_theta: ∀h,o,a,G,L,V1,V,T,U. ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}V.T ➡*[h, o] U →
∀V2. ⬆[0, 1] V1 ≡ V2 → ⓐV1.ⓓ{a}V.T ≂ U ∨
- ⦃G, L⦄ ⊢ ⓓ{a}V.ⓐV2.T ➡*[h, g] U.
-#h #g #a #G #L #V1 #V #T #U #H #V2 #HV12
+ ⦃G, L⦄ ⊢ ⓓ{a}V.ⓐV2.T ➡*[h, o] U.
+#h #o #a #G #L #V1 #V #T #U #H #V2 #HV12
elim (cpxs_inv_appl1 … H) -H *
[ -HV12 #V0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or_introl/
| #b #W #T0 #HT0 #HU
]
qed-.
-lemma cpxs_fwd_cast: ∀h,g,G,L,W,T,U. ⦃G, L⦄ ⊢ ⓝW.T ➡*[h, g] U →
- ∨∨ ⓝW. T ≂ U | ⦃G, L⦄ ⊢ T ➡*[h, g] U | ⦃G, L⦄ ⊢ W ➡*[h, g] U.
-#h #g #G #L #W #T #U #H
+lemma cpxs_fwd_cast: ∀h,o,G,L,W,T,U. ⦃G, L⦄ ⊢ ⓝW.T ➡*[h, o] U →
+ ∨∨ ⓝW. T ≂ U | ⦃G, L⦄ ⊢ T ➡*[h, o] U | ⦃G, L⦄ ⊢ W ➡*[h, o] U.
+#h #o #G #L #W #T #U #H
elim (cpxs_inv_cast1 … H) -H /2 width=1 by or3_intro1, or3_intro2/ *
#W0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or3_intro0/
qed-.
(* Vector form of forward lemmas involving same top term structure **********)
(* Basic_1: was just: nf2_iso_appls_lref *)
-lemma cpxs_fwd_cnx_vector: ∀h,g,G,L,T. 𝐒⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T⦄ →
- ∀Vs,U. ⦃G, L⦄ ⊢ ⒶVs.T ➡*[h, g] U → ⒶVs.T ≂ U.
-#h #g #G #L #T #H1T #H2T #Vs elim Vs -Vs [ @(cpxs_fwd_cnx … H2T) ] (**) (* /2 width=3 by cpxs_fwd_cnx/ does not work *)
+lemma cpxs_fwd_cnx_vector: ∀h,o,G,L,T. 𝐒⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ →
+ ∀Vs,U. ⦃G, L⦄ ⊢ ⒶVs.T ➡*[h, o] U → ⒶVs.T ≂ U.
+#h #o #G #L #T #H1T #H2T #Vs elim Vs -Vs [ @(cpxs_fwd_cnx … H2T) ] (**) (* /2 width=3 by cpxs_fwd_cnx/ does not work *)
#V #Vs #IHVs #U #H
elim (cpxs_inv_appl1 … H) -H *
[ -IHVs #V0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair/
]
qed-.
-lemma cpxs_fwd_sort_vector: ∀h,g,G,L,k,Vs,U. ⦃G, L⦄ ⊢ ⒶVs.⋆k ➡*[h, g] U →
- ⒶVs.⋆k ≂ U ∨ ⦃G, L⦄ ⊢ ⒶVs.⋆(next h k) ➡*[h, g] U.
-#h #g #G #L #k #Vs elim Vs -Vs /2 width=1 by cpxs_fwd_sort/
+lemma cpxs_fwd_sort_vector: ∀h,o,G,L,s,Vs,U. ⦃G, L⦄ ⊢ ⒶVs.⋆s ➡*[h, o] U →
+ ⒶVs.⋆s ≂ U ∨ ⦃G, L⦄ ⊢ ⒶVs.⋆(next h s) ➡*[h, o] U.
+#h #o #G #L #s #Vs elim Vs -Vs /2 width=1 by cpxs_fwd_sort/
#V #Vs #IHVs #U #H
elim (cpxs_inv_appl1 … H) -H *
[ -IHVs #V1 #T1 #_ #_ #H destruct /2 width=1 by tsts_pair, or_introl/
(* Basic_1: was just: pr3_iso_appls_beta *)
-lemma cpxs_fwd_beta_vector: ∀h,g,a,G,L,Vs,V,W,T,U. ⦃G, L⦄ ⊢ ⒶVs.ⓐV.ⓛ{a}W.T ➡*[h, g] U →
- ⒶVs. ⓐV. ⓛ{a}W. T ≂ U ∨ ⦃G, L⦄ ⊢ ⒶVs.ⓓ{a}ⓝW.V.T ➡*[h, g] U.
-#h #g #a #G #L #Vs elim Vs -Vs /2 width=1 by cpxs_fwd_beta/
+lemma cpxs_fwd_beta_vector: ∀h,o,a,G,L,Vs,V,W,T,U. ⦃G, L⦄ ⊢ ⒶVs.ⓐV.ⓛ{a}W.T ➡*[h, o] U →
+ ⒶVs. ⓐV. ⓛ{a}W. T ≂ U ∨ ⦃G, L⦄ ⊢ ⒶVs.ⓓ{a}ⓝW.V.T ➡*[h, o] U.
+#h #o #a #G #L #Vs elim Vs -Vs /2 width=1 by cpxs_fwd_beta/
#V0 #Vs #IHVs #V #W #T #U #H
elim (cpxs_inv_appl1 … H) -H *
[ -IHVs #V1 #T1 #_ #_ #H destruct /2 width=1 by tsts_pair, or_introl/
]
qed-.
-lemma cpxs_fwd_delta_vector: ∀h,g,I,G,L,K,V1,i. ⬇[i] L ≡ K.ⓑ{I}V1 →
+lemma cpxs_fwd_delta_vector: ∀h,o,I,G,L,K,V1,i. ⬇[i] L ≡ K.ⓑ{I}V1 →
∀V2. ⬆[0, i + 1] V1 ≡ V2 →
- ∀Vs,U. ⦃G, L⦄ ⊢ ⒶVs.#i ➡*[h, g] U →
- ⒶVs.#i ≂ U ∨ ⦃G, L⦄ ⊢ ⒶVs.V2 ➡*[h, g] U.
-#h #g #I #G #L #K #V1 #i #HLK #V2 #HV12 #Vs elim Vs -Vs /2 width=5 by cpxs_fwd_delta/
+ ∀Vs,U. ⦃G, L⦄ ⊢ ⒶVs.#i ➡*[h, o] U →
+ ⒶVs.#i ≂ U ∨ ⦃G, L⦄ ⊢ ⒶVs.V2 ➡*[h, o] U.
+#h #o #I #G #L #K #V1 #i #HLK #V2 #HV12 #Vs elim Vs -Vs /2 width=5 by cpxs_fwd_delta/
#V #Vs #IHVs #U #H -K -V1
elim (cpxs_inv_appl1 … H) -H *
[ -IHVs #V0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or_introl/
qed-.
(* Basic_1: was just: pr3_iso_appls_abbr *)
-lemma cpxs_fwd_theta_vector: ∀h,g,G,L,V1s,V2s. ⬆[0, 1] V1s ≡ V2s →
- ∀a,V,T,U. ⦃G, L⦄ ⊢ ⒶV1s.ⓓ{a}V.T ➡*[h, g] U →
- ⒶV1s. ⓓ{a}V. T ≂ U ∨ ⦃G, L⦄ ⊢ ⓓ{a}V.ⒶV2s.T ➡*[h, g] U.
-#h #g #G #L #V1s #V2s * -V1s -V2s /3 width=1 by or_intror/
-#V1s #V2s #V1a #V2a #HV12a #HV12s #a
+lemma cpxs_fwd_theta_vector: ∀h,o,G,L,V1c,V2c. ⬆[0, 1] V1c ≡ V2c →
+ ∀a,V,T,U. ⦃G, L⦄ ⊢ ⒶV1c.ⓓ{a}V.T ➡*[h, o] U →
+ ⒶV1c. ⓓ{a}V. T ≂ U ∨ ⦃G, L⦄ ⊢ ⓓ{a}V.ⒶV2c.T ➡*[h, o] U.
+#h #o #G #L #V1c #V2c * -V1c -V2c /3 width=1 by or_intror/
+#V1c #V2c #V1a #V2a #HV12a #HV12c #a
generalize in match HV12a; -HV12a
generalize in match V2a; -V2a
generalize in match V1a; -V1a
-elim HV12s -V1s -V2s /2 width=1 by cpxs_fwd_theta/
-#V1s #V2s #V1b #V2b #HV12b #_ #IHV12s #V1a #V2a #HV12a #V #T #U #H
+elim HV12c -V1c -V2c /2 width=1 by cpxs_fwd_theta/
+#V1c #V2c #V1b #V2b #HV12b #_ #IHV12c #V1a #V2a #HV12a #V #T #U #H
elim (cpxs_inv_appl1 … H) -H *
-[ -IHV12s -HV12a -HV12b #V0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or_introl/
+[ -IHV12c -HV12a -HV12b #V0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or_introl/
| #b #W0 #T0 #HT0 #HU
- elim (IHV12s … HV12b … HT0) -IHV12s -HT0 #HT0
+ elim (IHV12c … HV12b … HT0) -IHV12c -HT0 #HT0
[ -HV12a -HV12b -HU
elim (tsts_inv_pair1 … HT0) #V1 #T1 #H destruct
- | @or_intror -V1s (**) (* explicit constructor *)
+ | @or_intror -V1c (**) (* explicit constructor *)
@(cpxs_trans … HU) -U
elim (cpxs_inv_abbr1 … HT0) -HT0 *
[ -HV12a -HV12b #V1 #T1 #_ #_ #H destruct
| -V1b #X #HT1 #H #H0 destruct
elim (lift_inv_bind1 … H) -H #W1 #T1 #HW01 #HT01 #H destruct
- @(cpxs_trans … (+ⓓV.ⓐV2a.ⓛ{b}W1.T1)) [ /3 width=1 by cpxs_flat_dx, cpxs_bind_dx/ ] -T -V2b -V2s
+ @(cpxs_trans … (+ⓓV.ⓐV2a.ⓛ{b}W1.T1)) [ /3 width=1 by cpxs_flat_dx, cpxs_bind_dx/ ] -T -V2b -V2c
@(cpxs_strap2 … (ⓐV1a.ⓛ{b}W0.T0))
/4 width=7 by cpxs_beta_dx, cpx_zeta, lift_bind, lift_flat/
]
]
| #b #V0a #Va #V0 #T0 #HV10a #HV0a #HT0 #HU
- elim (IHV12s … HV12b … HT0) -HV12b -IHV12s -HT0 #HT0
+ elim (IHV12c … HV12b … HT0) -HV12b -IHV12c -HT0 #HT0
[ -HV12a -HV10a -HV0a -HU
elim (tsts_inv_pair1 … HT0) #V1 #T1 #H destruct
- | @or_intror -V1s -V1b (**) (* explicit constructor *)
+ | @or_intror -V1c -V1b (**) (* explicit constructor *)
@(cpxs_trans … HU) -U
elim (cpxs_inv_abbr1 … HT0) -HT0 *
[ #V1 #T1 #HV1 #HT1 #H destruct
| #X #HT1 #H #H0 destruct
elim (lift_inv_bind1 … H) -H #V1 #T1 #HW01 #HT01 #H destruct
lapply (cpxs_lift … HV10a (L.ⓓV0) (Ⓕ) … HV12a … HV0a) -V0a [ /2 width=1 by drop_drop/ ] #HV2a
- @(cpxs_trans … (+ⓓV.ⓐV2a.ⓓ{b}V1.T1)) [ /3 width=1 by cpxs_flat_dx, cpxs_bind_dx/ ] -T -V2b -V2s
+ @(cpxs_trans … (+ⓓV.ⓐV2a.ⓓ{b}V1.T1)) [ /3 width=1 by cpxs_flat_dx, cpxs_bind_dx/ ] -T -V2b -V2c
@(cpxs_strap2 … (ⓐV1a.ⓓ{b}V0.T0)) [ /4 width=7 by cpx_zeta, lift_bind, lift_flat/ ] -V -V1 -T1
@(cpxs_strap2 … (ⓓ{b}V0.ⓐV2a.T0)) /3 width=3 by cpxs_pair_sn, cpxs_bind_dx, cpr_cpx, cpr_theta/
]
qed-.
(* Basic_1: was just: pr3_iso_appls_cast *)
-lemma cpxs_fwd_cast_vector: ∀h,g,G,L,Vs,W,T,U. ⦃G, L⦄ ⊢ ⒶVs.ⓝW.T ➡*[h, g] U →
+lemma cpxs_fwd_cast_vector: ∀h,o,G,L,Vs,W,T,U. ⦃G, L⦄ ⊢ ⒶVs.ⓝW.T ➡*[h, o] U →
∨∨ ⒶVs. ⓝW. T ≂ U
- | ⦃G, L⦄ ⊢ ⒶVs.T ➡*[h, g] U
- | ⦃G, L⦄ ⊢ ⒶVs.W ➡*[h, g] U.
-#h #g #G #L #Vs elim Vs -Vs /2 width=1 by cpxs_fwd_cast/
+ | ⦃G, L⦄ ⊢ ⒶVs.T ➡*[h, o] U
+ | ⦃G, L⦄ ⊢ ⒶVs.W ➡*[h, o] U.
+#h #o #G #L #Vs elim Vs -Vs /2 width=1 by cpxs_fwd_cast/
#V #Vs #IHVs #W #T #U #H
elim (cpxs_inv_appl1 … H) -H *
[ -IHVs #V0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or3_intro0/
(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERMS ********************)
definition csx: ∀h. sd h → relation3 genv lenv term ≝
- λh,g,G,L. SN … (cpx h g G L) (eq …).
+ λh,o,G,L. SN … (cpx h o G L) (eq …).
interpretation
"context-sensitive extended strong normalization (term)"
- 'SN h g G L T = (csx h g G L T).
+ 'SN h o G L T = (csx h o G L T).
(* Basic eliminators ********************************************************)
-lemma csx_ind: ∀h,g,G,L. ∀R:predicate term.
- (∀T1. ⦃G, L⦄ ⊢ ⬊*[h, g] T1 →
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → (T1 = T2 → ⊥) → R T2) →
+lemma csx_ind: ∀h,o,G,L. ∀R:predicate term.
+ (∀T1. ⦃G, L⦄ ⊢ ⬊*[h, o] T1 →
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → (T1 = T2 → ⊥) → R T2) →
R T1
) →
- ∀T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → R T.
-#h #g #G #L #R #H0 #T1 #H elim H -T1
+ ∀T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → R T.
+#h #o #G #L #R #H0 #T1 #H elim H -T1
/5 width=1 by SN_intro/
qed-.
(* Basic properties *********************************************************)
(* Basic_1: was just: sn3_pr2_intro *)
-lemma csx_intro: ∀h,g,G,L,T1.
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊*[h, g] T2) →
- ⦃G, L⦄ ⊢ ⬊*[h, g] T1.
+lemma csx_intro: ∀h,o,G,L,T1.
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊*[h, o] T2) →
+ ⦃G, L⦄ ⊢ ⬊*[h, o] T1.
/4 width=1 by SN_intro/ qed.
-lemma csx_cpx_trans: ∀h,g,G,L,T1. ⦃G, L⦄ ⊢ ⬊*[h, g] T1 →
- ∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → ⦃G, L⦄ ⊢ ⬊*[h, g] T2.
-#h #g #G #L #T1 #H @(csx_ind … H) -T1 #T1 #HT1 #IHT1 #T2 #HLT12
+lemma csx_cpx_trans: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬊*[h, o] T1 →
+ ∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → ⦃G, L⦄ ⊢ ⬊*[h, o] T2.
+#h #o #G #L #T1 #H @(csx_ind … H) -T1 #T1 #HT1 #IHT1 #T2 #HLT12
elim (eq_term_dec T1 T2) #HT12 destruct /3 width=4 by/
qed-.
(* Basic_1: was just: sn3_nf2 *)
-lemma cnx_csx: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T⦄ → ⦃G, L⦄ ⊢ ⬊*[h, g] T.
+lemma cnx_csx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → ⦃G, L⦄ ⊢ ⬊*[h, o] T.
/2 width=1 by NF_to_SN/ qed.
-lemma csx_sort: ∀h,g,G,L,k. ⦃G, L⦄ ⊢ ⬊*[h, g] ⋆k.
-#h #g #G #L #k elim (deg_total h g k)
-#d generalize in match k; -k @(nat_ind_plus … d) -d /3 width=6 by cnx_csx, cnx_sort/
-#d #IHd #k #Hkd lapply (deg_next_SO … Hkd) -Hkd
+lemma csx_sort: ∀h,o,G,L,s. ⦃G, L⦄ ⊢ ⬊*[h, o] ⋆s.
+#h #o #G #L #s elim (deg_total h o s)
+#d generalize in match s; -s @(nat_ind_plus … d) -d /3 width=6 by cnx_csx, cnx_sort/
+#d #IHd #s #Hkd lapply (deg_next_SO … Hkd) -Hkd
#Hkd @csx_intro #X #H #HX elim (cpx_inv_sort1 … H) -H
[ #H destruct elim HX //
| -HX * #d0 #_ #H destruct -d0 /2 width=1 by/
qed.
(* Basic_1: was just: sn3_cast *)
-lemma csx_cast: ∀h,g,G,L,W. ⦃G, L⦄ ⊢ ⬊*[h, g] W →
- ∀T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓝW.T.
-#h #g #G #L #W #HW @(csx_ind … HW) -W #W #HW #IHW #T #HT @(csx_ind … HT) -T #T #HT #IHT
+lemma csx_cast: ∀h,o,G,L,W. ⦃G, L⦄ ⊢ ⬊*[h, o] W →
+ ∀T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓝW.T.
+#h #o #G #L #W #HW @(csx_ind … HW) -W #W #HW #IHW #T #HT @(csx_ind … HT) -T #T #HT #IHT
@csx_intro #X #H1 #H2
elim (cpx_inv_cast1 … H1) -H1
[ * #W0 #T0 #HLW0 #HLT0 #H destruct
(* Basic forward lemmas *****************************************************)
-fact csx_fwd_pair_sn_aux: ∀h,g,G,L,U. ⦃G, L⦄ ⊢ ⬊*[h, g] U →
- ∀I,V,T. U = ②{I}V.T → ⦃G, L⦄ ⊢ ⬊*[h, g] V.
-#h #g #G #L #U #H elim H -H #U0 #_ #IH #I #V #T #H destruct
+fact csx_fwd_pair_sn_aux: ∀h,o,G,L,U. ⦃G, L⦄ ⊢ ⬊*[h, o] U →
+ ∀I,V,T. U = ②{I}V.T → ⦃G, L⦄ ⊢ ⬊*[h, o] V.
+#h #o #G #L #U #H elim H -H #U0 #_ #IH #I #V #T #H destruct
@csx_intro #V2 #HLV2 #HV2
@(IH (②{I}V2.T)) -IH /2 width=3 by cpx_pair_sn/ -HLV2
#H destruct /2 width=1 by/
qed-.
(* Basic_1: was just: sn3_gen_head *)
-lemma csx_fwd_pair_sn: ∀h,g,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, g] ②{I}V.T → ⦃G, L⦄ ⊢ ⬊*[h, g] V.
+lemma csx_fwd_pair_sn: ∀h,o,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ②{I}V.T → ⦃G, L⦄ ⊢ ⬊*[h, o] V.
/2 width=5 by csx_fwd_pair_sn_aux/ qed-.
-fact csx_fwd_bind_dx_aux: ∀h,g,G,L,U. ⦃G, L⦄ ⊢ ⬊*[h, g] U →
- ∀a,I,V,T. U = ⓑ{a,I}V.T → ⦃G, L.ⓑ{I}V⦄ ⊢ ⬊*[h, g] T.
-#h #g #G #L #U #H elim H -H #U0 #_ #IH #a #I #V #T #H destruct
+fact csx_fwd_bind_dx_aux: ∀h,o,G,L,U. ⦃G, L⦄ ⊢ ⬊*[h, o] U →
+ ∀a,I,V,T. U = ⓑ{a,I}V.T → ⦃G, L.ⓑ{I}V⦄ ⊢ ⬊*[h, o] T.
+#h #o #G #L #U #H elim H -H #U0 #_ #IH #a #I #V #T #H destruct
@csx_intro #T2 #HLT2 #HT2
@(IH (ⓑ{a,I}V.T2)) -IH /2 width=3 by cpx_bind/ -HLT2
#H destruct /2 width=1 by/
qed-.
(* Basic_1: was just: sn3_gen_bind *)
-lemma csx_fwd_bind_dx: ∀h,g,a,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, g] ⓑ{a,I}V.T → ⦃G, L.ⓑ{I}V⦄ ⊢ ⬊*[h, g] T.
+lemma csx_fwd_bind_dx: ∀h,o,a,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⓑ{a,I}V.T → ⦃G, L.ⓑ{I}V⦄ ⊢ ⬊*[h, o] T.
/2 width=4 by csx_fwd_bind_dx_aux/ qed-.
-fact csx_fwd_flat_dx_aux: ∀h,g,G,L,U. ⦃G, L⦄ ⊢ ⬊*[h, g] U →
- ∀I,V,T. U = ⓕ{I}V.T → ⦃G, L⦄ ⊢ ⬊*[h, g] T.
-#h #g #G #L #U #H elim H -H #U0 #_ #IH #I #V #T #H destruct
+fact csx_fwd_flat_dx_aux: ∀h,o,G,L,U. ⦃G, L⦄ ⊢ ⬊*[h, o] U →
+ ∀I,V,T. U = ⓕ{I}V.T → ⦃G, L⦄ ⊢ ⬊*[h, o] T.
+#h #o #G #L #U #H elim H -H #U0 #_ #IH #I #V #T #H destruct
@csx_intro #T2 #HLT2 #HT2
@(IH (ⓕ{I}V.T2)) -IH /2 width=3 by cpx_flat/ -HLT2
#H destruct /2 width=1 by/
qed-.
(* Basic_1: was just: sn3_gen_flat *)
-lemma csx_fwd_flat_dx: ∀h,g,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, g] ⓕ{I}V.T → ⦃G, L⦄ ⊢ ⬊*[h, g] T.
+lemma csx_fwd_flat_dx: ∀h,o,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⓕ{I}V.T → ⦃G, L⦄ ⊢ ⬊*[h, o] T.
/2 width=5 by csx_fwd_flat_dx_aux/ qed-.
-lemma csx_fwd_bind: ∀h,g,a,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, g] ⓑ{a,I}V.T →
- ⦃G, L⦄ ⊢ ⬊*[h, g] V ∧ ⦃G, L.ⓑ{I}V⦄ ⊢ ⬊*[h, g] T.
+lemma csx_fwd_bind: ∀h,o,a,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⓑ{a,I}V.T →
+ ⦃G, L⦄ ⊢ ⬊*[h, o] V ∧ ⦃G, L.ⓑ{I}V⦄ ⊢ ⬊*[h, o] T.
/3 width=3 by csx_fwd_pair_sn, csx_fwd_bind_dx, conj/ qed-.
-lemma csx_fwd_flat: ∀h,g,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, g] ⓕ{I}V.T →
- ⦃G, L⦄ ⊢ ⬊*[h, g] V ∧ ⦃G, L⦄ ⊢ ⬊*[h, g] T.
+lemma csx_fwd_flat: ∀h,o,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⓕ{I}V.T →
+ ⦃G, L⦄ ⊢ ⬊*[h, o] V ∧ ⦃G, L⦄ ⊢ ⬊*[h, o] T.
/3 width=3 by csx_fwd_pair_sn, csx_fwd_flat_dx, conj/ qed-.
(* Basic_1: removed theorems 14:
(* Main properties on atomic arity assignment *******************************)
-theorem aaa_csx: ∀h,g,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ⦃G, L⦄ ⊢ ⬊*[h, g] T.
-#h #g #G #L #T #A #H
-@(gcr_aaa … (csx_gcp h g) (csx_gcr h g) … H)
+theorem aaa_csx: ∀h,o,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ⦃G, L⦄ ⊢ ⬊*[h, o] T.
+#h #o #G #L #T #A #H
+@(gcr_aaa … (csx_gcp h o) (csx_gcr h o) … H)
qed.
(* Advanced eliminators *****************************************************)
-fact aaa_ind_csx_aux: ∀h,g,G,L,A. ∀R:predicate term.
+fact aaa_ind_csx_aux: ∀h,o,G,L,A. ∀R:predicate term.
(∀T1. ⦃G, L⦄ ⊢ T1 ⁝ A →
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → (T1 = T2 → ⊥) → R T2) → R T1
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → (T1 = T2 → ⊥) → R T2) → R T1
) →
- ∀T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → ⦃G, L⦄ ⊢ T ⁝ A → R T.
-#h #g #G #L #A #R #IH #T #H @(csx_ind … H) -T /4 width=5 by cpx_aaa_conf/
+ ∀T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ⦃G, L⦄ ⊢ T ⁝ A → R T.
+#h #o #G #L #A #R #IH #T #H @(csx_ind … H) -T /4 width=5 by cpx_aaa_conf/
qed-.
-lemma aaa_ind_csx: ∀h,g,G,L,A. ∀R:predicate term.
+lemma aaa_ind_csx: ∀h,o,G,L,A. ∀R:predicate term.
(∀T1. ⦃G, L⦄ ⊢ T1 ⁝ A →
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → (T1 = T2 → ⊥) → R T2) → R T1
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → (T1 = T2 → ⊥) → R T2) → R T1
) →
∀T. ⦃G, L⦄ ⊢ T ⁝ A → R T.
/5 width=9 by aaa_ind_csx_aux, aaa_csx/ qed-.
-fact aaa_ind_csx_alt_aux: ∀h,g,G,L,A. ∀R:predicate term.
+fact aaa_ind_csx_alt_aux: ∀h,o,G,L,A. ∀R:predicate term.
(∀T1. ⦃G, L⦄ ⊢ T1 ⁝ A →
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → (T1 = T2 → ⊥) → R T2) → R T1
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 = T2 → ⊥) → R T2) → R T1
) →
- ∀T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → ⦃G, L⦄ ⊢ T ⁝ A → R T.
-#h #g #G #L #A #R #IH #T #H @(csx_ind_alt … H) -T /4 width=5 by cpxs_aaa_conf/
+ ∀T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ⦃G, L⦄ ⊢ T ⁝ A → R T.
+#h #o #G #L #A #R #IH #T #H @(csx_ind_alt … H) -T /4 width=5 by cpxs_aaa_conf/
qed-.
-lemma aaa_ind_csx_alt: ∀h,g,G,L,A. ∀R:predicate term.
+lemma aaa_ind_csx_alt: ∀h,o,G,L,A. ∀R:predicate term.
(∀T1. ⦃G, L⦄ ⊢ T1 ⁝ A →
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → (T1 = T2 → ⊥) → R T2) → R T1
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 = T2 → ⊥) → R T2) → R T1
) →
∀T. ⦃G, L⦄ ⊢ T ⁝ A → R T.
/5 width=9 by aaa_ind_csx_alt_aux, aaa_csx/ qed-.
(* alternative definition of csx *)
definition csxa: ∀h. sd h → relation3 genv lenv term ≝
- λh,g,G,L. SN … (cpxs h g G L) (eq …).
+ λh,o,G,L. SN … (cpxs h o G L) (eq …).
interpretation
"context-sensitive extended strong normalization (term) alternative"
- 'SNAlt h g G L T = (csxa h g G L T).
+ 'SNAlt h o G L T = (csxa h o G L T).
(* Basic eliminators ********************************************************)
-lemma csxa_ind: ∀h,g,G,L. ∀R:predicate term.
- (∀T1. ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T1 →
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → (T1 = T2 → ⊥) → R T2) → R T1
+lemma csxa_ind: ∀h,o,G,L. ∀R:predicate term.
+ (∀T1. ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T1 →
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 = T2 → ⊥) → R T2) → R T1
) →
- ∀T. ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T → R T.
-#h #g #G #L #R #H0 #T1 #H elim H -T1 /5 width=1 by SN_intro/
+ ∀T. ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T → R T.
+#h #o #G #L #R #H0 #T1 #H elim H -T1 /5 width=1 by SN_intro/
qed-.
(* Basic properties *********************************************************)
-lemma csx_intro_cpxs: ∀h,g,G,L,T1.
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊*[h, g] T2) →
- ⦃G, L⦄ ⊢ ⬊*[h, g] T1.
+lemma csx_intro_cpxs: ∀h,o,G,L,T1.
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊*[h, o] T2) →
+ ⦃G, L⦄ ⊢ ⬊*[h, o] T1.
/4 width=1 by cpx_cpxs, csx_intro/ qed.
(* Basic_1: was just: sn3_intro *)
-lemma csxa_intro: ∀h,g,G,L,T1.
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T2) →
- ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T1.
+lemma csxa_intro: ∀h,o,G,L,T1.
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T2) →
+ ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T1.
/4 width=1 by SN_intro/ qed.
-fact csxa_intro_aux: ∀h,g,G,L,T1. (
- ∀T,T2. ⦃G, L⦄ ⊢ T ➡*[h, g] T2 → T1 = T → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T2
- ) → ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T1.
+fact csxa_intro_aux: ∀h,o,G,L,T1. (
+ ∀T,T2. ⦃G, L⦄ ⊢ T ➡*[h, o] T2 → T1 = T → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T2
+ ) → ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T1.
/4 width=3 by csxa_intro/ qed-.
(* Basic_1: was just: sn3_pr3_trans (old version) *)
-lemma csxa_cpxs_trans: ∀h,g,G,L,T1. ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T1 →
- ∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T2.
-#h #g #G #L #T1 #H elim H -T1 #T1 #HT1 #IHT1 #T2 #HLT12
+lemma csxa_cpxs_trans: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T1 →
+ ∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T2.
+#h #o #G #L #T1 #H elim H -T1 #T1 #HT1 #IHT1 #T2 #HLT12
@csxa_intro #T #HLT2 #HT2
elim (eq_term_dec T1 T2) #HT12
[ -IHT1 -HLT12 destruct /3 width=1 by/
qed.
(* Basic_1: was just: sn3_pr2_intro (old version) *)
-lemma csxa_intro_cpx: ∀h,g,G,L,T1. (
- ∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T2
- ) → ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T1.
-#h #g #G #L #T1 #H
+lemma csxa_intro_cpx: ∀h,o,G,L,T1. (
+ ∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T2
+ ) → ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T1.
+#h #o #G #L #T1 #H
@csxa_intro_aux #T #T2 #H @(cpxs_ind_dx … H) -T
[ -H #H destruct #H
elim H //
(* Main properties **********************************************************)
-theorem csx_csxa: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T.
-#h #g #G #L #T #H @(csx_ind … H) -T /4 width=1 by csxa_intro_cpx/
+theorem csx_csxa: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T.
+#h #o #G #L #T #H @(csx_ind … H) -T /4 width=1 by csxa_intro_cpx/
qed.
-theorem csxa_csx: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T → ⦃G, L⦄ ⊢ ⬊*[h, g] T.
-#h #g #G #L #T #H @(csxa_ind … H) -T /4 width=1 by cpx_cpxs, csx_intro/
+theorem csxa_csx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ⬊⬊*[h, o] T → ⦃G, L⦄ ⊢ ⬊*[h, o] T.
+#h #o #G #L #T #H @(csxa_ind … H) -T /4 width=1 by cpx_cpxs, csx_intro/
qed.
(* Basic_1: was just: sn3_pr3_trans *)
-lemma csx_cpxs_trans: ∀h,g,G,L,T1. ⦃G, L⦄ ⊢ ⬊*[h, g] T1 →
- ∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → ⦃G, L⦄ ⊢ ⬊*[h, g] T2.
-#h #g #G #L #T1 #HT1 #T2 #H @(cpxs_ind … H) -T2 /2 width=3 by csx_cpx_trans/
+lemma csx_cpxs_trans: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬊*[h, o] T1 →
+ ∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L⦄ ⊢ ⬊*[h, o] T2.
+#h #o #G #L #T1 #HT1 #T2 #H @(cpxs_ind … H) -T2 /2 width=3 by csx_cpx_trans/
qed-.
(* Main eliminators *********************************************************)
-lemma csx_ind_alt: ∀h,g,G,L. ∀R:predicate term.
- (∀T1. ⦃G, L⦄ ⊢ ⬊*[h, g] T1 →
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → (T1 = T2 → ⊥) → R T2) → R T1
+lemma csx_ind_alt: ∀h,o,G,L. ∀R:predicate term.
+ (∀T1. ⦃G, L⦄ ⊢ ⬊*[h, o] T1 →
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 = T2 → ⊥) → R T2) → R T1
) →
- ∀T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → R T.
-#h #g #G #L #R #H0 #T1 #H @(csxa_ind … (csx_csxa … H)) -T1 /4 width=1 by csxa_csx/
+ ∀T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → R T.
+#h #o #G #L #R #H0 #T1 #H @(csxa_ind … (csx_csxa … H)) -T1 /4 width=1 by csxa_csx/
qed-.
(* Advanced properties ******************************************************)
-lemma csx_fpb_conf: ∀h,g,G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, g] T1 →
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≽[h, g] ⦃G2, L2, T2⦄ → ⦃G2, L2⦄ ⊢ ⬊*[h, g] T2.
-#h #g #G1 #L1 #T1 #HT1 #G2 #L2 #T2 *
+lemma csx_fpb_conf: ∀h,o,G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 →
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ → ⦃G2, L2⦄ ⊢ ⬊*[h, o] T2.
+#h #o #G1 #L1 #T1 #HT1 #G2 #L2 #T2 *
/2 width=5 by csx_cpx_trans, csx_fquq_conf, csx_lpx_conf, csx_lleq_conf/
qed-.
-lemma csx_fpbs_conf: ∀h,g,G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, g] T1 →
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ → ⦃G2, L2⦄ ⊢ ⬊*[h, g] T2.
-#h #g #G1 #L1 #T1 #HT1 #G2 #L2 #T2 #H @(fpbs_ind … H) -G2 -L2 -T2
+lemma csx_fpbs_conf: ∀h,o,G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 →
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → ⦃G2, L2⦄ ⊢ ⬊*[h, o] T2.
+#h #o #G1 #L1 #T1 #HT1 #G2 #L2 #T2 #H @(fpbs_ind … H) -G2 -L2 -T2
/2 width=5 by csx_fpb_conf/
qed-.
(* Relocation properties ****************************************************)
(* Basic_1: was just: sn3_lift *)
-lemma csx_lift: ∀h,g,G,L2,L1,T1,s,l,m. ⦃G, L1⦄ ⊢ ⬊*[h, g] T1 →
- ∀T2. ⬇[s, l, m] L2 ≡ L1 → ⬆[l, m] T1 ≡ T2 → ⦃G, L2⦄ ⊢ ⬊*[h, g] T2.
-#h #g #G #L2 #L1 #T1 #s #l #m #H elim H -T1 #T1 #_ #IHT1 #T2 #HL21 #HT12
+lemma csx_lift: ∀h,o,G,L2,L1,T1,c,l,k. ⦃G, L1⦄ ⊢ ⬊*[h, o] T1 →
+ ∀T2. ⬇[c, l, k] L2 ≡ L1 → ⬆[l, k] T1 ≡ T2 → ⦃G, L2⦄ ⊢ ⬊*[h, o] T2.
+#h #o #G #L2 #L1 #T1 #c #l #k #H elim H -T1 #T1 #_ #IHT1 #T2 #HL21 #HT12
@csx_intro #T #HLT2 #HT2
elim (cpx_inv_lift1 … HLT2 … HL21 … HT12) -HLT2 #T0 #HT0 #HLT10
@(IHT1 … HLT10) // -L1 -L2 #H destruct
qed.
(* Basic_1: was just: sn3_gen_lift *)
-lemma csx_inv_lift: ∀h,g,G,L2,L1,T1,s,l,m. ⦃G, L1⦄ ⊢ ⬊*[h, g] T1 →
- ∀T2. ⬇[s, l, m] L1 ≡ L2 → ⬆[l, m] T2 ≡ T1 → ⦃G, L2⦄ ⊢ ⬊*[h, g] T2.
-#h #g #G #L2 #L1 #T1 #s #l #m #H elim H -T1 #T1 #_ #IHT1 #T2 #HL12 #HT21
+lemma csx_inv_lift: ∀h,o,G,L2,L1,T1,c,l,k. ⦃G, L1⦄ ⊢ ⬊*[h, o] T1 →
+ ∀T2. ⬇[c, l, k] L1 ≡ L2 → ⬆[l, k] T2 ≡ T1 → ⦃G, L2⦄ ⊢ ⬊*[h, o] T2.
+#h #o #G #L2 #L1 #T1 #c #l #k #H elim H -T1 #T1 #_ #IHT1 #T2 #HL12 #HT21
@csx_intro #T #HLT2 #HT2
-elim (lift_total T l m) #T0 #HT0
+elim (lift_total T l k) #T0 #HT0
lapply (cpx_lift … HLT2 … HL12 … HT21 … HT0) -HLT2 #HLT10
@(IHT1 … HLT10) // -L1 -L2 #H destruct
>(lift_inj … HT0 … HT21) in HT2; -T1 /2 width=1 by/
(* Advanced inversion lemmas ************************************************)
(* Basic_1: was: sn3_gen_def *)
-lemma csx_inv_lref_bind: ∀h,g,I,G,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V →
- ⦃G, L⦄ ⊢ ⬊*[h, g] #i → ⦃G, K⦄ ⊢ ⬊*[h, g] V.
-#h #g #I #G #L #K #V #i #HLK #Hi
+lemma csx_inv_lref_bind: ∀h,o,I,G,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V →
+ ⦃G, L⦄ ⊢ ⬊*[h, o] #i → ⦃G, K⦄ ⊢ ⬊*[h, o] V.
+#h #o #I #G #L #K #V #i #HLK #Hi
elim (lift_total V 0 (i+1))
/4 width=9 by csx_inv_lift, csx_cpx_trans, cpx_delta, drop_fwd_drop2/
qed-.
(* Advanced properties ******************************************************)
(* Basic_1: was just: sn3_abbr *)
-lemma csx_lref_bind: ∀h,g,I,G,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → ⦃G, K⦄ ⊢ ⬊*[h, g] V → ⦃G, L⦄ ⊢ ⬊*[h, g] #i.
-#h #g #I #G #L #K #V #i #HLK #HV
+lemma csx_lref_bind: ∀h,o,I,G,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → ⦃G, K⦄ ⊢ ⬊*[h, o] V → ⦃G, L⦄ ⊢ ⬊*[h, o] #i.
+#h #o #I #G #L #K #V #i #HLK #HV
@csx_intro #X #H #Hi
elim (cpx_inv_lref1 … H) -H
[ #H destruct elim Hi //
]
qed.
-lemma csx_appl_simple: ∀h,g,G,L,V. ⦃G, L⦄ ⊢ ⬊*[h, g] V → ∀T1.
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓐV.T2) →
- 𝐒⦃T1⦄ → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓐV.T1.
-#h #g #G #L #V #H @(csx_ind … H) -V #V #_ #IHV #T1 #IHT1 #HT1
+lemma csx_appl_simple: ∀h,o,G,L,V. ⦃G, L⦄ ⊢ ⬊*[h, o] V → ∀T1.
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV.T2) →
+ 𝐒⦃T1⦄ → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV.T1.
+#h #o #G #L #V #H @(csx_ind … H) -V #V #_ #IHV #T1 #IHT1 #HT1
@csx_intro #X #H1 #H2
elim (cpx_inv_appl1_simple … H1) // -H1
#V0 #T0 #HLV0 #HLT10 #H destruct
]
qed.
-lemma csx_fqu_conf: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ⦃G1, L1⦄ ⊢ ⬊*[h, g] T1 → ⦃G2, L2⦄ ⊢ ⬊*[h, g] T2.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
+lemma csx_fqu_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 → ⦃G2, L2⦄ ⊢ ⬊*[h, o] T2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
/2 width=8 by csx_inv_lref_bind, csx_inv_lift, csx_fwd_flat_dx, csx_fwd_bind_dx, csx_fwd_pair_sn/
qed-.
-lemma csx_fquq_conf: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ⦃G1, L1⦄ ⊢ ⬊*[h, g] T1 → ⦃G2, L2⦄ ⊢ ⬊*[h, g] T2.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H12 #H elim (fquq_inv_gen … H12) -H12
+lemma csx_fquq_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 → ⦃G2, L2⦄ ⊢ ⬊*[h, o] T2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #H elim (fquq_inv_gen … H12) -H12
[ /2 width=5 by csx_fqu_conf/
| * #HG #HL #HT destruct //
]
qed-.
-lemma csx_fqup_conf: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ⦃G1, L1⦄ ⊢ ⬊*[h, g] T1 → ⦃G2, L2⦄ ⊢ ⬊*[h, g] T2.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
+lemma csx_fqup_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
+ ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 → ⦃G2, L2⦄ ⊢ ⬊*[h, o] T2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
/3 width=5 by csx_fqu_conf/
qed-.
-lemma csx_fqus_conf: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
- ⦃G1, L1⦄ ⊢ ⬊*[h, g] T1 → ⦃G2, L2⦄ ⊢ ⬊*[h, g] T2.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H12 #H elim (fqus_inv_gen … H12) -H12
+lemma csx_fqus_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+ ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 → ⦃G2, L2⦄ ⊢ ⬊*[h, o] T2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #H elim (fqus_inv_gen … H12) -H12
[ /2 width=5 by csx_fqup_conf/
| * #HG #HL #HT destruct //
]
(* Main properties **********************************************************)
-theorem csx_gcp: ∀h,g. gcp (cpx h g) (eq …) (csx h g).
-#h #g @mk_gcp
+theorem csx_gcp: ∀h,o. gcp (cpx h o) (eq …) (csx h o).
+#h #o @mk_gcp
[ normalize /3 width=13 by cnx_lift/
-| #G #L elim (deg_total h g 0) /3 width=8 by cnx_sort_iter, ex_intro/
+| #G #L elim (deg_total h o 0) /3 width=8 by cnx_sort_iter, ex_intro/
| /2 width=8 by csx_lift/
| /2 width=3 by csx_fwd_flat_dx/
]
(* Properties on lazy equivalence for local environments ********************)
-lemma csx_lleq_conf: ∀h,g,G,L1,T. ⦃G, L1⦄ ⊢ ⬊*[h, g] T →
- ∀L2. L1 ≡[T, 0] L2 → ⦃G, L2⦄ ⊢ ⬊*[h, g] T.
-#h #g #G #L1 #T #H @(csx_ind … H) -T
+lemma csx_lleq_conf: ∀h,o,G,L1,T. ⦃G, L1⦄ ⊢ ⬊*[h, o] T →
+ ∀L2. L1 ≡[T, 0] L2 → ⦃G, L2⦄ ⊢ ⬊*[h, o] T.
+#h #o #G #L1 #T #H @(csx_ind … H) -T
/4 width=6 by csx_intro, cpx_lleq_conf_dx, lleq_cpx_trans/
qed-.
-lemma csx_lleq_trans: ∀h,g,G,L1,L2,T.
- L1 ≡[T, 0] L2 → ⦃G, L2⦄ ⊢ ⬊*[h, g] T → ⦃G, L1⦄ ⊢ ⬊*[h, g] T.
+lemma csx_lleq_trans: ∀h,o,G,L1,L2,T.
+ L1 ≡[T, 0] L2 → ⦃G, L2⦄ ⊢ ⬊*[h, o] T → ⦃G, L1⦄ ⊢ ⬊*[h, o] T.
/3 width=3 by csx_lleq_conf, lleq_sym/ qed-.
(* Advanced properties ******************************************************)
-lemma csx_lpx_conf: ∀h,g,G,L1,L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 →
- ∀T. ⦃G, L1⦄ ⊢ ⬊*[h, g] T → ⦃G, L2⦄ ⊢ ⬊*[h, g] T.
-#h #g #G #L1 #L2 #HL12 #T #H @(csx_ind_alt … H) -T
+lemma csx_lpx_conf: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 →
+ ∀T. ⦃G, L1⦄ ⊢ ⬊*[h, o] T → ⦃G, L2⦄ ⊢ ⬊*[h, o] T.
+#h #o #G #L1 #L2 #HL12 #T #H @(csx_ind_alt … H) -T
/4 width=3 by csx_intro, lpx_cpx_trans/
qed-.
-lemma csx_abst: ∀h,g,a,G,L,W. ⦃G, L⦄ ⊢ ⬊*[h, g] W →
- ∀T. ⦃G, L.ⓛW⦄ ⊢ ⬊*[h, g] T → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓛ{a}W.T.
-#h #g #a #G #L #W #HW @(csx_ind … HW) -W #W #_ #IHW #T #HT @(csx_ind … HT) -T #T #HT #IHT
+lemma csx_abst: ∀h,o,a,G,L,W. ⦃G, L⦄ ⊢ ⬊*[h, o] W →
+ ∀T. ⦃G, L.ⓛW⦄ ⊢ ⬊*[h, o] T → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓛ{a}W.T.
+#h #o #a #G #L #W #HW @(csx_ind … HW) -W #W #_ #IHW #T #HT @(csx_ind … HT) -T #T #HT #IHT
@csx_intro #X #H1 #H2
elim (cpx_inv_abst1 … H1) -H1
#W0 #T0 #HLW0 #HLT0 #H destruct
]
qed.
-lemma csx_abbr: ∀h,g,a,G,L,V. ⦃G, L⦄ ⊢ ⬊*[h, g] V →
- ∀T. ⦃G, L.ⓓV⦄ ⊢ ⬊*[h, g] T → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓓ{a}V. T.
-#h #g #a #G #L #V #HV elim HV -V #V #_ #IHV #T #HT @(csx_ind_alt … HT) -T #T #HT #IHT
+lemma csx_abbr: ∀h,o,a,G,L,V. ⦃G, L⦄ ⊢ ⬊*[h, o] V →
+ ∀T. ⦃G, L.ⓓV⦄ ⊢ ⬊*[h, o] T → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓓ{a}V. T.
+#h #o #a #G #L #V #HV elim HV -V #V #_ #IHV #T #HT @(csx_ind_alt … HT) -T #T #HT #IHT
@csx_intro #X #H1 #H2
elim (cpx_inv_abbr1 … H1) -H1 *
[ #V1 #T1 #HLV1 #HLT1 #H destruct
]
qed.
-fact csx_appl_beta_aux: ∀h,g,a,G,L,U1. ⦃G, L⦄ ⊢ ⬊*[h, g] U1 →
- ∀V,W,T1. U1 = ⓓ{a}ⓝW.V.T1 → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓐV.ⓛ{a}W.T1.
-#h #g #a #G #L #X #H @(csx_ind … H) -X
+fact csx_appl_beta_aux: ∀h,o,a,G,L,U1. ⦃G, L⦄ ⊢ ⬊*[h, o] U1 →
+ ∀V,W,T1. U1 = ⓓ{a}ⓝW.V.T1 → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV.ⓛ{a}W.T1.
+#h #o #a #G #L #X #H @(csx_ind … H) -X
#X #HT1 #IHT1 #V #W #T1 #H1 destruct
@csx_intro #X #H1 #H2
elim (cpx_inv_appl1 … H1) -H1 *
qed-.
(* Basic_1: was just: sn3_beta *)
-lemma csx_appl_beta: ∀h,g,a,G,L,V,W,T. ⦃G, L⦄ ⊢ ⬊*[h, g] ⓓ{a}ⓝW.V.T → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓐV.ⓛ{a}W.T.
+lemma csx_appl_beta: ∀h,o,a,G,L,V,W,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⓓ{a}ⓝW.V.T → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV.ⓛ{a}W.T.
/2 width=3 by csx_appl_beta_aux/ qed.
-fact csx_appl_theta_aux: ∀h,g,a,G,L,U. ⦃G, L⦄ ⊢ ⬊*[h, g] U → ∀V1,V2. ⬆[0, 1] V1 ≡ V2 →
- ∀V,T. U = ⓓ{a}V.ⓐV2.T → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓐV1.ⓓ{a}V.T.
-#h #g #a #G #L #X #H @(csx_ind_alt … H) -X #X #HVT #IHVT #V1 #V2 #HV12 #V #T #H destruct
+fact csx_appl_theta_aux: ∀h,o,a,G,L,U. ⦃G, L⦄ ⊢ ⬊*[h, o] U → ∀V1,V2. ⬆[0, 1] V1 ≡ V2 →
+ ∀V,T. U = ⓓ{a}V.ⓐV2.T → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV1.ⓓ{a}V.T.
+#h #o #a #G #L #X #H @(csx_ind_alt … H) -X #X #HVT #IHVT #V1 #V2 #HV12 #V #T #H destruct
lapply (csx_fwd_pair_sn … HVT) #HV
lapply (csx_fwd_bind_dx … HVT) -HVT #HVT
@csx_intro #X #HL #H
]
qed-.
-lemma csx_appl_theta: ∀h,g,a,V1,V2. ⬆[0, 1] V1 ≡ V2 →
- ∀G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, g] ⓓ{a}V.ⓐV2.T → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓐV1.ⓓ{a}V.T.
+lemma csx_appl_theta: ∀h,o,a,V1,V2. ⬆[0, 1] V1 ≡ V2 →
+ ∀G,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⓓ{a}V.ⓐV2.T → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV1.ⓓ{a}V.T.
/2 width=5 by csx_appl_theta_aux/ qed.
(* Basic_1: was just: sn3_appl_appl *)
-lemma csx_appl_simple_tsts: ∀h,g,G,L,V. ⦃G, L⦄ ⊢ ⬊*[h, g] V → ∀T1. ⦃G, L⦄ ⊢ ⬊*[h, g] T1 →
- (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → (T1 ≂ T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓐV.T2) →
- 𝐒⦃T1⦄ → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓐV.T1.
-#h #g #G #L #V #H @(csx_ind … H) -V #V #_ #IHV #T1 #H @(csx_ind … H) -T1 #T1 #H1T1 #IHT1 #H2T1 #H3T1
+lemma csx_appl_simple_tsts: ∀h,o,G,L,V. ⦃G, L⦄ ⊢ ⬊*[h, o] V → ∀T1. ⦃G, L⦄ ⊢ ⬊*[h, o] T1 →
+ (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → (T1 ≂ T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV.T2) →
+ 𝐒⦃T1⦄ → ⦃G, L⦄ ⊢ ⬊*[h, o] ⓐV.T1.
+#h #o #G #L #V #H @(csx_ind … H) -V #V #_ #IHV #T1 #H @(csx_ind … H) -T1 #T1 #H1T1 #IHT1 #H2T1 #H3T1
@csx_intro #X #HL #H
elim (cpx_inv_appl1_simple … HL) -HL //
#V0 #T0 #HLV0 #HLT10 #H0 destruct
(* Properties on sn extended parallel computation for local environments ****)
-lemma csx_lpxs_conf: ∀h,g,G,L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 →
- ∀T. ⦃G, L1⦄ ⊢ ⬊*[h, g] T → ⦃G, L2⦄ ⊢ ⬊*[h, g] T.
-#h #g #G #L1 #L2 #H @(lpxs_ind … H) -L2 /3 by lpxs_strap1, csx_lpx_conf/
+lemma csx_lpxs_conf: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 →
+ ∀T. ⦃G, L1⦄ ⊢ ⬊*[h, o] T → ⦃G, L2⦄ ⊢ ⬊*[h, o] T.
+#h #o #G #L1 #L2 #H @(lpxs_ind … H) -L2 /3 by lpxs_strap1, csx_lpx_conf/
qed-.
(* Advanced properties ******************************************************)
(* Basic_1: was just: sn3_appls_lref *)
-lemma csx_applv_cnx: ∀h,g,G,L,T. 𝐒⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T⦄ →
- ∀Vs. ⦃G, L⦄ ⊢ ⬊*[h, g] Vs → ⦃G, L⦄ ⊢ ⬊*[h, g] ⒶVs.T.
-#h #g #G #L #T #H1T #H2T #Vs elim Vs -Vs [ #_ @(cnx_csx … H2T) ] (**) (* /2 width=1/ does not work *)
+lemma csx_applv_cnx: ∀h,o,G,L,T. 𝐒⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ →
+ ∀Vs. ⦃G, L⦄ ⊢ ⬊*[h, o] Vs → ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶVs.T.
+#h #o #G #L #T #H1T #H2T #Vs elim Vs -Vs [ #_ @(cnx_csx … H2T) ] (**) (* /2 width=1/ does not work *)
#V #Vs #IHV #H
elim (csxv_inv_cons … H) -H #HV #HVs
@csx_appl_simple_tsts /2 width=1 by applv_simple/ -IHV -HV -HVs
elim (H0) -H0 //
qed.
-lemma csx_applv_sort: ∀h,g,G,L,k,Vs. ⦃G, L⦄ ⊢ ⬊*[h, g] Vs → ⦃G, L⦄ ⊢ ⬊*[h, g] ⒶVs.⋆k.
-#h #g #G #L #k elim (deg_total h g k)
-#d generalize in match k; -k @(nat_ind_plus … d) -d [ /3 width=6 by csx_applv_cnx, cnx_sort, simple_atom/ ]
-#d #IHd #k #Hkd lapply (deg_next_SO … Hkd) -Hkd
+lemma csx_applv_sort: ∀h,o,G,L,s,Vs. ⦃G, L⦄ ⊢ ⬊*[h, o] Vs → ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶVs.⋆s.
+#h #o #G #L #s elim (deg_total h o s)
+#d generalize in match s; -s @(nat_ind_plus … d) -d [ /3 width=6 by csx_applv_cnx, cnx_sort, simple_atom/ ]
+#d #IHd #s #Hkd lapply (deg_next_SO … Hkd) -Hkd
#Hkd #Vs elim Vs -Vs /2 width=1 by/
#V #Vs #IHVs #HVVs
elim (csxv_inv_cons … HVVs) #HV #HVs
#X #H #H0
elim (cpxs_fwd_sort_vector … H) -H #H
[ elim H0 -H0 //
-| -H0 @(csx_cpxs_trans … (Ⓐ(V@Vs).⋆(next h k))) /2 width=1 by cpxs_flat_dx/
+| -H0 @(csx_cpxs_trans … (Ⓐ(V@Vs).⋆(next h s))) /2 width=1 by cpxs_flat_dx/
]
qed.
(* Basic_1: was just: sn3_appls_beta *)
-lemma csx_applv_beta: ∀h,g,a,G,L,Vs,V,W,T. ⦃G, L⦄ ⊢ ⬊*[h, g] ⒶVs.ⓓ{a}ⓝW.V.T →
- ⦃G, L⦄ ⊢ ⬊*[h, g] ⒶVs. ⓐV.ⓛ{a}W.T.
-#h #g #a #G #L #Vs elim Vs -Vs /2 width=1 by csx_appl_beta/
+lemma csx_applv_beta: ∀h,o,a,G,L,Vs,V,W,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶVs.ⓓ{a}ⓝW.V.T →
+ ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶVs. ⓐV.ⓛ{a}W.T.
+#h #o #a #G #L #Vs elim Vs -Vs /2 width=1 by csx_appl_beta/
#V0 #Vs #IHV #V #W #T #H1T
lapply (csx_fwd_pair_sn … H1T) #HV0
lapply (csx_fwd_flat_dx … H1T) #H2T
]
qed.
-lemma csx_applv_delta: ∀h,g,I,G,L,K,V1,i. ⬇[i] L ≡ K.ⓑ{I}V1 →
+lemma csx_applv_delta: ∀h,o,I,G,L,K,V1,i. ⬇[i] L ≡ K.ⓑ{I}V1 →
∀V2. ⬆[0, i + 1] V1 ≡ V2 →
- ∀Vs. ⦃G, L⦄ ⊢ ⬊*[h, g] (ⒶVs.V2) → ⦃G, L⦄ ⊢ ⬊*[h, g] (ⒶVs.#i).
-#h #g #I #G #L #K #V1 #i #HLK #V2 #HV12 #Vs elim Vs -Vs
+ ∀Vs. ⦃G, L⦄ ⊢ ⬊*[h, o] (ⒶVs.V2) → ⦃G, L⦄ ⊢ ⬊*[h, o] (ⒶVs.#i).
+#h #o #I #G #L #K #V1 #i #HLK #V2 #HV12 #Vs elim Vs -Vs
[ /4 width=12 by csx_inv_lift, csx_lref_bind, drop_fwd_drop2/
| #V #Vs #IHV #H1T
lapply (csx_fwd_pair_sn … H1T) #HV
qed.
(* Basic_1: was just: sn3_appls_abbr *)
-lemma csx_applv_theta: ∀h,g,a,G,L,V1s,V2s. ⬆[0, 1] V1s ≡ V2s →
- ∀V,T. ⦃G, L⦄ ⊢ ⬊*[h, g] ⓓ{a}V.ⒶV2s.T →
- ⦃G, L⦄ ⊢ ⬊*[h, g] ⒶV1s.ⓓ{a}V.T.
-#h #g #a #G #L #V1s #V2s * -V1s -V2s /2 width=1 by/
-#V1s #V2s #V1 #V2 #HV12 #H
+lemma csx_applv_theta: ∀h,o,a,G,L,V1c,V2c. ⬆[0, 1] V1c ≡ V2c →
+ ∀V,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⓓ{a}V.ⒶV2c.T →
+ ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶV1c.ⓓ{a}V.T.
+#h #o #a #G #L #V1c #V2c * -V1c -V2c /2 width=1 by/
+#V1c #V2c #V1 #V2 #HV12 #H
generalize in match HV12; -HV12 generalize in match V2; -V2 generalize in match V1; -V1
-elim H -V1s -V2s /2 width=3 by csx_appl_theta/
-#V1s #V2s #V1 #V2 #HV12 #HV12s #IHV12s #W1 #W2 #HW12 #V #T #H
+elim H -V1c -V2c /2 width=3 by csx_appl_theta/
+#V1c #V2c #V1 #V2 #HV12 #HV12c #IHV12c #W1 #W2 #HW12 #V #T #H
lapply (csx_appl_theta … HW12 … H) -H -HW12 #H
lapply (csx_fwd_pair_sn … H) #HW1
lapply (csx_fwd_flat_dx … H) #H1
-@csx_appl_simple_tsts /2 width=3 by simple_flat/ -IHV12s -HW1 -H1 #X #H1 #H2
-elim (cpxs_fwd_theta_vector … (V2@V2s) … H1) -H1 /2 width=1 by liftv_cons/ -HV12s -HV12
+@csx_appl_simple_tsts /2 width=3 by simple_flat/ -IHV12c -HW1 -H1 #X #H1 #H2
+elim (cpxs_fwd_theta_vector … (V2@V2c) … H1) -H1 /2 width=1 by liftv_cons/ -HV12c -HV12
[ -H #H elim H2 -H2 //
| -H2 /3 width=5 by csx_cpxs_trans, cpxs_flat_dx/
]
qed.
(* Basic_1: was just: sn3_appls_cast *)
-lemma csx_applv_cast: ∀h,g,G,L,Vs,W,T. ⦃G, L⦄ ⊢ ⬊*[h, g] ⒶVs.W → ⦃G, L⦄ ⊢ ⬊*[h, g] ⒶVs.T →
- ⦃G, L⦄ ⊢ ⬊*[h, g] ⒶVs.ⓝW.T.
-#h #g #G #L #Vs elim Vs -Vs /2 width=1 by csx_cast/
+lemma csx_applv_cast: ∀h,o,G,L,Vs,W,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶVs.W → ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶVs.T →
+ ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶVs.ⓝW.T.
+#h #o #G #L #Vs elim Vs -Vs /2 width=1 by csx_cast/
#V #Vs #IHV #W #T #H1W #H1T
lapply (csx_fwd_pair_sn … H1W) #HV
lapply (csx_fwd_flat_dx … H1W) #H2W
]
qed.
-theorem csx_gcr: ∀h,g. gcr (cpx h g) (eq …) (csx h g) (csx h g).
-#h #g @mk_gcr //
+theorem csx_gcr: ∀h,o. gcr (cpx h o) (eq …) (csx h o) (csx h o).
+#h #o @mk_gcr //
[ /3 width=1 by csx_applv_cnx/
|2,3,6: /2 width=1 by csx_applv_beta, csx_applv_sort, csx_applv_cast/
| /2 width=7 by csx_applv_delta/
-| #G #L #V1s #V2s #HV12s #a #V #T #H #HV
- @(csx_applv_theta … HV12s) -HV12s
+| #G #L #V1c #V2c #HV12c #a #V #T #H #HV
+ @(csx_applv_theta … HV12c) -HV12c
@csx_abbr //
]
qed.
(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERM VECTORS *************)
definition csxv: ∀h. sd h → relation3 genv lenv (list term) ≝
- λh,g,G,L. all … (csx h g G L).
+ λh,o,G,L. all … (csx h o G L).
interpretation
"context-sensitive strong normalization (term vector)"
- 'SN h g G L Ts = (csxv h g G L Ts).
+ 'SN h o G L Ts = (csxv h o G L Ts).
(* Basic inversion lemmas ***************************************************)
-lemma csxv_inv_cons: ∀h,g,G,L,T,Ts. ⦃G, L⦄ ⊢ ⬊*[h, g] T @ Ts →
- ⦃G, L⦄ ⊢ ⬊*[h, g] T ∧ ⦃G, L⦄ ⊢ ⬊*[h, g] Ts.
+lemma csxv_inv_cons: ∀h,o,G,L,T,Ts. ⦃G, L⦄ ⊢ ⬊*[h, o] T @ Ts →
+ ⦃G, L⦄ ⊢ ⬊*[h, o] T ∧ ⦃G, L⦄ ⊢ ⬊*[h, o] Ts.
normalize // qed-.
(* Basic forward lemmas *****************************************************)
-lemma csx_fwd_applv: ∀h,g,G,L,T,Vs. ⦃G, L⦄ ⊢ ⬊*[h, g] Ⓐ Vs.T →
- ⦃G, L⦄ ⊢ ⬊*[h, g] Vs ∧ ⦃G, L⦄ ⊢ ⬊*[h, g] T.
-#h #g #G #L #T #Vs elim Vs -Vs /2 width=1 by conj/
+lemma csx_fwd_applv: ∀h,o,G,L,T,Vs. ⦃G, L⦄ ⊢ ⬊*[h, o] Ⓐ Vs.T →
+ ⦃G, L⦄ ⊢ ⬊*[h, o] Vs ∧ ⦃G, L⦄ ⊢ ⬊*[h, o] T.
+#h #o #G #L #T #Vs elim Vs -Vs /2 width=1 by conj/
#V #Vs #IHVs #HVs
lapply (csx_fwd_pair_sn … HVs) #HV
lapply (csx_fwd_flat_dx … HVs) -HVs #HVs
(* "QRST" PROPER PARALLEL COMPUTATION FOR CLOSURES **************************)
definition fpbg: ∀h. sd h → tri_relation genv lenv term ≝
- λh,g,G1,L1,T1,G2,L2,T2.
- ∃∃G,L,T. ⦃G1, L1, T1⦄ ≻[h, g] ⦃G, L, T⦄ & ⦃G, L, T⦄ ≥[h, g] ⦃G2, L2, T2⦄.
+ λh,o,G1,L1,T1,G2,L2,T2.
+ ∃∃G,L,T. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G, L, T⦄ & ⦃G, L, T⦄ ≥[h, o] ⦃G2, L2, T2⦄.
interpretation "'qrst' proper parallel computation (closure)"
- 'LazyBTPRedStarProper h g G1 L1 T1 G2 L2 T2 = (fpbg h g G1 L1 T1 G2 L2 T2).
+ 'LazyBTPRedStarProper h o G1 L1 T1 G2 L2 T2 = (fpbg h o G1 L1 T1 G2 L2 T2).
(* Basic properties *********************************************************)
-lemma fpb_fpbg: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄.
+lemma fpb_fpbg: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
/2 width=5 by ex2_3_intro/ qed.
-lemma fpbg_fpbq_trans: ∀h,g,G1,G,G2,L1,L,L2,T1,T,T2.
- ⦃G1, L1, T1⦄ >≡[h, g] ⦃G, L, T⦄ → ⦃G, L, T⦄ ≽[h, g] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄.
-#h #g #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 *
+lemma fpbg_fpbq_trans: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2.
+ ⦃G1, L1, T1⦄ >≡[h, o] ⦃G, L, T⦄ → ⦃G, L, T⦄ ≽[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 *
/3 width=9 by fpbs_strap1, ex2_3_intro/
qed-.
(* Properties on lazy equivalence for closures ******************************)
-lemma fpbg_fleq_trans: ∀h,g,G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ >≡[h, g] ⦃G, L, T⦄ →
- ∀G2,L2,T2. ⦃G, L, T⦄ ≡[0] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄.
+lemma fpbg_fleq_trans: ∀h,o,G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G, L, T⦄ →
+ ∀G2,L2,T2. ⦃G, L, T⦄ ≡[0] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
/3 width=5 by fpbg_fpbq_trans, fleq_fpbq/ qed-.
-lemma fleq_fpbg_trans: ∀h,g,G,G2,L,L2,T,T2. ⦃G, L, T⦄ >≡[h, g] ⦃G2, L2, T2⦄ →
- ∀G1,L1,T1. ⦃G1, L1, T1⦄ ≡[0] ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄.
-#h #g #G #G2 #L #L2 #T #T2 * #G0 #L0 #T0 #H0 #H02 #G1 #L1 #T1 #H1
+lemma fleq_fpbg_trans: ∀h,o,G,G2,L,L2,T,T2. ⦃G, L, T⦄ >≡[h, o] ⦃G2, L2, T2⦄ →
+ ∀G1,L1,T1. ⦃G1, L1, T1⦄ ≡[0] ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
+#h #o #G #G2 #L #L2 #T #T2 * #G0 #L0 #T0 #H0 #H02 #G1 #L1 #T1 #H1
elim (fleq_fpb_trans … H1 … H0) -G -L -T
/4 width=9 by fpbs_strap2, fleq_fpbq, ex2_3_intro/
qed-.
(* alternative definition of fpbs *******************************************)
-lemma fleq_fpbs: ∀h,g,G1,G2,L1,L2,T1,T2.
- ⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 * /2 width=1 by lleq_fpbs/
+lemma fleq_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2.
+ ⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 * /2 width=1 by lleq_fpbs/
qed.
-lemma fpbg_fwd_fpbs: ∀h,g,G1,G2,L1,L2,T1,T2.
- ⦃G1, L1, T1⦄ >≡[h,g] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 *
+lemma fpbg_fwd_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2.
+ ⦃G1, L1, T1⦄ >≡[h,o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 *
/3 width=5 by fpbs_strap2, fpb_fpbq/
qed-.
-lemma fpbs_fpbg: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ →
+lemma fpbs_fpbg: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ ∨
- ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H @(fpbs_ind … H) -G2 -L2 -T2
+ ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fpbs_ind … H) -G2 -L2 -T2
[ /2 width=1 by or_introl/
| #G #G2 #L #L2 #T #T2 #_ #H2 * #H1 @(fpbq_ind_alt … H2) -H2 #H2
[ /3 width=5 by fleq_trans, or_introl/
(* Advanced properties of "qrst" parallel computation on closures ***********)
-lemma fpbs_fpb_trans: ∀h,g,F1,F2,K1,K2,T1,T2. ⦃F1, K1, T1⦄ ≥[h, g] ⦃F2, K2, T2⦄ →
- ∀G2,L2,U2. ⦃F2, K2, T2⦄ ≻[h, g] ⦃G2, L2, U2⦄ →
- ∃∃G1,L1,U1. ⦃F1, K1, T1⦄ ≻[h, g] ⦃G1, L1, U1⦄ & ⦃G1, L1, U1⦄ ≥[h, g] ⦃G2, L2, U2⦄.
-#h #g #F1 #F2 #K1 #K2 #T1 #T2 #H elim (fpbs_fpbg … H) -H
+lemma fpbs_fpb_trans: ∀h,o,F1,F2,K1,K2,T1,T2. ⦃F1, K1, T1⦄ ≥[h, o] ⦃F2, K2, T2⦄ →
+ ∀G2,L2,U2. ⦃F2, K2, T2⦄ ≻[h, o] ⦃G2, L2, U2⦄ →
+ ∃∃G1,L1,U1. ⦃F1, K1, T1⦄ ≻[h, o] ⦃G1, L1, U1⦄ & ⦃G1, L1, U1⦄ ≥[h, o] ⦃G2, L2, U2⦄.
+#h #o #F1 #F2 #K1 #K2 #T1 #T2 #H elim (fpbs_fpbg … H) -H
[ #H12 #G2 #L2 #U2 #H2 elim (fleq_fpb_trans … H12 … H2) -F2 -K2 -T2
/3 width=5 by fleq_fpbs, ex2_3_intro/
| * #H1 #H2 #H3 #H4 #H5 #H6 #H7 #H8 #H9
(* Main properties **********************************************************)
-theorem fpbg_trans: ∀h,g. tri_transitive … (fpbg h g).
+theorem fpbg_trans: ∀h,o. tri_transitive … (fpbg h o).
/3 width=5 by fpbg_fpbs_trans, fpbg_fwd_fpbs/ qed-.
(* Properties on "qrst" parallel reduction on closures **********************)
-lemma fpb_fpbg_trans: ∀h,g,G1,G,G2,L1,L,L2,T1,T,T2.
- ⦃G1, L1, T1⦄ ≻[h, g] ⦃G, L, T⦄ → ⦃G, L, T⦄ >≡[h, g] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄.
+lemma fpb_fpbg_trans: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2.
+ ⦃G1, L1, T1⦄ ≻[h, o] ⦃G, L, T⦄ → ⦃G, L, T⦄ >≡[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
/3 width=5 by fpbg_fwd_fpbs, ex2_3_intro/ qed-.
-lemma fpbq_fpbg_trans: ∀h,g,G1,G,G2,L1,L,L2,T1,T,T2.
- ⦃G1, L1, T1⦄ ≽[h, g] ⦃G, L, T⦄ → ⦃G, L, T⦄ >≡[h, g] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄.
-#h #g #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H2 @(fpbq_ind_alt … H1) -H1
+lemma fpbq_fpbg_trans: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2.
+ ⦃G1, L1, T1⦄ ≽[h, o] ⦃G, L, T⦄ → ⦃G, L, T⦄ >≡[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H2 @(fpbq_ind_alt … H1) -H1
/2 width=5 by fleq_fpbg_trans, fpb_fpbg_trans/
qed-.
(* Properties on "qrst" parallel compuutation on closures *******************)
-lemma fpbs_fpbg_trans: ∀h,g,G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G, L, T⦄ →
- ∀G2,L2,T2. ⦃G, L, T⦄ >≡[h, g] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄.
-#h #g #G1 #G #L1 #L #T1 #T #H @(fpbs_ind … H) -G -L -T /3 width=5 by fpbq_fpbg_trans/
+lemma fpbs_fpbg_trans: ∀h,o,G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ →
+ ∀G2,L2,T2. ⦃G, L, T⦄ >≡[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G #L1 #L #T1 #T #H @(fpbs_ind … H) -G -L -T /3 width=5 by fpbq_fpbg_trans/
qed-.
(* Note: this is used in the closure proof *)
-lemma fpbg_fpbs_trans: ∀h,g,G,G2,L,L2,T,T2. ⦃G, L, T⦄ ≥[h, g] ⦃G2, L2, T2⦄ →
- ∀G1,L1,T1. ⦃G1, L1, T1⦄ >≡[h, g] ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄.
-#h #g #G #G2 #L #L2 #T #T2 #H @(fpbs_ind_dx … H) -G -L -T /3 width=5 by fpbg_fpbq_trans/
+lemma fpbg_fpbs_trans: ∀h,o,G,G2,L,L2,T,T2. ⦃G, L, T⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
+ ∀G1,L1,T1. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
+#h #o #G #G2 #L #L2 #T #T2 #H @(fpbs_ind_dx … H) -G -L -T /3 width=5 by fpbg_fpbq_trans/
qed-.
(* Note: this is used in the closure proof *)
-lemma fqup_fpbg: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H elim (fqup_inv_step_sn … H) -H
+lemma fqup_fpbg: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim (fqup_inv_step_sn … H) -H
/3 width=5 by fqus_fpbs, fpb_fqu, ex2_3_intro/
qed.
-lemma cpxs_fpbg: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 →
- (T1 = T2 → ⊥) → ⦃G, L, T1⦄ >≡[h, g] ⦃G, L, T2⦄.
-#h #g #G #L #T1 #T2 #H #H0 elim (cpxs_neq_inv_step_sn … H … H0) -H -H0
+lemma cpxs_fpbg: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 →
+ (T1 = T2 → ⊥) → ⦃G, L, T1⦄ >≡[h, o] ⦃G, L, T2⦄.
+#h #o #G #L #T1 #T2 #H #H0 elim (cpxs_neq_inv_step_sn … H … H0) -H -H0
/4 width=5 by cpxs_fpbs, fpb_cpx, ex2_3_intro/
qed.
-lemma lstas_fpbg: ∀h,g,G,L,T1,T2,d2. ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 → (T1 = T2 → ⊥) →
- ∀d1. d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 ▪[h, g] d1 → ⦃G, L, T1⦄ >≡[h, g] ⦃G, L, T2⦄.
+lemma lstas_fpbg: ∀h,o,G,L,T1,T2,d2. ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 → (T1 = T2 → ⊥) →
+ ∀d1. d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 ▪[h, o] d1 → ⦃G, L, T1⦄ >≡[h, o] ⦃G, L, T2⦄.
/3 width=5 by lstas_cpxs, cpxs_fpbg/ qed.
-lemma lpxs_fpbg: ∀h,g,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 →
- (L1 ≡[T, 0] L2 → ⊥) → ⦃G, L1, T⦄ >≡[h, g] ⦃G, L2, T⦄.
-#h #g #G #L1 #L2 #T #H #H0 elim (lpxs_nlleq_inv_step_sn … H … H0) -H -H0
+lemma lpxs_fpbg: ∀h,o,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 →
+ (L1 ≡[T, 0] L2 → ⊥) → ⦃G, L1, T⦄ >≡[h, o] ⦃G, L2, T⦄.
+#h #o #G #L1 #L2 #T #H #H0 elim (lpxs_nlleq_inv_step_sn … H … H0) -H -H0
/4 width=5 by fpb_lpx, lpxs_lleq_fpbs, ex2_3_intro/
qed.
(* Advanced properties ******************************************************)
-lemma sta_fpbg: ∀h,g,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 ▪[h, g] d+1 →
- ⦃G, L⦄ ⊢ T1 •*[h, 1] T2 → ⦃G, L, T1⦄ >≡[h, g] ⦃G, L, T2⦄.
+lemma sta_fpbg: ∀h,o,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 ▪[h, o] d+1 →
+ ⦃G, L⦄ ⊢ T1 •*[h, 1] T2 → ⦃G, L, T1⦄ >≡[h, o] ⦃G, L, T2⦄.
/4 width=2 by fpb_fpbg, sta_fpb/ qed.
(* "QRST" PARALLEL COMPUTATION FOR CLOSURES *********************************)
definition fpbs: ∀h. sd h → tri_relation genv lenv term ≝
- λh,g. tri_TC … (fpbq h g).
+ λh,o. tri_TC … (fpbq h o).
interpretation "'qrst' parallel computation (closure)"
- 'BTPRedStar h g G1 L1 T1 G2 L2 T2 = (fpbs h g G1 L1 T1 G2 L2 T2).
+ 'BTPRedStar h o G1 L1 T1 G2 L2 T2 = (fpbs h o G1 L1 T1 G2 L2 T2).
(* Basic eliminators ********************************************************)
-lemma fpbs_ind: ∀h,g,G1,L1,T1. ∀R:relation3 genv lenv term. R G1 L1 T1 →
- (∀G,G2,L,L2,T,T2. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G, L, T⦄ → ⦃G, L, T⦄ ≽[h, g] ⦃G2, L2, T2⦄ → R G L T → R G2 L2 T2) →
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ → R G2 L2 T2.
+lemma fpbs_ind: ∀h,o,G1,L1,T1. ∀R:relation3 genv lenv term. R G1 L1 T1 →
+ (∀G,G2,L,L2,T,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ → ⦃G, L, T⦄ ≽[h, o] ⦃G2, L2, T2⦄ → R G L T → R G2 L2 T2) →
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2.
/3 width=8 by tri_TC_star_ind/ qed-.
-lemma fpbs_ind_dx: ∀h,g,G2,L2,T2. ∀R:relation3 genv lenv term. R G2 L2 T2 →
- (∀G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ≽[h, g] ⦃G, L, T⦄ → ⦃G, L, T⦄ ≥[h, g] ⦃G2, L2, T2⦄ → R G L T → R G1 L1 T1) →
- ∀G1,L1,T1. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ → R G1 L1 T1.
+lemma fpbs_ind_dx: ∀h,o,G2,L2,T2. ∀R:relation3 genv lenv term. R G2 L2 T2 →
+ (∀G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G, L, T⦄ → ⦃G, L, T⦄ ≥[h, o] ⦃G2, L2, T2⦄ → R G L T → R G1 L1 T1) →
+ ∀G1,L1,T1. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → R G1 L1 T1.
/3 width=8 by tri_TC_star_ind_dx/ qed-.
(* Basic properties *********************************************************)
-lemma fpbs_refl: ∀h,g. tri_reflexive … (fpbs h g).
+lemma fpbs_refl: ∀h,o. tri_reflexive … (fpbs h o).
/2 width=1 by tri_inj/ qed.
-lemma fpbq_fpbs: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≽[h, g] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄.
+lemma fpbq_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
/2 width=1 by tri_inj/ qed.
-lemma fpbs_strap1: ∀h,g,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G, L, T⦄ →
- ⦃G, L, T⦄ ≽[h, g] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄.
+lemma fpbs_strap1: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ →
+ ⦃G, L, T⦄ ≽[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
/2 width=5 by tri_step/ qed-.
-lemma fpbs_strap2: ∀h,g,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ≽[h, g] ⦃G, L, T⦄ →
- ⦃G, L, T⦄ ≥[h, g] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄.
+lemma fpbs_strap2: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G, L, T⦄ →
+ ⦃G, L, T⦄ ≥[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
/2 width=5 by tri_TC_strap/ qed-.
-lemma fqup_fpbs: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
+lemma fqup_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
/4 width=5 by fqu_fquq, fpbq_fquq, tri_step/
qed.
-lemma fqus_fpbs: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqus_ind … H) -G2 -L2 -T2
+lemma fqus_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqus_ind … H) -G2 -L2 -T2
/3 width=5 by fpbq_fquq, tri_step/
qed.
-lemma cpxs_fpbs: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → ⦃G, L, T1⦄ ≥[h, g] ⦃G, L, T2⦄.
-#h #g #G #L #T1 #T2 #H @(cpxs_ind … H) -T2
+lemma cpxs_fpbs: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2 → ⦃G, L, T1⦄ ≥[h, o] ⦃G, L, T2⦄.
+#h #o #G #L #T1 #T2 #H @(cpxs_ind … H) -T2
/3 width=5 by fpbq_cpx, fpbs_strap1/
qed.
-lemma lpxs_fpbs: ∀h,g,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → ⦃G, L1, T⦄ ≥[h, g] ⦃G, L2, T⦄.
-#h #g #G #L1 #L2 #T #H @(lpxs_ind … H) -L2
+lemma lpxs_fpbs: ∀h,o,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → ⦃G, L1, T⦄ ≥[h, o] ⦃G, L2, T⦄.
+#h #o #G #L1 #L2 #T #H @(lpxs_ind … H) -L2
/3 width=5 by fpbq_lpx, fpbs_strap1/
qed.
-lemma lleq_fpbs: ∀h,g,G,L1,L2,T. L1 ≡[T, 0] L2 → ⦃G, L1, T⦄ ≥[h, g] ⦃G, L2, T⦄.
+lemma lleq_fpbs: ∀h,o,G,L1,L2,T. L1 ≡[T, 0] L2 → ⦃G, L1, T⦄ ≥[h, o] ⦃G, L2, T⦄.
/3 width=1 by fpbq_fpbs, fpbq_lleq/ qed.
-lemma cprs_fpbs: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → ⦃G, L, T1⦄ ≥[h, g] ⦃G, L, T2⦄.
+lemma cprs_fpbs: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → ⦃G, L, T1⦄ ≥[h, o] ⦃G, L, T2⦄.
/3 width=1 by cprs_cpxs, cpxs_fpbs/ qed.
-lemma lprs_fpbs: ∀h,g,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡* L2 → ⦃G, L1, T⦄ ≥[h, g] ⦃G, L2, T⦄.
+lemma lprs_fpbs: ∀h,o,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡* L2 → ⦃G, L1, T⦄ ≥[h, o] ⦃G, L2, T⦄.
/3 width=1 by lprs_lpxs, lpxs_fpbs/ qed.
-lemma fpbs_fqus_trans: ∀h,g,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G, L, T⦄ →
- ⦃G, L, T⦄ ⊐* ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄.
-#h #g #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H @(fqus_ind … H) -G2 -L2 -T2
+lemma fpbs_fqus_trans: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ →
+ ⦃G, L, T⦄ ⊐* ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H @(fqus_ind … H) -G2 -L2 -T2
/3 width=5 by fpbs_strap1, fpbq_fquq/
qed-.
-lemma fpbs_fqup_trans: ∀h,g,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G, L, T⦄ →
- ⦃G, L, T⦄ ⊐+ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄.
+lemma fpbs_fqup_trans: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ →
+ ⦃G, L, T⦄ ⊐+ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
/3 width=5 by fpbs_fqus_trans, fqup_fqus/ qed-.
-lemma fpbs_cpxs_trans: ∀h,g,G1,G,L1,L,T1,T,T2. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G, L, T⦄ →
- ⦃G, L⦄ ⊢ T ➡*[h, g] T2 → ⦃G1, L1, T1⦄ ≥[h, g] ⦃G, L, T2⦄.
-#h #g #G1 #G #L1 #L #T1 #T #T2 #H1 #H @(cpxs_ind … H) -T2
+lemma fpbs_cpxs_trans: ∀h,o,G1,G,L1,L,T1,T,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ →
+ ⦃G, L⦄ ⊢ T ➡*[h, o] T2 → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T2⦄.
+#h #o #G1 #G #L1 #L #T1 #T #T2 #H1 #H @(cpxs_ind … H) -T2
/3 width=5 by fpbs_strap1, fpbq_cpx/
qed-.
-lemma fpbs_lpxs_trans: ∀h,g,G1,G,L1,L,L2,T1,T. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G, L, T⦄ →
- ⦃G, L⦄ ⊢ ➡*[h, g] L2 → ⦃G1, L1, T1⦄ ≥[h, g] ⦃G, L2, T⦄.
-#h #g #G1 #G #L1 #L #L2 #T1 #T #H1 #H @(lpxs_ind … H) -L2
+lemma fpbs_lpxs_trans: ∀h,o,G1,G,L1,L,L2,T1,T. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ →
+ ⦃G, L⦄ ⊢ ➡*[h, o] L2 → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L2, T⦄.
+#h #o #G1 #G #L1 #L #L2 #T1 #T #H1 #H @(lpxs_ind … H) -L2
/3 width=5 by fpbs_strap1, fpbq_lpx/
qed-.
-lemma fpbs_lleq_trans: ∀h,g,G1,G,L1,L,L2,T1,T. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G, L, T⦄ →
- L ≡[T, 0] L2 → ⦃G1, L1, T1⦄ ≥[h, g] ⦃G, L2, T⦄.
+lemma fpbs_lleq_trans: ∀h,o,G1,G,L1,L,L2,T1,T. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L, T⦄ →
+ L ≡[T, 0] L2 → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G, L2, T⦄.
/3 width=5 by fpbs_strap1, fpbq_lleq/ qed-.
-lemma fqus_fpbs_trans: ∀h,g,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G, L, T⦄ ≥[h, g] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ⊐* ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄.
-#h #g #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H @(fqus_ind_dx … H) -G1 -L1 -T1
+lemma fqus_fpbs_trans: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G, L, T⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ⊐* ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H @(fqus_ind_dx … H) -G1 -L1 -T1
/3 width=5 by fpbs_strap2, fpbq_fquq/
qed-.
-lemma cpxs_fpbs_trans: ∀h,g,G1,G2,L1,L2,T1,T,T2. ⦃G1, L1, T⦄ ≥[h, g] ⦃G2, L2, T2⦄ →
- ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] T → ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄.
-#h #g #G1 #G2 #L1 #L2 #T1 #T #T2 #H1 #H @(cpxs_ind_dx … H) -T1
+lemma cpxs_fpbs_trans: ∀h,o,G1,G2,L1,L2,T1,T,T2. ⦃G1, L1, T⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] T → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T #T2 #H1 #H @(cpxs_ind_dx … H) -T1
/3 width=5 by fpbs_strap2, fpbq_cpx/
qed-.
-lemma lpxs_fpbs_trans: ∀h,g,G1,G2,L1,L,L2,T1,T2. ⦃G1, L, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ →
- ⦃G1, L1⦄ ⊢ ➡*[h, g] L → ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄.
-#h #g #G1 #G2 #L1 #L #L2 #T1 #T2 #H1 #H @(lpxs_ind_dx … H) -L1
+lemma lpxs_fpbs_trans: ∀h,o,G1,G2,L1,L,L2,T1,T2. ⦃G1, L, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1⦄ ⊢ ➡*[h, o] L → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L #L2 #T1 #T2 #H1 #H @(lpxs_ind_dx … H) -L1
/3 width=5 by fpbs_strap2, fpbq_lpx/
qed-.
-lemma lleq_fpbs_trans: ∀h,g,G1,G2,L1,L,L2,T1,T2. ⦃G1, L, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ →
- L1 ≡[T1, 0] L → ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄.
+lemma lleq_fpbs_trans: ∀h,o,G1,G2,L1,L,L2,T1,T2. ⦃G1, L, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
+ L1 ≡[T1, 0] L → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
/3 width=5 by fpbs_strap2, fpbq_lleq/ qed-.
-lemma cpxs_fqus_fpbs: ∀h,g,G1,G2,L1,L2,T1,T,T2. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] T →
- ⦃G1, L1, T⦄ ⊐* ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄.
+lemma cpxs_fqus_fpbs: ∀h,o,G1,G2,L1,L2,T1,T,T2. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] T →
+ ⦃G1, L1, T⦄ ⊐* ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
/3 width=5 by fpbs_fqus_trans, cpxs_fpbs/ qed.
-lemma cpxs_fqup_fpbs: ∀h,g,G1,G2,L1,L2,T1,T,T2. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] T →
- ⦃G1, L1, T⦄ ⊐+ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄.
+lemma cpxs_fqup_fpbs: ∀h,o,G1,G2,L1,L2,T1,T,T2. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] T →
+ ⦃G1, L1, T⦄ ⊐+ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
/3 width=5 by fpbs_fqup_trans, cpxs_fpbs/ qed.
-lemma fqus_lpxs_fpbs: ∀h,g,G1,G2,L1,L,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L, T2⦄ →
- ⦃G2, L⦄ ⊢ ➡*[h, g] L2 → ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄.
+lemma fqus_lpxs_fpbs: ∀h,o,G1,G2,L1,L,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L, T2⦄ →
+ ⦃G2, L⦄ ⊢ ➡*[h, o] L2 → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
/3 width=3 by fpbs_lpxs_trans, fqus_fpbs/ qed.
-lemma cpxs_fqus_lpxs_fpbs: ∀h,g,G1,G2,L1,L,L2,T1,T,T2. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] T →
- ⦃G1, L1, T⦄ ⊐* ⦃G2, L, T2⦄ → ⦃G2, L⦄ ⊢ ➡*[h, g] L2 → ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄.
+lemma cpxs_fqus_lpxs_fpbs: ∀h,o,G1,G2,L1,L,L2,T1,T,T2. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] T →
+ ⦃G1, L1, T⦄ ⊐* ⦃G2, L, T2⦄ → ⦃G2, L⦄ ⊢ ➡*[h, o] L2 → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
/3 width=5 by cpxs_fqus_fpbs, fpbs_lpxs_trans/ qed.
-lemma lpxs_lleq_fpbs: ∀h,g,G,L1,L,L2,T. ⦃G, L1⦄ ⊢ ➡*[h, g] L →
- L ≡[T, 0] L2 → ⦃G, L1, T⦄ ≥[h, g] ⦃G, L2, T⦄.
+lemma lpxs_lleq_fpbs: ∀h,o,G,L1,L,L2,T. ⦃G, L1⦄ ⊢ ➡*[h, o] L →
+ L ≡[T, 0] L2 → ⦃G, L1, T⦄ ≥[h, o] ⦃G, L2, T⦄.
/3 width=3 by lpxs_fpbs_trans, lleq_fpbs/ qed.
(* Note: this is used in the closure proof *)
-lemma cpr_lpr_fpbs: ∀h,g,G,L1,L2,T1,T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ⦃G, L1⦄ ⊢ ➡ L2 →
- ⦃G, L1, T1⦄ ≥[h, g] ⦃G, L2, T2⦄.
+lemma cpr_lpr_fpbs: ∀h,o,G,L1,L2,T1,T2. ⦃G, L1⦄ ⊢ T1 ➡ T2 → ⦃G, L1⦄ ⊢ ➡ L2 →
+ ⦃G, L1, T1⦄ ≥[h, o] ⦃G, L2, T2⦄.
/4 width=5 by fpbs_strap1, fpbq_fpbs, lpr_fpbq, cpr_fpbq/
qed.
(* Properties on atomic arity assignment for terms **************************)
-lemma fpbs_aaa_conf: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ →
+lemma fpbs_aaa_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
∀A1. ⦃G1, L1⦄ ⊢ T1 ⁝ A1 → ∃A2. ⦃G2, L2⦄ ⊢ T2 ⁝ A2.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H @(fpbs_ind … H) -G2 -L2 -T2 /2 width=2 by ex_intro/
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fpbs_ind … H) -G2 -L2 -T2 /2 width=2 by ex_intro/
#G #G2 #L #L2 #T #T2 #_ #H2 #IH1 #A #HA elim (IH1 … HA) -IH1 -A
/2 width=8 by fpbq_aaa_conf/
qed-.
(* Note: alternative definition of fpbs *)
definition fpbsa: ∀h. sd h → tri_relation genv lenv term ≝
- λh,g,G1,L1,T1,G2,L2,T2.
- ∃∃L0,L,T. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] T &
+ λh,o,G1,L1,T1,G2,L2,T2.
+ ∃∃L0,L,T. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] T &
⦃G1, L1, T⦄ ⊐* ⦃G2, L0, T2⦄ &
- ⦃G2, L0⦄ ⊢ ➡*[h, g] L & L ≡[T2, 0] L2.
+ ⦃G2, L0⦄ ⊢ ➡*[h, o] L & L ≡[T2, 0] L2.
interpretation "'big tree' parallel computation (closure) alternative"
- 'BTPRedStarAlt h g G1 L1 T1 G2 L2 T2 = (fpbsa h g G1 L1 T1 G2 L2 T2).
+ 'BTPRedStarAlt h o G1 L1 T1 G2 L2 T2 = (fpbsa h o G1 L1 T1 G2 L2 T2).
(* Basic properties *********************************************************)
-lemma fpb_fpbsa_trans: ∀h,g,G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ≽[h, g] ⦃G, L, T⦄ →
- ∀G2,L2,T2. ⦃G, L, T⦄ ≥≥[h, g] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥≥[h, g] ⦃G2, L2, T2⦄.
-#h #g #G1 #G #L1 #L #T1 #T * -G -L -T [ #G #L #T #HG1 | #T #HT1 | #L #HL1 | #L #HL1 ]
+lemma fpb_fpbsa_trans: ∀h,o,G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G, L, T⦄ →
+ ∀G2,L2,T2. ⦃G, L, T⦄ ≥≥[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥≥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G #L1 #L #T1 #T * -G -L -T [ #G #L #T #HG1 | #T #HT1 | #L #HL1 | #L #HL1 ]
#G2 #L2 #T2 * #L00 #L0 #T0 #HT0 #HG2 #HL00 #HL02
[ elim (fquq_cpxs_trans … HT0 … HG1) -T
/3 width=7 by fqus_strap2, ex4_3_intro/
(* Main properties **********************************************************)
-theorem fpbs_fpbsa: ∀h,g,G1,G2,L1,L2,T1,T2.
- ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥≥[h, g] ⦃G2, L2, T2⦄.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H @(fpbs_ind_dx … H) -G1 -L1 -T1
+theorem fpbs_fpbsa: ∀h,o,G1,G2,L1,L2,T1,T2.
+ ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥≥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fpbs_ind_dx … H) -G1 -L1 -T1
/2 width=7 by fpb_fpbsa_trans, ex4_3_intro/
qed.
(* Main inversion lemmas ****************************************************)
-theorem fpbsa_inv_fpbs: ∀h,g,G1,G2,L1,L2,T1,T2.
- ⦃G1, L1, T1⦄ ≥≥[h, g] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 *
+theorem fpbsa_inv_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2.
+ ⦃G1, L1, T1⦄ ≥≥[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 *
/3 width=5 by cpxs_fqus_lpxs_fpbs, fpbs_strap1, fpbq_lleq/
qed-.
(* Advanced properties ******************************************************)
-lemma fpbs_intro_alt: ∀h,g,G1,G2,L1,L0,L,L2,T1,T,T2.
- ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] T → ⦃G1, L1, T⦄ ⊐* ⦃G2, L0, T2⦄ →
- ⦃G2, L0⦄ ⊢ ➡*[h, g] L → L ≡[T2, 0] L2 → ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ .
+lemma fpbs_intro_alt: ∀h,o,G1,G2,L1,L0,L,L2,T1,T,T2.
+ ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] T → ⦃G1, L1, T⦄ ⊐* ⦃G2, L0, T2⦄ →
+ ⦃G2, L0⦄ ⊢ ➡*[h, o] L → L ≡[T2, 0] L2 → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ .
/3 width=7 by fpbsa_inv_fpbs, ex4_3_intro/ qed.
(* Advanced inversion lemmas *************************************************)
-lemma fpbs_inv_alt: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ →
- ∃∃L0,L,T. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] T &
+lemma fpbs_inv_alt: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
+ ∃∃L0,L,T. ⦃G1, L1⦄ ⊢ T1 ➡*[h, o] T &
⦃G1, L1, T⦄ ⊐* ⦃G2, L0, T2⦄ &
- ⦃G2, L0⦄ ⊢ ➡*[h, g] L & L ≡[T2, 0] L2.
+ ⦃G2, L0⦄ ⊢ ➡*[h, o] L & L ≡[T2, 0] L2.
/2 width=1 by fpbs_fpbsa/ qed-.
(* Properties on extended context-sensitive parallel computation for terms **)
-lemma fpbs_cpx_trans_neq: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ →
- ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, g] U2 → (T2 = U2 → ⊥) →
- ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, g] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ≥[h, g] ⦃G2, L2, U2⦄.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H #U2 #HTU2 #HnTU2 elim (fpbs_inv_alt … H) -H
+lemma fpbs_cpx_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
+ ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, o] U2 → (T2 = U2 → ⊥) →
+ ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ≥[h, o] ⦃G2, L2, U2⦄.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #U2 #HTU2 #HnTU2 elim (fpbs_inv_alt … H) -H
#L00 #L0 #T0 #HT10 #H10 #HL00 #HL02 lapply (lleq_cpx_trans … HTU2 … HL02) -HTU2
#HTU2 lapply (cpx_lleq_conf_sn … HTU2 … HL02) -HL02
#HL02 lapply (lpxs_cpx_trans … HTU2 … HL00) -HTU2
(* Properties on "rst" proper parallel reduction on closures ****************)
-lemma fpb_fpbs: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄.
+lemma fpb_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.
/3 width=1 by fpbq_fpbs, fpb_fpbq/ qed.
(* Main properties **********************************************************)
-theorem fpbs_trans: ∀h,g. tri_transitive … (fpbs h g).
+theorem fpbs_trans: ∀h,o. tri_transitive … (fpbs h o).
/2 width=5 by tri_TC_transitive/ qed-.
(* Advanced properties ******************************************************)
-lemma lstas_fpbs: ∀h,g,G,L,T1,T2,d2. ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 →
- ∀d1. d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 ▪[h, g] d1 → ⦃G, L, T1⦄ ≥[h, g] ⦃G, L, T2⦄.
+lemma lstas_fpbs: ∀h,o,G,L,T1,T2,d2. ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 →
+ ∀d1. d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 ▪[h, o] d1 → ⦃G, L, T1⦄ ≥[h, o] ⦃G, L, T2⦄.
/3 width=5 by cpxs_fpbs, lstas_cpxs/ qed.
-lemma sta_fpbs: ∀h,g,G,L,T,U,d.
- ⦃G, L⦄ ⊢ T ▪[h, g] d+1 → ⦃G, L⦄ ⊢ T •*[h, 1] U →
- ⦃G, L, T⦄ ≥[h, g] ⦃G, L, U⦄.
+lemma sta_fpbs: ∀h,o,G,L,T,U,d.
+ ⦃G, L⦄ ⊢ T ▪[h, o] d+1 → ⦃G, L⦄ ⊢ T •*[h, 1] U →
+ ⦃G, L, T⦄ ≥[h, o] ⦃G, L, U⦄.
/2 width=5 by lstas_fpbs/ qed.
(* Note: this is used in the closure proof *)
-lemma cpr_lpr_sta_fpbs: ∀h,g,G,L1,L2,T1,T2,U2,d2.
+lemma cpr_lpr_sta_fpbs: ∀h,o,G,L1,L2,T1,T2,U2,d2.
⦃G, L1⦄ ⊢ T1 ➡ T2 → ⦃G, L1⦄ ⊢ ➡ L2 →
- ⦃G, L2⦄ ⊢ T2 ▪[h, g] d2+1 → ⦃G, L2⦄ ⊢ T2 •*[h, 1] U2 →
- ⦃G, L1, T1⦄ ≥[h, g] ⦃G, L2, U2⦄.
+ ⦃G, L2⦄ ⊢ T2 ▪[h, o] d2+1 → ⦃G, L2⦄ ⊢ T2 •*[h, 1] U2 →
+ ⦃G, L1, T1⦄ ≥[h, o] ⦃G, L2, U2⦄.
/4 width=5 by fpbs_strap1, cpr_lpr_fpbs, sta_cpx, fpbq_cpx/ qed.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/suptermplus_6.ma".
-include "basic_2/substitution/fqu.ma".
-
-(* PLUS-ITERATED SUPCLOSURE *************************************************)
-
-definition fqup: tri_relation genv lenv term ≝ tri_TC … fqu.
-
-interpretation "plus-iterated structural successor (closure)"
- 'SupTermPlus G1 L1 T1 G2 L2 T2 = (fqup G1 L1 T1 G2 L2 T2).
-
-(* Basic properties *********************************************************)
-
-lemma fqu_fqup: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄.
-/2 width=1 by tri_inj/ qed.
-
-lemma fqup_strap1: ∀G1,G,G2,L1,L,L2,T1,T,T2.
- ⦃G1, L1, T1⦄ ⊐+ ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐ ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄.
-/2 width=5 by tri_step/ qed.
-
-lemma fqup_strap2: ∀G1,G,G2,L1,L,L2,T1,T,T2.
- ⦃G1, L1, T1⦄ ⊐ ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄.
-/2 width=5 by tri_TC_strap/ qed.
-
-lemma fqup_drop: ∀G1,G2,L1,K1,K2,T1,T2,U1,k. ⬇[k] L1 ≡ K1 → ⬆[0, k] T1 ≡ U1 →
- ⦃G1, K1, T1⦄ ⊐+ ⦃G2, K2, T2⦄ → ⦃G1, L1, U1⦄ ⊐+ ⦃G2, K2, T2⦄.
-#G1 #G2 #L1 #K1 #K2 #T1 #T2 #U1 #k #HLK1 #HTU1 #HT12 elim (eq_or_gt … k) #H destruct
-[ >(drop_inv_O2 … HLK1) -L1 <(lift_inv_O2 … HTU1) -U1 //
-| /3 width=5 by fqup_strap2, fqu_drop_lt/
-]
-qed-.
-
-lemma fqup_lref: ∀I,G,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → ⦃G, L, #i⦄ ⊐+ ⦃G, K, V⦄.
-/3 width=6 by fqu_lref_O, fqu_fqup, lift_lref_ge, fqup_drop/ qed.
-
-lemma fqup_pair_sn: ∀I,G,L,V,T. ⦃G, L, ②{I}V.T⦄ ⊐+ ⦃G, L, V⦄.
-/2 width=1 by fqu_pair_sn, fqu_fqup/ qed.
-
-lemma fqup_bind_dx: ∀a,I,G,L,V,T. ⦃G, L, ⓑ{a,I}V.T⦄ ⊐+ ⦃G, L.ⓑ{I}V, T⦄.
-/2 width=1 by fqu_bind_dx, fqu_fqup/ qed.
-
-lemma fqup_flat_dx: ∀I,G,L,V,T. ⦃G, L, ⓕ{I}V.T⦄ ⊐+ ⦃G, L, T⦄.
-/2 width=1 by fqu_flat_dx, fqu_fqup/ qed.
-
-lemma fqup_flat_dx_pair_sn: ∀I1,I2,G,L,V1,V2,T. ⦃G, L, ⓕ{I1}V1.②{I2}V2.T⦄ ⊐+ ⦃G, L, V2⦄.
-/2 width=5 by fqu_pair_sn, fqup_strap1/ qed.
-
-lemma fqup_bind_dx_flat_dx: ∀a,G,I1,I2,L,V1,V2,T. ⦃G, L, ⓑ{a,I1}V1.ⓕ{I2}V2.T⦄ ⊐+ ⦃G, L.ⓑ{I1}V1, T⦄.
-/2 width=5 by fqu_flat_dx, fqup_strap1/ qed.
-
-lemma fqup_flat_dx_bind_dx: ∀a,I1,I2,G,L,V1,V2,T. ⦃G, L, ⓕ{I1}V1.ⓑ{a,I2}V2.T⦄ ⊐+ ⦃G, L.ⓑ{I2}V2, T⦄.
-/2 width=5 by fqu_bind_dx, fqup_strap1/ qed.
-
-(* Basic eliminators ********************************************************)
-
-lemma fqup_ind: ∀G1,L1,T1. ∀R:relation3 ….
- (∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → R G2 L2 T2) →
- (∀G,G2,L,L2,T,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐ ⦃G2, L2, T2⦄ → R G L T → R G2 L2 T2) →
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → R G2 L2 T2.
-#G1 #L1 #T1 #R #IH1 #IH2 #G2 #L2 #T2 #H
-@(tri_TC_ind … IH1 IH2 G2 L2 T2 H)
-qed-.
-
-lemma fqup_ind_dx: ∀G2,L2,T2. ∀R:relation3 ….
- (∀G1,L1,T1. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → R G1 L1 T1) →
- (∀G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ⊐ ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐+ ⦃G2, L2, T2⦄ → R G L T → R G1 L1 T1) →
- ∀G1,L1,T1. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → R G1 L1 T1.
-#G2 #L2 #T2 #R #IH1 #IH2 #G1 #L1 #T1 #H
-@(tri_TC_ind_dx … IH1 IH2 G1 L1 T1 H)
-qed-.
-
-(* Basic forward lemmas *****************************************************)
-
-lemma fqup_fwd_fw: ∀G1,G2,L1,L2,T1,T2.
- ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → ♯{G2, L2, T2} < ♯{G1, L1, T1}.
-#G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
-/3 width=3 by fqu_fwd_fw, transitive_lt/
-qed-.
-
-(* Advanced eliminators *****************************************************)
-
-lemma fqup_wf_ind: ∀R:relation3 …. (
- ∀G1,L1,T1. (∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → R G2 L2 T2) →
- R G1 L1 T1
- ) → ∀G1,L1,T1. R G1 L1 T1.
-#R #HR @(f3_ind … fw) #x #IHx #G1 #L1 #T1 #H destruct /4 width=1 by fqup_fwd_fw/
-qed-.
-
-lemma fqup_wf_ind_eq: ∀R:relation3 …. (
- ∀G1,L1,T1. (∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → R G2 L2 T2) →
- ∀G2,L2,T2. G1 = G2 → L1 = L2 → T1 = T2 → R G2 L2 T2
- ) → ∀G1,L1,T1. R G1 L1 T1.
-#R #HR @(f3_ind … fw) #x #IHx #G1 #L1 #T1 #H destruct /4 width=7 by fqup_fwd_fw/
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/multiple/fqup.ma".
-
-(* PLUS-ITERATED SUPCLOSURE *************************************************)
-
-(* Main properties **********************************************************)
-
-theorem fqup_trans: tri_transitive … fqup.
-/2 width=5 by tri_TC_transitive/ qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/suptermstar_6.ma".
-include "basic_2/substitution/fquq.ma".
-include "basic_2/multiple/fqup.ma".
-
-(* STAR-ITERATED SUPCLOSURE *************************************************)
-
-definition fqus: tri_relation genv lenv term ≝ tri_TC … fquq.
-
-interpretation "star-iterated structural successor (closure)"
- 'SupTermStar G1 L1 T1 G2 L2 T2 = (fqus G1 L1 T1 G2 L2 T2).
-
-(* Basic eliminators ********************************************************)
-
-lemma fqus_ind: ∀G1,L1,T1. ∀R:relation3 …. R G1 L1 T1 →
- (∀G,G2,L,L2,T,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐⸮ ⦃G2, L2, T2⦄ → R G L T → R G2 L2 T2) →
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ → R G2 L2 T2.
-#G1 #L1 #T1 #R #IH1 #IH2 #G2 #L2 #T2 #H
-@(tri_TC_star_ind … IH1 IH2 G2 L2 T2 H) //
-qed-.
-
-lemma fqus_ind_dx: ∀G2,L2,T2. ∀R:relation3 …. R G2 L2 T2 →
- (∀G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐* ⦃G2, L2, T2⦄ → R G L T → R G1 L1 T1) →
- ∀G1,L1,T1. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ → R G1 L1 T1.
-#G2 #L2 #T2 #R #IH1 #IH2 #G1 #L1 #T1 #H
-@(tri_TC_star_ind_dx … IH1 IH2 G1 L1 T1 H) //
-qed-.
-
-(* Basic properties *********************************************************)
-
-lemma fqus_refl: tri_reflexive … fqus.
-/2 width=1 by tri_inj/ qed.
-
-lemma fquq_fqus: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄.
-/2 width=1 by tri_inj/ qed.
-
-lemma fqus_strap1: ∀G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄.
-/2 width=5 by tri_step/ qed-.
-
-lemma fqus_strap2: ∀G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐* ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄.
-/2 width=5 by tri_TC_strap/ qed-.
-
-lemma fqus_drop: ∀G1,G2,K1,K2,T1,T2. ⦃G1, K1, T1⦄ ⊐* ⦃G2, K2, T2⦄ →
- ∀L1,U1,k. ⬇[k] L1 ≡ K1 → ⬆[0, k] T1 ≡ U1 →
- ⦃G1, L1, U1⦄ ⊐* ⦃G2, K2, T2⦄.
-#G1 #G2 #K1 #K2 #T1 #T2 #H @(fqus_ind … H) -G2 -K2 -T2
-/3 width=5 by fqus_strap1, fquq_fqus, fquq_drop/
-qed-.
-
-lemma fqup_fqus: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄.
-#G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
-/3 width=5 by fqus_strap1, fquq_fqus, fqu_fquq/
-qed.
-
-(* Basic forward lemmas *****************************************************)
-
-lemma fqus_fwd_fw: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ → ♯{G2, L2, T2} ≤ ♯{G1, L1, T1}.
-#G1 #G2 #L1 #L2 #T1 #T2 #H @(fqus_ind … H) -L2 -T2
-/3 width=3 by fquq_fwd_fw, transitive_le/
-qed-.
-
-(* Basic inversion lemmas ***************************************************)
-
-lemma fqup_inv_step_sn: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ∃∃G,L,T. ⦃G1, L1, T1⦄ ⊐ ⦃G, L, T⦄ & ⦃G, L, T⦄ ⊐* ⦃G2, L2, T2⦄.
-#G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind_dx … H) -G1 -L1 -T1 /2 width=5 by ex2_3_intro/
-#G1 #G #L1 #L #T1 #T #H1 #_ * /4 width=9 by fqus_strap2, fqu_fquq, ex2_3_intro/
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/substitution/fquq_alt.ma".
-include "basic_2/multiple/fqus.ma".
-
-(* STAR-ITERATED SUPCLOSURE *************************************************)
-
-(* Advanced inversion lemmas ************************************************)
-
-lemma fqus_inv_gen: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ ∨ (∧∧ G1 = G2 & L1 = L2 & T1 = T2).
-#G1 #G2 #L1 #L2 #T1 #T2 #H @(fqus_ind … H) -G2 -L2 -T2 //
-#G #G2 #L #L2 #T #T2 #_ #H2 * elim (fquq_inv_gen … H2) -H2
-[ /3 width=5 by fqup_strap1, or_introl/
-| * #HG #HL #HT destruct /2 width=1 by or_introl/
-| #H2 * #HG #HL #HT destruct /3 width=1 by fqu_fqup, or_introl/
-| * #H1G #H1L #H1T * #H2G #H2L #H2T destruct /2 width=1 by or_intror/
-]
-qed-.
-
-(* Advanced properties ******************************************************)
-
-lemma fqus_strap1_fqu: ∀G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐ ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄.
-#G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H2 elim (fqus_inv_gen … H1) -H1
-[ /2 width=5 by fqup_strap1/
-| * #HG #HL #HT destruct /2 width=1 by fqu_fqup/
-]
-qed-.
-
-lemma fqus_strap2_fqu: ∀G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐* ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄.
-#G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H2 elim (fqus_inv_gen … H2) -H2
-[ /2 width=5 by fqup_strap2/
-| * #HG #HL #HT destruct /2 width=1 by fqu_fqup/
-]
-qed-.
-
-lemma fqus_fqup_trans: ∀G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄.
-#G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H2 @(fqup_ind … H2) -H2 -G2 -L2 -T2
-/2 width=5 by fqus_strap1_fqu, fqup_strap1/
-qed-.
-
-lemma fqup_fqus_trans: ∀G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G, L, T⦄ →
- ⦃G, L, T⦄ ⊐* ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄.
-#G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 @(fqup_ind_dx … H1) -H1 -G1 -L1 -T1
-/3 width=5 by fqus_strap2_fqu, fqup_strap2/
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/multiple/fqus.ma".
-
-(* STAR-ITERATED SUPCLOSURE *************************************************)
-
-(* Main properties **********************************************************)
-
-theorem fqus_trans: tri_transitive … fqus.
-/2 width=5 by tri_TC_transitive/ qed-.
(* "QRST" STRONGLY NORMALIZING CLOSURES *************************************)
-inductive fsb (h) (g): relation3 genv lenv term ≝
+inductive fsb (h) (o): relation3 genv lenv term ≝
| fsb_intro: ∀G1,L1,T1. (
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄ → fsb h g G2 L2 T2
- ) → fsb h g G1 L1 T1
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → fsb h o G2 L2 T2
+ ) → fsb h o G1 L1 T1
.
interpretation
"'qrst' strong normalization (closure)"
- 'BTSN h g G L T = (fsb h g G L T).
+ 'BTSN h o G L T = (fsb h o G L T).
(* Basic eliminators ********************************************************)
-lemma fsb_ind_alt: ∀h,g. ∀R: relation3 …. (
- ∀G1,L1,T1. ⦥[h,g] ⦃G1, L1, T1⦄ → (
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄ → R G2 L2 T2
+lemma fsb_ind_alt: ∀h,o. ∀R: relation3 …. (
+ ∀G1,L1,T1. ⦥[h,o] ⦃G1, L1, T1⦄ → (
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2
) → R G1 L1 T1
) →
- ∀G,L,T. ⦥[h, g] ⦃G, L, T⦄ → R G L T.
-#h #g #R #IH #G #L #T #H elim H -G -L -T
+ ∀G,L,T. ⦥[h, o] ⦃G, L, T⦄ → R G L T.
+#h #o #R #IH #G #L #T #H elim H -G -L -T
/4 width=1 by fsb_intro/
qed-.
(* Basic inversion lemmas ***************************************************)
-lemma fsb_inv_csx: ∀h,g,G,L,T. ⦥[h, g] ⦃G, L, T⦄ → ⦃G, L⦄ ⊢ ⬊*[h, g] T.
-#h #g #G #L #T #H elim H -G -L -T /5 width=1 by csx_intro, fpb_cpx/
+lemma fsb_inv_csx: ∀h,o,G,L,T. ⦥[h, o] ⦃G, L, T⦄ → ⦃G, L⦄ ⊢ ⬊*[h, o] T.
+#h #o #G #L #T #H elim H -G -L -T /5 width=1 by csx_intro, fpb_cpx/
qed-.
(* Main properties **********************************************************)
(* Note: this is the "big tree" theorem ("RST" version) *)
-theorem aaa_fsb: ∀h,g,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ⦥[h, g] ⦃G, L, T⦄.
+theorem aaa_fsb: ∀h,o,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ⦥[h, o] ⦃G, L, T⦄.
/3 width=2 by aaa_csx, csx_fsb/ qed.
(* Note: this is the "big tree" theorem ("QRST" version) *)
-theorem aaa_fsba: ∀h,g,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ⦥⦥[h, g] ⦃G, L, T⦄.
+theorem aaa_fsba: ∀h,o,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ⦥⦥[h, o] ⦃G, L, T⦄.
/3 width=2 by fsb_fsba, aaa_fsb/ qed.
(* Advanced eliminators on atomica arity assignment for terms ***************)
-fact aaa_ind_fpb_aux: ∀h,g. ∀R:relation3 genv lenv term.
+fact aaa_ind_fpb_aux: ∀h,o. ∀R:relation3 genv lenv term.
(∀G1,L1,T1,A. ⦃G1, L1⦄ ⊢ T1 ⁝ A →
- (∀G2,L2,T2. ⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
+ (∀G2,L2,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
R G1 L1 T1
) →
- ∀G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → ∀A. ⦃G, L⦄ ⊢ T ⁝ A → R G L T.
-#h #g #R #IH #G #L #T #H @(csx_ind_fpb … H) -G -L -T
+ ∀G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ∀A. ⦃G, L⦄ ⊢ T ⁝ A → R G L T.
+#h #o #R #IH #G #L #T #H @(csx_ind_fpb … H) -G -L -T
#G1 #L1 #T1 #H1 #IH1 #A1 #HTA1 @IH -IH //
-#G2 #L2 #T2 #H12 elim (fpbs_aaa_conf h g … G2 … L2 … T2 … HTA1) -A1
+#G2 #L2 #T2 #H12 elim (fpbs_aaa_conf h o … G2 … L2 … T2 … HTA1) -A1
/2 width=2 by fpb_fpbs/
qed-.
-lemma aaa_ind_fpb: ∀h,g. ∀R:relation3 genv lenv term.
+lemma aaa_ind_fpb: ∀h,o. ∀R:relation3 genv lenv term.
(∀G1,L1,T1,A. ⦃G1, L1⦄ ⊢ T1 ⁝ A →
- (∀G2,L2,T2. ⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
+ (∀G2,L2,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
R G1 L1 T1
) →
∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → R G L T.
/4 width=4 by aaa_ind_fpb_aux, aaa_csx/ qed-.
-fact aaa_ind_fpbg_aux: ∀h,g. ∀R:relation3 genv lenv term.
+fact aaa_ind_fpbg_aux: ∀h,o. ∀R:relation3 genv lenv term.
(∀G1,L1,T1,A. ⦃G1, L1⦄ ⊢ T1 ⁝ A →
- (∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
+ (∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
R G1 L1 T1
) →
- ∀G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → ∀A. ⦃G, L⦄ ⊢ T ⁝ A → R G L T.
-#h #g #R #IH #G #L #T #H @(csx_ind_fpbg … H) -G -L -T
+ ∀G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ∀A. ⦃G, L⦄ ⊢ T ⁝ A → R G L T.
+#h #o #R #IH #G #L #T #H @(csx_ind_fpbg … H) -G -L -T
#G1 #L1 #T1 #H1 #IH1 #A1 #HTA1 @IH -IH //
-#G2 #L2 #T2 #H12 elim (fpbs_aaa_conf h g … G2 … L2 … T2 … HTA1) -A1
+#G2 #L2 #T2 #H12 elim (fpbs_aaa_conf h o … G2 … L2 … T2 … HTA1) -A1
/2 width=2 by fpbg_fwd_fpbs/
qed-.
-lemma aaa_ind_fpbg: ∀h,g. ∀R:relation3 genv lenv term.
+lemma aaa_ind_fpbg: ∀h,o. ∀R:relation3 genv lenv term.
(∀G1,L1,T1,A. ⦃G1, L1⦄ ⊢ T1 ⁝ A →
- (∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
+ (∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
R G1 L1 T1
) →
∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → R G L T.
(* "QRST" STRONGLY NORMALIZING CLOSURES *************************************)
(* Note: alternative definition of fsb *)
-inductive fsba (h) (g): relation3 genv lenv term ≝
+inductive fsba (h) (o): relation3 genv lenv term ≝
| fsba_intro: ∀G1,L1,T1. (
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄ → fsba h g G2 L2 T2
- ) → fsba h g G1 L1 T1.
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄ → fsba h o G2 L2 T2
+ ) → fsba h o G1 L1 T1.
interpretation
"'big tree' strong normalization (closure) alternative"
- 'BTSNAlt h g G L T = (fsba h g G L T).
+ 'BTSNAlt h o G L T = (fsba h o G L T).
(* Basic eliminators ********************************************************)
-lemma fsba_ind_alt: ∀h,g. ∀R: relation3 …. (
- ∀G1,L1,T1. ⦥⦥[h,g] ⦃G1, L1, T1⦄ → (
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄ → R G2 L2 T2
+lemma fsba_ind_alt: ∀h,o. ∀R: relation3 …. (
+ ∀G1,L1,T1. ⦥⦥[h,o] ⦃G1, L1, T1⦄ → (
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2
) → R G1 L1 T1
) →
- ∀G,L,T. ⦥⦥[h, g] ⦃G, L, T⦄ → R G L T.
-#h #g #R #IH #G #L #T #H elim H -G -L -T
+ ∀G,L,T. ⦥⦥[h, o] ⦃G, L, T⦄ → R G L T.
+#h #o #R #IH #G #L #T #H elim H -G -L -T
/4 width=1 by fsba_intro/
qed-.
(* Basic properties *********************************************************)
-lemma fsba_fpbs_trans: ∀h,g,G1,L1,T1. ⦥⦥[h, g] ⦃G1, L1, T1⦄ →
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ → ⦥⦥[h, g] ⦃G2, L2, T2⦄.
-#h #g #G1 #L1 #T1 #H @(fsba_ind_alt … H) -G1 -L1 -T1
+lemma fsba_fpbs_trans: ∀h,o,G1,L1,T1. ⦥⦥[h, o] ⦃G1, L1, T1⦄ →
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → ⦥⦥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #L1 #T1 #H @(fsba_ind_alt … H) -G1 -L1 -T1
/4 width=5 by fsba_intro, fpbs_fpbg_trans/
qed-.
(* Main properties **********************************************************)
-theorem fsb_fsba: ∀h,g,G,L,T. ⦥[h, g] ⦃G, L, T⦄ → ⦥⦥[h, g] ⦃G, L, T⦄.
-#h #g #G #L #T #H @(fsb_ind_alt … H) -G -L -T
+theorem fsb_fsba: ∀h,o,G,L,T. ⦥[h, o] ⦃G, L, T⦄ → ⦥⦥[h, o] ⦃G, L, T⦄.
+#h #o #G #L #T #H @(fsb_ind_alt … H) -G -L -T
#G1 #L1 #T1 #_ #IH @fsba_intro
#G2 #L2 #T2 * /3 width=5 by fsba_fpbs_trans/
qed.
(* Main inversion lemmas ****************************************************)
-theorem fsba_inv_fsb: ∀h,g,G,L,T. ⦥⦥[h, g] ⦃G, L, T⦄ → ⦥[h, g] ⦃G, L, T⦄.
-#h #g #G #L #T #H @(fsba_ind_alt … H) -G -L -T
+theorem fsba_inv_fsb: ∀h,o,G,L,T. ⦥⦥[h, o] ⦃G, L, T⦄ → ⦥[h, o] ⦃G, L, T⦄.
+#h #o #G #L #T #H @(fsba_ind_alt … H) -G -L -T
/4 width=1 by fsb_intro, fpb_fpbg/
qed-.
(* Advanced properties ******************************************************)
-lemma fsb_fpbs_trans: ∀h,g,G1,L1,T1. ⦥[h, g] ⦃G1, L1, T1⦄ →
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ → ⦥[h, g] ⦃G2, L2, T2⦄.
+lemma fsb_fpbs_trans: ∀h,o,G1,L1,T1. ⦥[h, o] ⦃G1, L1, T1⦄ →
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → ⦥[h, o] ⦃G2, L2, T2⦄.
/4 width=5 by fsba_inv_fsb, fsb_fsba, fsba_fpbs_trans/ qed-.
(* Advanced eliminators *****************************************************)
-lemma fsb_ind_fpbg: ∀h,g. ∀R:relation3 genv lenv term.
- (∀G1,L1,T1. ⦥[h, g] ⦃G1, L1, T1⦄ →
- (∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
+lemma fsb_ind_fpbg: ∀h,o. ∀R:relation3 genv lenv term.
+ (∀G1,L1,T1. ⦥[h, o] ⦃G1, L1, T1⦄ →
+ (∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
R G1 L1 T1
) →
- ∀G1,L1,T1. ⦥[h, g] ⦃G1, L1, T1⦄ → R G1 L1 T1.
-#h #g #R #IH #G1 #L1 #T1 #H @(fsba_ind_alt h g … G1 L1 T1)
+ ∀G1,L1,T1. ⦥[h, o] ⦃G1, L1, T1⦄ → R G1 L1 T1.
+#h #o #R #IH #G1 #L1 #T1 #H @(fsba_ind_alt h o … G1 L1 T1)
/3 width=1 by fsba_inv_fsb, fsb_fsba/
qed-.
(* Advanced propreties on context-sensitive extended normalizing terms ******)
-lemma csx_fsb_fpbs: ∀h,g,G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, g] T1 →
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ → ⦥[h, g] ⦃G2, L2, T2⦄.
-#h #g #G1 #L1 #T1 #H @(csx_ind … H) -T1
+lemma csx_fsb_fpbs: ∀h,o,G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 →
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → ⦥[h, o] ⦃G2, L2, T2⦄.
+#h #o #G1 #L1 #T1 #H @(csx_ind … H) -T1
#T1 #HT1 #IHc #G2 #L2 #T2 @(fqup_wf_ind … G2 L2 T2) -G2 -L2 -T2
#G0 #L0 #T0 #IHu #H10 lapply (csx_fpbs_conf … H10) // -HT1
#HT0 generalize in match IHu; -IHu generalize in match H10; -H10
]
qed.
-lemma csx_fsb: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → ⦥[h, g] ⦃G, L, T⦄.
+lemma csx_fsb: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ⦥[h, o] ⦃G, L, T⦄.
/2 width=5 by csx_fsb_fpbs/ qed.
(* Advanced eliminators *****************************************************)
-lemma csx_ind_fpb: ∀h,g. ∀R:relation3 genv lenv term.
- (∀G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, g] T1 →
- (∀G2,L2,T2. ⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
+lemma csx_ind_fpb: ∀h,o. ∀R:relation3 genv lenv term.
+ (∀G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 →
+ (∀G2,L2,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
R G1 L1 T1
) →
- ∀G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → R G L T.
+ ∀G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → R G L T.
/4 width=4 by fsb_inv_csx, csx_fsb, fsb_ind_alt/ qed-.
-lemma csx_ind_fpbg: ∀h,g. ∀R:relation3 genv lenv term.
- (∀G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, g] T1 →
- (∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
+lemma csx_ind_fpbg: ∀h,o. ∀R:relation3 genv lenv term.
+ (∀G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, o] T1 →
+ (∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, o] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
R G1 L1 T1
) →
- ∀G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → R G L T.
+ ∀G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → R G L T.
/4 width=4 by fsb_inv_csx, csx_fsb, fsb_ind_fpbg/ qed-.
∀G. d_liftable1 (nf RR RS G) (Ⓕ).
definition CP1 ≝ λRR:relation4 genv lenv term term. λRS:relation term.
- ∀G,L. ∃k. NF … (RR G L) RS (⋆k).
+ ∀G,L. ∃s. NF … (RR G L) RS (⋆s).
definition CP2 ≝ λRP:candidate. ∀G. d_liftable1 (RP G) (Ⓕ).
definition CP3 ≝ λRP:candidate.
- ∀G,L,T,k. RP G L (ⓐ⋆k.T) → RP G L T.
+ ∀G,L,T,s. RP G L (ⓐ⋆s.T) → RP G L T.
(* requirements for generic computation properties *)
record gcp (RR:relation4 genv lenv term term) (RS:relation term) (RP:candidate) : Prop ≝
∀T0. ⬆*[cs] T ≡ T0 → ∀L2. G ⊢ L2 ⫃[RP] L0 →
⦃G, L2, T0⦄ ϵ[RP] 〚A〛.
#RR #RS #RP #H1RP #H2RP #G #L1 #T #A #H elim H -G -L1 -T -A
-[ #G #L #k #L0 #cs #HL0 #X #H #L2 #HL20
+[ #G #L #s #L0 #cs #HL0 #X #H #L2 #HL20
>(lifts_inv_sort1 … H) -H
lapply (acr_gcr … H1RP H2RP (⓪)) #HAtom
- lapply (s4 … HAtom G L2 (◊)) /2 width=1 by/
+ lapply (c4 … HAtom G L2 (◊)) /2 width=1 by/
| #I #G #L1 #K1 #V1 #B #i #HLK1 #HKV1B #IHB #L0 #cs #HL01 #X #H #L2 #HL20
lapply (acr_gcr … H1RP H2RP B) #HB
elim (lifts_inv_lref1 … H) -H #i1 #Hi1 #H destruct
[ #K2 #HK20 #H destruct
elim (lift_total V0 0 (i0 +1)) #V #HV0
elim (lifts_lift_trans … Hi0 … Hcs0 … HV10 … HV0) -HV10 #V2 #HV12 #HV2
- lapply (s5 … HB ? G ? ? (◊) … HV0 HLK2) /3 width=7 by drops_cons, lifts_cons/ (* Note: uses IHB HL20 V2 HV0 *)
+ lapply (c5 … HB ? G ? ? (◊) … HV0 HLK2) /3 width=7 by drops_cons, lifts_cons/ (* Note: uses IHB HL20 V2 HV0 *)
| -HLK1 -IHB -HL01 -HL20 -HK1b -Hi0 -Hcs0
#K2 #V2 #A2 #HKV2A #H1KV0A #H2KV0A #_ #H1 #H2 destruct
lapply (drop_fwd_drop2 … HLK2) #HLK2b
lapply (aaa_mono … H2KV0A … HKV0B) #H destruct -H2KV0A -HKV0B
elim (lift_total V0 0 (i0 +1)) #V3 #HV03
elim (lift_total V2 0 (i0 +1)) #V #HV2
- lapply (s5 … HB ? G ? ? (◊) … (ⓝV3.V) … HLK2) /2 width=1 by lift_flat/
- lapply (s7 … HB G L2 (◊)) /3 width=7 by gcr_lift/
+ lapply (c5 … HB ? G ? ? (◊) … (ⓝV3.V) … HLK2) /2 width=1 by lift_flat/
+ lapply (c7 … HB G L2 (◊)) /3 width=7 by gcr_lift/
]
| #a #G #L #V #T #B #A #_ #_ #IHB #IHA #L0 #cs #HL0 #X #H #L2 #HL20
elim (lifts_inv_bind1 … H) -H #V0 #T0 #HV0 #HT0 #H destruct
lapply (acr_gcr … H1RP H2RP A) #HA
lapply (acr_gcr … H1RP H2RP B) #HB
- lapply (s1 … HB) -HB #HB
- lapply (s6 … HA G L2 (◊) (◊)) /4 width=5 by lsubc_pair, drops_skip, liftv_nil/
+ lapply (c1 … HB) -HB #HB
+ lapply (c6 … HA G L2 (◊) (◊)) /4 width=5 by lsubc_pair, drops_skip, liftv_nil/
| #a #G #L #W #T #B #A #HLWB #_ #IHB #IHA #L0 #cs #HL0 #X #H #L2 #HL02
elim (lifts_inv_bind1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct
@(acr_abst … H1RP H2RP) /2 width=5 by/
| #G #L #V #T #A #_ #_ #IH1A #IH2A #L0 #cs #HL0 #X #H #L2 #HL20
elim (lifts_inv_flat1 … H) -H #V0 #T0 #HV0 #HT0 #H destruct
lapply (acr_gcr … H1RP H2RP A) #HA
- lapply (s7 … HA G L2 (◊)) /3 width=5 by/
+ lapply (c7 … HA G L2 (◊)) /3 width=5 by/
]
qed.
∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → RP G L T.
#RR #RS #RP #H1RP #H2RP #G #L #T #A #HT
lapply (acr_gcr … H1RP H2RP A) #HA
-@(s1 … HA) /2 width=4 by acr_aaa/
+@(c1 … HA) /2 width=4 by acr_aaa/
qed.
C G L (ⒶVs.ⓓ{a}ⓝW.V.T) → C G L (ⒶVs.ⓐV.ⓛ{a}W.T).
definition S4 ≝ λRP,C:candidate.
- ∀G,L,Vs. all … (RP G L) Vs → ∀k. C G L (ⒶVs.⋆k).
+ ∀G,L,Vs. all … (RP G L) Vs → ∀s. C G L (ⒶVs.⋆s).
definition S5 ≝ λC:candidate. ∀I,G,L,K,Vs,V1,V2,i.
C G L (ⒶVs.V2) → ⬆[0, i+1] V1 ≡ V2 →
⬇[i] L ≡ K.ⓑ{I}V1 → C G L (ⒶVs.#i).
definition S6 ≝ λRP,C:candidate.
- ∀G,L,V1s,V2s. ⬆[0, 1] V1s ≡ V2s →
- ∀a,V,T. C G (L.ⓓV) (ⒶV2s.T) → RP G L V → C G L (ⒶV1s.ⓓ{a}V.T).
+ ∀G,L,V1c,V2c. ⬆[0, 1] V1c ≡ V2c →
+ ∀a,V,T. C G (L.ⓓV) (ⒶV2c.T) → RP G L V → C G L (ⒶV1c.ⓓ{a}V.T).
definition S7 ≝ λC:candidate.
∀G,L,Vs,T,W. C G L (ⒶVs.T) → C G L (ⒶVs.W) → C G L (ⒶVs.ⓝW.T).
(* requirements for the generic reducibility candidate *)
record gcr (RR:relation4 genv lenv term term) (RS:relation term) (RP,C:candidate) : Prop ≝
-{ s1: S1 RP C;
- s2: S2 RR RS RP C;
- s3: S3 C;
- s4: S4 RP C;
- s5: S5 C;
- s6: S6 RP C;
- s7: S7 C
+{ c1: S1 RP C;
+ c2: S2 RR RS RP C;
+ c3: S3 C;
+ c4: S4 RP C;
+ c5: S5 C;
+ c6: S6 RP C;
+ c7: S7 C
}.
(* the functional construction for candidates *)
#RR #RS #RP #H1RP #H2RP #A elim A -A //
#B #A #IHB #IHA @mk_gcr
[ #G #L #T #H
- elim (cp1 … H1RP G L) #k #HK
- lapply (H L (⋆k) T (◊) ? ? ?) -H //
- [ lapply (s2 … IHB G L (◊) … HK) //
- | /3 width=6 by s1, cp3/
+ elim (cp1 … H1RP G L) #s #HK
+ lapply (H L (⋆s) T (◊) ? ? ?) -H //
+ [ lapply (c2 … IHB G L (◊) … HK) //
+ | /3 width=6 by c1, cp3/
]
| #G #L #Vs #HVs #T #H1T #H2T #L0 #V0 #X #cs #HL0 #H #HB
- elim (lifts_inv_applv1 … H) -H #V0s #T0 #HV0s #HT0 #H destruct
- lapply (s1 … IHB … HB) #HV0
- @(s2 … IHA … (V0 @ V0s))
+ elim (lifts_inv_applv1 … H) -H #V0c #T0 #HV0c #HT0 #H destruct
+ lapply (c1 … IHB … HB) #HV0
+ @(c2 … IHA … (V0 @ V0c))
/3 width=14 by gcp2_lifts_all, gcp2_lifts, gcp0_lifts, lifts_simple_dx, conj/
| #a #G #L #Vs #U #T #W #HA #L0 #V0 #X #cs #HL0 #H #HB
- elim (lifts_inv_applv1 … H) -H #V0s #Y #HV0s #HY #H destruct
+ elim (lifts_inv_applv1 … H) -H #V0c #Y #HV0c #HY #H destruct
elim (lifts_inv_flat1 … HY) -HY #U0 #X #HU0 #HX #H destruct
elim (lifts_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT0 #H destruct
- @(s3 … IHA … (V0 @ V0s)) /5 width=6 by lifts_applv, lifts_flat, lifts_bind/
-| #G #L #Vs #HVs #k #L0 #V0 #X #cs #HL0 #H #HB
- elim (lifts_inv_applv1 … H) -H #V0s #Y #HV0s #HY #H destruct
+ @(c3 … IHA … (V0 @ V0c)) /5 width=6 by lifts_applv, lifts_flat, lifts_bind/
+| #G #L #Vs #HVs #s #L0 #V0 #X #cs #HL0 #H #HB
+ elim (lifts_inv_applv1 … H) -H #V0c #Y #HV0c #HY #H destruct
>(lifts_inv_sort1 … HY) -Y
- lapply (s1 … IHB … HB) #HV0
- @(s4 … IHA … (V0 @ V0s)) /3 width=7 by gcp2_lifts_all, conj/
+ lapply (c1 … IHB … HB) #HV0
+ @(c4 … IHA … (V0 @ V0c)) /3 width=7 by gcp2_lifts_all, conj/
| #I #G #L #K #Vs #V1 #V2 #i #HA #HV12 #HLK #L0 #V0 #X #cs #HL0 #H #HB
- elim (lifts_inv_applv1 … H) -H #V0s #Y #HV0s #HY #H destruct
+ elim (lifts_inv_applv1 … H) -H #V0c #Y #HV0c #HY #H destruct
elim (lifts_inv_lref1 … HY) -HY #i0 #Hi0 #H destruct
elim (drops_drop_trans … HL0 … HLK) #X #cs0 #i1 #HL02 #H #Hi1 #Hcs0
>(at_mono … Hi1 … Hi0) in HL02; -i1 #HL02
elim (lift_total W1 0 (i0 + 1)) #W2 #HW12
elim (lifts_lift_trans … Hcs0 … HVW1 … HW12) // -Hcs0 -Hi0 #V3 #HV13 #HVW2
>(lift_mono … HV13 … HV12) in HVW2; -V3 #HVW2
- @(s5 … IHA … (V0 @ V0s) … HW12 HL02) /3 width=5 by lifts_applv/
-| #G #L #V1s #V2s #HV12s #a #V #T #HA #HV #L0 #V10 #X #cs #HL0 #H #HB
- elim (lifts_inv_applv1 … H) -H #V10s #Y #HV10s #HY #H destruct
+ @(c5 … IHA … (V0 @ V0c) … HW12 HL02) /3 width=5 by lifts_applv/
+| #G #L #V1c #V2c #HV12c #a #V #T #HA #HV #L0 #V10 #X #cs #HL0 #H #HB
+ elim (lifts_inv_applv1 … H) -H #V10c #Y #HV10c #HY #H destruct
elim (lifts_inv_bind1 … HY) -HY #V0 #T0 #HV0 #HT0 #H destruct
elim (lift_total V10 0 1) #V20 #HV120
- elim (liftv_total 0 1 V10s) #V20s #HV120s
- @(s6 … IHA … (V10 @ V10s) (V20 @ V20s)) /3 width=7 by gcp2_lifts, liftv_cons/
+ elim (liftv_total 0 1 V10c) #V20c #HV120c
+ @(c6 … IHA … (V10 @ V10c) (V20 @ V20c)) /3 width=7 by gcp2_lifts, liftv_cons/
@(HA … (cs + 1)) /2 width=2 by drops_skip/
[ @lifts_applv //
- elim (liftsv_liftv_trans_le … HV10s … HV120s) -V10s #V10s #HV10s #HV120s
- >(liftv_mono … HV12s … HV10s) -V1s //
+ elim (liftsv_liftv_trans_le … HV10c … HV120c) -V10c #V10c #HV10c #HV120c
+ >(liftv_mono … HV12c … HV10c) -V1c //
| @(gcr_lift … H1RP … HB … HV120) /2 width=2 by drop_drop/
]
| #G #L #Vs #T #W #HA #HW #L0 #V0 #X #cs #HL0 #H #HB
- elim (lifts_inv_applv1 … H) -H #V0s #Y #HV0s #HY #H destruct
+ elim (lifts_inv_applv1 … H) -H #V0c #Y #HV0c #HY #H destruct
elim (lifts_inv_flat1 … HY) -HY #W0 #T0 #HW0 #HT0 #H destruct
- @(s7 … IHA … (V0 @ V0s)) /3 width=5 by lifts_applv/
+ @(c7 … IHA … (V0 @ V0c)) /3 width=5 by lifts_applv/
]
qed.
lapply (acr_gcr … H1RP H2RP B) #HCB
elim (lifts_inv_bind1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct
lapply (gcr_lifts … H1RP … HL0 … HW0 HW) -HW #HW0
-lapply (s3 … HCA … a G L0 (◊)) #H @H -H
-lapply (s6 … HCA G L0 (◊) (◊) ?) // #H @H -H
+lapply (c3 … HCA … a G L0 (◊)) #H @H -H
+lapply (c6 … HCA G L0 (◊) (◊) ?) // #H @H -H
[ @(HA … HL0) //
-| lapply (s1 … HCB) -HCB #HCB
- lapply (s7 … H2RP G L0 (◊)) /3 width=1 by/
+| lapply (c1 … HCB) -HCB #HCB
+ lapply (c7 … H2RP G L0 (◊)) /3 width=1 by/
]
qed.
(* SN EXTENDED STRONGLY CONORMALIZING LOCAL ENVIRONMENTS ********************)
-inductive lcosx (h) (g) (G): relation2 ynat lenv ≝
-| lcosx_sort: ∀l. lcosx h g G l (⋆)
-| lcosx_skip: ∀I,L,T. lcosx h g G 0 L → lcosx h g G 0 (L.ⓑ{I}T)
-| lcosx_pair: ∀I,L,T,l. G ⊢ ⬊*[h, g, T, l] L →
- lcosx h g G l L → lcosx h g G (⫯l) (L.ⓑ{I}T)
+inductive lcosx (h) (o) (G): relation2 ynat lenv ≝
+| lcosx_sort: ∀l. lcosx h o G l (⋆)
+| lcosx_skip: ∀I,L,T. lcosx h o G 0 L → lcosx h o G 0 (L.ⓑ{I}T)
+| lcosx_pair: ∀I,L,T,l. G ⊢ ⬊*[h, o, T, l] L →
+ lcosx h o G l L → lcosx h o G (⫯l) (L.ⓑ{I}T)
.
interpretation
"sn extended strong conormalization (local environment)"
- 'CoSN h g l G L = (lcosx h g G l L).
+ 'CoSN h o l G L = (lcosx h o G l L).
(* Basic properties *********************************************************)
-lemma lcosx_O: ∀h,g,G,L. G ⊢ ~⬊*[h, g, 0] L.
-#h #g #G #L elim L /2 width=1 by lcosx_skip/
+lemma lcosx_O: ∀h,o,G,L. G ⊢ ~⬊*[h, o, 0] L.
+#h #o #G #L elim L /2 width=1 by lcosx_skip/
qed.
-lemma lcosx_drop_trans_lt: ∀h,g,G,L,l. G ⊢ ~⬊*[h, g, l] L →
+lemma lcosx_drop_trans_lt: ∀h,o,G,L,l. G ⊢ ~⬊*[h, o, l] L →
∀I,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → i < l →
- G ⊢ ~⬊*[h, g, ⫰(l-i)] K ∧ G ⊢ ⬊*[h, g, V, ⫰(l-i)] K.
-#h #g #G #L #l #H elim H -L -l
+ G ⊢ ~⬊*[h, o, ⫰(l-i)] K ∧ G ⊢ ⬊*[h, o, V, ⫰(l-i)] K.
+#h #o #G #L #l #H elim H -L -l
[ #l #J #K #V #i #H elim (drop_inv_atom1 … H) -H #H destruct
| #I #L #T #_ #_ #J #K #V #i #_ #H elim (ylt_yle_false … H) -H //
| #I #L #T #l #HT #HL #IHL #J #K #V #i #H #Hil
(* Basic inversion lemmas ***************************************************)
-fact lcosx_inv_succ_aux: ∀h,g,G,L,x. G ⊢ ~⬊*[h, g, x] L → ∀l. x = ⫯l →
+fact lcosx_inv_succ_aux: ∀h,o,G,L,x. G ⊢ ~⬊*[h, o, x] L → ∀l. x = ⫯l →
L = ⋆ ∨
- ∃∃I,K,V. L = K.ⓑ{I}V & G ⊢ ~⬊*[h, g, l] K &
- G ⊢ ⬊*[h, g, V, l] K.
-#h #g #G #L #l * -L -l /2 width=1 by or_introl/
+ ∃∃I,K,V. L = K.ⓑ{I}V & G ⊢ ~⬊*[h, o, l] K &
+ G ⊢ ⬊*[h, o, V, l] K.
+#h #o #G #L #l * -L -l /2 width=1 by or_introl/
[ #I #L #T #_ #x #H elim (ysucc_inv_O_sn … H)
| #I #L #T #l #HT #HL #x #H <(ysucc_inv_inj … H) -x
/3 width=6 by ex3_3_intro, or_intror/
]
qed-.
-lemma lcosx_inv_succ: ∀h,g,G,L,l. G ⊢ ~⬊*[h, g, ⫯l] L → L = ⋆ ∨
- ∃∃I,K,V. L = K.ⓑ{I}V & G ⊢ ~⬊*[h, g, l] K &
- G ⊢ ⬊*[h, g, V, l] K.
+lemma lcosx_inv_succ: ∀h,o,G,L,l. G ⊢ ~⬊*[h, o, ⫯l] L → L = ⋆ ∨
+ ∃∃I,K,V. L = K.ⓑ{I}V & G ⊢ ~⬊*[h, o, l] K &
+ G ⊢ ⬊*[h, o, V, l] K.
/2 width=3 by lcosx_inv_succ_aux/ qed-.
-lemma lcosx_inv_pair: ∀h,g,I,G,L,T,l. G ⊢ ~⬊*[h, g, ⫯l] L.ⓑ{I}T →
- G ⊢ ~⬊*[h, g, l] L ∧ G ⊢ ⬊*[h, g, T, l] L.
-#h #g #I #G #L #T #l #H elim (lcosx_inv_succ … H) -H
+lemma lcosx_inv_pair: ∀h,o,I,G,L,T,l. G ⊢ ~⬊*[h, o, ⫯l] L.ⓑ{I}T →
+ G ⊢ ~⬊*[h, o, l] L ∧ G ⊢ ⬊*[h, o, T, l] L.
+#h #o #I #G #L #T #l #H elim (lcosx_inv_succ … H) -H
[ #H destruct
| * #Z #Y #X #H destruct /2 width=1 by conj/
]
(* *)
(**************************************************************************)
-include "ground_2/ynat/ynat_max.ma".
include "basic_2/computation/lsx_drop.ma".
include "basic_2/computation/lsx_lpx.ma".
include "basic_2/computation/lsx_lpxs.ma".
(* Properties on extended context-sensitive parallel reduction for term *****)
-lemma lsx_cpx_trans_lcosx: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 →
- ∀l. G ⊢ ~⬊*[h, g, l] L →
- G ⊢ ⬊*[h, g, T1, l] L → G ⊢ ⬊*[h, g, T2, l] L.
-#h #g #G #L #T1 #T2 #H elim H -G -L -T1 -T2 //
+lemma lsx_cpx_trans_lcosx: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 →
+ ∀l. G ⊢ ~⬊*[h, o, l] L →
+ G ⊢ ⬊*[h, o, T1, l] L → G ⊢ ⬊*[h, o, T2, l] L.
+#h #o #G #L #T1 #T2 #H elim H -G -L -T1 -T2 //
[ #I #G #L #K #V1 #V2 #W2 #i #HLK #_ #HVW2 #IHV12 #l #HL #H
elim (ylt_split i l) #Hli [ -H | -HL ]
[ <(ymax_pre_sn l (⫯i)) /2 width=1 by ylt_fwd_le_succ1/
]
qed-.
-lemma lsx_cpx_trans_O: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 →
- G ⊢ ⬊*[h, g, T1, 0] L → G ⊢ ⬊*[h, g, T2, 0] L.
+lemma lsx_cpx_trans_O: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 →
+ G ⊢ ⬊*[h, o, T1, 0] L → G ⊢ ⬊*[h, o, T2, 0] L.
/2 width=3 by lsx_cpx_trans_lcosx/ qed-.
(* Properties on context-sensitive parallel computation for terms ***********)
-lemma lprs_cpr_trans: ∀G. s_r_transitive … (cpr G) (λ_. lprs G).
-/3 width=5 by s_r_trans_LTC2, lpr_cprs_trans/ qed-.
+lemma lprs_cpr_trans: ∀G. c_r_transitive … (cpr G) (λ_. lprs G).
+/3 width=5 by c_r_trans_LTC2, lpr_cprs_trans/ qed-.
(* Basic_1: was just: pr3_pr3_pr3_t *)
(* Note: alternative proof /3 width=5 by s_r_trans_LTC1, lprs_cpr_trans/ *)
-lemma lprs_cprs_trans: ∀G. s_rs_transitive … (cpr G) (λ_. lprs G).
-#G @s_r_to_s_rs_trans @s_r_trans_LTC2
-@s_rs_trans_TC1 /2 width=3 by lpr_cprs_trans/ (**) (* full auto too slow *)
+lemma lprs_cprs_trans: ∀G. c_rs_transitive … (cpr G) (λ_. lprs G).
+#G @c_r_to_c_rs_trans @c_r_trans_LTC2
+@c_rs_trans_TC1 /2 width=3 by lpr_cprs_trans/ (**) (* full auto too slow *)
qed-.
lemma lprs_cprs_conf_dx: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡* T1 →
(* SN EXTENDED PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS *******************)
definition lpxs: ∀h. sd h → relation3 genv lenv lenv ≝
- λh,g,G. TC … (lpx h g G).
+ λh,o,G. TC … (lpx h o G).
interpretation "extended parallel computation (local environment, sn variant)"
- 'PRedSnStar h g G L1 L2 = (lpxs h g G L1 L2).
+ 'PRedSnStar h o G L1 L2 = (lpxs h o G L1 L2).
(* Basic eliminators ********************************************************)
-lemma lpxs_ind: ∀h,g,G,L1. ∀R:predicate lenv. R L1 →
- (∀L,L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L → ⦃G, L⦄ ⊢ ➡[h, g] L2 → R L → R L2) →
- ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → R L2.
-#h #g #G #L1 #R #HL1 #IHL1 #L2 #HL12
+lemma lpxs_ind: ∀h,o,G,L1. ∀R:predicate lenv. R L1 →
+ (∀L,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L → ⦃G, L⦄ ⊢ ➡[h, o] L2 → R L → R L2) →
+ ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → R L2.
+#h #o #G #L1 #R #HL1 #IHL1 #L2 #HL12
@(TC_star_ind … HL1 IHL1 … HL12) //
qed-.
-lemma lpxs_ind_dx: ∀h,g,G,L2. ∀R:predicate lenv. R L2 →
- (∀L1,L. ⦃G, L1⦄ ⊢ ➡[h, g] L → ⦃G, L⦄ ⊢ ➡*[h, g] L2 → R L → R L1) →
- ∀L1. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → R L1.
-#h #g #G #L2 #R #HL2 #IHL2 #L1 #HL12
+lemma lpxs_ind_dx: ∀h,o,G,L2. ∀R:predicate lenv. R L2 →
+ (∀L1,L. ⦃G, L1⦄ ⊢ ➡[h, o] L → ⦃G, L⦄ ⊢ ➡*[h, o] L2 → R L → R L1) →
+ ∀L1. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → R L1.
+#h #o #G #L2 #R #HL2 #IHL2 #L1 #HL12
@(TC_star_ind_dx … HL2 IHL2 … HL12) //
qed-.
(* Basic properties *********************************************************)
-lemma lprs_lpxs: ∀h,g,G,L1,L2. ⦃G, L1⦄ ⊢ ➡* L2 → ⦃G, L1⦄ ⊢ ➡*[h, g] L2.
+lemma lprs_lpxs: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡* L2 → ⦃G, L1⦄ ⊢ ➡*[h, o] L2.
/3 width=3 by lpr_lpx, monotonic_TC/ qed.
-lemma lpx_lpxs: ∀h,g,G,L1,L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 → ⦃G, L1⦄ ⊢ ➡*[h, g] L2.
+lemma lpx_lpxs: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → ⦃G, L1⦄ ⊢ ➡*[h, o] L2.
/2 width=1 by inj/ qed.
-lemma lpxs_refl: ∀h,g,G,L. ⦃G, L⦄ ⊢ ➡*[h, g] L.
+lemma lpxs_refl: ∀h,o,G,L. ⦃G, L⦄ ⊢ ➡*[h, o] L.
/2 width=1 by lprs_lpxs/ qed.
-lemma lpxs_strap1: ∀h,g,G,L1,L,L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L → ⦃G, L⦄ ⊢ ➡[h, g] L2 → ⦃G, L1⦄ ⊢ ➡*[h, g] L2.
+lemma lpxs_strap1: ∀h,o,G,L1,L,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L → ⦃G, L⦄ ⊢ ➡[h, o] L2 → ⦃G, L1⦄ ⊢ ➡*[h, o] L2.
/2 width=3 by step/ qed.
-lemma lpxs_strap2: ∀h,g,G,L1,L,L2. ⦃G, L1⦄ ⊢ ➡[h, g] L → ⦃G, L⦄ ⊢ ➡*[h, g] L2 → ⦃G, L1⦄ ⊢ ➡*[h, g] L2.
+lemma lpxs_strap2: ∀h,o,G,L1,L,L2. ⦃G, L1⦄ ⊢ ➡[h, o] L → ⦃G, L⦄ ⊢ ➡*[h, o] L2 → ⦃G, L1⦄ ⊢ ➡*[h, o] L2.
/2 width=3 by TC_strap/ qed.
-lemma lpxs_pair_refl: ∀h,g,G,L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → ∀I,V. ⦃G, L1.ⓑ{I}V⦄ ⊢ ➡*[h, g] L2.ⓑ{I}V.
+lemma lpxs_pair_refl: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → ∀I,V. ⦃G, L1.ⓑ{I}V⦄ ⊢ ➡*[h, o] L2.ⓑ{I}V.
/2 width=1 by TC_lpx_sn_pair_refl/ qed.
(* Basic inversion lemmas ***************************************************)
-lemma lpxs_inv_atom1: ∀h,g,G,L2. ⦃G, ⋆⦄ ⊢ ➡*[h, g] L2 → L2 = ⋆.
+lemma lpxs_inv_atom1: ∀h,o,G,L2. ⦃G, ⋆⦄ ⊢ ➡*[h, o] L2 → L2 = ⋆.
/2 width=2 by TC_lpx_sn_inv_atom1/ qed-.
-lemma lpxs_inv_atom2: ∀h,g,G,L1. ⦃G, L1⦄ ⊢ ➡*[h, g] ⋆ → L1 = ⋆.
+lemma lpxs_inv_atom2: ∀h,o,G,L1. ⦃G, L1⦄ ⊢ ➡*[h, o] ⋆ → L1 = ⋆.
/2 width=2 by TC_lpx_sn_inv_atom2/ qed-.
(* Basic forward lemmas *****************************************************)
-lemma lpxs_fwd_length: ∀h,g,G,L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → |L1| = |L2|.
+lemma lpxs_fwd_length: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → |L1| = |L2|.
/2 width=2 by TC_lpx_sn_fwd_length/ qed-.
(* Properties about atomic arity assignment on terms ************************)
-lemma lpxs_aaa_conf: ∀h,g,G,L1,T,A. ⦃G, L1⦄ ⊢ T ⁝ A →
- ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → ⦃G, L2⦄ ⊢ T ⁝ A.
-#h #g #G #L1 #T #A #HT #L2 #HL12
+lemma lpxs_aaa_conf: ∀h,o,G,L1,T,A. ⦃G, L1⦄ ⊢ T ⁝ A →
+ ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → ⦃G, L2⦄ ⊢ T ⁝ A.
+#h #o #G #L1 #T #A #HT #L2 #HL12
@(TC_Conf3 … (λL,A. ⦃G, L⦄ ⊢ T ⁝ A) … HT ? HL12) /2 width=5 by lpx_aaa_conf/
qed-.
(* Advanced properties ******************************************************)
-lemma lpxs_pair: ∀h,g,I,G,L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 →
- ∀V1,V2. ⦃G, L1⦄ ⊢ V1 ➡*[h, g] V2 →
- ⦃G, L1.ⓑ{I}V1⦄ ⊢ ➡*[h, g] L2.ⓑ{I}V2.
+lemma lpxs_pair: ∀h,o,I,G,L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 →
+ ∀V1,V2. ⦃G, L1⦄ ⊢ V1 ➡*[h, o] V2 →
+ ⦃G, L1.ⓑ{I}V1⦄ ⊢ ➡*[h, o] L2.ⓑ{I}V2.
/2 width=1 by TC_lpx_sn_pair/ qed.
(* Advanced inversion lemmas ************************************************)
-lemma lpxs_inv_pair1: ∀h,g,I,G,K1,L2,V1. ⦃G, K1.ⓑ{I}V1⦄ ⊢ ➡*[h, g] L2 →
- ∃∃K2,V2. ⦃G, K1⦄ ⊢ ➡*[h, g] K2 & ⦃G, K1⦄ ⊢ V1 ➡*[h, g] V2 & L2 = K2.ⓑ{I}V2.
+lemma lpxs_inv_pair1: ∀h,o,I,G,K1,L2,V1. ⦃G, K1.ⓑ{I}V1⦄ ⊢ ➡*[h, o] L2 →
+ ∃∃K2,V2. ⦃G, K1⦄ ⊢ ➡*[h, o] K2 & ⦃G, K1⦄ ⊢ V1 ➡*[h, o] V2 & L2 = K2.ⓑ{I}V2.
/3 width=3 by TC_lpx_sn_inv_pair1, lpx_cpxs_trans/ qed-.
-lemma lpxs_inv_pair2: ∀h,g,I,G,L1,K2,V2. ⦃G, L1⦄ ⊢ ➡*[h, g] K2.ⓑ{I}V2 →
- ∃∃K1,V1. ⦃G, K1⦄ ⊢ ➡*[h, g] K2 & ⦃G, K1⦄ ⊢ V1 ➡*[h, g] V2 & L1 = K1.ⓑ{I}V1.
+lemma lpxs_inv_pair2: ∀h,o,I,G,L1,K2,V2. ⦃G, L1⦄ ⊢ ➡*[h, o] K2.ⓑ{I}V2 →
+ ∃∃K1,V1. ⦃G, K1⦄ ⊢ ➡*[h, o] K2 & ⦃G, K1⦄ ⊢ V1 ➡*[h, o] V2 & L1 = K1.ⓑ{I}V1.
/3 width=3 by TC_lpx_sn_inv_pair2, lpx_cpxs_trans/ qed-.
(* Advanced eliminators *****************************************************)
-lemma lpxs_ind_alt: ∀h,g,G. ∀R:relation lenv.
+lemma lpxs_ind_alt: ∀h,o,G. ∀R:relation lenv.
R (⋆) (⋆) → (
∀I,K1,K2,V1,V2.
- ⦃G, K1⦄ ⊢ ➡*[h, g] K2 → ⦃G, K1⦄ ⊢ V1 ➡*[h, g] V2 →
+ ⦃G, K1⦄ ⊢ ➡*[h, o] K2 → ⦃G, K1⦄ ⊢ V1 ➡*[h, o] V2 →
R K1 K2 → R (K1.ⓑ{I}V1) (K2.ⓑ{I}V2)
) →
- ∀L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → R L1 L2.
+ ∀L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → R L1 L2.
/3 width=4 by TC_lpx_sn_ind, lpx_cpxs_trans/ qed-.
(* Properties on context-sensitive extended parallel computation for terms **)
-lemma lpxs_cpx_trans: ∀h,g,G. s_r_transitive … (cpx h g G) (λ_.lpxs h g G).
-/3 width=5 by s_r_trans_LTC2, lpx_cpxs_trans/ qed-.
+lemma lpxs_cpx_trans: ∀h,o,G. c_r_transitive … (cpx h o G) (λ_.lpxs h o G).
+/3 width=5 by c_r_trans_LTC2, lpx_cpxs_trans/ qed-.
(* Note: alternative proof: /3 width=5 by s_r_trans_TC1, lpxs_cpx_trans/ *)
-lemma lpxs_cpxs_trans: ∀h,g,G. s_rs_transitive … (cpx h g G) (λ_.lpxs h g G).
-#h #g #G @s_r_to_s_rs_trans @s_r_trans_LTC2
-@s_rs_trans_TC1 /2 width=3 by lpx_cpxs_trans/ (**) (* full auto too slow *)
+lemma lpxs_cpxs_trans: ∀h,o,G. c_rs_transitive … (cpx h o G) (λ_.lpxs h o G).
+#h #o #G @c_r_to_c_rs_trans @c_r_trans_LTC2
+@c_rs_trans_TC1 /2 width=3 by lpx_cpxs_trans/ (**) (* full auto too slow *)
qed-.
-lemma cpxs_bind2: ∀h,g,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡*[h, g] V2 →
- ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ➡*[h, g] T2 →
- ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, g] ⓑ{a,I}V2.T2.
+lemma cpxs_bind2: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 →
+ ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ➡*[h, o] T2 →
+ ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, o] ⓑ{a,I}V2.T2.
/4 width=5 by lpxs_cpxs_trans, lpxs_pair, cpxs_bind/ qed.
(* Inversion lemmas on context-sensitive ext parallel computation for terms *)
-lemma cpxs_inv_abst1: ∀h,g,a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{a}V1.T1 ➡*[h, g] U2 →
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡*[h, g] V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 ➡*[h, g] T2 &
+lemma cpxs_inv_abst1: ∀h,o,a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{a}V1.T1 ➡*[h, o] U2 →
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 ➡*[h, o] T2 &
U2 = ⓛ{a}V2.T2.
-#h #g #a #G #L #V1 #T1 #U2 #H @(cpxs_ind … H) -U2 /2 width=5 by ex3_2_intro/
+#h #o #a #G #L #V1 #T1 #U2 #H @(cpxs_ind … H) -U2 /2 width=5 by ex3_2_intro/
#U0 #U2 #_ #HU02 * #V0 #T0 #HV10 #HT10 #H destruct
elim (cpx_inv_abst1 … HU02) -HU02 #V2 #T2 #HV02 #HT02 #H destruct
lapply (lpxs_cpx_trans … HT02 (L.ⓛV1) ?)
/3 width=5 by lpxs_pair, cpxs_trans, cpxs_strap1, ex3_2_intro/
qed-.
-lemma cpxs_inv_abbr1: ∀h,g,a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{a}V1.T1 ➡*[h, g] U2 → (
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡*[h, g] V2 & ⦃G, L.ⓓV1⦄ ⊢ T1 ➡*[h, g] T2 &
+lemma cpxs_inv_abbr1: ∀h,o,a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{a}V1.T1 ➡*[h, o] U2 → (
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 & ⦃G, L.ⓓV1⦄ ⊢ T1 ➡*[h, o] T2 &
U2 = ⓓ{a}V2.T2
) ∨
- ∃∃T2. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡*[h, g] T2 & ⬆[0, 1] U2 ≡ T2 & a = true.
-#h #g #a #G #L #V1 #T1 #U2 #H @(cpxs_ind … H) -U2 /3 width=5 by ex3_2_intro, or_introl/
+ ∃∃T2. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡*[h, o] T2 & ⬆[0, 1] U2 ≡ T2 & a = true.
+#h #o #a #G #L #V1 #T1 #U2 #H @(cpxs_ind … H) -U2 /3 width=5 by ex3_2_intro, or_introl/
#U0 #U2 #_ #HU02 * *
[ #V0 #T0 #HV10 #HT10 #H destruct
elim (cpx_inv_abbr1 … HU02) -HU02 *
(* More advanced properties *************************************************)
-lemma lpxs_pair2: ∀h,g,I,G,L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 →
- ∀V1,V2. ⦃G, L2⦄ ⊢ V1 ➡*[h, g] V2 → ⦃G, L1.ⓑ{I}V1⦄ ⊢ ➡*[h, g] L2.ⓑ{I}V2.
+lemma lpxs_pair2: ∀h,o,I,G,L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 →
+ ∀V1,V2. ⦃G, L2⦄ ⊢ V1 ➡*[h, o] V2 → ⦃G, L1.ⓑ{I}V1⦄ ⊢ ➡*[h, o] L2.ⓑ{I}V2.
/3 width=3 by lpxs_pair, lpxs_cpxs_trans/ qed.
(* Properties on supclosure *************************************************)
-lemma lpx_fqup_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, g] L1 →
- ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡*[h, g] T & ⦃G1, K1, T⦄ ⊐+ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, g] L2.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
+lemma lpx_fqup_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 →
+ ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡*[h, o] T & ⦃G1, K1, T⦄ ⊐+ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
[ #G2 #L2 #T2 #H12 #K1 #HKL1 elim (lpx_fqu_trans … H12 … HKL1) -L1
/3 width=5 by cpx_cpxs, fqu_fqup, ex3_2_intro/
| #G #G2 #L #L2 #T #T2 #_ #H2 #IH1 #K1 #HLK1 elim (IH1 … HLK1) -L1
]
qed-.
-lemma lpx_fqus_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, g] L1 →
- ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡*[h, g] T & ⦃G1, K1, T⦄ ⊐* ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, g] L2.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqus_ind … H) -G2 -L2 -T2 [ /2 width=5 by ex3_2_intro/ ]
+lemma lpx_fqus_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 →
+ ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡*[h, o] T & ⦃G1, K1, T⦄ ⊐* ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqus_ind … H) -G2 -L2 -T2 [ /2 width=5 by ex3_2_intro/ ]
#G #G2 #L #L2 #T #T2 #_ #H2 #IH1 #K1 #HLK1 elim (IH1 … HLK1) -L1
#L0 #T0 #HT10 #HT0 #HL0 elim (lpx_fquq_trans … H2 … HL0) -L
#L #T3 #HT3 #HT32 #HL2 elim (fqus_cpx_trans … HT0 … HT3) -T
/3 width=7 by cpxs_strap1, fqus_strap1, ex3_2_intro/
qed-.
-lemma lpxs_fquq_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, g] L1 →
- ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡*[h, g] T & ⦃G1, K1, T⦄ ⊐⸮ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, g] L2.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #HT12 #K1 #H @(lpxs_ind_dx … H) -K1
+lemma lpxs_fquq_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, o] L1 →
+ ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡*[h, o] T & ⦃G1, K1, T⦄ ⊐⸮ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, o] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #HT12 #K1 #H @(lpxs_ind_dx … H) -K1
[ /2 width=5 by ex3_2_intro/
| #K1 #K #HK1 #_ * #L #T #HT1 #HT2 #HL2 -HT12
lapply (lpx_cpxs_trans … HT1 … HK1) -HT1
]
qed-.
-lemma lpxs_fqup_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, g] L1 →
- ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡*[h, g] T & ⦃G1, K1, T⦄ ⊐+ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, g] L2.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #HT12 #K1 #H @(lpxs_ind_dx … H) -K1
+lemma lpxs_fqup_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, o] L1 →
+ ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡*[h, o] T & ⦃G1, K1, T⦄ ⊐+ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, o] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #HT12 #K1 #H @(lpxs_ind_dx … H) -K1
[ /2 width=5 by ex3_2_intro/
| #K1 #K #HK1 #_ * #L #T #HT1 #HT2 #HL2 -HT12
lapply (lpx_cpxs_trans … HT1 … HK1) -HT1
]
qed-.
-lemma lpxs_fqus_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, g] L1 →
- ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡*[h, g] T & ⦃G1, K1, T⦄ ⊐* ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, g] L2.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqus_ind … H) -G2 -L2 -T2 /2 width=5 by ex3_2_intro/
+lemma lpxs_fqus_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, o] L1 →
+ ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡*[h, o] T & ⦃G1, K1, T⦄ ⊐* ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, o] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqus_ind … H) -G2 -L2 -T2 /2 width=5 by ex3_2_intro/
#G #G2 #L #L2 #T #T2 #_ #H2 #IH1 #K1 #HLK1 elim (IH1 … HLK1) -L1
#L0 #T0 #HT10 #HT0 #HL0 elim (lpxs_fquq_trans … H2 … HL0) -L
#L #T3 #HT3 #HT32 #HL2 elim (fqus_cpxs_trans … HT3 … HT0) -T
(* Properties on local environment slicing ***********************************)
-lemma lpxs_drop_conf: ∀h,g,G. dropable_sn (lpxs h g G).
+lemma lpxs_drop_conf: ∀h,o,G. dropable_sn (lpxs h o G).
/3 width=3 by dropable_sn_TC, lpx_drop_conf/ qed-.
-lemma drop_lpxs_trans: ∀h,g,G. dedropable_sn (lpxs h g G).
+lemma drop_lpxs_trans: ∀h,o,G. dedropable_sn (lpxs h o G).
/3 width=3 by dedropable_sn_TC, drop_lpx_trans/ qed-.
-lemma lpxs_drop_trans_O1: ∀h,g,G. dropable_dx (lpxs h g G).
+lemma lpxs_drop_trans_O1: ∀h,o,G. dropable_dx (lpxs h o G).
/3 width=3 by dropable_dx_TC, lpx_drop_trans_O1/ qed-.
(* Properties on lazy equivalence for local environments ********************)
-lemma lleq_lpxs_trans: ∀h,g,G,L2,K2. ⦃G, L2⦄ ⊢ ➡*[h, g] K2 →
+lemma lleq_lpxs_trans: ∀h,o,G,L2,K2. ⦃G, L2⦄ ⊢ ➡*[h, o] K2 →
∀L1,T,l. L1 ≡[T, l] L2 →
- ∃∃K1. ⦃G, L1⦄ ⊢ ➡*[h, g] K1 & K1 ≡[T, l] K2.
-#h #g #G #L2 #K2 #H @(lpxs_ind … H) -K2 /2 width=3 by ex2_intro/
+ ∃∃K1. ⦃G, L1⦄ ⊢ ➡*[h, o] K1 & K1 ≡[T, l] K2.
+#h #o #G #L2 #K2 #H @(lpxs_ind … H) -K2 /2 width=3 by ex2_intro/
#K #K2 #_ #HK2 #IH #L1 #T #l #HT elim (IH … HT) -L2
#L #HL1 #HT elim (lleq_lpx_trans … HK2 … HT) -K
/3 width=3 by lpxs_strap1, ex2_intro/
qed-.
-lemma lpxs_nlleq_inv_step_sn: ∀h,g,G,L1,L2,T,l. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → (L1 ≡[T, l] L2 → ⊥) →
- ∃∃L,L0. ⦃G, L1⦄ ⊢ ➡[h, g] L & L1 ≡[T, l] L → ⊥ & ⦃G, L⦄ ⊢ ➡*[h, g] L0 & L0 ≡[T, l] L2.
-#h #g #G #L1 #L2 #T #l #H @(lpxs_ind_dx … H) -L1
+lemma lpxs_nlleq_inv_step_sn: ∀h,o,G,L1,L2,T,l. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) →
+ ∃∃L,L0. ⦃G, L1⦄ ⊢ ➡[h, o] L & L1 ≡[T, l] L → ⊥ & ⦃G, L⦄ ⊢ ➡*[h, o] L0 & L0 ≡[T, l] L2.
+#h #o #G #L1 #L2 #T #l #H @(lpxs_ind_dx … H) -L1
[ #H elim H -H //
| #L1 #L #H1 #H2 #IH2 #H12 elim (lleq_dec T L1 L l) #H
[ -H1 -H2 elim IH2 -IH2 /3 width=3 by lleq_trans/ -H12
]
qed-.
-lemma lpxs_lleq_fqu_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, g] L1 → K1 ≡[T1, 0] L1 →
- ∃∃K2. ⦃G1, K1, T1⦄ ⊐ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, g] L2 & K2 ≡[T2, 0] L2.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
+lemma lpxs_lleq_fqu_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, o] L1 → K1 ≡[T1, 0] L1 →
+ ∃∃K2. ⦃G1, K1, T1⦄ ⊐ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, o] L2 & K2 ≡[T2, 0] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
[ #I #G1 #L1 #V1 #X #H1 #H2 elim (lpxs_inv_pair2 … H1) -H1
#K0 #V0 #H1KL1 #_ #H destruct
elim (lleq_inv_lref_ge_dx … H2 ? I L1 V1) -H2 //
/3 width=4 by lpxs_pair, fqu_bind_dx, ex3_intro/
| #I #G1 #L1 #V1 #T1 #K1 #HLK1 #H elim (lleq_inv_flat … H) -H
/2 width=4 by fqu_flat_dx, ex3_intro/
-| #G1 #L1 #L #T1 #U1 #m #HL1 #HTU1 #K1 #H1KL1 #H2KL1
- elim (drop_O1_le (Ⓕ) (m+1) K1)
+| #G1 #L1 #L #T1 #U1 #k #HL1 #HTU1 #K1 #H1KL1 #H2KL1
+ elim (drop_O1_le (Ⓕ) (k+1) K1)
[ #K #HK1 lapply (lleq_inv_lift_le … H2KL1 … HK1 HL1 … HTU1 ?) -H2KL1 //
#H2KL elim (lpxs_drop_trans_O1 … H1KL1 … HL1) -L1
#K0 #HK10 #H1KL lapply (drop_mono … HK10 … HK1) -HK10 #H destruct
/3 width=4 by fqu_drop, ex3_intro/
- | lapply (drop_fwd_length_le2 … HL1) -L -T1 -g
+ | lapply (drop_fwd_length_le2 … HL1) -L -T1 -o
lapply (lleq_fwd_length … H2KL1) //
]
]
qed-.
-lemma lpxs_lleq_fquq_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, g] L1 → K1 ≡[T1, 0] L1 →
- ∃∃K2. ⦃G1, K1, T1⦄ ⊐⸮ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, g] L2 & K2 ≡[T2, 0] L2.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #H1KL1 #H2KL1
+lemma lpxs_lleq_fquq_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, o] L1 → K1 ≡[T1, 0] L1 →
+ ∃∃K2. ⦃G1, K1, T1⦄ ⊐⸮ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, o] L2 & K2 ≡[T2, 0] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #H1KL1 #H2KL1
elim (fquq_inv_gen … H) -H
[ #H elim (lpxs_lleq_fqu_trans … H … H1KL1 H2KL1) -L1
/3 width=4 by fqu_fquq, ex3_intro/
]
qed-.
-lemma lpxs_lleq_fqup_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, g] L1 → K1 ≡[T1, 0] L1 →
- ∃∃K2. ⦃G1, K1, T1⦄ ⊐+ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, g] L2 & K2 ≡[T2, 0] L2.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
+lemma lpxs_lleq_fqup_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, o] L1 → K1 ≡[T1, 0] L1 →
+ ∃∃K2. ⦃G1, K1, T1⦄ ⊐+ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, o] L2 & K2 ≡[T2, 0] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
[ #G2 #L2 #T2 #H #K1 #H1KL1 #H2KL1 elim (lpxs_lleq_fqu_trans … H … H1KL1 H2KL1) -L1
/3 width=4 by fqu_fqup, ex3_intro/
| #G #G2 #L #L2 #T #T2 #_ #HT2 #IHT1 #K1 #H1KL1 #H2KL1 elim (IHT1 … H2KL1) // -L1
]
qed-.
-lemma lpxs_lleq_fqus_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
- ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, g] L1 → K1 ≡[T1, 0] L1 →
- ∃∃K2. ⦃G1, K1, T1⦄ ⊐* ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, g] L2 & K2 ≡[T2, 0] L2.
-#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #H1KL1 #H2KL1
+lemma lpxs_lleq_fqus_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+ ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, o] L1 → K1 ≡[T1, 0] L1 →
+ ∃∃K2. ⦃G1, K1, T1⦄ ⊐* ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, o] L2 & K2 ≡[T2, 0] L2.
+#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #H1KL1 #H2KL1
elim (fqus_inv_gen … H) -H
[ #H elim (lpxs_lleq_fqup_trans … H … H1KL1 H2KL1) -L1
/3 width=4 by fqup_fqus, ex3_intro/
]
qed-.
-fact lreq_lpxs_trans_lleq_aux: ∀h,g,G,L1,L0,l,m. L1 ⩬[l, m] L0 → m = ∞ →
- ∀L2. ⦃G, L0⦄ ⊢ ➡*[h, g] L2 →
- ∃∃L. L ⩬[l, m] L2 & ⦃G, L1⦄ ⊢ ➡*[h, g] L &
+fact lreq_lpxs_trans_lleq_aux: ∀h,o,G,L1,L0,l,k. L1 ⩬[l, k] L0 → k = ∞ →
+ ∀L2. ⦃G, L0⦄ ⊢ ➡*[h, o] L2 →
+ ∃∃L. L ⩬[l, k] L2 & ⦃G, L1⦄ ⊢ ➡*[h, o] L &
(∀T. L0 ≡[T, l] L2 ↔ L1 ≡[T, l] L).
-#h #g #G #L1 #L0 #l #m #H elim H -L1 -L0 -l -m
-[ #l #m #_ #L2 #H >(lpxs_inv_atom1 … H) -H
+#h #o #G #L1 #L0 #l #k #H elim H -L1 -L0 -l -k
+[ #l #k #_ #L2 #H >(lpxs_inv_atom1 … H) -H
/3 width=5 by ex3_intro, conj/
| #I1 #I0 #L1 #L0 #V1 #V0 #_ #_ #Hm destruct
-| #I #L1 #L0 #V1 #m #HL10 #IHL10 #Hm #Y #H
+| #I #L1 #L0 #V1 #k #HL10 #IHL10 #Hm #Y #H
elim (lpxs_inv_pair1 … H) -H #L2 #V2 #HL02 #HV02 #H destruct
lapply (ysucc_inv_Y_dx … Hm) -Hm #Hm
elim (IHL10 … HL02) // -IHL10 -HL02 #L #HL2 #HL1 #IH
@(ex3_intro … (L.ⓑ{I}V2)) /3 width=3 by lpxs_pair, lreq_cpxs_trans, lreq_pair/
#T elim (IH T) #HL0dx #HL0sn
@conj #H @(lleq_lreq_repl … H) -H /3 width=1 by lreq_sym, lreq_pair_O_Y/
-| #I1 #I0 #L1 #L0 #V1 #V0 #l #m #HL10 #IHL10 #Hm #Y #H
+| #I1 #I0 #L1 #L0 #V1 #V0 #l #k #HL10 #IHL10 #Hm #Y #H
elim (lpxs_inv_pair1 … H) -H #L2 #V2 #HL02 #HV02 #H destruct
elim (IHL10 … HL02) // -IHL10 -HL02 #L #HL2 #HL1 #IH
@(ex3_intro … (L.ⓑ{I1}V1)) /3 width=1 by lpxs_pair, lreq_succ/
]
qed-.
-lemma lreq_lpxs_trans_lleq: ∀h,g,G,L1,L0,l. L1 ⩬[l, ∞] L0 →
- ∀L2. ⦃G, L0⦄ ⊢ ➡*[h, g] L2 →
- ∃∃L. L ⩬[l, ∞] L2 & ⦃G, L1⦄ ⊢ ➡*[h, g] L &
+lemma lreq_lpxs_trans_lleq: ∀h,o,G,L1,L0,l. L1 ⩬[l, ∞] L0 →
+ ∀L2. ⦃G, L0⦄ ⊢ ➡*[h, o] L2 →
+ ∃∃L. L ⩬[l, ∞] L2 & ⦃G, L1⦄ ⊢ ➡*[h, o] L &
(∀T. L0 ≡[T, l] L2 ↔ L1 ≡[T, l] L).
/2 width=1 by lreq_lpxs_trans_lleq_aux/ qed-.
(* Main properties **********************************************************)
-theorem lpxs_trans: ∀h,g,G. Transitive … (lpxs h g G).
+theorem lpxs_trans: ∀h,o,G. Transitive … (lpxs h o G).
/2 width=3 by trans_TC/ qed-.
(* Basic_1: was: csubc_drop_conf_O *)
(* Note: the constant 0 can not be generalized *)
-lemma lsubc_drop_O1_trans: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → ∀K2,s,m. ⬇[s, 0, m] L2 ≡ K2 →
- ∃∃K1. ⬇[s, 0, m] L1 ≡ K1 & G ⊢ K1 ⫃[RP] K2.
+lemma lsubc_drop_O1_trans: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → ∀K2,c,k. ⬇[c, 0, k] L2 ≡ K2 →
+ ∃∃K1. ⬇[c, 0, k] L1 ≡ K1 & G ⊢ K1 ⫃[RP] K2.
#RP #G #L1 #L2 #H elim H -L1 -L2
-[ #X #s #m #H elim (drop_inv_atom1 … H) -H /4 width=3 by drop_atom, ex2_intro/
-| #I #L1 #L2 #V #_ #IHL12 #X #s #m #H
+[ #X #c #k #H elim (drop_inv_atom1 … H) -H /4 width=3 by drop_atom, ex2_intro/
+| #I #L1 #L2 #V #_ #IHL12 #X #c #k #H
elim (drop_inv_O1_pair1 … H) -H * #Hm #H destruct
- [ elim (IHL12 L2 s 0) -IHL12 // #X #H <(drop_inv_O2 … H) -H
+ [ elim (IHL12 L2 c 0) -IHL12 // #X #H <(drop_inv_O2 … H) -H
/3 width=3 by lsubc_pair, drop_pair, ex2_intro/
| elim (IHL12 … H) -L2 /3 width=3 by drop_drop_lt, ex2_intro/
]
-| #L1 #L2 #V #W #A #HV #H1W #H2W #_ #IHL12 #X #s #m #H
+| #L1 #L2 #V #W #A #HV #H1W #H2W #_ #IHL12 #X #c #k #H
elim (drop_inv_O1_pair1 … H) -H * #Hm #H destruct
- [ elim (IHL12 L2 s 0) -IHL12 // #X #H <(drop_inv_O2 … H) -H
+ [ elim (IHL12 L2 c 0) -IHL12 // #X #H <(drop_inv_O2 … H) -H
/3 width=8 by lsubc_beta, drop_pair, ex2_intro/
| elim (IHL12 … H) -L2 /3 width=3 by drop_drop_lt, ex2_intro/
]
(* Basic_1: was: csubc_drop_conf_rev *)
lemma drop_lsubc_trans: ∀RR,RS,RP. gcp RR RS RP →
- ∀G,L1,K1,l,m. ⬇[Ⓕ, l, m] L1 ≡ K1 → ∀K2. G ⊢ K1 ⫃[RP] K2 →
- ∃∃L2. G ⊢ L1 ⫃[RP] L2 & ⬇[Ⓕ, l, m] L2 ≡ K2.
-#RR #RS #RP #Hgcp #G #L1 #K1 #l #m #H elim H -L1 -K1 -l -m
-[ #l #m #Hm #X #H elim (lsubc_inv_atom1 … H) -H
+ ∀G,L1,K1,l,k. ⬇[Ⓕ, l, k] L1 ≡ K1 → ∀K2. G ⊢ K1 ⫃[RP] K2 →
+ ∃∃L2. G ⊢ L1 ⫃[RP] L2 & ⬇[Ⓕ, l, k] L2 ≡ K2.
+#RR #RS #RP #Hgcp #G #L1 #K1 #l #k #H elim H -L1 -K1 -l -k
+[ #l #k #Hm #X #H elim (lsubc_inv_atom1 … H) -H
>Hm /2 width=3 by ex2_intro/
| #L1 #I #V1 #X #H
elim (lsubc_inv_pair1 … H) -H *
| #K1 #V #W1 #A #HV1 #H1W1 #H2W1 #HLK1 #H1 #H2 #H3 destruct
/3 width=4 by lsubc_beta, drop_pair, ex2_intro/
]
-| #I #L1 #K1 #V1 #m #_ #IHLK1 #K2 #HK12
+| #I #L1 #K1 #V1 #k #_ #IHLK1 #K2 #HK12
elim (IHLK1 … HK12) -K1 /3 width=5 by lsubc_pair, drop_drop, ex2_intro/
-| #I #L1 #K1 #V1 #V2 #l #m #HLK1 #HV21 #IHLK1 #X #H
+| #I #L1 #K1 #V1 #V2 #l #k #HLK1 #HV21 #IHLK1 #X #H
elim (lsubc_inv_pair1 … H) -H *
[ #K2 #HK12 #H destruct
elim (IHLK1 … HK12) -K1 /3 width=5 by lsubc_pair, drop_skip, ex2_intro/
∃∃L2. G ⊢ L1 ⫃[RP] L2 & ⬇*[Ⓕ, cs] L2 ≡ K2.
#RR #RS #RP #Hgcp #G #L1 #K1 #cs #H elim H -L1 -K1 -cs
[ /2 width=3 by drops_nil, ex2_intro/
-| #L1 #L #K1 #cs #l #m #_ #HLK1 #IHL #K2 #HK12
+| #L1 #L #K1 #cs #l #k #_ #HLK1 #IHL #K2 #HK12
elim (drop_lsubc_trans … Hgcp … HLK1 … HK12) -Hgcp -K1 #K #HLK #HK2
elim (IHL … HLK) -IHL -HLK /3 width=5 by drops_cons, ex2_intro/
]
(* SN EXTENDED STRONGLY NORMALIZING LOCAL ENVIRONMENTS **********************)
definition lsx: ∀h. sd h → relation4 ynat term genv lenv ≝
- λh,g,l,T,G. SN … (lpx h g G) (lleq l T).
+ λh,o,l,T,G. SN … (lpx h o G) (lleq l T).
interpretation
"extended strong normalization (local environment)"
- 'SN h g l T G L = (lsx h g T l G L).
+ 'SN h o l T G L = (lsx h o T l G L).
(* Basic eliminators ********************************************************)
-lemma lsx_ind: ∀h,g,G,T,l. ∀R:predicate lenv.
- (∀L1. G ⊢ ⬊*[h, g, T, l] L1 →
- (∀L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 → (L1 ≡[T, l] L2 → ⊥) → R L2) →
+lemma lsx_ind: ∀h,o,G,T,l. ∀R:predicate lenv.
+ (∀L1. G ⊢ ⬊*[h, o, T, l] L1 →
+ (∀L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) → R L2) →
R L1
) →
- ∀L. G ⊢ ⬊*[h, g, T, l] L → R L.
-#h #g #G #T #l #R #H0 #L1 #H elim H -L1
+ ∀L. G ⊢ ⬊*[h, o, T, l] L → R L.
+#h #o #G #T #l #R #H0 #L1 #H elim H -L1
/5 width=1 by lleq_sym, SN_intro/
qed-.
(* Basic properties *********************************************************)
-lemma lsx_intro: ∀h,g,G,L1,T,l.
- (∀L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊*[h, g, T, l] L2) →
- G ⊢ ⬊*[h, g, T, l] L1.
+lemma lsx_intro: ∀h,o,G,L1,T,l.
+ (∀L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊*[h, o, T, l] L2) →
+ G ⊢ ⬊*[h, o, T, l] L1.
/5 width=1 by lleq_sym, SN_intro/ qed.
-lemma lsx_atom: ∀h,g,G,T,l. G ⊢ ⬊*[h, g, T, l] ⋆.
-#h #g #G #T #l @lsx_intro
+lemma lsx_atom: ∀h,o,G,T,l. G ⊢ ⬊*[h, o, T, l] ⋆.
+#h #o #G #T #l @lsx_intro
#X #H #HT lapply (lpx_inv_atom1 … H) -H
#H destruct elim HT -HT //
qed.
-lemma lsx_sort: ∀h,g,G,L,l,k. G ⊢ ⬊*[h, g, ⋆k, l] L.
-#h #g #G #L1 #l #k @lsx_intro
+lemma lsx_sort: ∀h,o,G,L,l,s. G ⊢ ⬊*[h, o, ⋆s, l] L.
+#h #o #G #L1 #l #s @lsx_intro
#L2 #HL12 #H elim H -H
/3 width=4 by lpx_fwd_length, lleq_sort/
qed.
-lemma lsx_gref: ∀h,g,G,L,l,p. G ⊢ ⬊*[h, g, §p, l] L.
-#h #g #G #L1 #l #p @lsx_intro
+lemma lsx_gref: ∀h,o,G,L,l,p. G ⊢ ⬊*[h, o, §p, l] L.
+#h #o #G #L1 #l #p @lsx_intro
#L2 #HL12 #H elim H -H
/3 width=4 by lpx_fwd_length, lleq_gref/
qed.
-lemma lsx_ge_up: ∀h,g,G,L,T,U,lt,l,m. lt ≤ yinj l + yinj m →
- ⬆[l, m] T ≡ U → G ⊢ ⬊*[h, g, U, lt] L → G ⊢ ⬊*[h, g, U, l] L.
-#h #g #G #L #T #U #lt #l #m #Hltlm #HTU #H @(lsx_ind … H) -L
+lemma lsx_ge_up: ∀h,o,G,L,T,U,lt,l,k. lt ≤ yinj l + yinj k →
+ ⬆[l, k] T ≡ U → G ⊢ ⬊*[h, o, U, lt] L → G ⊢ ⬊*[h, o, U, l] L.
+#h #o #G #L #T #U #lt #l #k #Hltlm #HTU #H @(lsx_ind … H) -L
/5 width=7 by lsx_intro, lleq_ge_up/
qed-.
-lemma lsx_ge: ∀h,g,G,L,T,l1,l2. l1 ≤ l2 →
- G ⊢ ⬊*[h, g, T, l1] L → G ⊢ ⬊*[h, g, T, l2] L.
-#h #g #G #L #T #l1 #l2 #Hl12 #H @(lsx_ind … H) -L
+lemma lsx_ge: ∀h,o,G,L,T,l1,l2. l1 ≤ l2 →
+ G ⊢ ⬊*[h, o, T, l1] L → G ⊢ ⬊*[h, o, T, l2] L.
+#h #o #G #L #T #l1 #l2 #Hl12 #H @(lsx_ind … H) -L
/5 width=7 by lsx_intro, lleq_ge/
qed-.
(* Basic forward lemmas *****************************************************)
-lemma lsx_fwd_bind_sn: ∀h,g,a,I,G,L,V,T,l. G ⊢ ⬊*[h, g, ⓑ{a,I}V.T, l] L →
- G ⊢ ⬊*[h, g, V, l] L.
-#h #g #a #I #G #L #V #T #l #H @(lsx_ind … H) -L
+lemma lsx_fwd_bind_sn: ∀h,o,a,I,G,L,V,T,l. G ⊢ ⬊*[h, o, ⓑ{a,I}V.T, l] L →
+ G ⊢ ⬊*[h, o, V, l] L.
+#h #o #a #I #G #L #V #T #l #H @(lsx_ind … H) -L
#L1 #_ #IHL1 @lsx_intro
#L2 #HL12 #HV @IHL1 /3 width=4 by lleq_fwd_bind_sn/
qed-.
-lemma lsx_fwd_flat_sn: ∀h,g,I,G,L,V,T,l. G ⊢ ⬊*[h, g, ⓕ{I}V.T, l] L →
- G ⊢ ⬊*[h, g, V, l] L.
-#h #g #I #G #L #V #T #l #H @(lsx_ind … H) -L
+lemma lsx_fwd_flat_sn: ∀h,o,I,G,L,V,T,l. G ⊢ ⬊*[h, o, ⓕ{I}V.T, l] L →
+ G ⊢ ⬊*[h, o, V, l] L.
+#h #o #I #G #L #V #T #l #H @(lsx_ind … H) -L
#L1 #_ #IHL1 @lsx_intro
#L2 #HL12 #HV @IHL1 /3 width=3 by lleq_fwd_flat_sn/
qed-.
-lemma lsx_fwd_flat_dx: ∀h,g,I,G,L,V,T,l. G ⊢ ⬊*[h, g, ⓕ{I}V.T, l] L →
- G ⊢ ⬊*[h, g, T, l] L.
-#h #g #I #G #L #V #T #l #H @(lsx_ind … H) -L
+lemma lsx_fwd_flat_dx: ∀h,o,I,G,L,V,T,l. G ⊢ ⬊*[h, o, ⓕ{I}V.T, l] L →
+ G ⊢ ⬊*[h, o, T, l] L.
+#h #o #I #G #L #V #T #l #H @(lsx_ind … H) -L
#L1 #_ #IHL1 @lsx_intro
#L2 #HL12 #HV @IHL1 /3 width=3 by lleq_fwd_flat_dx/
qed-.
-lemma lsx_fwd_pair_sn: ∀h,g,I,G,L,V,T,l. G ⊢ ⬊*[h, g, ②{I}V.T, l] L →
- G ⊢ ⬊*[h, g, V, l] L.
-#h #g * /2 width=4 by lsx_fwd_bind_sn, lsx_fwd_flat_sn/
+lemma lsx_fwd_pair_sn: ∀h,o,I,G,L,V,T,l. G ⊢ ⬊*[h, o, ②{I}V.T, l] L →
+ G ⊢ ⬊*[h, o, V, l] L.
+#h #o * /2 width=4 by lsx_fwd_bind_sn, lsx_fwd_flat_sn/
qed-.
(* Basic inversion lemmas ***************************************************)
-lemma lsx_inv_flat: ∀h,g,I,G,L,V,T,l. G ⊢ ⬊*[h, g, ⓕ{I}V.T, l] L →
- G ⊢ ⬊*[h, g, V, l] L ∧ G ⊢ ⬊*[h, g, T, l] L.
+lemma lsx_inv_flat: ∀h,o,I,G,L,V,T,l. G ⊢ ⬊*[h, o, ⓕ{I}V.T, l] L →
+ G ⊢ ⬊*[h, o, V, l] L ∧ G ⊢ ⬊*[h, o, T, l] L.
/3 width=3 by lsx_fwd_flat_sn, lsx_fwd_flat_dx, conj/ qed-.
(* alternative definition of lsx *)
definition lsxa: ∀h. sd h → relation4 ynat term genv lenv ≝
- λh,g,l,T,G. SN … (lpxs h g G) (lleq l T).
+ λh,o,l,T,G. SN … (lpxs h o G) (lleq l T).
interpretation
"extended strong normalization (local environment) alternative"
- 'SNAlt h g l T G L = (lsxa h g T l G L).
+ 'SNAlt h o l T G L = (lsxa h o T l G L).
(* Basic eliminators ********************************************************)
-lemma lsxa_ind: ∀h,g,G,T,l. ∀R:predicate lenv.
- (∀L1. G ⊢ ⬊⬊*[h, g, T, l] L1 →
- (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → (L1 ≡[T, l] L2 → ⊥) → R L2) →
+lemma lsxa_ind: ∀h,o,G,T,l. ∀R:predicate lenv.
+ (∀L1. G ⊢ ⬊⬊*[h, o, T, l] L1 →
+ (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) → R L2) →
R L1
) →
- ∀L. G ⊢ ⬊⬊*[h, g, T, l] L → R L.
-#h #g #G #T #l #R #H0 #L1 #H elim H -L1
+ ∀L. G ⊢ ⬊⬊*[h, o, T, l] L → R L.
+#h #o #G #T #l #R #H0 #L1 #H elim H -L1
/5 width=1 by lleq_sym, SN_intro/
qed-.
(* Basic properties *********************************************************)
-lemma lsxa_intro: ∀h,g,G,L1,T,l.
- (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊⬊*[h, g, T, l] L2) →
- G ⊢ ⬊⬊*[h, g, T, l] L1.
+lemma lsxa_intro: ∀h,o,G,L1,T,l.
+ (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊⬊*[h, o, T, l] L2) →
+ G ⊢ ⬊⬊*[h, o, T, l] L1.
/5 width=1 by lleq_sym, SN_intro/ qed.
-fact lsxa_intro_aux: ∀h,g,G,L1,T,l.
- (∀L,L2. ⦃G, L⦄ ⊢ ➡*[h, g] L2 → L1 ≡[T, l] L → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊⬊*[h, g, T, l] L2) →
- G ⊢ ⬊⬊*[h, g, T, l] L1.
+fact lsxa_intro_aux: ∀h,o,G,L1,T,l.
+ (∀L,L2. ⦃G, L⦄ ⊢ ➡*[h, o] L2 → L1 ≡[T, l] L → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊⬊*[h, o, T, l] L2) →
+ G ⊢ ⬊⬊*[h, o, T, l] L1.
/4 width=3 by lsxa_intro/ qed-.
-lemma lsxa_lleq_trans: ∀h,g,T,G,L1,l. G ⊢ ⬊⬊*[h, g, T, l] L1 →
- ∀L2. L1 ≡[T, l] L2 → G ⊢ ⬊⬊*[h, g, T, l] L2.
-#h #g #T #G #L1 #l #H @(lsxa_ind … H) -L1
+lemma lsxa_lleq_trans: ∀h,o,T,G,L1,l. G ⊢ ⬊⬊*[h, o, T, l] L1 →
+ ∀L2. L1 ≡[T, l] L2 → G ⊢ ⬊⬊*[h, o, T, l] L2.
+#h #o #T #G #L1 #l #H @(lsxa_ind … H) -L1
#L1 #_ #IHL1 #L2 #HL12 @lsxa_intro
#K2 #HLK2 #HnLK2 elim (lleq_lpxs_trans … HLK2 … HL12) -HLK2
/5 width=4 by lleq_canc_sn, lleq_trans/
qed-.
-lemma lsxa_lpxs_trans: ∀h,g,T,G,L1,l. G ⊢ ⬊⬊*[h, g, T, l] L1 →
- ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → G ⊢ ⬊⬊*[h, g, T, l] L2.
-#h #g #T #G #L1 #l #H @(lsxa_ind … H) -L1 #L1 #HL1 #IHL1 #L2 #HL12
+lemma lsxa_lpxs_trans: ∀h,o,T,G,L1,l. G ⊢ ⬊⬊*[h, o, T, l] L1 →
+ ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → G ⊢ ⬊⬊*[h, o, T, l] L2.
+#h #o #T #G #L1 #l #H @(lsxa_ind … H) -L1 #L1 #HL1 #IHL1 #L2 #HL12
elim (lleq_dec T L1 L2 l) /3 width=4 by lsxa_lleq_trans/
qed-.
-lemma lsxa_intro_lpx: ∀h,g,G,L1,T,l.
- (∀L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊⬊*[h, g, T, l] L2) →
- G ⊢ ⬊⬊*[h, g, T, l] L1.
-#h #g #G #L1 #T #l #IH @lsxa_intro_aux
+lemma lsxa_intro_lpx: ∀h,o,G,L1,T,l.
+ (∀L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊⬊*[h, o, T, l] L2) →
+ G ⊢ ⬊⬊*[h, o, T, l] L1.
+#h #o #G #L1 #T #l #IH @lsxa_intro_aux
#L #L2 #H @(lpxs_ind_dx … H) -L
[ #H destruct #H elim H //
| #L0 #L elim (lleq_dec T L1 L l) /3 width=1 by/
(* Main properties **********************************************************)
-theorem lsx_lsxa: ∀h,g,G,L,T,l. G ⊢ ⬊*[h, g, T, l] L → G ⊢ ⬊⬊*[h, g, T, l] L.
-#h #g #G #L #T #l #H @(lsx_ind … H) -L
+theorem lsx_lsxa: ∀h,o,G,L,T,l. G ⊢ ⬊*[h, o, T, l] L → G ⊢ ⬊⬊*[h, o, T, l] L.
+#h #o #G #L #T #l #H @(lsx_ind … H) -L
/4 width=1 by lsxa_intro_lpx/
qed.
(* Main inversion lemmas ****************************************************)
-theorem lsxa_inv_lsx: ∀h,g,G,L,T,l. G ⊢ ⬊⬊*[h, g, T, l] L → G ⊢ ⬊*[h, g, T, l] L.
-#h #g #G #L #T #l #H @(lsxa_ind … H) -L
+theorem lsxa_inv_lsx: ∀h,o,G,L,T,l. G ⊢ ⬊⬊*[h, o, T, l] L → G ⊢ ⬊*[h, o, T, l] L.
+#h #o #G #L #T #l #H @(lsxa_ind … H) -L
/4 width=1 by lsx_intro, lpx_lpxs/
qed-.
(* Advanced properties ******************************************************)
-lemma lsx_intro_alt: ∀h,g,G,L1,T,l.
- (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊*[h, g, T, l] L2) →
- G ⊢ ⬊*[h, g, T, l] L1.
+lemma lsx_intro_alt: ∀h,o,G,L1,T,l.
+ (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) → G ⊢ ⬊*[h, o, T, l] L2) →
+ G ⊢ ⬊*[h, o, T, l] L1.
/6 width=1 by lsxa_inv_lsx, lsx_lsxa, lsxa_intro/ qed.
-lemma lsx_lpxs_trans: ∀h,g,G,L1,T,l. G ⊢ ⬊*[h, g, T, l] L1 →
- ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → G ⊢ ⬊*[h, g, T, l] L2.
+lemma lsx_lpxs_trans: ∀h,o,G,L1,T,l. G ⊢ ⬊*[h, o, T, l] L1 →
+ ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → G ⊢ ⬊*[h, o, T, l] L2.
/4 width=3 by lsxa_inv_lsx, lsx_lsxa, lsxa_lpxs_trans/ qed-.
(* Advanced eliminators *****************************************************)
-lemma lsx_ind_alt: ∀h,g,G,T,l. ∀R:predicate lenv.
- (∀L1. G ⊢ ⬊*[h, g, T, l] L1 →
- (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → (L1 ≡[T, l] L2 → ⊥) → R L2) →
+lemma lsx_ind_alt: ∀h,o,G,T,l. ∀R:predicate lenv.
+ (∀L1. G ⊢ ⬊*[h, o, T, l] L1 →
+ (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → (L1 ≡[T, l] L2 → ⊥) → R L2) →
R L1
) →
- ∀L. G ⊢ ⬊*[h, g, T, l] L → R L.
-#h #g #G #T #l #R #IH #L #H @(lsxa_ind h g G T l … L)
+ ∀L. G ⊢ ⬊*[h, o, T, l] L → R L.
+#h #o #G #T #l #R #IH #L #H @(lsxa_ind h o G T l … L)
/4 width=1 by lsxa_inv_lsx, lsx_lsxa/
qed-.
(* Advanced properties ******************************************************)
-lemma lsx_lref_be_lpxs: ∀h,g,I,G,K1,V,i,l. l ≤ yinj i → ⦃G, K1⦄ ⊢ ⬊*[h, g] V →
- ∀K2. G ⊢ ⬊*[h, g, V, 0] K2 → ⦃G, K1⦄ ⊢ ➡*[h, g] K2 →
- ∀L2. ⬇[i] L2 ≡ K2.ⓑ{I}V → G ⊢ ⬊*[h, g, #i, l] L2.
-#h #g #I #G #K1 #V #i #l #Hli #H @(csx_ind_alt … H) -V
+lemma lsx_lref_be_lpxs: ∀h,o,I,G,K1,V,i,l. l ≤ yinj i → ⦃G, K1⦄ ⊢ ⬊*[h, o] V →
+ ∀K2. G ⊢ ⬊*[h, o, V, 0] K2 → ⦃G, K1⦄ ⊢ ➡*[h, o] K2 →
+ ∀L2. ⬇[i] L2 ≡ K2.ⓑ{I}V → G ⊢ ⬊*[h, o, #i, l] L2.
+#h #o #I #G #K1 #V #i #l #Hli #H @(csx_ind_alt … H) -V
#V0 #_ #IHV0 #K2 #H @(lsx_ind … H) -K2
#K0 #HK0 #IHK0 #HK10 #L0 #HLK0 @lsx_intro
#L2 #HL02 #HnL02 elim (lpx_drop_conf … HLK0 … HL02) -HL02
]
qed.
-lemma lsx_lref_be: ∀h,g,I,G,K,V,i,l. l ≤ yinj i → ⦃G, K⦄ ⊢ ⬊*[h, g] V →
- G ⊢ ⬊*[h, g, V, 0] K →
- ∀L. ⬇[i] L ≡ K.ⓑ{I}V → G ⊢ ⬊*[h, g, #i, l] L.
+lemma lsx_lref_be: ∀h,o,I,G,K,V,i,l. l ≤ yinj i → ⦃G, K⦄ ⊢ ⬊*[h, o] V →
+ G ⊢ ⬊*[h, o, V, 0] K →
+ ∀L. ⬇[i] L ≡ K.ⓑ{I}V → G ⊢ ⬊*[h, o, #i, l] L.
/2 width=8 by lsx_lref_be_lpxs/ qed.
(* Main properties **********************************************************)
-theorem csx_lsx: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → ∀l. G ⊢ ⬊*[h, g, T, l] L.
-#h #g #G #L #T @(fqup_wf_ind_eq … G L T) -G -L -T
+theorem csx_lsx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ∀l. G ⊢ ⬊*[h, o, T, l] L.
+#h #o #G #L #T @(fqup_wf_ind_eq … G L T) -G -L -T
#Z #Y #X #IH #G #L * * //
[ #i #HG #HL #HT #H #l destruct
elim (lt_or_ge i (|L|)) /2 width=1 by lsx_lref_free/
(* Advanced properties ******************************************************)
-lemma lsx_lref_free: ∀h,g,G,L,l,i. |L| ≤ i → G ⊢ ⬊*[h, g, #i, l] L.
-#h #g #G #L1 #l #i #HL1 @lsx_intro
+lemma lsx_lref_free: ∀h,o,G,L,l,i. |L| ≤ i → G ⊢ ⬊*[h, o, #i, l] L.
+#h #o #G #L1 #l #i #HL1 @lsx_intro
#L2 #HL12 #H elim H -H
/4 width=6 by lpx_fwd_length, lleq_free, le_repl_sn_conf_aux/
qed.
-lemma lsx_lref_skip: ∀h,g,G,L,l,i. yinj i < l → G ⊢ ⬊*[h, g, #i, l] L.
-#h #g #G #L1 #l #i #HL1 @lsx_intro
+lemma lsx_lref_skip: ∀h,o,G,L,l,i. yinj i < l → G ⊢ ⬊*[h, o, #i, l] L.
+#h #o #G #L1 #l #i #HL1 @lsx_intro
#L2 #HL12 #H elim H -H
/3 width=4 by lpx_fwd_length, lleq_skip/
qed.
(* Advanced forward lemmas **************************************************)
-lemma lsx_fwd_lref_be: ∀h,g,I,G,L,l,i. l ≤ yinj i → G ⊢ ⬊*[h, g, #i, l] L →
- ∀K,V. ⬇[i] L ≡ K.ⓑ{I}V → G ⊢ ⬊*[h, g, V, 0] K.
-#h #g #I #G #L #l #i #Hli #H @(lsx_ind … H) -L
+lemma lsx_fwd_lref_be: ∀h,o,I,G,L,l,i. l ≤ yinj i → G ⊢ ⬊*[h, o, #i, l] L →
+ ∀K,V. ⬇[i] L ≡ K.ⓑ{I}V → G ⊢ ⬊*[h, o, V, 0] K.
+#h #o #I #G #L #l #i #Hli #H @(lsx_ind … H) -L
#L1 #_ #IHL1 #K1 #V #HLK1 @lsx_intro
#K2 #HK12 #HnK12 lapply (drop_fwd_drop2 … HLK1)
#H2LK1 elim (drop_lpx_trans … H2LK1 … HK12) -H2LK1 -HK12
(* Properties on relocation *************************************************)
-lemma lsx_lift_le: ∀h,g,G,K,T,U,lt,l,m. lt ≤ yinj l →
- ⬆[l, m] T ≡ U → G ⊢ ⬊*[h, g, T, lt] K →
- ∀L. ⬇[Ⓕ, l, m] L ≡ K → G ⊢ ⬊*[h, g, U, lt] L.
-#h #g #G #K #T #U #lt #l #m #Hltl #HTU #H @(lsx_ind … H) -K
+lemma lsx_lift_le: ∀h,o,G,K,T,U,lt,l,k. lt ≤ yinj l →
+ ⬆[l, k] T ≡ U → G ⊢ ⬊*[h, o, T, lt] K →
+ ∀L. ⬇[Ⓕ, l, k] L ≡ K → G ⊢ ⬊*[h, o, U, lt] L.
+#h #o #G #K #T #U #lt #l #k #Hltl #HTU #H @(lsx_ind … H) -K
#K1 #_ #IHK1 #L1 #HLK1 @lsx_intro
#L2 #HL12 #HnU elim (lpx_drop_conf … HLK1 … HL12) -HL12
/4 width=10 by lleq_lift_le/
qed-.
-lemma lsx_lift_ge: ∀h,g,G,K,T,U,lt,l,m. yinj l ≤ lt →
- ⬆[l, m] T ≡ U → G ⊢ ⬊*[h, g, T, lt] K →
- ∀L. ⬇[Ⓕ, l, m] L ≡ K → G ⊢ ⬊*[h, g, U, lt + m] L.
-#h #g #G #K #T #U #lt #l #m #Hllt #HTU #H @(lsx_ind … H) -K
+lemma lsx_lift_ge: ∀h,o,G,K,T,U,lt,l,k. yinj l ≤ lt →
+ ⬆[l, k] T ≡ U → G ⊢ ⬊*[h, o, T, lt] K →
+ ∀L. ⬇[Ⓕ, l, k] L ≡ K → G ⊢ ⬊*[h, o, U, lt + k] L.
+#h #o #G #K #T #U #lt #l #k #Hllt #HTU #H @(lsx_ind … H) -K
#K1 #_ #IHK1 #L1 #HLK1 @lsx_intro
#L2 #HL12 #HnU elim (lpx_drop_conf … HLK1 … HL12) -HL12
/4 width=9 by lleq_lift_ge/
(* Inversion lemmas on relocation *******************************************)
-lemma lsx_inv_lift_le: ∀h,g,G,L,T,U,lt,l,m. lt ≤ yinj l →
- ⬆[l, m] T ≡ U → G ⊢ ⬊*[h, g, U, lt] L →
- ∀K. ⬇[Ⓕ, l, m] L ≡ K → G ⊢ ⬊*[h, g, T, lt] K.
-#h #g #G #L #T #U #lt #l #m #Hltl #HTU #H @(lsx_ind … H) -L
+lemma lsx_inv_lift_le: ∀h,o,G,L,T,U,lt,l,k. lt ≤ yinj l →
+ ⬆[l, k] T ≡ U → G ⊢ ⬊*[h, o, U, lt] L →
+ ∀K. ⬇[Ⓕ, l, k] L ≡ K → G ⊢ ⬊*[h, o, T, lt] K.
+#h #o #G #L #T #U #lt #l #k #Hltl #HTU #H @(lsx_ind … H) -L
#L1 #_ #IHL1 #K1 #HLK1 @lsx_intro
#K2 #HK12 #HnT elim (drop_lpx_trans … HLK1 … HK12) -HK12
/4 width=10 by lleq_inv_lift_le/
qed-.
-lemma lsx_inv_lift_be: ∀h,g,G,L,T,U,lt,l,m. yinj l ≤ lt → lt ≤ l + m →
- ⬆[l, m] T ≡ U → G ⊢ ⬊*[h, g, U, lt] L →
- ∀K. ⬇[Ⓕ, l, m] L ≡ K → G ⊢ ⬊*[h, g, T, l] K.
-#h #g #G #L #T #U #lt #l #m #Hllt #Hltlm #HTU #H @(lsx_ind … H) -L
+lemma lsx_inv_lift_be: ∀h,o,G,L,T,U,lt,l,k. yinj l ≤ lt → lt ≤ l + k →
+ ⬆[l, k] T ≡ U → G ⊢ ⬊*[h, o, U, lt] L →
+ ∀K. ⬇[Ⓕ, l, k] L ≡ K → G ⊢ ⬊*[h, o, T, l] K.
+#h #o #G #L #T #U #lt #l #k #Hllt #Hltlm #HTU #H @(lsx_ind … H) -L
#L1 #_ #IHL1 #K1 #HLK1 @lsx_intro
#K2 #HK12 #HnT elim (drop_lpx_trans … HLK1 … HK12) -HK12
/4 width=11 by lleq_inv_lift_be/
qed-.
-lemma lsx_inv_lift_ge: ∀h,g,G,L,T,U,lt,l,m. yinj l + yinj m ≤ lt →
- ⬆[l, m] T ≡ U → G ⊢ ⬊*[h, g, U, lt] L →
- ∀K. ⬇[Ⓕ, l, m] L ≡ K → G ⊢ ⬊*[h, g, T, lt-m] K.
-#h #g #G #L #T #U #lt #l #m #Hlmlt #HTU #H @(lsx_ind … H) -L
+lemma lsx_inv_lift_ge: ∀h,o,G,L,T,U,lt,l,k. yinj l + yinj k ≤ lt →
+ ⬆[l, k] T ≡ U → G ⊢ ⬊*[h, o, U, lt] L →
+ ∀K. ⬇[Ⓕ, l, k] L ≡ K → G ⊢ ⬊*[h, o, T, lt-k] K.
+#h #o #G #L #T #U #lt #l #k #Hlmlt #HTU #H @(lsx_ind … H) -L
#L1 #_ #IHL1 #K1 #HLK1 @lsx_intro
#K2 #HK12 #HnT elim (drop_lpx_trans … HLK1 … HK12) -HK12
/4 width=9 by lleq_inv_lift_ge/
(* Advanced properties ******************************************************)
-lemma lsx_lleq_trans: ∀h,g,T,G,L1,l. G ⊢ ⬊*[h, g, T, l] L1 →
- ∀L2. L1 ≡[T, l] L2 → G ⊢ ⬊*[h, g, T, l] L2.
-#h #g #T #G #L1 #l #H @(lsx_ind … H) -L1
+lemma lsx_lleq_trans: ∀h,o,T,G,L1,l. G ⊢ ⬊*[h, o, T, l] L1 →
+ ∀L2. L1 ≡[T, l] L2 → G ⊢ ⬊*[h, o, T, l] L2.
+#h #o #T #G #L1 #l #H @(lsx_ind … H) -L1
#L1 #_ #IHL1 #L2 #HL12 @lsx_intro
#K2 #HLK2 #HnLK2 elim (lleq_lpx_trans … HLK2 … HL12) -HLK2
/5 width=4 by lleq_canc_sn, lleq_trans/
qed-.
-lemma lsx_lpx_trans: ∀h,g,T,G,L1,l. G ⊢ ⬊*[h, g, T, l] L1 →
- ∀L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 → G ⊢ ⬊*[h, g, T, l] L2.
-#h #g #T #G #L1 #l #H @(lsx_ind … H) -L1 #L1 #HL1 #IHL1 #L2 #HL12
+lemma lsx_lpx_trans: ∀h,o,T,G,L1,l. G ⊢ ⬊*[h, o, T, l] L1 →
+ ∀L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → G ⊢ ⬊*[h, o, T, l] L2.
+#h #o #T #G #L1 #l #H @(lsx_ind … H) -L1 #L1 #HL1 #IHL1 #L2 #HL12
elim (lleq_dec T L1 L2 l) /3 width=4 by lsx_lleq_trans/
qed-.
-lemma lsx_lreq_conf: ∀h,g,G,L1,T,l. G ⊢ ⬊*[h, g, T, l] L1 →
- ∀L2. L1 ⩬[l, ∞] L2 → G ⊢ ⬊*[h, g, T, l] L2.
-#h #g #G #L1 #T #l #H @(lsx_ind … H) -L1
+lemma lsx_lreq_conf: ∀h,o,G,L1,T,l. G ⊢ ⬊*[h, o, T, l] L1 →
+ ∀L2. L1 ⩬[l, ∞] L2 → G ⊢ ⬊*[h, o, T, l] L2.
+#h #o #G #L1 #T #l #H @(lsx_ind … H) -L1
#L1 #_ #IHL1 #L2 #HL12 @lsx_intro
#L3 #HL23 #HnL23 elim (lreq_lpx_trans_lleq … HL12 … HL23) -HL12 -HL23
#L0 #HL03 #HL10 #H elim (H T) -H /4 width=4 by/
(* Advanced forward lemmas **************************************************)
-lemma lsx_fwd_bind_dx: ∀h,g,a,I,G,L,V,T,l. G ⊢ ⬊*[h, g, ⓑ{a,I}V.T, l] L →
- G ⊢ ⬊*[h, g, T, ⫯l] L.ⓑ{I}V.
-#h #g #a #I #G #L #V1 #T #l #H @(lsx_ind … H) -L
+lemma lsx_fwd_bind_dx: ∀h,o,a,I,G,L,V,T,l. G ⊢ ⬊*[h, o, ⓑ{a,I}V.T, l] L →
+ G ⊢ ⬊*[h, o, T, ⫯l] L.ⓑ{I}V.
+#h #o #a #I #G #L #V1 #T #l #H @(lsx_ind … H) -L
#L1 #_ #IHL1 @lsx_intro
#Y #H #HT elim (lpx_inv_pair1 … H) -H
#L2 #V2 #HL12 #_ #H destruct
(* Advanced inversion lemmas ************************************************)
-lemma lsx_inv_bind: ∀h,g,a,I,G,L,V,T,l. G ⊢ ⬊*[h, g, ⓑ{a, I}V.T, l] L →
- G ⊢ ⬊*[h, g, V, l] L ∧ G ⊢ ⬊*[h, g, T, ⫯l] L.ⓑ{I}V.
+lemma lsx_inv_bind: ∀h,o,a,I,G,L,V,T,l. G ⊢ ⬊*[h, o, ⓑ{a, I}V.T, l] L →
+ G ⊢ ⬊*[h, o, V, l] L ∧ G ⊢ ⬊*[h, o, T, ⫯l] L.ⓑ{I}V.
/3 width=4 by lsx_fwd_bind_sn, lsx_fwd_bind_dx, conj/ qed-.
(* Advanced properties ******************************************************)
-fact lsx_bind_lpxs_aux: ∀h,g,a,I,G,L1,V,l. G ⊢ ⬊*[h, g, V, l] L1 →
- ∀Y,T. G ⊢ ⬊*[h, g, T, ⫯l] Y →
- ∀L2. Y = L2.ⓑ{I}V → ⦃G, L1⦄ ⊢ ➡*[h, g] L2 →
- G ⊢ ⬊*[h, g, ⓑ{a,I}V.T, l] L2.
-#h #g #a #I #G #L1 #V #l #H @(lsx_ind_alt … H) -L1
+fact lsx_bind_lpxs_aux: ∀h,o,a,I,G,L1,V,l. G ⊢ ⬊*[h, o, V, l] L1 →
+ ∀Y,T. G ⊢ ⬊*[h, o, T, ⫯l] Y →
+ ∀L2. Y = L2.ⓑ{I}V → ⦃G, L1⦄ ⊢ ➡*[h, o] L2 →
+ G ⊢ ⬊*[h, o, ⓑ{a,I}V.T, l] L2.
+#h #o #a #I #G #L1 #V #l #H @(lsx_ind_alt … H) -L1
#L1 #HL1 #IHL1 #Y #T #H @(lsx_ind_alt … H) -Y
#Y #HY #IHY #L2 #H #HL12 destruct @lsx_intro_alt
#L0 #HL20 lapply (lpxs_trans … HL12 … HL20)
]
qed-.
-lemma lsx_bind: ∀h,g,a,I,G,L,V,l. G ⊢ ⬊*[h, g, V, l] L →
- ∀T. G ⊢ ⬊*[h, g, T, ⫯l] L.ⓑ{I}V →
- G ⊢ ⬊*[h, g, ⓑ{a,I}V.T, l] L.
+lemma lsx_bind: ∀h,o,a,I,G,L,V,l. G ⊢ ⬊*[h, o, V, l] L →
+ ∀T. G ⊢ ⬊*[h, o, T, ⫯l] L.ⓑ{I}V →
+ G ⊢ ⬊*[h, o, ⓑ{a,I}V.T, l] L.
/2 width=3 by lsx_bind_lpxs_aux/ qed.
-lemma lsx_flat_lpxs: ∀h,g,I,G,L1,V,l. G ⊢ ⬊*[h, g, V, l] L1 →
- ∀L2,T. G ⊢ ⬊*[h, g, T, l] L2 → ⦃G, L1⦄ ⊢ ➡*[h, g] L2 →
- G ⊢ ⬊*[h, g, ⓕ{I}V.T, l] L2.
-#h #g #I #G #L1 #V #l #H @(lsx_ind_alt … H) -L1
+lemma lsx_flat_lpxs: ∀h,o,I,G,L1,V,l. G ⊢ ⬊*[h, o, V, l] L1 →
+ ∀L2,T. G ⊢ ⬊*[h, o, T, l] L2 → ⦃G, L1⦄ ⊢ ➡*[h, o] L2 →
+ G ⊢ ⬊*[h, o, ⓕ{I}V.T, l] L2.
+#h #o #I #G #L1 #V #l #H @(lsx_ind_alt … H) -L1
#L1 #HL1 #IHL1 #L2 #T #H @(lsx_ind_alt … H) -L2
#L2 #HL2 #IHL2 #HL12 @lsx_intro_alt
#L0 #HL20 lapply (lpxs_trans … HL12 … HL20)
]
qed-.
-lemma lsx_flat: ∀h,g,I,G,L,V,l. G ⊢ ⬊*[h, g, V, l] L →
- ∀T. G ⊢ ⬊*[h, g, T, l] L → G ⊢ ⬊*[h, g, ⓕ{I}V.T, l] L.
+lemma lsx_flat: ∀h,o,I,G,L,V,l. G ⊢ ⬊*[h, o, V, l] L →
+ ∀T. G ⊢ ⬊*[h, o, T, l] L → G ⊢ ⬊*[h, o, ⓕ{I}V.T, l] L.
/2 width=3 by lsx_flat_lpxs/ qed.
(* STRATIFIED DECOMPOSED PARALLEL COMPUTATION ON TERMS **********************)
definition scpds: ∀h. sd h → nat → relation4 genv lenv term term ≝
- λh,g,d2,G,L,T1,T2.
- ∃∃T,d1. d2 ≤ d1 & ⦃G, L⦄ ⊢ T1 ▪[h, g] d1 & ⦃G, L⦄ ⊢ T1 •*[h, d2] T & ⦃G, L⦄ ⊢ T ➡* T2.
+ λh,o,d2,G,L,T1,T2.
+ ∃∃T,d1. d2 ≤ d1 & ⦃G, L⦄ ⊢ T1 ▪[h, o] d1 & ⦃G, L⦄ ⊢ T1 •*[h, d2] T & ⦃G, L⦄ ⊢ T ➡* T2.
interpretation "stratified decomposed parallel computation (term)"
- 'DPRedStar h g d G L T1 T2 = (scpds h g d G L T1 T2).
+ 'DPRedStar h o d G L T1 T2 = (scpds h o d G L T1 T2).
(* Basic properties *********************************************************)
-lemma sta_cprs_scpds: ∀h,g,G,L,T1,T,T2,d. ⦃G, L⦄ ⊢ T1 ▪[h, g] d+1 → ⦃G, L⦄ ⊢ T1 •*[h, 1] T →
- ⦃G, L⦄ ⊢ T ➡* T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g, 1] T2.
+lemma sta_cprs_scpds: ∀h,o,G,L,T1,T,T2,d. ⦃G, L⦄ ⊢ T1 ▪[h, o] d+1 → ⦃G, L⦄ ⊢ T1 •*[h, 1] T →
+ ⦃G, L⦄ ⊢ T ➡* T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, o, 1] T2.
/2 width=6 by ex4_2_intro/ qed.
-lemma lstas_scpds: ∀h,g,G,L,T1,T2,d1. ⦃G, L⦄ ⊢ T1 ▪[h, g] d1 →
- ∀d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d2] T2.
+lemma lstas_scpds: ∀h,o,G,L,T1,T2,d1. ⦃G, L⦄ ⊢ T1 ▪[h, o] d1 →
+ ∀d2. d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 •*[h, d2] T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d2] T2.
/2 width=6 by ex4_2_intro/ qed.
-lemma scpds_strap1: ∀h,g,G,L,T1,T,T2,d.
- ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d] T → ⦃G, L⦄ ⊢ T ➡ T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d] T2.
-#h #g #G #L #T1 #T #T2 #d * /3 width=8 by cprs_strap1, ex4_2_intro/
+lemma scpds_strap1: ∀h,o,G,L,T1,T,T2,d.
+ ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d] T → ⦃G, L⦄ ⊢ T ➡ T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d] T2.
+#h #o #G #L #T1 #T #T2 #d * /3 width=8 by cprs_strap1, ex4_2_intro/
qed.
(* Basic forward lemmas *****************************************************)
-lemma scpds_fwd_cprs: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 •*➡*[h, g, 0] T2 →
+lemma scpds_fwd_cprs: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 •*➡*[h, o, 0] T2 →
⦃G, L⦄ ⊢ T1 ➡* T2.
-#h #g #G #L #T1 #T2 * /3 width=3 by cprs_strap2, lstas_cpr/
+#h #o #G #L #T1 #T2 * /3 width=3 by cprs_strap2, lstas_cpr/
qed-.
(* Properties on atomic arity assignment for terms **************************)
-lemma scpds_aaa_conf: ∀h,g,G,L,d. Conf3 … (aaa G L) (scpds h g d G L).
-#h #g #G #L #d #A #T #HT #U * /3 width=6 by lstas_aaa_conf, cprs_aaa_conf/
+lemma scpds_aaa_conf: ∀h,o,G,L,d. Conf3 … (aaa G L) (scpds h o d G L).
+#h #o #G #L #d #A #T #HT #U * /3 width=6 by lstas_aaa_conf, cprs_aaa_conf/
qed.
(* Relocation properties ****************************************************)
-lemma scpds_lift: ∀h,g,G,d. d_liftable (scpds h g d G).
-#h #g #G #d2 #K #T1 #T2 * #T #d1 #Hd21 #Hd1 #HT1 #HT2 #L #s #l #m
-elim (lift_total T l m)
+lemma scpds_lift: ∀h,o,G,d. d_liftable (scpds h o d G).
+#h #o #G #d2 #K #T1 #T2 * #T #d1 #Hd21 #Hd1 #HT1 #HT2 #L #c #l #k
+elim (lift_total T l k)
/3 width=15 by cprs_lift, da_lift, lstas_lift, ex4_2_intro/
qed.
-lemma scpds_inv_lift1: ∀h,g,G,d. d_deliftable_sn (scpds h g d G).
-#h #g #G #d2 #L #U1 #U2 * #U #d1 #Hd21 #Hd1 #HU1 #HU2 #K #s #l #m #HLK #T1 #HTU1
+lemma scpds_inv_lift1: ∀h,o,G,d. d_deliftable_sn (scpds h o d G).
+#h #o #G #d2 #L #U1 #U2 * #U #d1 #Hd21 #Hd1 #HU1 #HU2 #K #c #l #k #HLK #T1 #HTU1
lapply (da_inv_lift … Hd1 … HLK … HTU1) -Hd1 #Hd1
elim (lstas_inv_lift1 … HU1 … HLK … HTU1) -U1 #T #HTU #HT1
elim (cprs_inv_lift1 … HU2 … HLK … HTU) -U -L
(* Advanced properties ******************************************************)
-lemma scpds_strap2: ∀h,g,G,L,T1,T,T2,d1,d. ⦃G, L⦄ ⊢ T1 ▪[h, g] d1+1 →
- ⦃G, L⦄ ⊢ T1 •*[h, 1] T → ⦃G, L⦄ ⊢ T •*➡*[h, g, d] T2 →
- ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d+1] T2.
-#h #g #G #L #T1 #T #T2 #d1 #d #Hd1 #HT1 *
+lemma scpds_strap2: ∀h,o,G,L,T1,T,T2,d1,d. ⦃G, L⦄ ⊢ T1 ▪[h, o] d1+1 →
+ ⦃G, L⦄ ⊢ T1 •*[h, 1] T → ⦃G, L⦄ ⊢ T •*➡*[h, o, d] T2 →
+ ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d+1] T2.
+#h #o #G #L #T1 #T #T2 #d1 #d #Hd1 #HT1 *
#T0 #d0 #Hd0 #HTd0 #HT0 #HT02
-lapply (lstas_da_conf … HT1 … Hd1) <minus_plus_m_m #HTd1
+lapply (lstas_da_conf … HT1 … Hd1) <minus_plus_k_k #HTd1
lapply (da_mono … HTd0 … HTd1) -HTd0 -HTd1 #H destruct
lapply (lstas_trans … HT1 … HT0) -T >commutative_plus
/3 width=6 by le_S_S, ex4_2_intro/
qed.
-lemma scpds_cprs_trans: ∀h,g,G,L,T1,T,T2,d.
- ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d] T → ⦃G, L⦄ ⊢ T ➡* T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d] T2.
-#h #g #G #L #T1 #T #T2 #d * /3 width=8 by cprs_trans, ex4_2_intro/
+lemma scpds_cprs_trans: ∀h,o,G,L,T1,T,T2,d.
+ ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d] T → ⦃G, L⦄ ⊢ T ➡* T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d] T2.
+#h #o #G #L #T1 #T #T2 #d * /3 width=8 by cprs_trans, ex4_2_intro/
qed-.
-lemma lstas_scpds_trans: ∀h,g,G,L,T1,T,T2,d1,d2,d.
- d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 ▪[h, g] d1 →
- ⦃G, L⦄ ⊢ T1 •*[h, d2] T → ⦃G, L⦄ ⊢ T •*➡*[h, g, d] T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d2+d] T2.
-#h #g #G #L #T1 #T #T2 #d1 #d2 #d #Hd21 #HTd1 #HT1 * #T0 #d0 #Hd0 #HTd0 #HT0 #HT02
+lemma lstas_scpds_trans: ∀h,o,G,L,T1,T,T2,d1,d2,d.
+ d2 ≤ d1 → ⦃G, L⦄ ⊢ T1 ▪[h, o] d1 →
+ ⦃G, L⦄ ⊢ T1 •*[h, d2] T → ⦃G, L⦄ ⊢ T •*➡*[h, o, d] T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d2+d] T2.
+#h #o #G #L #T1 #T #T2 #d1 #d2 #d #Hd21 #HTd1 #HT1 * #T0 #d0 #Hd0 #HTd0 #HT0 #HT02
lapply (lstas_da_conf … HT1 … HTd1) #HTd12
lapply (da_mono … HTd12 … HTd0) -HTd12 -HTd0 #H destruct
lapply (le_minus_to_plus_r … Hd21 Hd0) -Hd21 -Hd0
(* Advanced inversion lemmas ************************************************)
-lemma scpds_inv_abst1: ∀h,g,a,G,L,V1,T1,U2,d. ⦃G, L⦄ ⊢ ⓛ{a}V1.T1 •*➡*[h, g, d] U2 →
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡* V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 •*➡*[h, g, d] T2 &
+lemma scpds_inv_abst1: ∀h,o,a,G,L,V1,T1,U2,d. ⦃G, L⦄ ⊢ ⓛ{a}V1.T1 •*➡*[h, o, d] U2 →
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡* V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 •*➡*[h, o, d] T2 &
U2 = ⓛ{a}V2.T2.
-#h #g #a #G #L #V1 #T1 #U2 #d2 * #X #d1 #Hd21 #Hd1 #H1 #H2
+#h #o #a #G #L #V1 #T1 #U2 #d2 * #X #d1 #Hd21 #Hd1 #H1 #H2
lapply (da_inv_bind … Hd1) -Hd1 #Hd1
elim (lstas_inv_bind1 … H1) -H1 #U #HTU1 #H destruct
elim (cprs_inv_abst1 … H2) -H2 #V2 #T2 #HV12 #HUT2 #H destruct
/3 width=6 by ex4_2_intro, ex3_2_intro/
qed-.
-lemma scpds_inv_abbr_abst: ∀h,g,a1,a2,G,L,V1,W2,T1,T2,d. ⦃G, L⦄ ⊢ ⓓ{a1}V1.T1 •*➡*[h, g, d] ⓛ{a2}W2.T2 →
- ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 •*➡*[h, g, d] T & ⬆[0, 1] ⓛ{a2}W2.T2 ≡ T & a1 = true.
-#h #g #a1 #a2 #G #L #V1 #W2 #T1 #T2 #d2 * #X #d1 #Hd21 #Hd1 #H1 #H2
+lemma scpds_inv_abbr_abst: ∀h,o,a1,a2,G,L,V1,W2,T1,T2,d. ⦃G, L⦄ ⊢ ⓓ{a1}V1.T1 •*➡*[h, o, d] ⓛ{a2}W2.T2 →
+ ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 •*➡*[h, o, d] T & ⬆[0, 1] ⓛ{a2}W2.T2 ≡ T & a1 = true.
+#h #o #a1 #a2 #G #L #V1 #W2 #T1 #T2 #d2 * #X #d1 #Hd21 #Hd1 #H1 #H2
lapply (da_inv_bind … Hd1) -Hd1 #Hd1
elim (lstas_inv_bind1 … H1) -H1 #U1 #HTU1 #H destruct
elim (cprs_inv_abbr1 … H2) -H2 *
]
qed-.
-lemma scpds_inv_lstas_eq: ∀h,g,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d] T2 →
+lemma scpds_inv_lstas_eq: ∀h,o,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d] T2 →
∀T. ⦃G, L⦄ ⊢ T1 •*[h, d] T → ⦃G, L⦄ ⊢ T ➡* T2.
-#h #g #G #L #T1 #T2 #d2 *
+#h #o #G #L #T1 #T2 #d2 *
#T0 #d1 #_ #_ #HT10 #HT02 #T #HT1
lapply (lstas_mono … HT10 … HT1) #H destruct //
qed-.
(* Advanced forward lemmas **************************************************)
-lemma scpds_fwd_cpxs: ∀h,g,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 •*➡*[h, g, d] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2.
-#h #g #G #L #T1 #T2 #d * /3 width=5 by cpxs_trans, lstas_cpxs, cprs_cpxs/
+lemma scpds_fwd_cpxs: ∀h,o,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, o] T2.
+#h #o #G #L #T1 #T2 #d * /3 width=5 by cpxs_trans, lstas_cpxs, cprs_cpxs/
qed-.
(* Main properties **********************************************************)
-theorem scpds_conf_eq: ∀h,g,G,L,T0,T1,d. ⦃G, L⦄ ⊢ T0 •*➡*[h, g, d] T1 →
- ∀T2. ⦃G, L⦄ ⊢ T0 •*➡*[h, g, d] T2 →
+theorem scpds_conf_eq: ∀h,o,G,L,T0,T1,d. ⦃G, L⦄ ⊢ T0 •*➡*[h, o, d] T1 →
+ ∀T2. ⦃G, L⦄ ⊢ T0 •*➡*[h, o, d] T2 →
∃∃T. ⦃G, L⦄ ⊢ T1 ➡* T & ⦃G, L⦄ ⊢ T2 ➡* T.
-#h #g #G #L #T0 #T1 #d0 * #U1 #d1 #_ #_ #H1 #HUT1 #T2 * #U2 #d2 #_ #_ #H2 #HUT2 -d1 -d2
+#h #o #G #L #T0 #T1 #d0 * #U1 #d1 #_ #_ #H1 #HUT1 #T2 * #U2 #d2 #_ #_ #H2 #HUT2 -d1 -d2
lapply (lstas_mono … H1 … H2) #H destruct -h -d0 /2 width=3 by cprs_conf/
qed-.
-lemma drop_split: ∀L1,L2,l,m2,s. ⬇[s, l, m2] L1 ≡ L2 → ∀m1. m1 ≤ m2 →
- ∃∃L. ⬇[s, l, m2 - m1] L1 ≡ L & ⬇[s, l, m1] L ≡ L2.
-#L1 #L2 #l #m2 #s #H elim H -L1 -L2 -l -m2
-[ #l #m2 #Hs #m1 #Hm12 @(ex2_intro … (⋆))
- @drop_atom #H lapply (Hs H) -s #H destruct /2 width=1 by le_n_O_to_eq/
-| #I #L1 #V #m1 #Hm1 lapply (le_n_O_to_eq … Hm1) -Hm1
- #H destruct /2 width=3 by ex2_intro/
-| #I #L1 #L2 #V #m2 #HL12 #IHL12 #m1 @(nat_ind_plus … m1) -m1
- [ /3 width=3 by drop_drop, ex2_intro/
- | -HL12 #m1 #_ #Hm12 lapply (le_plus_to_le_r … Hm12) -Hm12
- #Hm12 elim (IHL12 … Hm12) -IHL12 >minus_plus_plus_l
- #L #HL1 #HL2 elim (lt_or_ge (|L1|) (m2-m1)) #H0
- [ elim (drop_inv_O1_gt … HL1 H0) -HL1 #H1 #H2 destruct
- elim (drop_inv_atom1 … HL2) -HL2 #H #_ destruct
- @(ex2_intro … (⋆)) [ @drop_O1_ge normalize // ]
- @drop_atom #H destruct
- | elim (drop_O1_pair … HL1 H0 I V) -HL1 -H0 /3 width=5 by drop_drop, ex2_intro/
- ]
- ]
-| #I #L1 #L2 #V1 #V2 #l #m2 #_ #HV21 #IHL12 #m1 #Hm12 elim (IHL12 … Hm12) -IHL12
- #L #HL1 #HL2 elim (lift_split … HV21 l m1) -HV21 /3 width=5 by drop_skip, ex2_intro/
-]
-qed-.
-
lemma drop_inv_refl: ∀L,l,m. ⬇[Ⓕ, l, m] L ≡ L → m = 0.
/2 width=5 by drop_inv_length_eq/ qed-.
(* Inversion lemmas with test for uniformity ********************************)
+lemma drops_inv_isuni: ∀L1,L2,f. ⬇*[Ⓣ, f] L1 ≡ L2 → 𝐔⦃f⦄ →
+ (𝐈⦃f⦄ ∧ L2 = L1) ∨
+ ∃∃I,K,V,g. ⬇*[Ⓣ, g] K ≡ L2 & L1 = K.ⓑ{I}V & f = ⫯g.
+#L1 #L2 #f * -L1 -L2 -f
+[ /4 width=1 by or_introl, conj/
+| /4 width=7 by ex3_4_intro, or_intror/
+| /7 width=6 by drops_fwd_isid, lifts_fwd_isid, isuni_inv_push, isid_push, or_introl, conj, eq_f3, sym_eq/
+]
+qed-.
+
(* Basic_2A1: was: drop_inv_O1_pair1 *)
lemma drops_inv_pair1_isuni: ∀I,K,L2,V,c,f. 𝐔⦃f⦄ → ⬇*[c, f] K.ⓑ{I}V ≡ L2 →
(𝐈⦃f⦄ ∧ L2 = K.ⓑ{I}V) ∨
]
qed-.
+(* Inversion lemmas with uniform relocations ********************************)
+
+lemma drops_inv_succ: ∀L1,L2,l. ⬇*[⫯l] L1 ≡ L2 →
+ ∃∃I,K,V. ⬇*[l] K ≡ L2 & L1 = K.ⓑ{I}V.
+#L1 #L2 #l #H elim (drops_inv_isuni … H) -H // *
+[ #H elim (isid_inv_next … H) -H //
+| /2 width=5 by ex2_3_intro/
+]
+qed-.
+
(* Basic_2A1: removed theorems 12:
drops_inv_nil drops_inv_cons d1_liftable_liftables
drop_refl_atom_O2 drop_inv_pair1
+++ /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/relocation/rtmap_sor.ma".
-include "basic_2/notation/relations/freestar_3.ma".
-include "basic_2/grammar/lenv.ma".
-
-(* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
-
-inductive frees: relation3 lenv term rtmap ≝
-| frees_atom: ∀I,f. 𝐈⦃f⦄ → frees (⋆) (⓪{I}) f
-| frees_sort: ∀I,L,V,s,f. frees L (⋆s) f →
- frees (L.ⓑ{I}V) (⋆s) (↑f)
-| frees_zero: ∀I,L,V,f. frees L V f →
- frees (L.ⓑ{I}V) (#0) (⫯f)
-| frees_lref: ∀I,L,V,i,f. frees L (#i) f →
- frees (L.ⓑ{I}V) (#⫯i) (↑f)
-| frees_gref: ∀I,L,V,p,f. frees L (§p) f →
- frees (L.ⓑ{I}V) (§p) (↑f)
-| frees_bind: ∀I,L,V,T,a,f1,f2,f. frees L V f1 → frees (L.ⓑ{I}V) T f2 →
- f1 ⋓ ⫱f2 ≡ f → frees L (ⓑ{a,I}V.T) f
-| frees_flat: ∀I,L,V,T,f1,f2,f. frees L V f1 → frees L T f2 →
- f1 ⋓ f2 ≡ f → frees L (ⓕ{I}V.T) f
-.
-
-interpretation
- "context-sensitive free variables (term)"
- 'FreeStar L T t = (frees L T t).
-
-(* Basic inversion lemmas ***************************************************)
-
-fact frees_inv_atom_aux: ∀L,X,f. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀J. L = ⋆ → X = ⓪{J} → 𝐈⦃f⦄.
-#L #X #f #H elim H -L -X -f /3 width=3 by isid_push/
-[5,6: #I #L #V #T [ #p ] #f1 #f2 #f #_ #_ #_ #_ #_ #J #_ #H destruct
-|*: #I #L #V [1,3,4: #x ] #f #_ #_ #J #H destruct
-]
-qed-.
-
-lemma frees_inv_atom: ∀I,f. ⋆ ⊢ 𝐅*⦃⓪{I}⦄ ≡ f → 𝐈⦃f⦄.
-/2 width=6 by frees_inv_atom_aux/ qed-.
-
-fact frees_inv_sort_aux: ∀L,X,f. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀x. X = ⋆x → 𝐈⦃f⦄.
-#L #X #f #H elim H -L -X -f /3 width=3 by isid_push/
-[ #_ #L #V #f #_ #_ #x #H destruct
-| #_ #L #_ #i #f #_ #_ #x #H destruct
-| #I #L #V #T #p #f1 #f2 #f #_ #_ #_ #_ #_ #x #H destruct
-| #I #L #V #T #f1 #f2 #f #_ #_ #_ #_ #_ #x #H destruct
-]
-qed-.
-
-lemma frees_inv_sort: ∀L,s,f. L ⊢ 𝐅*⦃⋆s⦄ ≡ f → 𝐈⦃f⦄.
-/2 width=5 by frees_inv_sort_aux/ qed-.
-
-fact frees_inv_gref_aux: ∀L,X,f. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀x. X = §x → 𝐈⦃f⦄.
-#L #X #f #H elim H -L -X -f /3 width=3 by isid_push/
-[ #_ #L #V #f #_ #_ #x #H destruct
-| #_ #L #_ #i #f #_ #_ #x #H destruct
-| #I #L #V #T #p #f1 #f2 #f #_ #_ #_ #_ #_ #x #H destruct
-| #I #L #V #T #f1 #f2 #f #_ #_ #_ #_ #_ #x #H destruct
-]
-qed-.
-
-lemma frees_inv_gref: ∀L,l,f. L ⊢ 𝐅*⦃§l⦄ ≡ f → 𝐈⦃f⦄.
-/2 width=5 by frees_inv_gref_aux/ qed-.
-
-fact frees_inv_zero_aux: ∀L,X,f. L ⊢ 𝐅*⦃X⦄ ≡ f → X = #0 →
- (L = ⋆ ∧ 𝐈⦃f⦄) ∨
- ∃∃I,K,V,g. K ⊢ 𝐅*⦃V⦄ ≡ g & L = K.ⓑ{I}V & f = ⫯g.
-#L #X #f * -L -X -f
-[ /3 width=1 by or_introl, conj/
-| #I #L #V #s #f #_ #H destruct
-| /3 width=7 by ex3_4_intro, or_intror/
-| #I #L #V #i #f #_ #H destruct
-| #I #L #V #l #f #_ #H destruct
-| #I #L #V #T #p #f1 #f2 #f #_ #_ #_ #H destruct
-| #I #L #V #T #f1 #f2 #f #_ #_ #_ #H destruct
-]
-qed-.
-
-lemma frees_inv_zero: ∀L,f. L ⊢ 𝐅*⦃#0⦄ ≡ f →
- (L = ⋆ ∧ 𝐈⦃f⦄) ∨
- ∃∃I,K,V,g. K ⊢ 𝐅*⦃V⦄ ≡ g & L = K.ⓑ{I}V & f = ⫯g.
-/2 width=3 by frees_inv_zero_aux/ qed-.
-
-fact frees_inv_lref_aux: ∀L,X,f. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀j. X = #(⫯j) →
- (L = ⋆ ∧ 𝐈⦃f⦄) ∨
- ∃∃I,K,V,g. K ⊢ 𝐅*⦃#j⦄ ≡ g & L = K.ⓑ{I}V & f = ↑g.
-#L #X #f * -L -X -f
-[ /3 width=1 by or_introl, conj/
-| #I #L #V #s #f #_ #j #H destruct
-| #I #L #V #f #_ #j #H destruct
-| #I #L #V #i #f #Ht #j #H destruct /3 width=7 by ex3_4_intro, or_intror/
-| #I #L #V #l #f #_ #j #H destruct
-| #I #L #V #T #p #f1 #f2 #f #_ #_ #_ #j #H destruct
-| #I #L #V #T #f1 #f2 #f #_ #_ #_ #j #H destruct
-]
-qed-.
-
-lemma frees_inv_lref: ∀L,i,f. L ⊢ 𝐅*⦃#(⫯i)⦄ ≡ f →
- (L = ⋆ ∧ 𝐈⦃f⦄) ∨
- ∃∃I,K,V,g. K ⊢ 𝐅*⦃#i⦄ ≡ g & L = K.ⓑ{I}V & f = ↑g.
-/2 width=3 by frees_inv_lref_aux/ qed-.
-
-fact frees_inv_bind_aux: ∀L,X,f. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀I,V,T,a. X = ⓑ{a,I}V.T →
- ∃∃f1,f2. L ⊢ 𝐅*⦃V⦄ ≡ f1 & L.ⓑ{I}V ⊢ 𝐅*⦃T⦄ ≡ f2 & f1 ⋓ ⫱f2 ≡ f.
-#L #X #f * -L -X -f
-[ #I #f #_ #J #W #U #b #H destruct
-| #I #L #V #s #f #_ #J #W #U #b #H destruct
-| #I #L #V #f #_ #J #W #U #b #H destruct
-| #I #L #V #i #f #_ #J #W #U #b #H destruct
-| #I #L #V #l #f #_ #J #W #U #b #H destruct
-| #I #L #V #T #p #f1 #f2 #f #HV #HT #Hf #J #W #U #b #H destruct /2 width=5 by ex3_2_intro/
-| #I #L #V #T #f1 #f2 #f #_ #_ #_ #J #W #U #b #H destruct
-]
-qed-.
-
-lemma frees_inv_bind: ∀I,L,V,T,a,f. L ⊢ 𝐅*⦃ⓑ{a,I}V.T⦄ ≡ f →
- ∃∃f1,f2. L ⊢ 𝐅*⦃V⦄ ≡ f1 & L.ⓑ{I}V ⊢ 𝐅*⦃T⦄ ≡ f2 & f1 ⋓ ⫱f2 ≡ f.
-/2 width=4 by frees_inv_bind_aux/ qed-.
-
-fact frees_inv_flat_aux: ∀L,X,f. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀I,V,T. X = ⓕ{I}V.T →
- ∃∃f1,f2. L ⊢ 𝐅*⦃V⦄ ≡ f1 & L ⊢ 𝐅*⦃T⦄ ≡ f2 & f1 ⋓ f2 ≡ f.
-#L #X #f * -L -X -f
-[ #I #f #_ #J #W #U #H destruct
-| #I #L #V #s #f #_ #J #W #U #H destruct
-| #I #L #V #f #_ #J #W #U #H destruct
-| #I #L #V #i #f #_ #J #W #U #H destruct
-| #I #L #V #l #f #_ #J #W #U #H destruct
-| #I #L #V #T #p #f1 #f2 #f #_ #_ #_ #J #W #U #H destruct
-| #I #L #V #T #f1 #f2 #f #HV #HT #Hf #J #W #U #H destruct /2 width=5 by ex3_2_intro/
-]
-qed-.
-
-lemma frees_inv_flat: ∀I,L,V,T,f. L ⊢ 𝐅*⦃ⓕ{I}V.T⦄ ≡ f →
- ∃∃f1,f2. L ⊢ 𝐅*⦃V⦄ ≡ f1 & L ⊢ 𝐅*⦃T⦄ ≡ f2 & f1 ⋓ f2 ≡ f.
-/2 width=4 by frees_inv_flat_aux/ qed-.
-
-(* Basic forward lemmas ****************************************************)
-
-lemma frees_fwd_isfin: ∀L,T,f. L ⊢ 𝐅*⦃T⦄ ≡ f → 𝐅⦃f⦄.
-#L #T #f #H elim H -L -T -f
-/3 width=5 by sor_isfin, isfin_isid, isfin_tl, isfin_push, isfin_next/
-qed-.
-
-(* Basic properties ********************************************************)
-
-lemma frees_eq_repl_back: ∀L,T. eq_repl_back … (λf. L ⊢ 𝐅*⦃T⦄ ≡ f).
-#L #T #f1 #H elim H -L -T -f1
-[ /3 width=3 by frees_atom, isid_eq_repl_back/
-| #I #L #V #s #f1 #_ #IH #f2 #Hf12
- elim (eq_inv_px … Hf12) -Hf12 /3 width=3 by frees_sort/
-| #I #L #V #f1 #_ #IH #f2 #Hf12
- elim (eq_inv_nx … Hf12) -Hf12 /3 width=3 by frees_zero/
-| #I #L #V #i #f1 #_ #IH #f2 #Hf12
- elim (eq_inv_px … Hf12) -Hf12 /3 width=3 by frees_lref/
-| #I #L #V #l #f1 #_ #IH #f2 #Hf12
- elim (eq_inv_px … Hf12) -Hf12 /3 width=3 by frees_gref/
-| /3 width=7 by frees_bind, sor_eq_repl_back3/
-| /3 width=7 by frees_flat, sor_eq_repl_back3/
-]
-qed-.
-
-lemma frees_eq_repl_fwd: ∀L,T. eq_repl_fwd … (λf. L ⊢ 𝐅*⦃T⦄ ≡ f).
-#L #T @eq_repl_sym /2 width=3 by frees_eq_repl_back/
-qed-.
-
-lemma frees_sort_gen: ∀L,s,f. 𝐈⦃f⦄ → L ⊢ 𝐅*⦃⋆s⦄ ≡ f.
-#L elim L -L
-/4 width=3 by frees_eq_repl_back, frees_sort, frees_atom, eq_push_inv_isid/
-qed.
-
-lemma frees_gref_gen: ∀L,p,f. 𝐈⦃f⦄ → L ⊢ 𝐅*⦃§p⦄ ≡ f.
-#L elim L -L
-/4 width=3 by frees_eq_repl_back, frees_gref, frees_atom, eq_push_inv_isid/
-qed.
-
-(* Basic_2A1: removed theorems 27:
- frees_eq frees_be frees_inv
- frees_inv_sort frees_inv_gref frees_inv_lref frees_inv_lref_free
- frees_inv_lref_skip frees_inv_lref_ge frees_inv_lref_lt
- frees_inv_bind frees_inv_flat frees_inv_bind_O
- frees_lref_eq frees_lref_be frees_weak
- frees_bind_sn frees_bind_dx frees_flat_sn frees_flat_dx
- lreq_frees_trans frees_lreq_conf
- llor_atom llor_skip llor_total
- llor_tail_frees llor_tail_cofrees
-*)
+++ /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/frees.ma".
-
-(* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
-
-(* Main inversion lemmas ****************************************************)
-
-theorem frees_mono: ∀L,T,f1. L ⊢ 𝐅*⦃T⦄ ≡ f1 → ∀f2. L ⊢ 𝐅*⦃T⦄ ≡ f2 → f1 ≗ f2.
-#L #T #f1 #H elim H -L -T -f1
-[ /3 width=2 by frees_inv_atom, isid_inv_eq_repl/
-| /4 width=5 by frees_inv_sort, eq_push_inv_isid, isid_inv_eq_repl, eq_trans/
-| #I #L #V #f1 #_ #IH #x #H elim (frees_inv_zero … H) -H *
- [ #H destruct
- | #Z #Y #X #f2 #Hf2 #H1 #H2 destruct /3 width=5 by eq_next/
- ]
-| #I #L #V #i #f1 #_ #IH #x #H elim (frees_inv_lref … H) -H *
- [ #H destruct
- | #Z #Y #X #f2 #Hf2 #H1 #H2 destruct /3 width=5 by eq_push/
- ]
-| /4 width=5 by frees_inv_gref, eq_push_inv_isid, isid_inv_eq_repl, eq_trans/
-| #I #L #V #T #p #f1 #f2 #f #_ #_ #Hf #IHV #IHT #g #H elim (frees_inv_bind … H) -H
- #g1 #g2 #HV #HT #Hg @(sor_mono … Hf) -Hf
- /5 width=3 by sor_eq_repl_fwd2, sor_eq_repl_fwd1, tl_eq_repl/ (**) (* full auto too slow *)
-| #I #L #V #T #f1 #f2 #f #_ #_ #Hf #IHV #IHT #g #H elim (frees_inv_flat … H) -H
- #g1 #g2 #HV #HT #Hg @(sor_mono … Hf) -Hf
- /4 width=3 by sor_eq_repl_fwd2, sor_eq_repl_fwd1/ (**) (* full auto too slow *)
-]
-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/lreq.ma".
-include "basic_2/relocation/frees.ma".
-
-(* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
-
-(* Properties with ranged equivalence for local environments ****************)
-
-lemma frees_lreq_conf: ∀L1,T,f. L1 ⊢ 𝐅*⦃T⦄ ≡ f → ∀L2. L1 ≡[f] L2 → L2 ⊢ 𝐅*⦃T⦄ ≡ f.
-#L1 #T #f #H elim H -L1 -T -f
-[ #I #f #Hf #X #H lapply (lreq_inv_atom1 … H) -H
- #H destruct /2 width=1 by frees_atom/
-| #I #L1 #V1 #s #f #_ #IH #X #H elim (lreq_inv_push1 … H) -H
- /3 width=1 by frees_sort/
-| #I #L1 #V1 #f #_ #IH #X #H elim (lreq_inv_next1 … H) -H
- /3 width=1 by frees_zero/
-| #I #L1 #V1 #i #f #_ #IH #X #H elim (lreq_inv_push1 … H) -H
- /3 width=1 by frees_lref/
-| #I #L1 #V1 #l #f #_ #IH #X #H elim (lreq_inv_push1 … H) -H
- /3 width=1 by frees_gref/
-| /6 width=5 by frees_bind, lreq_inv_tl, sle_lreq_trans, sor_inv_sle_dx, sor_inv_sle_sn/
-| /5 width=5 by frees_flat, sle_lreq_trans, sor_inv_sle_dx, sor_inv_sle_sn/
-]
-qed-.
-
-lemma lreq_frees_trans: ∀L1,T,f. L1 ⊢ 𝐅*⦃T⦄ ≡ f → ∀L2. L2 ≡[f] L1 → L2 ⊢ 𝐅*⦃T⦄ ≡ f.
-/3 width=3 by frees_lreq_conf, lreq_sym/ 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 "ground_2/relocation/rtmap_id.ma".
-include "basic_2/grammar/cl_restricted_weight.ma".
-include "basic_2/relocation/frees.ma".
-
-(* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
-
-(* Advanced properties ******************************************************)
-
-(* Notes: this replaces lemma 1400 concluding the "big tree" theorem *)
-lemma frees_total: ∀L,T. ∃f. L ⊢ 𝐅*⦃T⦄ ≡ f.
-@(f2_ind … rfw) #n #IH #L *
-[ cases L -L /3 width=2 by frees_atom, ex_intro/
- #L #I #V *
- [ #s #Hn destruct elim (IH L (⋆s)) -IH /3 width=2 by frees_sort, ex_intro/
- | * [2: #i ] #Hn destruct
- [ elim (IH L (#i)) -IH /3 width=2 by frees_lref, ex_intro/
- | elim (IH L V) -IH /3 width=2 by frees_zero, ex_intro/
- ]
- | #l #Hn destruct elim (IH L (§l)) -IH /3 width=2 by frees_gref, ex_intro/
- ]
-| * [ #p ] #I #V #T #Hn destruct elim (IH L V) // #f1 #HV
- [ elim (IH (L.ⓑ{I}V) T) -IH // #f2 #HT
- elim (sor_isfin_ex f1 (⫱f2))
- /3 width=6 by frees_fwd_isfin, frees_bind, isfin_tl, ex_intro/
- | elim (IH L T) -IH // #f2 #HT
- elim (sor_isfin_ex f1 f2)
- /3 width=6 by frees_fwd_isfin, frees_flat, ex_intro/
- ]
-]
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/notation/relations/lazyeq_6.ma".
-include "basic_2/grammar/genv.ma".
-include "basic_2/relocation/frees_weight.ma".
-include "basic_2/relocation/frees_lreq.ma".
-
-(* RANGED EQUIVALENCE FOR CLOSURES ******************************************)
-
-inductive freq (G) (L1) (T): relation3 genv lenv term ≝
-| fleq_intro: ∀L2,f. L1 ⊢ 𝐅*⦃T⦄ ≡ f → L1 ≡[f] L2 → freq G L1 T G L2 T
-.
-
-interpretation
- "ranged equivalence (closure)"
- 'LazyEq G1 L1 T1 G2 L2 T2 = (freq G1 L1 T1 G2 L2 T2).
-
-(* Basic properties *********************************************************)
-
-lemma freq_refl: tri_reflexive … freq.
-#G #L #T elim (frees_total L T) /2 width=3 by fleq_intro/
-qed.
-
-lemma freq_sym: tri_symmetric … freq.
-#G1 #L1 #T1 #G2 #L2 #T2 * /4 width=3 by fleq_intro, frees_lreq_conf, lreq_sym/
-qed-.
-
-(* Basic inversion lemmas ***************************************************)
-
-lemma freq_inv_gen: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≡ ⦃G2, L2, T2⦄ →
- ∃∃f. G1 = G2 & L1 ⊢ 𝐅*⦃T1⦄ ≡ f & L1 ≡[f] L2 & T1 = T2.
-#G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2 /2 width=3 by ex4_intro/
-qed-.
-
-(* Basic_2A1: removed theorems 6:
- fleq_refl fleq_sym fleq_inv_gen
- fleq_trans fleq_canc_sn fleq_canc_dx
-*)
+++ /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/lreq_lreq.ma".
-include "basic_2/relocation/frees_frees.ma".
-include "basic_2/relocation/freq.ma".
-
-(* RANGED EQUIVALENCE FOR CLOSURES *****************************************)
-
-(* Main properties **********************************************************)
-
-theorem freq_trans: tri_transitive … freq.
-#G1 #G #L1 #L #T1 #T * -G -L -T
-#L #f1 #H1T11 #Hf1 #G2 #L2 #T2 * -G2 -L2 -T2 #L2 #f2 #HT12 #Hf2
-lapply (frees_lreq_conf … H1T11 … Hf1) #HT11
-lapply (frees_mono … HT12 … HT11) -HT11 -HT12
-/4 width=7 by fleq_intro, lreq_eq_repl_back, lreq_trans/
-qed-.
-
-theorem freq_canc_sn: ∀G,G1,G2,L,L1,L2,T,T1,T2.
- ⦃G, L, T⦄ ≡ ⦃G1, L1, T1⦄→ ⦃G, L, T⦄ ≡ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≡ ⦃G2, L2, T2⦄.
-/3 width=5 by freq_trans, freq_sym/ qed-.
-
-theorem freq_canc_dx: ∀G1,G2,G,L1,L2,L,T1,T2,T.
- ⦃G1, L1, T1⦄ ≡ ⦃G, L, T⦄ → ⦃G2, L2, T2⦄ ≡ ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ ≡ ⦃G2, L2, T2⦄.
-/3 width=5 by freq_trans, freq_sym/ qed-.
+
(**************************************************************************)
(* ___ *)
(* ||M|| *)
]
qed-.
-lemma lift_SO: ∀i. ⬆*[1] #i ≡ #(⫯i).
-/2 width=1 by lifts_lref/ qed.
+lemma lift_lref_uni: ∀l,i. ⬆*[l] #i ≡ #(l+i).
+#l elim l -l /2 width=1 by lifts_lref/
+qed.
(* Basic_1: includes: lift_free (right to left) *)
(* Basic_2A1: includes: lift_split *)
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/suptermplus_6.ma".
+include "basic_2/s_transition/fqu.ma".
+
+(* PLUS-ITERATED SUPCLOSURE *************************************************)
+
+definition fqup: tri_relation genv lenv term ≝ tri_TC … fqu.
+
+interpretation "plus-iterated structural successor (closure)"
+ 'SupTermPlus G1 L1 T1 G2 L2 T2 = (fqup G1 L1 T1 G2 L2 T2).
+
+(* Basic properties *********************************************************)
+
+lemma fqu_fqup: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄.
+/2 width=1 by tri_inj/ qed.
+
+lemma fqup_strap1: ∀G1,G,G2,L1,L,L2,T1,T,T2.
+ ⦃G1, L1, T1⦄ ⊐+ ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄.
+/2 width=5 by tri_step/ qed.
+
+lemma fqup_strap2: ∀G1,G,G2,L1,L,L2,T1,T,T2.
+ ⦃G1, L1, T1⦄ ⊐ ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐+ ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄.
+/2 width=5 by tri_TC_strap/ qed.
+
+lemma fqup_pair_sn: ∀I,G,L,V,T. ⦃G, L, ②{I}V.T⦄ ⊐+ ⦃G, L, V⦄.
+/2 width=1 by fqu_pair_sn, fqu_fqup/ qed.
+
+lemma fqup_bind_dx: ∀a,I,G,L,V,T. ⦃G, L, ⓑ{a,I}V.T⦄ ⊐+ ⦃G, L.ⓑ{I}V, T⦄.
+/2 width=1 by fqu_bind_dx, fqu_fqup/ qed.
+
+lemma fqup_flat_dx: ∀I,G,L,V,T. ⦃G, L, ⓕ{I}V.T⦄ ⊐+ ⦃G, L, T⦄.
+/2 width=1 by fqu_flat_dx, fqu_fqup/ qed.
+
+lemma fqup_flat_dx_pair_sn: ∀I1,I2,G,L,V1,V2,T. ⦃G, L, ⓕ{I1}V1.②{I2}V2.T⦄ ⊐+ ⦃G, L, V2⦄.
+/2 width=5 by fqu_pair_sn, fqup_strap1/ qed.
+
+lemma fqup_bind_dx_flat_dx: ∀a,G,I1,I2,L,V1,V2,T. ⦃G, L, ⓑ{a,I1}V1.ⓕ{I2}V2.T⦄ ⊐+ ⦃G, L.ⓑ{I1}V1, T⦄.
+/2 width=5 by fqu_flat_dx, fqup_strap1/ qed.
+
+lemma fqup_flat_dx_bind_dx: ∀a,I1,I2,G,L,V1,V2,T. ⦃G, L, ⓕ{I1}V1.ⓑ{a,I2}V2.T⦄ ⊐+ ⦃G, L.ⓑ{I2}V2, T⦄.
+/2 width=5 by fqu_bind_dx, fqup_strap1/ qed.
+
+(* Basic eliminators ********************************************************)
+
+lemma fqup_ind: ∀G1,L1,T1. ∀R:relation3 ….
+ (∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → R G2 L2 T2) →
+ (∀G,G2,L,L2,T,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐ ⦃G2, L2, T2⦄ → R G L T → R G2 L2 T2) →
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → R G2 L2 T2.
+#G1 #L1 #T1 #R #IH1 #IH2 #G2 #L2 #T2 #H
+@(tri_TC_ind … IH1 IH2 G2 L2 T2 H)
+qed-.
+
+lemma fqup_ind_dx: ∀G2,L2,T2. ∀R:relation3 ….
+ (∀G1,L1,T1. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → R G1 L1 T1) →
+ (∀G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ⊐ ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐+ ⦃G2, L2, T2⦄ → R G L T → R G1 L1 T1) →
+ ∀G1,L1,T1. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → R G1 L1 T1.
+#G2 #L2 #T2 #R #IH1 #IH2 #G1 #L1 #T1 #H
+@(tri_TC_ind_dx … IH1 IH2 G1 L1 T1 H)
+qed-.
+
+(* Basic_2A1: removed theorems 1: fqup_drop *)
--- /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/s_computation/fqup.ma".
+
+(* PLUS-ITERATED SUPCLOSURE *************************************************)
+
+(* Properties with generic slicing for local environments *******************)
+
+lemma fqup_drops_succ: ∀G,K,T,l,L,U. ⬇*[⫯l] L ≡ K → ⬆*[⫯l] T ≡ U →
+ ⦃G, L, U⦄ ⊐+ ⦃G, K, T⦄.
+#G #K #T #l elim l -l
+[ #L #U #HLK #HTU elim (drops_inv_succ … HLK) -HLK
+ #I #Y #V #HY #H destruct <(drops_fwd_isid … HY) -K
+ /3 width=1 by fqu_fqup, fqu_drop/
+| #l #IH #L #U #HLK #HTU elim (drops_inv_succ … HLK) -HLK
+ #I #Y #V #HY #H destruct
+ elim (lifts_split_trans … HTU … (𝐔❴⫯l❵) (𝐔❴1❵)) -HTU
+ /4 width=5 by fqup_strap2, fqu_drop/
+]
+qed.
+
+lemma fqup_drops_strap1: ∀G1,G2,L1,K1,K2,T1,T2,U1,l. ⬇*[l] L1 ≡ K1 → ⬆*[l] T1 ≡ U1 →
+ ⦃G1, K1, T1⦄ ⊐ ⦃G2, K2, T2⦄ → ⦃G1, L1, U1⦄ ⊐+ ⦃G2, K2, T2⦄.
+#G1 #G2 #L1 #K1 #K2 #T1 #T2 #U1 *
+[ #HLK1 #HTU1 #HT12
+ >(drops_fwd_isid … HLK1) -L1 //
+ <(lifts_fwd_isid … HTU1) -U1 /2 width=1 by fqu_fqup/
+| /3 width=5 by fqup_strap1, fqup_drops_succ/
+]
+qed-.
+
+lemma fqup_lref: ∀I,G,L,K,V,i. ⬇*[i] L ≡ K.ⓑ{I}V → ⦃G, L, #i⦄ ⊐+ ⦃G, K, V⦄.
+/2 width=6 by fqup_drops_strap1/ 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/s_computation/fqup.ma".
+
+(* PLUS-ITERATED SUPCLOSURE *************************************************)
+
+(* Main properties **********************************************************)
+
+theorem fqup_trans: tri_transitive … fqup.
+/2 width=5 by tri_TC_transitive/ 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/s_transition/fqu_weight.ma".
+include "basic_2/s_computation/fqup.ma".
+
+(* PLUS-ITERATED SUPCLOSURE *************************************************)
+
+(* Forward lemmas with weight for closures **********************************)
+
+lemma fqup_fwd_fw: ∀G1,G2,L1,L2,T1,T2.
+ ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → ♯{G2, L2, T2} < ♯{G1, L1, T1}.
+#G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
+/3 width=3 by fqu_fwd_fw, transitive_lt/
+qed-.
+
+(* Advanced eliminators *****************************************************)
+
+lemma fqup_wf_ind: ∀R:relation3 …. (
+ ∀G1,L1,T1. (∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → R G2 L2 T2) →
+ R G1 L1 T1
+ ) → ∀G1,L1,T1. R G1 L1 T1.
+#R #HR @(f3_ind … fw) #x #IHx #G1 #L1 #T1 #H destruct /4 width=1 by fqup_fwd_fw/
+qed-.
+
+lemma fqup_wf_ind_eq: ∀R:relation3 …. (
+ ∀G1,L1,T1. (∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → R G2 L2 T2) →
+ ∀G2,L2,T2. G1 = G2 → L1 = L2 → T1 = T2 → R G2 L2 T2
+ ) → ∀G1,L1,T1. R G1 L1 T1.
+#R #HR @(f3_ind … fw) #x #IHx #G1 #L1 #T1 #H destruct /4 width=7 by fqup_fwd_fw/
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/suptermstar_6.ma".
+include "basic_2/s_transition/fquq.ma".
+
+(* STAR-ITERATED SUPCLOSURE *************************************************)
+
+definition fqus: tri_relation genv lenv term ≝ tri_TC … fquq.
+
+interpretation "star-iterated structural successor (closure)"
+ 'SupTermStar G1 L1 T1 G2 L2 T2 = (fqus G1 L1 T1 G2 L2 T2).
+
+(* Basic eliminators ********************************************************)
+
+lemma fqus_ind: ∀G1,L1,T1. ∀R:relation3 …. R G1 L1 T1 →
+ (∀G,G2,L,L2,T,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐⸮ ⦃G2, L2, T2⦄ → R G L T → R G2 L2 T2) →
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ → R G2 L2 T2.
+#G1 #L1 #T1 #R #IH1 #IH2 #G2 #L2 #T2 #H
+@(tri_TC_star_ind … IH1 IH2 G2 L2 T2 H) //
+qed-.
+
+lemma fqus_ind_dx: ∀G2,L2,T2. ∀R:relation3 …. R G2 L2 T2 →
+ (∀G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐* ⦃G2, L2, T2⦄ → R G L T → R G1 L1 T1) →
+ ∀G1,L1,T1. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ → R G1 L1 T1.
+#G2 #L2 #T2 #R #IH1 #IH2 #G1 #L1 #T1 #H
+@(tri_TC_star_ind_dx … IH1 IH2 G1 L1 T1 H) //
+qed-.
+
+(* Basic properties *********************************************************)
+
+lemma fqus_refl: tri_reflexive … fqus.
+/2 width=1 by tri_inj/ qed.
+
+lemma fquq_fqus: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄.
+/2 width=1 by tri_inj/ qed.
+
+lemma fqus_strap1: ∀G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄.
+/2 width=5 by tri_step/ qed-.
+
+lemma fqus_strap2: ∀G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐* ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄.
+/2 width=5 by tri_TC_strap/ qed-.
+
+(* Basic_2A1: removed theorems 1: fqus_drop *)
--- /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/s_computation/fqup_drops.ma".
+include "basic_2/s_computation/fqus_fqup.ma".
+
+(* STAR-ITERATED SUPCLOSURE *************************************************)
+
+(* Properties with generic slicing for local environments *******************)
+
+lemma fqus_drops: ∀G,L,K,T,U,l. ⬇*[l] L ≡ K → ⬆*[l] T ≡ U →
+ ⦃G, L, U⦄ ⊐* ⦃G, K, T⦄.
+#G #L #K #T #U * /3 width=3 by fqup_drops_succ, fqup_fqus/
+#HLK #HTU <(lifts_fwd_isid … HTU) -U // <(drops_fwd_isid … HLK) -K //
+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/s_computation/fqup.ma".
+include "basic_2/s_computation/fqus.ma".
+
+(* STAR-ITERATED SUPCLOSURE *************************************************)
+
+(* Alternative definition with plus-iterated supclosure *********************)
+
+lemma fqup_fqus: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄.
+#G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
+/3 width=5 by fqus_strap1, fquq_fqus, fqu_fquq/
+qed.
+
+(* Basic_2A1: was: fqus_inv_gen *)
+lemma fqus_inv_fqup: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ ∨ (∧∧ G1 = G2 & L1 = L2 & T1 = T2).
+#G1 #G2 #L1 #L2 #T1 #T2 #H @(fqus_ind … H) -G2 -L2 -T2 //
+#G #G2 #L #L2 #T #T2 #_ *
+[ #H2 * /3 width=5 by fqup_strap1, or_introl/
+ * /3 width=1 by fqu_fqup, or_introl/
+| * #HG #HL #HT destruct * /2 width=1 by or_introl/
+ * /2 width=1 by or_intror/
+]
+qed-.
+
+(* Advanced properties ******************************************************)
+
+lemma fqus_strap1_fqu: ∀G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄.
+#G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H2 elim (fqus_inv_fqup … H1) -H1
+[ /2 width=5 by fqup_strap1/
+| * /2 width=1 by fqu_fqup/
+]
+qed-.
+
+lemma fqus_strap2_fqu: ∀G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐* ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄.
+#G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H2 elim (fqus_inv_fqup … H2) -H2
+[ /2 width=5 by fqup_strap2/
+| * /2 width=1 by fqu_fqup/
+]
+qed-.
+
+lemma fqus_fqup_trans: ∀G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐+ ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄.
+#G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H2 @(fqup_ind … H2) -H2 -G2 -L2 -T2
+/2 width=5 by fqus_strap1_fqu, fqup_strap1/
+qed-.
+
+lemma fqup_fqus_trans: ∀G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G, L, T⦄ →
+ ⦃G, L, T⦄ ⊐* ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄.
+#G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 @(fqup_ind_dx … H1) -H1 -G1 -L1 -T1
+/3 width=5 by fqus_strap2_fqu, fqup_strap2/
+qed-.
+
+(* Advanced inversion lemmas for plus-iterated supclosure *******************)
+
+lemma fqup_inv_step_sn: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
+ ∃∃G,L,T. ⦃G1, L1, T1⦄ ⊐ ⦃G, L, T⦄ & ⦃G, L, T⦄ ⊐* ⦃G2, L2, T2⦄.
+#G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind_dx … H) -G1 -L1 -T1 /2 width=5 by ex2_3_intro/
+#G1 #G #L1 #L #T1 #T #H1 #_ * /4 width=9 by fqus_strap2, fqu_fquq, ex2_3_intro/
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/s_computation/fqus.ma".
+
+(* STAR-ITERATED SUPCLOSURE *************************************************)
+
+(* Main properties **********************************************************)
+
+theorem fqus_trans: tri_transitive … fqus.
+/2 width=5 by tri_TC_transitive/ 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/s_transition/fquq_weight.ma".
+include "basic_2/s_computation/fqus.ma".
+
+(* STAR-ITERATED SUPCLOSURE *************************************************)
+
+(* Forward lemmas with weight for closures **********************************)
+
+lemma fqus_fwd_fw: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ → ♯{G2, L2, T2} ≤ ♯{G1, L1, T1}.
+#G1 #G2 #L1 #L2 #T1 #T2 #H @(fqus_ind … H) -L2 -T2
+/3 width=3 by fquq_fwd_fw, transitive_le/
+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 "ground_2/relocation/rtmap_sor.ma".
+include "basic_2/notation/relations/freestar_3.ma".
+include "basic_2/grammar/lenv.ma".
+
+(* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
+
+inductive frees: relation3 lenv term rtmap ≝
+| frees_atom: ∀I,f. 𝐈⦃f⦄ → frees (⋆) (⓪{I}) f
+| frees_sort: ∀I,L,V,s,f. frees L (⋆s) f →
+ frees (L.ⓑ{I}V) (⋆s) (↑f)
+| frees_zero: ∀I,L,V,f. frees L V f →
+ frees (L.ⓑ{I}V) (#0) (⫯f)
+| frees_lref: ∀I,L,V,i,f. frees L (#i) f →
+ frees (L.ⓑ{I}V) (#⫯i) (↑f)
+| frees_gref: ∀I,L,V,p,f. frees L (§p) f →
+ frees (L.ⓑ{I}V) (§p) (↑f)
+| frees_bind: ∀I,L,V,T,a,f1,f2,f. frees L V f1 → frees (L.ⓑ{I}V) T f2 →
+ f1 ⋓ ⫱f2 ≡ f → frees L (ⓑ{a,I}V.T) f
+| frees_flat: ∀I,L,V,T,f1,f2,f. frees L V f1 → frees L T f2 →
+ f1 ⋓ f2 ≡ f → frees L (ⓕ{I}V.T) f
+.
+
+interpretation
+ "context-sensitive free variables (term)"
+ 'FreeStar L T t = (frees L T t).
+
+(* Basic inversion lemmas ***************************************************)
+
+fact frees_inv_atom_aux: ∀L,X,f. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀J. L = ⋆ → X = ⓪{J} → 𝐈⦃f⦄.
+#L #X #f #H elim H -L -X -f /3 width=3 by isid_push/
+[5,6: #I #L #V #T [ #p ] #f1 #f2 #f #_ #_ #_ #_ #_ #J #_ #H destruct
+|*: #I #L #V [1,3,4: #x ] #f #_ #_ #J #H destruct
+]
+qed-.
+
+lemma frees_inv_atom: ∀I,f. ⋆ ⊢ 𝐅*⦃⓪{I}⦄ ≡ f → 𝐈⦃f⦄.
+/2 width=6 by frees_inv_atom_aux/ qed-.
+
+fact frees_inv_sort_aux: ∀L,X,f. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀x. X = ⋆x → 𝐈⦃f⦄.
+#L #X #f #H elim H -L -X -f /3 width=3 by isid_push/
+[ #_ #L #V #f #_ #_ #x #H destruct
+| #_ #L #_ #i #f #_ #_ #x #H destruct
+| #I #L #V #T #p #f1 #f2 #f #_ #_ #_ #_ #_ #x #H destruct
+| #I #L #V #T #f1 #f2 #f #_ #_ #_ #_ #_ #x #H destruct
+]
+qed-.
+
+lemma frees_inv_sort: ∀L,s,f. L ⊢ 𝐅*⦃⋆s⦄ ≡ f → 𝐈⦃f⦄.
+/2 width=5 by frees_inv_sort_aux/ qed-.
+
+fact frees_inv_gref_aux: ∀L,X,f. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀x. X = §x → 𝐈⦃f⦄.
+#L #X #f #H elim H -L -X -f /3 width=3 by isid_push/
+[ #_ #L #V #f #_ #_ #x #H destruct
+| #_ #L #_ #i #f #_ #_ #x #H destruct
+| #I #L #V #T #p #f1 #f2 #f #_ #_ #_ #_ #_ #x #H destruct
+| #I #L #V #T #f1 #f2 #f #_ #_ #_ #_ #_ #x #H destruct
+]
+qed-.
+
+lemma frees_inv_gref: ∀L,l,f. L ⊢ 𝐅*⦃§l⦄ ≡ f → 𝐈⦃f⦄.
+/2 width=5 by frees_inv_gref_aux/ qed-.
+
+fact frees_inv_zero_aux: ∀L,X,f. L ⊢ 𝐅*⦃X⦄ ≡ f → X = #0 →
+ (L = ⋆ ∧ 𝐈⦃f⦄) ∨
+ ∃∃I,K,V,g. K ⊢ 𝐅*⦃V⦄ ≡ g & L = K.ⓑ{I}V & f = ⫯g.
+#L #X #f * -L -X -f
+[ /3 width=1 by or_introl, conj/
+| #I #L #V #s #f #_ #H destruct
+| /3 width=7 by ex3_4_intro, or_intror/
+| #I #L #V #i #f #_ #H destruct
+| #I #L #V #l #f #_ #H destruct
+| #I #L #V #T #p #f1 #f2 #f #_ #_ #_ #H destruct
+| #I #L #V #T #f1 #f2 #f #_ #_ #_ #H destruct
+]
+qed-.
+
+lemma frees_inv_zero: ∀L,f. L ⊢ 𝐅*⦃#0⦄ ≡ f →
+ (L = ⋆ ∧ 𝐈⦃f⦄) ∨
+ ∃∃I,K,V,g. K ⊢ 𝐅*⦃V⦄ ≡ g & L = K.ⓑ{I}V & f = ⫯g.
+/2 width=3 by frees_inv_zero_aux/ qed-.
+
+fact frees_inv_lref_aux: ∀L,X,f. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀j. X = #(⫯j) →
+ (L = ⋆ ∧ 𝐈⦃f⦄) ∨
+ ∃∃I,K,V,g. K ⊢ 𝐅*⦃#j⦄ ≡ g & L = K.ⓑ{I}V & f = ↑g.
+#L #X #f * -L -X -f
+[ /3 width=1 by or_introl, conj/
+| #I #L #V #s #f #_ #j #H destruct
+| #I #L #V #f #_ #j #H destruct
+| #I #L #V #i #f #Ht #j #H destruct /3 width=7 by ex3_4_intro, or_intror/
+| #I #L #V #l #f #_ #j #H destruct
+| #I #L #V #T #p #f1 #f2 #f #_ #_ #_ #j #H destruct
+| #I #L #V #T #f1 #f2 #f #_ #_ #_ #j #H destruct
+]
+qed-.
+
+lemma frees_inv_lref: ∀L,i,f. L ⊢ 𝐅*⦃#(⫯i)⦄ ≡ f →
+ (L = ⋆ ∧ 𝐈⦃f⦄) ∨
+ ∃∃I,K,V,g. K ⊢ 𝐅*⦃#i⦄ ≡ g & L = K.ⓑ{I}V & f = ↑g.
+/2 width=3 by frees_inv_lref_aux/ qed-.
+
+fact frees_inv_bind_aux: ∀L,X,f. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀I,V,T,a. X = ⓑ{a,I}V.T →
+ ∃∃f1,f2. L ⊢ 𝐅*⦃V⦄ ≡ f1 & L.ⓑ{I}V ⊢ 𝐅*⦃T⦄ ≡ f2 & f1 ⋓ ⫱f2 ≡ f.
+#L #X #f * -L -X -f
+[ #I #f #_ #J #W #U #b #H destruct
+| #I #L #V #s #f #_ #J #W #U #b #H destruct
+| #I #L #V #f #_ #J #W #U #b #H destruct
+| #I #L #V #i #f #_ #J #W #U #b #H destruct
+| #I #L #V #l #f #_ #J #W #U #b #H destruct
+| #I #L #V #T #p #f1 #f2 #f #HV #HT #Hf #J #W #U #b #H destruct /2 width=5 by ex3_2_intro/
+| #I #L #V #T #f1 #f2 #f #_ #_ #_ #J #W #U #b #H destruct
+]
+qed-.
+
+lemma frees_inv_bind: ∀I,L,V,T,a,f. L ⊢ 𝐅*⦃ⓑ{a,I}V.T⦄ ≡ f →
+ ∃∃f1,f2. L ⊢ 𝐅*⦃V⦄ ≡ f1 & L.ⓑ{I}V ⊢ 𝐅*⦃T⦄ ≡ f2 & f1 ⋓ ⫱f2 ≡ f.
+/2 width=4 by frees_inv_bind_aux/ qed-.
+
+fact frees_inv_flat_aux: ∀L,X,f. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀I,V,T. X = ⓕ{I}V.T →
+ ∃∃f1,f2. L ⊢ 𝐅*⦃V⦄ ≡ f1 & L ⊢ 𝐅*⦃T⦄ ≡ f2 & f1 ⋓ f2 ≡ f.
+#L #X #f * -L -X -f
+[ #I #f #_ #J #W #U #H destruct
+| #I #L #V #s #f #_ #J #W #U #H destruct
+| #I #L #V #f #_ #J #W #U #H destruct
+| #I #L #V #i #f #_ #J #W #U #H destruct
+| #I #L #V #l #f #_ #J #W #U #H destruct
+| #I #L #V #T #p #f1 #f2 #f #_ #_ #_ #J #W #U #H destruct
+| #I #L #V #T #f1 #f2 #f #HV #HT #Hf #J #W #U #H destruct /2 width=5 by ex3_2_intro/
+]
+qed-.
+
+lemma frees_inv_flat: ∀I,L,V,T,f. L ⊢ 𝐅*⦃ⓕ{I}V.T⦄ ≡ f →
+ ∃∃f1,f2. L ⊢ 𝐅*⦃V⦄ ≡ f1 & L ⊢ 𝐅*⦃T⦄ ≡ f2 & f1 ⋓ f2 ≡ f.
+/2 width=4 by frees_inv_flat_aux/ qed-.
+
+(* Basic forward lemmas ****************************************************)
+
+lemma frees_fwd_isfin: ∀L,T,f. L ⊢ 𝐅*⦃T⦄ ≡ f → 𝐅⦃f⦄.
+#L #T #f #H elim H -L -T -f
+/3 width=5 by sor_isfin, isfin_isid, isfin_tl, isfin_push, isfin_next/
+qed-.
+
+(* Basic properties ********************************************************)
+
+lemma frees_eq_repl_back: ∀L,T. eq_repl_back … (λf. L ⊢ 𝐅*⦃T⦄ ≡ f).
+#L #T #f1 #H elim H -L -T -f1
+[ /3 width=3 by frees_atom, isid_eq_repl_back/
+| #I #L #V #s #f1 #_ #IH #f2 #Hf12
+ elim (eq_inv_px … Hf12) -Hf12 /3 width=3 by frees_sort/
+| #I #L #V #f1 #_ #IH #f2 #Hf12
+ elim (eq_inv_nx … Hf12) -Hf12 /3 width=3 by frees_zero/
+| #I #L #V #i #f1 #_ #IH #f2 #Hf12
+ elim (eq_inv_px … Hf12) -Hf12 /3 width=3 by frees_lref/
+| #I #L #V #l #f1 #_ #IH #f2 #Hf12
+ elim (eq_inv_px … Hf12) -Hf12 /3 width=3 by frees_gref/
+| /3 width=7 by frees_bind, sor_eq_repl_back3/
+| /3 width=7 by frees_flat, sor_eq_repl_back3/
+]
+qed-.
+
+lemma frees_eq_repl_fwd: ∀L,T. eq_repl_fwd … (λf. L ⊢ 𝐅*⦃T⦄ ≡ f).
+#L #T @eq_repl_sym /2 width=3 by frees_eq_repl_back/
+qed-.
+
+lemma frees_sort_gen: ∀L,s,f. 𝐈⦃f⦄ → L ⊢ 𝐅*⦃⋆s⦄ ≡ f.
+#L elim L -L
+/4 width=3 by frees_eq_repl_back, frees_sort, frees_atom, eq_push_inv_isid/
+qed.
+
+lemma frees_gref_gen: ∀L,p,f. 𝐈⦃f⦄ → L ⊢ 𝐅*⦃§p⦄ ≡ f.
+#L elim L -L
+/4 width=3 by frees_eq_repl_back, frees_gref, frees_atom, eq_push_inv_isid/
+qed.
+
+(* Basic_2A1: removed theorems 27:
+ frees_eq frees_be frees_inv
+ frees_inv_sort frees_inv_gref frees_inv_lref frees_inv_lref_free
+ frees_inv_lref_skip frees_inv_lref_ge frees_inv_lref_lt
+ frees_inv_bind frees_inv_flat frees_inv_bind_O
+ frees_lref_eq frees_lref_be frees_weak
+ frees_bind_sn frees_bind_dx frees_flat_sn frees_flat_dx
+ lreq_frees_trans frees_lreq_conf
+ llor_atom llor_skip llor_total
+ llor_tail_frees llor_tail_cofrees
+*)
--- /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/frees.ma".
+
+(* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
+
+(* Main inversion lemmas ****************************************************)
+
+theorem frees_mono: ∀L,T,f1. L ⊢ 𝐅*⦃T⦄ ≡ f1 → ∀f2. L ⊢ 𝐅*⦃T⦄ ≡ f2 → f1 ≗ f2.
+#L #T #f1 #H elim H -L -T -f1
+[ /3 width=2 by frees_inv_atom, isid_inv_eq_repl/
+| /4 width=5 by frees_inv_sort, eq_push_inv_isid, isid_inv_eq_repl, eq_trans/
+| #I #L #V #f1 #_ #IH #x #H elim (frees_inv_zero … H) -H *
+ [ #H destruct
+ | #Z #Y #X #f2 #Hf2 #H1 #H2 destruct /3 width=5 by eq_next/
+ ]
+| #I #L #V #i #f1 #_ #IH #x #H elim (frees_inv_lref … H) -H *
+ [ #H destruct
+ | #Z #Y #X #f2 #Hf2 #H1 #H2 destruct /3 width=5 by eq_push/
+ ]
+| /4 width=5 by frees_inv_gref, eq_push_inv_isid, isid_inv_eq_repl, eq_trans/
+| #I #L #V #T #p #f1 #f2 #f #_ #_ #Hf #IHV #IHT #g #H elim (frees_inv_bind … H) -H
+ #g1 #g2 #HV #HT #Hg @(sor_mono … Hf) -Hf
+ /5 width=3 by sor_eq_repl_fwd2, sor_eq_repl_fwd1, tl_eq_repl/ (**) (* full auto too slow *)
+| #I #L #V #T #f1 #f2 #f #_ #_ #Hf #IHV #IHT #g #H elim (frees_inv_flat … H) -H
+ #g1 #g2 #HV #HT #Hg @(sor_mono … Hf) -Hf
+ /4 width=3 by sor_eq_repl_fwd2, sor_eq_repl_fwd1/ (**) (* full auto too slow *)
+]
+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/lreq.ma".
+include "basic_2/relocation/frees.ma".
+
+(* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
+
+(* Properties with ranged equivalence for local environments ****************)
+
+lemma frees_lreq_conf: ∀L1,T,f. L1 ⊢ 𝐅*⦃T⦄ ≡ f → ∀L2. L1 ≡[f] L2 → L2 ⊢ 𝐅*⦃T⦄ ≡ f.
+#L1 #T #f #H elim H -L1 -T -f
+[ #I #f #Hf #X #H lapply (lreq_inv_atom1 … H) -H
+ #H destruct /2 width=1 by frees_atom/
+| #I #L1 #V1 #s #f #_ #IH #X #H elim (lreq_inv_push1 … H) -H
+ /3 width=1 by frees_sort/
+| #I #L1 #V1 #f #_ #IH #X #H elim (lreq_inv_next1 … H) -H
+ /3 width=1 by frees_zero/
+| #I #L1 #V1 #i #f #_ #IH #X #H elim (lreq_inv_push1 … H) -H
+ /3 width=1 by frees_lref/
+| #I #L1 #V1 #l #f #_ #IH #X #H elim (lreq_inv_push1 … H) -H
+ /3 width=1 by frees_gref/
+| /6 width=5 by frees_bind, lreq_inv_tl, sle_lreq_trans, sor_inv_sle_dx, sor_inv_sle_sn/
+| /5 width=5 by frees_flat, sle_lreq_trans, sor_inv_sle_dx, sor_inv_sle_sn/
+]
+qed-.
+
+lemma lreq_frees_trans: ∀L1,T,f. L1 ⊢ 𝐅*⦃T⦄ ≡ f → ∀L2. L2 ≡[f] L1 → L2 ⊢ 𝐅*⦃T⦄ ≡ f.
+/3 width=3 by frees_lreq_conf, lreq_sym/ 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 "ground_2/relocation/rtmap_id.ma".
+include "basic_2/grammar/cl_restricted_weight.ma".
+include "basic_2/relocation/frees.ma".
+
+(* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
+
+(* Advanced properties ******************************************************)
+
+(* Notes: this replaces lemma 1400 concluding the "big tree" theorem *)
+lemma frees_total: ∀L,T. ∃f. L ⊢ 𝐅*⦃T⦄ ≡ f.
+@(f2_ind … rfw) #n #IH #L *
+[ cases L -L /3 width=2 by frees_atom, ex_intro/
+ #L #I #V *
+ [ #s #Hn destruct elim (IH L (⋆s)) -IH /3 width=2 by frees_sort, ex_intro/
+ | * [2: #i ] #Hn destruct
+ [ elim (IH L (#i)) -IH /3 width=2 by frees_lref, ex_intro/
+ | elim (IH L V) -IH /3 width=2 by frees_zero, ex_intro/
+ ]
+ | #l #Hn destruct elim (IH L (§l)) -IH /3 width=2 by frees_gref, ex_intro/
+ ]
+| * [ #p ] #I #V #T #Hn destruct elim (IH L V) // #f1 #HV
+ [ elim (IH (L.ⓑ{I}V) T) -IH // #f2 #HT
+ elim (sor_isfin_ex f1 (⫱f2))
+ /3 width=6 by frees_fwd_isfin, frees_bind, isfin_tl, ex_intro/
+ | elim (IH L T) -IH // #f2 #HT
+ elim (sor_isfin_ex f1 f2)
+ /3 width=6 by frees_fwd_isfin, frees_flat, ex_intro/
+ ]
+]
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/notation/relations/lazyeq_6.ma".
+include "basic_2/grammar/genv.ma".
+include "basic_2/relocation/frees_weight.ma".
+include "basic_2/relocation/frees_lreq.ma".
+
+(* RANGED EQUIVALENCE FOR CLOSURES ******************************************)
+
+inductive freq (G) (L1) (T): relation3 genv lenv term ≝
+| fleq_intro: ∀L2,f. L1 ⊢ 𝐅*⦃T⦄ ≡ f → L1 ≡[f] L2 → freq G L1 T G L2 T
+.
+
+interpretation
+ "ranged equivalence (closure)"
+ 'LazyEq G1 L1 T1 G2 L2 T2 = (freq G1 L1 T1 G2 L2 T2).
+
+(* Basic properties *********************************************************)
+
+lemma freq_refl: tri_reflexive … freq.
+#G #L #T elim (frees_total L T) /2 width=3 by fleq_intro/
+qed.
+
+lemma freq_sym: tri_symmetric … freq.
+#G1 #L1 #T1 #G2 #L2 #T2 * /4 width=3 by fleq_intro, frees_lreq_conf, lreq_sym/
+qed-.
+
+(* Basic inversion lemmas ***************************************************)
+
+lemma freq_inv_gen: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≡ ⦃G2, L2, T2⦄ →
+ ∃∃f. G1 = G2 & L1 ⊢ 𝐅*⦃T1⦄ ≡ f & L1 ≡[f] L2 & T1 = T2.
+#G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2 /2 width=3 by ex4_intro/
+qed-.
+
+(* Basic_2A1: removed theorems 6:
+ fleq_refl fleq_sym fleq_inv_gen
+ fleq_trans fleq_canc_sn fleq_canc_dx
+*)
--- /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/lreq_lreq.ma".
+include "basic_2/relocation/frees_frees.ma".
+include "basic_2/relocation/freq.ma".
+
+(* RANGED EQUIVALENCE FOR CLOSURES *****************************************)
+
+(* Main properties **********************************************************)
+
+theorem freq_trans: tri_transitive … freq.
+#G1 #G #L1 #L #T1 #T * -G -L -T
+#L #f1 #H1T11 #Hf1 #G2 #L2 #T2 * -G2 -L2 -T2 #L2 #f2 #HT12 #Hf2
+lapply (frees_lreq_conf … H1T11 … Hf1) #HT11
+lapply (frees_mono … HT12 … HT11) -HT11 -HT12
+/4 width=7 by fleq_intro, lreq_eq_repl_back, lreq_trans/
+qed-.
+
+theorem freq_canc_sn: ∀G,G1,G2,L,L1,L2,T,T1,T2.
+ ⦃G, L, T⦄ ≡ ⦃G1, L1, T1⦄→ ⦃G, L, T⦄ ≡ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≡ ⦃G2, L2, T2⦄.
+/3 width=5 by freq_trans, freq_sym/ qed-.
+
+theorem freq_canc_dx: ∀G1,G2,G,L1,L2,L,T1,T2,T.
+ ⦃G1, L1, T1⦄ ≡ ⦃G, L, T⦄ → ⦃G2, L2, T2⦄ ≡ ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ ≡ ⦃G2, L2, T2⦄.
+/3 width=5 by freq_trans, freq_sym/ qed-.
<page xmlns="http://lambdadelta.info/"
description = "\lambda\delta home page"
title = "\lambda\delta home page"
+ logo = "crux"
head = "cic:/matita/lambdadelta/basic_2/ (core λδ version 2)"
>
<sitemap name="sitemap"/>
</news>
<subsection name="A2">Stage "A2": "Extending the Applicability Condition"</subsection>
+
+ <news class="alpha" date="2016 March 25.">
+ Relocation with reference transforming maps (rtmap).
+ </news>
+
<news class="alpha" date="2015 October 9.">
λδ version 2A2 is started.
</news>
]
class "grass"
[ { "s-computation" * } {
- [ { "" * } {
- [ * ]
+ [ { "iterated structural successor for closures" * } {
+ [ "fqus ( ⦃?,?,?⦄ ⊐* ⦃?,?,?⦄ )" "fqus_weight" + "fqus_drops" + "fqus_fqup" + "fqus_fqus" * ]
+ [ "fqup ( ⦃?,?,?⦄ ⊐+ ⦃?,?,?⦄ )" "fqup_weight" + "fqup_drops" + "fqup_fqup" * ]
}
]
}
[ "cpys ( ⦃?,?⦄ ⊢ ? ▶*[?,?] ? )" "cpys_alt ( ⦃?,?⦄ ⊢ ? ▶▶*[?,?] ? )" "cpys_lift" + "cpys_cpys" * ]
}
]
- [ { "iterated structural successor for closures" * } {
- [ "fqus ( ⦃?,?,?⦄ ⊐* ⦃?,?,?⦄ )" "fqus_alt" + "fqus_fqus" * ]
- [ "fqup ( ⦃?,?,?⦄ ⊐+ ⦃?,?,?⦄ )" "fqup_fqup" * ]
- }
- ]
[ { "pointwise union for local environments" * } {
[ "llor ( ? ⋓[?,?] ? ≡ ? )" "llor_alt" + "llor_drop" * ]
}
/5 width=7 by isid_after_dx, after_eq_repl_back_2, after_next, after_push, eq_push_inv_isid/
qed.
+lemma after_uni_next2: ∀f2. 𝐔⦃f2⦄ → ∀f1,f. ⫯f2 ⊚ f1 ≡ f → f2 ⊚ ⫯f1 ≡ f.
+#f2 #H elim H -f2
+[ #f2 #Hf2 #f1 #f #Hf
+ elim (after_inv_nxx … Hf) -Hf [2,3: // ] #g #Hg #H0 destruct
+ /4 width=7 by after_isid_inv_sn, isid_after_sn, after_eq_repl_back_0, eq_next/
+| #f2 #_ #g2 #H2 #IH #f1 #f #Hf
+ elim (after_inv_nxx … Hf) -Hf [2,3: // ] #g #Hg #H0 destruct
+ /3 width=5 by after_next/
+]
+qed.
+
+(* Properties on uni ********************************************************)
+
+lemma after_uni: ∀n1,n2. 𝐔❴n1❵ ⊚ 𝐔❴n2❵ ≡ 𝐔❴n1+n2❵.
+@nat_elim2
+/4 width=5 by after_uni_next2, isid_after_sn, isid_after_dx, after_next/
+qed.
+
(* Forward lemmas on at *****************************************************)
lemma after_at_fwd: ∀i,i1,f. @⦃i1, f⦄ ≡ i → ∀f2,f1. f2 ⊚ f1 ≡ f →
<page xmlns="http://lambdadelta.info/"
description = "\lambda\delta home page"
title = "\lambda\delta home page"
+ logo = "crux"
head = "cic:/matita/lambdadelta/ground_2/ (background for λδ version 2)"
>
<sitemap name="sitemap"/>