| GRef _ ⇒ U
]
| TPair I V T ⇒ match I with
- [ Bind2 I ⇒ ⓑ{I} (flift d e V). (flift (d+1) e T)
- | Flat2 I ⇒ ⓕ{I} (flift d e V). (flift d e T)
+ [ Bind2 a I ⇒ ⓑ{a,I} (flift d e V). (flift (d+1) e T)
+ | Flat2 I ⇒ ⓕ{I} (flift d e V). (flift d e T)
]
].
theorem flift_lift: ∀T,d,e. ⇧[d, e] T ≡ ↑[d, e] T.
#T elim T -T
[ * #i #d #e //
- elim (lt_or_eq_or_gt i d) #Hid normalize
+ elim (lt_or_eq_or_gt i d) #Hid normalize
[ >(tri_lt ?????? Hid) /2 width=1/
| /2 width=1/
| >(tri_gt ?????? Hid) /3 width=2/
rtm_step (mk_rtm G u E S (ⓐV. T))
(mk_rtm G u E ({E, V} @ S) T)
| rtm_beta : ∀G,u,E,F,V,S,W,T.
- rtm_step (mk_rtm G u E ({F, V} @ S) (ⓛW. T))
+ rtm_step (mk_rtm G u E ({F, V} @ S) (+ⓛW. T))
(mk_rtm G u (E. ④{Abbr} {u, F, V}) S T)
| rtm_push : ∀G,u,E,W,T.
- rtm_step (mk_rtm G u E ⟠ (ⓛW. T))
+ rtm_step (mk_rtm G u E ⟠ (+ⓛW. T))
(mk_rtm G (u + 1) (E. ④{Abst} {u, E, W}) ⟠ T)
| rtm_theta : ∀G,u,E,S,V,T.
- rtm_step (mk_rtm G u E S (ⓓV. T))
+ rtm_step (mk_rtm G u E S (+ⓓV. T))
(mk_rtm G u (E. ④{Abbr} {u, E, V}) S T)
.
| GRef _ ⇒ U
]
| TPair I V T ⇒ match I with
- [ Bind2 I ⇒ ⓑ{I} (fsubst W d V). (fsubst W (d+1) T)
- | Flat2 I ⇒ ⓕ{I} (fsubst W d V). (fsubst W d T)
+ [ Bind2 a I ⇒ ⓑ{a,I} (fsubst W d V). (fsubst W (d+1) T)
+ | Flat2 I ⇒ ⓕ{I} (fsubst W d V). (fsubst W d T)
]
].
]
| -HLK >(delift_inv_gref1 … H) -H //
]
-| * #I #V1 #T1 #IHV1 #IHT1 #L #X #d #HLK #H
+| * [ #a ] #I #V1 #T1 #IHV1 #IHT1 #L #X #d #HLK #H
[ elim (delift_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
<(IHV1 … HV12) -IHV1 -HV12 // <(IHT1 … HT12) -IHT1 -HT12 // /2 width=1/
| elim (delift_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
@(s4 … HB … ◊ … HV2 HLK2)
@(s7 … HB … HKV2B) //
]
-| #L #V #T #B #A #_ #_ #IHB #IHA #L0 #des #HL0 #X #H #L2 #HL20
+| #a #L #V #T #B #A #_ #_ #IHB #IHA #L0 #des #HL0 #X #H #L2 #HL20
elim (lifts_inv_bind1 … H) -H #V0 #T0 #HV0 #HT0 #H destruct
lapply (aacr_acr … H1RP H2RP A) #HA
lapply (aacr_acr … H1RP H2RP B) #HB
lapply (s1 … HB) -HB #HB
@(s5 … HA … ◊ ◊) // /3 width=5/
-| #L #W #T #B #A #HLWB #_ #IHB #IHA #L0 #des #HL0 #X #H #L2 #HL02
+| #a #L #W #T #B #A #HLWB #_ #IHB #IHA #L0 #des #HL0 #X #H #L2 #HL02
elim (lifts_inv_bind1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct
@(aacr_abst … H1RP H2RP)
[ lapply (aacr_acr … H1RP H2RP B) #HB
(* Note: this is Tait's ii *)
definition S3 ≝ λRP,C:lenv→predicate term.
- ∀L,Vs,V,T,W. C L (ⒶVs. ⓓV. T) → RP L W → C L (ⒶVs. ⓐV. ⓛW. T).
+ ∀a,L,Vs,V,T,W. C L (ⒶVs. ⓓ{a}V. T) → RP L W → C L (ⒶVs. ⓐV. ⓛ{a}W. T).
definition S4 ≝ λRP,C:lenv→predicate term. ∀L,K,Vs,V1,V2,i.
C L (ⒶVs. V2) → ⇧[0, i + 1] V1 ≡ V2 →
definition S5 ≝ λRP,C:lenv→predicate term.
∀L,V1s,V2s. ⇧[0, 1] V1s ≡ V2s →
- ∀V,T. C (L. ⓓV) (ⒶV2s. T) → RP L V → C L (ⒶV1s. ⓓV. T).
+ ∀a,V,T. C (L. ⓓV) (ⒶV2s. T) → RP L V → C L (ⒶV1s. ⓓ{a}V. T).
definition S6 ≝ λRP,C:lenv→predicate term.
∀L,Vs,T,W. C L (ⒶVs. T) → RP L W → C L (ⒶVs. ⓝW. T).
elim (lifts_inv_applv1 … H) -H #V0s #T0 #HV0s #HT0 #H destruct
lapply (s1 … IHB … HB) #HV0
@(s2 … IHA … (V0 @ V0s)) /2 width=4 by lifts_simple_dx/ /3 width=6/
-| #L #Vs #U #T #W #HA #HW #L0 #V0 #X #des #HB #HL0 #H
+| #a #L #Vs #U #T #W #HA #HW #L0 #V0 #X #des #HB #HL0 #H
elim (lifts_inv_applv1 … H) -H #V0s #Y #HV0s #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
elim (lifts_lift_trans … Hdes0 … HVW1 … HW12) // -Hdes0 -Hi0 #V3 #HV13 #HVW2
>(lift_mono … HV13 … HV12) in HVW2; -V3 #HVW2
@(s4 … IHA … (V0 @ V0s) … HW12 HL02) /3 width=4/
-| #L #V1s #V2s #HV12s #V #T #HA #HV #L0 #V10 #X #des #HB #HL0 #H
+| #L #V1s #V2s #HV12s #a #V #T #HA #HV #L0 #V10 #X #des #HB #HL0 #H
elim (lifts_inv_applv1 … H) -H #V10s #Y #HV10s #HY #H destruct
elim (lifts_inv_bind1 … HY) -HY #V0 #T0 #HV0 #HT0 #H destruct
elim (lift_total V10 0 1) #V20 #HV120
qed.
lemma aacr_abst: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λL,T. RP L T) →
- ∀L,W,T,A,B. RP L W → (
+ ∀a,L,W,T,A,B. RP L W → (
∀L0,V0,T0,des. ⇩*[des] L0 ≡ L → ⇧*[des + 1] T ≡ T0 →
⦃L0, V0⦄ ϵ[RP] 〚B〛 → ⦃L0. ⓓV0, T0⦄ ϵ[RP] 〚A〛
) →
- ⦃L, ⓛW. T⦄ ϵ[RP] 〚②B. A〛.
-#RR #RS #RP #H1RP #H2RP #L #W #T #A #B #HW #HA #L0 #V0 #X #des #HB #HL0 #H
+ ⦃L, ⓛ{a}W. T⦄ ϵ[RP] 〚②B. A〛.
+#RR #RS #RP #H1RP #H2RP #a #L #W #T #A #B #HW #HA #L0 #V0 #X #des #HB #HL0 #H
lapply (aacr_acr … H1RP H2RP A) #HCA
lapply (aacr_acr … H1RP H2RP B) #HCB
elim (lifts_inv_bind1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct
(* Advanced properties ******************************************************)
lemma cprs_abst_dx: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀V,T1,T2.
- L.ⓛV ⊢ T1 ➡* T2 → L ⊢ ⓛV1. T1 ➡* ⓛV2. T2.
-#L #V1 #V2 #HV12 #V #T1 #T2 #HT12 @(cprs_ind … HT12) -T2
+ L.ⓛV ⊢ T1 ➡* T2 → ∀a. L ⊢ ⓛ{a}V1. T1 ➡* ⓛ{a}V2. T2.
+#L #V1 #V2 #HV12 #V #T1 #T2 #HT12 #a @(cprs_ind … HT12) -T2
[ /3 width=2/
| /3 width=6 by cprs_strap1, cpr_abst/ (**) (* /3 width=6/ is too slow *)
]
qed.
lemma cprs_abbr1_dx: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀T1,T2. L. ⓓV1 ⊢ T1 ➡* T2 →
- L ⊢ ⓓV1. T1 ➡* ⓓV2. T2.
-#L #V1 #V2 #HV12 #T1 #T2 #HT12 @(cprs_ind_dx … HT12) -T1
+ ∀a. L ⊢ ⓓ{a}V1. T1 ➡* ⓓ{a}V2. T2.
+#L #V1 #V2 #HV12 #T1 #T2 #HT12 #a @(cprs_ind_dx … HT12) -T1
[ /3 width=5/
| #T1 #T #HT1 #_ #IHT1
@(cprs_strap2 … IHT1) -IHT1 /2 width=1/
qed.
lemma cpr_abbr1: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀T1,T2. L. ⓓV1 ⊢ T1 ➡ T2 →
- L ⊢ ⓓV1. T1 ➡* ⓓV2. T2.
+ ∀a. L ⊢ ⓓ{a}V1. T1 ➡* ⓓ{a}V2. T2.
/3 width=1/ qed.
lemma cpr_abbr2: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀T1,T2. L. ⓓV2 ⊢ T1 ➡ T2 →
- L ⊢ ⓓV1. T1 ➡* ⓓV2. T2.
+ ∀a. L ⊢ ⓓ{a}V1. T1 ➡* ⓓ{a}V2. T2.
#L #V1 #V2 #HV12 #T1 #T2 #HT12
lapply (lcpr_cpr_trans (L. ⓓV1) … HT12) /2 width=1/
qed.
(* Basic_1: was pr3_gen_appl *)
lemma cprs_inv_appl1: ∀L,V1,T1,U2. L ⊢ ⓐV1. T1 ➡* U2 →
- ∨∨ ∃∃V2,T2. L ⊢ V1 ➡* V2 & L ⊢ T1 ➡* T2 &
- U2 = ⓐV2. T2
- | ∃∃V2,W,T. L ⊢ V1 ➡* V2 &
- L ⊢ T1 ➡* ⓛW. T & L ⊢ ⓓV2. T ➡* U2
- | ∃∃V0,V2,V,T. L ⊢ V1 ➡* V0 & ⇧[0,1] V0 ≡ V2 &
- L ⊢ T1 ➡* ⓓV. T & L ⊢ ⓓV. ⓐV2. T ➡* U2.
+ ∨∨ ∃∃V2,T2. L ⊢ V1 ➡* V2 & L ⊢ T1 ➡* T2 &
+ U2 = ⓐV2. T2
+ | ∃∃a,V2,W,T. L ⊢ V1 ➡* V2 &
+ L ⊢ T1 ➡* ⓛ{a}W. T & L ⊢ ⓓ{a}V2. T ➡* U2
+ | ∃∃a,V0,V2,V,T. L ⊢ V1 ➡* V0 & ⇧[0,1] V0 ≡ V2 &
+ L ⊢ T1 ➡* ⓓ{a}V. T & L ⊢ ⓓ{a}V. ⓐV2. T ➡* U2.
#L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5/
#U #U2 #_ #HU2 * *
[ #V0 #T0 #HV10 #HT10 #H destruct
elim (cpr_inv_appl1 … HU2) -HU2 *
[ #V2 #T2 #HV02 #HT02 #H destruct /4 width=5/
- | #V2 #W2 #T #T2 #HV02 #HT2 #H1 #H2 destruct /4 width=7/
- | #V #V2 #W0 #W2 #T #T2 #HV0 #HW02 #HT2 #HV2 #H1 #H2 destruct
- @or3_intro2 @(ex4_4_intro … HV2 HT10) /2 width=3/ /3 width=1/ (**) (* explicit constructor. /5 width=8/ is too slow because TC_transitive gets in the way *)
+ | #a #V2 #W2 #T #T2 #HV02 #HT2 #H1 #H2 destruct /4 width=7/
+ | #a #V #V2 #W0 #W2 #T #T2 #HV0 #HW02 #HT2 #HV2 #H1 #H2 destruct
+ @or3_intro2 @(ex4_5_intro … HV2 HT10) /2 width=3/ /3 width=1/ (**) (* explicit constructor. /5 width=8/ is too slow because TC_transitive gets in the way *)
]
-| /4 width=8/
-| /4 width=10/
+| /4 width=9/
+| /4 width=11/
]
qed-.
L ⊢ ⓕ{I} V1. T1 ➡* ⓕ{I} V2. T2.
#I #L #T1 #T2 #HT12 #V1 #V2 #HV12 @(cprs_ind … HV12) -V2 /2 width=1/
#V #V2 #_ #HV2 #IHV1
-@(cprs_trans … IHV1) -IHV1 /2 width=1/
+@(cprs_trans … IHV1) -IHV1 /2 width=1/
qed.
lemma cprs_abst: ∀L,V1,V2. L ⊢ V1 ➡* V2 → ∀V,T1,T2.
- L.ⓛV ⊢ T1 ➡* T2 → L ⊢ ⓛV1. T1 ➡* ⓛV2. T2.
-#L #V1 #V2 #HV12 #V #T1 #T2 #HT12 @(cprs_ind … HV12) -V2
+ L.ⓛV ⊢ T1 ➡* T2 → ∀a. L ⊢ ⓛ{a}V1. T1 ➡* ⓛ{a}V2. T2.
+#L #V1 #V2 #HV12 #V #T1 #T2 #HT12 #a @(cprs_ind … HV12) -V2
[ lapply (cprs_lsubs_trans … HT12 (L.ⓛV1) ?) -HT12 /2 width=2/
| #V0 #V2 #_ #HV02 #IHV01
@(cprs_trans … IHV01) -V1 /2 width=2/
qed.
lemma cprs_abbr1: ∀L,V1,T1,T2. L. ⓓV1 ⊢ T1 ➡* T2 → ∀V2. L ⊢ V1 ➡* V2 →
- L ⊢ ⓓV1. T1 ➡* ⓓV2. T2.
-#L #V1 #T1 #T2 #HT12 #V2 #HV12 @(cprs_ind … HV12) -V2 /2 width=1/
+ ∀a.L ⊢ ⓓ{a}V1. T1 ➡* ⓓ{a}V2. T2.
+#L #V1 #T1 #T2 #HT12 #V2 #HV12 #a @(cprs_ind … HV12) -V2 /2 width=1/
#V #V2 #_ #HV2 #IHV1
@(cprs_trans … IHV1) -IHV1 /2 width=1/
qed.
lemma cprs_abbr2_dx: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀T1,T2. L. ⓓV2 ⊢ T1 ➡* T2 →
- L ⊢ ⓓV1. T1 ➡* ⓓV2. T2.
-#L #V1 #V2 #HV12 #T1 #T2 #HT12 @(cprs_ind_dx … HT12) -T1
+ ∀a. L ⊢ ⓓ{a}V1. T1 ➡* ⓓ{a}V2. T2.
+#L #V1 #V2 #HV12 #T1 #T2 #HT12 #a @(cprs_ind_dx … HT12) -T1
[ /2 width=1/
| #T1 #T #HT1 #_ #IHT1
lapply (lcpr_cpr_trans (L. ⓓV1) … HT1) -HT1 /2 width=1/ #HT1
qed.
lemma cprs_abbr2: ∀L,V1,V2. L ⊢ V1 ➡* V2 → ∀T1,T2. L. ⓓV2 ⊢ T1 ➡* T2 →
- L ⊢ ⓓV1. T1 ➡* ⓓV2. T2.
+ ∀a. L ⊢ ⓓ{a}V1. T1 ➡* ⓓ{a}V2. T2.
#L #V1 #V2 #HV12 @(cprs_ind … HV12) -V2 /2 width=1/
-#V #V2 #_ #HV2 #IHV1 #T1 #T2 #HT12
-lapply (IHV1 T1 T1 ?) -IHV1 // #HV1
+#V #V2 #_ #HV2 #IHV1 #T1 #T2 #HT12 #a
+lapply (IHV1 T1 T1 ? a) -IHV1 // #HV1
@(cprs_trans … HV1) -HV1 /2 width=1/
qed.
lemma cprs_beta_dx: ∀L,V1,V2,W,T1,T2.
- L ⊢ V1 ➡ V2 → L.ⓛW ⊢ T1 ➡* T2 → L ⊢ ⓐV1.ⓛW.T1 ➡* ⓓV2.T2.
-#L #V1 #V2 #W #T1 #T2 #HV12 #HT12 @(cprs_ind … HT12) -T2
+ L ⊢ V1 ➡ V2 → L.ⓛW ⊢ T1 ➡* T2 →
+ ∀a.L ⊢ ⓐV1.ⓛ{a}W.T1 ➡* ⓓ{a}V2.T2.
+#L #V1 #V2 #W #T1 #T2 #HV12 #HT12 #a @(cprs_ind … HT12) -T2
[ /3 width=1/
| -HV12 #T #T2 #_ #HT2 #IHT1
lapply (cpr_lsubs_trans … HT2 (L.ⓓV2) ?) -HT2 /2 width=1/ #HT2
(* Properties on inverse basic term relocation ******************************)
(* Note: this should be stated with tprs *)
-lemma cprs_zeta_delift: ∀L,V,T1,T2. L.ⓓV ⊢ ▼*[O, 1] T1 ≡ T2 → L ⊢ ⓓV.T1 ➡* T2.
+lemma cprs_zeta_delift: ∀L,V,T1,T2. L.ⓓV ⊢ ▼*[O, 1] T1 ≡ T2 → L ⊢ +ⓓV.T1 ➡* T2.
#L #V #T1 #T2 * #T #HT1 #HT2
-@(cprs_strap2 … (ⓓV.T)) [ /3 width=3/ | @inj /3 width=3/ ] (**) (* explicit constructor, /5 width=3/ is too slow *)
+@(cprs_strap2 … (+ⓓV.T)) [ /3 width=3/ | @inj /3 width=3/ ] (**) (* explicit constructor, /5 width=3/ is too slow *)
qed.
(* Basic_1: was only: pr3_gen_cabbr *)
(* Advanced inversion lemmas ************************************************)
(* Basic_1: was pr3_gen_abbr *)
-lemma cprs_inv_abbr1: ∀L,V1,T1,U2. L ⊢ ⓓV1. T1 ➡* U2 →
+lemma cprs_inv_abbr1: ∀a,L,V1,T1,U2. L ⊢ ⓓ{a}V1. T1 ➡* U2 →
(∃∃V2,T2. L ⊢ V1 ➡* V2 & L. ⓓV1 ⊢ T1 ➡* T2 &
- U2 = ⓓV2. T2
+ U2 = ⓓ{a}V2. T2
) ∨
- ∃∃T2. L. ⓓV1 ⊢ T1 ➡* T2 & ⇧[0, 1] U2 ≡ T2.
-#L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5/
+ ∃∃T2. L. ⓓV1 ⊢ T1 ➡* T2 & ⇧[0, 1] U2 ≡ T2 & a = true.
+#a #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5/
#U0 #U2 #_ #HU02 * *
[ #V0 #T0 #HV10 #HT10 #H destruct
elim (cpr_inv_abbr1 … HU02) -HU02 *
qed-.
(* Basic_1: was: pr3_gen_abst *)
-lemma cprs_inv_abst1: ∀I,W,L,V1,T1,U2. L ⊢ ⓛV1. T1 ➡* U2 →
+lemma cprs_inv_abst1: ∀I,W,a,L,V1,T1,U2. L ⊢ ⓛ{a}V1. T1 ➡* U2 →
∃∃V2,T2. L ⊢ V1 ➡* V2 & L. ⓑ{I} W ⊢ T1 ➡* T2 &
- U2 = ⓛV2. T2.
-#I #W #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /2 width=5/
+ U2 = ⓛ{a}V2. T2.
+#I #W #a #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /2 width=5/
#U #U2 #_ #HU2 * #V #T #HV1 #HT1 #H destruct
elim (cpr_inv_abst1 … HU2 I W) -HU2 #V2 #T2 #HV2 #HT2 #H destruct /3 width=5/
qed-.
-lemma cprs_inv_abst: ∀L,V1,V2,T1,T2. L ⊢ ⓛV1. T1 ➡* ⓛV2. T2 → ∀I,W.
+lemma cprs_inv_abst: ∀a,L,V1,V2,T1,T2. L ⊢ ⓛ{a}V1. T1 ➡* ⓛ{a}V2. T2 → ∀I,W.
L ⊢ V1 ➡* V2 ∧ L. ⓑ{I} W ⊢ T1 ➡* T2.
-#L #V1 #V2 #T1 #T2 #H #I #W
+#a #L #V1 #V2 #T1 #T2 #H #I #W
elim (cprs_inv_abst1 I W … H) -H #V #T #HV1 #HT1 #H destruct /2 width=1/
qed-.
qed-.
(* Basic_1: was: pr3_iso_beta *)
-lemma cprs_fwd_beta: ∀L,V,W,T,U. L ⊢ ⓐV. ⓛW. T ➡* U →
- ⓐV. ⓛW. T ≃ U ∨ L ⊢ ⓓV. T ➡* U.
-#L #V #W #T #U #H
+lemma cprs_fwd_beta: ∀a,L,V,W,T,U. L ⊢ ⓐV. ⓛ{a}W. T ➡* U →
+ ⓐV. ⓛ{a}W. T ≃ U ∨ L ⊢ ⓓ{a}V. T ➡* U.
+#a #L #V #W #T #U #H
elim (cprs_inv_appl1 … H) -H *
[ #V0 #T0 #_ #_ #H destruct /2 width=1/
-| #V0 #W0 #T0 #HV0 #HT0 #HU
+| #b #V0 #W0 #T0 #HV0 #HT0 #HU
elim (cprs_inv_abst1 Abbr V … HT0) -HT0 #W1 #T1 #_ #HT1 #H destruct -W1
@or_intror -W
@(cprs_trans … HU) -U /2 width=1/ (**) (* explicit constructor *)
-| #V1 #V2 #V0 #T1 #_ #_ #HT1 #_
+| #b #V1 #V2 #V0 #T1 #_ #_ #HT1 #_
elim (cprs_inv_abst1 Abbr V … HT1) -HT1 #W2 #T2 #_ #_ #H destruct
]
qed-.
lapply (ldrop_fwd_ldrop2 … HLK) -HLK /3 width=9/
qed-.
-lemma cprs_fwd_theta: ∀L,V1,V,T,U. L ⊢ ⓐV1. ⓓV. T ➡* U →
- ∀V2. ⇧[0, 1] V1 ≡ V2 → ⓐV1. ⓓV. T ≃ U ∨
- L ⊢ ⓓV. ⓐV2. T ➡* U.
-#L #V1 #V #T #U #H #V2 #HV12
+lemma cprs_fwd_theta: ∀a,L,V1,V,T,U. L ⊢ ⓐV1. ⓓ{a}V. T ➡* U →
+ ∀V2. ⇧[0, 1] V1 ≡ V2 → ⓐV1. ⓓ{a}V. T ≃ U ∨
+ L ⊢ ⓓ{a}V. ⓐV2. T ➡* U.
+#a #L #V1 #V #T #U #H #V2 #HV12
elim (cprs_inv_appl1 … H) -H *
[ -HV12 #V0 #T0 #_ #_ #H destruct /2 width=1/
-| #V0 #W #T0 #HV10 #HT0 #HU
+| #b #V0 #W #T0 #HV10 #HT0 #HU
elim (cprs_inv_abbr1 … HT0) -HT0 *
[ #V3 #T3 #_ #_ #H destruct
- | #X #HT2 #H
+ | #X #HT2 #H #H0 destruct
elim (lift_inv_bind1 … H) -H #W2 #T2 #HW2 #HT02 #H destruct
@or_intror @(cprs_trans … HU) -U (**) (* explicit constructor *)
- @(cprs_trans … (ⓓV.ⓐV2.ⓛW2.T2)) [ /3 width=1/ ] -T
- @(cprs_strap2 … (ⓐV1.ⓛW.T0)) [ /5 width=7/ ] -V -V2 -W2 -T2
- @(cprs_strap2 … (ⓓV1.T0)) [ /3 width=1/ ] -W /2 width=1/
+ @(cprs_trans … (+ⓓV.ⓐV2.ⓛ{b}W2.T2)) [ /3 width=1/ ] -T
+ @(cprs_strap2 … (ⓐV1.ⓛ{b}W.T0)) [ /5 width=7/ ] -V -V2 -W2 -T2
+ @(cprs_strap2 … (ⓓ{b}V1.T0)) [ /3 width=1/ ] -W /2 width=1/
]
-| #V3 #V4 #V0 #T0 #HV13 #HV34 #HT0 #HU
+| #b #V3 #V4 #V0 #T0 #HV13 #HV34 #HT0 #HU
@or_intror @(cprs_trans … HU) -U (**) (* explicit constructor *)
elim (cprs_inv_abbr1 … HT0) -HT0 *
[ #V5 #T5 #HV5 #HT5 #H destruct
lapply (cprs_lift (L.ⓓV) … HV12 … HV13 … HV34) -V1 -V3 /2 width=1/
/3 width=1/
- | #X #HT1 #H
+ | #X #HT1 #H #H0 destruct
elim (lift_inv_bind1 … H) -H #V5 #T5 #HV05 #HT05 #H destruct
lapply (cprs_lift (L.ⓓV0) … HV12 … HV13 … HV34) -V3 /2 width=1/ #HV24
- @(cprs_trans … (ⓓV.ⓐV2.ⓓV5.T5)) [ /3 width=1/ ] -T
- @(cprs_strap2 … (ⓐV1.ⓓV0.T0)) [ /5 width=7/ ] -V -V5 -T5
- @(cprs_strap2 … (ⓓV0.ⓐV2.T0)) [ /3 width=3/ ] -V1 /3 width=1/
+ @(cprs_trans … (+ⓓV.ⓐV2.ⓓ{b}V5.T5)) [ /3 width=1/ ] -T
+ @(cprs_strap2 … (ⓐV1.ⓓ{b}V0.T0)) [ /5 width=7/ ] -V -V5 -T5
+ @(cprs_strap2 … (ⓓ{b}V0.ⓐV2.T0)) [ /3 width=3/ ] -V1 /3 width=1/
]
]
qed-.
#V #Vs #IHVs #U #H
elim (cprs_inv_appl1 … H) -H *
[ -IHVs #V0 #T0 #_ #_ #H destruct /2 width=1/
-| #V0 #W0 #T0 #HV0 #HT0 #HU
+| #a #V0 #W0 #T0 #HV0 #HT0 #HU
lapply (IHVs … HT0) -IHVs -HT0 #HT0
elim (tstc_inv_bind_appls_simple … HT0 ?) //
-| #V1 #V2 #V0 #T0 #HV1 #HV12 #HT0 #HU
+| #a #V1 #V2 #V0 #T0 #HV1 #HV12 #HT0 #HU
lapply (IHVs … HT0) -IHVs -HT0 #HT0
elim (tstc_inv_bind_appls_simple … HT0 ?) //
]
qed-.
(* Basic_1: was: pr3_iso_appls_beta *)
-lemma cprs_fwd_beta_vector: ∀L,Vs,V,W,T,U. L ⊢ ⒶVs. ⓐV. ⓛW. T ➡* U →
- ⒶVs. ⓐV. ⓛW. T ≃ U ∨ L ⊢ ⒶVs. ⓓV. T ➡* U.
-#L #Vs elim Vs -Vs /2 width=1 by cprs_fwd_beta/
+lemma cprs_fwd_beta_vector: ∀a,L,Vs,V,W,T,U. L ⊢ ⒶVs. ⓐV. ⓛ{a}W. T ➡* U →
+ ⒶVs. ⓐV. ⓛ{a}W. T ≃ U ∨ L ⊢ ⒶVs. ⓓ{a}V. T ➡* U.
+#a #L #Vs elim Vs -Vs /2 width=1 by cprs_fwd_beta/
#V0 #Vs #IHVs #V #W #T #U #H
elim (cprs_inv_appl1 … H) -H *
[ -IHVs #V1 #T1 #_ #_ #H destruct /2 width=1/
-| #V1 #W1 #T1 #HV01 #HT1 #HU
+| #b #V1 #W1 #T1 #HV01 #HT1 #HU
elim (IHVs … HT1) -IHVs -HT1 #HT1
[ elim (tstc_inv_bind_appls_simple … HT1 ?) //
| @or_intror -W (**) (* explicit constructor *)
@(cprs_trans … HU) -U
- @(cprs_strap1 … (ⓐV1.ⓛW1.T1)) [ /2 width=1/ ] -V -V0 -Vs -T /3 width=1/
+ @(cprs_strap1 … (ⓐV1.ⓛ{b}W1.T1)) [ /2 width=1/ ] -V -V0 -Vs -T /3 width=1/
]
-| #V1 #V2 #V3 #T1 #HV01 #HV12 #HT1 #HU
+| #b #V1 #V2 #V3 #T1 #HV01 #HV12 #HT1 #HU
elim (IHVs … HT1) -IHVs -HT1 #HT1
[ elim (tstc_inv_bind_appls_simple … HT1 ?) //
| @or_intror -W (**) (* explicit constructor *)
@(cprs_trans … HU) -U
- @(cprs_strap1 … (ⓐV1.ⓓV3.T1)) [ /2 width=1/ ] -V -V0 -Vs -T /3 width=3/
+ @(cprs_strap1 … (ⓐV1.ⓓ{b}V3.T1)) [ /2 width=1/ ] -V -V0 -Vs -T /3 width=3/
]
]
qed-.
#V #Vs #IHVs #U #H -K -V1
elim (cprs_inv_appl1 … H) -H *
[ -IHVs #V0 #T0 #_ #_ #H destruct /2 width=1/
-| #V0 #W0 #T0 #HV0 #HT0 #HU
+| #b #V0 #W0 #T0 #HV0 #HT0 #HU
elim (IHVs … HT0) -IHVs -HT0 #HT0
[ elim (tstc_inv_bind_appls_simple … HT0 ?) //
| @or_intror -i (**) (* explicit constructor *)
@(cprs_trans … HU) -U
- @(cprs_strap1 … (ⓐV0.ⓛW0.T0)) [ /2 width=1/ ] -V -V2 -Vs /3 width=1/
+ @(cprs_strap1 … (ⓐV0.ⓛ{b}W0.T0)) [ /2 width=1/ ] -V -V2 -Vs /3 width=1/
]
-| #V0 #V1 #V3 #T0 #HV0 #HV01 #HT0 #HU
+| #b #V0 #V1 #V3 #T0 #HV0 #HV01 #HT0 #HU
elim (IHVs … HT0) -IHVs -HT0 #HT0
[ elim (tstc_inv_bind_appls_simple … HT0 ?) //
| @or_intror -i (**) (* explicit constructor *)
@(cprs_trans … HU) -U
- @(cprs_strap1 … (ⓐV0.ⓓV3.T0)) [ /2 width=1/ ] -V -V2 -Vs /3 width=3/
+ @(cprs_strap1 … (ⓐV0.ⓓ{b}V3.T0)) [ /2 width=1/ ] -V -V2 -Vs /3 width=3/
]
]
qed-.
(* Basic_1: was: pr3_iso_appls_abbr *)
lemma cprs_fwd_theta_vector: ∀L,V1s,V2s. ⇧[0, 1] V1s ≡ V2s →
- ∀V,T,U. L ⊢ ⒶV1s. ⓓV. T ➡* U →
- ⒶV1s. ⓓV. T ≃ U ∨ L ⊢ ⓓV. ⒶV2s. T ➡* U.
+ ∀a,V,T,U. L ⊢ ⒶV1s. ⓓ{a}V. T ➡* U →
+ ⒶV1s. ⓓ{a}V. T ≃ U ∨ L ⊢ ⓓ{a}V. ⒶV2s. T ➡* U.
#L #V1s #V2s * -V1s -V2s /3 width=1/
-#V1s #V2s #V1a #V2a #HV12a #HV12s
+#V1s #V2s #V1a #V2a #HV12a #HV12s #a
generalize in match HV12a; -HV12a
generalize in match V2a; -V2a
generalize in match V1a; -V1a
#V1s #V2s #V1b #V2b #HV12b #_ #IHV12s #V1a #V2a #HV12a #V #T #U #H
elim (cprs_inv_appl1 … H) -H *
[ -IHV12s -HV12a -HV12b #V0 #T0 #_ #_ #H destruct /2 width=1/
-| #V0a #W0 #T0 #HV10a #HT0 #HU
+| #b #V0a #W0 #T0 #HV10a #HT0 #HU
elim (IHV12s … HV12b … HT0) -IHV12s -HT0 #HT0
[ -HV12a -HV12b -HV10a -HU
elim (tstc_inv_pair1 … HT0) #V1 #T1 #H destruct
@(cprs_trans … HU) -U
elim (cprs_inv_abbr1 … HT0) -HT0 *
[ -HV12a -HV12b -HV10a #V1 #T1 #_ #_ #H destruct
- | -V1b #X #HT1 #H
+ | -V1b #X #HT1 #H #H0 destruct
elim (lift_inv_bind1 … H) -H #W1 #T1 #HW01 #HT01 #H destruct
- @(cprs_trans … (ⓓV.ⓐV2a.ⓛW1.T1)) [ /3 width=1/ ] -T -V2b -V2s
- @(cprs_strap2 … (ⓐV1a.ⓛW0.T0)) [ /5 width=7/ ] -V -V2a -W1 -T1
- @(cprs_strap2 … (ⓓV1a.T0)) [ /3 width=1/ ] -W0 /2 width=1/
+ @(cprs_trans … (+ⓓV.ⓐV2a.ⓛ{b}W1.T1)) [ /3 width=1/ ] -T -V2b -V2s
+ @(cprs_strap2 … (ⓐV1a.ⓛ{b}W0.T0)) [ /5 width=7/ ] -V -V2a -W1 -T1
+ @(cprs_strap2 … (ⓓ{b}V1a.T0)) [ /3 width=1/ ] -W0 /2 width=1/
]
]
-| #V0a #Va #V0 #T0 #HV10a #HV0a #HT0 #HU
+| #b #V0a #Va #V0 #T0 #HV10a #HV0a #HT0 #HU
elim (IHV12s … HV12b … HT0) -HV12b -IHV12s -HT0 #HT0
[ -HV12a -HV10a -HV0a -HU
elim (tstc_inv_pair1 … HT0) #V1 #T1 #H destruct
elim (cprs_inv_abbr1 … HT0) -HT0 *
[ #V1 #T1 #HV1 #HT1 #H destruct
lapply (cprs_lift (L.ⓓV) … HV12a … HV10a … HV0a) -V1a -V0a [ /2 width=1/ ] #HV2a
- @(cprs_trans … (ⓓV.ⓐV2a.T1)) [ /3 width=1/ ] -T -V2b -V2s /3 width=1/
- | #X #HT1 #H
+ @(cprs_trans … (ⓓ{a}V.ⓐV2a.T1)) [ /3 width=1/ ] -T -V2b -V2s /3 width=1/
+ | #X #HT1 #H #H0 destruct
elim (lift_inv_bind1 … H) -H #V1 #T1 #HW01 #HT01 #H destruct
lapply (cprs_lift (L.ⓓV0) … HV12a … HV10a … HV0a) -V0a [ /2 width=1/ ] #HV2a
- @(cprs_trans … (ⓓV.ⓐV2a.ⓓV1.T1)) [ /3 width=1/ ] -T -V2b -V2s
- @(cprs_strap2 … (ⓐV1a.ⓓV0.T0)) [ /5 width=7/ ] -V -V1 -T1
- @(cprs_strap2 … (ⓓV0.ⓐV2a.T0)) [ /3 width=3/ ] -V1a /3 width=1/
+ @(cprs_trans … (+ⓓV.ⓐV2a.ⓓ{b}V1.T1)) [ /3 width=1/ ] -T -V2b -V2s
+ @(cprs_strap2 … (ⓐV1a.ⓓ{b}V0.T0)) [ /5 width=7/ ] -V -V1 -T1
+ @(cprs_strap2 … (ⓓ{b}V0.ⓐV2a.T0)) [ /3 width=3/ ] -V1a /3 width=1/
]
]
]
#V #Vs #IHVs #W #T #U #H
elim (cprs_inv_appl1 … H) -H *
[ -IHVs #V0 #T0 #_ #_ #H destruct /2 width=1/
-| #V0 #W0 #T0 #HV0 #HT0 #HU
+| #b #V0 #W0 #T0 #HV0 #HT0 #HU
elim (IHVs … HT0) -IHVs -HT0 #HT0
[ elim (tstc_inv_bind_appls_simple … HT0 ?) //
| @or_intror -W (**) (* explicit constructor *)
@(cprs_trans … HU) -U
- @(cprs_strap1 … (ⓐV0.ⓛW0.T0)) [ /2 width=1/ ] -V -Vs -T /3 width=1/
+ @(cprs_strap1 … (ⓐV0.ⓛ{b}W0.T0)) [ /2 width=1/ ] -V -Vs -T /3 width=1/
]
-| #V0 #V1 #V2 #T0 #HV0 #HV01 #HT0 #HU
+| #b #V0 #V1 #V2 #T0 #HV0 #HV01 #HT0 #HU
elim (IHVs … HT0) -IHVs -HT0 #HT0
[ elim (tstc_inv_bind_appls_simple … HT0 ?) //
| @or_intror -W (**) (* explicit constructor *)
@(cprs_trans … HU) -U
- @(cprs_strap1 … (ⓐV0.ⓓV2.T0)) [ /2 width=1/ ] -V -Vs -T /3 width=3/
+ @(cprs_strap1 … (ⓐV0.ⓓ{b}V2.T0)) [ /2 width=1/ ] -V -Vs -T /3 width=3/
]
qed-.
/2 width=5/ qed.
fact csn_fwd_bind_dx_aux: ∀L,U. L ⊢ ⬊* U →
- ∀I,V,T. U = ⓑ{I} V. T → L. ⓑ{I} V ⊢ ⬊* T.
-#L #U #H elim H -H #U0 #_ #IH #I #V #T #H destruct
+ ∀a,I,V,T. U = ⓑ{a,I} V. T → L. ⓑ{I} V ⊢ ⬊* T.
+#L #U #H elim H -H #U0 #_ #IH #a #I #V #T #H destruct
@csn_intro #T2 #HLT2 #HT2
-@(IH (ⓑ{I} V. T2)) -IH // /2 width=1/ -HLT2 #H destruct /2 width=1/
+@(IH (ⓑ{a,I} V. T2)) -IH // /2 width=1/ -HLT2 #H destruct /2 width=1/
qed.
(* Basic_1: was: sn3_gen_bind *)
-lemma csn_fwd_bind_dx: ∀I,L,V,T. L ⊢ ⬊* ⓑ{I} V. T → L. ⓑ{I} V ⊢ ⬊* T.
-/2 width=3/ qed.
+lemma csn_fwd_bind_dx: ∀a,I,L,V,T. L ⊢ ⬊* ⓑ{a,I} V. T → L. ⓑ{I} V ⊢ ⬊* T.
+/2 width=4/ qed.
@IHT /2 width=2/ -IHT -HT0 /2 width=3/
qed.
-lemma csn_abbr: ∀L,V. L ⊢ ⬊* V → ∀T. L. ⓓV ⊢ ⬊* T → L ⊢ ⬊* ⓓV. T.
-#L #V #HV elim HV -V #V #_ #IHV #T #HT @(csn_ind_alt … HT) -T #T #HT #IHT
+lemma csn_abbr: ∀a,L,V. L ⊢ ⬊* V → ∀T. L. ⓓV ⊢ ⬊* T → L ⊢ ⬊* ⓓ{a}V. T.
+#a #L #V #HV elim HV -V #V #_ #IHV #T #HT @(csn_ind_alt … HT) -T #T #HT #IHT
@csn_intro #X #H1 #H2
elim (cpr_inv_abbr1 … H1) -H1 *
[ #V0 #V1 #T1 #HLV0 #HLV01 #HLT1 #H destruct
]
qed.
-fact csn_appl_beta_aux: ∀L,W. L ⊢ ⬊* W → ∀U. L ⊢ ⬊* U →
- ∀V,T. U = ⓓV. T → L ⊢ ⬊* ⓐV. ⓛW. T.
-#L #W #H elim H -W #W #_ #IHW #X #H @(csn_ind_alt … H) -X #X #HVT #IHVT #V #T #H destruct
+fact csn_appl_beta_aux: ∀a,L,W. L ⊢ ⬊* W → ∀U. L ⊢ ⬊* U →
+ ∀V,T. U = ⓓ{a}V. T → L ⊢ ⬊* ⓐV. ⓛ{a}W. T.
+#a #L #W #H elim H -W #W #_ #IHW #X #H @(csn_ind_alt … H) -X #X #HVT #IHVT #V #T #H destruct
lapply (csn_fwd_pair_sn … HVT) #HV
lapply (csn_fwd_bind_dx … HVT) #HT -HVT
@csn_intro #X #H #H2
| -IHW -HLW0 -HV -HT * #H #HVT0 destruct
@(IHVT … HVT0) -IHVT -HVT0 // /2 width=1/
]
-| -IHW -IHVT -H2 #V0 #W0 #T0 #T1 #HLV0 #HLT01 #H1 #H2 destruct
+| -IHW -IHVT -H2 #b #V0 #W0 #T0 #T1 #HLV0 #HLT01 #H1 #H2 destruct
lapply (lcpr_cpr_trans (L. ⓓV) … HLT01) -HLT01 /2 width=1/ #HLT01
@csn_abbr /2 width=3/ -HV
@(csn_lcpr_conf (L. ⓓV)) /2 width=1/ -V0 /2 width=3/
-| -IHW -IHVT -HV -HT -H2 #V0 #V1 #W0 #W1 #T0 #T1 #_ #_ #_ #_ #H destruct
+| -IHW -IHVT -HV -HT -H2 #b #V0 #V1 #W0 #W1 #T0 #T1 #_ #_ #_ #_ #H destruct
]
qed.
(* Basic_1: was: sn3_beta *)
-lemma csn_appl_beta: ∀L,W. L ⊢ ⬊* W → ∀V,T. L ⊢ ⬊* ⓓV. T →
- L ⊢ ⬊* ⓐV. ⓛW. T.
+lemma csn_appl_beta: ∀a,L,W. L ⊢ ⬊* W → ∀V,T. L ⊢ ⬊* ⓓ{a}V. T →
+ L ⊢ ⬊* ⓐV. ⓛ{a}W. T.
/2 width=3/ qed.
-fact csn_appl_theta_aux: ∀L,U. L ⊢ ⬊* U → ∀V1,V2. ⇧[0, 1] V1 ≡ V2 →
- ∀V,T. U = ⓓV. ⓐV2. T → L ⊢ ⬊* ⓐV1. ⓓV. T.
-#L #X #H @(csn_ind_alt … H) -X #X #HVT #IHVT #V1 #V2 #HV12 #V #T #H destruct
+fact csn_appl_theta_aux: ∀a,L,U. L ⊢ ⬊* U → ∀V1,V2. ⇧[0, 1] V1 ≡ V2 →
+ ∀V,T. U = ⓓ{a}V. ⓐV2. T → L ⊢ ⬊* ⓐV1. ⓓ{a}V. T.
+#a #L #X #H @(csn_ind_alt … H) -X #X #HVT #IHVT #V1 #V2 #HV12 #V #T #H destruct
lapply (csn_fwd_pair_sn … HVT) #HV
lapply (csn_fwd_bind_dx … HVT) -HVT #HVT
@csn_intro #X #HL #H
[ #V3 #V4 #T3 #HV3 #HLV34 #HLT3 #H0 destruct
lapply (cpr_intro … HV3 HLV34) -HLV34 #HLV34
elim (lift_total V0 0 1) #V5 #HV05
- elim (term_eq_dec (ⓓV.ⓐV2.T) (ⓓV4.ⓐV5.T3))
+ elim (term_eq_dec (ⓓ{a}V.ⓐV2.T) (ⓓ{a}V4.ⓐV5.T3))
[ -IHVT #H0 destruct
elim (eq_false_inv_tpair_sn … H) -H
[ -HLV10 -HLV34 -HV3 -HLT3 -HVT
lapply (ltpr_cpr_trans (L. ⓓV) … HLT3) /2 width=1/ -HLT3 #HLT3
@(IHVT … H … HV05) -IHVT // -H -HV05 /3 width=1/
]
- | -H -IHVT #T0 #HLT0 #HT0
+ | -H -IHVT #T0 #HLT0 #HT0 #H0 destruct
lapply (csn_cpr_trans … HVT (ⓐV2.T0) ?) /2 width=1/ -T #HVT0
lapply (csn_inv_lift L … 1 HVT0 ? ? ?) -HVT0 [ /2 width=4/ |2,3: skip | /2 width=1/ ] -V2 -T0 #HVY
@(csn_cpr_trans … HVY) /2 width=1/
]
-| -HV -HV12 -HVT -IHVT -H #V0 #W0 #T0 #T1 #_ #_ #H destruct
-| -IHVT -H #V0 #V3 #W0 #W1 #T0 #T1 #HLV10 #HLW01 #HLT01 #HV03 #H1 #H2 destruct
+| -HV -HV12 -HVT -IHVT -H #b #V0 #W0 #T0 #T1 #_ #_ #H destruct
+| -IHVT -H #b #V0 #V3 #W0 #W1 #T0 #T1 #HLV10 #HLW01 #HLT01 #HV03 #H1 #H2 destruct
lapply (cpr_lift (L. ⓓW0) … HV12 … HV03 HLV10) -HLV10 -HV12 -HV03 /2 width=1/ #HLV23
lapply (lcpr_cpr_trans (L. ⓓW0) … HLT01) -HLT01 /2 width=1/ #HLT01
@csn_abbr /2 width=3/ -HV
]
qed.
-lemma csn_appl_theta: ∀V1,V2. ⇧[0, 1] V1 ≡ V2 →
- ∀L,V,T. L ⊢ ⬊* ⓓV. ⓐV2. T → L ⊢ ⬊* ⓐV1. ⓓV. T.
+lemma csn_appl_theta: ∀a,V1,V2. ⇧[0, 1] V1 ≡ V2 →
+ ∀L,V,T. L ⊢ ⬊* ⓓ{a}V. ⓐV2. T → L ⊢ ⬊* ⓐV1. ⓓ{a}V. T.
/2 width=5/ qed.
(* Basic_1: was only: sn3_appl_appl *)
]
qed.
-lemma csn_abst: ∀L,W. L ⊢ ⬊* W → ∀I,V,T. L. ⓑ{I} V ⊢ ⬊* T → L ⊢ ⬊* ⓛW. T.
-#L #W #HW elim HW -W #W #_ #IHW #I #V #T #HT @(csn_ind … HT) -T #T #HT #IHT
+lemma csn_abst: ∀a,L,W. L ⊢ ⬊* W → ∀I,V,T. L. ⓑ{I} V ⊢ ⬊* T → L ⊢ ⬊* ⓛ{a}W. T.
+#a #L #W #HW elim HW -W #W #_ #IHW #I #V #T #HT @(csn_ind … HT) -T #T #HT #IHT
@csn_intro #X #H1 #H2
elim (cpr_inv_abst1 … H1 I V) -H1
#W0 #T0 #HLW0 #HLT0 #H destruct
qed.
(* Basic_1: was: sn3_appls_beta *)
-lemma csn_applv_beta: ∀L,W. L ⊢ ⬊* W →
- ∀Vs,V,T. L ⊢ ⬊* ⒶVs.ⓓV.T →
- L ⊢ ⬊* ⒶVs. ⓐV.ⓛW. T.
-#L #W #HW #Vs elim Vs -Vs /2 width=1/ -HW
+lemma csn_applv_beta: ∀a,L,W. L ⊢ ⬊* W →
+ ∀Vs,V,T. L ⊢ ⬊* ⒶVs.ⓓ{a}V.T →
+ L ⊢ ⬊* ⒶVs. ⓐV.ⓛ{a}W. T.
+#a #L #W #HW #Vs elim Vs -Vs /2 width=1/ -HW
#V0 #Vs #IHV #V #T #H1T
lapply (csn_fwd_pair_sn … H1T) #HV0
lapply (csn_fwd_flat_dx … H1T) #H2T
qed.
(* Basic_1: was: sn3_appls_abbr *)
-lemma csn_applv_theta: ∀L,V1s,V2s. ⇧[0, 1] V1s ≡ V2s →
- ∀V,T. L ⊢ ⬊* ⓓV. ⒶV2s. T → L ⊢ ⬊* V →
- L ⊢ ⬊* ⒶV1s. ⓓV. T.
-#L #V1s #V2s * -V1s -V2s /2 width=1/
+lemma csn_applv_theta: ∀a,L,V1s,V2s. ⇧[0, 1] V1s ≡ V2s →
+ ∀V,T. L ⊢ ⬊* ⓓ{a}V. ⒶV2s. T → L ⊢ ⬊* V →
+ L ⊢ ⬊* ⒶV1s. ⓓ{a}V. T.
+#a #L #V1s #V2s * -V1s -V2s /2 width=1/
#V1s #V2s #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/
[ /3 width=1/
| /2 width=1/
| /2 width=6/
-| #L #V1 #V2 #HV12 #V #T #H #HVT
+| #L #V1 #V2 #HV12 #a #V #T #H #HVT
@(csn_applv_theta … HV12) -HV12 //
@(csn_abbr) //
| /2 width=1/
lemma xprs_refl: ∀h,g,L,T. ⦃h, L⦄ ⊢ T ➸*[g] T.
/2 width=1/ qed.
-lemma xprs_strap1: ∀h,g,L,T1,T,T2.
+axiom xprs_strap1: ∀h,g,L,T1,T,T2.
⦃h, L⦄ ⊢ T1 ➸*[g] T → ⦃h, L⦄ ⊢ T ➸[g] T2 → ⦃h, L⦄ ⊢ T1 ➸*[g] T2.
+(**) (* NTypeChecker failure
/2 width=3/ qed.
-
-lemma xprs_strap2: ∀h,g,L,T1,T,T2.
+*)
+axiom xprs_strap2: ∀h,g,L,T1,T,T2.
⦃h, L⦄ ⊢ T1 ➸[g] T → ⦃h, L⦄ ⊢ T ➸*[g] T2 → ⦃h, L⦄ ⊢ T1 ➸*[g] T2.
+(**) (* NTypeChecker failure
/2 width=3/ qed.
-
+*)
(* Basic inversion lemmas ***************************************************)
(*
axiom xprs_inv_cnf1: ∀L,T,U. L ⊢ T ➡* U → L ⊢ 𝐍⦃T⦄ → T = U.
--- /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/reducibility/xpr_aaa.ma".
+include "basic_2/computation/xprs.ma".
+
+(* EXTENDED PARALLEL COMPUTATION ON TERMS ***********************************)
+
+(* Properties on atomic arity assignment for terms **************************)
+
+lemma xprs_aaa: ∀h,g,L,T,A. L ⊢ T ⁝ A → ∀U. ⦃h, L⦄ ⊢ T ➸*[g] U → L ⊢ U ⁝ A.
+#h #g #L #T #A #HT #U #H @(xprs_ind … H) -U // /2 width=5/
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/computation/cprs.ma".
+include "basic_2/computation/xprs.ma".
+
+(* EXTENDED PARALLEL COMPUTATION ON TERMS ***********************************)
+
+(* properties on context sensitive parallel computation for terms ***********)
+
+lemma cprs_xprs: ∀h,g,L,T1,T2. L ⊢ T1 ➡* T2 → ⦃h, L⦄ ⊢ T1 ➸*[g] T2.
+#h #g #L #T1 #T2 #H @(cprs_ind … H) -T2 //
+/3 width=3 by xprs_strap1, cpr_xpr/ (**) (* NTypeChecker failure without trace *)
+qed.
(* Advanced forward lemmas **************************************************)
-lemma xprs_fwd_abst1: ∀h,g,L,V1,T1,U2. ⦃h, L⦄ ⊢ ⓛV1. T1 ➸*[g] U2 →
- ∃∃V2,T2. L ⊢ V1 ➡* V2 & U2 = ⓛV2. T2.
-#h #g #L #V1 #T1 #U2 #H @(xprs_ind … H) -U2 /2 width=4/
+lemma xprs_fwd_abst1: ∀h,g,a,L,V1,T1,U2. ⦃h, L⦄ ⊢ ⓛ{a}V1. T1 ➸*[g] U2 →
+ ∃∃V2,T2. L ⊢ V1 ➡* V2 & U2 = ⓛ{a}V2. T2.
+#h #g #a #L #V1 #T1 #U2 #H @(xprs_ind … H) -U2 /2 width=4/
#U #U2 #_ #HU2 * #V #T #HV1 #H destruct
elim (xpr_inv_abst1 … HU2) -HU2 #V2 #T2 #HV2 #_ #H destruct /3 width=4/
qed-.
inductive snv (h:sh) (g:sd h): lenv → predicate term ≝
| snv_sort: ∀L,k. snv h g L (⋆k)
| snv_lref: ∀I,L,K,V,i. ⇩[0, i] L ≡ K.ⓑ{I}V → snv h g K V → snv h g L (#i)
-| snv_bind: ∀I,L,V,T. snv h g L V → snv h g (L.ⓑ{I}V) T → snv h g L (ⓑ{I}V.T)
-| snv_appl: ∀L,V,W,W0,T,U,l. snv h g L V → snv h g L T →
+| snv_bind: ∀a,I,L,V,T. snv h g L V → snv h g (L.ⓑ{I}V) T → snv h g L (ⓑ{a,I}V.T)
+| snv_appl: ∀a,L,V,W,W0,T,U,l. snv h g L V → snv h g L T →
⦃h, L⦄ ⊢ V •[g, l + 1] W → L ⊢ W ➡* W0 →
- ⦃h, L⦄ ⊢ T ➸*[g] ⓛW0.U → snv h g L (ⓐV.T)
+ ⦃h, L⦄ ⊢ T ➸*[g] ⓛ{a}W0.U → snv h g L (ⓐV.T)
| snv_cast: ∀L,W,T,U,l. snv h g L W → snv h g L T →
⦃h, L⦄ ⊢ T •[g, l + 1] U → L ⊢ U ⬌* W → snv h g L (ⓝW.T)
.
(* Basic inversion lemmas ***************************************************)
-lemma snv_inv_bind_aux: ∀h,g,L,X. ⦃h, L⦄ ⊩ X :[g] → ∀I,V,T. X = ⓑ{I}V.T →
+lemma snv_inv_bind_aux: ∀h,g,L,X. ⦃h, L⦄ ⊩ X :[g] → ∀a,I,V,T. X = ⓑ{a,I}V.T →
⦃h, L⦄ ⊩ V :[g] ∧ ⦃h, L.ⓑ{I}V⦄ ⊩ T :[g].
#h #g #L #X * -L -X
-[ #L #k #I #V #T #H destruct
-| #I0 #L #K #V0 #i #_ #_ #I #V #T #H destruct
-| #I0 #L #V0 #T0 #HV0 #HT0 #I #V #T #H destruct /2 width=1/
-| #L #V0 #W0 #W00 #T0 #U0 #l #_ #_ #_ #_ #_ #I #V #T #H destruct
-| #L #W0 #T0 #U0 #l #_ #_ #_ #_ #I #V #T #H destruct
+[ #L #k #a #I #V #T #H destruct
+| #I0 #L #K #V0 #i #_ #_ #a #I #V #T #H destruct
+| #b #I0 #L #V0 #T0 #HV0 #HT0 #a #I #V #T #H destruct /2 width=1/
+| #b #L #V0 #W0 #W00 #T0 #U0 #l #_ #_ #_ #_ #_ #a #I #V #T #H destruct
+| #L #W0 #T0 #U0 #l #_ #_ #_ #_ #a #I #V #T #H destruct
]
qed.
-lemma snv_inv_bind: ∀h,g,I,L,V,T. ⦃h, L⦄ ⊩ ⓑ{I}V.T :[g] →
+lemma snv_inv_bind: ∀h,g,a,I,L,V,T. ⦃h, L⦄ ⊩ ⓑ{a,I}V.T :[g] →
⦃h, L⦄ ⊩ V :[g] ∧ ⦃h, L.ⓑ{I}V⦄ ⊩ T :[g].
-/2 width=3/ qed-.
+/2 width=4/ qed-.
lemma snv_inv_appl_aux: ∀h,g,L,X. ⦃h, L⦄ ⊩ X :[g] → ∀V,T. X = ⓐV.T →
- ∃∃W,W0,U,l. ⦃h, L⦄ ⊩ V :[g] & ⦃h, L⦄ ⊩ T :[g] &
+ ∃∃a,W,W0,U,l. ⦃h, L⦄ ⊩ V :[g] & ⦃h, L⦄ ⊩ T :[g] &
⦃h, L⦄ ⊢ V •[g, l + 1] W & L ⊢ W ➡* W0 &
- ⦃h, L⦄ ⊢ T ➸*[g] ⓛW0.U.
+ ⦃h, L⦄ ⊢ T ➸*[g] ⓛ{a}W0.U.
#h #g #L #X * -L -X
[ #L #k #V #T #H destruct
| #I #L #K #V0 #i #_ #_ #V #T #H destruct
-| #I #L #V0 #T0 #_ #_ #V #T #H destruct
-| #L #V0 #W0 #W00 #T0 #U0 #l #HV0 #HT0 #HVW0 #HW00 #HTU0 #V #T #H destruct /2 width=7/
+| #a #I #L #V0 #T0 #_ #_ #V #T #H destruct
+| #a #L #V0 #W0 #W00 #T0 #U0 #l #HV0 #HT0 #HVW0 #HW00 #HTU0 #V #T #H destruct /2 width=8/
| #L #W0 #T0 #U0 #l #_ #_ #_ #_ #V #T #H destruct
]
qed.
lemma snv_inv_appl: ∀h,g,L,V,T. ⦃h, L⦄ ⊩ ⓐV.T :[g] →
- ∃∃W,W0,U,l. ⦃h, L⦄ ⊩ V :[g] & ⦃h, L⦄ ⊩ T :[g] &
+ ∃∃a,W,W0,U,l. ⦃h, L⦄ ⊩ V :[g] & ⦃h, L⦄ ⊩ T :[g] &
⦃h, L⦄ ⊢ V •[g, l + 1] W & L ⊢ W ➡* W0 &
- ⦃h, L⦄ ⊢ T ➸*[g] ⓛW0.U.
+ ⦃h, L⦄ ⊢ T ➸*[g] ⓛ{a}W0.U.
/2 width=3/ qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/computation/csn_aaa.ma".
+include "basic_2/computation/xprs_aaa.ma".
+include "basic_2/computation/xprs_cprs.ma".
+include "basic_2/equivalence/lcpcs_aaa.ma".
+include "basic_2/dynamic/snv.ma".
+
+(* STRATIFIED NATIVE VALIDITY FOR TERMS *************************************)
+
+(* Properties on atomic arity assignment for terms **************************)
+
+lemma snv_aaa: ∀h,g,L,T. ⦃h, L⦄ ⊩ T :[g] → ∃A. L ⊢ T ⁝ A.
+#h #g #L #T #H elim H -L -T
+[ /2 width=2/
+| #I #L #K #V #i #HLK #_ * /3 width=6/
+| #a * #L #V #T #_ #_ * #B #HV * #A #HA /3 width=2/
+| #a #L #V #W #W0 #T #U #l #_ #_ #HVW #HW0 #HTU * #B #HV * #X #HT
+ lapply (xprs_aaa h g … HV W0 ?)
+ [ /3 width=3 by xprs_strap2, ssta_xpr, cprs_xprs/ ] -W #HW0 (**) (* NTypeChecker failure without trace *)
+ lapply (xprs_aaa … HT … HTU) -HTU #H
+ elim (aaa_inv_abst … H) -H #B0 #A #H1 #HU #H2 destruct
+ lapply (aaa_mono … H1 … HW0) -W0 #H destruct /3 width=4/
+| #L #W #T #U #l #_ #_ #HTU #HUW * #B #HW * #A #HT
+ lapply (aaa_cpcs_mono … HUW A … HW) -HUW /2 width=7/ -HTU #H destruct /3 width=3/
+]
+qed-.
+
+lemma snv_csn: ∀h,g,L,T. ⦃h, L⦄ ⊩ T :[g] → L ⊢ ⬊* T.
+#h #g #L #T #H elim (snv_aaa … H) -H /2 width=2/
+qed.
/3 width=8/
| lapply (ldrop_trans_ge … HLK … HK0 ?) -K // -Hid /3 width=8/
]
-| #I #K #V #T #_ #_ #IHV #IHT #L #d #e #HLK #X #H
+| #a #I #K #V #T #_ #_ #IHV #IHT #L #d #e #HLK #X #H
elim (lift_inv_bind1 … H) -H #W #U #HVW #HTU #H destruct
/4 width=4/
-| #K #V #V0 #V1 #T #T1 #l #_ #_ #HV0 #HV01 #HT1 #IHV #IHT #L #d #e #HLK #X #H
+| #a #K #V #V0 #V1 #T #T1 #l #_ #_ #HV0 #HV01 #HT1 #IHV #IHT #L #d #e #HLK #X #H
elim (lift_inv_flat1 … H) -H #W #U #HVW #HTU #H destruct
elim (lift_total V0 d e) #W0 #HVW0
elim (lift_total V1 d e) #W1 #HVW1
elim (lift_total T1 (d+1) e) #U1 #HTU1
- @(snv_appl … W0 … W1 … U1 l)
+ @(snv_appl … a … W0 … W1 … U1 l)
[ /2 width=4/ | /2 width=4/ | /2 width=9/ | /2 width=9/ ]
@(xprs_lift … HLK … HTU … HT1) /2 width=1/
| #K #V0 #T #V #l #_ #_ #HTV #HV0 #IHV0 #IHT #L #d #e #HLK #X #H
/3 width=8/
| lapply (ldrop_conf_ge … HLK … HL0 ?) -L // -Hid /3 width=8/
]
-| #I #L #W #U #_ #_ #IHW #IHU #K #d #e #HLK #X #H
+| #a #I #L #W #U #_ #_ #IHW #IHU #K #d #e #HLK #X #H
elim (lift_inv_bind2 … H) -H #V #T #HVW #HTU #H destruct /4 width=4/
-| #L #W #W0 #W1 #U #U1 #l #_ #_ #HW0 #HW01 #HU1 #IHW #IHU #K #d #e #HLK #X #H
+| #a #L #W #W0 #W1 #U #U1 #l #_ #_ #HW0 #HW01 #HU1 #IHW #IHU #K #d #e #HLK #X #H
elim (lift_inv_flat2 … H) -H #V #T #HVW #HTU #H destruct
elim (ssta_inv_lift1 … HW0 … HLK … HVW) -HW0 #V0 #HV0 #HVW0
elim (cprs_inv_lift1 … HLK … HVW0 … HW01) -W0 #V1 #HVW1 #HV01
elim (xprs_inv_lift1 … HLK … HTU … HU1) -HU1 #X #H #HTU
elim (lift_inv_bind2 … H) -H #Y #T1 #HY #HTU1 #H destruct
- lapply (lift_inj … HY … HVW1) -HY #H destruct /3 width=7/
+ lapply (lift_inj … HY … HVW1) -HY #H destruct /3 width=8/
| #L #W0 #U #W #l #_ #_ #HUW #HW0 #IHW0 #IHU #K #d #e #HLK #X #H
elim (lift_inv_flat2 … H) -H #V0 #T #HVW0 #HTU #H destruct
elim (ssta_inv_lift1 … HUW … HLK … HTU) -HUW #V #HTV #HVW
qed-.
(* Basic_1: was: pc3_gen_sort_abst *)
-lemma cpcs_inv_sort_abst: ∀L,W,T,k. L ⊢ ⋆k ⬌* ⓛW.T → ⊥.
-#L #W #T #k #H
+lemma cpcs_inv_sort_abst: ∀a,L,W,T,k. L ⊢ ⋆k ⬌* ⓛ{a}W.T → ⊥.
+#a #L #W #T #k #H
elim (cpcs_inv_cprs … H) -H #X #H1
>(cprs_inv_sort1 … H1) -X #H2
elim (cprs_inv_abst1 Abst W … H2) -H2 #W0 #T0 #_ #_ #H destruct
qed-.
(* Basic_1: was: pc3_gen_abst *)
-lemma cpcs_inv_abst: ∀L,W1,W2,T1,T2. L ⊢ ⓛW1.T1 ⬌* ⓛW2.T2 → ∀I,V.
- L ⊢ W1 ⬌* W2 ∧ L. ②{I}V ⊢ T1 ⬌* T2.
-#L #W1 #W2 #T1 #T2 #H #I #V
+lemma cpcs_inv_abst: ∀a1,a2,L,W1,W2,T1,T2. L ⊢ ⓛ{a1}W1.T1 ⬌* ⓛ{a2}W2.T2 → ∀I,V.
+ ∧∧ L ⊢ W1 ⬌* W2 & L. ②{I}V ⊢ T1 ⬌* T2 & a1 = a2.
+#a1 #a2 #L #W1 #W2 #T1 #T2 #H #I #V
elim (cpcs_inv_cprs … H) -H #T #H1 #H2
elim (cprs_inv_abst1 I V … H1) -H1 #W0 #T0 #HW10 #HT10 #H destruct
elim (cprs_inv_abst1 I V … H2) -H2 #W #T #HW2 #HT2 #H destruct /3 width=3/
qed-.
(* Basic_1: was: pc3_gen_abst_shift *)
-lemma cpcs_inv_abst_shift: ∀L,W1,W2,T1,T2. L ⊢ ⓛW1.T1 ⬌* ⓛW2.T2 → ∀W.
- L ⊢ W1 ⬌* W2 ∧ L. ⓛW ⊢ T1 ⬌* T2.
-#L #W1 #W2 #T1 #T2 #H #W
+lemma cpcs_inv_abst_shift: ∀a1,a2,L,W1,W2,T1,T2. L ⊢ ⓛ{a1}W1.T1 ⬌* ⓛ{a2}W2.T2 → ∀W.
+ ∧∧ L ⊢ W1 ⬌* W2 & L. ⓛW ⊢ T1 ⬌* T2 & a1 = a2.
+#a1 #a2 #L #W1 #W2 #T1 #T2 #H #W
lapply (cpcs_inv_abst … H Abst W) -H //
qed.
-lemma cpcs_inv_abst1: ∀L,W1,T1,T. L ⊢ ⓛW1.T1 ⬌* T →
- ∃∃W2,T2. L ⊢ T ➡* ⓛW2.T2 & L ⊢ ⓛW1.T1 ➡* ⓛW2.T2.
-#L #W1 #T1 #T #H
+lemma cpcs_inv_abst1: ∀a,L,W1,T1,T. L ⊢ ⓛ{a}W1.T1 ⬌* T →
+ ∃∃W2,T2. L ⊢ T ➡* ⓛ{a}W2.T2 & L ⊢ ⓛ{a}W1.T1 ➡* ⓛ{a}W2.T2.
+#a #L #W1 #T1 #T #H
elim (cpcs_inv_cprs … H) -H #X #H1 #H2
elim (cprs_inv_abst1 Abst W1 … H1) -H1 #W2 #T2 #HW12 #HT12 #H destruct
@(ex2_2_intro … H2) -H2 /2 width=2/ (**) (* explicit constructor, /3 width=6/ is slow *)
qed-.
-lemma cpcs_inv_abst2: ∀L,W1,T1,T. L ⊢ T ⬌* ⓛW1.T1 →
- ∃∃W2,T2. L ⊢ T ➡* ⓛW2.T2 & L ⊢ ⓛW1.T1 ➡* ⓛW2.T2.
+lemma cpcs_inv_abst2: ∀a,L,W1,T1,T. L ⊢ T ⬌* ⓛ{a}W1.T1 →
+ ∃∃W2,T2. L ⊢ T ➡* ⓛ{a}W2.T2 & L ⊢ ⓛ{a}W1.T1 ➡* ⓛ{a}W2.T2.
/3 width=1 by cpcs_inv_abst1, cpcs_sym/ qed-.
(* Basic_1: was: pc3_gen_lift *)
∀I. L ⊢ ⓕ{I}V1. T1 ⬌* ⓕ{I}V2. T2.
/3 width=1/ qed.
-lemma cpcs_abst: ∀L,V1,V2. L ⊢ V1 ⬌* V2 →
- ∀V,T1,T2. L.ⓛV ⊢ T1 ⬌* T2 → L ⊢ ⓛV1. T1 ⬌* ⓛV2. T2.
-#L #V1 #V2 #HV12 #V #T1 #T2 #HT12
+lemma cpcs_abst: ∀a,L,V1,V2. L ⊢ V1 ⬌* V2 →
+ ∀V,T1,T2. L.ⓛV ⊢ T1 ⬌* T2 → L ⊢ ⓛ{a}V1. T1 ⬌* ⓛ{a}V2. T2.
+#a #L #V1 #V2 #HV12 #V #T1 #T2 #HT12
elim (cpcs_inv_cprs … HV12) -HV12
elim (cpcs_inv_cprs … HT12) -HT12
/3 width=6 by cprs_div, cprs_abst/ (**) (* /3 width=6/ is a bit slow *)
qed.
-lemma cpcs_abbr_dx: ∀L,V,T1,T2. L.ⓓV ⊢ T1 ⬌* T2 → L ⊢ ⓓV. T1 ⬌* ⓓV. T2.
-#L #V #T1 #T2 #HT12
+lemma cpcs_abbr_dx: ∀a,L,V,T1,T2. L.ⓓV ⊢ T1 ⬌* T2 → L ⊢ ⓓ{a}V. T1 ⬌* ⓓ{a}V. T2.
+#a #L #V #T1 #T2 #HT12
elim (cpcs_inv_cprs … HT12) -HT12 /3 width=5 by cprs_div, cprs_abbr1/ (**) (* /3 width=5/ is a bit slow *)
qed.
-lemma cpcs_bind_dx: ∀I,L,V,T1,T2. L.ⓑ{I}V ⊢ T1 ⬌* T2 →
- L ⊢ ⓑ{I}V. T1 ⬌* ⓑ{I}V. T2.
-* /2 width=1/ /2 width=2/ qed.
+lemma cpcs_bind_dx: ∀a,I,L,V,T1,T2. L.ⓑ{I}V ⊢ T1 ⬌* T2 →
+ L ⊢ ⓑ{a,I}V. T1 ⬌* ⓑ{a,I}V. T2.
+#a * /2 width=1/ /2 width=2/ qed.
-lemma cpcs_abbr_sn: ∀L,V1,V2,T. L ⊢ V1 ⬌* V2 → L ⊢ ⓓV1. T ⬌* ⓓV2. T.
-#L #V1 #V2 #T #HV12
+lemma cpcs_abbr_sn: ∀a,L,V1,V2,T. L ⊢ V1 ⬌* V2 → L ⊢ ⓓ{a}V1. T ⬌* ⓓ{a}V2. T.
+#a #L #V1 #V2 #T #HV12
elim (cpcs_inv_cprs … HV12) -HV12 /3 width=5 by cprs_div, cprs_abbr1/ (**) (* /3 width=5/ is a bit slow *)
qed.
-lemma cpcs_bind_sn: ∀I,L,V1,V2,T. L ⊢ V1 ⬌* V2 → L ⊢ ⓑ{I}V1. T ⬌* ⓑ{I}V2. T.
-* /2 width=1/ /2 width=2/ qed.
+lemma cpcs_bind_sn: ∀a,I,L,V1,V2,T. L ⊢ V1 ⬌* V2 → L ⊢ ⓑ{a,I}V1. T ⬌* ⓑ{a,I}V2. T.
+#a * /2 width=1/ /2 width=2/ qed.
-lemma cpcs_beta_dx: ∀L,V1,V2,W,T1,T2.
- L ⊢ V1 ➡ V2 → L.ⓛW ⊢ T1 ⬌* T2 → L ⊢ ⓐV1.ⓛW.T1 ⬌* ⓓV2.T2.
-#L #V1 #V2 #W #T1 #T2 #HV12 #HT12
+lemma cpcs_beta_dx: ∀a,L,V1,V2,W,T1,T2.
+ L ⊢ V1 ➡ V2 → L.ⓛW ⊢ T1 ⬌* T2 → L ⊢ ⓐV1.ⓛ{a}W.T1 ⬌* ⓓ{a}V2.T2.
+#a #L #V1 #V2 #W #T1 #T2 #HV12 #HT12
elim (cpcs_inv_cprs … HT12) -HT12 #T #HT1 #HT2
-lapply (cprs_beta_dx … HV12 HT1) -HV12 -HT1 #HT1
+lapply (cprs_beta_dx … HV12 HT1 a) -HV12 -HT1 #HT1
lapply (cprs_lsubs_trans … HT2 (L.ⓓV2) ?) -HT2 /2 width=1/ #HT2
@(cprs_div … HT1) /2 width=1/
qed.
-lemma cpcs_beta_dx_tpr_rev: ∀L,V1,V2,W,T1,T2.
+lemma cpcs_beta_dx_tpr_rev: ∀a,L,V1,V2,W,T1,T2.
V1 ➡ V2 → L.ⓛW ⊢ T2 ⬌* T1 →
- L ⊢ ⓓV2.T2 ⬌* ⓐV1.ⓛW.T1.
+ L ⊢ ⓓ{a}V2.T2 ⬌* ⓐV1.ⓛ{a}W.T1.
/4 width=1/ qed.
(* Note: it does not hold replacing |L1| with |L2| *)
theorem cpcs_canc_dx: ∀L,T,T1,T2. L ⊢ T1 ⬌* T → L ⊢ T2 ⬌* T → L ⊢ T1 ⬌* T2.
/3 width=3 by cpcs_trans, cprs_comm/ qed. (**) (* /3 width=3/ is too slow *)
-lemma cpcs_abbr1: ∀L,V1,V2. L ⊢ V1 ⬌* V2 → ∀T1,T2. L.ⓓV1 ⊢ T1 ⬌* T2 →
- L ⊢ ⓓV1. T1 ⬌* ⓓV2. T2.
-#L #V1 #V2 #HV12 #T1 #T2 #HT12
-@(cpcs_trans … (ⓓV1.T2)) /2 width=1/
+lemma cpcs_abbr1: ∀a,L,V1,V2. L ⊢ V1 ⬌* V2 → ∀T1,T2. L.ⓓV1 ⊢ T1 ⬌* T2 →
+ L ⊢ ⓓ{a}V1. T1 ⬌* ⓓ{a}V2. T2.
+#a #L #V1 #V2 #HV12 #T1 #T2 #HT12
+@(cpcs_trans … (ⓓ{a}V1.T2)) /2 width=1/
qed.
-lemma cpcs_abbr2: ∀L,V1,V2. L ⊢ V1 ⬌* V2 → ∀T1,T2. L.ⓓV2 ⊢ T1 ⬌* T2 →
- L ⊢ ⓓV1. T1 ⬌* ⓓV2. T2.
-#L #V1 #V2 #HV12 #T1 #T2 #HT12
-@(cpcs_trans … (ⓓV2.T1)) /2 width=1/
+lemma cpcs_abbr2: ∀a,L,V1,V2. L ⊢ V1 ⬌* V2 → ∀T1,T2. L.ⓓV2 ⊢ T1 ⬌* T2 →
+ L ⊢ ⓓ{a}V1. T1 ⬌* ⓓ{a}V2. T2.
+#a #L #V1 #V2 #HV12 #T1 #T2 #HT12
+@(cpcs_trans … (ⓓ{a}V2.T1)) /2 width=1/
qed.
(* CONTEXT-SENSITIVE PARALLEL EQUIVALENCE ON TERMS **************************)
(* Advanced inversion lemmas ************************************************)
-
-lemma cprs_inv_appl1_cpcs: ∀L,V1,T1,U2. L ⊢ ⓐV1. T1 ➡* U2 → (
- ∃∃V2,T2. L ⊢ V1 ➡* V2 & L ⊢ T1 ➡* T2 &
- L ⊢ U2 ➡* ⓐV2. T2
+(* does not holds
+axiom cprs_inv_appl1_cpcs: ∀L,V1,T1,U2. L ⊢ ⓐV1. T1 ➡* U2 → (
+ ∃∃V2,T2. L ⊢ V1 ➡* V2 & L ⊢ T1 ➡* T2 &
+ L ⊢ U2 ➡* ⓐV2. T2
) ∨
- ∃∃V2,W,T. L ⊢ V1 ➡* V2 &
- L ⊢ T1 ➡* ⓛW. T & L ⊢ ⓓV2. T ⬌* U2.
+ ∃∃a,V2,W,T. L ⊢ V1 ➡* V2 &
+ L ⊢ T1 ➡* ⓛ{a}W. T & L ⊢ ⓓ{a}V2. T ⬌* U2.
#L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5/
#U #U2 #_ #HU2 * *
[ #V0 #T0 #HV10 #HT10 #HUT0
elim (cprs_strip … HUT0 … HU2) -U #U #H #HU2
elim (cpr_inv_appl1 … H) -H *
[ #V2 #T2 #HV02 #HT02 #H destruct /4 width=5/
- | #V2 #W2 #T #T2 #HV02 #HT2 #H1 #H2 destruct
+ | #b #V2 #W2 #T #T2 #HV02 #HT2 #H1 #H2 destruct
lapply (cprs_strap1 … HV10 HV02) -V0 #HV12
- lapply (cprs_div ? (ⓓV2.T) ? ? ? HU2) -HU2 /2 width=1/ /3 width=6/
- | #V #V2 #W0 #W2 #T #T2 #HV0 #HW02 #HT2 #HV2 #H1 #H2 destruct
+ lapply (cprs_div ? (ⓓ{b}V2.T) ? ? ? HU2) -HU2 /2 width=1/ /3 width=7/
+ | #b #V #V2 #W0 #W2 #T #T2 #HV0 #HW02 #HT2 #HV2 #H1 #H2 destruct
lapply (cprs_strap1 … HV10 HV0) -V0 #HV1
- lapply (cprs_trans … HT10 (ⓓW2.T2) ?) -HT10 /2 width=1/ -W0 -T #HT1
+ lapply (cprs_trans … HT10 (ⓓ{b}W2.T2) ?) -HT10 /2 width=1/ -W0 -T #HT1
elim (sfr_delift (L.ⓓW2) (ⓐV2.T2) 0 1 ? ?) // #X #H1
lapply (cprs_zeta_delift … H1) #H2
lapply (cprs_trans … HU2 … H2) -HU2 -H2 #HU2T3
| /4 width=8/
]
qed-.
-
-lemma cprs_inv_appl_abst: ∀L,V,T,W,U. L ⊢ ⓐV.T ➡* ⓛW.U →
+*)
+(* maybe holds
+axiom cprs_inv_appl_abst: ∀L,V,T,W,U. L ⊢ ⓐV.T ➡* ⓛW.U →
∃∃W0,T0,V1,T1. L ⊢ T ➡* ⓛW0.T0 &
L ⊢ ⓓV.T0 ➡* ⓛV1.T1 &
L ⊢ ⓛW.U ➡* ⓛV1.T1.
lapply (cprs_trans … (ⓓV.T0) … H1) -H1 /2 width=1/ -V0 /2 width=7/
]
qed-.
-
+*)
(* Properties on inverse basic term relocation ******************************)
lemma cpcs_zeta_delift_comm: ∀L,V,T1,T2. L.ⓓV ⊢ ▼*[O, 1] T1 ≡ T2 →
- L ⊢ T2 ⬌* ⓓV.T1.
+ L ⊢ T2 ⬌* +ⓓV.T1.
/3 width=1/ qed.
(* Basic_1: was only: pc3_gen_cabbr *)
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+notation "hvbox( ⦃ h , break L ⦄ ⊢ break term 46 T1 •* break [ g ] break term 46 T2 )"
+ non associative with precedence 45
+ for @{ 'StaticTypeStar $h $g $L $T1 $T2 }.
+
+include "basic_2/static/ssta.ma".
+
+(* ITERATED STRATIFIED STATIC TYPE ASSIGNMENTON TERMS ***********************)
+
+inductive sstas (h:sh) (g:sd h) (L:lenv): relation term ≝
+| sstas_refl: ∀T,U. ⦃h, L⦄ ⊢ T •[g, 0] U → sstas h g L T T
+| sstas_step: ∀T,U1,U2,l. ⦃h, L⦄ ⊢ T •[g, l+1] U1 → sstas h g L U1 U2 →
+ sstas h g L T U2.
+
+interpretation "stratified unwind (term)"
+ 'StaticTypeStar h g L T U = (sstas h g L T U).
+
+(* Basic eliminators ********************************************************)
+
+fact sstas_ind_alt_aux: ∀h,g,L,U2. ∀R:predicate term.
+ (∀T. ⦃h, L⦄ ⊢ U2 •[g , 0] T → R U2) →
+ (∀T,U1,l. ⦃h, L⦄ ⊢ T •[g, l + 1] U1 →
+ ⦃h, L⦄ ⊢ U1 •* [g] U2 → R U1 → R T
+ ) →
+ ∀T,U. ⦃h, L⦄ ⊢ T •*[g] U → U = U2 → R T.
+#h #g #L #U2 #R #H1 #H2 #T #U #H elim H -H -T -U /2 width=2/ /3 width=5/
+qed-.
+
+lemma sstas_ind_alt: ∀h,g,L,U2. ∀R:predicate term.
+ (∀T. ⦃h, L⦄ ⊢ U2 •[g , 0] T → R U2) →
+ (∀T,U1,l. ⦃h, L⦄ ⊢ T •[g, l + 1] U1 →
+ ⦃h, L⦄ ⊢ U1 •* [g] U2 → R U1 → R T
+ ) →
+ ∀T. ⦃h, L⦄ ⊢ T •*[g] U2 → R T.
+/3 width=9 by sstas_ind_alt_aux/ qed-.
+
+(* Basic inversion lemmas ***************************************************)
+
+fact sstas_inv_sort1_aux: ∀h,g,L,T,U. ⦃h, L⦄ ⊢ T •*[g] U → ∀k. T = ⋆k →
+ ∀l. deg h g k l → U = ⋆((next h)^l k).
+#h #g #L #T #U #H @(sstas_ind_alt … H) -T
+[ #U0 #HU0 #k #H #l #Hkl destruct
+ elim (ssta_inv_sort1 … HU0) -L #HkO #_ -U0
+ >(deg_mono … Hkl HkO) -g -l //
+| #T0 #U0 #l0 #HTU0 #_ #IHU0 #k #H #l #Hkl destruct
+ elim (ssta_inv_sort1 … HTU0) -L #HkS #H destruct
+ lapply (deg_mono … Hkl HkS) -Hkl #H destruct
+ >(IHU0 (next h k) ? l0) -IHU0 // /2 width=1/ >iter_SO >iter_n_Sm //
+]
+qed.
+
+lemma sstas_inv_sort1: ∀h,g,L,U,k. ⦃h, L⦄ ⊢ ⋆k •*[g] U → ∀l. deg h g k l →
+ U = ⋆((next h)^l k).
+/2 width=6/ qed-.
+
+fact sstas_inv_bind1_aux: ∀h,g,L,T,U. ⦃h, L⦄ ⊢ T •*[g] U →
+ ∀J,X,Y. T = ⓑ{J}Y.X →
+ ∃∃Z. ⦃h, L.ⓑ{J}Y⦄ ⊢ X •*[g] Z & U = ⓑ{J}Y.Z.
+#h #g #L #T #U #H @(sstas_ind_alt … H) -T
+[ #U0 #HU0 #J #X #Y #H destruct
+ elim (ssta_inv_bind1 … HU0) -HU0 #X0 #HX0 #H destruct /3 width=3/
+| #T0 #U0 #l #HTU0 #_ #IHU0 #J #X #Y #H destruct
+ elim (ssta_inv_bind1 … HTU0) -HTU0 #X0 #HX0 #H destruct
+ elim (IHU0 J X0 Y ?) -IHU0 // #X1 #HX01 #H destruct /3 width=4/
+]
+qed.
+
+lemma sstas_inv_bind1: ∀h,g,J,L,Y,X,U. ⦃h, L⦄ ⊢ ⓑ{J}Y.X •*[g] U →
+ ∃∃Z. ⦃h, L.ⓑ{J}Y⦄ ⊢ X •*[g] Z & U = ⓑ{J}Y.Z.
+/2 width=3/ qed-.
+
+fact sstas_inv_appl1_aux: ∀h,g,L,T,U. ⦃h, L⦄ ⊢ T •*[g] U → ∀X,Y. T = ⓐY.X →
+ ∃∃Z. ⦃h, L⦄ ⊢ X •*[g] Z & U = ⓐY.Z.
+#h #g #L #T #U #H @(sstas_ind_alt … H) -T
+[ #U0 #HU0 #X #Y #H destruct
+ elim (ssta_inv_appl1 … HU0) -HU0 #X0 #HX0 #H destruct /3 width=3/
+| #T0 #U0 #l #HTU0 #_ #IHU0 #X #Y #H destruct
+ elim (ssta_inv_appl1 … HTU0) -HTU0 #X0 #HX0 #H destruct
+ elim (IHU0 X0 Y ?) -IHU0 // #X1 #HX01 #H destruct /3 width=4/
+]
+qed.
+
+lemma sstas_inv_appl1: ∀h,g,L,Y,X,U. ⦃h, L⦄ ⊢ ⓐY.X •*[g] U →
+ ∃∃Z. ⦃h, L⦄ ⊢ X •*[g] Z & U = ⓐY.Z.
+/2 width=3/ qed-.
+
+(* Basic forward lemmas *****************************************************)
+
+lemma sstas_fwd_correct: ∀h,g,L,T,U. ⦃h, L⦄ ⊢ T •*[g] U →
+ ∃∃W. ⦃h, L⦄ ⊢ U •[g, 0] W & ⦃h, L⦄ ⊢ U •*[g] U.
+#h #g #L #T #U #H @(sstas_ind_alt … H) -T /2 width=1/ /3 width=2/
+qed-.
+
+(* Basic_1: removed theorems 7:
+ sty1_bind sty1_abbr sty1_appl sty1_cast2
+ sty1_lift sty1_correct sty1_trans
+*)
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/static/ssta_lift.ma".
+include "basic_2/unwind/sstas.ma".
+
+(* ITERATED STRATIFIED STATIC TYPE ASSIGNMENTON TERMS ***********************)
+
+(* Advanced properties ******************************************************)
+
+lemma sstas_total_S: ∀h,g,L,l,T,U. ⦃h, L⦄ ⊢ T•[g, l + 1]U →
+ ∃∃W. ⦃h, L⦄ ⊢ T •*[g] W & ⦃h, L⦄ ⊢ U •*[g] W.
+#h #g #L #l @(nat_ind_plus … l) -l
+[ #T #U #HTU
+ elim (ssta_fwd_correct … HTU) /4 width=4/
+| #l #IHl #T #U #HTU
+ elim (ssta_fwd_correct … HTU) <minus_plus_m_m #V #HUV
+ elim (IHl … HUV) -IHl -HUV /3 width=4/
+]
+qed-.
+
+(* Properties on relocation *************************************************)
+
+lemma sstas_lift: ∀h,g,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 •*[g] U1 →
+ ∀L2,d,e. ⇩[d, e] L2 ≡ L1 → ∀T2. ⇧[d, e] T1 ≡ T2 →
+ ∀U2. ⇧[d, e] U1 ≡ U2 → ⦃h, L2⦄ ⊢ T2 •*[g] U2.
+#h #g #L1 #T1 #U1 #H @(sstas_ind_alt … H) -T1
+[ #T1 #HUT1 #L2 #d #e #HL21 #X #HX #U2 #HU12
+ >(lift_mono … HX … HU12) -X
+ elim (lift_total T1 d e) /3 width=10/
+| #T0 #U0 #l0 #HTU0 #_ #IHU01 #L2 #d #e #HL21 #T2 #HT02 #U2 #HU12
+ elim (lift_total U0 d e) /3 width=10/
+]
+qed.
+
+lemma sstas_inv_lift1: ∀h,g,L2,T2,U2. ⦃h, L2⦄ ⊢ T2 •*[g] U2 →
+ ∀L1,d,e. ⇩[d, e] L2 ≡ L1 → ∀T1. ⇧[d, e] T1 ≡ T2 →
+ ∃∃U1. ⦃h, L1⦄ ⊢ T1 •*[g] U1 & ⇧[d, e] U1 ≡ U2.
+#h #g #L2 #T2 #U2 #H @(sstas_ind_alt … H) -T2
+[ #T2 #HUT2 #L1 #d #e #HL21 #U1 #HU12
+ elim (ssta_inv_lift1 … HUT2 … HL21 … HU12) -HUT2 -HL21 /3 width=3/
+| #T0 #U0 #l0 #HTU0 #_ #IHU01 #L1 #d #e #HL21 #U1 #HU12
+ elim (ssta_inv_lift1 … HTU0 … HL21 … HU12) -HTU0 -HU12 #U #HU1 #HU0
+ elim (IHU01 … HL21 … HU0) -IHU01 -HL21 -U0 /3 width=4/
+]
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/static/ssta_ltpss.ma".
+include "basic_2/unwind/sstas.ma".
+
+(* ITERATED STRATIFIED STATIC TYPE ASSIGNMENTON TERMS ***********************)
+
+(* Properties about parallel unfold *****************************************)
+
+lemma sstas_ltpss_tpss_conf: ∀h,g,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 •*[g] U1 →
+ ∀L2,d,e. L1 ▶* [d, e] L2 →
+ ∀T2. L2 ⊢ T1 ▶* [d, e] T2 →
+ ∃∃U2. ⦃h, L2⦄ ⊢ T2 •*[g] U2 &
+ L2 ⊢ U1 ▶* [d, e] U2.
+#h #g #L1 #T1 #U1 #H @(sstas_ind_alt … H) -T1
+[ #T1 #HUT1 #L2 #d #e #HL12 #U2 #HU12
+ elim (ssta_ltpss_tpss_conf … HUT1 … HL12 … HU12) -HUT1 -HL12 /3 width=3/
+| #T0 #U0 #l0 #HTU0 #_ #IHU01 #L2 #d #e #HL12 #T #HT0
+ elim (ssta_ltpss_tpss_conf … HTU0 … HL12 … HT0) -HTU0 -HT0 #U #HTU #HU0
+ elim (IHU01 … HL12 … HU0) -IHU01 -HL12 -U0 /3 width=4/
+]
+qed.
+
+lemma sstas_ltpss_tps_conf: ∀h,g,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 •*[g] U1 →
+ ∀L2,d,e. L1 ▶* [d, e] L2 →
+ ∀T2. L2 ⊢ T1 ▶ [d, e] T2 →
+ ∃∃U2. ⦃h, L2⦄ ⊢ T2 •*[g] U2 & L2 ⊢ U1 ▶* [d, e] U2.
+/3 width=5/ qed.
+
+lemma sstas_ltpss_conf: ∀h,g,L1,T,U1. ⦃h, L1⦄ ⊢ T •*[g] U1 →
+ ∀L2,d,e. L1 ▶* [d, e] L2 →
+ ∃∃U2. ⦃h, L2⦄ ⊢ T •*[g] U2 & L2 ⊢ U1 ▶* [d, e] U2.
+/2 width=5/ qed.
+
+lemma sstas_tpss_conf: ∀h,g,L,T1,U1. ⦃h, L⦄ ⊢ T1 •*[g] U1 →
+ ∀T2,d,e. L ⊢ T1 ▶* [d, e] T2 →
+ ∃∃U2. ⦃h, L⦄ ⊢ T2 •*[g] U2 & L ⊢ U1 ▶* [d, e] U2.
+/2 width=5/ qed.
+
+lemma sstas_tps_conf: ∀h,g,L,T1,U1. ⦃h, L⦄ ⊢ T1 •*[g] U1 →
+ ∀T2,d,e. L ⊢ T1 ▶ [d, e] T2 →
+ ∃∃U2. ⦃h, L⦄ ⊢ T2 •*[g] U2 & L ⊢ U1 ▶* [d, e] U2.
+/2 width=5/ qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/unfold/delift_lift.ma".
+include "basic_2/static/ssta_ssta.ma".
+include "basic_2/unwind/sstas_lift.ma".
+
+(* ITERATED STRATIFIED STATIC TYPE ASSIGNMENTON TERMS ***********************)
+
+(* Advanced inversion lemmas ************************************************)
+
+lemma sstas_inv_O: ∀h,g,L,T,U. ⦃h, L⦄ ⊢ T •*[g] U →
+ ∀T0. ⦃h, L⦄ ⊢ T •[g , 0] T0 → U = T.
+#h #g #L #T #U #H @(sstas_ind_alt … H) -T //
+#T0 #U0 #l0 #HTU0 #_ #_ #T1 #HT01
+elim (ssta_mono … HTU0 … HT01) <plus_n_Sm #H destruct
+qed-.
+
+lemma sstas_inv_S: ∀h,g,L,T,U. ⦃h, L⦄ ⊢ T •*[g] U →
+ ∀T0,l. ⦃h, L⦄ ⊢ T •[g , l+1] T0 → ⦃h, L⦄ ⊢ T0 •*[g] U.
+#h #g #L #T #U #H @(sstas_ind_alt … H) -T
+[ #U0 #HU0 #T #l #HUT
+ elim (ssta_mono … HUT … HU0) <plus_n_Sm #H destruct
+| #T0 #U0 #l0 #HTU0 #HU0 #_ #T #l #HT0
+ elim (ssta_mono … HT0 … HTU0) -T0 #_ #H destruct -l0 //
+]
+qed-.
+
+(* Main properties **********************************************************)
+
+theorem sstas_mono: ∀h,g,L,T,U1. ⦃h, L⦄ ⊢ T •*[g] U1 →
+ ∀U2. ⦃h, L⦄ ⊢ T •*[g] U2 → U1 = U2.
+#h #g #L #T #U1 #H @(sstas_ind_alt … H) -T
+[ #T1 #HUT1 #U2 #HU12
+ >(sstas_inv_O … HU12 … HUT1) -h -L -T1 -U2 //
+| #T0 #U0 #l0 #HTU0 #_ #IHU01 #U2 #HU12
+ lapply (sstas_inv_S … HU12 … HTU0) -T0 -l0 /2 width=1/
+]
+qed-.
+
+(* More advancd inversion lemmas ********************************************)
+
+fact sstas_inv_lref1_aux: ∀h,g,L,T,U. ⦃h, L⦄ ⊢ T •*[g] U → ∀j. T = #j →
+ ∃∃I,K,V,W. ⇩[0, j] L ≡ K. ⓑ{I}V & ⦃h, K⦄ ⊢ V •*[g] W &
+ L ⊢ ▼*[0, j + 1] U ≡ W.
+#h #g #L #T #U #H @(sstas_ind_alt … H) -T
+[ #T #HUT #j #H destruct
+ elim (ssta_inv_lref1 … HUT) -HUT * #K #V #W [2: #l] #HLK #HVW #HVT
+ [ <plus_n_Sm #H destruct
+ | /3 width=12/
+ ]
+| #T0 #U0 #l0 #HTU0 #HU0 #_ #j #H destruct
+ elim (ssta_inv_lref1 … HTU0) -HTU0 * #K #V #W [2: #l] #HLK #HVW #HVU0
+ [ #_ -HVW
+ lapply (ldrop_fwd_ldrop2 … HLK) #H
+ elim (sstas_inv_lift1 … HU0 … H … HVU0) -HU0 -H -HVU0 /3 width=7/
+ | elim (sstas_total_S … HVW) -HVW #T #HVT #HWT
+ lapply (ldrop_fwd_ldrop2 … HLK) #H
+ elim (sstas_inv_lift1 … HU0 … H … HVU0) -HU0 -H -HVU0 #X #HWX
+ >(sstas_mono … HWX … HWT) -X -W /3 width=7/
+ ]
+]
+qed-.
let rec shift L T on L ≝ match L with
[ LAtom ⇒ T
-| LPair L I V ⇒ shift L (ⓑ{I} V. T)
+| LPair L I V ⇒ shift L (-ⓑ{I} V. T)
].
interpretation "shift (closure)" 'Append L T = (shift L T).
∀L,T. R L T.
(* Basic_1: was: flt_shift *)
-lemma cw_shift: ∀K,I,V,T. #[K. ⓑ{I} V, T] < #[K, ②{I} V. T].
+lemma cw_shift: ∀a,K,I,V,T. #[K. ⓑ{I} V, T] < #[K, ⓑ{a,I} V. T].
normalize //
qed.
#I1 #I2 #L #V1 #V2 #T normalize in ⊢ (? % %); /2 width=1/
qed.
-lemma cw_tpair_dx_sn_shift: ∀I1,I2,L,V1,V2,T. #[L.ⓑ{I2}V2, T] < #[L, ②{I1}V1.②{I2}V2.T].
-#I1 #I2 #L #V1 #V2 #T normalize in ⊢ (? % %); /2 width=1/
+lemma cw_tpair_dx_sn_shift: ∀a2,I1,I2,L,V1,V2,T. #[L.ⓑ{I2}V2, T] < #[L, ②{I1}V1.ⓑ{a2,I2}V2.T].
+#a2 #I1 #I2 #L #V1 #V2 #T normalize in ⊢ (? % %); /2 width=1/
qed.
(* Basic_1: removed theorems 6:
(* binary items *)
inductive item2: Type[0] ≝
- | Bind2: bind2 → item2 (* binding item *)
- | Flat2: flat2 → item2 (* non-binding item *)
+ | Bind2: bool → bind2 → item2 (* polarized binding item *)
+ | Flat2: flat2 → item2 (* non-binding item *)
.
-coercion Bind2.
-
-coercion Flat2.
-
(* Basic properties *********************************************************)
axiom item0_eq_dec: ∀I1,I2:item0. Decidable (I1 = I2).
'SnItem2 I T1 T2 = (TPair I T1 T2).
interpretation "term binding construction (binary)"
- 'SnBind2 I T1 T2 = (TPair (Bind2 I) T1 T2).
+ 'SnBind2 a I T1 T2 = (TPair (Bind2 a I) T1 T2).
+
+interpretation "term positive binding construction (binary)"
+ 'SnBind2Pos I T1 T2 = (TPair (Bind2 true I) T1 T2).
+
+interpretation "term negative binding construction (binary)"
+ 'SnBind2Neg I T1 T2 = (TPair (Bind2 false I) T1 T2).
interpretation "term flat construction (binary)"
'SnFlat2 I T1 T2 = (TPair (Flat2 I) T1 T2).
'GRef p = (TAtom (GRef p)).
interpretation "abbreviation (term)"
- 'SnAbbr T1 T2 = (TPair (Bind2 Abbr) T1 T2).
+ 'SnAbbr a T1 T2 = (TPair (Bind2 a Abbr) T1 T2).
+
+interpretation "positive abbreviation (term)"
+ 'SnAbbrPos T1 T2 = (TPair (Bind2 true Abbr) T1 T2).
+
+interpretation "negative abbreviation (term)"
+ 'SnAbbrNeg T1 T2 = (TPair (Bind2 false Abbr) T1 T2).
interpretation "abstraction (term)"
- 'SnAbst T1 T2 = (TPair (Bind2 Abst) T1 T2).
+ 'SnAbst a T1 T2 = (TPair (Bind2 a Abst) T1 T2).
+
+interpretation "positive abstraction (term)"
+ 'SnAbstPos T1 T2 = (TPair (Bind2 true Abst) T1 T2).
+
+interpretation "negative abstraction (term)"
+ 'SnAbstNeg T1 T2 = (TPair (Bind2 false Abst) T1 T2).
interpretation "application (term)"
'SnAppl T1 T2 = (TPair (Flat2 Appl) T1 T2).
@or_intror @conj // #HT12 destruct /2 width=1/
qed-.
-lemma eq_false_inv_beta: ∀V1,V2,W1,W2,T1,T2.
- (ⓐV1. ⓛW1. T1 = ⓐV2. ⓛW2 .T2 →⊥) →
+lemma eq_false_inv_beta: ∀a,V1,V2,W1,W2,T1,T2.
+ (ⓐV1. ⓛ{a}W1. T1 = ⓐV2. ⓛ{a}W2 .T2 → ⊥) →
(W1 = W2 → ⊥) ∨
- (W1 = W2 ∧ (ⓓV1. T1 = ⓓV2. T2 → ⊥)).
-#V1 #V2 #W1 #W2 #T1 #T2 #H
+ (W1 = W2 ∧ (ⓓ{a}V1. T1 = ⓓ{a}V2. T2 → ⊥)).
+#a #V1 #V2 #W1 #W2 #T1 #T2 #H
elim (eq_false_inv_tpair_sn … H) -H
[ #HV12 elim (term_eq_dec W1 W2) /3 width=1/
#H destruct @or_intror @conj // #H destruct /2 width=1/
interpretation "simple (term)" 'Simple T = (simple T).
(* Basic inversion lemmas ***************************************************)
-
-fact simple_inv_bind_aux: ∀T. 𝐒⦃T⦄ → ∀J,W,U. T = ⓑ{J} W. U → ⊥.
+(*
+lemma mt: ∀R1,R2:Prop. (R1 → R2) → (R2 → ⊥) → R1 → ⊥.
+/3 width=1/ qed.
+*)
+fact simple_inv_bind_aux: ∀T. 𝐒⦃T⦄ → ∀a,J,W,U. T = ⓑ{a,J} W. U → ⊥.
#T * -T
-[ #I #J #W #U #H destruct
-| #I #V #T #J #W #U #H destruct
+[ #I #a #J #W #U #H destruct
+| #I #V #T #a #J #W #U #H destruct
]
qed.
-lemma simple_inv_bind: ∀I,V,T. 𝐒⦃ⓑ{I} V. T⦄ → ⊥.
-/2 width=6/ qed-.
+lemma simple_inv_bind: ∀a,I,V,T. 𝐒⦃ⓑ{a,I} V. T⦄ → ⊥.
+/2 width=7/ qed-. (**) (* auto fails if mt is enabled *)
lemma simple_inv_pair: ∀I,V,T. 𝐒⦃②{I}V.T⦄ → ∃J. I = Flat2 J.
-* /2 width=2/ #I #V #T #H
+* /2 width=2/ #a #I #V #T #H
elim (simple_inv_bind … H)
qed-.
-
-(*
-lemma mt: ∀R1,R2:Prop. (R1 → R2) → (R2 → ⊥) → R1 → ⊥.
-/3 width=1/ qed-.
-*)
let rec applv Vs T on Vs ≝
match Vs with
[ nil ⇒ T
- | cons hd tl ⇒ ⓐhd. (applv tl T)
+ | cons hd tl ⇒ ⓐhd. (applv tl T)
].
interpretation "application o vevtor (term)"
(* properties concerning simple terms ***************************************)
-lemma applv_simple: ∀T,Vs. 𝐒⦃T⦄ -> 𝐒⦃ⒶVs.T⦄.
+lemma applv_simple: ∀T,Vs. 𝐒⦃T⦄ → 𝐒⦃ⒶVs.T⦄.
#T * //
qed.
inductive tshf: relation term ≝
| tshf_atom: ∀I. tshf (⓪{I}) (⓪{I})
- | tshf_abst: ∀V1,V2,T1,T2. tshf (ⓛV1. T1) (ⓛV2. T2)
+ | tshf_abbr: ∀V1,V2,T1,T2. tshf (-ⓓV1. T1) (-ⓓV2. T2)
+ | tshf_abst: ∀a,V1,V2,T1,T2. tshf (ⓛ{a}V1. T1) (ⓛ{a}V2. T2)
| tshf_appl: ∀V1,V2,T1,T2. tshf T1 T2 → 𝐒⦃T1⦄ → 𝐒⦃T2⦄ →
tshf (ⓐV1. T1) (ⓐV2. T2)
.
lemma simple_tshf_repl_dx: ∀T1,T2. T1 ≈ T2 → 𝐒⦃T1⦄ → 𝐒⦃T2⦄.
#T1 #T2 #H elim H -T1 -T2 //
-#V1 #V2 #T1 #T2 #H
-elim (simple_inv_bind … H)
+[ #V1 #V2 #T1 #T2 #H
+ elim (simple_inv_bind … H)
+| #a #V1 #V2 #T1 #T2 #H
+ elim (simple_inv_bind … H)
+]
qed. (**) (* remove from index *)
lemma simple_tshf_repl_sn: ∀T1,T2. T1 ≈ T2 → 𝐒⦃T2⦄ → 𝐒⦃T1⦄.
(* Basic inversion lemmas ***************************************************)
-fact tshf_inv_bind1_aux: ∀T1,T2. T1 ≈ T2 → ∀I,W1,U1. T1 = ⓑ{I}W1.U1 →
- ∃∃W2,U2. I = Abst & T2 = ⓛW2. U2.
+fact tshf_inv_bind1_aux: ∀T1,T2. T1 ≈ T2 → ∀a,I,W1,U1. T1 = ⓑ{a,I}W1.U1 →
+ ∃∃W2,U2. T2 = ⓑ{a,I}W2. U2 &
+ (Bind2 a I = Bind2 false Abbr ∨ I = Abst).
#T1 #T2 * -T1 -T2
-[ #J #I #W1 #U1 #H destruct
-| #V1 #V2 #T1 #T2 #I #W1 #U1 #H destruct /2 width=3/
-| #V1 #V2 #T1 #T2 #H_ #_ #_ #I #W1 #U1 #H destruct
+[ #J #a #I #W1 #U1 #H destruct
+| #V1 #V2 #T1 #T2 #a #I #W1 #U1 #H destruct /3 width=3/
+| #b #V1 #V2 #T1 #T2 #a #I #W1 #U1 #H destruct /3 width=3/
+| #V1 #V2 #T1 #T2 #_ #_ #_ #a #I #W1 #U1 #H destruct
]
qed.
-lemma tshf_inv_bind1: ∀I,W1,U1,T2. ⓑ{I}W1.U1 ≈ T2 →
- ∃∃W2,U2. I = Abst & T2 = ⓛW2. U2.
+lemma tshf_inv_bind1: ∀a,I,W1,U1,T2. ⓑ{a,I}W1.U1 ≈ T2 →
+ ∃∃W2,U2. T2 = ⓑ{a,I}W2. U2 &
+ (Bind2 a I = Bind2 false Abbr ∨ I = Abst).
/2 width=5/ qed-.
fact tshf_inv_flat1_aux: ∀T1,T2. T1 ≈ T2 → ∀I,W1,U1. T1 = ⓕ{I}W1.U1 →
#T1 #T2 * -T1 -T2
[ #J #I #W1 #U1 #H destruct
| #V1 #V2 #T1 #T2 #I #W1 #U1 #H destruct
+| #a #V1 #V2 #T1 #T2 #I #W1 #U1 #H destruct
| #V1 #V2 #T1 #T2 #HT12 #HT1 #HT2 #I #W1 #U1 #H destruct /2 width=5/
]
qed.
(* Advanced inversion lemmas ************************************************)
(* Basic_1: was only: iso_flats_lref_bind_false iso_flats_flat_bind_false *)
-lemma tstc_inv_bind_appls_simple: ∀I,Vs,V2,T1,T2. ⒶVs.T1 ≃ ⓑ{I} V2. T2 →
+lemma tstc_inv_bind_appls_simple: ∀a,I,Vs,V2,T1,T2. ⒶVs.T1 ≃ ⓑ{a,I} V2. T2 →
𝐒⦃T1⦄ → ⊥.
-#I #Vs #V2 #T1 #T2 #H
+#a #I #Vs #V2 #T1 #T2 #H
elim (tstc_inv_pair2 … H) -H #V0 #T0
elim Vs -Vs normalize
[ #H destruct #H
I,J : item
K,L : local environment
M,N : reserved: future use
-O :
-P,Q : reserved: future use (lambda_delta 4)
+O,P,Q :
R : generic predicate (relation)
S : RTM stack
T,U,V,W: term
X,Y,Z : reserved: transient objet denoted by a capital letter
-a,b : reserved: future use (lambda_delta 4)
+a,b : binder polarity
c : reserved: future use (lambda_delta 3)
d : relocation depth
e : relocation height
non associative with precedence 55
for @{ 'SnItem2 $I $T1 $T }.
-notation "hvbox( ⓑ { I } break term 55 T1 . break term 55 T )"
+notation "hvbox( ⓑ { a , I } break term 55 T1 . break term 55 T )"
non associative with precedence 55
- for @{ 'SnBind2 $I $T1 $T }.
+ for @{ 'SnBind2 $a $I $T1 $T }.
+
+notation "hvbox( + ⓑ { I } break term 55 T1 . break term 55 T )"
+ non associative with precedence 55
+ for @{ 'SnBind2Pos $I $T1 $T }.
+
+notation "hvbox( - ⓑ { I } break term 55 T1 . break term 55 T )"
+ non associative with precedence 55
+ for @{ 'SnBind2Neg $I $T1 $T }.
notation "hvbox( ⓕ { I } break term 55 T1 . break term 55 T )"
non associative with precedence 55
for @{ 'SnFlat2 $I $T1 $T }.
-notation "hvbox( ⓓ term 55 T1 . break term 55 T2 )"
+notation "hvbox( ⓓ { a } term 55 T1 . break term 55 T2 )"
+ non associative with precedence 55
+ for @{ 'SnAbbr $a $T1 $T2 }.
+
+notation "hvbox( + ⓓ term 55 T1 . break term 55 T2 )"
+ non associative with precedence 55
+ for @{ 'SnAbbrPos $T1 $T2 }.
+
+notation "hvbox( - ⓓ term 55 T1 . break term 55 T2 )"
non associative with precedence 55
- for @{ 'SnAbbr $T1 $T2 }.
+ for @{ 'SnAbbrNeg $T1 $T2 }.
-notation "hvbox( ⓛ term 55 T1 . break term 55 T2 )"
+notation "hvbox( ⓛ { a } term 55 T1 . break term 55 T2 )"
non associative with precedence 55
- for @{ 'SnAbst $T1 $T2 }.
+ for @{ 'SnAbst $a $T1 $T2 }.
+
+notation "hvbox( + ⓛ term 55 T1 . break term 55 T2 )"
+ non associative with precedence 55
+ for @{ 'SnAbstPos $T1 $T2 }.
+
+notation "hvbox( - ⓛ term 55 T1 . break term 55 T2 )"
+ non associative with precedence 55
+ for @{ 'SnAbstNeg $T1 $T2 }.
notation "hvbox( ⓐ term 55 T1 . break term 55 T2 )"
non associative with precedence 55
(* Unwind *******************************************************************)
-notation "hvbox( ⦃ h , break L ⦄ ⊢ break term 46 T1 •* break [ g ] break term 46 T2 )"
+notation "hvbox( L1 ⊢ ⧫* break term 46 T ≡ break term 46 L2 )"
non associative with precedence 45
- for @{ 'StaticTypeStar $h $g $L $T1 $T2 }.
+ for @{ 'Unwind $L1 $T $L2 }.
(* Reducibility *************************************************************)
non associative with precedence 45
for @{ 'CPRed $L1 $L2 }.
+notation "hvbox( ⦃ L1 ⦄ ➡ break ⦃ L2 ⦄ )"
+ non associative with precedence 45
+ for @{ 'FocalizedPRed $L1 $L2 }.
+
+notation "hvbox( ⦃ L1, break T1 ⦄ ➡ break ⦃ L2 , break T2 ⦄ )"
+ non associative with precedence 45
+ for @{ 'FocalizedPRed $L1 $T1 $L2 $T2 }.
+
notation "hvbox( ⦃ h , break L ⦄ ⊢ break term 46 T1 ➸ break [ g ] break term 46 T2 )"
non associative with precedence 45
for @{ 'XPRed $h $g $L $T1 $T2 }.
qed.
(* Basic_1: was: nf2_abst *)
-lemma cnf_abst: ∀I,L,V,W,T. L ⊢ 𝐍⦃W⦄ → L. ⓑ{I} V ⊢ 𝐍⦃T⦄ → L ⊢ 𝐍⦃ⓛW.T⦄.
-#I #L #V #W #T #HW #HT #X #H
+lemma cnf_abst: ∀a,I,L,V,W,T. L ⊢ 𝐍⦃W⦄ → L. ⓑ{I} V ⊢ 𝐍⦃T⦄ → L ⊢ 𝐍⦃ⓛ{a}W.T⦄.
+#a #I #L #V #W #T #HW #HT #X #H
elim (cpr_inv_abst1 … H I V) -H #W0 #T0 #HW0 #HT0 #H destruct
>(HW … HW0) -W0 >(HT … HT0) -T0 //
qed.
(* Advanced properties ******************************************************)
-lemma cpr_bind_sn: ∀I,L,V1,V2,T1,T2. L ⊢ V1 ➡ V2 → T1 ➡ T2 →
- L ⊢ ⓑ{I} V1. T1 ➡ ⓑ{I} V2. T2.
-#I #L #V1 #V2 #T1 #T2 * #V #HV1 #HV2 #HT12
+lemma cpr_bind_sn: ∀a,I,L,V1,V2,T1,T2. L ⊢ V1 ➡ V2 → T1 ➡ T2 →
+ L ⊢ ⓑ{a,I} V1. T1 ➡ ⓑ{a,I} V2. T2.
+#a #I #L #V1 #V2 #T1 #T2 * #V #HV1 #HV2 #HT12
@ex2_1_intro [2: @(tpr_delta … HV1 HT12) | skip ] /2 width=3/ (* /3 width=5/ is too slow *)
qed.
(* Basic_1: was only: pr2_gen_cbind *)
-lemma cpr_bind_dx: ∀I,L,V1,V2,T1,T2. V1 ➡ V2 → L. ⓑ{I} V2 ⊢ T1 ➡ T2 →
- L ⊢ ⓑ{I} V1. T1 ➡ ⓑ{I} V2. T2.
-#I #L #V1 #V2 #T1 #T2 #HV12 * #T #HT1 normalize #HT2
+lemma cpr_bind_dx: ∀a,I,L,V1,V2,T1,T2. V1 ➡ V2 → L. ⓑ{I} V2 ⊢ T1 ➡ T2 →
+ L ⊢ ⓑ{a,I} V1. T1 ➡ ⓑ{a,I} V2. T2.
+#a #I #L #V1 #V2 #T1 #T2 #HV12 * #T #HT1 normalize #HT2
elim (tpss_split_up … HT2 1 ? ?) -HT2 // #T0 <minus_n_O #HT0 normalize <minus_plus_m_m #HT02
lapply (tpss_lsubs_trans … HT0 (⋆. ⓑ{I} V2) ?) -HT0 /2 width=1/ #HT0
lapply (tpss_inv_SO2 … HT0) -HT0 #HT0
qed.
lemma cpr_abst: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀V,T1,T2.
- L.ⓛV ⊢ T1 ➡ T2 → L ⊢ ⓛV1. T1 ➡ ⓛV2. T2.
-#L #V1 #V2 * #V0 #HV10 #HV02 #V #T1 #T2 * #T0 #HT10 #HT02
+ L.ⓛV ⊢ T1 ➡ T2 → ∀a. L ⊢ ⓛ{a}V1. T1 ➡ ⓛ{a}V2. T2.
+#L #V1 #V2 * #V0 #HV10 #HV02 #V #T1 #T2 * #T0 #HT10 #HT02 #a
lapply (tpss_inv_S2 … HT02 L V ?) -HT02 // #HT02
lapply (tpss_lsubs_trans … HT02 (L.ⓛV2) ?) -HT02 /2 width=1/ #HT02
-@(ex2_1_intro … (ⓛV0.T0)) /2 width=1/ (* explicit constructors *)
+@(ex2_1_intro … (ⓛ{a}V0.T0)) /2 width=1/ (* explicit constructors *)
qed.
-lemma cpr_beta: ∀L,V1,V2,W,T1,T2.
- L ⊢ V1 ➡ V2 → L.ⓛW ⊢ T1 ➡ T2 → L ⊢ ⓐV1.ⓛW.T1 ➡ ⓓV2.T2.
-#L #V1 #V2 #W #T1 #T2 * #V #HV1 #HV2 * #T #HT1 #HT2
+lemma cpr_beta: ∀a,L,V1,V2,W,T1,T2.
+ L ⊢ V1 ➡ V2 → L.ⓛW ⊢ T1 ➡ T2 → L ⊢ ⓐV1.ⓛ{a}W.T1 ➡ ⓓ{a}V2.T2.
+#a #L #V1 #V2 #W #T1 #T2 * #V #HV1 #HV2 * #T #HT1 #HT2
lapply (tpss_inv_S2 … HT2 L W ?) -HT2 // #HT2
lapply (tpss_lsubs_trans … HT2 (L.ⓓV2) ?) -HT2 /2 width=1/ #HT2
-@(ex2_1_intro … (ⓓV.T)) /2 width=1/ (**) (* explicit constructor, /3/ is too slow *)
+@(ex2_1_intro … (ⓓ{a}V.T)) /2 width=1/ (**) (* explicit constructor, /3/ is too slow *)
qed.
-lemma cpr_beta_dx: ∀L,V1,V2,W,T1,T2.
- V1 ➡ V2 → L.ⓛW ⊢ T1 ➡ T2 → L ⊢ ⓐV1.ⓛW.T1 ➡ ⓓV2.T2.
+lemma cpr_beta_dx: ∀a,L,V1,V2,W,T1,T2.
+ V1 ➡ V2 → L.ⓛW ⊢ T1 ➡ T2 → L ⊢ ⓐV1.ⓛ{a}W.T1 ➡ ⓓ{a}V2.T2.
/3 width=1/ qed.
(* Advanced inversion lemmas ************************************************)
qed-.
(* Basic_1: was pr2_gen_abbr *)
-lemma cpr_inv_abbr1: ∀L,V1,T1,U2. L ⊢ ⓓV1. T1 ➡ U2 →
+lemma cpr_inv_abbr1: ∀a,L,V1,T1,U2. L ⊢ ⓓ{a}V1. T1 ➡ U2 →
(∃∃V,V2,T2. V1 ➡ V & L ⊢ V ▶* [O, |L|] V2 &
L. ⓓV ⊢ T1 ➡ T2 &
- U2 = ⓓV2. T2
+ U2 = ⓓ{a}V2. T2
) ∨
- ∃∃T2. L.ⓓV1 ⊢ T1 ➡ T2 & ⇧[0,1] U2 ≡ T2.
-#L #V1 #T1 #Y * #X #H1 #H2
+ ∃∃T2. L.ⓓV1 ⊢ T1 ➡ T2 & ⇧[0,1] U2 ≡ T2 & a = true.
+#a #L #V1 #T1 #Y * #X #H1 #H2
elim (tpr_inv_abbr1 … H1) -H1 *
[ #V #T #T0 #HV1 #HT1 #HT0 #H destruct
elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV2 #HT02 #H destruct
lapply (tpss_lsubs_trans … HT02 (L. ⓓV) ?) -HT02 /2 width=1/ #HT02
lapply (tpss_weak_all … HT02) -HT02 #HT02
lapply (tpss_strap2 … HT0 HT02) -T0 /4 width=7/
-| #T2 #HT12 #HXT2
+| #T2 #HT12 #HXT2 #H destruct
elim (lift_total Y 0 1) #Z #HYZ
lapply (tpss_lift_ge … H2 (L.ⓓV1) … HXT2 … HYZ) -X // /2 width=1/ #H
lapply (cpr_intro … HT12 … H) -T2 /3 width=3/
qed-.
(* Basic_1: was: pr2_gen_abst *)
-lemma cpr_inv_abst1: ∀L,V1,T1,U2. L ⊢ ⓛV1. T1 ➡ U2 → ∀I,W.
- ∃∃V2,T2. L ⊢ V1 ➡ V2 & L. ⓑ{I} W ⊢ T1 ➡ T2 & U2 = ⓛV2. T2.
-#L #V1 #T1 #Y * #X #H1 #H2 #I #W
+lemma cpr_inv_abst1: ∀a,L,V1,T1,U2. L ⊢ ⓛ{a}V1. T1 ➡ U2 → ∀I,W.
+ ∃∃V2,T2. L ⊢ V1 ➡ V2 & L. ⓑ{I} W ⊢ T1 ➡ T2 & U2 = ⓛ{a}V2. T2.
+#a #L #V1 #T1 #Y * #X #H1 #H2 #I #W
elim (tpr_inv_abst1 … H1) -H1 #V #T #HV1 #HT1 #H destruct
elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV2 #HT2 #H destruct
lapply (tpss_lsubs_trans … HT2 (L. ⓑ{I} W) ?) -HT2 /2 width=1/ /4 width=5/
(* Basic_1: was pr2_gen_appl *)
lemma cpr_inv_appl1: ∀L,V1,U0,U2. L ⊢ ⓐV1. U0 ➡ U2 →
- ∨∨ ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ U0 ➡ T2 &
- U2 = ⓐV2. T2
- | ∃∃V2,W,T1,T2. L ⊢ V1 ➡ V2 & L. ⓓV2 ⊢ T1 ➡ T2 &
- U0 = ⓛW. T1 &
- U2 = ⓓV2. T2
- | ∃∃V2,V,W1,W2,T1,T2. L ⊢ V1 ➡ V2 & L ⊢ W1 ➡ W2 & L. ⓓW2 ⊢ T1 ➡ T2 &
- ⇧[0,1] V2 ≡ V &
- U0 = ⓓW1. T1 &
- U2 = ⓓW2. ⓐV. T2.
+ ∨∨ ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ U0 ➡ T2 &
+ U2 = ⓐV2. T2
+ | ∃∃a,V2,W,T1,T2. L ⊢ V1 ➡ V2 & L. ⓓV2 ⊢ T1 ➡ T2 &
+ U0 = ⓛ{a}W. T1 &
+ U2 = ⓓ{a}V2. T2
+ | ∃∃a,V2,V,W1,W2,T1,T2. L ⊢ V1 ➡ V2 & L ⊢ W1 ➡ W2 & L. ⓓW2 ⊢ T1 ➡ T2 &
+ ⇧[0,1] V2 ≡ V &
+ U0 = ⓓ{a}W1. T1 &
+ U2 = ⓓ{a}W2. ⓐV. T2.
#L #V1 #U0 #Y * #X #H1 #H2
elim (tpr_inv_appl1 … H1) -H1 *
[ #V #U #HV1 #HU0 #H destruct
elim (tpss_inv_flat1 … H2) -H2 #V2 #U2 #HV2 #HU2 #H destruct /4 width=5/
-| #V #W #T0 #T #HV1 #HT0 #H #H1 destruct
+| #a #V #W #T0 #T #HV1 #HT0 #H #H1 destruct
elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV2 #HT2 #H destruct
- lapply (tpss_weak … HT2 0 (|L|+1) ? ?) -HT2 // /4 width=8/
-| #V0 #V #W #W0 #T #T0 #HV10 #HW0 #HT0 #HV0 #H #H1 destruct
+ lapply (tpss_weak … HT2 0 (|L|+1) ? ?) -HT2 // /4 width=9/
+| #a #V0 #V #W #W0 #T #T0 #HV10 #HW0 #HT0 #HV0 #H #H1 destruct
elim (tpss_inv_bind1 … H2) -H2 #W2 #X #HW02 #HX #HY destruct
elim (tpss_inv_flat1 … HX) -HX #V2 #T2 #HV2 #HT2 #H destruct
elim (tpss_inv_lift1_ge … HV2 … HV0 ?) -V // [3: /2 width=1/ |2: skip ] #V <minus_plus_m_m
- lapply (tpss_weak … HT2 0 (|L|+1) ? ?) -HT2 // /4 width=12/
+ lapply (tpss_weak … HT2 0 (|L|+1) ? ?) -HT2 // /4 width=13/
]
qed-.
#L #V1 #T1 #U #H #HT1
elim (cpr_inv_appl1 … H) -H *
[ /2 width=5/
-| #V2 #W #W1 #W2 #_ #_ #H #_ destruct
+| #a #V2 #W #W1 #W2 #_ #_ #H #_ destruct
elim (simple_inv_bind … HT1)
-| #V2 #V #W1 #W2 #U1 #U2 #_ #_ #_ #_ #H #_ destruct
+| #a #V2 #V #W1 #W2 #U1 #U2 #_ #_ #_ #_ #H #_ destruct
elim (simple_inv_bind … HT1)
]
qed-.
∀T1,U1. ⇧[d, e] T1 ≡ U1 → ∀T2,U2. ⇧[d, e] T2 ≡ U2 →
K ⊢ T1 ➡ T2 → L ⊢ U1 ➡ U2.
#L #K #d #e #HLK #T1 #U1 #HTU1 #T2 #U2 #HTU2 * #T #HT1 #HT2
-elim (lift_total T d e) #U #HTU
+elim (lift_total T d e) #U #HTU
lapply (tpr_lift … HT1 … HTU1 … HTU) -T1 #HU1
elim (lt_or_ge (|K|) d) #HKd
[ lapply (tpss_lift_le … HT2 … HLK HTU … HTU2) -T2 -T -HLK [ /2 width=2/ | /3 width=4/ ]
elim (ltpr_ldrop_conf … HLK1 … HL12) -HLK1 -HL12 #X #HLK2 #H
elim (ltpr_inv_pair1 … H) -H #K2 #V2 #HK12 #HV12 #H destruct
lapply (IH … HKV1 … HK1 … HK12 … HV12) // -L1 -K1 -V1 /2 width=5/
-| #L1 #V1 #T1 #B #A #HB #HA #H1 #H2 #L2 #HL12 #X #H destruct
+| #a #L1 #V1 #T1 #B #A #HB #HA #H1 #H2 #L2 #HL12 #X #H destruct
elim (tpr_inv_abbr1 … H) -H *
[ #V2 #T #T2 #HV12 #HT1 #HT2 #H destruct
lapply (tps_lsubs_trans … HT2 (L2.ⓓV2) ?) -HT2 /2 width=1/ #HT2
lapply (IH … HB … HL12 … HV12) -HB /width=5/ #HB
lapply (IH … HA … (L2.ⓓV2) … HT1) -IH -HA -HT1 /width=5/ -T1 /2 width=1/ -L1 -V1 /3 width=5/
- | -B #T #HT1 #HXT
+ | -B #T #HT1 #HXT #H destruct
lapply (IH … HA … (L2.ⓓV1) … HT1) /width=5/ -T1 /2 width=1/ -L1 #HA
@(aaa_inv_lift … HA … HXT) /2 width=1/
]
-| #L1 #V1 #T1 #B #A #HB #HA #H1 #H2 #L2 #HL12 #X #H destruct
+| #a #L1 #V1 #T1 #B #A #HB #HA #H1 #H2 #L2 #HL12 #X #H destruct
elim (tpr_inv_abst1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
lapply (IH … HB … HL12 … HV12) -HB /width=5/ #HB
lapply (IH … HA … (L2.ⓛV2) … HT12) -IH -HA -HT12 /width=5/ -T1 /2 width=1/
[ #V2 #T2 #HV12 #HT12 #H destruct
lapply (IH … HV1 … HL12 … HV12) -HV1 -HV12 /width=5/ #HB
lapply (IH … HT1 … HL12 … HT12) -IH -HT1 -HL12 -HT12 /width=5/ /2 width=3/
- | #V2 #W2 #T0 #T2 #HV12 #HT02 #H1 #H2 destruct
+ | #a #V2 #W2 #T0 #T2 #HV12 #HT02 #H1 #H2 destruct
elim (aaa_inv_abst … HT1) -HT1 #B0 #A0 #HB0 #HA0 #H destruct
lapply (IH … HV1 … HL12 … HV12) -HV1 -HV12 /width=5/ #HB
lapply (IH … HB0 … HL12 W2 ?) -HB0 /width=5/ #HB0
lapply (IH … HA0 … (L2.ⓛW2) … HT02) -IH -HA0 -HT02 /width=5/ -T0 /2 width=1/ -L1 -V1 /4 width=7/
- | #V0 #V2 #W0 #W2 #T0 #T2 #HV10 #HW02 #HT02 #HV02 #H1 #H2 destruct
+ | #a #V0 #V2 #W0 #W2 #T0 #T2 #HV10 #HW02 #HT02 #HV02 #H1 #H2 destruct
elim (aaa_inv_abbr … HT1) -HT1 #B0 #HW0 #HT0
lapply (IH … HW0 … HL12 … HW02) -HW0 /width=5/ #HW2
lapply (IH … HV1 … HL12 … HV10) -HV1 -HV10 /width=5/ #HV0
elim (ltpr_inv_pair2 … H) -H #K1 #V1 #HK12 #HV12 #H destruct -K2
elim (lift_total V1 0 (i+1)) #W1 #HVW1
lapply (tpr_lift … HV12 … HVW1 … HVW2) -V2 /3 width=4/
-| #L2 #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L1 #HL12
+| #L2 #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L1 #HL12
elim (IHV12 … HL12) -IHV12 #V #HV1 #HV2
elim (IHT12 (L1.ⓑ{I}V) ?) /2 width=1/ -L2 /3 width=5/
| #L2 #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L1 #HL12
elim (ltpr_inv_pair1 … H) -H #K2 #V2 #HK12 #HV12 #H destruct -K1
elim (lift_total V2 0 (i+1)) #W2 #HVW2
lapply (tpr_lift … HV12 … HVW1 … HVW2) -V1 /3 width=4/
-| #L1 #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L2 #HL12
+| #L1 #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L2 #HL12
elim (IHV12 … HL12) -IHV12 #V #HV1 #HV2
elim (IHT12 (L2.ⓑ{I}V) ?) /2 width=1/ -L1 /3 width=5/
| #L1 #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L2 #HL12
(* Basic inversion lemmas ***************************************************)
-lemma tnf_inv_abst: ∀V,T. 𝐍⦃ⓛV.T⦄ → 𝐍⦃V⦄ ∧ 𝐍⦃T⦄.
-#V1 #T1 #HVT1 @conj
-[ #V2 #HV2 lapply (HVT1 (ⓛV2.T1) ?) -HVT1 /2 width=1/ -HV2 #H destruct //
-| #T2 #HT2 lapply (HVT1 (ⓛV1.T2) ?) -HVT1 /2 width=1/ -HT2 #H destruct //
+lemma tnf_inv_abst: ∀a,V,T. 𝐍⦃ⓛ{a}V.T⦄ → 𝐍⦃V⦄ ∧ 𝐍⦃T⦄.
+#a #V1 #T1 #HVT1 @conj
+[ #V2 #HV2 lapply (HVT1 (ⓛ{a}V2.T1) ?) -HVT1 /2 width=1/ -HV2 #H destruct //
+| #T2 #HT2 lapply (HVT1 (ⓛ{a}V1.T2) ?) -HVT1 /2 width=1/ -HT2 #H destruct //
]
qed-.
#V1 #T1 #HVT1 @and3_intro
[ #V2 #HV2 lapply (HVT1 (ⓐV2.T1) ?) -HVT1 /2 width=1/ -HV2 #H destruct //
| #T2 #HT2 lapply (HVT1 (ⓐV1.T2) ?) -HVT1 /2 width=1/ -HT2 #H destruct //
-| generalize in match HVT1; -HVT1 elim T1 -T1 * // * #W1 #U1 #_ #_ #H
+| generalize in match HVT1; -HVT1 elim T1 -T1 * // #a * #W1 #U1 #_ #_ #H
[ elim (lift_total V1 0 1) #V2 #HV12
- lapply (H (ⓓW1.ⓐV2.U1) ?) -H /2 width=3/ -HV12 #H destruct
- | lapply (H (ⓓV1.U1) ?) -H /2 width=1/ #H destruct
+ lapply (H (ⓓ{a}W1.ⓐV2.U1) ?) -H /2 width=3/ -HV12 #H destruct
+ | lapply (H (ⓓ{a}V1.U1) ?) -H /2 width=1/ #H destruct
]
qed-.
-lemma tnf_inv_abbr: ∀V,T. 𝐍⦃ⓓV.T⦄ → ⊥.
+lemma tnf_inv_abbr: ∀V,T. 𝐍⦃+ⓓV.T⦄ → ⊥.
#V #T #H elim (is_lift_dec T 0 1)
[ * #U #HTU
lapply (H U ?) -H /2 width=3/ #H destruct
elim (lift_inv_pair_xy_y … HTU)
| #HT
elim (tps_full (⋆) V T (⋆. ⓓV) 0 ?) // #T2 #T1 #HT2 #HT12
- lapply (H (ⓓV.T2) ?) -H /2 width=3/ -HT2 #H destruct /3 width=2/
+ lapply (H (+ⓓV.T2) ?) -H /2 width=3/ -HT2 #H destruct /3 width=2/
]
qed.
| tpr_atom : ∀I. tpr (⓪{I}) (⓪{I})
| tpr_flat : ∀I,V1,V2,T1,T2. tpr V1 V2 → tpr T1 T2 →
tpr (ⓕ{I} V1. T1) (ⓕ{I} V2. T2)
-| tpr_beta : ∀V1,V2,W,T1,T2.
- tpr V1 V2 → tpr T1 T2 → tpr (ⓐV1. ⓛW. T1) (ⓓV2. T2)
-| tpr_delta: ∀I,V1,V2,T1,T,T2.
+| tpr_beta : ∀a,V1,V2,W,T1,T2.
+ tpr V1 V2 → tpr T1 T2 → tpr (ⓐV1. ⓛ{a}W. T1) (ⓓ{a}V2. T2)
+| tpr_delta: ∀a,I,V1,V2,T1,T,T2.
tpr V1 V2 → tpr T1 T → ⋆. ⓑ{I} V2 ⊢ T ▶ [0, 1] T2 →
- tpr (ⓑ{I} V1. T1) (ⓑ{I} V2. T2)
-| tpr_theta: ∀V,V1,V2,W1,W2,T1,T2.
+ tpr (ⓑ{a,I} V1. T1) (ⓑ{a,I} V2. T2)
+| tpr_theta: ∀a,V,V1,V2,W1,W2,T1,T2.
tpr V1 V2 → ⇧[0,1] V2 ≡ V → tpr W1 W2 → tpr T1 T2 →
- tpr (ⓐV1. ⓓW1. T1) (ⓓW2. ⓐV. T2)
-| tpr_zeta : ∀V,T1,T,T2. tpr T1 T → ⇧[0, 1] T2 ≡ T → tpr (ⓓV. T1) T2
+ tpr (ⓐV1. ⓓ{a}W1. T1) (ⓓ{a}W2. ⓐV. T2)
+| tpr_zeta : ∀V,T1,T,T2. tpr T1 T → ⇧[0, 1] T2 ≡ T → tpr (+ⓓV. T1) T2
| tpr_tau : ∀V,T1,T2. tpr T1 T2 → tpr (ⓝV. T1) T2
.
(* Basic properties *********************************************************)
-lemma tpr_bind: ∀I,V1,V2,T1,T2. V1 ➡ V2 → T1 ➡ T2 → ⓑ{I} V1. T1 ➡ ⓑ{I} V2. T2.
+lemma tpr_bind: ∀a,I,V1,V2,T1,T2. V1 ➡ V2 → T1 ➡ T2 → ⓑ{a,I} V1. T1 ➡ ⓑ{a,I} V2. T2.
/2 width=3/ qed.
(* Basic_1: was by definition: pr0_refl *)
#U1 #U2 * -U1 -U2
[ //
| #I #V1 #V2 #T1 #T2 #_ #_ #k #H destruct
-| #V1 #V2 #W #T1 #T2 #_ #_ #k #H destruct
-| #I #V1 #V2 #T1 #T #T2 #_ #_ #_ #k #H destruct
-| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #k #H destruct
+| #a #V1 #V2 #W #T1 #T2 #_ #_ #k #H destruct
+| #a #I #V1 #V2 #T1 #T #T2 #_ #_ #_ #k #H destruct
+| #a #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #k #H destruct
| #V #T1 #T #T2 #_ #_ #k #H destruct
| #V #T1 #T2 #_ #k #H destruct
]
lemma tpr_inv_atom1: ∀I,U2. ⓪{I} ➡ U2 → U2 = ⓪{I}.
/2 width=3/ qed-.
-fact tpr_inv_bind1_aux: ∀U1,U2. U1 ➡ U2 → ∀I,V1,T1. U1 = ⓑ{I} V1. T1 →
+fact tpr_inv_bind1_aux: ∀U1,U2. U1 ➡ U2 → ∀a,I,V1,T1. U1 = ⓑ{a,I} V1. T1 →
(∃∃V2,T,T2. V1 ➡ V2 & T1 ➡ T &
⋆. ⓑ{I} V2 ⊢ T ▶ [0, 1] T2 &
- U2 = ⓑ{I} V2. T2
+ U2 = ⓑ{a,I} V2. T2
) ∨
- ∃∃T. T1 ➡ T & ⇧[0, 1] U2 ≡ T & I = Abbr.
+ ∃∃T. T1 ➡ T & ⇧[0, 1] U2 ≡ T & a = true & I = Abbr.
#U1 #U2 * -U1 -U2
-[ #J #I #V #T #H destruct
-| #I1 #V1 #V2 #T1 #T2 #_ #_ #I #V #T #H destruct
-| #V1 #V2 #W #T1 #T2 #_ #_ #I #V #T #H destruct
-| #I1 #V1 #V2 #T1 #T #T2 #HV12 #HT1 #HT2 #I0 #V0 #T0 #H destruct /3 width=7/
-| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #I0 #V0 #T0 #H destruct
-| #V #T1 #T #T2 #HT1 #HT2 #I0 #V0 #T0 #H destruct /3 width=3/
-| #V #T1 #T2 #_ #I0 #V0 #T0 #H destruct
+[ #J #a #I #V #T #H destruct
+| #I1 #V1 #V2 #T1 #T2 #_ #_ #a #I #V #T #H destruct
+| #b #V1 #V2 #W #T1 #T2 #_ #_ #a #I #V #T #H destruct
+| #b #I1 #V1 #V2 #T1 #T #T2 #HV12 #HT1 #HT2 #a #I0 #V0 #T0 #H destruct /3 width=7/
+| #b #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #a #I0 #V0 #T0 #H destruct
+| #V #T1 #T #T2 #HT1 #HT2 #a #I0 #V0 #T0 #H destruct /3 width=3/
+| #V #T1 #T2 #_ #a #I0 #V0 #T0 #H destruct
]
qed.
-lemma tpr_inv_bind1: ∀V1,T1,U2,I. ⓑ{I} V1. T1 ➡ U2 →
+lemma tpr_inv_bind1: ∀V1,T1,U2,a,I. ⓑ{a,I} V1. T1 ➡ U2 →
(∃∃V2,T,T2. V1 ➡ V2 & T1 ➡ T &
⋆. ⓑ{I} V2 ⊢ T ▶ [0, 1] T2 &
- U2 = ⓑ{I} V2. T2
+ U2 = ⓑ{a,I} V2. T2
) ∨
- ∃∃T. T1 ➡ T & ⇧[0,1] U2 ≡ T & I = Abbr.
+ ∃∃T. T1 ➡ T & ⇧[0,1] U2 ≡ T & a = true & I = Abbr.
/2 width=3/ qed-.
(* Basic_1: was pr0_gen_abbr *)
-lemma tpr_inv_abbr1: ∀V1,T1,U2. ⓓV1. T1 ➡ U2 →
+lemma tpr_inv_abbr1: ∀a,V1,T1,U2. ⓓ{a}V1. T1 ➡ U2 →
(∃∃V2,T,T2. V1 ➡ V2 & T1 ➡ T &
⋆. ⓓV2 ⊢ T ▶ [0, 1] T2 &
- U2 = ⓓV2. T2
+ U2 = ⓓ{a}V2. T2
) ∨
- ∃∃T. T1 ➡ T & ⇧[0, 1] U2 ≡ T.
-#V1 #T1 #U2 #H
+ ∃∃T. T1 ➡ T & ⇧[0, 1] U2 ≡ T & a = true.
+#a #V1 #T1 #U2 #H
elim (tpr_inv_bind1 … H) -H * /3 width=7/
qed-.
fact tpr_inv_flat1_aux: ∀U1,U2. U1 ➡ U2 → ∀I,V1,U0. U1 = ⓕ{I} V1. U0 →
- ∨∨ ∃∃V2,T2. V1 ➡ V2 & U0 ➡ T2 &
- U2 = ⓕ{I} V2. T2
- | ∃∃V2,W,T1,T2. V1 ➡ V2 & T1 ➡ T2 &
- U0 = ⓛW. T1 &
- U2 = ⓓV2. T2 & I = Appl
- | ∃∃V2,V,W1,W2,T1,T2. V1 ➡ V2 & W1 ➡ W2 & T1 ➡ T2 &
- ⇧[0,1] V2 ≡ V &
- U0 = ⓓW1. T1 &
- U2 = ⓓW2. ⓐV. T2 &
- I = Appl
- | (U0 ➡ U2 ∧ I = Cast).
+ ∨∨ ∃∃V2,T2. V1 ➡ V2 & U0 ➡ T2 &
+ U2 = ⓕ{I} V2. T2
+ | ∃∃a,V2,W,T1,T2. V1 ➡ V2 & T1 ➡ T2 &
+ U0 = ⓛ{a}W. T1 &
+ U2 = ⓓ{a}V2. T2 & I = Appl
+ | ∃∃a,V2,V,W1,W2,T1,T2. V1 ➡ V2 & W1 ➡ W2 & T1 ➡ T2 &
+ ⇧[0,1] V2 ≡ V &
+ U0 = ⓓ{a}W1. T1 &
+ U2 = ⓓ{a}W2. ⓐV. T2 &
+ I = Appl
+ | (U0 ➡ U2 ∧ I = Cast).
#U1 #U2 * -U1 -U2
[ #I #J #V #T #H destruct
| #I #V1 #V2 #T1 #T2 #HV12 #HT12 #J #V #T #H destruct /3 width=5/
-| #V1 #V2 #W #T1 #T2 #HV12 #HT12 #J #V #T #H destruct /3 width=8/
-| #I #V1 #V2 #T1 #T #T2 #_ #_ #_ #J #V0 #T0 #H destruct
-| #V #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HV2 #HW12 #HT12 #J #V0 #T0 #H destruct /3 width=12/
+| #a #V1 #V2 #W #T1 #T2 #HV12 #HT12 #J #V #T #H destruct /3 width=9/
+| #a #I #V1 #V2 #T1 #T #T2 #_ #_ #_ #J #V0 #T0 #H destruct
+| #a #V #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HV2 #HW12 #HT12 #J #V0 #T0 #H destruct /3 width=13/
| #V #T1 #T #T2 #_ #_ #J #V0 #T0 #H destruct
| #V #T1 #T2 #HT12 #J #V0 #T0 #H destruct /3 width=1/
]
qed.
lemma tpr_inv_flat1: ∀V1,U0,U2,I. ⓕ{I} V1. U0 ➡ U2 →
- ∨∨ ∃∃V2,T2. V1 ➡ V2 & U0 ➡ T2 &
- U2 = ⓕ{I} V2. T2
- | ∃∃V2,W,T1,T2. V1 ➡ V2 & T1 ➡ T2 &
- U0 = ⓛW. T1 &
- U2 = ⓓV2. T2 & I = Appl
- | ∃∃V2,V,W1,W2,T1,T2. V1 ➡ V2 & W1 ➡ W2 & T1 ➡ T2 &
- ⇧[0,1] V2 ≡ V &
- U0 = ⓓW1. T1 &
- U2 = ⓓW2. ⓐV. T2 &
- I = Appl
- | (U0 ➡ U2 ∧ I = Cast).
+ ∨∨ ∃∃V2,T2. V1 ➡ V2 & U0 ➡ T2 &
+ U2 = ⓕ{I} V2. T2
+ | ∃∃a,V2,W,T1,T2. V1 ➡ V2 & T1 ➡ T2 &
+ U0 = ⓛ{a}W. T1 &
+ U2 = ⓓ{a}V2. T2 & I = Appl
+ | ∃∃a,V2,V,W1,W2,T1,T2. V1 ➡ V2 & W1 ➡ W2 & T1 ➡ T2 &
+ ⇧[0,1] V2 ≡ V &
+ U0 = ⓓ{a}W1. T1 &
+ U2 = ⓓ{a}W2. ⓐV. T2 &
+ I = Appl
+ | (U0 ➡ U2 ∧ I = Cast).
/2 width=3/ qed-.
(* Basic_1: was pr0_gen_appl *)
lemma tpr_inv_appl1: ∀V1,U0,U2. ⓐV1. U0 ➡ U2 →
- ∨∨ ∃∃V2,T2. V1 ➡ V2 & U0 ➡ T2 &
- U2 = ⓐV2. T2
- | ∃∃V2,W,T1,T2. V1 ➡ V2 & T1 ➡ T2 &
- U0 = ⓛW. T1 &
- U2 = ⓓV2. T2
- | ∃∃V2,V,W1,W2,T1,T2. V1 ➡ V2 & W1 ➡ W2 & T1 ➡ T2 &
- ⇧[0,1] V2 ≡ V &
- U0 = ⓓW1. T1 &
- U2 = ⓓW2. ⓐV. T2.
+ ∨∨ ∃∃V2,T2. V1 ➡ V2 & U0 ➡ T2 &
+ U2 = ⓐV2. T2
+ | ∃∃a,V2,W,T1,T2. V1 ➡ V2 & T1 ➡ T2 &
+ U0 = ⓛ{a}W. T1 &
+ U2 = ⓓ{a}V2. T2
+ | ∃∃a,V2,V,W1,W2,T1,T2. V1 ➡ V2 & W1 ➡ W2 & T1 ➡ T2 &
+ ⇧[0,1] V2 ≡ V &
+ U0 = ⓓ{a}W1. T1 &
+ U2 = ⓓ{a}W2. ⓐV. T2.
#V1 #U0 #U2 #H
-elim (tpr_inv_flat1 … H) -H * /3 width=12/ #_ #H destruct
+elim (tpr_inv_flat1 … H) -H *
+/3 width=5/ /3 width=9/ /3 width=13/
+#_ #H destruct
qed-.
(* Note: the main property of simple terms *)
#V1 #T1 #U #H #HT1
elim (tpr_inv_appl1 … H) -H *
[ /2 width=5/
-| #V2 #W #W1 #W2 #_ #_ #H #_ destruct
+| #a #V2 #W #W1 #W2 #_ #_ #H #_ destruct
elim (simple_inv_bind … HT1)
-| #V2 #V #W1 #W2 #U1 #U2 #_ #_ #_ #_ #H #_ destruct
+| #a #V2 #V #W1 #W2 #U1 #U2 #_ #_ #_ #_ #H #_ destruct
elim (simple_inv_bind … HT1)
]
qed-.
(∃∃V2,T2. V1 ➡ V2 & T1 ➡ T2 & U2 = ⓝV2. T2)
∨ T1 ➡ U2.
#V1 #T1 #U2 #H
-elim (tpr_inv_flat1 … H) -H * /3 width=5/
-[ #V2 #W #W1 #W2 #_ #_ #_ #_ #H destruct
-| #V2 #W #W1 #W2 #T2 #U1 #_ #_ #_ #_ #_ #_ #H destruct
+elim (tpr_inv_flat1 … H) -H * /3 width=5/ #a #V2 #W #W1 #W2
+[ #_ #_ #_ #_ #H destruct
+| #T2 #U1 #_ #_ #_ #_ #_ #_ #H destruct
]
qed-.
fact tpr_inv_lref2_aux: ∀T1,T2. T1 ➡ T2 → ∀i. T2 = #i →
∨∨ T1 = #i
- | ∃∃V,T. T ➡ #(i+1) & T1 = ⓓV. T
+ | ∃∃V,T. T ➡ #(i+1) & T1 = +ⓓV. T
| ∃∃V,T. T ➡ #i & T1 = ⓝV. T.
#T1 #T2 * -T1 -T2
[ #I #i #H destruct /2 width=1/
| #I #V1 #V2 #T1 #T2 #_ #_ #i #H destruct
-| #V1 #V2 #W #T1 #T2 #_ #_ #i #H destruct
-| #I #V1 #V2 #T1 #T #T2 #_ #_ #_ #i #H destruct
-| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #i #H destruct
+| #a #V1 #V2 #W #T1 #T2 #_ #_ #i #H destruct
+| #a #I #V1 #V2 #T1 #T #T2 #_ #_ #_ #i #H destruct
+| #a #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #i #H destruct
| #V #T1 #T #T2 #HT1 #HT2 #i #H destruct
lapply (lift_inv_lref1_ge … HT2 ?) -HT2 // #H destruct /3 width=4/
| #V #T1 #T2 #HT12 #i #H destruct /3 width=4/
lemma tpr_inv_lref2: ∀T1,i. T1 ➡ #i →
∨∨ T1 = #i
- | ∃∃V,T. T ➡ #(i+1) & T1 = ⓓV. T
+ | ∃∃V,T. T ➡ #(i+1) & T1 = +ⓓV. T
| ∃∃V,T. T ➡ #i & T1 = ⓝV. T.
/2 width=3/ qed-.
| #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #X1 #HX1 #X2 #HX2
elim (lift_inv_flat1 … HX1) -HX1 #W1 #U1 #HVW1 #HTU1 #HX1 destruct
elim (lift_inv_flat1 … HX2) -HX2 #W2 #U2 #HVW2 #HTU2 #HX2 destruct /3 width=4/
-| #V1 #V2 #W #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #X1 #HX1 #X2 #HX2
+| #a #V1 #V2 #W #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #X1 #HX1 #X2 #HX2
elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct
elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct
elim (lift_inv_bind1 … HX2) -HX2 #V3 #T3 #HV23 #HT23 #HX2 destruct /3 width=4/
-| #I #V1 #V2 #T1 #T #T2 #_ #_ #HT2 #IHV12 #IHT1 #d #e #X1 #HX1 #X2 #HX2
+| #a #I #V1 #V2 #T1 #T #T2 #_ #_ #HT2 #IHV12 #IHT1 #d #e #X1 #HX1 #X2 #HX2
elim (lift_inv_bind1 … HX1) -HX1 #W1 #U1 #HVW1 #HTU1 #HX1 destruct
elim (lift_inv_bind1 … HX2) -HX2 #W2 #U0 #HVW2 #HTU0 #HX2 destruct
elim (lift_total T (d + 1) e) #U #HTU
@tpr_delta
[4: @(tps_lift_le … HT2 … HTU HTU0 ?) /2 width=1/ |1: skip |2: /2 width=4/ |3: /2 width=4/ ] (**) (*/3. is too slow *)
-| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #d #e #X1 #HX1 #X2 #HX2
+| #a #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #d #e #X1 #HX1 #X2 #HX2
elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct
elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct
elim (lift_inv_bind1 … HX2) -HX2 #W3 #X #HW23 #HX #HX2 destruct
elim (lift_inv_flat2 … HX) -HX #V0 #T0 #HV01 #HT01 #HX destruct
elim (IHV12 … HV01) -V1
elim (IHT12 … HT01) -T1 /3 width=5/
-| #V1 #V2 #W1 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #X #HX
+| #a #V1 #V2 #W1 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #X #HX
elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct
elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct
elim (IHV12 … HV01) -V1
elim (IHT12 … HT01) -T1 /3 width=5/
-| #I #V1 #V2 #T1 #T #T2 #_ #_ #HT2 #IHV12 #IHT1 #d #e #X #HX
+| #a #I #V1 #V2 #T1 #T #T2 #_ #_ #HT2 #IHV12 #IHT1 #d #e #X #HX
elim (lift_inv_bind2 … HX) -HX #W1 #U1 #HWV1 #HUT1 #HX destruct
elim (IHV12 … HWV1) -V1 #W2 #HWV2 #HW12
elim (IHT1 … HUT1) -T1 #U #HUT #HU1
elim (tps_inv_lift1_le … HT2 … HUT ?) -T // [3: /2 width=5/ |2: skip ] #U2 #HU2 #HUT2
@ex2_1_intro [2: /2 width=2/ |1: skip |3: /2 width=3/ ] (**) (* /3 width=5/ is slow *)
-| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #d #e #X #HX
+| #a #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #d #e #X #HX
elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct
elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct
elim (IHV12 … HV01) -V1 #V3 #HV32 #HV03
(* Advanced inversion lemmas ************************************************)
-fact tpr_inv_abst1_aux: ∀U1,U2. U1 ➡ U2 → ∀V1,T1. U1 = ⓛV1. T1 →
- ∃∃V2,T2. V1 ➡ V2 & T1 ➡ T2 & U2 = ⓛV2. T2.
+fact tpr_inv_abst1_aux: ∀U1,U2. U1 ➡ U2 → ∀a,V1,T1. U1 = ⓛ{a}V1. T1 →
+ ∃∃V2,T2. V1 ➡ V2 & T1 ➡ T2 & U2 = ⓛ{a}V2. T2.
#U1 #U2 * -U1 -U2
-[ #I #V #T #H destruct
-| #I #V1 #V2 #T1 #T2 #_ #_ #V #T #H destruct
-| #V1 #V2 #W #T1 #T2 #_ #_ #V #T #H destruct
-| #I #V1 #V2 #T1 #T #T2 #HV12 #HT1 #HT2 #V0 #T0 #H destruct
+[ #I #a #V #T #H destruct
+| #I #V1 #V2 #T1 #T2 #_ #_ #a #V #T #H destruct
+| #b #V1 #V2 #W #T1 #T2 #_ #_ #a #V #T #H destruct
+| #b #I #V1 #V2 #T1 #T #T2 #HV12 #HT1 #HT2 #a #V0 #T0 #H destruct
<(tps_inv_refl_SO2 … HT2 ? ? ?) -T2 /2 width=5/
-| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #V0 #T0 #H destruct
-| #V #T1 #T #T2 #_ #_ #V0 #T0 #H destruct
-| #V #T1 #T2 #_ #V0 #T0 #H destruct
+| #b #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #a #V0 #T0 #H destruct
+| #V #T1 #T #T2 #_ #_ #a #V0 #T0 #H destruct
+| #V #T1 #T2 #_ #a #V0 #T0 #H destruct
]
qed.
(* Basic_1: was pr0_gen_abst *)
-lemma tpr_inv_abst1: ∀V1,T1,U2. ⓛV1. T1 ➡ U2 →
- ∃∃V2,T2. V1 ➡ V2 & T1 ➡ T2 & U2 = ⓛV2. T2.
+lemma tpr_inv_abst1: ∀a,V1,T1,U2. ⓛ{a}V1. T1 ➡ U2 →
+ ∃∃V2,T2. V1 ➡ V2 & T1 ➡ T2 & U2 = ⓛ{a}V2. T2.
/2 width=3/ qed-.
qed.
fact tpr_conf_flat_beta:
- ∀V0,V1,T1,V2,W0,U0,T2. (
+ ∀a,V0,V1,T1,V2,W0,U0,T2. (
∀X0:term. #[X0] < #[V0] + (#[W0] + #[U0] + 1) + 1 →
∀X1,X2. X0 ➡ X1 → X0 ➡ X2 →
∃∃X. X1 ➡ X & X2 ➡ X
) →
V0 ➡ V1 → V0 ➡ V2 →
- U0 ➡ T2 → ⓛW0. U0 ➡ T1 →
- ∃∃X. ⓐV1. T1 ➡ X & ⓓV2. T2 ➡ X.
-#V0 #V1 #T1 #V2 #W0 #U0 #T2 #IH #HV01 #HV02 #HT02 #H
+ U0 ➡ T2 → ⓛ{a}W0. U0 ➡ T1 →
+ ∃∃X. ⓐV1. T1 ➡ X & ⓓ{a}V2. T2 ➡ X.
+#a #V0 #V1 #T1 #V2 #W0 #U0 #T2 #IH #HV01 #HV02 #HT02 #H
elim (tpr_inv_abst1 … H) -H #W1 #U1 #HW01 #HU01 #H destruct
elim (IH … HV01 … HV02) -HV01 -HV02 /2 width=1/ #V #HV1 #HV2
elim (IH … HT02 … HU01) -HT02 -HU01 -IH /2 width=1/ /3 width=5/
pr0_cong_upsilon_cong pr0_cong_upsilon_delta
*)
fact tpr_conf_flat_theta:
- ∀V0,V1,T1,V2,V,W0,W2,U0,U2. (
+ ∀a,V0,V1,T1,V2,V,W0,W2,U0,U2. (
∀X0:term. #[X0] < #[V0] + (#[W0] + #[U0] + 1) + 1 →
∀X1,X2. X0 ➡ X1 → X0 ➡ X2 →
∃∃X. X1 ➡ X & X2 ➡ X
) →
V0 ➡ V1 → V0 ➡ V2 → ⇧[O,1] V2 ≡ V →
- W0 ➡ W2 → U0 ➡ U2 → ⓓW0. U0 ➡ T1 →
- ∃∃X. ⓐV1. T1 ➡ X & ⓓW2. ⓐV. U2 ➡ X.
-#V0 #V1 #T1 #V2 #V #W0 #W2 #U0 #U2 #IH #HV01 #HV02 #HV2 #HW02 #HU02 #H
+ W0 ➡ W2 → U0 ➡ U2 → ⓓ{a}W0. U0 ➡ T1 →
+ ∃∃X. ⓐV1. T1 ➡ X & ⓓ{a}W2. ⓐV. U2 ➡ X.
+#a #V0 #V1 #T1 #V2 #V #W0 #W2 #U0 #U2 #IH #HV01 #HV02 #HV2 #HW02 #HU02 #H
elim (IH … HV01 … HV02) -HV01 -HV02 /2 width=1/ #VV #HVV1 #HVV2
elim (lift_total VV 0 1) #VVV #HVV
lapply (tpr_lift … HVV2 … HV2 … HVV) #HVVV
(* case 3: zeta *)
| -HV2 -HW02 -HVV2 #U1 #HU01 #HTU1
elim (IH … HU01 … HU02) -HU01 -HU02 -IH // -U0 #U #HU1 #HU2
- elim (tpr_inv_lift1 … HU1 … HTU1) -U1 #UU #HUU #HT1UU
+ elim (tpr_inv_lift1 … HU1 … HTU1) -U1 #UU #HUU #HT1UU #H destruct
@(ex2_1_intro … (ⓐVV.UU)) /2 width=1/ /3 width=5/ (**) (* /4 width=9/ is too slow *)
]
qed.
qed.
fact tpr_conf_beta_beta:
- ∀W0:term. ∀V0,V1,T0,T1,V2,T2. (
+ ∀a. ∀W0:term. ∀V0,V1,T0,T1,V2,T2. (
∀X0:term. #[X0] < #[V0] + (#[W0] + #[T0] + 1) + 1 →
∀X1,X2. X0 ➡ X1 → X0 ➡ X2 →
∃∃X. X1 ➡ X & X2 ➡ X
) →
V0 ➡ V1 → V0 ➡ V2 → T0 ➡ T1 → T0 ➡ T2 →
- ∃∃X. ⓓV1. T1 ➡X & ⓓV2. T2 ➡ X.
-#W0 #V0 #V1 #T0 #T1 #V2 #T2 #IH #HV01 #HV02 #HT01 #HT02
+ ∃∃X. ⓓ{a}V1. T1 ➡X & ⓓ{a}V2. T2 ➡ X.
+#a #W0 #V0 #V1 #T0 #T1 #V2 #T2 #IH #HV01 #HV02 #HT01 #HT02
elim (IH … HV01 … HV02) -HV01 -HV02 /2 width=1/
elim (IH … HT01 … HT02) -HT01 -HT02 -IH /2 width=1/ /3 width=5/
qed.
(* Basic_1: was: pr0_cong_delta pr0_delta_delta *)
fact tpr_conf_delta_delta:
- ∀I1,V0,V1,T0,T1,TT1,V2,T2,TT2. (
+ ∀a,I1,V0,V1,T0,T1,TT1,V2,T2,TT2. (
∀X0:term. #[X0] < #[V0] + #[T0] + 1 →
∀X1,X2. X0 ➡ X1 → X0 ➡ X2 →
∃∃X. X1 ➡ X & X2 ➡ X
V0 ➡ V1 → V0 ➡ V2 → T0 ➡ T1 → T0 ➡ T2 →
⋆. ⓑ{I1} V1 ⊢ T1 ▶ [O, 1] TT1 →
⋆. ⓑ{I1} V2 ⊢ T2 ▶ [O, 1] TT2 →
- ∃∃X. ⓑ{I1} V1. TT1 ➡ X & ⓑ{I1} V2. TT2 ➡ X.
-#I1 #V0 #V1 #T0 #T1 #TT1 #V2 #T2 #TT2 #IH #HV01 #HV02 #HT01 #HT02 #HTT1 #HTT2
+ ∃∃X. ⓑ{a,I1} V1. TT1 ➡ X & ⓑ{a,I1} V2. TT2 ➡ X.
+#a #I1 #V0 #V1 #T0 #T1 #TT1 #V2 #T2 #TT2 #IH #HV01 #HV02 #HT01 #HT02 #HTT1 #HTT2
elim (IH … HV01 … HV02) -HV01 -HV02 // #V #HV1 #HV2
elim (IH … HT01 … HT02) -HT01 -HT02 -IH // #T #HT1 #HT2
elim (tpr_tps_bind … HV1 HT1 … HTT1) -T1 #U1 #TTU1 #HTU1
) →
V0 ➡ V1 → T0 ➡ T1 → ⋆. ⓓV1 ⊢ T1 ▶ [O,1] TT1 →
T0 ➡ T2 → ⇧[O, 1] X2 ≡ T2 →
- ∃∃X. ⓓV1. TT1 ➡ X & X2 ➡ X.
+ ∃∃X. +ⓓV1. TT1 ➡ X & X2 ➡ X.
#X2 #V0 #V1 #T0 #T1 #TT1 #T2 #IH #_ #HT01 #HTT1 #HT02 #HXT2
elim (IH … HT01 … HT02) -IH -HT01 -HT02 // -V0 -T0 #T #HT1 #HT2
elim (tpr_tps_bind ? ? V1 … HT1 HTT1) -T1 // #TT #HTT1 #HTT
(* Basic_1: was: pr0_upsilon_upsilon *)
fact tpr_conf_theta_theta:
- ∀VV1,V0,V1,W0,W1,T0,T1,V2,VV2,W2,T2. (
+ ∀a,VV1,V0,V1,W0,W1,T0,T1,V2,VV2,W2,T2. (
∀X0:term. #[X0] < #[V0] + (#[W0] + #[T0] + 1) + 1 →
∀X1,X2. X0 ➡ X1 → X0 ➡ X2 →
∃∃X. X1 ➡ X & X2 ➡ X
) →
V0 ➡ V1 → V0 ➡ V2 → W0 ➡ W1 → W0 ➡ W2 → T0 ➡ T1 → T0 ➡ T2 →
⇧[O, 1] V1 ≡ VV1 → ⇧[O, 1] V2 ≡ VV2 →
- ∃∃X. ⓓW1. ⓐVV1. T1 ➡ X & ⓓW2. ⓐVV2. T2 ➡ X.
-#VV1 #V0 #V1 #W0 #W1 #T0 #T1 #V2 #VV2 #W2 #T2 #IH #HV01 #HV02 #HW01 #HW02 #HT01 #HT02 #HVV1 #HVV2
+ ∃∃X. ⓓ{a}W1. ⓐVV1. T1 ➡ X & ⓓ{a}W2. ⓐVV2. T2 ➡ X.
+#a #VV1 #V0 #V1 #W0 #W1 #T0 #T1 #V2 #VV2 #W2 #T2 #IH #HV01 #HV02 #HW01 #HW02 #HT01 #HT02 #HVV1 #HVV2
elim (IH … HV01 … HV02) -HV01 -HV02 /2 width=1/ #V #HV1 #HV2
elim (IH … HW01 … HW02) -HW01 -HW02 /2 width=1/ #W #HW1 #HW2
elim (IH … HT01 … HT02) -HT01 -HT02 -IH /2 width=1/ #T #HT1 #HT2
[ #V2 #T2 #HV02 #HT02 #H destruct
/3 width=7 by tpr_conf_flat_flat/ (**) (* /3 width=7/ is too slow *)
(* case 3: flat, beta *)
- | #V2 #W #U0 #T2 #HV02 #HT02 #H1 #H2 #H3 destruct
+ | #b #V2 #W #U0 #T2 #HV02 #HT02 #H1 #H2 #H3 destruct
/3 width=8 by tpr_conf_flat_beta/ (**) (* /3 width=8/ is too slow *)
(* case 4: flat, theta *)
- | #V2 #V #W0 #W2 #U0 #U2 #HV02 #HW02 #HT02 #HV2 #H1 #H2 #H3 destruct
+ | #b #V2 #V #W0 #W2 #U0 #U2 #HV02 #HW02 #HT02 #HV2 #H1 #H2 #H3 destruct
/3 width=11 by tpr_conf_flat_theta/ (**) (* /3 width=11/ is too slow *)
(* case 5: flat, tau *)
| #HT02 #H destruct
/3 width=6 by tpr_conf_flat_cast/ (**) (* /3 width=6/ is too slow *)
]
-| #V0 #V1 #W0 #T0 #T1 #HV01 #HT01 #H1 #H2 destruct
+| #a #V0 #V1 #W0 #T0 #T1 #HV01 #HT01 #H1 #H2 destruct
elim (tpr_inv_appl1 … H1) -H1 *
(* case 6: beta, flat (repeated) *)
[ #V2 #T2 #HV02 #HT02 #H destruct
@ex2_1_comm /3 width=8 by tpr_conf_flat_beta/
(* case 7: beta, beta *)
- | #V2 #WW0 #TT0 #T2 #HV02 #HT02 #H1 #H2 destruct
+ | #b #V2 #WW0 #TT0 #T2 #HV02 #HT02 #H1 #H2 destruct
/3 width=8 by tpr_conf_beta_beta/ (**) (* /3 width=8/ is too slow *)
(* case 8, beta, theta (excluded) *)
- | #V2 #VV2 #WW0 #W2 #TT0 #T2 #_ #_ #_ #_ #H destruct
+ | #b #V2 #VV2 #WW0 #W2 #TT0 #T2 #_ #_ #_ #_ #H destruct
]
-| #I1 #V0 #V1 #T0 #T1 #TT1 #HV01 #HT01 #HTT1 #H1 #H2 destruct
+| #a #I1 #V0 #V1 #T0 #T1 #TT1 #HV01 #HT01 #HTT1 #H1 #H2 destruct
elim (tpr_inv_bind1 … H1) -H1 *
(* case 9: delta, delta *)
[ #V2 #T2 #TT2 #HV02 #HT02 #HTT2 #H destruct
/3 width=11 by tpr_conf_delta_delta/ (**) (* /3 width=11/ is too slow *)
(* case 10: delta, zeta *)
- | #T2 #HT20 #HTX2 #H destruct
+ | #T2 #HT20 #HTX2 #H1 #H2 destruct
/3 width=10 by tpr_conf_delta_zeta/ (**) (* /3 width=10/ is too slow *)
]
-| #VV1 #V0 #V1 #W0 #W1 #T0 #T1 #HV01 #HVV1 #HW01 #HT01 #H1 #H2 destruct
+| #a #VV1 #V0 #V1 #W0 #W1 #T0 #T1 #HV01 #HVV1 #HW01 #HT01 #H1 #H2 destruct
elim (tpr_inv_appl1 … H1) -H1 *
(* case 11: theta, flat (repeated) *)
[ #V2 #T2 #HV02 #HT02 #H destruct
@ex2_1_comm /3 width=11 by tpr_conf_flat_theta/
(* case 12: theta, beta (repeated) *)
- | #V2 #WW0 #TT0 #T2 #_ #_ #H destruct
+ | #b #V2 #WW0 #TT0 #T2 #_ #_ #H destruct
(* case 13: theta, theta *)
- | #V2 #VV2 #WW0 #W2 #TT0 #T2 #V02 #HW02 #HT02 #HVV2 #H1 #H2 destruct
+ | #b #V2 #VV2 #WW0 #W2 #TT0 #T2 #V02 #HW02 #HT02 #HVV2 #H1 #H2 destruct
/3 width=14 by tpr_conf_theta_theta/ (**) (* /3 width=14/ is too slow *)
]
| #V0 #TT0 #T0 #T1 #HTT0 #HT01 #H1 #H2 destruct
(* case 15: zeta, zeta *)
| #TT2 #HTT02 #HXTT2
/3 width=9 by tpr_conf_zeta_zeta/ (**) (* /3 width=9/ is too slow *)
- ]
+ ]
| #V0 #T0 #T1 #HT01 #H1 #H2 destruct
elim (tpr_inv_cast1 … H1) -H1
(* case 16: tau, flat (repeated) *)
elim (tps_inv_flat1 … H) -H #W1 #U1 #HVW1 #HTU1 #H destruct
elim (IHV12 … HVW1 … HL12) -V1
elim (IHT12 … HTU1 … HL12) -T1 -HL12 /3 width=5/
-| #V1 #V2 #W #T1 #T2 #_ #_ #IHV12 #IHT12 #L1 #d #e #X #H #L2 #HL12
+| #a #V1 #V2 #W #T1 #T2 #_ #_ #IHV12 #IHT12 #L1 #d #e #X #H #L2 #HL12
elim (tps_inv_flat1 … H) -H #VV1 #Y #HVV1 #HY #HX destruct
elim (tps_inv_bind1 … HY) -HY #WW #TT1 #_ #HTT1 #H destruct
elim (IHV12 … HVV1 … HL12) -V1 #VV2 #HVV12 #HVV2
elim (IHT12 … HTT1 (L2. ⓛWW) ?) -T1 /2 width=1/ -HL12 #TT2 #HTT12 #HTT2
lapply (tpss_lsubs_trans … HTT2 (L2. ⓓVV2) ?) -HTT2 /3 width=5/
-| #I #V1 #V2 #T1 #T #T2 #HV12 #_ #HT2 #IHV12 #IHT1 #L1 #d #e #X #H #L2 #HL12
+| #a #I #V1 #V2 #T1 #T #T2 #HV12 #_ #HT2 #IHV12 #IHT1 #L1 #d #e #X #H #L2 #HL12
elim (tps_inv_bind1 … H) -H #W1 #U1 #HVW1 #HTU1 #H destruct
elim (IHV12 … HVW1 … HL12) -V1 #W2 #HW12 #HVW2
elim (IHT1 … HTU1 (L2. ⓑ{I} W2) ?) -T1 /2 width=1/ -HL12 #U #HU1 #HTU
lapply (tps_lsubs_trans … HU3 (⋆. ⓑ{I} W2) ?) -HU3 /2 width=1/ #HU3
lapply (tpss_lsubs_trans … HU23 (L2. ⓑ{I} W2) ?) -HU23 /2 width=1/ #HU23
lapply (tpss_trans_eq … HTU2 … HU23) -U2 /3 width=5/
-| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #L1 #d #e #X #H #L2 #HL12
+| #a #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #L1 #d #e #X #H #L2 #HL12
elim (tps_inv_flat1 … H) -H #VV1 #Y #HVV1 #HY #HX destruct
elim (tps_inv_bind1 … HY) -HY #WW1 #TT1 #HWW1 #HTT1 #H destruct
elim (IHV12 … HVV1 … HL12) -V1 #VV2 #HVV12 #HVV2
elim (tshf_inv_flat1 … H) -H #W2 #U2 #HT1U2 #HT1 #_ #H1 #H2 destruct
lapply (IHT12 HT1U2) -IHT12 -HT1U2 #HUT2
lapply (simple_tshf_repl_dx … HUT2 HT1) /2 width=1/
-| #V1 #V2 #W #T1 #T2 #_ #_ #_ #_ #H
+| #a #V1 #V2 #W #T1 #T2 #_ #_ #_ #_ #H
elim (tshf_inv_flat1 … H) -H #W2 #U2 #_ #H
elim (simple_inv_bind … H)
-| #I #V1 #V2 #T1 #T #T2 #_ #_ #_ #_ #_ #H
- elim (tshf_inv_bind1 … H) -H #W2 #U2 #H destruct //
-| #V2 #V1 #V #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #_ #H
+| #a #I #V1 #V2 #T1 #T #T2 #_ #_ #_ #_ #_ #H
+ elim (tshf_inv_bind1 … H) -H #W2 #U2 #H1 * #H2 destruct //
+| #a #V2 #V1 #V #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #_ #H
elim (tshf_inv_flat1 … H) -H #U1 #U2 #_ #H
elim (simple_inv_bind … H)
| #V #T #T1 #T2 #_ #_ #_ #H
- elim (tshf_inv_bind1 … H) -H #W2 #U2 #H destruct
+ elim (tshf_inv_bind1 … H) -H #W2 #U2 #H1 * #H2 destruct
| #V #T1 #T2 #_ #_ #H
elim (tshf_inv_flat1 … H) -H #W2 #U2 #_ #_ #_ #H destruct
]
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/static/ssta_aaa.ma".
+include "basic_2/reducibility/lcpr_aaa.ma".
+include "basic_2/reducibility/xpr.ma".
+
+(* EXTENDED PARALLEL REDUCTION ON TERMS *************************************)
+
+(* Properties on atomic arity assignment for terms **************************)
+
+lemma xpr_aaa: ∀h,g,L,T,A. L ⊢ T ⁝ A → ∀U. ⦃h, L⦄ ⊢ T ➸[g] U → L ⊢ U ⁝ A.
+#h #g #L #T #A #HT #U * [2: * ] /2 width=3/ /2 width=6/
+qed.
(* Advanced inversion lemmas ************************************************)
-lemma xpr_inv_abst1: ∀h,g,L,V1,T1,U2. ⦃h, L⦄ ⊢ ⓛV1.T1 ➸[g] U2 →
+lemma xpr_inv_abst1: ∀h,g,a,L,V1,T1,U2. ⦃h, L⦄ ⊢ ⓛ{a}V1.T1 ➸[g] U2 →
∃∃V2,T2. L ⊢ V1 ➡ V2 & ⦃h, L. ⓛV1⦄ ⊢ T1 ➸[g] T2 &
- U2 = ⓛV2. T2.
-#h #g #L #V1 #T1 #U2 *
+ U2 = ⓛ{a}V2. T2.
+#h #g #a #L #V1 #T1 #U2 *
[ #H elim (cpr_inv_abst1 … H Abst V1) /3 width=5/
| * #l #H elim (ssta_inv_bind1 … H) /3 width=5/
]
inductive aaa: lenv → term → predicate aarity ≝
| aaa_sort: ∀L,k. aaa L (⋆k) ⓪
| aaa_lref: ∀I,L,K,V,B,i. ⇩[0, i] L ≡ K. ⓑ{I} V → aaa K V B → aaa L (#i) B
-| aaa_abbr: ∀L,V,T,B,A.
- aaa L V B → aaa (L. ⓓV) T A → aaa L (ⓓV. T) A
-| aaa_abst: ∀L,V,T,B,A.
- aaa L V B → aaa (L. ⓛV) T A → aaa L (ⓛV. T) (②B. A)
+| aaa_abbr: ∀a,L,V,T,B,A.
+ aaa L V B → aaa (L. ⓓV) T A → aaa L (ⓓ{a}V. T) A
+| aaa_abst: ∀a,L,V,T,B,A.
+ aaa L V B → aaa (L. ⓛV) T A → aaa L (ⓛ{a}V. T) (②B. A)
| aaa_appl: ∀L,V,T,B,A. aaa L V B → aaa L T (②B. A) → aaa L (ⓐV. T) A
| aaa_cast: ∀L,V,T,A. aaa L V A → aaa L T A → aaa L (ⓝV. T) A
.
#L #T #A * -L -T -A
[ //
| #I #L #K #V #B #i #_ #_ #k #H destruct
-| #L #V #T #B #A #_ #_ #k #H destruct
-| #L #V #T #B #A #_ #_ #k #H destruct
+| #a #L #V #T #B #A #_ #_ #k #H destruct
+| #a #L #V #T #B #A #_ #_ #k #H destruct
| #L #V #T #B #A #_ #_ #k #H destruct
| #L #V #T #A #_ #_ #k #H destruct
]
#L #T #A * -L -T -A
[ #L #k #i #H destruct
| #I #L #K #V #B #j #HLK #HB #i #H destruct /2 width=5/
-| #L #V #T #B #A #_ #_ #i #H destruct
-| #L #V #T #B #A #_ #_ #i #H destruct
+| #a #L #V #T #B #A #_ #_ #i #H destruct
+| #a #L #V #T #B #A #_ #_ #i #H destruct
| #L #V #T #B #A #_ #_ #i #H destruct
| #L #V #T #A #_ #_ #i #H destruct
]
∃∃I,K,V. ⇩[0, i] L ≡ K. ⓑ{I} V & K ⊢ V ⁝ A.
/2 width=3/ qed-.
-fact aaa_inv_abbr_aux: ∀L,T,A. L ⊢ T ⁝ A → ∀W,U. T = ⓓW. U →
+fact aaa_inv_abbr_aux: ∀L,T,A. L ⊢ T ⁝ A → ∀a,W,U. T = ⓓ{a}W. U →
∃∃B. L ⊢ W ⁝ B & L. ⓓW ⊢ U ⁝ A.
#L #T #A * -L -T -A
-[ #L #k #W #U #H destruct
-| #I #L #K #V #B #i #_ #_ #W #U #H destruct
-| #L #V #T #B #A #HV #HT #W #U #H destruct /2 width=2/
-| #L #V #T #B #A #_ #_ #W #U #H destruct
-| #L #V #T #B #A #_ #_ #W #U #H destruct
-| #L #V #T #A #_ #_ #W #U #H destruct
+[ #L #k #a #W #U #H destruct
+| #I #L #K #V #B #i #_ #_ #a #W #U #H destruct
+| #b #L #V #T #B #A #HV #HT #a #W #U #H destruct /2 width=2/
+| #b #L #V #T #B #A #_ #_ #a #W #U #H destruct
+| #L #V #T #B #A #_ #_ #a #W #U #H destruct
+| #L #V #T #A #_ #_ #a #W #U #H destruct
]
qed.
-lemma aaa_inv_abbr: ∀L,V,T,A. L ⊢ ⓓV. T ⁝ A →
+lemma aaa_inv_abbr: ∀a,L,V,T,A. L ⊢ ⓓ{a}V. T ⁝ A →
∃∃B. L ⊢ V ⁝ B & L. ⓓV ⊢ T ⁝ A.
-/2 width=3/ qed-.
+/2 width=4/ qed-.
-fact aaa_inv_abst_aux: ∀L,T,A. L ⊢ T ⁝ A → ∀W,U. T = ⓛW. U →
+fact aaa_inv_abst_aux: ∀L,T,A. L ⊢ T ⁝ A → ∀a,W,U. T = ⓛ{a}W. U →
∃∃B1,B2. L ⊢ W ⁝ B1 & L. ⓛW ⊢ U ⁝ B2 & A = ②B1. B2.
#L #T #A * -L -T -A
-[ #L #k #W #U #H destruct
-| #I #L #K #V #B #i #_ #_ #W #U #H destruct
-| #L #V #T #B #A #_ #_ #W #U #H destruct
-| #L #V #T #B #A #HV #HT #W #U #H destruct /2 width=5/
-| #L #V #T #B #A #_ #_ #W #U #H destruct
-| #L #V #T #A #_ #_ #W #U #H destruct
+[ #L #k #a #W #U #H destruct
+| #I #L #K #V #B #i #_ #_ #a #W #U #H destruct
+| #b #L #V #T #B #A #_ #_ #a #W #U #H destruct
+| #b #L #V #T #B #A #HV #HT #a #W #U #H destruct /2 width=5/
+| #L #V #T #B #A #_ #_ #a #W #U #H destruct
+| #L #V #T #A #_ #_ #a #W #U #H destruct
]
qed.
-lemma aaa_inv_abst: ∀L,W,T,A. L ⊢ ⓛW. T ⁝ A →
+lemma aaa_inv_abst: ∀a,L,W,T,A. L ⊢ ⓛ{a}W. T ⁝ A →
∃∃B1,B2. L ⊢ W ⁝ B1 & L. ⓛW ⊢ T ⁝ B2 & A = ②B1. B2.
-/2 width=3/ qed-.
+/2 width=4/ qed-.
fact aaa_inv_appl_aux: ∀L,T,A. L ⊢ T ⁝ A → ∀W,U. T = ⓐW. U →
∃∃B. L ⊢ W ⁝ B & L ⊢ U ⁝ ②B. A.
#L #T #A * -L -T -A
[ #L #k #W #U #H destruct
| #I #L #K #V #B #i #_ #_ #W #U #H destruct
-| #L #V #T #B #A #_ #_ #W #U #H destruct
-| #L #V #T #B #A #_ #_ #W #U #H destruct
+| #a #L #V #T #B #A #_ #_ #W #U #H destruct
+| #a #L #V #T #B #A #_ #_ #W #U #H destruct
| #L #V #T #B #A #HV #HT #W #U #H destruct /2 width=3/
| #L #V #T #A #_ #_ #W #U #H destruct
]
#L #T #A * -L -T -A
[ #L #k #W #U #H destruct
| #I #L #K #V #B #i #_ #_ #W #U #H destruct
-| #L #V #T #B #A #_ #_ #W #U #H destruct
-| #L #V #T #B #A #_ #_ #W #U #H destruct
+| #a #L #V #T #B #A #_ #_ #W #U #H destruct
+| #a #L #V #T #B #A #_ #_ #W #U #H destruct
| #L #V #T #B #A #_ #_ #W #U #H destruct
| #L #V #T #A #HV #HT #W #U #H destruct /2 width=1/
]
| #I1 #L #K1 #V1 #B #i #HLK1 #_ #IHA1 #A2 #H
elim (aaa_inv_lref … H) -H #I2 #K2 #V2 #HLK2 #HA2
lapply (ldrop_mono … HLK1 … HLK2) -L #H destruct /2 width=1/
-| #L #V #T #B1 #A1 #_ #_ #_ #IHA1 #A2 #H
+| #a #L #V #T #B1 #A1 #_ #_ #_ #IHA1 #A2 #H
elim (aaa_inv_abbr … H) -H /2 width=1/
-| #L #V1 #T1 #B1 #A1 #_ #_ #IHB1 #IHA1 #X #H
+| #a #L #V1 #T1 #B1 #A1 #_ #_ #IHB1 #IHA1 #X #H
elim (aaa_inv_abst … H) -H #B2 #A2 #HB2 #HA2 #H destruct /3 width=1/
| #L #V1 #T1 #B1 #A1 #_ #_ #_ #IHA1 #A2 #H
elim (aaa_inv_appl … H) -H #B2 #_ #HA2
/3 width=8/
| lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 // -Hid /3 width=8/
]
-| #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #d #e #HL21 #X #H
+| #a #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #d #e #HL21 #X #H
elim (lift_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
/4 width=4/
-| #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #d #e #HL21 #X #H
+| #a #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #d #e #HL21 #X #H
elim (lift_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
/4 width=4/
| #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #d #e #HL21 #X #H
[ elim (ldrop_conf_lt … HL21 … HLK2 ?) -L2 // -Hid /3 width=8/
| lapply (ldrop_conf_ge … HL21 … HLK2 ?) -L2 // -Hid /3 width=8/
]
-| #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #d #e #HL21 #X #H
+| #a #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #d #e #HL21 #X #H
elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
/4 width=4/
-| #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #d #e #HL21 #X #H
+| #a #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #d #e #HL21 #X #H
elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
/4 width=4/
| #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #d #e #HL21 #X #H
elim (lt_or_ge i d) #Hdi
[ elim (ltpss_ldrop_conf_le … HL12 … HLK1 ?) -L1 /2 width=2/ #X #H #HLK2
elim (ltpss_inv_tpss11 … H ?) -H /2 width=1/ -Hdi #K2 #V2 #HK12 #HV12 #H destruct
- /3 width=8/
+ /3 width=8 by aaa_lref/ (**) (* too slow without trace *)
| elim (lt_or_ge i (d + e)) #Hide
[ elim (ltpss_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HLK2
elim (ltpss_inv_tpss21 … H ?) -H /2 width=1/ -Hdi -Hide #K2 #V2 #HK12 #HV12 #H destruct
- /3 width=8/
+ /3 width=8 by aaa_lref/ (**) (* too slow without trace *)
| -Hdi
- lapply (ltpss_ldrop_conf_ge … HL12 … HLK1 ?) -L1 // -Hide /3 width=8/
+ lapply (ltpss_ldrop_conf_ge … HL12 … HLK1 ?) -L1 // -Hide
+ /3 width=8 by aaa_lref/ (**) (* too slow without trace *)
]
]
| * #K2 #V2 #W2 #Hdi #Hide #HLK2 #HVW2 #HWT2
lapply (ldrop_fwd_ldrop2 … HLK2) -HLK2 #HLK2
lapply (tpss_trans_eq … HV12 HVW2) -V2 /3 width=7/
]
-| #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #L2 #d #e #HL12 #X #H
+| #a #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #L2 #d #e #HL12 #X #H
elim (tpss_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct /4 width=4/
-| #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #L2 #d #e #HL12 #X #H
+| #a #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #L2 #d #e #HL12 #X #H
elim (tpss_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct /4 width=4/
| #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #L2 #d #e #HL12 #X #H
elim (tpss_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct /3 width=4/
⇧[0, i + 1] W ≡ U → ssta h g l L (#i) U
| ssta_ldec: ∀L,K,W,V,U,i,l. ⇩[0, i] L ≡ K. ⓛW → ssta h g l K W V →
⇧[0, i + 1] W ≡ U → ssta h g (l+1) L (#i) U
-| ssta_bind: ∀I,L,V,T,U,l. ssta h g l (L. ⓑ{I} V) T U →
- ssta h g l L (ⓑ{I}V.T) (ⓑ{I}V.U)
+| ssta_bind: ∀a,I,L,V,T,U,l. ssta h g l (L. ⓑ{I} V) T U →
+ ssta h g l L (ⓑ{a,I}V.T) (ⓑ{a,I}V.U)
| ssta_appl: ∀L,V,T,U,l. ssta h g l L T U →
ssta h g l L (ⓐV.T) (ⓐV.U)
| ssta_cast: ∀L,V,W,T,U,l. ssta h g (l - 1) L V W → ssta h g l L T U →
[ #L #k #l #Hkl #k0 #H destruct /2 width=1/
| #L #K #V #W #U #i #l #_ #_ #_ #k0 #H destruct
| #L #K #W #V #U #i #l #_ #_ #_ #k0 #H destruct
-| #I #L #V #T #U #l #_ #k0 #H destruct
+| #a #I #L #V #T #U #l #_ #k0 #H destruct
| #L #V #T #U #l #_ #k0 #H destruct
| #L #V #W #T #U #l #_ #_ #k0 #H destruct
qed.
[ #L #k #l #_ #j #H destruct
| #L #K #V #W #U #i #l #HLK #HVW #HWU #j #H destruct /3 width=6/
| #L #K #W #V #U #i #l #HLK #HWV #HWU #j #H destruct /3 width=8/
-| #I #L #V #T #U #l #_ #j #H destruct
+| #a #I #L #V #T #U #l #_ #j #H destruct
| #L #V #T #U #l #_ #j #H destruct
| #L #V #W #T #U #l #_ #_ #j #H destruct
]
/2 width=3/ qed-.
fact ssta_inv_bind1_aux: ∀h,g,L,T,U,l. ⦃h, L⦄ ⊢ T •[g, l] U →
- ∀J,X,Y. T = ⓑ{J}Y.X →
- ∃∃Z. ⦃h, L.ⓑ{J}Y⦄ ⊢ X •[g, l] Z & U = ⓑ{J}Y.Z.
+ ∀a,I,X,Y. T = ⓑ{a,I}Y.X →
+ ∃∃Z. ⦃h, L.ⓑ{I}Y⦄ ⊢ X •[g, l] Z & U = ⓑ{a,I}Y.Z.
#h #g #L #T #U #l * -L -T -U -l
-[ #L #k #l #_ #J #X #Y #H destruct
-| #L #K #V #W #U #i #l #_ #_ #_ #J #X #Y #H destruct
-| #L #K #W #V #U #i #l #_ #_ #_ #J #X #Y #H destruct
-| #I #L #V #T #U #l #HTU #J #X #Y #H destruct /2 width=3/
-| #L #V #T #U #l #_ #J #X #Y #H destruct
-| #L #V #W #T #U #l #_ #_ #J #X #Y #H destruct
+[ #L #k #l #_ #a #I #X #Y #H destruct
+| #L #K #V #W #U #i #l #_ #_ #_ #a #I #X #Y #H destruct
+| #L #K #W #V #U #i #l #_ #_ #_ #a #I #X #Y #H destruct
+| #b #J #L #V #T #U #l #HTU #a #I #X #Y #H destruct /2 width=3/
+| #L #V #T #U #l #_ #a #I #X #Y #H destruct
+| #L #V #W #T #U #l #_ #_ #a #I #X #Y #H destruct
]
qed.
(* Basic_1: was just: sty0_gen_bind *)
-lemma ssta_inv_bind1: ∀h,g,J,L,Y,X,U,l. ⦃h, L⦄ ⊢ ⓑ{J}Y.X •[g, l] U →
- ∃∃Z. ⦃h, L.ⓑ{J}Y⦄ ⊢ X •[g, l] Z & U = ⓑ{J}Y.Z.
+lemma ssta_inv_bind1: ∀h,g,a,I,L,Y,X,U,l. ⦃h, L⦄ ⊢ ⓑ{a,I}Y.X •[g, l] U →
+ ∃∃Z. ⦃h, L.ⓑ{I}Y⦄ ⊢ X •[g, l] Z & U = ⓑ{a,I}Y.Z.
/2 width=3/ qed-.
fact ssta_inv_appl1_aux: ∀h,g,L,T,U,l. ⦃h, L⦄ ⊢ T •[g, l] U → ∀X,Y. T = ⓐY.X →
[ #L #k #l #_ #X #Y #H destruct
| #L #K #V #W #U #i #l #_ #_ #_ #X #Y #H destruct
| #L #K #W #V #U #i #l #_ #_ #_ #X #Y #H destruct
-| #I #L #V #T #U #l #_ #X #Y #H destruct
+| #a #I #L #V #T #U #l #_ #X #Y #H destruct
| #L #V #T #U #l #HTU #X #Y #H destruct /2 width=3/
| #L #V #W #T #U #l #_ #_ #X #Y #H destruct
]
[ #L #k #l #_ #X #Y #H destruct
| #L #K #V #W #U #l #i #_ #_ #_ #X #Y #H destruct
| #L #K #W #V #U #l #i #_ #_ #_ #X #Y #H destruct
-| #I #L #V #T #U #l #_ #X #Y #H destruct
+| #a #I #L #V #T #U #l #_ #X #Y #H destruct
| #L #V #T #U #l #_ #X #Y #H destruct
| #L #V #W #T #U #l #HVW #HTU #X #Y #H destruct /2 width=5/
]
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/static/aaa_lift.ma".
+include "basic_2/static/ssta.ma".
+
+(* STRATIFIED STATIC TYPE ASSIGNMENT ON TERMS *******************************)
+
+(* Properties on atomic arity assignment for terms **************************)
+
+lemma ssta_aaa: ∀h,g,L,T,A. L ⊢ T ⁝ A → ∀U,l. ⦃h, L⦄ ⊢ T •[g, l] U → L ⊢ U ⁝ A.
+#h #g #L #T #A #H elim H -L -T -A
+[ #L #k #U #l #H
+ elim (ssta_inv_sort1 … H) -H #_ #H destruct //
+| #I #L #K #V #B #i #HLK #HV #IHV #U #l #H
+ elim (ssta_inv_lref1 … H) -H * #K0 #V0 #W0 [2: #l0 ] #HLK0 #HVW0 #HU [ #H ]
+ lapply (ldrop_mono … HLK0 … HLK) -HLK0 #H0 destruct
+ lapply (ldrop_fwd_ldrop2 … HLK) -HLK #HLK
+ @(aaa_lift … HLK … HU) -HU -L // -HV /2 width=2/
+| #a #L #V #T #B #A #HV #_ #_ #IHT #X #l #H
+ elim (ssta_inv_bind1 … H) -H #U #HTU #H destruct /3 width=2/
+| #a #L #V #T #B #A #HV #_ #_ #IHT #X #l #H
+ elim (ssta_inv_bind1 … H) -H #U #HTU #H destruct /3 width=2/
+| #L #V #T #B #A #HV #_ #_ #IHT #X #l #H
+ elim (ssta_inv_appl1 … H) -H #U #HTU #H destruct /3 width=3/
+| #L #V #T #A #_ #_ #IHV #IHT #X #l #H
+ elim (ssta_inv_cast1 … H) -H #W #U #HVW #HTU #H destruct /3 width=2/
+]
+qed.
+
| lapply (lift_trans_be … HW1 … HWU2 ? ?) -W // /2 width=1/ #HW1U2
lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 // -Hid /3 width=8/
]
-| #I #L1 #V1 #T1 #U1 #l #_ #IHTU1 #L2 #d #e #HL21 #X1 #H1 #X2 #H2
+| #a #I #L1 #V1 #T1 #U1 #l #_ #IHTU1 #L2 #d #e #HL21 #X1 #H1 #X2 #H2
elim (lift_inv_bind1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H destruct
elim (lift_inv_bind1 … H2) -H2 #X #U2 #H1 #HU12 #H2 destruct
lapply (lift_mono … H1 … HV12) -H1 #H destruct /4 width=5/
| <le_plus_minus_comm // /2 width=1/
]
]
-| #I #L2 #V2 #T2 #U2 #l #_ #IHTU2 #L1 #d #e #HL21 #X #H
+| #a #I #L2 #V2 #T2 #U2 #l #_ #IHTU2 #L1 #d #e #HL21 #X #H
elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
elim (IHTU2 (L1.ⓑ{I}V1) … HT12) -IHTU2 -HT12 /2 width=1/ -HL21 /3 width=5/
| #L2 #V2 #T2 #U2 #l #_ #IHTU2 #L1 #d #e #HL21 #X #H
| #L #K #W #V #V0 #i #l #HLK #HWV #HWV0 #_
lapply (ldrop_fwd_ldrop2 … HLK) -HLK #HLK
elim (lift_total V 0 (i+1)) /3 width=10/
-| #I #L #V #T #U #l #_ * /3 width=2/
+| #a #I #L #V #T #U #l #_ * /3 width=2/
| #L #V #T #U #l #_ * #T0 #HUT0 /3 width=2/
| #L #V #W #T #U #l #_ #_ * #W0 #HW0 * /3 width=2/
]
elim (ltpss_inv_tpss21 … H ?) -H /2 width=1/ -Hdi -Hide #K0 #V0 #_ #_ #H destruct
lapply (ldrop_mono … HL2K0 … HLK2) -HL2K0 -HLK2 #H destruct
]
-| #I #L1 #V1 #T1 #U1 #l #_ #IHTU1 #L2 #d #e #HL12 #X #H
+| #a #I #L1 #V1 #T1 #U1 #l #_ #IHTU1 #L2 #d #e #HL12 #X #H
elim (tpss_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
elim (IHTU1 … HT12) -IHTU1 -HT12 /2 width=1/ -HL12 /3 width=5/
| #L1 #V1 #T1 #U1 #l #_ #IHTU1 #L2 #d #e #HL12 #X #H
lapply (ldrop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct
lapply (IHWV … HWV0) -IHWV -HWV0 * #H1 #H2 destruct
>(lift_mono … HWU1 … HV0U2) -W -U1 /2 width=1/
-| #I #L #V #T #U1 #l1 #_ #IHTU1 #X #l2 #H
+| #a #I #L #V #T #U1 #l1 #_ #IHTU1 #X #l2 #H
elim (ssta_inv_bind1 … H) -H #U2 #HTU2 #H destruct
elim (IHTU1 … HTU2) -T /3 width=1/
| #L #V #T #U1 #l1 #_ #IHTU1 #X #l2 #H
| lift_lref_lt: ∀i,d,e. i < d → lift d e (#i) (#i)
| lift_lref_ge: ∀i,d,e. d ≤ i → lift d e (#i) (#(i + e))
| lift_gref : ∀p,d,e. lift d e (§p) (§p)
-| lift_bind : ∀I,V1,V2,T1,T2,d,e.
+| lift_bind : ∀a,I,V1,V2,T1,T2,d,e.
lift d e V1 V2 → lift (d + 1) e T1 T2 →
- lift d e (ⓑ{I} V1. T1) (ⓑ{I} V2. T2)
+ lift d e (ⓑ{a,I} V1. T1) (ⓑ{a,I} V2. T2)
| lift_flat : ∀I,V1,V2,T1,T2,d,e.
lift d e V1 V2 → lift d e T1 T2 →
lift d e (ⓕ{I} V1. T1) (ⓕ{I} V2. T2)
fact lift_inv_sort1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀k. T1 = ⋆k → T2 = ⋆k.
#d #e #T1 #T2 * -d -e -T1 -T2 //
[ #i #d #e #_ #k #H destruct
-| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct
+| #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct
| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct
]
qed.
| #j #d #e #Hj #i #Hi destruct /3 width=1/
| #j #d #e #Hj #i #Hi destruct /3 width=1/
| #p #d #e #i #H destruct
-| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct
+| #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct
| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct
]
qed.
fact lift_inv_gref1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀p. T1 = §p → T2 = §p.
#d #e #T1 #T2 * -d -e -T1 -T2 //
[ #i #d #e #_ #k #H destruct
-| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct
+| #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct
| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct
]
qed.
/2 width=5/ qed-.
fact lift_inv_bind1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 →
- ∀I,V1,U1. T1 = ⓑ{I} V1.U1 →
+ ∀a,I,V1,U1. T1 = ⓑ{a,I} V1.U1 →
∃∃V2,U2. ⇧[d,e] V1 ≡ V2 & ⇧[d+1,e] U1 ≡ U2 &
- T2 = ⓑ{I} V2. U2.
+ T2 = ⓑ{a,I} V2. U2.
#d #e #T1 #T2 * -d -e -T1 -T2
-[ #k #d #e #I #V1 #U1 #H destruct
-| #i #d #e #_ #I #V1 #U1 #H destruct
-| #i #d #e #_ #I #V1 #U1 #H destruct
-| #p #d #e #I #V1 #U1 #H destruct
-| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V1 #U1 #H destruct /2 width=5/
-| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V1 #U1 #H destruct
+[ #k #d #e #a #I #V1 #U1 #H destruct
+| #i #d #e #_ #a #I #V1 #U1 #H destruct
+| #i #d #e #_ #a #I #V1 #U1 #H destruct
+| #p #d #e #a #I #V1 #U1 #H destruct
+| #b #J #W1 #W2 #T1 #T2 #d #e #HW #HT #a #I #V1 #U1 #H destruct /2 width=5/
+| #J #W1 #W2 #T1 #T2 #d #e #_ #HT #a #I #V1 #U1 #H destruct
]
qed.
-lemma lift_inv_bind1: ∀d,e,T2,I,V1,U1. ⇧[d,e] ⓑ{I} V1. U1 ≡ T2 →
+lemma lift_inv_bind1: ∀d,e,T2,a,I,V1,U1. ⇧[d,e] ⓑ{a,I} V1. U1 ≡ T2 →
∃∃V2,U2. ⇧[d,e] V1 ≡ V2 & ⇧[d+1,e] U1 ≡ U2 &
- T2 = ⓑ{I} V2. U2.
+ T2 = ⓑ{a,I} V2. U2.
/2 width=3/ qed-.
fact lift_inv_flat1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 →
| #i #d #e #_ #I #V1 #U1 #H destruct
| #i #d #e #_ #I #V1 #U1 #H destruct
| #p #d #e #I #V1 #U1 #H destruct
-| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V1 #U1 #H destruct
+| #a #J #W1 #W2 #T1 #T2 #d #e #_ #_ #I #V1 #U1 #H destruct
| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V1 #U1 #H destruct /2 width=5/
]
qed.
fact lift_inv_sort2_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀k. T2 = ⋆k → T1 = ⋆k.
#d #e #T1 #T2 * -d -e -T1 -T2 //
[ #i #d #e #_ #k #H destruct
-| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct
+| #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct
| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct
]
qed.
| #j #d #e #Hj #i #Hi destruct /3 width=1/
| #j #d #e #Hj #i #Hi destruct <minus_plus_m_m /4 width=1/
| #p #d #e #i #H destruct
-| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct
+| #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct
| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct
]
qed.
fact lift_inv_gref2_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀p. T2 = §p → T1 = §p.
#d #e #T1 #T2 * -d -e -T1 -T2 //
[ #i #d #e #_ #k #H destruct
-| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct
+| #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct
| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct
]
qed.
/2 width=5/ qed-.
fact lift_inv_bind2_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 →
- ∀I,V2,U2. T2 = ⓑ{I} V2.U2 →
+ ∀a,I,V2,U2. T2 = ⓑ{a,I} V2.U2 →
∃∃V1,U1. ⇧[d,e] V1 ≡ V2 & ⇧[d+1,e] U1 ≡ U2 &
- T1 = ⓑ{I} V1. U1.
+ T1 = ⓑ{a,I} V1. U1.
#d #e #T1 #T2 * -d -e -T1 -T2
-[ #k #d #e #I #V2 #U2 #H destruct
-| #i #d #e #_ #I #V2 #U2 #H destruct
-| #i #d #e #_ #I #V2 #U2 #H destruct
-| #p #d #e #I #V2 #U2 #H destruct
-| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V2 #U2 #H destruct /2 width=5/
-| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V2 #U2 #H destruct
+[ #k #d #e #a #I #V2 #U2 #H destruct
+| #i #d #e #_ #a #I #V2 #U2 #H destruct
+| #i #d #e #_ #a #I #V2 #U2 #H destruct
+| #p #d #e #a #I #V2 #U2 #H destruct
+| #b #J #W1 #W2 #T1 #T2 #d #e #HW #HT #a #I #V2 #U2 #H destruct /2 width=5/
+| #J #W1 #W2 #T1 #T2 #d #e #_ #_ #a #I #V2 #U2 #H destruct
]
qed.
(* Basic_1: was: lift_gen_bind *)
-lemma lift_inv_bind2: ∀d,e,T1,I,V2,U2. ⇧[d,e] T1 ≡ ⓑ{I} V2. U2 →
+lemma lift_inv_bind2: ∀d,e,T1,a,I,V2,U2. ⇧[d,e] T1 ≡ ⓑ{a,I} V2. U2 →
∃∃V1,U1. ⇧[d,e] V1 ≡ V2 & ⇧[d+1,e] U1 ≡ U2 &
- T1 = ⓑ{I} V1. U1.
+ T1 = ⓑ{a,I} V1. U1.
/2 width=3/ qed-.
fact lift_inv_flat2_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 →
| #i #d #e #_ #I #V2 #U2 #H destruct
| #i #d #e #_ #I #V2 #U2 #H destruct
| #p #d #e #I #V2 #U2 #H destruct
-| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V2 #U2 #H destruct
+| #a #J #W1 #W2 #T1 #T2 #d #e #_ #_ #I #V2 #U2 #H destruct
| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V2 #U2 #H destruct /2 width=5/
]
qed.
| elim (lift_inv_lref2 … H) -H * #_ #H destruct
| lapply (lift_inv_gref2 … H) -H #H destruct
]
-| * #I #W2 #U2 #IHW2 #_ #T #H
+| * [ #a ] #I #W2 #U2 #IHW2 #_ #T #H
[ elim (lift_inv_bind2 … H) -H #W1 #U1 #HW12 #_ #H destruct /2 width=2/
| elim (lift_inv_flat2 … H) -H #W1 #U1 #HW12 #_ #H destruct /2 width=2/
]
| elim (lift_inv_lref2 … H) -H * #_ #H destruct
| lapply (lift_inv_gref2 … H) -H #H destruct
]
-| * #I #W2 #U2 #_ #IHU2 #V #d #e #H
+| * [ #a ] #I #W2 #U2 #_ #IHU2 #V #d #e #H
[ elim (lift_inv_bind2 … H) -H #W1 #U1 #_ #HU12 #H destruct /2 width=4/
| elim (lift_inv_flat2 … H) -H #W1 #U1 #_ #HU12 #H destruct /2 width=4/
]
lemma lift_simple_dx: ∀d,e,T1,T2. ⇧[d, e] T1 ≡ T2 → 𝐒⦃T1⦄ → 𝐒⦃T2⦄.
#d #e #T1 #T2 #H elim H -d -e -T1 -T2 //
-#I #V1 #V2 #T1 #T2 #d #e #_ #_ #_ #_ #H
+#a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #_ #_ #H
elim (simple_inv_bind … H)
qed-.
lemma lift_simple_sn: ∀d,e,T1,T2. ⇧[d, e] T1 ≡ T2 → 𝐒⦃T2⦄ → 𝐒⦃T1⦄.
#d #e #T1 #T2 #H elim H -d -e -T1 -T2 //
-#I #V1 #V2 #T1 #T2 #d #e #_ #_ #_ #_ #H
+#a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #_ #_ #H
elim (simple_inv_bind … H)
qed-.
lemma lift_total: ∀T1,d,e. ∃T2. ⇧[d,e] T1 ≡ T2.
#T1 elim T1 -T1
[ * #i /2 width=2/ #d #e elim (lt_or_ge i d) /3 width=2/
-| * #I #V1 #T1 #IHV1 #IHT1 #d #e
+| * [ #a ] #I #V1 #T1 #IHV1 #IHT1 #d #e
elim (IHV1 d e) -IHV1 #V2 #HV12
[ elim (IHT1 (d+1) e) -IHT1 /3 width=2/
| elim (IHT1 d e) -IHT1 /3 width=2/
lapply (transitive_le … (i+e1) Hd21 ?) /2 width=1/ -Hd21 #Hd21
>(plus_minus_m_m e2 e1 ?) // /3 width=3/
| /3 width=3/
-| #I #V1 #V2 #T1 #T2 #d1 #e2 #_ #_ #IHV #IHT #d2 #e1 #Hd12 #Hd21 #He12
+| #a #I #V1 #V2 #T1 #T2 #d1 #e2 #_ #_ #IHV #IHT #d2 #e1 #Hd12 #Hd21 #He12
elim (IHV … Hd12 Hd21 He12) -IHV #V0 #HV0a #HV0b
elim (IHT (d2+1) … ? ? He12) /2 width=1/ /3 width=5/
| #I #V1 #V2 #T1 #T2 #d1 #e2 #_ #_ #IHV #IHT #d2 #e1 #Hd12 #Hd21 #He12
| lapply (false_lt_to_le … Hide) -Hide /4 width=2/
]
]
-| * #I #V2 #T2 #IHV2 #IHT2 #d #e
+| * [ #a ] #I #V2 #T2 #IHV2 #IHT2 #d #e
[ elim (IHV2 d e) -IHV2
[ * #V1 #HV12 elim (IHT2 (d+1) e) -IHT2
[ * #T1 #HT12 @or_introl /3 width=2/
lapply (lift_inv_lref2_ge … HX ?) -HX // /2 width=1/
| #p #d #e #X #HX
lapply (lift_inv_gref2 … HX) -HX //
-| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
+| #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
elim (lift_inv_bind2 … HX) -HX #V #T #HV1 #HT1 #HX destruct /3 width=1/
| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
elim (lift_inv_flat2 … HX) -HX #V #T #HV1 #HT1 #HX destruct /3 width=1/
]
| #p #d1 #e1 #d2 #e2 #T2 #Hk #Hd12
lapply (lift_inv_gref2 … Hk) -Hk #Hk destruct /3 width=3/
-| #I #W1 #W #U1 #U #d1 #e1 #_ #_ #IHW #IHU #d2 #e2 #T2 #H #Hd12
+| #a #I #W1 #W #U1 #U #d1 #e1 #_ #_ #IHW #IHU #d2 #e2 #T2 #H #Hd12
lapply (lift_inv_bind2 … H) -H * #W2 #U2 #HW2 #HU2 #H destruct
elim (IHW … HW2 ?) // -IHW -HW2 #W0 #HW2 #HW1
>plus_plus_comm_23 in HU2; #HU2 elim (IHU … HU2 ?) /2 width=1/ /3 width=5/
@lift_lref_ge_minus_eq [ >plus_minus_commutative // | /2 width=1/ ]
]
| #p #d1 #e1 #e #e2 #T2 #H >(lift_inv_gref2 … H) -H /2 width=3/
-| #I #V1 #V #T1 #T #d1 #e1 #_ #_ #IHV1 #IHT1 #e #e2 #X #H #He1 #He1e2
+| #a #I #V1 #V #T1 #T #d1 #e1 #_ #_ #IHV1 #IHT1 #e #e2 #X #H #He1 #He1e2
elim (lift_inv_bind2 … H) -H #V2 #T2 #HV2 #HT2 #H destruct
elim (IHV1 … HV2 ? ?) -V // >plus_plus_comm_23 in HT2; #HT2
elim (IHT1 … HT2 ? ?) -T // -He1 -He1e2 /3 width=5/
lapply (lift_inv_lref1_ge … HX ?) -HX //
| #p #d #e #X #HX
lapply (lift_inv_gref1 … HX) -HX //
-| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
+| #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
elim (lift_inv_bind1 … HX) -HX #V #T #HV1 #HT1 #HX destruct /3 width=1/
| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
elim (lift_inv_flat1 … HX) -HX #V #T #HV1 #HT1 #HX destruct /3 width=1/
]
| #p #d1 #e1 #d2 #e2 #T2 #HT2 #_ #_
>(lift_inv_gref1 … HT2) -HT2 //
-| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd12 #Hd21
+| #a #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd12 #Hd21
elim (lift_inv_bind1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct
lapply (IHV12 … HV20 ? ?) // -IHV12 -HV20 #HV10
lapply (IHT12 … HT20 ? ?) /2 width=1/
>plus_plus_comm_23 /4 width=3/
| #p #d1 #e1 #d2 #e2 #X #HX #_
>(lift_inv_gref1 … HX) -HX /2 width=3/
-| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd21
+| #a #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd21
elim (lift_inv_bind1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct
elim (IHV12 … HV20 ?) -IHV12 -HV20 //
elim (IHT12 … HT20 ?) -IHT12 -HT20 /2 width=1/ /3 width=5/
elim (lift_inv_lref1 … HX) -HX * #Hied #HX destruct /4 width=3/
| #p #d1 #e1 #d2 #e2 #X #HX #_
>(lift_inv_gref1 … HX) -HX /2 width=3/
-| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hded
+| #a #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hded
elim (lift_inv_bind1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct
elim (IHV12 … HV20 ?) -IHV12 -HV20 //
elim (IHT12 … HT20 ?) -IHT12 -HT20 /2 width=1/ #T
| tps_atom : ∀L,I,d,e. tps d e L (⓪{I}) (⓪{I})
| tps_subst: ∀L,K,V,W,i,d,e. d ≤ i → i < d + e →
⇩[0, i] L ≡ K. ⓓV → ⇧[0, i + 1] V ≡ W → tps d e L (#i) W
-| tps_bind : ∀L,I,V1,V2,T1,T2,d,e.
+| tps_bind : ∀L,a,I,V1,V2,T1,T2,d,e.
tps d e L V1 V2 → tps (d + 1) e (L. ⓑ{I} V2) T1 T2 →
- tps d e L (ⓑ{I} V1. T1) (ⓑ{I} V2. T2)
+ tps d e L (ⓑ{a,I} V1. T1) (ⓑ{a,I} V2. T2)
| tps_flat : ∀L,I,V1,V2,T1,T2,d,e.
tps d e L V1 V2 → tps d e L T1 T2 →
tps d e L (ⓕ{I} V1. T1) (ⓕ{I} V2. T2)
destruct
elim (lift_total V 0 (i+1)) #W #HVW
elim (lift_split … HVW i i ? ? ?) // /3 width=4/
-| * #I #W1 #U1 #IHW1 #IHU1 #L #d #HLK
+| * [ #a ] #I #W1 #U1 #IHW1 #IHU1 #L #d #HLK
elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2
- [ elim (IHU1 (L. ⓑ{I} W2) (d+1) ?) -IHU1 /2 width=1/ -HLK /3 width=8/
+ [ elim (IHU1 (L. ⓑ{I} W2) (d+1) ?) -IHU1 /2 width=1/ -HLK /3 width=9/
| elim (IHU1 … HLK) -IHU1 -HLK /3 width=8/
]
]
generalize in match Hide; -Hide (**) (* rewriting in the premises, rewrites in the goal too *)
>(plus_minus_m_m … Hjde) in ⊢ (% → ?); -Hjde /4 width=4/
]
-| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide
+| #L #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide
elim (IHV12 i ? ?) -IHV12 // #V #HV1 #HV2
elim (IHT12 (i + 1) ? ?) -IHT12 /2 width=1/
-Hdi -Hide >arith_c1x #T #HT1 #HT2
| -Hdi -Hdj
>(plus_minus_m_m (d+e) j) in Hide; // -Hjde /3 width=4/
]
-| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide
+| #L #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide
elim (IHV12 i ? ?) -IHV12 // #V #HV1 #HV2
elim (IHT12 (i + 1) ? ?) -IHT12 /2 width=1/
-Hdi -Hide >arith_c1x #T #HT1 #HT2
#L #T1 #T2 #d #e * -L -T1 -T2 -d -e
[ #L #I #d #e #J #H destruct /2 width=1/
| #L #K #V #T2 #i #d #e #Hdi #Hide #HLK #HVT2 #I #H destruct /3 width=8/
-| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct
+| #L #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct
| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct
]
qed.
qed-.
fact tps_inv_bind1_aux: ∀d,e,L,U1,U2. L ⊢ U1 ▶ [d, e] U2 →
- ∀I,V1,T1. U1 = ⓑ{I} V1. T1 →
+ ∀a,I,V1,T1. U1 = ⓑ{a,I} V1. T1 →
∃∃V2,T2. L ⊢ V1 ▶ [d, e] V2 &
L. ⓑ{I} V2 ⊢ T1 ▶ [d + 1, e] T2 &
- U2 = ⓑ{I} V2. T2.
+ U2 = ⓑ{a,I} V2. T2.
#d #e #L #U1 #U2 * -d -e -L -U1 -U2
-[ #L #k #d #e #I #V1 #T1 #H destruct
-| #L #K #V #W #i #d #e #_ #_ #_ #_ #I #V1 #T1 #H destruct
-| #L #J #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #I #V #T #H destruct /2 width=5/
-| #L #J #V1 #V2 #T1 #T2 #d #e #_ #_ #I #V #T #H destruct
+[ #L #k #d #e #a #I #V1 #T1 #H destruct
+| #L #K #V #W #i #d #e #_ #_ #_ #_ #a #I #V1 #T1 #H destruct
+| #L #b #J #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #a #I #V #T #H destruct /2 width=5/
+| #L #J #V1 #V2 #T1 #T2 #d #e #_ #_ #a #I #V #T #H destruct
]
qed.
-lemma tps_inv_bind1: ∀d,e,L,I,V1,T1,U2. L ⊢ ⓑ{I} V1. T1 ▶ [d, e] U2 →
+lemma tps_inv_bind1: ∀d,e,L,a,I,V1,T1,U2. L ⊢ ⓑ{a,I} V1. T1 ▶ [d, e] U2 →
∃∃V2,T2. L ⊢ V1 ▶ [d, e] V2 &
L. ⓑ{I} V2 ⊢ T1 ▶ [d + 1, e] T2 &
- U2 = ⓑ{I} V2. T2.
+ U2 = ⓑ{a,I} V2. T2.
/2 width=3/ qed-.
fact tps_inv_flat1_aux: ∀d,e,L,U1,U2. L ⊢ U1 ▶ [d, e] U2 →
#d #e #L #U1 #U2 * -d -e -L -U1 -U2
[ #L #k #d #e #I #V1 #T1 #H destruct
| #L #K #V #W #i #d #e #_ #_ #_ #_ #I #V1 #T1 #H destruct
-| #L #J #V1 #V2 #T1 #T2 #d #e #_ #_ #I #V #T #H destruct
+| #L #a #J #V1 #V2 #T1 #T2 #d #e #_ #_ #I #V #T #H destruct
| #L #J #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #I #V #T #H destruct /2 width=5/
]
qed.
elim (ldrop_trans_le … HLK … HKV ?) -K /2 width=2/ #X #HLK #H
elim (ldrop_inv_skip2 … H ?) -H /2 width=1/ -Hid #K #Y #_ #HVY
>(lift_mono … HVY … HVW) -Y -HVW #H destruct /2 width=4/
-| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdetd
+| #K #a #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdetd
elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1
elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct
@tps_bind [ /2 width=6/ | @IHT12 /2 width=6/ ] (**) (* /3 width=6/ is too slow, arith3 needed to avoid crash *)
lapply (lift_trans_be … HVW … HWU2 ? ?) -W // /2 width=1/ >plus_plus_comm_23 #HVU2
lapply (ldrop_trans_ge_comm … HLK … HKV ?) -K // -Hid /3 width=4/
]
-| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdtd #Hddet
+| #K #a #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdtd #Hddet
elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1
elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct
@tps_bind [ /2 width=6/ | @IHT12 [3,4: // | skip |5,6: /2 width=1/ | /2 width=1/ ]
lapply (lift_inv_lref1_ge … H … Hid) -H #H destruct
lapply (lift_trans_be … HVW … HWU2 ? ?) -W // /2 width=1/ >plus_plus_comm_23 #HVU2
lapply (ldrop_trans_ge_comm … HLK … HKV ?) -K // -Hid /3 width=4/
-| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hddt
+| #K #a #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hddt
elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1
elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct
@tps_bind [ /2 width=5/ | /3 width=5/ ] (**) (* explicit constructor *)
lapply (lift_inv_lref2_lt … H … Hid) -H #H destruct
elim (ldrop_conf_lt … HLK … HLKV ?) -L // #L #U #HKL #_ #HUV
elim (lift_trans_le … HUV … HVW ?) -V // >minus_plus <plus_minus_m_m // -Hid /3 width=4/
-| #L #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdetd
+| #L #a #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdetd
elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
elim (IHV12 … HLK … HWV1 ?) -V1 // #W2 #HW12 #HWV2
elim (IHU12 … HTU1 ?) -IHU12 -HTU1 [3: /2 width=1/ |4: @ldrop_skip // |2: skip ] -HLK -Hdetd (**) (* /3 width=5/ is too slow *)
#V1 #HV1 >plus_minus // <minus_minus // /2 width=1/ <minus_n_n <plus_n_O #H
@ex2_1_intro [3: @H | skip | @tps_subst [3,5,6: // |1,2: skip | >commutative_plus >plus_minus // /2 width=1/ ] ] (**) (* explicit constructor, uses monotonic_lt_minus_l *)
]
-| #L #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdtd #Hdedet
+| #L #a #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdtd #Hdedet
elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
elim (IHV12 … HLK … HWV1 ? ?) -V1 // #W2 #HW12 #HWV2
elim (IHU12 … HTU1 ? ?) -U1 [5: @ldrop_skip // |2: skip |3: >plus_plus_comm_23 >(plus_plus_comm_23 dt) /2 width=1/ |4: /2 width=1/ ] (**) (* 29s without the rewrites *)
elim (lift_split … HVW d (i - e + 1) ? ? ?) -HVW [4: // |3: /2 width=1/ |2: /3 width=1/ ] -Hdei -Hdie
#V0 #HV10 >plus_minus // <minus_minus // /2 width=1/ <minus_n_n <plus_n_O #H
@ex2_1_intro [3: @H | skip | @tps_subst [5,6: // |1,2: skip | /2 width=1/ | >plus_minus // /2 width=1/ ] ] (**) (* explicit constructor, uses monotonic_lt_minus_l *)
-| #L #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdetd
+| #L #a #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdetd
elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
elim (le_inv_plus_l … Hdetd) #_ #Hedt
elim (IHV12 … HLK … HWV1 ?) -V1 // #W2 #HW12 #HWV2
| lapply (lt_to_le_to_lt … Hide … H) -Hide -H #H
elim (lt_refl_false … H)
]
-| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
+| #L #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
elim (lift_inv_bind2 … HX) -HX #V #T #HV1 #HT1 #H destruct
>IHV12 // >IHT12 //
| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
lapply (ldrop_mono … HLK1 … HLK2) -HLK1 -HLK2 #H destruct
>(lift_mono … HVT1 … HVT2) -HVT1 -HVT2 /2 width=3/
]
-| #L #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #X #d2 #e2 #HX
+| #L #a #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #X #d2 #e2 #HX
elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct
lapply (tps_lsubs_trans … HT02 (L. ⓑ{I} V1) ?) -HT02 /2 width=1/ #HT02
elim (IHV01 … HV02) -V0 #V #HV1 #HV2
elim (lt_refl_false … H)
]
]
-| #L1 #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #L2 #X #d2 #e2 #HX #H
+| #L1 #a #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #L2 #X #d2 #e2 #HX #H
elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct
elim (IHV01 … HV02 H) -V0 #V #HV1 #HV2
elim (IHT01 … HT02 ?) -T0
| #L #K #V #V2 #i #d #e #Hdi #Hide #HLK #HVW #T2 #HVT2 #He
lapply (tps_weak … HVT2 0 (i +1) ? ?) -HVT2 /2 width=1/ #HVT2
<(tps_inv_lift1_eq … HVT2 … HVW) -HVT2 /2 width=4/
-| #L #I #V1 #V0 #T1 #T0 #d #e #_ #_ #IHV10 #IHT10 #X #H #He
+| #L #a #I #V1 #V0 #T1 #T0 #d #e #_ #_ #IHV10 #IHT10 #X #H #He
elim (tps_inv_bind1 … H) -H #V2 #T2 #HV02 #HT02 #H destruct
lapply (tps_lsubs_trans … HT02 (L. ⓑ{I} V0) ?) -HT02 /2 width=1/ #HT02
lapply (IHT10 … HT02 He) -T0 #HT12
lapply (transitive_le … Hde2d1 Hdi1) -Hde2d1 #Hde2i1
lapply (tps_weak … HWT2 0 (i1 + 1) ? ?) -HWT2 normalize /2 width=1/ -Hde2i1 #HWT2
<(tps_inv_lift1_eq … HWT2 … HVW) -HWT2 /4 width=4/
-| #L #I #V1 #V0 #T1 #T0 #d1 #e1 #_ #_ #IHV10 #IHT10 #X #d2 #e2 #HX #de2d1
+| #L #a #I #V1 #V0 #T1 #T0 #d1 #e1 #_ #_ #IHV10 #IHT10 #X #d2 #e2 #HX #de2d1
elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct
lapply (tps_lsubs_trans … HT02 (L. ⓑ{I} V0) ?) -HT02 /2 width=1/ #HT02
elim (IHV10 … HV02 ?) -IHV10 -HV02 // #V
lemma delift_gref: ∀L,d,e,p. L ⊢ ▼*[d, e] §p ≡ §p.
/2 width=3/ qed.
-lemma delift_bind: ∀I,L,V1,V2,T1,T2,d,e.
+lemma delift_bind: ∀a,I,L,V1,V2,T1,T2,d,e.
L ⊢ ▼*[d, e] V1 ≡ V2 → L. ⓑ{I} V2 ⊢ ▼*[d+1, e] T1 ≡ T2 →
- L ⊢ ▼*[d, e] ⓑ{I} V1. T1 ≡ ⓑ{I} V2. T2.
-#I #L #V1 #V2 #T1 #T2 #d #e * #V #HV1 #HV2 * #T #HT1 #HT2
+ L ⊢ ▼*[d, e] ⓑ{a,I} V1. T1 ≡ ⓑ{a,I} V2. T2.
+#a #I #L #V1 #V2 #T1 #T2 #d #e * #V #HV1 #HV2 * #T #HT1 #HT2
lapply (tpss_lsubs_trans … HT1 (L. ⓑ{I} V) ?) -HT1 /2 width=1/ /3 width=5/
qed.
>(lift_inv_gref2 … HU2) -HU2 //
qed-.
-lemma delift_inv_bind1: ∀I,L,V1,T1,U2,d,e. L ⊢ ▼*[d, e] ⓑ{I} V1. T1 ≡ U2 →
+lemma delift_inv_bind1: ∀a,I,L,V1,T1,U2,d,e. L ⊢ ▼*[d, e] ⓑ{a,I} V1. T1 ≡ U2 →
∃∃V2,T2. L ⊢ ▼*[d, e] V1 ≡ V2 &
L. ⓑ{I} V2 ⊢ ▼*[d+1, e] T1 ≡ T2 &
- U2 = ⓑ{I} V2. T2.
-#I #L #V1 #T1 #U2 #d #e * #U #HU #HU2
+ U2 = ⓑ{a,I} V2. T2.
+#a #I #L #V1 #T1 #U2 #d #e * #U #HU #HU2
elim (tpss_inv_bind1 … HU) -HU #V #T #HV1 #HT1 #X destruct
elim (lift_inv_bind2 … HU2) -HU2 #V2 #T2 #HV2 #HT2
lapply (tpss_lsubs_trans … HT1 (L. ⓑ{I} V2) ?) -HT1 /2 width=1/ /3 width=5/
⇧[0, d] V2 ≡ W2 → delifta d e L (#i) W2
| delifta_lref_ge: ∀L,d,e,i. d + e ≤ i → delifta d e L (#i) (#(i - e))
| delifta_gref : ∀L,d,e,p. delifta d e L (§p) (§p)
-| delifta_bind : ∀L,I,V1,V2,T1,T2,d,e.
+| delifta_bind : ∀L,a,I,V1,V2,T1,T2,d,e.
delifta d e L V1 V2 → delifta (d + 1) e (L. ⓑ{I} V2) T1 T2 →
- delifta d e L (ⓑ{I} V1. T1) (ⓑ{I} V2. T2)
+ delifta d e L (ⓑ{a,I} V1. T1) (ⓑ{a,I} V2. T2)
| delifta_flat : ∀L,I,V1,V2,T1,T2,d,e.
delifta d e L V1 V2 → delifta d e L T1 T2 →
delifta d e L (ⓕ{I} V1. T1) (ⓕ{I} V2. T2)
lapply (IH … HV12) // -H /2 width=6/
| >(delift_inv_gref1 … H) -H //
]
-| * #I #V1 #T1 #_ #_ #IH #X #d #e #H
+| * [ #a ] #I #V1 #T1 #_ #_ #IH #X #d #e #H
[ elim (delift_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
lapply (delift_lsubs_trans … HT12 (L.ⓑ{I}V1) ?) -HT12 /2 width=1/ #HT12
lapply (IH … HV12) -HV12 // #HV12
) →
(∀L,d,e,i. d + e ≤ i → R d e L (#i) (#(i - e))) →
(∀L,d,e,p. R d e L (§p) (§p)) →
- (∀L,I,V1,V2,T1,T2,d,e. L ⊢ ▼*[d, e] V1 ≡ V2 →
+ (∀L,a,I,V1,V2,T1,T2,d,e. L ⊢ ▼*[d, e] V1 ≡ V2 →
L.ⓑ{I}V2 ⊢ ▼*[d + 1, e] T1 ≡ T2 → R d e L V1 V2 →
- R (d+1) e (L.ⓑ{I}V2) T1 T2 → R d e L (ⓑ{I}V1.T1) (ⓑ{I}V2.T2)
+ R (d+1) e (L.ⓑ{I}V2) T1 T2 → R d e L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2)
) →
(∀L,I,V1,V2,T1,T2,d,e. L ⊢ ▼*[d, e] V1 ≡ V2 →
L⊢ ▼*[d, e] T1 ≡ T2 → R d e L V1 V2 →
elim (IH … HKL … HK ?) -IH -HKL -HK
[3: // |2: skip |4: >H -H /2 width=1/ ] -Hde -H #V2 #V12 (**) (* H erased two times *)
elim (lift_total V2 0 d) /3 width=7/
-| #I #V1 #T1 #d #e #Hde #HL #H destruct
+| #a #I #V1 #T1 #d #e #Hde #HL #H destruct
elim (IH … V1 … Hde HL ?) [2,4: // |3: skip ] #V2 #HV12
elim (IH (L.ⓑ{I}V1) T1 ? ? (d+1) e ? ? ?) -IH [3,6: // |2: skip |4,5: /2 width=1/ ] -Hde -HL #T2 #HT12
lapply (delift_lsubs_trans … HT12 (L.ⓑ{I}V2) ?) -HT12 /2 width=1/ /3 width=4/
qed-.
(* Basic_1: was: lift1_bind *)
-lemma lifts_inv_bind1: ∀I,T2,des,V1,U1. ⇧*[des] ⓑ{I} V1. U1 ≡ T2 →
+lemma lifts_inv_bind1: ∀a,I,T2,des,V1,U1. ⇧*[des] ⓑ{a,I} V1. U1 ≡ T2 →
∃∃V2,U2. ⇧*[des] V1 ≡ V2 & ⇧*[des + 1] U1 ≡ U2 &
- T2 = ⓑ{I} V2. U2.
-#I #T2 #des elim des -des
+ T2 = ⓑ{a,I} V2. U2.
+#a #I #T2 #des elim des -des
[ #V1 #U1 #H
<(lifts_inv_nil … H) -H /2 width=5/
| #d #e #des #IHdes #V1 #U1 #H
(* Basic properties *********************************************************)
-lemma lifts_bind: ∀I,T2,V1,V2,des. ⇧*[des] V1 ≡ V2 →
+lemma lifts_bind: ∀a,I,T2,V1,V2,des. ⇧*[des] V1 ≡ V2 →
∀T1. ⇧*[des + 1] T1 ≡ T2 →
- ⇧*[des] ⓑ{I} V1. T1 ≡ ⓑ{I} V2. T2.
-#I #T2 #V1 #V2 #des #H elim H -V1 -V2 -des
+ ⇧*[des] ⓑ{a,I} V1. T1 ≡ ⓑ{a,I} V2. T2.
+#a #I #T2 #V1 #V2 #des #H elim H -V1 -V2 -des
[ #V #T1 #H >(lifts_inv_nil … H) -H //
| #V1 #V #V2 #des #d #e #HV1 #_ #IHV #T1 #H
elim (lifts_inv_cons … H) -H /3 width=3/
lapply (tpss_trans_eq … HV01 HV12) -V1 #HV02
lapply (IH … HV02 … HK01 ? ?) -IH -HV02 -HK01
[1,3: // |2,4: skip | normalize /2 width=1/ | /2 width=6/ ]
-| #L #I #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #_ #_ #L0 #HL0 #H1 #H2 destruct
+| #L #a #I #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #_ #_ #L0 #HL0 #H1 #H2 destruct
lapply (tpss_lsubs_trans … HT12 (L. ⓑ{I} V1) ?) -HT12 /2 width=1/ #HT12
lapply (IH … HV12 … HL0 ? ?) -HV12 [1,3: // |2,4: skip |5: /2 width=2/ ] #HV12
lapply (IH … HT12 (L0. ⓑ{I} V1) ? ? ?) -IH -HT12 [1,3,5: /2 width=2/ |2,4: skip | normalize // ] -HL0 #HT12
| #L0 #K0 #V0 #W0 #i2 #d2 #e2 #Hdi2 #Hide2 #HLK0 #HVW0 #L1 #d1 #e1 #HL01 #Hde1d2
lapply (transitive_le … Hde1d2 Hdi2) -Hde1d2 #Hde1i2
lapply (ltpss_ldrop_conf_ge … HL01 … HLK0 ?) -L0 // /2 width=4/
-| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL01 #Hde1d2
+| #L0 #a #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL01 #Hde1d2
@tps_bind [ /2 width=4/ | @IHTU2 -IHTU2 -IHVW2 [3: /2 by ltpss_tpss1/ |1,2: skip | /2 width=1/ ] ] (**) (* explicit constructor *)
| /3 width=4/
]
| #L0 #K0 #V0 #W0 #i2 #d2 #e2 #Hdi2 #Hide2 #HLK0 #HVW0 #L1 #d1 #e1 #HL10 #Hde1d2
lapply (transitive_le … Hde1d2 Hdi2) -Hde1d2 #Hde1i2
lapply (ltpss_ldrop_trans_ge … HL10 … HLK0 ?) -L0 // /2 width=4/
-| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL10 #Hde1d2
+| #L0 #a #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL10 #Hde1d2
@tps_bind [ /2 width=4/ | @IHTU2 -IHTU2 -IHVW2 [3: /2 by ltpss_tpss1/ |1,2: skip | /2 width=1/ ] ] (**) (* explicit constructor *)
| /3 width=4/
]
| lapply (ltpss_ldrop_conf_ge … HL01 … HLK0 ?) -L0 // /3 width=4/
]
]
-| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL01
+| #L0 #a #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL01
elim (IHVW2 … HL01) -IHVW2 #V #HV2 #HVW2
elim (IHTU2 (L1. ⓑ{I} V) (d1 + 1) e1 ?) -IHTU2 /2 width=1/ -HL01 /3 width=5/
| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL01
| lapply (ltpss_ldrop_trans_ge … HL10 … HLK0 ?) -HL10 -HLK0 // /3 width=4/
]
]
-| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL10
+| #L0 #a #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL10
elim (IHVW2 … HL10) -IHVW2 #V #HV2 #HVW2
elim (IHTU2 (L1. ⓑ{I} V) (d1 + 1) e1 ?) -IHTU2 /2 width=1/ -HL10 /3 width=5/
| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL10
/2 width=1/ qed.
lemma tpss_bind: ∀L,V1,V2,d,e. L ⊢ V1 ▶* [d, e] V2 →
- ∀I,T1,T2. L. ⓑ{I} V2 ⊢ T1 ▶* [d + 1, e] T2 →
- L ⊢ ⓑ{I} V1. T1 ▶* [d, e] ⓑ{I} V2. T2.
+ ∀a,I,T1,T2. L. ⓑ{I} V2 ⊢ T1 ▶* [d + 1, e] T2 →
+ L ⊢ ⓑ{a,I} V1. T1 ▶* [d, e] ⓑ{a,I} V2. T2.
#L #V1 #V2 #d #e #HV12 elim HV12 -V2
-[ #V2 #HV12 #I #T1 #T2 #HT12 elim HT12 -T2
+[ #V2 #HV12 #a #I #T1 #T2 #HT12 elim HT12 -T2
[ /3 width=5/
| #T #T2 #_ #HT2 #IHT @step /2 width=5/ (**) (* /3 width=5/ is too slow *)
]
-| #V #V2 #_ #HV12 #IHV #I #T1 #T2 #HT12
+| #V #V2 #_ #HV12 #IHV #a #I #T1 #T2 #HT12
lapply (tpss_lsubs_trans … HT12 (L. ⓑ{I} V) ?) -HT12 /2 width=1/ #HT12
- lapply (IHV … HT12) -IHV -HT12 #HT12 @step /2 width=5/ (**) (* /3 width=5/ is too slow *)
+ lapply (IHV a … HT12) -IHV -HT12 #HT12 @step /2 width=5/ (**) (* /3 width=5/ is too slow *)
]
qed.
]
qed-.
-lemma tpss_inv_bind1: ∀d,e,L,I,V1,T1,U2. L ⊢ ⓑ{I} V1. T1 ▶* [d, e] U2 →
+lemma tpss_inv_bind1: ∀d,e,L,a,I,V1,T1,U2. L ⊢ ⓑ{a,I} V1. T1 ▶* [d, e] U2 →
∃∃V2,T2. L ⊢ V1 ▶* [d, e] V2 &
L. ⓑ{I} V2 ⊢ T1 ▶* [d + 1, e] T2 &
- U2 = ⓑ{I} V2. T2.
-#d #e #L #I #V1 #T1 #U2 #H @(tpss_ind … H) -U2
+ U2 = ⓑ{a,I} V2. T2.
+#d #e #L #a #I #V1 #T1 #U2 #H @(tpss_ind … H) -U2
[ /2 width=5/
| #U #U2 #_ #HU2 * #V #T #HV1 #HT1 #H destruct
elim (tps_inv_bind1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H
| tpssa_subst: ∀L,K,V1,V2,W2,i,d,e. d ≤ i → i < d + e →
⇩[0, i] L ≡ K. ⓓV1 → tpssa 0 (d + e - i - 1) K V1 V2 →
⇧[0, i + 1] V2 ≡ W2 → tpssa d e L (#i) W2
-| tpssa_bind : ∀L,I,V1,V2,T1,T2,d,e.
+| tpssa_bind : ∀L,a,I,V1,V2,T1,T2,d,e.
tpssa d e L V1 V2 → tpssa (d + 1) e (L. ⓑ{I} V2) T1 T2 →
- tpssa d e L (ⓑ{I} V1. T1) (ⓑ{I} V2. T2)
+ tpssa d e L (ⓑ{a,I} V1. T1) (ⓑ{a,I} V2. T2)
| tpssa_flat : ∀L,I,V1,V2,T1,T2,d,e.
tpssa d e L V1 V2 → tpssa d e L T1 T2 →
tpssa d e L (ⓕ{I} V1. T1) (ⓕ{I} V2. T2)
lapply (ldrop_fwd_ldrop2 … HLK) #H0LK
lapply (tps_weak … H 0 (d+e) ? ?) -H // #H
elim (tps_inv_lift1_be … H … H0LK … HVW2 ? ?) -H -H0LK -HVW2 // /3 width=6/
-| #L #I #V1 #V #T1 #T #d #e #_ #_ #IHV1 #IHT1 #X #H
+| #L #a #I #V1 #V #T1 #T #d #e #_ #_ #IHV1 #IHT1 #X #H
elim (tps_inv_bind1 … H) -H #V2 #T2 #HV2 #HT2 #H destruct
lapply (tps_lsubs_trans … HT2 (L.ⓑ{I}V) ?) -HT2 /2 width=1/ #HT2
lapply (IHV1 … HV2) -IHV1 -HV2 #HV12
⇩[O, i] L ≡ K.ⓓV1 → K ⊢ V1 ▶* [O, d + e - i - 1] V2 →
⇧[O, i + 1] V2 ≡ W2 → R O (d+e-i-1) K V1 V2 → R d e L #i W2
) →
- (∀L,I,V1,V2,T1,T2,d,e. L ⊢ V1 ▶* [d, e] V2 →
+ (∀L,a,I,V1,V2,T1,T2,d,e. L ⊢ V1 ▶* [d, e] V2 →
L.ⓑ{I}V2 ⊢ T1 ▶* [d + 1, e] T2 → R d e L V1 V2 →
- R (d+1) e (L.ⓑ{I}V2) T1 T2 → R d e L (ⓑ{I}V1.T1) (ⓑ{I}V2.T2)
+ R (d+1) e (L.ⓑ{I}V2) T1 T2 → R d e L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2)
) →
(∀L,I,V1,V2,T1,T2,d,e. L ⊢ V1 ▶* [d, e] V2 →
L ⊢ T1 ▶* [d, e] T2 → R d e L V1 V2 →
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/static/ssta.ma".
-
-(* STRATIFIED UNWIND ON TERMS ***********************************************)
-
-inductive sstas (h:sh) (g:sd h) (L:lenv): relation term ≝
-| sstas_refl: ∀T,U. ⦃h, L⦄ ⊢ T •[g, 0] U → sstas h g L T T
-| sstas_step: ∀T,U1,U2,l. ⦃h, L⦄ ⊢ T •[g, l+1] U1 → sstas h g L U1 U2 →
- sstas h g L T U2.
-
-interpretation "stratified unwind (term)"
- 'StaticTypeStar h g L T U = (sstas h g L T U).
-
-(* Basic eliminators ********************************************************)
-
-fact sstas_ind_alt_aux: ∀h,g,L,U2. ∀R:predicate term.
- (∀T. ⦃h, L⦄ ⊢ U2 •[g , 0] T → R U2) →
- (∀T,U1,l. ⦃h, L⦄ ⊢ T •[g, l + 1] U1 →
- ⦃h, L⦄ ⊢ U1 •* [g] U2 → R U1 → R T
- ) →
- ∀T,U. ⦃h, L⦄ ⊢ T •*[g] U → U = U2 → R T.
-#h #g #L #U2 #R #H1 #H2 #T #U #H elim H -H -T -U /2 width=2/ /3 width=5/
-qed-.
-
-lemma sstas_ind_alt: ∀h,g,L,U2. ∀R:predicate term.
- (∀T. ⦃h, L⦄ ⊢ U2 •[g , 0] T → R U2) →
- (∀T,U1,l. ⦃h, L⦄ ⊢ T •[g, l + 1] U1 →
- ⦃h, L⦄ ⊢ U1 •* [g] U2 → R U1 → R T
- ) →
- ∀T. ⦃h, L⦄ ⊢ T •*[g] U2 → R T.
-/3 width=9 by sstas_ind_alt_aux/ qed-.
-
-(* Basic inversion lemmas ***************************************************)
-
-fact sstas_inv_sort1_aux: ∀h,g,L,T,U. ⦃h, L⦄ ⊢ T •*[g] U → ∀k. T = ⋆k →
- ∀l. deg h g k l → U = ⋆((next h)^l k).
-#h #g #L #T #U #H @(sstas_ind_alt … H) -T
-[ #U0 #HU0 #k #H #l #Hkl destruct
- elim (ssta_inv_sort1 … HU0) -L #HkO #_ -U0
- >(deg_mono … Hkl HkO) -g -l //
-| #T0 #U0 #l0 #HTU0 #_ #IHU0 #k #H #l #Hkl destruct
- elim (ssta_inv_sort1 … HTU0) -L #HkS #H destruct
- lapply (deg_mono … Hkl HkS) -Hkl #H destruct
- >(IHU0 (next h k) ? l0) -IHU0 // /2 width=1/ >iter_SO >iter_n_Sm //
-]
-qed.
-
-lemma sstas_inv_sort1: ∀h,g,L,U,k. ⦃h, L⦄ ⊢ ⋆k •*[g] U → ∀l. deg h g k l →
- U = ⋆((next h)^l k).
-/2 width=6/ qed-.
-
-fact sstas_inv_bind1_aux: ∀h,g,L,T,U. ⦃h, L⦄ ⊢ T •*[g] U →
- ∀J,X,Y. T = ⓑ{J}Y.X →
- ∃∃Z. ⦃h, L.ⓑ{J}Y⦄ ⊢ X •*[g] Z & U = ⓑ{J}Y.Z.
-#h #g #L #T #U #H @(sstas_ind_alt … H) -T
-[ #U0 #HU0 #J #X #Y #H destruct
- elim (ssta_inv_bind1 … HU0) -HU0 #X0 #HX0 #H destruct /3 width=3/
-| #T0 #U0 #l #HTU0 #_ #IHU0 #J #X #Y #H destruct
- elim (ssta_inv_bind1 … HTU0) -HTU0 #X0 #HX0 #H destruct
- elim (IHU0 J X0 Y ?) -IHU0 // #X1 #HX01 #H destruct /3 width=4/
-]
-qed.
-
-lemma sstas_inv_bind1: ∀h,g,J,L,Y,X,U. ⦃h, L⦄ ⊢ ⓑ{J}Y.X •*[g] U →
- ∃∃Z. ⦃h, L.ⓑ{J}Y⦄ ⊢ X •*[g] Z & U = ⓑ{J}Y.Z.
-/2 width=3/ qed-.
-
-fact sstas_inv_appl1_aux: ∀h,g,L,T,U. ⦃h, L⦄ ⊢ T •*[g] U → ∀X,Y. T = ⓐY.X →
- ∃∃Z. ⦃h, L⦄ ⊢ X •*[g] Z & U = ⓐY.Z.
-#h #g #L #T #U #H @(sstas_ind_alt … H) -T
-[ #U0 #HU0 #X #Y #H destruct
- elim (ssta_inv_appl1 … HU0) -HU0 #X0 #HX0 #H destruct /3 width=3/
-| #T0 #U0 #l #HTU0 #_ #IHU0 #X #Y #H destruct
- elim (ssta_inv_appl1 … HTU0) -HTU0 #X0 #HX0 #H destruct
- elim (IHU0 X0 Y ?) -IHU0 // #X1 #HX01 #H destruct /3 width=4/
-]
-qed.
-
-lemma sstas_inv_appl1: ∀h,g,L,Y,X,U. ⦃h, L⦄ ⊢ ⓐY.X •*[g] U →
- ∃∃Z. ⦃h, L⦄ ⊢ X •*[g] Z & U = ⓐY.Z.
-/2 width=3/ qed-.
-
-(* Basic forward lemmas *****************************************************)
-
-lemma sstas_fwd_correct: ∀h,g,L,T,U. ⦃h, L⦄ ⊢ T •*[g] U →
- ∃∃W. ⦃h, L⦄ ⊢ U •[g, 0] W & ⦃h, L⦄ ⊢ U •*[g] U.
-#h #g #L #T #U #H @(sstas_ind_alt … H) -T /2 width=1/ /3 width=2/
-qed-.
-
-(* Basic_1: removed theorems 7:
- sty1_bind sty1_abbr sty1_appl sty1_cast2
- sty1_lift sty1_correct sty1_trans
-*)
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/static/ssta_lift.ma".
-include "basic_2/unwind/sstas.ma".
-
-(* STRATIFIED UNWIND ON TERMS ***********************************************)
-
-(* Advanced properties ******************************************************)
-
-lemma sstas_total_S: ∀h,g,L,l,T,U. ⦃h, L⦄ ⊢ T•[g, l + 1]U →
- ∃∃W. ⦃h, L⦄ ⊢ T •*[g] W & ⦃h, L⦄ ⊢ U •*[g] W.
-#h #g #L #l @(nat_ind_plus … l) -l
-[ #T #U #HTU
- elim (ssta_fwd_correct … HTU) /4 width=4/
-| #l #IHl #T #U #HTU
- elim (ssta_fwd_correct … HTU) <minus_plus_m_m #V #HUV
- elim (IHl … HUV) -IHl -HUV /3 width=4/
-]
-qed-.
-
-(* Properties on relocation *************************************************)
-
-lemma sstas_lift: ∀h,g,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 •*[g] U1 →
- ∀L2,d,e. ⇩[d, e] L2 ≡ L1 → ∀T2. ⇧[d, e] T1 ≡ T2 →
- ∀U2. ⇧[d, e] U1 ≡ U2 → ⦃h, L2⦄ ⊢ T2 •*[g] U2.
-#h #g #L1 #T1 #U1 #H @(sstas_ind_alt … H) -T1
-[ #T1 #HUT1 #L2 #d #e #HL21 #X #HX #U2 #HU12
- >(lift_mono … HX … HU12) -X
- elim (lift_total T1 d e) /3 width=10/
-| #T0 #U0 #l0 #HTU0 #_ #IHU01 #L2 #d #e #HL21 #T2 #HT02 #U2 #HU12
- elim (lift_total U0 d e) /3 width=10/
-]
-qed.
-
-lemma sstas_inv_lift1: ∀h,g,L2,T2,U2. ⦃h, L2⦄ ⊢ T2 •*[g] U2 →
- ∀L1,d,e. ⇩[d, e] L2 ≡ L1 → ∀T1. ⇧[d, e] T1 ≡ T2 →
- ∃∃U1. ⦃h, L1⦄ ⊢ T1 •*[g] U1 & ⇧[d, e] U1 ≡ U2.
-#h #g #L2 #T2 #U2 #H @(sstas_ind_alt … H) -T2
-[ #T2 #HUT2 #L1 #d #e #HL21 #U1 #HU12
- elim (ssta_inv_lift1 … HUT2 … HL21 … HU12) -HUT2 -HL21 /3 width=3/
-| #T0 #U0 #l0 #HTU0 #_ #IHU01 #L1 #d #e #HL21 #U1 #HU12
- elim (ssta_inv_lift1 … HTU0 … HL21 … HU12) -HTU0 -HU12 #U #HU1 #HU0
- elim (IHU01 … HL21 … HU0) -IHU01 -HL21 -U0 /3 width=4/
-]
-qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/static/ssta_ltpss.ma".
-include "basic_2/unwind/sstas.ma".
-
-(* STRATIFIED UNWIND ON TERMS ***********************************************)
-
-(* Properties about parallel unfold *****************************************)
-
-lemma sstas_ltpss_tpss_conf: ∀h,g,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 •*[g] U1 →
- ∀L2,d,e. L1 ▶* [d, e] L2 →
- ∀T2. L2 ⊢ T1 ▶* [d, e] T2 →
- ∃∃U2. ⦃h, L2⦄ ⊢ T2 •*[g] U2 &
- L2 ⊢ U1 ▶* [d, e] U2.
-#h #g #L1 #T1 #U1 #H @(sstas_ind_alt … H) -T1
-[ #T1 #HUT1 #L2 #d #e #HL12 #U2 #HU12
- elim (ssta_ltpss_tpss_conf … HUT1 … HL12 … HU12) -HUT1 -HL12 /3 width=3/
-| #T0 #U0 #l0 #HTU0 #_ #IHU01 #L2 #d #e #HL12 #T #HT0
- elim (ssta_ltpss_tpss_conf … HTU0 … HL12 … HT0) -HTU0 -HT0 #U #HTU #HU0
- elim (IHU01 … HL12 … HU0) -IHU01 -HL12 -U0 /3 width=4/
-]
-qed.
-
-lemma sstas_ltpss_tps_conf: ∀h,g,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 •*[g] U1 →
- ∀L2,d,e. L1 ▶* [d, e] L2 →
- ∀T2. L2 ⊢ T1 ▶ [d, e] T2 →
- ∃∃U2. ⦃h, L2⦄ ⊢ T2 •*[g] U2 & L2 ⊢ U1 ▶* [d, e] U2.
-/3 width=5/ qed.
-
-lemma sstas_ltpss_conf: ∀h,g,L1,T,U1. ⦃h, L1⦄ ⊢ T •*[g] U1 →
- ∀L2,d,e. L1 ▶* [d, e] L2 →
- ∃∃U2. ⦃h, L2⦄ ⊢ T •*[g] U2 & L2 ⊢ U1 ▶* [d, e] U2.
-/2 width=5/ qed.
-
-lemma sstas_tpss_conf: ∀h,g,L,T1,U1. ⦃h, L⦄ ⊢ T1 •*[g] U1 →
- ∀T2,d,e. L ⊢ T1 ▶* [d, e] T2 →
- ∃∃U2. ⦃h, L⦄ ⊢ T2 •*[g] U2 & L ⊢ U1 ▶* [d, e] U2.
-/2 width=5/ qed.
-
-lemma sstas_tps_conf: ∀h,g,L,T1,U1. ⦃h, L⦄ ⊢ T1 •*[g] U1 →
- ∀T2,d,e. L ⊢ T1 ▶ [d, e] T2 →
- ∃∃U2. ⦃h, L⦄ ⊢ T2 •*[g] U2 & L ⊢ U1 ▶* [d, e] U2.
-/2 width=5/ qed.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/unfold/delift_lift.ma".
-include "basic_2/static/ssta_ssta.ma".
-include "basic_2/unwind/sstas_lift.ma".
-
-(* STRATIFIED UNWIND ON TERMS ***********************************************)
-
-(* Advanced inversion lemmas ************************************************)
-
-lemma sstas_inv_O: ∀h,g,L,T,U. ⦃h, L⦄ ⊢ T •*[g] U →
- ∀T0. ⦃h, L⦄ ⊢ T •[g , 0] T0 → U = T.
-#h #g #L #T #U #H @(sstas_ind_alt … H) -T //
-#T0 #U0 #l0 #HTU0 #_ #_ #T1 #HT01
-elim (ssta_mono … HTU0 … HT01) <plus_n_Sm #H destruct
-qed-.
-
-lemma sstas_inv_S: ∀h,g,L,T,U. ⦃h, L⦄ ⊢ T •*[g] U →
- ∀T0,l. ⦃h, L⦄ ⊢ T •[g , l+1] T0 → ⦃h, L⦄ ⊢ T0 •*[g] U.
-#h #g #L #T #U #H @(sstas_ind_alt … H) -T
-[ #U0 #HU0 #T #l #HUT
- elim (ssta_mono … HUT … HU0) <plus_n_Sm #H destruct
-| #T0 #U0 #l0 #HTU0 #HU0 #_ #T #l #HT0
- elim (ssta_mono … HT0 … HTU0) -T0 #_ #H destruct -l0 //
-]
-qed-.
-
-(* Main properties **********************************************************)
-
-theorem sstas_mono: ∀h,g,L,T,U1. ⦃h, L⦄ ⊢ T •*[g] U1 →
- ∀U2. ⦃h, L⦄ ⊢ T •*[g] U2 → U1 = U2.
-#h #g #L #T #U1 #H @(sstas_ind_alt … H) -T
-[ #T1 #HUT1 #U2 #HU12
- >(sstas_inv_O … HU12 … HUT1) -h -L -T1 -U2 //
-| #T0 #U0 #l0 #HTU0 #_ #IHU01 #U2 #HU12
- lapply (sstas_inv_S … HU12 … HTU0) -T0 -l0 /2 width=1/
-]
-qed-.
-
-(* More advancd inversion lemmas ********************************************)
-
-fact sstas_inv_lref1_aux: ∀h,g,L,T,U. ⦃h, L⦄ ⊢ T •*[g] U → ∀j. T = #j →
- ∃∃I,K,V,W. ⇩[0, j] L ≡ K. ⓑ{I}V & ⦃h, K⦄ ⊢ V •*[g] W &
- L ⊢ ▼*[0, j + 1] U ≡ W.
-#h #g #L #T #U #H @(sstas_ind_alt … H) -T
-[ #T #HUT #j #H destruct
- elim (ssta_inv_lref1 … HUT) -HUT * #K #V #W [2: #l] #HLK #HVW #HVT
- [ <plus_n_Sm #H destruct
- | /3 width=12/
- ]
-| #T0 #U0 #l0 #HTU0 #HU0 #_ #j #H destruct
- elim (ssta_inv_lref1 … HTU0) -HTU0 * #K #V #W [2: #l] #HLK #HVW #HVU0
- [ #_ -HVW
- lapply (ldrop_fwd_ldrop2 … HLK) #H
- elim (sstas_inv_lift1 … HU0 … H … HVU0) -HU0 -H -HVU0 /3 width=7/
- | elim (sstas_total_S … HVW) -HVW #T #HVT #HWT
- lapply (ldrop_fwd_ldrop2 … HLK) #H
- elim (sstas_inv_lift1 … HU0 … H … HVU0) -HU0 -H -HVU0 #X #HWX
- >(sstas_mono … HWX … HWT) -X -W /3 width=7/
- ]
-]
-qed-.
<key name="ex">3 2</key>
<key name="ex">3 3</key>
<key name="ex">3 4</key>
+ <key name="ex">4 1</key>
<key name="ex">4 2</key>
<key name="ex">4 3</key>
<key name="ex">4 4</key>
+ <key name="ex">4 5</key>
<key name="ex">5 2</key>
<key name="ex">5 3</key>
<key name="ex">5 4</key>
+ <key name="ex">5 5</key>
<key name="ex">6 4</key>
<key name="ex">6 6</key>
- <key name="ex">7 6</key>
+ <key name="ex">6 7</key>
+ <key name="ex">7 7</key>
<key name="or">3</key>
<key name="or">4</key>
<key name="and">3</key>
interpretation "multiple existental quantifier (3, 4)" 'Ex P0 P1 P2 = (ex3_4 ? ? ? ? P0 P1 P2).
+(* multiple existental quantifier (4, 1) *)
+
+inductive ex4_1 (A0:Type[0]) (P0,P1,P2,P3:A0→Prop) : Prop ≝
+ | ex4_1_intro: ∀x0. P0 x0 → P1 x0 → P2 x0 → P3 x0 → ex4_1 ? ? ? ? ?
+.
+
+interpretation "multiple existental quantifier (4, 1)" 'Ex P0 P1 P2 P3 = (ex4_1 ? P0 P1 P2 P3).
+
(* multiple existental quantifier (4, 2) *)
inductive ex4_2 (A0,A1:Type[0]) (P0,P1,P2,P3:A0→A1→Prop) : Prop ≝
interpretation "multiple existental quantifier (4, 4)" 'Ex P0 P1 P2 P3 = (ex4_4 ? ? ? ? P0 P1 P2 P3).
+(* multiple existental quantifier (4, 5) *)
+
+inductive ex4_5 (A0,A1,A2,A3,A4:Type[0]) (P0,P1,P2,P3:A0→A1→A2→A3→A4→Prop) : Prop ≝
+ | ex4_5_intro: ∀x0,x1,x2,x3,x4. P0 x0 x1 x2 x3 x4 → P1 x0 x1 x2 x3 x4 → P2 x0 x1 x2 x3 x4 → P3 x0 x1 x2 x3 x4 → ex4_5 ? ? ? ? ? ? ? ? ?
+.
+
+interpretation "multiple existental quantifier (4, 5)" 'Ex P0 P1 P2 P3 = (ex4_5 ? ? ? ? ? P0 P1 P2 P3).
+
(* multiple existental quantifier (5, 2) *)
inductive ex5_2 (A0,A1:Type[0]) (P0,P1,P2,P3,P4:A0→A1→Prop) : Prop ≝
interpretation "multiple existental quantifier (5, 4)" 'Ex P0 P1 P2 P3 P4 = (ex5_4 ? ? ? ? P0 P1 P2 P3 P4).
+(* multiple existental quantifier (5, 5) *)
+
+inductive ex5_5 (A0,A1,A2,A3,A4:Type[0]) (P0,P1,P2,P3,P4:A0→A1→A2→A3→A4→Prop) : Prop ≝
+ | ex5_5_intro: ∀x0,x1,x2,x3,x4. P0 x0 x1 x2 x3 x4 → P1 x0 x1 x2 x3 x4 → P2 x0 x1 x2 x3 x4 → P3 x0 x1 x2 x3 x4 → P4 x0 x1 x2 x3 x4 → ex5_5 ? ? ? ? ? ? ? ? ? ?
+.
+
+interpretation "multiple existental quantifier (5, 5)" 'Ex P0 P1 P2 P3 P4 = (ex5_5 ? ? ? ? ? P0 P1 P2 P3 P4).
+
(* multiple existental quantifier (6, 4) *)
inductive ex6_4 (A0,A1,A2,A3:Type[0]) (P0,P1,P2,P3,P4,P5:A0→A1→A2→A3→Prop) : Prop ≝
interpretation "multiple existental quantifier (6, 6)" 'Ex P0 P1 P2 P3 P4 P5 = (ex6_6 ? ? ? ? ? ? P0 P1 P2 P3 P4 P5).
-(* multiple existental quantifier (7, 6) *)
+(* multiple existental quantifier (6, 7) *)
+
+inductive ex6_7 (A0,A1,A2,A3,A4,A5,A6:Type[0]) (P0,P1,P2,P3,P4,P5:A0→A1→A2→A3→A4→A5→A6→Prop) : Prop ≝
+ | ex6_7_intro: ∀x0,x1,x2,x3,x4,x5,x6. P0 x0 x1 x2 x3 x4 x5 x6 → P1 x0 x1 x2 x3 x4 x5 x6 → P2 x0 x1 x2 x3 x4 x5 x6 → P3 x0 x1 x2 x3 x4 x5 x6 → P4 x0 x1 x2 x3 x4 x5 x6 → P5 x0 x1 x2 x3 x4 x5 x6 → ex6_7 ? ? ? ? ? ? ? ? ? ? ? ? ?
+.
+
+interpretation "multiple existental quantifier (6, 7)" 'Ex P0 P1 P2 P3 P4 P5 = (ex6_7 ? ? ? ? ? ? ? P0 P1 P2 P3 P4 P5).
+
+(* multiple existental quantifier (7, 7) *)
-inductive ex7_6 (A0,A1,A2,A3,A4,A5:Type[0]) (P0,P1,P2,P3,P4,P5,P6:A0→A1→A2→A3→A4→A5→Prop) : Prop ≝
- | ex7_6_intro: ∀x0,x1,x2,x3,x4,x5. P0 x0 x1 x2 x3 x4 x5 → P1 x0 x1 x2 x3 x4 x5 → P2 x0 x1 x2 x3 x4 x5 → P3 x0 x1 x2 x3 x4 x5 → P4 x0 x1 x2 x3 x4 x5 → P5 x0 x1 x2 x3 x4 x5 → P6 x0 x1 x2 x3 x4 x5 → ex7_6 ? ? ? ? ? ? ? ? ? ? ? ? ?
+inductive ex7_7 (A0,A1,A2,A3,A4,A5,A6:Type[0]) (P0,P1,P2,P3,P4,P5,P6:A0→A1→A2→A3→A4→A5→A6→Prop) : Prop ≝
+ | ex7_7_intro: ∀x0,x1,x2,x3,x4,x5,x6. P0 x0 x1 x2 x3 x4 x5 x6 → P1 x0 x1 x2 x3 x4 x5 x6 → P2 x0 x1 x2 x3 x4 x5 x6 → P3 x0 x1 x2 x3 x4 x5 x6 → P4 x0 x1 x2 x3 x4 x5 x6 → P5 x0 x1 x2 x3 x4 x5 x6 → P6 x0 x1 x2 x3 x4 x5 x6 → ex7_7 ? ? ? ? ? ? ? ? ? ? ? ? ? ?
.
-interpretation "multiple existental quantifier (7, 6)" 'Ex P0 P1 P2 P3 P4 P5 P6 = (ex7_6 ? ? ? ? ? ? P0 P1 P2 P3 P4 P5 P6).
+interpretation "multiple existental quantifier (7, 7)" 'Ex P0 P1 P2 P3 P4 P5 P6 = (ex7_7 ? ? ? ? ? ? ? P0 P1 P2 P3 P4 P5 P6).
(* multiple disjunction connective (3) *)
non associative with precedence 20
for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P2) }.
+(* multiple existental quantifier (4, 1) *)
+
+notation > "hvbox(∃∃ ident x0 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3)"
+ non associative with precedence 20
+ for @{ 'Ex (λ${ident x0}.$P0) (λ${ident x0}.$P1) (λ${ident x0}.$P2) (λ${ident x0}.$P3) }.
+
+notation < "hvbox(∃∃ ident x0 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3)"
+ non associative with precedence 20
+ for @{ 'Ex (λ${ident x0}:$T0.$P0) (λ${ident x0}:$T0.$P1) (λ${ident x0}:$T0.$P2) (λ${ident x0}:$T0.$P3) }.
+
(* multiple existental quantifier (4, 2) *)
notation > "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3)"
non associative with precedence 20
for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P3) }.
+(* multiple existental quantifier (4, 5) *)
+
+notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3)"
+ non associative with precedence 20
+ for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.$P3) }.
+
+notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3)"
+ non associative with precedence 20
+ for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P3) }.
+
(* multiple existental quantifier (5, 2) *)
notation > "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4)"
non associative with precedence 20
for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P4) }.
+(* multiple existental quantifier (5, 5) *)
+
+notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4)"
+ non associative with precedence 20
+ for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.$P3) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.$P4) }.
+
+notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4)"
+ non associative with precedence 20
+ for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.$P4) }.
+
(* multiple existental quantifier (6, 4) *)
notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5)"
non associative with precedence 20
for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P4) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P5) }.
-(* multiple existental quantifier (7, 6) *)
+(* multiple existental quantifier (6, 7) *)
+
+notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 , ident x6 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5)"
+ non associative with precedence 20
+ for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.$P3) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.$P4) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.$P5) }.
+
+notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 , ident x6 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5)"
+ non associative with precedence 20
+ for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P4) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P5) }.
+
+(* multiple existental quantifier (7, 7) *)
-notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5 break & term 19 P6)"
+notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 , ident x6 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5 break & term 19 P6)"
non associative with precedence 20
- for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P3) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P4) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P5) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P6) }.
+ for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.$P3) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.$P4) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.$P5) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.λ${ident x6}.$P6) }.
-notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5 break & term 19 P6)"
+notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 , ident x6 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5 break & term 19 P6)"
non associative with precedence 20
- for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P4) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P5) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P6) }.
+ for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P4) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P5) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.λ${ident x6}:$T6.$P6) }.
(* multiple disjunction connective (3) *)
--- /dev/null
+for FILE in `find $1 -name "*.ma"`; do svn mv $FILE ${FILE/%.ma/.etc} ; done