(* Note: the constant 0 cannot be generalized *)
lemma lsubsv_drop_O1_conf: ∀h,o,G,L1,L2. G ⊢ L1 ⫃¡[h, o] L2 →
- ∀K1,c,k. ⬇[c, 0, k] L1 ≡ K1 →
- ∃∃K2. G ⊢ K1 ⫃¡[h, o] K2 & ⬇[c, 0, k] L2 ≡ K2.
+ ∀K1,b,k. ⬇[b, 0, k] L1 ≡ K1 →
+ ∃∃K2. G ⊢ K1 ⫃¡[h, o] K2 & ⬇[b, 0, k] L2 ≡ K2.
#h #o #G #L1 #L2 #H elim H -L1 -L2
[ /2 width=3 by ex2_intro/
-| #I #L1 #L2 #V #_ #IHL12 #K1 #c #k #H
+| #I #L1 #L2 #V #_ #IHL12 #K1 #b #k #H
elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK1
[ destruct
- elim (IHL12 L1 c 0) -IHL12 // #X #HL12 #H
+ elim (IHL12 L1 b 0) -IHL12 // #X #HL12 #H
<(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsubsv_pair, drop_pair, ex2_intro/
| elim (IHL12 … HLK1) -L1 /3 width=3 by drop_drop_lt, ex2_intro/
]
-| #L1 #L2 #W #V #d1 #HWV #HW #HVd1 #HWd1 #_ #IHL12 #K1 #c #k #H
+| #L1 #L2 #W #V #d1 #HWV #HW #HVd1 #HWd1 #_ #IHL12 #K1 #b #k #H
elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK1
[ destruct
- elim (IHL12 L1 c 0) -IHL12 // #X #HL12 #H
+ elim (IHL12 L1 b 0) -IHL12 // #X #HL12 #H
<(drop_inv_O2 … H) in HL12; -H /3 width=4 by lsubsv_beta, drop_pair, ex2_intro/
| elim (IHL12 … HLK1) -L1 /3 width=3 by drop_drop_lt, ex2_intro/
]
(* Note: the constant 0 cannot be generalized *)
lemma lsubsv_drop_O1_trans: ∀h,o,G,L1,L2. G ⊢ L1 ⫃¡[h, o] L2 →
- ∀K2,c, k. ⬇[c, 0, k] L2 ≡ K2 →
- ∃∃K1. G ⊢ K1 ⫃¡[h, o] K2 & ⬇[c, 0, k] L1 ≡ K1.
+ ∀K2,b, k. ⬇[b, 0, k] L2 ≡ K2 →
+ ∃∃K1. G ⊢ K1 ⫃¡[h, o] K2 & ⬇[b, 0, k] L1 ≡ K1.
#h #o #G #L1 #L2 #H elim H -L1 -L2
[ /2 width=3 by ex2_intro/
-| #I #L1 #L2 #V #_ #IHL12 #K2 #c #k #H
+| #I #L1 #L2 #V #_ #IHL12 #K2 #b #k #H
elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK2
[ destruct
- elim (IHL12 L2 c 0) -IHL12 // #X #HL12 #H
+ elim (IHL12 L2 b 0) -IHL12 // #X #HL12 #H
<(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsubsv_pair, drop_pair, ex2_intro/
| elim (IHL12 … HLK2) -L2 /3 width=3 by drop_drop_lt, ex2_intro/
]
-| #L1 #L2 #W #V #d1 #HWV #HW #HVd1 #HWd1 #_ #IHL12 #K2 #c #k #H
+| #L1 #L2 #W #V #d1 #HWV #HW #HVd1 #HWd1 #_ #IHL12 #K2 #b #k #H
elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK2
[ destruct
- elim (IHL12 L2 c 0) -IHL12 // #X #HL12 #H
+ elim (IHL12 L2 b 0) -IHL12 // #X #HL12 #H
<(drop_inv_O2 … H) in HL12; -H /3 width=4 by lsubsv_beta, drop_pair, ex2_intro/
| elim (IHL12 … HLK2) -L2 /3 width=3 by drop_drop_lt, ex2_intro/
]
| #G #L2 #K2 #V #W #U #i #d2 #HLK2 #_ #HWU #IHVW #d1 #Hd21 #Hd1 #L1 #HL12
elim (da_inv_lref … Hd1) -Hd1 * #K0 #V0 [| #d0 ] #HK0 #HV0 [| #H1 ]
lapply (drop_mono … HK0 … HLK2) -HK0 #H2 destruct
- lapply (le_plus_to_le_r … Hd21) -Hd21 #Hd21
+ lapply (le_plus_to_le_c … Hd21) -Hd21 #Hd21
elim (lsubsv_drop_O1_trans … HL12 … HLK2) -L2 #Y #H #HLK1
elim (lsubsv_inv_pair2 … H) -H * #K1
[ #HK12 #H destruct
(* Relocation properties ****************************************************)
-lemma snv_lift: ∀h,o,G,K,T. ⦃G, K⦄ ⊢ T ¡[h, o] → ∀L,c,l,k. ⬇[c, l, k] L ≡ K →
+lemma snv_lift: ∀h,o,G,K,T. ⦃G, K⦄ ⊢ T ¡[h, o] → ∀L,b,l,k. ⬇[b, l, k] L ≡ K →
∀U. ⬆[l, k] T ≡ U → ⦃G, L⦄ ⊢ U ¡[h, o].
#h #o #G #K #T #H elim H -G -K -T
-[ #G #K #s #L #c #l #k #_ #X #H
+[ #G #K #s #L #b #l #k #_ #X #H
>(lift_inv_sort1 … H) -X -K -l -k //
-| #I #G #K #K0 #V #i #HK0 #_ #IHV #L #c #l #k #HLK #X #H
+| #I #G #K #K0 #V #i #HK0 #_ #IHV #L #b #l #k #HLK #X #H
elim (lift_inv_lref1 … H) * #Hil #H destruct
[ elim (drop_trans_le … HLK … HK0) -K /2 width=2 by ylt_fwd_le/ #X #HL0 #H
elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ -Hil #L0 #W #HLK0 #HVW #H destruct
| lapply (drop_trans_ge … HLK … HK0 ?) -K
/3 width=9 by snv_lref, drop_inv_gen/
]
-| #a #I #G #K #V #T #_ #_ #IHV #IHT #L #c #l #k #HLK #X #H
+| #a #I #G #K #V #T #_ #_ #IHV #IHT #L #b #l #k #HLK #X #H
elim (lift_inv_bind1 … H) -H #W #U #HVW #HTU #H destruct
/4 width=5 by snv_bind, drop_skip/
-| #a #G #K #V #W0 #T #U0 #d #_ #_ #HVW0 #HTU0 #IHV #IHT #L #c #l #k #HLK #X #H
+| #a #G #K #V #W0 #T #U0 #d #_ #_ #HVW0 #HTU0 #IHV #IHT #L #b #l #k #HLK #X #H
elim (lift_inv_flat1 … H) -H #W #U #HVW #HTU #H destruct
elim (lift_total W0 l k)
elim (lift_total U0 (l+1) k)
/4 width=17 by snv_appl, scpds_lift, lift_bind/
-| #G #K #V #T #U0 #_ #_ #HVU0 #HTU0 #IHV #IHT #L #c #l #k #HLK #X #H
+| #G #K #V #T #U0 #_ #_ #HVU0 #HTU0 #IHV #IHT #L #b #l #k #HLK #X #H
elim (lift_inv_flat1 … H) -H #W #U #HVW #HTU #H destruct
elim (lift_total U0 l k)
/3 width=12 by snv_cast, scpds_lift/
]
qed.
-lemma snv_inv_lift: ∀h,o,G,L,U. ⦃G, L⦄ ⊢ U ¡[h, o] → ∀K,c,l,k. ⬇[c, l, k] L ≡ K →
+lemma snv_inv_lift: ∀h,o,G,L,U. ⦃G, L⦄ ⊢ U ¡[h, o] → ∀K,b,l,k. ⬇[b, l, k] L ≡ K →
∀T. ⬆[l, k] T ≡ U → ⦃G, K⦄ ⊢ T ¡[h, o].
#h #o #G #L #U #H elim H -G -L -U
-[ #G #L #s #K #c #l #k #_ #X #H
+[ #G #L #s #K #b #l #k #_ #X #H
>(lift_inv_sort2 … H) -X -L -l -k //
-| #I #G #L #L0 #W #i #HL0 #_ #IHW #K #c #l #k #HLK #X #H
+| #I #G #L #L0 #W #i #HL0 #_ #IHW #K #b #l #k #HLK #X #H
elim (lift_inv_lref2 … H) * #Hil #H destruct
[ elim (drop_conf_le … HLK … HL0) -L /2 width=2 by ylt_fwd_le/ #X #HK0 #H
elim (drop_inv_skip1 … H) -H /2 width=1 by ylt_to_minus/ -Hil #K0 #V #HLK0 #HVW #H destruct
/3 width=12 by snv_lref/
| lapply (drop_conf_ge … HLK … HL0 ?) -L /3 width=9 by snv_lref/
]
-| #a #I #G #L #W #U #_ #_ #IHW #IHU #K #c #l #k #HLK #X #H
+| #a #I #G #L #W #U #_ #_ #IHW #IHU #K #b #l #k #HLK #X #H
elim (lift_inv_bind2 … H) -H #V #T #HVW #HTU #H destruct
/4 width=5 by snv_bind, drop_skip/
-| #a #G #L #W #W1 #U #U1 #d #_ #_ #HW1 #HU1 #IHW #IHU #K #c #l #k #HLK #X #H
+| #a #G #L #W #W1 #U #U1 #d #_ #_ #HW1 #HU1 #IHW #IHU #K #b #l #k #HLK #X #H
elim (lift_inv_flat2 … H) -H #V #T #HVW #HTU #H destruct
elim (scpds_inv_lift1 … HW1 … HLK … HVW) -HW1 #W0 #HW01 #HVW0
elim (scpds_inv_lift1 … HU1 … HLK … HTU) -HU1 #X #H #HTU0
elim (lift_inv_bind2 … H) -H #Y #U0 #HY #HU01 #H destruct
lapply (lift_inj … HY … HW01) -HY #H destruct
/3 width=6 by snv_appl/
-| #G #L #W #U #U1 #_ #_ #HWU1 #HU1 #IHW #IHU #K #c #l #k #HLK #X #H
+| #G #L #W #U #U1 #_ #_ #HWU1 #HU1 #IHW #IHU #K #b #l #k #HLK #X #H
elim (lift_inv_flat2 … H) -H #V #T #HVW #HTU #H destruct
elim (scpds_inv_lift1 … HWU1 … HLK … HVW) -HWU1 #U0 #HU01 #HVU0
elim (scpds_inv_lift1 … HU1 … HLK … HTU) -HU1 #X #HX #HTU0
lapply (drop_mono … HLK0 … HLK1) -HLK0 #H0 destruct
elim (lstas_inv_lref1 … H2) -H2 * #K #Y #X [3,6: #d ] #HLK #HYX [1,2: #HXT #H0 |3,5: #HXT |4,6: #H1 #H2 ]
lapply (drop_mono … HLK … HLK1) -HLK #H destruct
- [ lapply (le_plus_to_le_r … Hd21) -Hd21 #Hd21 |3: -Hd21 ]
+ [ lapply (le_plus_to_le_c … Hd21) -Hd21 #Hd21 |3: -Hd21 ]
lapply (fqup_lref … G1 … HLK1) #H
lapply (drop_fwd_drop2 … HLK1) /4 width=8 by snv_lift, snv_lref, fqup_fpbg/
| #p #HG0 #HL0 #HT0 #H1 #d1 #d2 #Hd21 #Hd1 #X #H2 destruct -IH4 -IH3 -IH2 -IH1
lapply (drop_mono … HK0 … HLK1) -HK0 #H destruct
elim (lstas_inv_lref1 … H2) -H2 * #K0 #V0 #X0 [3,6: #d0 ] #HK0 #HVX0 [1,2: #HX02 #H |3,5: #HX02 |4,6: #H1 #H2 ] destruct
lapply (drop_mono … HK0 … HLK1) -HK0 #H destruct
- [ lapply (le_plus_to_le_r … Hd21) -Hd21 #Hd21 |3: -Hd21 ]
+ [ lapply (le_plus_to_le_c … Hd21) -Hd21 #Hd21 |3: -Hd21 ]
lapply (fqup_lref … G1 … HLK1) #HKV1
elim (lpr_drop_conf … HLK1 … HL12) -HL12 #X #H #HLK2
elim (lpr_inv_pair1 … H) -H #K2 [ #W2 | #W2 | #V2 ] #HK12 [ #HW12 | #HW12 | #HV12 ] #H destruct
qed-.
(* Basic_1: was: pc3_gen_lift *)
-lemma cpcs_inv_lift: ∀G,L,K,c,l,k. ⬇[c, l, k] L ≡ K →
+lemma cpcs_inv_lift: ∀G,L,K,b,l,k. ⬇[b, l, k] L ≡ K →
∀T1,U1. ⬆[l, k] T1 ≡ U1 → ∀T2,U2. ⬆[l, k] T2 ≡ U2 →
⦃G, L⦄ ⊢ U1 ⬌* U2 → ⦃G, K⦄ ⊢ T1 ⬌* T2.
-#G #L #K #c #l #k #HLK #T1 #U1 #HTU1 #T2 #U2 #HTU2 #HU12
+#G #L #K #b #l #k #HLK #T1 #U1 #HTU1 #T2 #U2 #HTU2 #HU12
elim (cpcs_inv_cprs … HU12) -HU12 #U #HU1 #HU2
elim (cprs_inv_lift1 … HU1 … HLK … HTU1) -U1 #T #HTU #HT1
elim (cprs_inv_lift1 … HU2 … HLK … HTU2) -L -U2 #X #HXU
qed-.
(* Basic_1: was: pc3_lift *)
-lemma cpcs_lift: ∀G,L,K,c,l,k. ⬇[c, l, k] L ≡ K →
+lemma cpcs_lift: ∀G,L,K,b,l,k. ⬇[b, l, k] L ≡ K →
∀T1,U1. ⬆[l, k] T1 ≡ U1 → ∀T2,U2. ⬆[l, k] T2 ≡ U2 →
⦃G, K⦄ ⊢ T1 ⬌* T2 → ⦃G, L⦄ ⊢ U1 ⬌* U2.
-#G #L #K #c #l #k #HLK #T1 #U1 #HTU1 #T2 #U2 #HTU2 #HT12
+#G #L #K #b #l #k #HLK #T1 #U1 #HTU1 #T2 #U2 #HTU2 #HT12
elim (cpcs_inv_cprs … HT12) -HT12 #T #HT1 #HT2
elim (lift_total T l k) /3 width=12 by cprs_div, cprs_lift/
qed.
T,U,V,W: term
X,Y,Z : reserved: transient objet denoted by a capital letter
-a,b :
-c : local dropping kind parameter (true = restricted, false = general)
+a :
+b : local dropping kind parameter (true = restricted, false = general)
+c : rt-reduction count parameter
d : term degree
e : reserved: future use (\lambda\delta 3)
f,g : local reference transforming map
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( ⦃ term 46 G, break term 46 L ⦄ ⊢ break term 46 T1 ➡ break [ term 46 h , break term 46 r ] break term 46 T2 )"
+notation "hvbox( ⦃ term 46 G, break term 46 L ⦄ ⊢ break term 46 T1 ➡ break [ term 46 c , break term 46 h ] break term 46 T2 )"
non associative with precedence 45
- for @{ 'PRed $h $r $G $L $T1 $T2 }.
+ for @{ 'PRed $c $h $G $L $T1 $T2 }.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-notation "hvbox( ⬇ * [ term 46 c , break term 46 f ] break term 46 L1 ≡ break term 46 L2 )"
+notation "hvbox( ⬇ * [ term 46 b , break term 46 f ] break term 46 L1 ≡ break term 46 L2 )"
non associative with precedence 45
- for @{ 'RDropStar $c $f $L1 $L2 }.
+ for @{ 'RDropStar $b $f $L1 $L2 }.
(* Basic_2A1: includes: drop_atom drop_pair drop_drop drop_skip
drop_refl_atom_O2 drop_drop_lt drop_skip_lt
*)
-inductive drops (c:bool): rtmap → relation lenv ≝
-| drops_atom: ∀f. (c = Ⓣ → 𝐈⦃f⦄) → drops c (f) (⋆) (⋆)
-| drops_drop: ∀I,L1,L2,V,f. drops c f L1 L2 → drops c (⫯f) (L1.ⓑ{I}V) L2
-| drops_skip: ∀I,L1,L2,V1,V2,f.
- drops c f L1 L2 → ⬆*[f] V2 ≡ V1 →
- drops c (↑f) (L1.ⓑ{I}V1) (L2.ⓑ{I}V2)
+inductive drops (b:bool): rtmap → relation lenv ≝
+| drops_atom: ∀f. (b = Ⓣ → 𝐈⦃f⦄) → drops b (f) (⋆) (⋆)
+| drops_drop: ∀f,I,L1,L2,V. drops b f L1 L2 → drops b (⫯f) (L1.ⓑ{I}V) L2
+| drops_skip: ∀f,I,L1,L2,V1,V2.
+ drops b f L1 L2 → ⬆*[f] V2 ≡ V1 →
+ drops b (↑f) (L1.ⓑ{I}V1) (L2.ⓑ{I}V2)
.
interpretation "uniform slicing (local environment)"
'RDropStar i L1 L2 = (drops true (uni i) L1 L2).
interpretation "generic slicing (local environment)"
- 'RDropStar c f L1 L2 = (drops c f L1 L2).
+ 'RDropStar b f L1 L2 = (drops b f L1 L2).
definition d_liftable1: relation2 lenv term → predicate bool ≝
- λR,c. ∀L,K,f. ⬇*[c, f] L ≡ K →
+ λR,b. ∀f,L,K. ⬇*[b, f] L ≡ K →
∀T,U. ⬆*[f] T ≡ U → R K T → R L U.
definition d_liftable2: predicate (lenv → relation term) ≝
- λR. ∀K,T1,T2. R K T1 T2 → ∀L,c,f. ⬇*[c, f] L ≡ K →
+ λR. ∀K,T1,T2. R K T1 T2 → ∀b,f,L. ⬇*[b, f] L ≡ K →
∀U1. ⬆*[f] T1 ≡ U1 →
∃∃U2. ⬆*[f] T2 ≡ U2 & R L U1 U2.
definition d_deliftable2_sn: predicate (lenv → relation term) ≝
- λR. ∀L,U1,U2. R L U1 U2 → ∀K,c,f. ⬇*[c, f] L ≡ K →
+ λR. ∀L,U1,U2. R L U1 U2 → ∀b,f,K. ⬇*[b, f] L ≡ K →
∀T1. ⬆*[f] T1 ≡ U1 →
∃∃T2. ⬆*[f] T2 ≡ U2 & R K T1 T2.
definition dropable_sn: predicate (rtmap → relation lenv) ≝
- λR. ∀L1,K1,c,f. ⬇*[c, f] L1 ≡ K1 → ∀L2,f2. R f2 L1 L2 →
+ λR. ∀b,f,L1,K1. ⬇*[b, f] L1 ≡ K1 → ∀f2,L2. R f2 L1 L2 →
∀f1. f ⊚ f1 ≡ f2 →
- ∃∃K2. R f1 K1 K2 & ⬇*[c, f] L2 ≡ K2.
+ ∃∃K2. R f1 K1 K2 & ⬇*[b, f] L2 ≡ K2.
definition dropable_dx: predicate (rtmap → relation lenv) ≝
- λR. ∀L1,L2,f2. R f2 L1 L2 →
- ∀K2,c,f. ⬇*[c, f] L2 ≡ K2 → 𝐔⦃f⦄ →
+ λR. ∀f2,L1,L2. R f2 L1 L2 →
+ ∀b,f,K2. ⬇*[b, f] L2 ≡ K2 → 𝐔⦃f⦄ →
∀f1. f ⊚ f1 ≡ f2 →
- ∃∃K1. ⬇*[c, f] L1 ≡ K1 & R f1 K1 K2.
+ ∃∃K1. ⬇*[b, f] L1 ≡ K1 & R f1 K1 K2.
definition dedropable_sn: predicate (rtmap → relation lenv) ≝
- λR. ∀L1,K1,c,f. ⬇*[c, f] L1 ≡ K1 → ∀K2,f1. R f1 K1 K2 →
+ λR. ∀b,f,L1,K1. ⬇*[b, f] L1 ≡ K1 → ∀f1,K2. R f1 K1 K2 →
∀f2. f ⊚ f1 ≡ f2 →
- ∃∃L2. R f2 L1 L2 & ⬇*[c, f] L2 ≡ K2 & L1 ≡[f] L2.
+ ∃∃L2. R f2 L1 L2 & ⬇*[b, f] L2 ≡ K2 & L1 ≡[f] L2.
(* Basic properties *********************************************************)
-lemma drops_eq_repl_back: ∀L1,L2,c. eq_repl_back … (λf. ⬇*[c, f] L1 ≡ L2).
-#L1 #L2 #c #f1 #H elim H -L1 -L2 -f1
+lemma drops_eq_repl_back: ∀b,L1,L2. eq_repl_back … (λf. ⬇*[b, f] L1 ≡ L2).
+#b #L1 #L2 #f1 #H elim H -f1 -L1 -L2
[ /4 width=3 by drops_atom, isid_eq_repl_back/
-| #I #L1 #L2 #V #f1 #_ #IH #f2 #H elim (eq_inv_nx … H) -H
+| #f1 #I #L1 #L2 #V #_ #IH #f2 #H elim (eq_inv_nx … H) -H
/3 width=3 by drops_drop/
-| #I #L1 #L2 #V1 #v2 #f1 #_ #HV #IH #f2 #H elim (eq_inv_px … H) -H
+| #f1 #I #L1 #L2 #V1 #v2 #_ #HV #IH #f2 #H elim (eq_inv_px … H) -H
/3 width=3 by drops_skip, lifts_eq_repl_back/
]
qed-.
-lemma drops_eq_repl_fwd: ∀L1,L2,c. eq_repl_fwd … (λf. ⬇*[c, f] L1 ≡ L2).
-#L1 #L2 #c @eq_repl_sym /2 width=3 by drops_eq_repl_back/ (**) (* full auto fails *)
+lemma drops_eq_repl_fwd: ∀b,L1,L2. eq_repl_fwd … (λf. ⬇*[b, f] L1 ≡ L2).
+#b #L1 #L2 @eq_repl_sym /2 width=3 by drops_eq_repl_back/ (**) (* full auto fails *)
qed-.
lemma drops_inv_tls_at: ∀f,i1,i2. @⦃i1,f⦄ ≡ i2 →
- ∀c, L1,L2. ⬇*[c,⫱*[i2]f] L1 ≡ L2 →
- ⬇*[c,↑⫱*[⫯i2]f] L1 ≡ L2.
+ ∀b,L1,L2. ⬇*[b,⫱*[i2]f] L1 ≡ L2 →
+ ⬇*[b,↑⫱*[⫯i2]f] L1 ≡ L2.
/3 width=3 by drops_eq_repl_fwd, at_inv_tls/ qed-.
(* Basic_2A1: includes: drop_FT *)
-lemma drops_TF: ∀L1,L2,f. ⬇*[Ⓣ, f] L1 ≡ L2 → ⬇*[Ⓕ, f] L1 ≡ L2.
-#L1 #L2 #f #H elim H -L1 -L2 -f
+lemma drops_TF: ∀f,L1,L2. ⬇*[Ⓣ, f] L1 ≡ L2 → ⬇*[Ⓕ, f] L1 ≡ L2.
+#f #L1 #L2 #H elim H -f -L1 -L2
/3 width=1 by drops_atom, drops_drop, drops_skip/
qed.
(* Basic_2A1: includes: drop_gen *)
-lemma drops_gen: ∀L1,L2,c,f. ⬇*[Ⓣ, f] L1 ≡ L2 → ⬇*[c, f] L1 ≡ L2.
-#L1 #L2 * /2 width=1 by drops_TF/
+lemma drops_gen: ∀b,f,L1,L2. ⬇*[Ⓣ, f] L1 ≡ L2 → ⬇*[b, f] L1 ≡ L2.
+* /2 width=1 by drops_TF/
qed-.
(* Basic_2A1: includes: drop_T *)
-lemma drops_F: ∀L1,L2,c,f. ⬇*[c, f] L1 ≡ L2 → ⬇*[Ⓕ, f] L1 ≡ L2.
-#L1 #L2 * /2 width=1 by drops_TF/
+lemma drops_F: ∀b,f,L1,L2. ⬇*[b, f] L1 ≡ L2 → ⬇*[Ⓕ, f] L1 ≡ L2.
+* /2 width=1 by drops_TF/
qed-.
(* Basic_2A1: includes: drop_refl *)
-lemma drops_refl: ∀c,L,f. 𝐈⦃f⦄ → ⬇*[c, f] L ≡ L.
-#c #L elim L -L /2 width=1 by drops_atom/
+lemma drops_refl: ∀b,L,f. 𝐈⦃f⦄ → ⬇*[b, f] L ≡ L.
+#b #L elim L -L /2 width=1 by drops_atom/
#L #I #V #IHL #f #Hf elim (isid_inv_gen … Hf) -Hf
/3 width=1 by drops_skip, lifts_refl/
qed.
(* Basic_2A1: includes: drop_split *)
-lemma drops_split_trans: ∀L1,L2,f,c. ⬇*[c, f] L1 ≡ L2 → ∀f1,f2. f1 ⊚ f2 ≡ f → 𝐔⦃f1⦄ →
- ∃∃L. ⬇*[c, f1] L1 ≡ L & ⬇*[c, f2] L ≡ L2.
-#L1 #L2 #f #c #H elim H -L1 -L2 -f
-[ #f #Hc #f1 #f2 #Hf #Hf1 @(ex2_intro … (⋆)) @drops_atom
- #H lapply (Hc H) -c
+lemma drops_split_trans: ∀b,f,L1,L2. ⬇*[b, f] L1 ≡ L2 → ∀f1,f2. f1 ⊚ f2 ≡ f → 𝐔⦃f1⦄ →
+ ∃∃L. ⬇*[b, f1] L1 ≡ L & ⬇*[b, f2] L ≡ L2.
+#b #f #L1 #L2 #H elim H -f -L1 -L2
+[ #f #H0f #f1 #f2 #Hf #Hf1 @(ex2_intro … (⋆)) @drops_atom
+ #H lapply (H0f H) -b
#H elim (after_inv_isid3 … Hf H) -f //
-| #I #L1 #L2 #V #f #HL12 #IHL12 #f1 #f2 #Hf #Hf1 elim (after_inv_xxn … Hf) -Hf [1,3: * |*: // ]
+| #f #I #L1 #L2 #V #HL12 #IHL12 #f1 #f2 #Hf #Hf1 elim (after_inv_xxn … Hf) -Hf [1,3: * |*: // ]
[ #g1 #g2 #Hf #H1 #H2 destruct
lapply (isuni_inv_push … Hf1 ??) -Hf1 [1,2: // ] #Hg1
elim (IHL12 … Hf) -f
| #g1 #Hf #H destruct elim (IHL12 … Hf) -f
/3 width=5 by ex2_intro, drops_drop, isuni_inv_next/
]
-| #I #L1 #L2 #V1 #V2 #f #_ #HV21 #IHL12 #f1 #f2 #Hf #Hf1 elim (after_inv_xxp … Hf) -Hf [2,3: // ]
+| #f #I #L1 #L2 #V1 #V2 #_ #HV21 #IHL12 #f1 #f2 #Hf #Hf1 elim (after_inv_xxp … Hf) -Hf [2,3: // ]
#g1 #g2 #Hf #H1 #H2 destruct elim (lifts_split_trans … HV21 … Hf) -HV21
elim (IHL12 … Hf) -f /3 width=5 by ex2_intro, drops_skip, isuni_fwd_push/
]
qed-.
-lemma drops_split_div: ∀L1,L,f1,c. ⬇*[c, f1] L1 ≡ L → ∀f2,f. f1 ⊚ f2 ≡ f → 𝐔⦃f2⦄ →
+lemma drops_split_div: ∀b,f1,L1,L. ⬇*[b, f1] L1 ≡ L → ∀f2,f. f1 ⊚ f2 ≡ f → 𝐔⦃f2⦄ →
∃∃L2. ⬇*[Ⓕ, f2] L ≡ L2 & ⬇*[Ⓕ, f] L1 ≡ L2.
-#L1 #L #f1 #c #H elim H -L1 -L -f1
-[ #f1 #Hc #f2 #f #Hf #Hf2 @(ex2_intro … (⋆)) @drops_atom #H destruct
-| #I #L1 #L #V #f1 #HL1 #IH #f2 #f #Hf #Hf2 elim (after_inv_nxx … Hf) -Hf [2,3: // ]
+#b #f1 #L1 #L #H elim H -f1 -L1 -L
+[ #f1 #Hf1 #f2 #f #Hf #Hf2 @(ex2_intro … (⋆)) @drops_atom #H destruct
+| #f1 #I #L1 #L #V #HL1 #IH #f2 #f #Hf #Hf2 elim (after_inv_nxx … Hf) -Hf [2,3: // ]
#g #Hg #H destruct elim (IH … Hg) -IH -Hg /3 width=5 by drops_drop, ex2_intro/
-| #I #L1 #L #V1 #V #f1 #HL1 #HV1 #IH #f2 #f #Hf #Hf2
+| #f1 #I #L1 #L #V1 #V #HL1 #HV1 #IH #f2 #f #Hf #Hf2
elim (after_inv_pxx … Hf) -Hf [1,3: * |*: // ]
#g2 #g #Hg #H2 #H0 destruct
[ lapply (isuni_inv_push … Hf2 ??) -Hf2 [1,2: // ] #Hg2 -IH
(* Basic_1: includes: drop_gen_refl *)
(* Basic_2A1: includes: drop_inv_O2 *)
-lemma drops_fwd_isid: ∀L1,L2,c,f. ⬇*[c, f] L1 ≡ L2 → 𝐈⦃f⦄ → L1 = L2.
-#L1 #L2 #c #f #H elim H -L1 -L2 -f //
-[ #I #L1 #L2 #V #f #_ #_ #H elim (isid_inv_next … H) //
+lemma drops_fwd_isid: ∀b,f,L1,L2. ⬇*[b, f] L1 ≡ L2 → 𝐈⦃f⦄ → L1 = L2.
+#b #f #L1 #L2 #H elim H -f -L1 -L2 //
+[ #f #I #L1 #L2 #V #_ #_ #H elim (isid_inv_next … H) //
| /5 width=5 by isid_inv_push, lifts_fwd_isid, eq_f3, sym_eq/
]
qed-.
-fact drops_fwd_drop2_aux: ∀X,Y,c,f2. ⬇*[c, f2] X ≡ Y → ∀I,K,V. Y = K.ⓑ{I}V →
- ∃∃f1,f. 𝐈⦃f1⦄ & f2 ⊚ ⫯f1 ≡ f & ⬇*[c, f] X ≡ K.
-#X #Y #c #f2 #H elim H -X -Y -f2
-[ #f2 #Ht2 #J #K #W #H destruct
-| #I #L1 #L2 #V #f2 #_ #IHL #J #K #W #H elim (IHL … H) -IHL
+fact drops_fwd_drop2_aux: ∀b,f2,X,Y. ⬇*[b, f2] X ≡ Y → ∀I,K,V. Y = K.ⓑ{I}V →
+ ∃∃f1,f. 𝐈⦃f1⦄ & f2 ⊚ ⫯f1 ≡ f & ⬇*[b, f] X ≡ K.
+#b #f2 #X #Y #H elim H -f2 -X -Y
+[ #f2 #Hf2 #J #K #W #H destruct
+| #f2 #I #L1 #L2 #V #_ #IHL #J #K #W #H elim (IHL … H) -IHL
/3 width=7 by after_next, ex3_2_intro, drops_drop/
-| #I #L1 #L2 #V1 #V2 #f2 #HL #_ #_ #J #K #W #H destruct
+| #f2 #I #L1 #L2 #V1 #V2 #HL #_ #_ #J #K #W #H destruct
lapply (isid_after_dx 𝐈𝐝 … f2) /3 width=9 by after_push, ex3_2_intro, drops_drop/
]
qed-.
-lemma drops_fwd_drop2: ∀I,X,K,V,c,f2. ⬇*[c, f2] X ≡ K.ⓑ{I}V →
- ∃∃f1,f. 𝐈⦃f1⦄ & f2 ⊚ ⫯f1 ≡ f & ⬇*[c, f] X ≡ K.
+lemma drops_fwd_drop2: ∀b,f2,I,X,K,V. ⬇*[b, f2] X ≡ K.ⓑ{I}V →
+ ∃∃f1,f. 𝐈⦃f1⦄ & f2 ⊚ ⫯f1 ≡ f & ⬇*[b, f] X ≡ K.
/2 width=5 by drops_fwd_drop2_aux/ qed-.
-lemma drops_after_fwd_drop2: ∀I,X,K,V,c,f2. ⬇*[c, f2] X ≡ K.ⓑ{I}V →
- ∀f1,f. 𝐈⦃f1⦄ → f2 ⊚ ⫯f1 ≡ f → ⬇*[c, f] X ≡ K.
-#I #X #K #V #c #f2 #H #f1 #f #Hf1 #Hf elim (drops_fwd_drop2 … H) -H
+lemma drops_after_fwd_drop2: ∀b,f2,I,X,K,V. ⬇*[b, f2] X ≡ K.ⓑ{I}V →
+ ∀f1,f. 𝐈⦃f1⦄ → f2 ⊚ ⫯f1 ≡ f → ⬇*[b, f] X ≡ K.
+#b #f2 #I #X #K #V #H #f1 #f #Hf1 #Hf elim (drops_fwd_drop2 … H) -H
#g1 #g #Hg1 #Hg #HK lapply (after_mono_eq … Hg … Hf ??) -Hg -Hf
/3 width=5 by drops_eq_repl_back, isid_inv_eq_repl, eq_next/
qed-.
(* Basic_1: was: drop_S *)
(* Basic_2A1: was: drop_fwd_drop2 *)
-lemma drops_isuni_fwd_drop2: ∀I,X,K,V,c,f. 𝐔⦃f⦄ → ⬇*[c, f] X ≡ K.ⓑ{I}V → ⬇*[c, ⫯f] X ≡ K.
+lemma drops_isuni_fwd_drop2: ∀b,f,I,X,K,V. 𝐔⦃f⦄ → ⬇*[b, f] X ≡ K.ⓑ{I}V → ⬇*[b, ⫯f] X ≡ K.
/3 width=7 by drops_after_fwd_drop2, after_isid_isuni/ qed-.
(* Forward lemmas with test for finite colength *****************************)
-lemma drops_fwd_isfin: ∀L1,L2,f. ⬇*[Ⓣ, f] L1 ≡ L2 → 𝐅⦃f⦄.
-#L1 #L2 #f #H elim H -L1 -L2 -f
+lemma drops_fwd_isfin: ∀f,L1,L2. ⬇*[Ⓣ, f] L1 ≡ L2 → 𝐅⦃f⦄.
+#f #L1 #L2 #H elim H -f -L1 -L2
/3 width=1 by isfin_next, isfin_push, isfin_isid/
qed-.
(* Basic inversion lemmas ***************************************************)
-fact drops_inv_atom1_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → X = ⋆ →
- Y = ⋆ ∧ (c = Ⓣ → 𝐈⦃f⦄).
-#X #Y #c #f * -X -Y -f
+fact drops_inv_atom1_aux: ∀b,f,X,Y. ⬇*[b, f] X ≡ Y → X = ⋆ →
+ Y = ⋆ ∧ (b = Ⓣ → 𝐈⦃f⦄).
+#b #f #X #Y * -f -X -Y
[ /3 width=1 by conj/
-| #I #L1 #L2 #V #f #_ #H destruct
-| #I #L1 #L2 #V1 #V2 #f #_ #_ #H destruct
+| #f #I #L1 #L2 #V #_ #H destruct
+| #f #I #L1 #L2 #V1 #V2 #_ #_ #H destruct
]
qed-.
(* Basic_1: includes: drop_gen_sort *)
(* Basic_2A1: includes: drop_inv_atom1 *)
-lemma drops_inv_atom1: ∀Y,c,f. ⬇*[c, f] ⋆ ≡ Y → Y = ⋆ ∧ (c = Ⓣ → 𝐈⦃f⦄).
+lemma drops_inv_atom1: ∀b,f,Y. ⬇*[b, f] ⋆ ≡ Y → Y = ⋆ ∧ (b = Ⓣ → 𝐈⦃f⦄).
/2 width=3 by drops_inv_atom1_aux/ qed-.
-fact drops_inv_drop1_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → ∀I,K,V,g. X = K.ⓑ{I}V → f = ⫯g →
- ⬇*[c, g] K ≡ Y.
-#X #Y #c #f * -X -Y -f
-[ #f #Ht #J #K #W #g #H destruct
-| #I #L1 #L2 #V #f #HL #J #K #W #g #H1 #H2 <(injective_next … H2) -g destruct //
-| #I #L1 #L2 #V1 #V2 #f #_ #_ #J #K #W #g #_ #H2 elim (discr_push_next … H2)
+fact drops_inv_drop1_aux: ∀b,f,X,Y. ⬇*[b, f] X ≡ Y → ∀g,I,K,V. X = K.ⓑ{I}V → f = ⫯g →
+ ⬇*[b, g] K ≡ Y.
+#b #f #X #Y * -f -X -Y
+[ #f #Hf #g #J #K #W #H destruct
+| #f #I #L1 #L2 #V #HL #g #J #K #W #H1 #H2 <(injective_next … H2) -g destruct //
+| #f #I #L1 #L2 #V1 #V2 #_ #_ #g #J #K #W #_ #H2 elim (discr_push_next … H2)
]
qed-.
(* Basic_1: includes: drop_gen_drop *)
(* Basic_2A1: includes: drop_inv_drop1_lt drop_inv_drop1 *)
-lemma drops_inv_drop1: ∀I,K,Y,V,c,f. ⬇*[c, ⫯f] K.ⓑ{I}V ≡ Y → ⬇*[c, f] K ≡ Y.
+lemma drops_inv_drop1: ∀b,f,I,K,Y,V. ⬇*[b, ⫯f] K.ⓑ{I}V ≡ Y → ⬇*[b, f] K ≡ Y.
/2 width=7 by drops_inv_drop1_aux/ qed-.
-
-fact drops_inv_skip1_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → ∀I,K1,V1,g. X = K1.ⓑ{I}V1 → f = ↑g →
- ∃∃K2,V2. ⬇*[c, g] K1 ≡ K2 & ⬆*[g] V2 ≡ V1 & Y = K2.ⓑ{I}V2.
-#X #Y #c #f * -X -Y -f
-[ #f #Ht #J #K1 #W1 #g #H destruct
-| #I #L1 #L2 #V #f #_ #J #K1 #W1 #g #_ #H2 elim (discr_next_push … H2)
-| #I #L1 #L2 #V1 #V2 #f #HL #HV #J #K1 #W1 #g #H1 #H2 <(injective_push … H2) -g destruct
+fact drops_inv_skip1_aux: ∀b,f,X,Y. ⬇*[b, f] X ≡ Y → ∀g,I,K1,V1. X = K1.ⓑ{I}V1 → f = ↑g →
+ ∃∃K2,V2. ⬇*[b, g] K1 ≡ K2 & ⬆*[g] V2 ≡ V1 & Y = K2.ⓑ{I}V2.
+#b #f #X #Y * -f -X -Y
+[ #f #Hf #g #J #K1 #W1 #H destruct
+| #f #I #L1 #L2 #V #_ #g #J #K1 #W1 #_ #H2 elim (discr_next_push … H2)
+| #f #I #L1 #L2 #V1 #V2 #HL #HV #g #J #K1 #W1 #H1 #H2 <(injective_push … H2) -g destruct
/2 width=5 by ex3_2_intro/
]
qed-.
(* Basic_1: includes: drop_gen_skip_l *)
(* Basic_2A1: includes: drop_inv_skip1 *)
-lemma drops_inv_skip1: ∀I,K1,V1,Y,c,f. ⬇*[c, ↑f] K1.ⓑ{I}V1 ≡ Y →
- ∃∃K2,V2. ⬇*[c, f] K1 ≡ K2 & ⬆*[f] V2 ≡ V1 & Y = K2.ⓑ{I}V2.
+lemma drops_inv_skip1: ∀b,f,I,K1,V1,Y. ⬇*[b, ↑f] K1.ⓑ{I}V1 ≡ Y →
+ ∃∃K2,V2. ⬇*[b, f] K1 ≡ K2 & ⬆*[f] V2 ≡ V1 & Y = K2.ⓑ{I}V2.
/2 width=5 by drops_inv_skip1_aux/ qed-.
-fact drops_inv_skip2_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → ∀I,K2,V2,g. Y = K2.ⓑ{I}V2 → f = ↑g →
- ∃∃K1,V1. ⬇*[c, g] K1 ≡ K2 & ⬆*[g] V2 ≡ V1 & X = K1.ⓑ{I}V1.
-#X #Y #c #f * -X -Y -f
-[ #f #Ht #J #K2 #W2 #g #H destruct
-| #I #L1 #L2 #V #f #_ #J #K2 #W2 #g #_ #H2 elim (discr_next_push … H2)
-| #I #L1 #L2 #V1 #V2 #f #HL #HV #J #K2 #W2 #g #H1 #H2 <(injective_push … H2) -g destruct
+fact drops_inv_skip2_aux: ∀b,f,X,Y. ⬇*[b, f] X ≡ Y → ∀g,I,K2,V2. Y = K2.ⓑ{I}V2 → f = ↑g →
+ ∃∃K1,V1. ⬇*[b, g] K1 ≡ K2 & ⬆*[g] V2 ≡ V1 & X = K1.ⓑ{I}V1.
+#b #f #X #Y * -f -X -Y
+[ #f #Hf #g #J #K2 #W2 #H destruct
+| #f #I #L1 #L2 #V #_ #g #J #K2 #W2 #_ #H2 elim (discr_next_push … H2)
+| #f #I #L1 #L2 #V1 #V2 #HL #HV #g #J #K2 #W2 #H1 #H2 <(injective_push … H2) -g destruct
/2 width=5 by ex3_2_intro/
]
qed-.
(* Basic_1: includes: drop_gen_skip_r *)
(* Basic_2A1: includes: drop_inv_skip2 *)
-lemma drops_inv_skip2: ∀I,X,K2,V2,c,f. ⬇*[c, ↑f] X ≡ K2.ⓑ{I}V2 →
- ∃∃K1,V1. ⬇*[c, f] K1 ≡ K2 & ⬆*[f] V2 ≡ V1 & X = K1.ⓑ{I}V1.
+lemma drops_inv_skip2: ∀b,f,I,X,K2,V2. ⬇*[b, ↑f] X ≡ K2.ⓑ{I}V2 →
+ ∃∃K1,V1. ⬇*[b, f] K1 ≡ K2 & ⬆*[f] V2 ≡ V1 & X = K1.ⓑ{I}V1.
/2 width=5 by drops_inv_skip2_aux/ qed-.
-fact drops_inv_TF_aux: ∀L1,L2,f. ⬇*[Ⓕ, f] L1 ≡ L2 → 𝐔⦃f⦄ →
+fact drops_inv_TF_aux: ∀f,L1,L2. ⬇*[Ⓕ, f] L1 ≡ L2 → 𝐔⦃f⦄ →
∀I,K,V. L2 = K.ⓑ{I}V →
⬇*[Ⓣ, f] L1 ≡ K.ⓑ{I}V.
-#L1 #L2 #f #H elim H -L1 -L2 -f
+#f #L1 #L2 #H elim H -f -L1 -L2
[ #f #_ #_ #J #K #W #H destruct
-| #I #L1 #L2 #V #f #_ #IH #Hf #J #K #W #H destruct
+| #f #I #L1 #L2 #V #_ #IH #Hf #J #K #W #H destruct
/4 width=3 by drops_drop, isuni_inv_next/
-| #I #L1 #L2 #V1 #V2 #f #HL12 #HV21 #_ #Hf #J #K #W #H destruct
+| #f #I #L1 #L2 #V1 #V2 #HL12 #HV21 #_ #Hf #J #K #W #H destruct
lapply (isuni_inv_push … Hf ??) -Hf [1,2: // ] #Hf
<(drops_fwd_isid … HL12) -K // <(lifts_fwd_isid … HV21) -V1
/3 width=3 by drops_refl, isid_push/
qed-.
(* Basic_2A1: includes: drop_inv_FT *)
-lemma drops_inv_TF: ∀I,L,K,V,f. ⬇*[Ⓕ, f] L ≡ K.ⓑ{I}V → 𝐔⦃f⦄ →
+lemma drops_inv_TF: ∀f,I,L,K,V. ⬇*[Ⓕ, f] L ≡ K.ⓑ{I}V → 𝐔⦃f⦄ →
⬇*[Ⓣ, f] L ≡ K.ⓑ{I}V.
/2 width=3 by drops_inv_TF_aux/ qed-.
(* Advanced inversion lemmas ************************************************)
(* Basic_2A1: includes: drop_inv_gen *)
-lemma drops_inv_gen: ∀I,L,K,V,c,f. ⬇*[c, f] L ≡ K.ⓑ{I}V → 𝐔⦃f⦄ →
+lemma drops_inv_gen: ∀b,f,I,L,K,V. ⬇*[b, f] L ≡ K.ⓑ{I}V → 𝐔⦃f⦄ →
⬇*[Ⓣ, f] L ≡ K.ⓑ{I}V.
-#I #L #K #V * /2 width=1 by drops_inv_TF/
+* /2 width=1 by drops_inv_TF/
qed-.
(* Basic_2A1: includes: drop_inv_T *)
-lemma drops_inv_F: ∀I,L,K,V,c,f. ⬇*[Ⓕ, f] L ≡ K.ⓑ{I}V → 𝐔⦃f⦄ →
- ⬇*[c, f] L ≡ K.ⓑ{I}V.
-#I #L #K #V * /2 width=1 by drops_inv_TF/
+lemma drops_inv_F: ∀b,f,I,L,K,V. ⬇*[Ⓕ, f] L ≡ K.ⓑ{I}V → 𝐔⦃f⦄ →
+ ⬇*[b, f] L ≡ K.ⓑ{I}V.
+* /2 width=1 by drops_inv_TF/
qed-.
(* Inversion lemmas with test for uniformity ********************************)
-lemma drops_inv_isuni: ∀L1,L2,f. ⬇*[Ⓣ, f] L1 ≡ L2 → 𝐔⦃f⦄ →
+lemma drops_inv_isuni: ∀f,L1,L2. ⬇*[Ⓣ, f] L1 ≡ L2 → 𝐔⦃f⦄ →
(𝐈⦃f⦄ ∧ L1 = L2) ∨
- ∃∃I,K,V,g. ⬇*[Ⓣ, g] K ≡ L2 & 𝐔⦃g⦄ & L1 = K.ⓑ{I}V & f = ⫯g.
-#L1 #L2 #f * -L1 -L2 -f
+ ∃∃g,I,K,V. ⬇*[Ⓣ, g] K ≡ L2 & 𝐔⦃g⦄ & L1 = K.ⓑ{I}V & f = ⫯g.
+#f #L1 #L2 * -f -L1 -L2
[ /4 width=1 by or_introl, conj/
| /4 width=8 by isuni_inv_next, ex4_4_intro, or_intror/
| /7 width=6 by drops_fwd_isid, lifts_fwd_isid, isuni_inv_push, isid_push, or_introl, conj, eq_f3, sym_eq/
qed-.
(* Basic_2A1: was: drop_inv_O1_pair1 *)
-lemma drops_inv_pair1_isuni: ∀I,K,L2,V,c,f. 𝐔⦃f⦄ → ⬇*[c, f] K.ⓑ{I}V ≡ L2 →
+lemma drops_inv_pair1_isuni: ∀b,f,I,K,L2,V. 𝐔⦃f⦄ → ⬇*[b, f] K.ⓑ{I}V ≡ L2 →
(𝐈⦃f⦄ ∧ L2 = K.ⓑ{I}V) ∨
- ∃∃g. 𝐔⦃g⦄ & ⬇*[c, g] K ≡ L2 & f = ⫯g.
-#I #K #L2 #V #c #f #Hf #H elim (isuni_split … Hf) -Hf * #g #Hg #H0 destruct
+ ∃∃g. 𝐔⦃g⦄ & ⬇*[b, g] K ≡ L2 & f = ⫯g.
+#b #f #I #K #L2 #V #Hf #H elim (isuni_split … Hf) -Hf * #g #Hg #H0 destruct
[ lapply (drops_inv_skip1 … H) -H * #Y #X #HY #HX #H destruct
<(drops_fwd_isid … HY Hg) -Y >(lifts_fwd_isid … HX Hg) -X
/4 width=3 by isid_push, or_introl, conj/
qed-.
(* Basic_2A1: was: drop_inv_O1_pair2 *)
-lemma drops_inv_pair2_isuni: ∀I,K,V,c,f,L1. 𝐔⦃f⦄ → ⬇*[c, f] L1 ≡ K.ⓑ{I}V →
+lemma drops_inv_pair2_isuni: ∀b,f,I,K,V,L1. 𝐔⦃f⦄ → ⬇*[b, f] L1 ≡ K.ⓑ{I}V →
(𝐈⦃f⦄ ∧ L1 = K.ⓑ{I}V) ∨
- ∃∃I1,K1,V1,g. 𝐔⦃g⦄ & ⬇*[c, g] K1 ≡ K.ⓑ{I}V & L1 = K1.ⓑ{I1}V1 & f = ⫯g.
-#I #K #V #c #f *
+ ∃∃g,I1,K1,V1. 𝐔⦃g⦄ & ⬇*[b, g] K1 ≡ K.ⓑ{I}V & L1 = K1.ⓑ{I1}V1 & f = ⫯g.
+#b #f #I #K #V *
[ #Hf #H elim (drops_inv_atom1 … H) -H #H destruct
| #L1 #I1 #V1 #Hf #H elim (drops_inv_pair1_isuni … Hf H) -Hf -H *
[ #Hf #H destruct /3 width=1 by or_introl, conj/
]
qed-.
-lemma drops_inv_pair2_isuni_next: ∀I,K,V,c,f,L1. 𝐔⦃f⦄ → ⬇*[c, ⫯f] L1 ≡ K.ⓑ{I}V →
- ∃∃I1,K1,V1. ⬇*[c, f] K1 ≡ K.ⓑ{I}V & L1 = K1.ⓑ{I1}V1.
-#I #K #V #c #f #L1 #Hf #H elim (drops_inv_pair2_isuni … H) -H /2 width=3 by isuni_next/ -Hf *
+lemma drops_inv_pair2_isuni_next: ∀b,f,I,K,V,L1. 𝐔⦃f⦄ → ⬇*[b, ⫯f] L1 ≡ K.ⓑ{I}V →
+ ∃∃I1,K1,V1. ⬇*[b, f] K1 ≡ K.ⓑ{I}V & L1 = K1.ⓑ{I1}V1.
+#b #f #I #K #V #L1 #Hf #H elim (drops_inv_pair2_isuni … H) -H /2 width=3 by isuni_next/ -Hf *
[ #H elim (isid_inv_next … H) -H //
| /2 width=5 by ex2_3_intro/
]
(* Inversion lemmas with uniform relocations ********************************)
-lemma drops_inv_succ: ∀L1,L2,l. ⬇*[⫯l] L1 ≡ L2 →
+lemma drops_inv_succ: ∀l,L1,L2. ⬇*[⫯l] L1 ≡ L2 →
∃∃I,K,V. ⬇*[l] K ≡ L2 & L1 = K.ⓑ{I}V.
-#L1 #L2 #l #H elim (drops_inv_isuni … H) -H // *
+#l #L1 #L2 #H elim (drops_inv_isuni … H) -H // *
[ #H elim (isid_inv_next … H) -H //
| /2 width=5 by ex2_3_intro/
]
(* Note: cfull_inv_lift does not hold *)
lemma cfull_lift: d_liftable2 cfull.
-#K #T1 #T2 #_ #L #c #f #_ #U1 #_ -K -T1 -c
+#K #T1 #T2 #_ #b #f #L #_ #U1 #_ -K -T1 -b
elim (lifts_total T2 f) /2 width=3 by ex2_intro/
qed-.
(* Main properties **********************************************************)
(* Basic_2A1: includes: drop_conf_ge drop_conf_be drop_conf_le *)
-theorem drops_conf: ∀L1,L,c1,f1. ⬇*[c1, f1] L1 ≡ L →
- ∀L2,c2,f. ⬇*[c2, f] L1 ≡ L2 →
- ∀f2. f1 ⊚ f2 ≡ f → ⬇*[c2, f2] L ≡ L2.
-#L1 #L #c1 #f1 #H elim H -L1 -L -f1
-[ #f1 #_ #L2 #c2 #f #HL2 #f2 #Hf12 elim (drops_inv_atom1 … HL2) -c1 -HL2
+theorem drops_conf: ∀b1,f1,L1,L. ⬇*[b1, f1] L1 ≡ L →
+ ∀b2,f,L2. ⬇*[b2, f] L1 ≡ L2 →
+ ∀f2. f1 ⊚ f2 ≡ f → ⬇*[b2, f2] L ≡ L2.
+#b1 #f1 #L1 #L #H elim H -f1 -L1 -L
+[ #f1 #_ #b2 #f #L2 #HL2 #f2 #Hf12 elim (drops_inv_atom1 … HL2) -b1 -HL2
#H #Hf destruct @drops_atom
#H elim (after_inv_isid3 … Hf12) -Hf12 /2 width=1 by/
-| #I #K1 #K #V1 #f1 #_ #IH #L2 #c2 #f #HL2 #f2 #Hf elim (after_inv_nxx … Hf) -Hf [2,3: // ]
+| #f1 #I #K1 #K #V1 #_ #IH #b2 #f #L2 #HL2 #f2 #Hf elim (after_inv_nxx … Hf) -Hf [2,3: // ]
#g #Hg #H destruct /3 width=3 by drops_inv_drop1/
-| #I #K1 #K #V1 #V #f1 #_ #HV1 #IH #L2 #c2 #f #HL2 #f2 #Hf elim (after_inv_pxx … Hf) -Hf [1,3: * |*:// ]
+| #f1 #I #K1 #K #V1 #V #_ #HV1 #IH #b2 #f #L2 #HL2 #f2 #Hf elim (after_inv_pxx … Hf) -Hf [1,3: * |*:// ]
#g2 #g #Hf #H1 #H2 destruct
[ elim (drops_inv_skip1 … HL2) -HL2 /3 width=6 by drops_skip, lifts_div/
| /4 width=3 by drops_inv_drop1, drops_drop/
(* Basic_2A1: includes: drop_trans_ge drop_trans_le drop_trans_ge_comm
drops_drop_trans
*)
-theorem drops_trans: ∀L1,L,c1,f1. ⬇*[c1, f1] L1 ≡ L →
- ∀L2,c2,f2. ⬇*[c2, f2] L ≡ L2 →
- ∀f. f1 ⊚ f2 ≡ f → ⬇*[c1∧c2, f] L1 ≡ L2.
-#L1 #L #c1 #f1 #H elim H -L1 -L -f1
-[ #f1 #Hf1 #L2 #c2 #f2 #HL2 #f #Hf elim (drops_inv_atom1 … HL2) -HL2
+theorem drops_trans: ∀b1,f1,L1,L. ⬇*[b1, f1] L1 ≡ L →
+ ∀b2,f2,L2. ⬇*[b2, f2] L ≡ L2 →
+ ∀f. f1 ⊚ f2 ≡ f → ⬇*[b1∧b2, f] L1 ≡ L2.
+#b1 #f1 #L1 #L #H elim H -f1 -L1 -L
+[ #f1 #Hf1 #b2 #f2 #L2 #HL2 #f #Hf elim (drops_inv_atom1 … HL2) -HL2
#H #Hf2 destruct @drops_atom #H elim (andb_inv_true_dx … H) -H
#H1 #H2 lapply (after_isid_inv_sn … Hf ?) -Hf
/3 width=3 by isid_eq_repl_back/
-| #I #K1 #K #V1 #f1 #_ #IH #L2 #c2 #f2 #HL2 #f #Hf elim (after_inv_nxx … Hf) -Hf
+| #f1 #I #K1 #K #V1 #_ #IH #b2 #f2 #L2 #HL2 #f #Hf elim (after_inv_nxx … Hf) -Hf
/3 width=3 by drops_drop/
-| #I #K1 #K #V1 #V #f1 #_ #HV1 #IH #L2 #c2 #f2 #HL2 #f #Hf elim (after_inv_pxx … Hf) -Hf [1,3: * |*: // ]
+| #f1 #I #K1 #K #V1 #V #_ #HV1 #IH #b2 #f2 #L2 #HL2 #f #Hf elim (after_inv_pxx … Hf) -Hf [1,3: * |*: // ]
#g2 #g #Hg #H1 #H2 destruct
[ elim (drops_inv_skip1 … HL2) -HL2 /3 width=6 by drops_skip, lifts_trans/
| /4 width=3 by drops_inv_drop1, drops_drop/
]
qed-.
-theorem drops_conf_div: ∀L,K,f1. ⬇*[Ⓣ,f1] L ≡ K → ∀f2. ⬇*[Ⓣ,f2] L ≡ K →
+theorem drops_conf_div: ∀f1,L,K. ⬇*[Ⓣ,f1] L ≡ K → ∀f2. ⬇*[Ⓣ,f2] L ≡ K →
𝐔⦃f1⦄ → 𝐔⦃f2⦄ → f1 ≗ f2.
-#L #K #f1 #H elim H -L -K -f1
+#f1 #L #K #H elim H -f1 -L -K
[ #f1 #Hf1 #f2 #Hf2 elim (drops_inv_atom1 … Hf2) -Hf2
/3 width=1 by isid_inv_eq_repl/
-| #I #L #K #V #f1 #Hf1 #IH #f2 elim (pn_split f2) *
+| #f1 #I #L #K #V #Hf1 #IH #f2 elim (pn_split f2) *
#g2 #H2 #Hf2 #HU1 #HU2 destruct
[ elim (drops_inv_skip1 … Hf2) -IH -HU1 -Hf2 #Y2 #X2 #HY2 #_ #H destruct
lapply (drops_fwd_isid … HY2 ?) -HY2 /2 width=3 by isuni_inv_push/ -HU2
#H destruct elim (drops_inv_x_pair_xy … Hf1)
| /4 width=5 by drops_inv_drop1, isuni_inv_next, eq_next/
]
-| #I #L #K #V #W #f1 #Hf1 #_ #IH #f2 elim (pn_split f2) *
+| #f1 #I #L #K #V #W #Hf1 #_ #IH #f2 elim (pn_split f2) *
#g2 #H2 #Hf2 #HU1 #HU2 destruct
[ elim (drops_inv_skip1 … Hf2) -Hf2 #Y2 #X2 #HY2 #_ #H destruct -Hf1
/4 width=5 by isuni_fwd_push, eq_push/
(* Advanced properties ******************************************************)
(* Basic_2A1: includes: drop_mono *)
-lemma drops_mono: ∀L,L1,c1,f. ⬇*[c1, f] L ≡ L1 →
- ∀L2,c2. ⬇*[c2, f] L ≡ L2 → L1 = L2.
-#L #L1 #c1 #f lapply (isid_after_dx 𝐈𝐝 … f)
+lemma drops_mono: ∀b1,f,L,L1. ⬇*[b1, f] L ≡ L1 →
+ ∀b2,L2. ⬇*[b2, f] L ≡ L2 → L1 = L2.
+#b1 #f #L #L1 lapply (isid_after_dx 𝐈𝐝 … f)
/3 width=8 by drops_conf, drops_fwd_isid/
qed-.
(* Basic_2A1: includes: drop_conf_lt *)
-lemma drops_conf_skip1: ∀L,L2,c2,f. ⬇*[c2, f] L ≡ L2 →
- ∀I,K1,V1,c1,f1. ⬇*[c1, f1] L ≡ K1.ⓑ{I}V1 →
+lemma drops_conf_skip1: ∀b2,f,L,L2. ⬇*[b2, f] L ≡ L2 →
+ ∀b1,f1,I,K1,V1. ⬇*[b1, f1] L ≡ K1.ⓑ{I}V1 →
∀f2. f1 ⊚ ↑f2 ≡ f →
∃∃K2,V2. L2 = K2.ⓑ{I}V2 &
- ⬇*[c2, f2] K1 ≡ K2 & ⬆*[f2] V2 ≡ V1.
-#L #L2 #c2 #f #H2 #I #K1 #V1 #c1 #f1 #H1 #f2 #Hf lapply (drops_conf … H1 … H2 … Hf) -L -Hf
+ ⬇*[b2, f2] K1 ≡ K2 & ⬆*[f2] V2 ≡ V1.
+#b2 #f #L #L2 #H2 #b1 #f1 #I #K1 #V1 #H1 #f2 #Hf lapply (drops_conf … H1 … H2 … Hf) -L -Hf
#H elim (drops_inv_skip1 … H) -H /2 width=5 by ex3_2_intro/
qed-.
(* Basic_2A1: includes: drop_trans_lt *)
-lemma drops_trans_skip2: ∀L1,L,c1,f1. ⬇*[c1, f1] L1 ≡ L →
- ∀I,K2,V2,c2,f2. ⬇*[c2, f2] L ≡ K2.ⓑ{I}V2 →
+lemma drops_trans_skip2: ∀b1,f1,L1,L. ⬇*[b1, f1] L1 ≡ L →
+ ∀b2,f2,I,K2,V2. ⬇*[b2, f2] L ≡ K2.ⓑ{I}V2 →
∀f. f1 ⊚ f2 ≡ ↑f →
∃∃K1,V1. L1 = K1.ⓑ{I}V1 &
- ⬇*[c1∧c2, f] K1 ≡ K2 & ⬆*[f] V2 ≡ V1.
-#L1 #L #c1 #f1 #H1 #I #K2 #V2 #c2 #f2 #H2 #f #Hf
+ ⬇*[b1∧b2, f] K1 ≡ K2 & ⬆*[f] V2 ≡ V1.
+#b1 #f1 #L1 #L #H1 #b2 #f2 #I #K2 #V2 #H2 #f #Hf
lapply (drops_trans … H1 … H2 … Hf) -L -Hf
#H elim (drops_inv_skip2 … H) -H /2 width=5 by ex3_2_intro/
qed-.
(* Basic_2A1: includes: drops_conf_div *)
-lemma drops_conf_div_pair: ∀I1,I2,L,K,V1,V2,f1,f2.
+lemma drops_conf_div_pair: ∀f1,f2,I1,I2,L,K,V1,V2.
⬇*[Ⓣ,f1] L ≡ K.ⓑ{I1}V1 → ⬇*[Ⓣ,f2] L ≡ K.ⓑ{I2}V2 →
𝐔⦃f1⦄ → 𝐔⦃f2⦄ → ∧∧ f1 ≗ f2 & I1 = I2 & V1 = V2.
-#I1 #I2 #L #K #V1 #V2 #f1 #f2 #Hf1 #Hf2 #HU1 #HU2
+#f1 #f2 #I1 #I2 #L #K #V1 #V2 #Hf1 #Hf2 #HU1 #HU2
lapply (drops_isuni_fwd_drop2 … Hf1) // #H1
lapply (drops_isuni_fwd_drop2 … Hf2) // #H2
lapply (drops_conf_div … H1 … H2 ??) /2 width=3 by isuni_next/ -H1 -H2 -HU1 -HU2 #H
(* Forward lemmas with length for local environments ************************)
(* Basic_2A1: includes: drop_fwd_length_le4 *)
-lemma drops_fwd_length_le4: ∀L1,L2,c,f. ⬇*[c, f] L1 ≡ L2 → |L2| ≤ |L1|.
-#L1 #L2 #c #f #H elim H -L1 -L2 -f /2 width=1 by le_S, le_S_S/
+lemma drops_fwd_length_le4: ∀b,f,L1,L2. ⬇*[b, f] L1 ≡ L2 → |L2| ≤ |L1|.
+#b #f #L1 #L2 #H elim H -f -L1 -L2 /2 width=1 by le_S, le_S_S/
qed-.
(* Basic_2A1: includes: drop_fwd_length_eq1 *)
-theorem drops_fwd_length_eq1: ∀L1,K1,c1,c2,f. ⬇*[c1, f] L1 ≡ K1 →
- ∀L2,K2. ⬇*[c2, f] L2 ≡ K2 →
+theorem drops_fwd_length_eq1: ∀b1,b2,f,L1,K1. ⬇*[b1, f] L1 ≡ K1 →
+ ∀L2,K2. ⬇*[b2, f] L2 ≡ K2 →
|L1| = |L2| → |K1| = |K2|.
-#L1 #K1 #c1 #c2 #f #HLK1 elim HLK1 -L1 -K1 -f
+#b1 #b2 #f #L1 #K1 #HLK1 elim HLK1 -f -L1 -K1
[ #f #_ #L2 #K2 #HLK2 #H lapply (length_inv_zero_sn … H) -H
#H destruct elim (drops_inv_atom1 … HLK2) -HLK2 //
-| #I1 #L1 #K1 #V1 #f #_ #IH #X2 #K2 #HX #H elim (length_inv_succ_sn … H) -H
+| #f #I1 #L1 #K1 #V1 #_ #IH #X2 #K2 #HX #H elim (length_inv_succ_sn … H) -H
#I2 #L2 #V2 #H12 #H destruct lapply (drops_inv_drop1 … HX) -HX
#HLK2 @(IH … HLK2 H12) (**) (* auto fails *)
-| #I1 #L1 #K1 #V1 #V2 #f #_ #_ #IH #X2 #Y2 #HX #H elim (length_inv_succ_sn … H) -H
+| #f #I1 #L1 #K1 #V1 #V2 #_ #_ #IH #X2 #Y2 #HX #H elim (length_inv_succ_sn … H) -H
#I2 #L2 #V2 #H12 #H destruct elim (drops_inv_skip1 … HX) -HX
#K2 #W2 #HLK2 #_ #H destruct
lapply (IH … HLK2 H12) -f /2 width=1 by/ (**) (* full auto fails *)
(* forward lemmas with finite colength assignment ***************************)
-lemma drops_fwd_fcla: ∀L1,L2,f. ⬇*[Ⓣ, f] L1 ≡ L2 →
+lemma drops_fwd_fcla: ∀f,L1,L2. ⬇*[Ⓣ, f] L1 ≡ L2 →
∃∃n. 𝐂⦃f⦄ ≡ n & |L1| = |L2| + n.
-#L1 #L2 #f #H elim H -L1 -L2 -f
+#f #L1 #L2 #H elim H -f -L1 -L2
[ /4 width=3 by fcla_isid, ex2_intro/
-| #I #L1 #L2 #V #f #_ * /3 width=3 by fcla_next, ex2_intro, eq_f/
-| #I #L1 #L2 #V1 #V2 #f #_ #_ * /3 width=3 by fcla_push, ex2_intro/
+| #f #I #L1 #L2 #V #_ * /3 width=3 by fcla_next, ex2_intro, eq_f/
+| #f #I #L1 #L2 #V1 #V2 #_ #_ * /3 width=3 by fcla_push, ex2_intro/
]
qed-.
(* Basic_2A1: includes: drop_fwd_length *)
-lemma drops_fcla_fwd: ∀L1,L2,f,n. ⬇*[Ⓣ, f] L1 ≡ L2 → 𝐂⦃f⦄ ≡ n →
+lemma drops_fcla_fwd: ∀f,L1,L2,n. ⬇*[Ⓣ, f] L1 ≡ L2 → 𝐂⦃f⦄ ≡ n →
|L1| = |L2| + n.
-#l1 #l2 #f #n #Hf #Hn elim (drops_fwd_fcla … Hf) -Hf
+#f #l1 #l2 #n #Hf #Hn elim (drops_fwd_fcla … Hf) -Hf
#m #Hm #H <(fcla_mono … Hm … Hn) -f //
qed-.
-lemma drops_fwd_fcla_le2: ∀L1,L2,f. ⬇*[Ⓣ, f] L1 ≡ L2 →
+lemma drops_fwd_fcla_le2: ∀f,L1,L2. ⬇*[Ⓣ, f] L1 ≡ L2 →
∃∃n. 𝐂⦃f⦄ ≡ n & n ≤ |L1|.
-#L1 #L2 #f #H elim (drops_fwd_fcla … H) -H /2 width=3 by ex2_intro/
+#f #L1 #L2 #H elim (drops_fwd_fcla … H) -H /2 width=3 by ex2_intro/
qed-.
(* Basic_2A1: includes: drop_fwd_length_le2 *)
-lemma drops_fcla_fwd_le2: ∀L1,L2,f,n. ⬇*[Ⓣ, f] L1 ≡ L2 → 𝐂⦃f⦄ ≡ n →
+lemma drops_fcla_fwd_le2: ∀f,L1,L2,n. ⬇*[Ⓣ, f] L1 ≡ L2 → 𝐂⦃f⦄ ≡ n →
n ≤ |L1|.
-#L1 #L2 #f #n #H #Hn elim (drops_fwd_fcla_le2 … H) -H
+#f #L1 #L2 #n #H #Hn elim (drops_fwd_fcla_le2 … H) -H
#m #Hm #H <(fcla_mono … Hm … Hn) -f //
qed-.
-lemma drops_fwd_fcla_lt2: ∀L1,I2,K2,V2,f. ⬇*[Ⓣ, f] L1 ≡ K2.ⓑ{I2}V2 →
+lemma drops_fwd_fcla_lt2: ∀f,L1,I2,K2,V2. ⬇*[Ⓣ, f] L1 ≡ K2.ⓑ{I2}V2 →
∃∃n. 𝐂⦃f⦄ ≡ n & n < |L1|.
-#L1 #I2 #K2 #V2 #f #H elim (drops_fwd_fcla … H) -H
+#f #L1 #I2 #K2 #V2 #H elim (drops_fwd_fcla … H) -H
#n #Hf #H >H -L1 /3 width=3 by le_S_S, ex2_intro/
qed-.
(* Basic_2A1: includes: drop_fwd_length_lt2 *)
-lemma drops_fcla_fwd_lt2: ∀L1,I2,K2,V2,f,n.
+lemma drops_fcla_fwd_lt2: ∀f,L1,I2,K2,V2,n.
⬇*[Ⓣ, f] L1 ≡ K2.ⓑ{I2}V2 → 𝐂⦃f⦄ ≡ n →
n < |L1|.
-#L1 #I2 #K2 #V2 #f #n #H #Hn elim (drops_fwd_fcla_lt2 … H) -H
+#f #L1 #I2 #K2 #V2 #n #H #Hn elim (drops_fwd_fcla_lt2 … H) -H
#m #Hm #H <(fcla_mono … Hm … Hn) -f //
qed-.
(* Basic_2A1: includes: drop_fwd_length_lt4 *)
-lemma drops_fcla_fwd_lt4: ∀L1,L2,f,n. ⬇*[Ⓣ, f] L1 ≡ L2 → 𝐂⦃f⦄ ≡ n → 0 < n →
+lemma drops_fcla_fwd_lt4: ∀f,L1,L2,n. ⬇*[Ⓣ, f] L1 ≡ L2 → 𝐂⦃f⦄ ≡ n → 0 < n →
|L2| < |L1|.
-#L1 #L2 #f #n #H #Hf #Hn lapply (drops_fcla_fwd … H Hf) -f
+#f #L1 #L2 #n #H #Hf #Hn lapply (drops_fcla_fwd … H Hf) -f
/2 width=1 by lt_minus_to_plus_r/ qed-.
(* Basic_2A1: includes: drop_inv_length_eq *)
-lemma drops_inv_length_eq: ∀L1,L2,f. ⬇*[Ⓣ, f] L1 ≡ L2 → |L1| = |L2| → 𝐈⦃f⦄.
-#L1 #L2 #f #H #HL12 elim (drops_fwd_fcla … H) -H
+lemma drops_inv_length_eq: ∀f,L1,L2. ⬇*[Ⓣ, f] L1 ≡ L2 → |L1| = |L2| → 𝐈⦃f⦄.
+#f #L1 #L2 #H #HL12 elim (drops_fwd_fcla … H) -H
#n #Hn <HL12 -L2 #H lapply (discr_plus_x_xy … H) -H
/2 width=3 by fcla_inv_xp/
qed-.
(* Basic_2A1: includes: drop_fwd_length_eq2 *)
-theorem drops_fwd_length_eq2: ∀L1,L2,K1,K2,f. ⬇*[Ⓣ, f] L1 ≡ K1 → ⬇*[Ⓣ, f] L2 ≡ K2 →
+theorem drops_fwd_length_eq2: ∀f,L1,L2,K1,K2. ⬇*[Ⓣ, f] L1 ≡ K1 → ⬇*[Ⓣ, f] L2 ≡ K2 →
|K1| = |K2| → |L1| = |L2|.
-#L1 #L2 #K1 #K2 #f #HLK1 #HLK2 #HL12
+#f #L1 #L2 #K1 #K2 #HLK1 #HLK2 #HL12
elim (drops_fwd_fcla … HLK1) -HLK1 #n1 #Hn1 #H1 >H1 -L1
elim (drops_fwd_fcla … HLK2) -HLK2 #n2 #Hn2 #H2 >H2 -L2
<(fcla_mono … Hn2 … Hn1) -f //
qed-.
-theorem drops_conf_div: ∀L1,L2,f1,f2. ⬇*[Ⓣ, f1] L1 ≡ L2 → ⬇*[Ⓣ, f2] L1 ≡ L2 →
+theorem drops_conf_div: ∀f1,f2,L1,L2. ⬇*[Ⓣ, f1] L1 ≡ L2 → ⬇*[Ⓣ, f2] L1 ≡ L2 →
∃∃n. 𝐂⦃f1⦄ ≡ n & 𝐂⦃f2⦄ ≡ n.
-#L1 #L2 #f1 #f2 #H1 #H2
+#f1 #f2 #L1 #L2 #H1 #H2
elim (drops_fwd_fcla … H1) -H1 #n1 #Hf1 #H1
elim (drops_fwd_fcla … H2) -H2 #n2 #Hf2 >H1 -L1 #H
lapply (injective_plus_r … H) -L2 #H destruct /2 width=3 by ex2_intro/
qed-.
-theorem drops_conf_div_fcla: ∀L1,L2,f1,f2,n1,n2.
+theorem drops_conf_div_fcla: ∀f1,f2,L1,L2,n1,n2.
⬇*[Ⓣ, f1] L1 ≡ L2 → ⬇*[Ⓣ, f2] L1 ≡ L2 → 𝐂⦃f1⦄ ≡ n1 → 𝐂⦃f2⦄ ≡ n2 →
n1 = n2.
-#L1 #L2 #f1 #f2 #n1 #n2 #Hf1 #Hf2 #Hn1 #Hn2
+#f1 #f2 #L1 #L2 #n1 #n2 #Hf1 #Hf2 #Hn1 #Hn2
lapply (drops_fcla_fwd … Hf1 Hn1) -f1 #H1
lapply (drops_fcla_fwd … Hf2 Hn2) -f2 >H1 -L1
/2 width=1 by injective_plus_r/
(* Basic_2A1: includes: lpx_sn_deliftable_dropable *)
lemma lexs_deliftable_dropable: ∀RN,RP. d_deliftable2_sn RN → d_deliftable2_sn RP →
dropable_sn (lexs RN RP).
-#RN #RP #HN #HP #L1 #K1 #c #f #H elim H -L1 -K1 -f
+#RN #RP #HN #HP #b #f #L1 #K1 #H elim H -f -L1 -K1
[ #f #Hf #X #f2 #H #f1 #Hf2 >(lexs_inv_atom1 … H) -X
/4 width=3 by lexs_atom, drops_atom, ex2_intro/
-| #I #L1 #K1 #V1 #f #_ #IH #X #f2 #H #f1 #Hf2 elim (after_inv_nxx … Hf2) -Hf2 [2,3: // ]
+| #f #I #L1 #K1 #V1 #_ #IH #X #f2 #H #f1 #Hf2 elim (after_inv_nxx … Hf2) -Hf2 [2,3: // ]
#g2 #Hg2 #H2 destruct elim (lexs_inv_next1 … H) -H
#L2 #V2 #HL12 #HV12 #H destruct elim (IH … HL12 … Hg2) -g2
/3 width=3 by drops_drop, ex2_intro/
-| #I #L1 #K1 #V1 #W1 #f #HLK #HWV #IH #X #f2 #H #f1 #Hf2 elim (after_inv_pxx … Hf2) -Hf2 [1,3:* |*: // ]
+| #f #I #L1 #K1 #V1 #W1 #HLK #HWV #IH #X #f2 #H #f1 #Hf2 elim (after_inv_pxx … Hf2) -Hf2 [1,3:* |*: // ]
#g1 #g2 #Hg2 #H1 #H2 destruct
[ elim (lexs_inv_push1 … H) | elim (lexs_inv_next1 … H) ] -H
#L2 #V2 #HL12 #HV12 #H destruct elim (IH … HL12 … Hg2) -g2
(* Basic_2A1: includes: lpx_sn_liftable_dedropable *)
lemma lexs_liftable_dedropable: ∀RN,RP. (∀L. reflexive ? (RN L)) → (∀L. reflexive ? (RP L)) →
d_liftable2 RN → d_liftable2 RP → dedropable_sn (lexs RN RP).
-#RN #RP #H1RN #H1RP #H2RN #H2RP #L1 #K1 #c #f #H elim H -L1 -K1 -f
+#RN #RP #H1RN #H1RP #H2RN #H2RP #b #f #L1 #K1 #H elim H -f -L1 -K1
[ #f #Hf #X #f1 #H #f2 #Hf2 >(lexs_inv_atom1 … H) -X
/4 width=4 by drops_atom, lexs_atom, ex3_intro/
-| #I #L1 #K1 #V1 #f #_ #IHLK1 #K2 #f1 #HK12 #f2 #Hf2
+| #f #I #L1 #K1 #V1 #_ #IHLK1 #K2 #f1 #HK12 #f2 #Hf2
elim (after_inv_nxx … Hf2) -Hf2 [2,3: // ] #g2 #Hg2 #H destruct
elim (IHLK1 … HK12 … Hg2) -K1
/3 width=6 by drops_drop, lexs_next, ex3_intro/
-| #I #L1 #K1 #V1 #W1 #f #HLK1 #HWV1 #IHLK1 #X #f1 #H #f2 #Hf2
+| #f #I #L1 #K1 #V1 #W1 #HLK1 #HWV1 #IHLK1 #X #f1 #H #f2 #Hf2
elim (after_inv_pxx … Hf2) -Hf2 [1,3: * |*: // ] #g1 #g2 #Hg2 #H1 #H2 destruct
[ elim (lexs_inv_push1 … H) | elim (lexs_inv_next1 … H) ] -H #K2 #W2 #HK12 #HW12 #H destruct
[ elim (H2RP … HW12 … HLK1 … HWV1) | elim (H2RN … HW12 … HLK1 … HWV1) ] -W1
]
qed-.
-fact lexs_dropable_aux: ∀RN,RP,L2,K2,c,f. ⬇*[c, f] L2 ≡ K2 → 𝐔⦃f⦄ →
- ∀L1,f2. L1 ⦻*[RN, RP, f2] L2 → ∀f1. f ⊚ f1 ≡ f2 →
- ∃∃K1. ⬇*[c, f] L1 ≡ K1 & K1 ⦻*[RN, RP, f1] K2.
-#RN #RP #L2 #K2 #c #f #H elim H -L2 -K2 -f
-[ #f #Hf #_ #X #f2 #H #f1 #Hf2 lapply (lexs_inv_atom2 … H) -H
+fact lexs_dropable_aux: ∀RN,RP,b,f,L2,K2. ⬇*[b, f] L2 ≡ K2 → 𝐔⦃f⦄ →
+ ∀f2,L1. L1 ⦻*[RN, RP, f2] L2 → ∀f1. f ⊚ f1 ≡ f2 →
+ ∃∃K1. ⬇*[b, f] L1 ≡ K1 & K1 ⦻*[RN, RP, f1] K2.
+#RN #RP #b #f #L2 #K2 #H elim H -f -L2 -K2
+[ #f #Hf #_ #f2 #X #H #f1 #Hf2 lapply (lexs_inv_atom2 … H) -H
#H destruct /4 width=3 by lexs_atom, drops_atom, ex2_intro/
-| #I #L2 #K2 #V2 #f #_ #IH #Hf #X #f2 #HX #f1 #Hf2
+| #f #I #L2 #K2 #V2 #_ #IH #Hf #f2 #X #HX #f1 #Hf2
elim (after_inv_nxx … Hf2) -Hf2 [2,3: // ] #g2 #Hg2 #H destruct
elim (lexs_inv_next2 … HX) -HX #L1 #V1 #HL12 #HV12 #H destruct
elim (IH … HL12 … Hg2) -L2 -V2 -g2
/3 width=3 by drops_drop, isuni_inv_next, ex2_intro/
-| #I #L2 #K2 #V2 #W2 #f #_ #HWV2 #IH #Hf #X #f2 #HX #f1 #Hf2
+| #f #I #L2 #K2 #V2 #W2 #_ #HWV2 #IH #Hf #f2 #X #HX #f1 #Hf2
elim (after_inv_pxx … Hf2) -Hf2 [1,3: * |*: // ] #g1 #g2 #Hg2 #H1 #H2 destruct
[ elim (lexs_inv_push2 … HX) | elim (lexs_inv_next2 … HX) ] -HX #L1 #V1 #HL12 #HV12 #H destruct
elim (IH … HL12 … Hg2) -L2 -g2 /2 width=3 by isuni_fwd_push/ #K1 #HLK1 #HK12
(* Basic_2A1: includes: lpx_sn_drop_conf *)
lemma lexs_drops_conf_next: ∀RN,RP. d_deliftable2_sn RN → d_deliftable2_sn RP →
- ∀L1,L2,f2. L1 ⦻*[RN, RP, f2] L2 →
- ∀I,K1,V1,c,f. ⬇*[c,f] L1 ≡ K1.ⓑ{I}V1 →
+ ∀f2,L1,L2. L1 ⦻*[RN, RP, f2] L2 →
+ ∀b,f,I,K1,V1. ⬇*[b,f] L1 ≡ K1.ⓑ{I}V1 →
∀f1. f ⊚ ⫯f1 ≡ f2 →
- ∃∃K2,V2. ⬇*[c,f] L2 ≡ K2.ⓑ{I}V2 & K1 ⦻*[RN, RP, f1] K2 & RN K1 V1 V2.
-#RN #RP #HRN #HRP #L1 #L2 #f2 #HL12 #I #K1 #V1 #c #f #HLK1 #f1 #Hf2
+ ∃∃K2,V2. ⬇*[b,f] L2 ≡ K2.ⓑ{I}V2 & K1 ⦻*[RN, RP, f1] K2 & RN K1 V1 V2.
+#RN #RP #HRN #HRP #f2 #L1 #L2 #HL12 #b #f #I #K1 #V1 #HLK1 #f1 #Hf2
elim (lexs_deliftable_dropable … HRN HRP … HLK1 … HL12 … Hf2) -L1 -f2 -HRN -HRP
#X #HX #HLK2 elim (lexs_inv_next1 … HX) -HX
#K2 #V2 #HK12 #HV12 #H destruct /2 width=5 by ex3_2_intro/
qed-.
lemma lexs_drops_conf_push: ∀RN,RP. d_deliftable2_sn RN → d_deliftable2_sn RP →
- ∀L1,L2,f2. L1 ⦻*[RN, RP, f2] L2 →
- ∀I,K1,V1,c,f. ⬇*[c,f] L1 ≡ K1.ⓑ{I}V1 →
+ ∀f2,L1,L2. L1 ⦻*[RN, RP, f2] L2 →
+ ∀b,f,I,K1,V1. ⬇*[b,f] L1 ≡ K1.ⓑ{I}V1 →
∀f1. f ⊚ ↑f1 ≡ f2 →
- ∃∃K2,V2. ⬇*[c,f] L2 ≡ K2.ⓑ{I}V2 & K1 ⦻*[RN, RP, f1] K2 & RP K1 V1 V2.
-#RN #RP #HRN #HRP #L1 #L2 #f2 #HL12 #I #K1 #V1 #c #f #HLK1 #f1 #Hf2
+ ∃∃K2,V2. ⬇*[b,f] L2 ≡ K2.ⓑ{I}V2 & K1 ⦻*[RN, RP, f1] K2 & RP K1 V1 V2.
+#RN #RP #HRN #HRP #f2 #L1 #L2 #HL12 #b #f #I #K1 #V1 #HLK1 #f1 #Hf2
elim (lexs_deliftable_dropable … HRN HRP … HLK1 … HL12 … Hf2) -L1 -f2 -HRN -HRP
#X #HX #HLK2 elim (lexs_inv_push1 … HX) -HX
#K2 #V2 #HK12 #HV12 #H destruct /2 width=5 by ex3_2_intro/
qed-.
(* Basic_2A1: includes: lpx_sn_drop_trans *)
-lemma lexs_drops_trans_next: ∀RN,RP,L1,L2,f2. L1 ⦻*[RN, RP, f2] L2 →
- ∀I,K2,V2,c,f. ⬇*[c,f] L2 ≡ K2.ⓑ{I}V2 → 𝐔⦃f⦄ →
+lemma lexs_drops_trans_next: ∀RN,RP,f2,L1,L2. L1 ⦻*[RN, RP, f2] L2 →
+ ∀b,f,I,K2,V2. ⬇*[b,f] L2 ≡ K2.ⓑ{I}V2 → 𝐔⦃f⦄ →
∀f1. f ⊚ ⫯f1 ≡ f2 →
- ∃∃K1,V1. ⬇*[c,f] L1 ≡ K1.ⓑ{I}V1 & K1 ⦻*[RN, RP, f1] K2 & RN K1 V1 V2.
-#RN #RP #L1 #L2 #f2 #HL12 #I #K1 #V1 #c #f #HLK1 #Hf #f1 #Hf2
+ ∃∃K1,V1. ⬇*[b,f] L1 ≡ K1.ⓑ{I}V1 & K1 ⦻*[RN, RP, f1] K2 & RN K1 V1 V2.
+#RN #RP #f2 #L1 #L2 #HL12 #b #f #I #K1 #V1 #HLK1 #Hf #f1 #Hf2
elim (lexs_dropable … HL12 … HLK1 … Hf … Hf2) -L2 -f2 -Hf
#X #HLK1 #HX elim (lexs_inv_next2 … HX) -HX
#K1 #V1 #HK12 #HV12 #H destruct /2 width=5 by ex3_2_intro/
qed-.
-lemma lexs_drops_trans_push: ∀RN,RP,L1,L2,f2. L1 ⦻*[RN, RP, f2] L2 →
- ∀I,K2,V2,c,f. ⬇*[c,f] L2 ≡ K2.ⓑ{I}V2 → 𝐔⦃f⦄ →
+lemma lexs_drops_trans_push: ∀RN,RP,f2,L1,L2. L1 ⦻*[RN, RP, f2] L2 →
+ ∀b,f,I,K2,V2. ⬇*[b,f] L2 ≡ K2.ⓑ{I}V2 → 𝐔⦃f⦄ →
∀f1. f ⊚ ↑f1 ≡ f2 →
- ∃∃K1,V1. ⬇*[c,f] L1 ≡ K1.ⓑ{I}V1 & K1 ⦻*[RN, RP, f1] K2 & RP K1 V1 V2.
-#RN #RP #L1 #L2 #f2 #HL12 #I #K1 #V1 #c #f #HLK1 #Hf #f1 #Hf2
+ ∃∃K1,V1. ⬇*[b,f] L1 ≡ K1.ⓑ{I}V1 & K1 ⦻*[RN, RP, f1] K2 & RP K1 V1 V2.
+#RN #RP #f2 #L1 #L2 #HL12 #b #f #I #K1 #V1 #HLK1 #Hf #f1 #Hf2
elim (lexs_dropable … HL12 … HLK1 … Hf … Hf2) -L2 -f2 -Hf
#X #HLK1 #HX elim (lexs_inv_push2 … HX) -HX
#K1 #V1 #HK12 #HV12 #H destruct /2 width=5 by ex3_2_intro/
lemma drops_lexs_trans_next: ∀RN,RP. (∀L. reflexive ? (RN L)) → (∀L. reflexive ? (RP L)) →
d_liftable2 RN → d_liftable2 RP →
- ∀K1,K2,f1. K1 ⦻*[RN, RP, f1] K2 →
- ∀I,L1,V1,c,f. ⬇*[c,f] L1.ⓑ{I}V1 ≡ K1 →
+ ∀f1,K1,K2. K1 ⦻*[RN, RP, f1] K2 →
+ ∀b,f,I,L1,V1. ⬇*[b,f] L1.ⓑ{I}V1 ≡ K1 →
∀f2. f ⊚ f1 ≡ ⫯f2 →
- ∃∃L2,V2. ⬇*[c,f] L2.ⓑ{I}V2 ≡ K2 & L1 ⦻*[RN, RP, f2] L2 & RN L1 V1 V2 & L1.ⓑ{I}V1≡[f]L2.ⓑ{I}V2.
-#RN #RP #H1RN #H1RP #H2RN #H2RP #K1 #K2 #f1 #HK12 #I #L1 #V1 #c #f #HLK1 #f2 #Hf2
+ ∃∃L2,V2. ⬇*[b,f] L2.ⓑ{I}V2 ≡ K2 & L1 ⦻*[RN, RP, f2] L2 & RN L1 V1 V2 & L1.ⓑ{I}V1≡[f]L2.ⓑ{I}V2.
+#RN #RP #H1RN #H1RP #H2RN #H2RP #f1 #K1 #K2 #HK12 #b #f #I #L1 #V1 #HLK1 #f2 #Hf2
elim (lexs_liftable_dedropable … H1RN H1RP H2RN H2RP … HLK1 … HK12 … Hf2) -K1 -f1 -H1RN -H1RP -H2RN -H2RP
#X #HX #HLK2 #H1L12 elim (lexs_inv_next1 … HX) -HX
#L2 #V2 #H2L12 #HV12 #H destruct /2 width=6 by ex4_2_intro/
lemma drops_lexs_trans_push: ∀RN,RP. (∀L. reflexive ? (RN L)) → (∀L. reflexive ? (RP L)) →
d_liftable2 RN → d_liftable2 RP →
- ∀K1,K2,f1. K1 ⦻*[RN, RP, f1] K2 →
- ∀I,L1,V1,c,f. ⬇*[c,f] L1.ⓑ{I}V1 ≡ K1 →
+ ∀f1,K1,K2. K1 ⦻*[RN, RP, f1] K2 →
+ ∀b,f,I,L1,V1. ⬇*[b,f] L1.ⓑ{I}V1 ≡ K1 →
∀f2. f ⊚ f1 ≡ ↑f2 →
- ∃∃L2,V2. ⬇*[c,f] L2.ⓑ{I}V2 ≡ K2 & L1 ⦻*[RN, RP, f2] L2 & RP L1 V1 V2 & L1.ⓑ{I}V1≡[f]L2.ⓑ{I}V2.
-#RN #RP #H1RN #H1RP #H2RN #H2RP #K1 #K2 #f1 #HK12 #I #L1 #V1 #c #f #HLK1 #f2 #Hf2
+ ∃∃L2,V2. ⬇*[b,f] L2.ⓑ{I}V2 ≡ K2 & L1 ⦻*[RN, RP, f2] L2 & RP L1 V1 V2 & L1.ⓑ{I}V1≡[f]L2.ⓑ{I}V2.
+#RN #RP #H1RN #H1RP #H2RN #H2RP #f1 #K1 #K2 #HK12 #b #f #I #L1 #V1 #HLK1 #f2 #Hf2
elim (lexs_liftable_dedropable … H1RN H1RP H2RN H2RP … HLK1 … HK12 … Hf2) -K1 -f1 -H1RN -H1RP -H2RN -H2RP
#X #HX #HLK2 #H1L12 elim (lexs_inv_push1 … HX) -HX
#L2 #V2 #H2L12 #HV12 #H destruct /2 width=6 by ex4_2_intro/
@lexs_dropable qed-.
(* Basic_2A1: includes: lreq_drop_trans_be *)
-lemma lreq_drops_trans_next: ∀L1,L2,f2. L1 ≡[f2] L2 →
- ∀I,K2,V,c,f. ⬇*[c,f] L2 ≡ K2.ⓑ{I}V → 𝐔⦃f⦄ →
+lemma lreq_drops_trans_next: ∀f2,L1,L2. L1 ≡[f2] L2 →
+ ∀b,f,I,K2,V. ⬇*[b,f] L2 ≡ K2.ⓑ{I}V → 𝐔⦃f⦄ →
∀f1. f ⊚ ⫯f1 ≡ f2 →
- ∃∃K1. ⬇*[c,f] L1 ≡ K1.ⓑ{I}V & K1 ≡[f1] K2.
-#L1 #L2 #f2 #HL12 #I #K1 #V #c #f #HLK1 #Hf #f1 #Hf2
-elim (lexs_drops_trans_next … HL12 … HLK1 Hf … Hf2) -L2 -f2 -Hf
+ ∃∃K1. ⬇*[b,f] L1 ≡ K1.ⓑ{I}V & K1 ≡[f1] K2.
+#f2 #L1 #L2 #HL12 #b #f #I #K1 #V #HLK1 #Hf #f1 #Hf2
+elim (lexs_drops_trans_next … HL12 … HLK1 Hf … Hf2) -f2 -L2 -Hf
/2 width=3 by ex2_intro/
qed-.
(* Basic_2A1: includes: lreq_drop_conf_be *)
-lemma lreq_drops_conf_next: ∀L1,L2,f2. L1 ≡[f2] L2 →
- ∀I,K1,V,c,f. ⬇*[c,f] L1 ≡ K1.ⓑ{I}V → 𝐔⦃f⦄ →
+lemma lreq_drops_conf_next: ∀f2,L1,L2. L1 ≡[f2] L2 →
+ ∀b,f,I,K1,V. ⬇*[b,f] L1 ≡ K1.ⓑ{I}V → 𝐔⦃f⦄ →
∀f1. f ⊚ ⫯f1 ≡ f2 →
- ∃∃K2. ⬇*[c,f] L2 ≡ K2.ⓑ{I}V & K1 ≡[f1] K2.
-#L1 #L2 #f2 #HL12 #I #K1 #V #c #f #HLK1 #Hf #f1 #Hf2
-elim (lreq_drops_trans_next … (lreq_sym … HL12) … HLK1 … Hf2) // -L1 -f2 -Hf
+ ∃∃K2. ⬇*[b,f] L2 ≡ K2.ⓑ{I}V & K1 ≡[f1] K2.
+#f2 #L1 #L2 #HL12 #b #f #I #K1 #V #HLK1 #Hf #f1 #Hf2
+elim (lreq_drops_trans_next … (lreq_sym … HL12) … HLK1 … Hf2) // -f2 -L1 -Hf
/3 width=3 by lreq_sym, ex2_intro/
qed-.
-lemma drops_lreq_trans_next: ∀K1,K2,f1. K1 ≡[f1] K2 →
- ∀I,L1,V,c,f. ⬇*[c,f] L1.ⓑ{I}V ≡ K1 →
+lemma drops_lreq_trans_next: ∀f1,K1,K2. K1 ≡[f1] K2 →
+ ∀b,f,I,L1,V. ⬇*[b,f] L1.ⓑ{I}V ≡ K1 →
∀f2. f ⊚ f1 ≡ ⫯f2 →
- ∃∃L2. ⬇*[c,f] L2.ⓑ{I}V ≡ K2 & L1 ≡[f2] L2 & L1.ⓑ{I}V≡[f]L2.ⓑ{I}V.
-#K1 #K2 #f1 #HK12 #I #L1 #V #c #f #HLK1 #f2 #Hf2
-elim (drops_lexs_trans_next … HK12 … HLK1 … Hf2) -K1 -f1
+ ∃∃L2. ⬇*[b,f] L2.ⓑ{I}V ≡ K2 & L1 ≡[f2] L2 & L1.ⓑ{I}V≡[f]L2.ⓑ{I}V.
+#f1 #K1 #K2 #HK12 #b #f #I #L1 #V #HLK1 #f2 #Hf2
+elim (drops_lexs_trans_next … HK12 … HLK1 … Hf2) -f1 -K1
/2 width=6 by cfull_lift, ceq_lift, cfull_refl, ceq_refl, ex3_intro/
qed-.
(* Basic_2A1: was: d_liftable_LTC *)
lemma d2_liftable_LTC: ∀R. d_liftable2 R → d_liftable2 (LTC … R).
#R #HR #K #T1 #T2 #H elim H -T2
-[ #T2 #HT12 #L #c #f #HLK #U1 #HTU1
+[ #T2 #HT12 #b #f #L #HLK #U1 #HTU1
elim (HR … HT12 … HLK … HTU1) /3 width=3 by inj, ex2_intro/
-| #T #T2 #_ #HT2 #IHT1 #L #c #f #HLK #U1 #HTU1
+| #T #T2 #_ #HT2 #IHT1 #b #f #L #HLK #U1 #HTU1
elim (IHT1 … HLK … HTU1) -T1 #U #HTU #HU1
elim (HR … HT2 … HLK … HTU) -HR -K -T /3 width=5 by step, ex2_intro/
]
(* Basic_2A1: was: d_deliftable_sn_LTC *)
lemma d2_deliftable_sn_LTC: ∀R. d_deliftable2_sn R → d_deliftable2_sn (LTC … R).
#R #HR #L #U1 #U2 #H elim H -U2
-[ #U2 #HU12 #K #c #f #HLK #T1 #HTU1
+[ #U2 #HU12 #b #f #K #HLK #T1 #HTU1
elim (HR … HU12 … HLK … HTU1) -HR -L -U1 /3 width=3 by inj, ex2_intro/
-| #U #U2 #_ #HU2 #IHU1 #K #c #f #HLK #T1 #HTU1
+| #U #U2 #_ #HU2 #IHU1 #b #f #K #HLK #T1 #HTU1
elim (IHU1 … HLK … HTU1) -IHU1 -U1 #T #HTU #HT1
elim (HR … HU2 … HLK … HTU) -HR -L -U /3 width=5 by step, ex2_intro/
]
qed-.
lemma dropable_sn_TC: ∀R. dropable_sn R → dropable_sn (LTC … R).
-#R #HR #L1 #K1 #c #f #HLK1 #L2 #f2 #H elim H -L2
-[ #L2 #HL12 #f1 #H elim (HR … HLK1 … HL12 … H) -HR -L1 -f2
+#R #HR #b #f #L1 #K1 #HLK1 #f2 #L2 #H elim H -L2
+[ #L2 #HL12 #f1 #H elim (HR … HLK1 … HL12 … H) -HR -f2 -L1
/3 width=3 by inj, ex2_intro/
| #L #L2 #_ #HL2 #IH #f1 #H elim (IH … H) -IH
- #K #HK1 #HLK elim (HR … HLK … HL2 … H) -HR -L -f2
+ #K #HK1 #HLK elim (HR … HLK … HL2 … H) -HR -f2 -L
/3 width=3 by step, ex2_intro/
]
qed-.
(* Basic_2A1: was: d_liftable_llstar *)
lemma d2_liftable_llstar: ∀R. d_liftable2 R → ∀d. d_liftable2 (llstar … R d).
#R #HR #d #K #T1 #T2 #H @(lstar_ind_r … d T2 H) -d -T2
-[ #L #c #f #_ #U1 #HTU1 -HR -K -c /2 width=3 by ex2_intro/
-| #d #T #T2 #_ #HT2 #IHT1 #L #c #f #HLK #U1 #HTU1
+[ #b #f #L #_ #U1 #HTU1 -HR -b -K /2 width=3 by ex2_intro/
+| #l #T #T2 #_ #HT2 #IHT1 #b #f #L #HLK #U1 #HTU1
elim (IHT1 … HLK … HTU1) -T1 #U #HTU #HU1
elim (HR … HT2 … HLK … HTU) -T /3 width=5 by lstar_dx, ex2_intro/
]
(* Basic_2A1: was: d_deliftable_sn_llstar *)
lemma d2_deliftable_sn_llstar: ∀R. d_deliftable2_sn R →
∀d. d_deliftable2_sn (llstar … R d).
-#R #HR #d #L #U1 #U2 #H @(lstar_ind_r … d U2 H) -d -U2
+#R #HR #l #L #U1 #U2 #H @(lstar_ind_r … l U2 H) -l -U2
[ /2 width=3 by lstar_O, ex2_intro/
-| #d #U #U2 #_ #HU2 #IHU1 #K #c #f #HLK #T1 #HTU1
+| #l #U #U2 #_ #HU2 #IHU1 #b #f #K #HLK #T1 #HTU1
elim (IHU1 … HLK … HTU1) -IHU1 -U1 #T #HTU #HT1
elim (HR … HU2 … HLK … HTU) -HR -L -U /3 width=5 by lstar_dx, ex2_intro/
]
qed-.
lemma dropable_dx_TC: ∀R. dropable_dx R → dropable_dx (LTC … R).
-#R #HR #L1 #L2 #f2 #H elim H -L2
-[ #L2 #HL12 #K2 #c #f #HLK2 #Hf #f1 #Hf2 elim (HR … HL12 … HLK2 … Hf … Hf2) -HR -L2 -f2 -Hf
+#R #HR #f2 #L1 #L2 #H elim H -L2
+[ #L2 #HL12 #b #f #K2 #HLK2 #Hf #f1 #Hf2 elim (HR … HL12 … HLK2 … Hf … Hf2) -HR -Hf -f2 -L2
/3 width=3 by inj, ex2_intro/
-| #L #L2 #_ #HL2 #IHL1 #K2 #c #f #HLK2 #Hf #f1 #Hf2 elim (HR … HL2 … HLK2 … Hf … Hf2) -HR -L2
- #K #HLK #HK2 elim (IHL1 … HLK … Hf … Hf2) -L -f2 -Hf
+| #L #L2 #_ #HL2 #IHL1 #b #f #K2 #HLK2 #Hf #f1 #Hf2 elim (HR … HL2 … HLK2 … Hf … Hf2) -HR -L2
+ #K #HLK #HK2 elim (IHL1 … HLK … Hf … Hf2) -Hf -f2 -L
/3 width=5 by step, ex2_intro/
]
qed-.
lemma dedropable_sn_TC: ∀R. dedropable_sn R → dedropable_sn (LTC … R).
-#R #HR #L1 #K1 #c #f #HLK1 #K2 #f1 #H elim H -K2
-[ #K2 #HK12 #f2 #Hf elim (HR … HLK1 … HK12 … Hf) -HR -K1 -f1
+#R #HR #b #f #L1 #K1 #HLK1 #f1 #K2 #H elim H -K2
+[ #K2 #HK12 #f2 #Hf elim (HR … HLK1 … HK12 … Hf) -HR -f1 -K1
/3 width=4 by inj, ex3_intro/
| #K #K2 #_ #HK2 #IH #f2 #Hf elim (IH … Hf) -IH -K1
- #L #H1L1 #HLK #H2L1 elim (HR … HLK … HK2 … Hf) -HR -K -f1
+ #L #H1L1 #HLK #H2L1 elim (HR … HLK … HK2 … Hf) -HR -f1 -K
/3 width=6 by lreq_trans, step, ex3_intro/
]
qed-.
(* GENERIC SLICING FOR LOCAL ENVIRONMENTS ***********************************)
definition d_liftable1_all: relation2 lenv term → predicate bool ≝
- λR,c. ∀L,K,f. ⬇*[c, f] L ≡ K →
+ λR,b. ∀f,L,K. ⬇*[b, f] L ≡ K →
∀Ts,Us. ⬆*[f] Ts ≡ Us →
all … (R K) Ts → all … (R L) Us.
(* Properties with generic relocation for term vectors **********************)
(* Basic_2A1: was: d1_liftables_liftables_all *)
-lemma d1_liftable_liftable_all: ∀R,c. d_liftable1 R c → d_liftable1_all R c.
-#R #c #HR #L #K #f #HLK #Ts #Us #H elim H -Ts -Us normalize //
+lemma d1_liftable_liftable_all: ∀R,b. d_liftable1 R b → d_liftable1_all R b.
+#R #b #HR #f #L #K #HLK #Ts #Us #H elim H -Ts -Us normalize //
#Ts #Us #T #U #HTU #_ #IHTUs * /3 width=7 by conj/
qed.
(* Forward lemmas with weight for local environments ************************)
(* Basic_2A1: includes: drop_fwd_lw *)
-lemma drops_fwd_lw: ∀L1,L2,c,f. ⬇*[c, f] L1 ≡ L2 → ♯{L2} ≤ ♯{L1}.
-#L1 #L2 #c #f #H elim H -L1 -L2 -f //
+lemma drops_fwd_lw: ∀b,f,L1,L2. ⬇*[b, f] L1 ≡ L2 → ♯{L2} ≤ ♯{L1}.
+#b #f #L1 #L2 #H elim H -f -L1 -L2 //
[ /2 width=3 by transitive_le/
-| #I #L1 #L2 #V1 #V2 #f #_ #HV21 #IHL12 normalize
+| #f #I #L1 #L2 #V1 #V2 #_ #HV21 #IHL12 normalize
>(lifts_fwd_tw … HV21) -HV21 /2 width=1 by monotonic_le_plus_l/
]
qed-.
(* Basic_2A1: includes: drop_fwd_lw_lt *)
-lemma drops_fwd_lw_lt: ∀L1,L2,f. ⬇*[Ⓣ, f] L1 ≡ L2 →
+lemma drops_fwd_lw_lt: ∀f,L1,L2. ⬇*[Ⓣ, f] L1 ≡ L2 →
(𝐈⦃f⦄ → ⊥) → ♯{L2} < ♯{L1}.
-#L1 #L2 #f #H elim H -L1 -L2 -f
+#f #L1 #L2 #H elim H -f -L1 -L2
[ #f #Hf #Hnf elim Hnf -Hnf /2 width=1 by/
| /3 width=3 by drops_fwd_lw, le_to_lt_to_lt/
-| #I #L1 #L2 #V1 #V2 #f #_ #HV21 #IHL12 #H normalize in ⊢ (?%%); -I
+| #f #I #L1 #L2 #V1 #V2 #_ #HV21 #IHL12 #H normalize in ⊢ (?%%); -I
>(lifts_fwd_tw … HV21) -V2 /5 width=3 by isid_push, monotonic_lt_plus_l/
]
qed-.
(* Forward lemmas with restricted weight for closures ***********************)
(* Basic_2A1: includes: drop_fwd_rfw *)
-lemma drops_pair2_fwd_rfw: ∀I,L,K,V,c,f. ⬇*[c, f] L ≡ K.ⓑ{I}V → ∀T. ♯{K, V} < ♯{L, T}.
-#I #L #K #V #c #f #HLK lapply (drops_fwd_lw … HLK) -HLK
+lemma drops_pair2_fwd_rfw: ∀b,f,I,L,K,V. ⬇*[b, f] L ≡ K.ⓑ{I}V → ∀T. ♯{K, V} < ♯{L, T}.
+#b #f #I #L #K #V #HLK lapply (drops_fwd_lw … HLK) -HLK
normalize in ⊢ (%→?→?%%); /3 width=3 by le_to_lt_to_lt/
qed-.
(* Advanced inversion lemma *************************************************)
-lemma drops_inv_x_pair_xy: ∀I,L,V,c,f. ⬇*[c,f] L ≡ L.ⓑ{I}V → ⊥.
-#I #L #V #c #f #H lapply (drops_fwd_lw … H) -c -f
+lemma drops_inv_x_pair_xy: ∀b,f,I,L,V. ⬇*[b,f] L ≡ L.ⓑ{I}V → ⊥.
+#b #f #I #L #V #H lapply (drops_fwd_lw … H) -b -f
/2 width=4 by lt_le_false/ (**) (* full auto is a bit slow: 19s *)
qed-.
(* Basic_2A1: includes: lpx_sn_atom lpx_sn_pair *)
inductive lexs (RN,RP:relation3 lenv term term): rtmap → relation lenv ≝
| lexs_atom: ∀f. lexs RN RP f (⋆) (⋆)
-| lexs_next: ∀I,L1,L2,V1,V2,f.
+| lexs_next: ∀f,I,L1,L2,V1,V2.
lexs RN RP f L1 L2 → RN L1 V1 V2 →
lexs RN RP (⫯f) (L1.ⓑ{I}V1) (L2.ⓑ{I}V2)
-| lexs_push: ∀I,L1,L2,V1,V2,f.
+| lexs_push: ∀f,I,L1,L2,V1,V2.
lexs RN RP f L1 L2 → RP L1 V1 V2 →
lexs RN RP (↑f) (L1.ⓑ{I}V1) (L2.ⓑ{I}V2)
.
(* Basic inversion lemmas ***************************************************)
-fact lexs_inv_atom1_aux: ∀RN,RP,X,Y,f. X ⦻*[RN, RP, f] Y → X = ⋆ → Y = ⋆.
-#RN #RP #X #Y #f * -X -Y -f //
-#I #L1 #L2 #V1 #V2 #f #_ #_ #H destruct
+fact lexs_inv_atom1_aux: ∀RN,RP,f,X,Y. X ⦻*[RN, RP, f] Y → X = ⋆ → Y = ⋆.
+#RN #RP #f #X #Y * -f -X -Y //
+#f #I #L1 #L2 #V1 #V2 #_ #_ #H destruct
qed-.
(* Basic_2A1: includes lpx_sn_inv_atom1 *)
-lemma lexs_inv_atom1: ∀RN,RP,Y,f. ⋆ ⦻*[RN, RP, f] Y → Y = ⋆.
+lemma lexs_inv_atom1: ∀RN,RP,f,Y. ⋆ ⦻*[RN, RP, f] Y → Y = ⋆.
/2 width=6 by lexs_inv_atom1_aux/ qed-.
-fact lexs_inv_next1_aux: ∀RN,RP,X,Y,f. X ⦻*[RN, RP, f] Y → ∀J,K1,W1,g. X = K1.ⓑ{J}W1 → f = ⫯g →
+fact lexs_inv_next1_aux: ∀RN,RP,f,X,Y. X ⦻*[RN, RP, f] Y → ∀g,J,K1,W1. X = K1.ⓑ{J}W1 → f = ⫯g →
∃∃K2,W2. K1 ⦻*[RN, RP, g] K2 & RN K1 W1 W2 & Y = K2.ⓑ{J}W2.
-#RN #RP #X #Y #f * -X -Y -f
-[ #f #J #K1 #W1 #g #H destruct
-| #I #L1 #L2 #V1 #V2 #f #HL #HV #J #K1 #W1 #g #H1 #H2 <(injective_next … H2) -g destruct
+#RN #RP #f #X #Y * -f -X -Y
+[ #f #g #J #K1 #W1 #H destruct
+| #f #I #L1 #L2 #V1 #V2 #HL #HV #g #J #K1 #W1 #H1 #H2 <(injective_next … H2) -g destruct
/2 width=5 by ex3_2_intro/
-| #I #L1 #L2 #V1 #V2 #f #_ #_ #J #K1 #W1 #g #_ #H elim (discr_push_next … H)
+| #f #I #L1 #L2 #V1 #V2 #_ #_ #g #J #K1 #W1 #_ #H elim (discr_push_next … H)
]
qed-.
(* Basic_2A1: includes lpx_sn_inv_pair1 *)
-lemma lexs_inv_next1: ∀RN,RP,J,K1,Y,W1,g. K1.ⓑ{J}W1 ⦻*[RN, RP, ⫯g] Y →
+lemma lexs_inv_next1: ∀RN,RP,g,J,K1,Y,W1. K1.ⓑ{J}W1 ⦻*[RN, RP, ⫯g] Y →
∃∃K2,W2. K1 ⦻*[RN, RP, g] K2 & RN K1 W1 W2 & Y = K2.ⓑ{J}W2.
/2 width=7 by lexs_inv_next1_aux/ qed-.
-fact lexs_inv_push1_aux: ∀RN,RP,X,Y,f. X ⦻*[RN, RP, f] Y → ∀J,K1,W1,g. X = K1.ⓑ{J}W1 → f = ↑g →
+fact lexs_inv_push1_aux: ∀RN,RP,f,X,Y. X ⦻*[RN, RP, f] Y → ∀g,J,K1,W1. X = K1.ⓑ{J}W1 → f = ↑g →
∃∃K2,W2. K1 ⦻*[RN, RP, g] K2 & RP K1 W1 W2 & Y = K2.ⓑ{J}W2.
-#RN #RP #X #Y #f * -X -Y -f
-[ #f #J #K1 #W1 #g #H destruct
-| #I #L1 #L2 #V1 #V2 #f #_ #_ #J #K1 #W1 #g #_ #H elim (discr_next_push … H)
-| #I #L1 #L2 #V1 #V2 #f #HL #HV #J #K1 #W1 #g #H1 #H2 <(injective_push … H2) -g destruct
+#RN #RP #f #X #Y * -f -X -Y
+[ #f #g #J #K1 #W1 #H destruct
+| #f #I #L1 #L2 #V1 #V2 #_ #_ #g #J #K1 #W1 #_ #H elim (discr_next_push … H)
+| #f #I #L1 #L2 #V1 #V2 #HL #HV #g #J #K1 #W1 #H1 #H2 <(injective_push … H2) -g destruct
/2 width=5 by ex3_2_intro/
]
qed-.
-lemma lexs_inv_push1: ∀RN,RP,J,K1,Y,W1,g. K1.ⓑ{J}W1 ⦻*[RN, RP, ↑g] Y →
+lemma lexs_inv_push1: ∀RN,RP,g,J,K1,Y,W1. K1.ⓑ{J}W1 ⦻*[RN, RP, ↑g] Y →
∃∃K2,W2. K1 ⦻*[RN, RP, g] K2 & RP K1 W1 W2 & Y = K2.ⓑ{J}W2.
/2 width=7 by lexs_inv_push1_aux/ qed-.
-fact lexs_inv_atom2_aux: ∀RN,RP,X,Y,f. X ⦻*[RN, RP, f] Y → Y = ⋆ → X = ⋆.
-#RN #RP #X #Y #f * -X -Y -f //
-#I #L1 #L2 #V1 #V2 #f #_ #_ #H destruct
+fact lexs_inv_atom2_aux: ∀RN,RP,f,X,Y. X ⦻*[RN, RP, f] Y → Y = ⋆ → X = ⋆.
+#RN #RP #f #X #Y * -f -X -Y //
+#f #I #L1 #L2 #V1 #V2 #_ #_ #H destruct
qed-.
(* Basic_2A1: includes lpx_sn_inv_atom2 *)
-lemma lexs_inv_atom2: ∀RN,RP,X,f. X ⦻*[RN, RP, f] ⋆ → X = ⋆.
+lemma lexs_inv_atom2: ∀RN,RP,f,X. X ⦻*[RN, RP, f] ⋆ → X = ⋆.
/2 width=6 by lexs_inv_atom2_aux/ qed-.
-fact lexs_inv_next2_aux: ∀RN,RP,X,Y,f. X ⦻*[RN, RP, f] Y → ∀J,K2,W2,g. Y = K2.ⓑ{J}W2 → f = ⫯g →
+fact lexs_inv_next2_aux: ∀RN,RP,f,X,Y. X ⦻*[RN, RP, f] Y → ∀g,J,K2,W2. Y = K2.ⓑ{J}W2 → f = ⫯g →
∃∃K1,W1. K1 ⦻*[RN, RP, g] K2 & RN K1 W1 W2 & X = K1.ⓑ{J}W1.
-#RN #RP #X #Y #f * -X -Y -f
-[ #f #J #K2 #W2 #g #H destruct
-| #I #L1 #L2 #V1 #V2 #f #HL #HV #J #K2 #W2 #g #H1 #H2 <(injective_next … H2) -g destruct
+#RN #RP #f #X #Y * -f -X -Y
+[ #f #g #J #K2 #W2 #H destruct
+| #f #I #L1 #L2 #V1 #V2 #HL #HV #g #J #K2 #W2 #H1 #H2 <(injective_next … H2) -g destruct
/2 width=5 by ex3_2_intro/
-| #I #L1 #L2 #V1 #V2 #f #_ #_ #J #K2 #W2 #g #_ #H elim (discr_push_next … H)
+| #f #I #L1 #L2 #V1 #V2 #_ #_ #g #J #K2 #W2 #_ #H elim (discr_push_next … H)
]
qed-.
(* Basic_2A1: includes lpx_sn_inv_pair2 *)
-lemma lexs_inv_next2: ∀RN,RP,J,X,K2,W2,g. X ⦻*[RN, RP, ⫯g] K2.ⓑ{J}W2 →
+lemma lexs_inv_next2: ∀RN,RP,g,J,X,K2,W2. X ⦻*[RN, RP, ⫯g] K2.ⓑ{J}W2 →
∃∃K1,W1. K1 ⦻*[RN, RP, g] K2 & RN K1 W1 W2 & X = K1.ⓑ{J}W1.
/2 width=7 by lexs_inv_next2_aux/ qed-.
-fact lexs_inv_push2_aux: ∀RN,RP,X,Y,f. X ⦻*[RN, RP, f] Y → ∀J,K2,W2,g. Y = K2.ⓑ{J}W2 → f = ↑g →
+fact lexs_inv_push2_aux: ∀RN,RP,f,X,Y. X ⦻*[RN, RP, f] Y → ∀g,J,K2,W2. Y = K2.ⓑ{J}W2 → f = ↑g →
∃∃K1,W1. K1 ⦻*[RN, RP, g] K2 & RP K1 W1 W2 & X = K1.ⓑ{J}W1.
-#RN #RP #X #Y #f * -X -Y -f
+#RN #RP #f #X #Y * -f -X -Y
[ #f #J #K2 #W2 #g #H destruct
-| #I #L1 #L2 #V1 #V2 #f #_ #_ #J #K2 #W2 #g #_ #H elim (discr_next_push … H)
-| #I #L1 #L2 #V1 #V2 #f #HL #HV #J #K2 #W2 #g #H1 #H2 <(injective_push … H2) -g destruct
+| #f #I #L1 #L2 #V1 #V2 #_ #_ #g #J #K2 #W2 #_ #H elim (discr_next_push … H)
+| #f #I #L1 #L2 #V1 #V2 #HL #HV #g #J #K2 #W2 #H1 #H2 <(injective_push … H2) -g destruct
/2 width=5 by ex3_2_intro/
]
qed-.
-lemma lexs_inv_push2: ∀RN,RP,J,X,K2,W2,g. X ⦻*[RN, RP, ↑g] K2.ⓑ{J}W2 →
+lemma lexs_inv_push2: ∀RN,RP,g,J,X,K2,W2. X ⦻*[RN, RP, ↑g] K2.ⓑ{J}W2 →
∃∃K1,W1. K1 ⦻*[RN, RP, g] K2 & RP K1 W1 W2 & X = K1.ⓑ{J}W1.
/2 width=7 by lexs_inv_push2_aux/ qed-.
(* Basic_2A1: includes lpx_sn_inv_pair *)
-lemma lexs_inv_next: ∀RN,RP,I1,I2,L1,L2,V1,V2,f.
+lemma lexs_inv_next: ∀RN,RP,f,I1,I2,L1,L2,V1,V2.
L1.ⓑ{I1}V1 ⦻*[RN, RP, ⫯f] (L2.ⓑ{I2}V2) →
∧∧ L1 ⦻*[RN, RP, f] L2 & RN L1 V1 V2 & I1 = I2.
-#RN #RP #I1 #I2 #L1 #L2 #V1 #V2 #f #H elim (lexs_inv_next1 … H) -H
+#RN #RP #f #I1 #I2 #L1 #L2 #V1 #V2 #H elim (lexs_inv_next1 … H) -H
#L0 #V0 #HL10 #HV10 #H destruct /2 width=1 by and3_intro/
qed-.
-lemma lexs_inv_push: ∀RN,RP,I1,I2,L1,L2,V1,V2,f.
+lemma lexs_inv_push: ∀RN,RP,f,I1,I2,L1,L2,V1,V2.
L1.ⓑ{I1}V1 ⦻*[RN, RP, ↑f] (L2.ⓑ{I2}V2) →
∧∧ L1 ⦻*[RN, RP, f] L2 & RP L1 V1 V2 & I1 = I2.
-#RN #RP #I1 #I2 #L1 #L2 #V1 #V2 #f #H elim (lexs_inv_push1 … H) -H
+#RN #RP #f #I1 #I2 #L1 #L2 #V1 #V2 #H elim (lexs_inv_push1 … H) -H
#L0 #V0 #HL10 #HV10 #H destruct /2 width=1 by and3_intro/
qed-.
-lemma lexs_inv_tl: ∀RN,RP,I,L1,L2,V1,V2,f. L1 ⦻*[RN, RP, ⫱f] L2 →
+lemma lexs_inv_tl: ∀RN,RP,f,I,L1,L2,V1,V2. L1 ⦻*[RN, RP, ⫱f] L2 →
RN L1 V1 V2 → RP L1 V1 V2 →
L1.ⓑ{I}V1 ⦻*[RN, RP, f] L2.ⓑ{I}V2.
-#RN #RP #I #L2 #L2 #V1 #V2 #f elim (pn_split f) *
+#RN #RP #f #I #L2 #L2 #V1 #V2 elim (pn_split f) *
/2 width=1 by lexs_next, lexs_push/
qed-.
(* Basic forward lemmas *****************************************************)
-lemma lexs_fwd_pair: ∀RN,RP,I1,I2,L1,L2,V1,V2,f.
+lemma lexs_fwd_pair: ∀RN,RP,f,I1,I2,L1,L2,V1,V2.
L1.ⓑ{I1}V1 ⦻*[RN, RP, f] L2.ⓑ{I2}V2 →
L1 ⦻*[RN, RP, ⫱f] L2 ∧ I1 = I2.
-#RN #RP #I1 #I2 #L2 #L2 #V1 #V2 #f #Hf
+#RN #RP #f #I1 #I2 #L2 #L2 #V1 #V2 #Hf
elim (pn_split f) * #g #H destruct
[ elim (lexs_inv_push … Hf) | elim (lexs_inv_next … Hf) ] -Hf
/2 width=1 by conj/
(* Basic properties *********************************************************)
lemma lexs_eq_repl_back: ∀RN,RP,L1,L2. eq_repl_back … (λf. L1 ⦻*[RN, RP, f] L2).
-#RN #RP #L1 #L2 #f1 #H elim H -L1 -L2 -f1 //
-#I #L1 #L2 #V1 #v2 #f1 #_ #HV #IH #f2 #H
+#RN #RP #L1 #L2 #f1 #H elim H -f1 -L1 -L2 //
+#f1 #I #L1 #L2 #V1 #v2 #_ #HV #IH #f2 #H
[ elim (eq_inv_nx … H) -H /3 width=3 by lexs_next/
| elim (eq_inv_px … H) -H /3 width=3 by lexs_push/
]
qed.
lemma sle_lexs_trans: ∀RN,RP. (∀L,T1,T2. RN L T1 T2 → RP L T1 T2) →
- ∀L1,L2,f2. L1 ⦻*[RN, RP, f2] L2 →
+ ∀f2,L1,L2. L1 ⦻*[RN, RP, f2] L2 →
∀f1. f1 ⊆ f2 → L1 ⦻*[RN, RP, f1] L2.
-#RN #RP #HR #L1 #L2 #f2 #H elim H -L1 -L2 -f2 //
-#I #L1 #L2 #V1 #V2 #f2 #_ #HV12 #IH
+#RN #RP #HR #f2 #L1 #L2 #H elim H -f2 -L1 -L2 //
+#f2 #I #L1 #L2 #V1 #V2 #_ #HV12 #IH
[ * * [2: #n1 ] ] #f1 #H
[ /4 width=5 by lexs_next, sle_inv_nn/
| /4 width=5 by lexs_push, sle_inv_pn/
qed-.
lemma sle_lexs_conf: ∀RN,RP. (∀L,T1,T2. RP L T1 T2 → RN L T1 T2) →
- ∀L1,L2,f1. L1 ⦻*[RN, RP, f1] L2 →
+ ∀f1,L1,L2. L1 ⦻*[RN, RP, f1] L2 →
∀f2. f1 ⊆ f2 → L1 ⦻*[RN, RP, f2] L2.
-#RN #RP #HR #L1 #L2 #f2 #H elim H -L1 -L2 -f2 //
-#I #L1 #L2 #V1 #V2 #f1 #_ #HV12 #IH
+#RN #RP #HR #f1 #L1 #L2 #H elim H -f1 -L1 -L2 //
+#f1 #I #L1 #L2 #V1 #V2 #_ #HV12 #IH
[2: * * [2: #n2 ] ] #f2 #H
[ /4 width=5 by lexs_next, sle_inv_pn/
| /4 width=5 by lexs_push, sle_inv_pp/
lemma lexs_co: ∀RN1,RP1,RN2,RP2.
(∀L1,T1,T2. RN1 L1 T1 T2 → RN2 L1 T1 T2) →
(∀L1,T1,T2. RP1 L1 T1 T2 → RP2 L1 T1 T2) →
- ∀L1,L2,f. L1 ⦻*[RN1, RP1, f] L2 → L1 ⦻*[RN2, RP2, f] L2.
-#RN1 #RP1 #RN2 #RP2 #HRN #HRP #L1 #L2 #f #H elim H -L1 -L2 -f
+ ∀f,L1,L2. L1 ⦻*[RN1, RP1, f] L2 → L1 ⦻*[RN2, RP2, f] L2.
+#RN1 #RP1 #RN2 #RP2 #HRN #HRP #f #L1 #L2 #H elim H -f -L1 -L2
/3 width=1 by lexs_atom, lexs_next, lexs_push/
qed-.
(* Forward lemmas with length for local environments ************************)
(* Basic_2A1: includes: lpx_sn_fwd_length *)
-lemma lexs_fwd_length: ∀RN,RP,L1,L2,f. L1 ⦻*[RN, RP, f] L2 → |L1| = |L2|.
-#RM #RP #L1 #L2 #f #H elim H -L1 -L2 -f //
+lemma lexs_fwd_length: ∀RN,RP,f,L1,L2. L1 ⦻*[RN, RP, f] L2 → |L1| = |L2|.
+#RM #RP #f #L1 #L2 #H elim H -f -L1 -L2 //
qed-.
∀L1,L0. L1 ⦻*[RN1, RP1, f] L0 →
∀L2. L0 ⦻*[RN2, RP2, f] L2 →
L1 ⦻*[RN, RP, f] L2.
-#RN1 #RP1 #RN2 #RP2 #RN #RP #f #HN #HP #L1 #L0 #H elim H -L1 -L0 -f
+#RN1 #RP1 #RN2 #RP2 #RN #RP #f #HN #HP #L1 #L0 #H elim H -f -L1 -L0
[ #f #L2 #H >(lexs_inv_atom1 … H) -L2 //
-| #I #K1 #K #V1 #V #f #HK1 #HV1 #IH #L2 #H elim (lexs_inv_next1 … H) -H
+| #f #I #K1 #K #V1 #V #HK1 #HV1 #IH #L2 #H elim (lexs_inv_next1 … H) -H
#K2 #V2 #HK2 #HV2 #H destruct /4 width=6 by lexs_next/
-| #I #K1 #K #V1 #V #f #HK1 #HV1 #IH #L2 #H elim (lexs_inv_push1 … H) -H
+| #f #I #K1 #K #V1 #V #HK1 #HV1 #IH #L2 #H elim (lexs_inv_push1 … H) -H
#K2 #V2 #HK2 #HV2 #H destruct /4 width=6 by lexs_push/
]
qed-.
∀f1. L1 ⦻*[RN, RP, f1] L2 →
∀f2. L1 ⦻*[RN, RP, f2] L2 →
∀f. f1 ⋒ f2 ≡ f → L1 ⦻*[RN, RP, f] L2.
-#RN #RP #L1 #L2 #f1 #H elim H -L1 -L2 -f1 //
-#I #L1 #L2 #V1 #V2 #f1 #_ #HV12 #IH #f2 #H #f #Hf
+#RN #RP #L1 #L2 #f1 #H elim H -f1 -L1 -L2 //
+#f1 #I #L1 #L2 #V1 #V2 #_ #HV12 #IH #f2 #H #f #Hf
elim (pn_split f2) * #g2 #H2 destruct
try elim (lexs_inv_push … H) try elim (lexs_inv_next … H) -H
[ elim (sand_inv_npx … Hf) | elim (sand_inv_nnx … Hf)
∀f1. L1 ⦻*[RN, RP, f1] L2 →
∀f2. L1 ⦻*[RN, RP, f2] L2 →
∀f. f1 ⋓ f2 ≡ f → L1 ⦻*[RN, RP, f] L2.
-#RN #RP #L1 #L2 #f1 #H elim H -L1 -L2 -f1 //
-#I #L1 #L2 #V1 #V2 #f1 #_ #HV12 #IH #f2 #H #f #Hf
+#RN #RP #L1 #L2 #f1 #H elim H -f1 -L1 -L2 //
+#f1 #I #L1 #L2 #V1 #V2 #_ #HV12 #IH #f2 #H #f #Hf
elim (pn_split f2) * #g2 #H2 destruct
try elim (lexs_inv_push … H) try elim (lexs_inv_next … H) -H
[ elim (sor_inv_npx … Hf) | elim (sor_inv_nnx … Hf)
lifts_nil lifts_cons
*)
inductive lifts: rtmap → relation term ≝
-| lifts_sort: ∀s,f. lifts f (⋆s) (⋆s)
-| lifts_lref: ∀i1,i2,f. @⦃i1, f⦄ ≡ i2 → lifts f (#i1) (#i2)
-| lifts_gref: ∀l,f. lifts f (§l) (§l)
-| lifts_bind: ∀p,I,V1,V2,T1,T2,f.
+| lifts_sort: ∀f,s. lifts f (⋆s) (⋆s)
+| lifts_lref: ∀f,i1,i2. @⦃i1, f⦄ ≡ i2 → lifts f (#i1) (#i2)
+| lifts_gref: ∀f,l. lifts f (§l) (§l)
+| lifts_bind: ∀f,p,I,V1,V2,T1,T2.
lifts f V1 V2 → lifts (↑f) T1 T2 →
lifts f (ⓑ{p,I}V1.T1) (ⓑ{p,I}V2.T2)
-| lifts_flat: ∀I,V1,V2,T1,T2,f.
+| lifts_flat: ∀f,I,V1,V2,T1,T2.
lifts f V1 V2 → lifts f T1 T2 →
lifts f (ⓕ{I}V1.T1) (ⓕ{I}V2.T2)
.
(* Basic inversion lemmas ***************************************************)
-fact lifts_inv_sort1_aux: ∀X,Y,f. ⬆*[f] X ≡ Y → ∀s. X = ⋆s → Y = ⋆s.
-#X #Y #f * -X -Y -f //
-[ #i1 #i2 #f #_ #x #H destruct
-| #p #I #V1 #V2 #T1 #T2 #f #_ #_ #x #H destruct
-| #I #V1 #V2 #T1 #T2 #f #_ #_ #x #H destruct
+fact lifts_inv_sort1_aux: ∀f,X,Y. ⬆*[f] X ≡ Y → ∀s. X = ⋆s → Y = ⋆s.
+#f #X #Y * -f -X -Y //
+[ #f #i1 #i2 #_ #x #H destruct
+| #f #p #I #V1 #V2 #T1 #T2 #_ #_ #x #H destruct
+| #f #I #V1 #V2 #T1 #T2 #_ #_ #x #H destruct
]
qed-.
(* Basic_1: was: lift1_sort *)
(* Basic_2A1: includes: lift_inv_sort1 *)
-lemma lifts_inv_sort1: ∀Y,s,f. ⬆*[f] ⋆s ≡ Y → Y = ⋆s.
+lemma lifts_inv_sort1: ∀f,Y,s. ⬆*[f] ⋆s ≡ Y → Y = ⋆s.
/2 width=4 by lifts_inv_sort1_aux/ qed-.
-fact lifts_inv_lref1_aux: ∀X,Y,f. ⬆*[f] X ≡ Y → ∀i1. X = #i1 →
+fact lifts_inv_lref1_aux: ∀f,X,Y. ⬆*[f] X ≡ Y → ∀i1. X = #i1 →
∃∃i2. @⦃i1, f⦄ ≡ i2 & Y = #i2.
-#X #Y #f * -X -Y -f
-[ #s #f #x #H destruct
-| #i1 #i2 #f #Hi12 #x #H destruct /2 width=3 by ex2_intro/
-| #l #f #x #H destruct
-| #p #I #V1 #V2 #T1 #T2 #f #_ #_ #x #H destruct
-| #I #V1 #V2 #T1 #T2 #f #_ #_ #x #H destruct
+#f #X #Y * -f -X -Y
+[ #f #s #x #H destruct
+| #f #i1 #i2 #Hi12 #x #H destruct /2 width=3 by ex2_intro/
+| #f #l #x #H destruct
+| #f #p #I #V1 #V2 #T1 #T2 #_ #_ #x #H destruct
+| #f #I #V1 #V2 #T1 #T2 #_ #_ #x #H destruct
]
qed-.
(* Basic_1: was: lift1_lref *)
(* Basic_2A1: includes: lift_inv_lref1 lift_inv_lref1_lt lift_inv_lref1_ge *)
-lemma lifts_inv_lref1: ∀Y,i1,f. ⬆*[f] #i1 ≡ Y →
+lemma lifts_inv_lref1: ∀f,Y,i1. ⬆*[f] #i1 ≡ Y →
∃∃i2. @⦃i1, f⦄ ≡ i2 & Y = #i2.
/2 width=3 by lifts_inv_lref1_aux/ qed-.
-fact lifts_inv_gref1_aux: ∀X,Y,f. ⬆*[f] X ≡ Y → ∀l. X = §l → Y = §l.
-#X #Y #f * -X -Y -f //
-[ #i1 #i2 #f #_ #x #H destruct
-| #p #I #V1 #V2 #T1 #T2 #f #_ #_ #x #H destruct
-| #I #V1 #V2 #T1 #T2 #f #_ #_ #x #H destruct
+fact lifts_inv_gref1_aux: ∀f,X,Y. ⬆*[f] X ≡ Y → ∀l. X = §l → Y = §l.
+#f #X #Y * -f -X -Y //
+[ #f #i1 #i2 #_ #x #H destruct
+| #f #p #I #V1 #V2 #T1 #T2 #_ #_ #x #H destruct
+| #f #I #V1 #V2 #T1 #T2 #_ #_ #x #H destruct
]
qed-.
(* Basic_2A1: includes: lift_inv_gref1 *)
-lemma lifts_inv_gref1: ∀Y,l,f. ⬆*[f] §l ≡ Y → Y = §l.
+lemma lifts_inv_gref1: ∀f,Y,l. ⬆*[f] §l ≡ Y → Y = §l.
/2 width=4 by lifts_inv_gref1_aux/ qed-.
-fact lifts_inv_bind1_aux: ∀X,Y,f. ⬆*[f] X ≡ Y →
+fact lifts_inv_bind1_aux: ∀f,X,Y. ⬆*[f] X ≡ Y →
∀p,I,V1,T1. X = ⓑ{p,I}V1.T1 →
∃∃V2,T2. ⬆*[f] V1 ≡ V2 & ⬆*[↑f] T1 ≡ T2 &
Y = ⓑ{p,I}V2.T2.
-#X #Y #f * -X -Y -f
-[ #s #f #q #J #W1 #U1 #H destruct
-| #i1 #i2 #f #_ #q #J #W1 #U1 #H destruct
-| #l #f #b #J #W1 #U1 #H destruct
-| #p #I #V1 #V2 #T1 #T2 #f #HV12 #HT12 #q #J #W1 #U1 #H destruct /2 width=5 by ex3_2_intro/
-| #I #V1 #V2 #T1 #T2 #f #_ #_ #q #J #W1 #U1 #H destruct
+#f #X #Y * -f -X -Y
+[ #f #s #q #J #W1 #U1 #H destruct
+| #f #i1 #i2 #_ #q #J #W1 #U1 #H destruct
+| #f #l #b #J #W1 #U1 #H destruct
+| #f #p #I #V1 #V2 #T1 #T2 #HV12 #HT12 #q #J #W1 #U1 #H destruct /2 width=5 by ex3_2_intro/
+| #f #I #V1 #V2 #T1 #T2 #_ #_ #q #J #W1 #U1 #H destruct
]
qed-.
(* Basic_1: was: lift1_bind *)
(* Basic_2A1: includes: lift_inv_bind1 *)
-lemma lifts_inv_bind1: ∀p,I,V1,T1,Y,f. ⬆*[f] ⓑ{p,I}V1.T1 ≡ Y →
+lemma lifts_inv_bind1: ∀f,p,I,V1,T1,Y. ⬆*[f] ⓑ{p,I}V1.T1 ≡ Y →
∃∃V2,T2. ⬆*[f] V1 ≡ V2 & ⬆*[↑f] T1 ≡ T2 &
Y = ⓑ{p,I}V2.T2.
/2 width=3 by lifts_inv_bind1_aux/ qed-.
-fact lifts_inv_flat1_aux: ∀X,Y. ∀f:rtmap. ⬆*[f] X ≡ Y →
+fact lifts_inv_flat1_aux: ∀f:rtmap. ∀X,Y. ⬆*[f] X ≡ Y →
∀I,V1,T1. X = ⓕ{I}V1.T1 →
∃∃V2,T2. ⬆*[f] V1 ≡ V2 & ⬆*[f] T1 ≡ T2 &
Y = ⓕ{I}V2.T2.
-#X #Y #f * -X -Y -f
-[ #s #f #J #W1 #U1 #H destruct
-| #i1 #i2 #f #_ #J #W1 #U1 #H destruct
-| #l #f #J #W1 #U1 #H destruct
-| #p #I #V1 #V2 #T1 #T2 #f #_ #_ #J #W1 #U1 #H destruct
-| #I #V1 #V2 #T1 #T2 #f #HV12 #HT12 #J #W1 #U1 #H destruct /2 width=5 by ex3_2_intro/
+#f #X #Y * -f -X -Y
+[ #f #s #J #W1 #U1 #H destruct
+| #f #i1 #i2 #_ #J #W1 #U1 #H destruct
+| #f #l #J #W1 #U1 #H destruct
+| #f #p #I #V1 #V2 #T1 #T2 #_ #_ #J #W1 #U1 #H destruct
+| #f #I #V1 #V2 #T1 #T2 #HV12 #HT12 #J #W1 #U1 #H destruct /2 width=5 by ex3_2_intro/
]
qed-.
(* Basic_1: was: lift1_flat *)
(* Basic_2A1: includes: lift_inv_flat1 *)
-lemma lifts_inv_flat1: ∀I,V1,T1,Y. ∀f:rtmap. ⬆*[f] ⓕ{I}V1.T1 ≡ Y →
+lemma lifts_inv_flat1: ∀f:rtmap. ∀I,V1,T1,Y. ⬆*[f] ⓕ{I}V1.T1 ≡ Y →
∃∃V2,T2. ⬆*[f] V1 ≡ V2 & ⬆*[f] T1 ≡ T2 &
Y = ⓕ{I}V2.T2.
/2 width=3 by lifts_inv_flat1_aux/ qed-.
-fact lifts_inv_sort2_aux: ∀X,Y,f. ⬆*[f] X ≡ Y → ∀s. Y = ⋆s → X = ⋆s.
-#X #Y #f * -X -Y -f //
-[ #i1 #i2 #f #_ #x #H destruct
-| #p #I #V1 #V2 #T1 #T2 #f #_ #_ #x #H destruct
-| #I #V1 #V2 #T1 #T2 #f #_ #_ #x #H destruct
+fact lifts_inv_sort2_aux: ∀f,X,Y. ⬆*[f] X ≡ Y → ∀s. Y = ⋆s → X = ⋆s.
+#f #X #Y * -f -X -Y //
+[ #f #i1 #i2 #_ #x #H destruct
+| #f #p #I #V1 #V2 #T1 #T2 #_ #_ #x #H destruct
+| #f #I #V1 #V2 #T1 #T2 #_ #_ #x #H destruct
]
qed-.
(* Basic_1: includes: lift_gen_sort *)
(* Basic_2A1: includes: lift_inv_sort2 *)
-lemma lifts_inv_sort2: ∀X,s,f. ⬆*[f] X ≡ ⋆s → X = ⋆s.
+lemma lifts_inv_sort2: ∀f,X,s. ⬆*[f] X ≡ ⋆s → X = ⋆s.
/2 width=4 by lifts_inv_sort2_aux/ qed-.
-fact lifts_inv_lref2_aux: ∀X,Y,f. ⬆*[f] X ≡ Y → ∀i2. Y = #i2 →
+fact lifts_inv_lref2_aux: ∀f,X,Y. ⬆*[f] X ≡ Y → ∀i2. Y = #i2 →
∃∃i1. @⦃i1, f⦄ ≡ i2 & X = #i1.
-#X #Y #f * -X -Y -f
-[ #s #f #x #H destruct
-| #i1 #i2 #f #Hi12 #x #H destruct /2 width=3 by ex2_intro/
-| #l #f #x #H destruct
-| #p #I #V1 #V2 #T1 #T2 #f #_ #_ #x #H destruct
-| #I #V1 #V2 #T1 #T2 #f #_ #_ #x #H destruct
+#f #X #Y * -f -X -Y
+[ #f #s #x #H destruct
+| #f #i1 #i2 #Hi12 #x #H destruct /2 width=3 by ex2_intro/
+| #f #l #x #H destruct
+| #f #p #I #V1 #V2 #T1 #T2 #_ #_ #x #H destruct
+| #f #I #V1 #V2 #T1 #T2 #_ #_ #x #H destruct
]
qed-.
(* Basic_1: includes: lift_gen_lref lift_gen_lref_lt lift_gen_lref_false lift_gen_lref_ge *)
(* Basic_2A1: includes: lift_inv_lref2 lift_inv_lref2_lt lift_inv_lref2_be lift_inv_lref2_ge lift_inv_lref2_plus *)
-lemma lifts_inv_lref2: ∀X,i2,f. ⬆*[f] X ≡ #i2 →
+lemma lifts_inv_lref2: ∀f,X,i2. ⬆*[f] X ≡ #i2 →
∃∃i1. @⦃i1, f⦄ ≡ i2 & X = #i1.
/2 width=3 by lifts_inv_lref2_aux/ qed-.
-fact lifts_inv_gref2_aux: ∀X,Y,f. ⬆*[f] X ≡ Y → ∀l. Y = §l → X = §l.
-#X #Y #f * -X -Y -f //
-[ #i1 #i2 #f #_ #x #H destruct
-| #p #I #V1 #V2 #T1 #T2 #f #_ #_ #x #H destruct
-| #I #V1 #V2 #T1 #T2 #f #_ #_ #x #H destruct
+fact lifts_inv_gref2_aux: ∀f,X,Y. ⬆*[f] X ≡ Y → ∀l. Y = §l → X = §l.
+#f #X #Y * -f -X -Y //
+[ #f #i1 #i2 #_ #x #H destruct
+| #f #p #I #V1 #V2 #T1 #T2 #_ #_ #x #H destruct
+| #f #I #V1 #V2 #T1 #T2 #_ #_ #x #H destruct
]
qed-.
(* Basic_2A1: includes: lift_inv_gref1 *)
-lemma lifts_inv_gref2: ∀X,l,f. ⬆*[f] X ≡ §l → X = §l.
+lemma lifts_inv_gref2: ∀f,X,l. ⬆*[f] X ≡ §l → X = §l.
/2 width=4 by lifts_inv_gref2_aux/ qed-.
-fact lifts_inv_bind2_aux: ∀X,Y,f. ⬆*[f] X ≡ Y →
+fact lifts_inv_bind2_aux: ∀f,X,Y. ⬆*[f] X ≡ Y →
∀p,I,V2,T2. Y = ⓑ{p,I}V2.T2 →
∃∃V1,T1. ⬆*[f] V1 ≡ V2 & ⬆*[↑f] T1 ≡ T2 &
X = ⓑ{p,I}V1.T1.
-#X #Y #f * -X -Y -f
-[ #s #f #q #J #W2 #U2 #H destruct
-| #i1 #i2 #f #_ #q #J #W2 #U2 #H destruct
-| #l #f #q #J #W2 #U2 #H destruct
-| #p #I #V1 #V2 #T1 #T2 #f #HV12 #HT12 #q #J #W2 #U2 #H destruct /2 width=5 by ex3_2_intro/
-| #I #V1 #V2 #T1 #T2 #f #_ #_ #q #J #W2 #U2 #H destruct
+#f #X #Y * -f -X -Y
+[ #f #s #q #J #W2 #U2 #H destruct
+| #f #i1 #i2 #_ #q #J #W2 #U2 #H destruct
+| #f #l #q #J #W2 #U2 #H destruct
+| #f #p #I #V1 #V2 #T1 #T2 #HV12 #HT12 #q #J #W2 #U2 #H destruct /2 width=5 by ex3_2_intro/
+| #f #I #V1 #V2 #T1 #T2 #_ #_ #q #J #W2 #U2 #H destruct
]
qed-.
(* Basic_1: includes: lift_gen_bind *)
(* Basic_2A1: includes: lift_inv_bind2 *)
-lemma lifts_inv_bind2: ∀p,I,V2,T2,X,f. ⬆*[f] X ≡ ⓑ{p,I}V2.T2 →
+lemma lifts_inv_bind2: ∀f,p,I,V2,T2,X. ⬆*[f] X ≡ ⓑ{p,I}V2.T2 →
∃∃V1,T1. ⬆*[f] V1 ≡ V2 & ⬆*[↑f] T1 ≡ T2 &
X = ⓑ{p,I}V1.T1.
/2 width=3 by lifts_inv_bind2_aux/ qed-.
-fact lifts_inv_flat2_aux: ∀X,Y. ∀f:rtmap. ⬆*[f] X ≡ Y →
+fact lifts_inv_flat2_aux: ∀f:rtmap. ∀X,Y. ⬆*[f] X ≡ Y →
∀I,V2,T2. Y = ⓕ{I}V2.T2 →
∃∃V1,T1. ⬆*[f] V1 ≡ V2 & ⬆*[f] T1 ≡ T2 &
X = ⓕ{I}V1.T1.
-#X #Y #f * -X -Y -f
-[ #s #f #J #W2 #U2 #H destruct
-| #i1 #i2 #f #_ #J #W2 #U2 #H destruct
-| #l #f #J #W2 #U2 #H destruct
-| #p #I #V1 #V2 #T1 #T2 #f #_ #_ #J #W2 #U2 #H destruct
-| #I #V1 #V2 #T1 #T2 #f #HV12 #HT12 #J #W2 #U2 #H destruct /2 width=5 by ex3_2_intro/
+#f #X #Y * -f -X -Y
+[ #f #s #J #W2 #U2 #H destruct
+| #f #i1 #i2 #_ #J #W2 #U2 #H destruct
+| #f #l #J #W2 #U2 #H destruct
+| #f #p #I #V1 #V2 #T1 #T2 #_ #_ #J #W2 #U2 #H destruct
+| #f #I #V1 #V2 #T1 #T2 #HV12 #HT12 #J #W2 #U2 #H destruct /2 width=5 by ex3_2_intro/
]
qed-.
(* Basic_1: includes: lift_gen_flat *)
(* Basic_2A1: includes: lift_inv_flat2 *)
-lemma lifts_inv_flat2: ∀I,V2,T2,X. ∀f:rtmap. ⬆*[f] X ≡ ⓕ{I}V2.T2 →
+lemma lifts_inv_flat2: ∀f:rtmap. ∀I,V2,T2,X. ⬆*[f] X ≡ ⓕ{I}V2.T2 →
∃∃V1,T1. ⬆*[f] V1 ≡ V2 & ⬆*[f] T1 ≡ T2 &
X = ⓕ{I}V1.T1.
/2 width=3 by lifts_inv_flat2_aux/ qed-.
(* Basic_2A1: includes: lift_inv_pair_xy_x *)
-lemma lifts_inv_pair_xy_x: ∀I,V,T,f. ⬆*[f] ②{I}V.T ≡ V → ⊥.
-#J #V elim V -V
-[ * #i #U #f #H
+lemma lifts_inv_pair_xy_x: ∀f,I,V,T. ⬆*[f] ②{I}V.T ≡ V → ⊥.
+#f #J #V elim V -V
+[ * #i #U #H
[ lapply (lifts_inv_sort2 … H) -H #H destruct
| elim (lifts_inv_lref2 … H) -H
#x #_ #H destruct
| lapply (lifts_inv_gref2 … H) -H #H destruct
]
-| * [ #p ] #I #V2 #T2 #IHV2 #_ #U #f #H
+| * [ #p ] #I #V2 #T2 #IHV2 #_ #U #H
[ elim (lifts_inv_bind2 … H) -H #V1 #T1 #HV12 #_ #H destruct /2 width=3 by/
| elim (lifts_inv_flat2 … H) -H #V1 #T1 #HV12 #_ #H destruct /2 width=3 by/
]
(* Basic forward lemmas *****************************************************)
(* Basic_2A1: includes: lift_inv_O2 *)
-lemma lifts_fwd_isid: ∀T1,T2,f. ⬆*[f] T1 ≡ T2 → 𝐈⦃f⦄ → T1 = T2.
-#T1 #T2 #f #H elim H -T1 -T2 -f
+lemma lifts_fwd_isid: ∀f,T1,T2. ⬆*[f] T1 ≡ T2 → 𝐈⦃f⦄ → T1 = T2.
+#f #T1 #T2 #H elim H -f -T1 -T2
/4 width=3 by isid_inv_at_mono, isid_push, eq_f2, eq_f/
qed-.
(* Basic_2A1: includes: lift_fwd_pair1 *)
-lemma lifts_fwd_pair1: ∀I,V1,T1,Y. ∀f:rtmap. ⬆*[f] ②{I}V1.T1 ≡ Y →
+lemma lifts_fwd_pair1: ∀f:rtmap. ∀I,V1,T1,Y. ⬆*[f] ②{I}V1.T1 ≡ Y →
∃∃V2,T2. ⬆*[f] V1 ≡ V2 & Y = ②{I}V2.T2.
-* [ #p ] #I #V1 #T1 #Y #f #H
+#f * [ #p ] #I #V1 #T1 #Y #H
[ elim (lifts_inv_bind1 … H) -H /2 width=4 by ex2_2_intro/
| elim (lifts_inv_flat1 … H) -H /2 width=4 by ex2_2_intro/
]
qed-.
(* Basic_2A1: includes: lift_fwd_pair2 *)
-lemma lifts_fwd_pair2: ∀I,V2,T2,X. ∀f:rtmap. ⬆*[f] X ≡ ②{I}V2.T2 →
+lemma lifts_fwd_pair2: ∀f:rtmap. ∀I,V2,T2,X. ⬆*[f] X ≡ ②{I}V2.T2 →
∃∃V1,T1. ⬆*[f] V1 ≡ V2 & X = ②{I}V1.T1.
-* [ #p ] #I #V2 #T2 #X #f #H
+#f * [ #p ] #I #V2 #T2 #X #H
[ elim (lifts_inv_bind2 … H) -H /2 width=4 by ex2_2_intro/
| elim (lifts_inv_flat2 … H) -H /2 width=4 by ex2_2_intro/
]
(* Basic_1: includes: lift_free (right to left) *)
(* Basic_2A1: includes: lift_split *)
-lemma lifts_split_trans: ∀T1,T2,f. ⬆*[f] T1 ≡ T2 →
+lemma lifts_split_trans: ∀f,T1,T2. ⬆*[f] T1 ≡ T2 →
∀f1,f2. f2 ⊚ f1 ≡ f →
∃∃T. ⬆*[f1] T1 ≡ T & ⬆*[f2] T ≡ T2.
-#T1 #T2 #f #H elim H -T1 -T2 -f
+#f #T1 #T2 #H elim H -f -T1 -T2
[ /3 width=3 by lifts_sort, ex2_intro/
-| #i1 #i2 #f #Hi #f1 #f2 #Ht elim (after_at_fwd … Hi … Ht) -Hi -Ht
+| #f #i1 #i2 #Hi #f1 #f2 #Ht elim (after_at_fwd … Hi … Ht) -Hi -Ht
/3 width=3 by lifts_lref, ex2_intro/
| /3 width=3 by lifts_gref, ex2_intro/
-| #p #I #V1 #V2 #T1 #T2 #f #_ #_ #IHV #IHT #f1 #f2 #Ht
+| #f #p #I #V1 #V2 #T1 #T2 #_ #_ #IHV #IHT #f1 #f2 #Ht
elim (IHV … Ht) elim (IHT (↑f1) (↑f2)) -IHV -IHT
/3 width=5 by lifts_bind, after_O2, ex2_intro/
-| #I #V1 #V2 #T1 #T2 #f #_ #_ #IHV #IHT #f1 #f2 #Ht
+| #f #I #V1 #V2 #T1 #T2 #_ #_ #IHV #IHT #f1 #f2 #Ht
elim (IHV … Ht) elim (IHT … Ht) -IHV -IHT -Ht
/3 width=5 by lifts_flat, ex2_intro/
]
qed-.
(* Note: apparently, this was missing in Basic_2A1 *)
-lemma lifts_split_div: ∀T1,T2,f1. ⬆*[f1] T1 ≡ T2 →
+lemma lifts_split_div: ∀f1,T1,T2. ⬆*[f1] T1 ≡ T2 →
∀f2,f. f2 ⊚ f1 ≡ f →
∃∃T. ⬆*[f2] T2 ≡ T & ⬆*[f] T1 ≡ T.
-#T1 #T2 #f1 #H elim H -T1 -T2 -f1
+#f1 #T1 #T2 #H elim H -f1 -T1 -T2
[ /3 width=3 by lifts_sort, ex2_intro/
-| #i1 #i2 #f1 #Hi #f2 #f #Ht elim (after_at1_fwd … Hi … Ht) -Hi -Ht
+| #f1 #i1 #i2 #Hi #f2 #f #Ht elim (after_at1_fwd … Hi … Ht) -Hi -Ht
/3 width=3 by lifts_lref, ex2_intro/
| /3 width=3 by lifts_gref, ex2_intro/
-| #p #I #V1 #V2 #T1 #T2 #f1 #_ #_ #IHV #IHT #f2 #f #Ht
+| #f1 #p #I #V1 #V2 #T1 #T2 #_ #_ #IHV #IHT #f2 #f #Ht
elim (IHV … Ht) elim (IHT (↑f2) (↑f)) -IHV -IHT
/3 width=5 by lifts_bind, after_O2, ex2_intro/
-| #I #V1 #V2 #T1 #T2 #f1 #_ #_ #IHV #IHT #f2 #f #Ht
+| #f1 #I #V1 #V2 #T1 #T2 #_ #_ #IHV #IHT #f2 #f #Ht
elim (IHV … Ht) elim (IHT … Ht) -IHV -IHT -Ht
/3 width=5 by lifts_flat, ex2_intro/
]
(* Basic_1: includes: lift_gen_lift *)
(* Basic_2A1: includes: lift_div_le lift_div_be *)
-theorem lifts_div: ∀T,T2,f2. ⬆*[f2] T2 ≡ T → ∀T1,f. ⬆*[f] T1 ≡ T →
+theorem lifts_div: ∀f2,T,T2. ⬆*[f2] T2 ≡ T → ∀T1,f. ⬆*[f] T1 ≡ T →
∀f1. f2 ⊚ f1 ≡ f → ⬆*[f1] T1 ≡ T2.
-#T #T2 #f2 #H elim H -T -T2 -f2
-[ #s #f2 #T1 #f #H >(lifts_inv_sort2 … H) -T1 //
-| #i2 #i #f2 #Hi2 #T1 #f #H #f1 #Ht21 elim (lifts_inv_lref2 … H) -H
+#f2 #T #T2 #H elim H -f2 -T -T2
+[ #f2 #s #T1 #f #H >(lifts_inv_sort2 … H) -T1 //
+| #f2 #i2 #i #Hi2 #T1 #f #H #f1 #Ht21 elim (lifts_inv_lref2 … H) -H
#i1 #Hi1 #H destruct /3 width=6 by lifts_lref, after_fwd_at1/
-| #l #f2 #T1 #f #H >(lifts_inv_gref2 … H) -T1 //
-| #p #I #W2 #W #U2 #U #f2 #_ #_ #IHW #IHU #T1 #f #H
+| #f2 #l #T1 #f #H >(lifts_inv_gref2 … H) -T1 //
+| #f2 #p #I #W2 #W #U2 #U #_ #_ #IHW #IHU #T1 #f #H
elim (lifts_inv_bind2 … H) -H #W1 #U1 #HW1 #HU1 #H destruct
/4 width=3 by lifts_bind, after_O2/
-| #I #W2 #W #U2 #U #f2 #_ #_ #IHW #IHU #T1 #f #H
+| #f2 #I #W2 #W #U2 #U #_ #_ #IHW #IHU #T1 #f #H
elim (lifts_inv_flat2 … H) -H #W1 #U1 #HW1 #HU1 #H destruct
/3 width=3 by lifts_flat/
]
(* Basic_1: was: lift1_lift1 (left to right) *)
(* Basic_1: includes: lift_free (left to right) lift_d lift1_xhg (right to left) lift1_free (right to left) *)
(* Basic_2A1: includes: lift_trans_be lift_trans_le lift_trans_ge lifts_lift_trans_le lifts_lift_trans *)
-theorem lifts_trans: ∀T1,T,f1. ⬆*[f1] T1 ≡ T → ∀T2,f2. ⬆*[f2] T ≡ T2 →
+theorem lifts_trans: ∀f1,T1,T. ⬆*[f1] T1 ≡ T → ∀T2,f2. ⬆*[f2] T ≡ T2 →
∀f. f2 ⊚ f1 ≡ f → ⬆*[f] T1 ≡ T2.
-#T1 #T #f1 #H elim H -T1 -T -f1
-[ #s #f1 #T2 #f2 #H >(lifts_inv_sort1 … H) -T2 //
-| #i1 #i #f1 #Hi1 #T2 #f2 #H #f #Ht21 elim (lifts_inv_lref1 … H) -H
+#f1 #T1 #T #H elim H -f1 -T1 -T
+[ #f1 #s #T2 #f2 #H >(lifts_inv_sort1 … H) -T2 //
+| #f1 #i1 #i #Hi1 #T2 #f2 #H #f #Ht21 elim (lifts_inv_lref1 … H) -H
#i2 #Hi2 #H destruct /3 width=6 by lifts_lref, after_fwd_at/
-| #l #f1 #T2 #f2 #H >(lifts_inv_gref1 … H) -T2 //
-| #p #I #W1 #W #U1 #U #f1 #_ #_ #IHW #IHU #T2 #f2 #H
+| #f1 #l #T2 #f2 #H >(lifts_inv_gref1 … H) -T2 //
+| #f1 #p #I #W1 #W #U1 #U #_ #_ #IHW #IHU #T2 #f2 #H
elim (lifts_inv_bind1 … H) -H #W2 #U2 #HW2 #HU2 #H destruct
/4 width=3 by lifts_bind, after_O2/
-| #I #W1 #W #U1 #U #f1 #_ #_ #IHW #IHU #T2 #f2 #H
+| #f1 #I #W1 #W #U1 #U #_ #_ #IHW #IHU #T2 #f2 #H
elim (lifts_inv_flat1 … H) -H #W2 #U2 #HW2 #HU2 #H destruct
/3 width=3 by lifts_flat/
]
qed-.
(* Basic_2A1: includes: lift_conf_O1 lift_conf_be *)
-theorem lifts_conf: ∀T,T1,f1. ⬆*[f1] T ≡ T1 → ∀T2,f. ⬆*[f] T ≡ T2 →
+theorem lifts_conf: ∀f1,T,T1. ⬆*[f1] T ≡ T1 → ∀T2,f. ⬆*[f] T ≡ T2 →
∀f2. f2 ⊚ f1 ≡ f → ⬆*[f2] T1 ≡ T2.
-#T #T1 #f1 #H elim H -T -T1 -f1
-[ #s #f1 #T2 #f #H >(lifts_inv_sort1 … H) -T2 //
-| #i #i1 #f1 #Hi1 #T2 #f #H #f2 #Ht21 elim (lifts_inv_lref1 … H) -H
+#f1 #T #T1 #H elim H -f1 -T -T1
+[ #f1 #s #T2 #f #H >(lifts_inv_sort1 … H) -T2 //
+| #f1 #i #i1 #Hi1 #T2 #f #H #f2 #Ht21 elim (lifts_inv_lref1 … H) -H
#i2 #Hi2 #H destruct /3 width=6 by lifts_lref, after_fwd_at2/
-| #l #f1 #T2 #f #H >(lifts_inv_gref1 … H) -T2 //
-| #p #I #W #W1 #U #U1 #f1 #_ #_ #IHW #IHU #T2 #f #H
+| #f1 #l #T2 #f #H >(lifts_inv_gref1 … H) -T2 //
+| #f1 #p #I #W #W1 #U #U1 #_ #_ #IHW #IHU #T2 #f #H
elim (lifts_inv_bind1 … H) -H #W2 #U2 #HW2 #HU2 #H destruct
/4 width=3 by lifts_bind, after_O2/
-| #I #W #W1 #U #U1 #f1 #_ #_ #IHW #IHU #T2 #f #H
+| #f1 #I #W #W1 #U #U1 #_ #_ #IHW #IHU #T2 #f #H
elim (lifts_inv_flat1 … H) -H #W2 #U2 #HW2 #HU2 #H destruct
/3 width=3 by lifts_flat/
]
(* Advanced proprerties *****************************************************)
(* Basic_2A1: includes: lift_inj *)
-lemma lifts_inj: ∀T1,U,f. ⬆*[f] T1 ≡ U → ∀T2. ⬆*[f] T2 ≡ U → T1 = T2.
-#T1 #U #f #H1 #T2 #H2 lapply (isid_after_dx 𝐈𝐝 … f)
+lemma lifts_inj: ∀f,T1,U. ⬆*[f] T1 ≡ U → ∀T2. ⬆*[f] T2 ≡ U → T1 = T2.
+#f #T1 #U #H1 #T2 #H2 lapply (isid_after_dx 𝐈𝐝 … f)
/3 width=6 by lifts_div, lifts_fwd_isid/
qed-.
(* Basic_2A1: includes: lift_mono *)
-lemma lifts_mono: ∀T,U1,f. ⬆*[f] T ≡ U1 → ∀U2. ⬆*[f] T ≡ U2 → U1 = U2.
-#T #U1 #f #H1 #U2 #H2 lapply (isid_after_sn 𝐈𝐝 … f)
+lemma lifts_mono: ∀f,T,U1. ⬆*[f] T ≡ U1 → ∀U2. ⬆*[f] T ≡ U2 → U1 = U2.
+#f #T #U1 #H1 #U2 #H2 lapply (isid_after_sn 𝐈𝐝 … f)
/3 width=6 by lifts_conf, lifts_fwd_isid/
qed-.
(* Main properties **********************************************************)
(* Basic_1: includes: lifts_inj *)
-theorem liftsv_inj: ∀T1s,Us,f. ⬆*[f] T1s ≡ Us →
+theorem liftsv_inj: ∀f,T1s,Us. ⬆*[f] T1s ≡ Us →
∀T2s. ⬆*[f] T2s ≡ Us → T1s = T2s.
-#T1s #Us #f #H elim H -T1s -Us
+#f #T1s #Us #H elim H -T1s -Us
[ #T2s #H >(liftsv_inv_nil2 … H) -H //
| #T1s #Us #T1 #U #HT1U #_ #IHT1Us #X #H destruct
elim (liftsv_inv_cons2 … H) -H #T2 #T2s #HT2U #HT2Us #H destruct
qed-.
(* Basic_2A1: includes: liftv_mono *)
-theorem liftsv_mono: ∀Ts,U1s,f. ⬆*[f] Ts ≡ U1s →
+theorem liftsv_mono: ∀f,Ts,U1s. ⬆*[f] Ts ≡ U1s →
∀U2s. ⬆*[f] Ts ≡ U2s → U1s = U2s.
-#Ts #U1s #f #H elim H -Ts -U1s
+#f #Ts #U1s #H elim H -Ts -U1s
[ #U2s #H >(liftsv_inv_nil1 … H) -H //
| #Ts #U1s #T #U1 #HTU1 #_ #IHTU1s #X #H destruct
elim (liftsv_inv_cons1 … H) -H #U2 #U2s #HTU2 #HTU2s #H destruct
(* Basic_1: includes: lifts1_xhg (right to left) *)
(* Basic_2A1: includes: liftsv_liftv_trans_le *)
-theorem liftsv_trans: ∀T1s,Ts,f1. ⬆*[f1] T1s ≡ Ts → ∀T2s,f2. ⬆*[f2] Ts ≡ T2s →
+theorem liftsv_trans: ∀f1,T1s,Ts. ⬆*[f1] T1s ≡ Ts → ∀T2s,f2. ⬆*[f2] Ts ≡ T2s →
∀f. f2 ⊚ f1 ≡ f → ⬆*[f] T1s ≡ T2s.
-#T1s #Ts #f1 #H elim H -T1s -Ts
+#f1 #T1s #Ts #H elim H -T1s -Ts
[ #T2s #f2 #H >(liftsv_inv_nil1 … H) -T2s /2 width=3 by liftsv_nil/
| #T1s #Ts #T1 #T #HT1 #_ #IHT1s #X #f2 #H elim (liftsv_inv_cons1 … H) -H
#T2 #T2s #HT2 #HT2s #H destruct /3 width=6 by lifts_trans, liftsv_cons/
(* Forward lemmas with simple terms *****************************************)
(* Basic_2A1: includes: lift_simple_dx *)
-lemma lifts_simple_dx: ∀T1,T2,f. ⬆*[f] T1 ≡ T2 → 𝐒⦃T1⦄ → 𝐒⦃T2⦄.
-#T1 #T2 #f #H elim H -T1 -T2 -f //
-#p #I #V1 #V2 #T1 #T2 #f #_ #_ #_ #_ #H elim (simple_inv_bind … H)
+lemma lifts_simple_dx: ∀f,T1,T2. ⬆*[f] T1 ≡ T2 → 𝐒⦃T1⦄ → 𝐒⦃T2⦄.
+#f #T1 #T2 #H elim H -f -T1 -T2 //
+#f #p #I #V1 #V2 #T1 #T2 #_ #_ #_ #_ #H elim (simple_inv_bind … H)
qed-.
(* Basic_2A1: includes: lift_simple_sn *)
-lemma lifts_simple_sn: ∀T1,T2,f. ⬆*[f] T1 ≡ T2 → 𝐒⦃T2⦄ → 𝐒⦃T1⦄.
-#T1 #T2 #f #H elim H -T1 -T2 -f //
-#p #I #V1 #V2 #T1 #T2 #f #_ #_ #_ #_ #H elim (simple_inv_bind … H)
+lemma lifts_simple_sn: ∀f,T1,T2. ⬆*[f] T1 ≡ T2 → 𝐒⦃T2⦄ → 𝐒⦃T1⦄.
+#f #T1 #T2 #H elim H -f -T1 -T2 //
+#f #p #I #V1 #V2 #T1 #T2 #_ #_ #_ #_ #H elim (simple_inv_bind … H)
qed-.
(* Basic inversion lemmas ***************************************************)
-fact liftsv_inv_nil1_aux: ∀X,Y,f. ⬆*[f] X ≡ Y → X = ◊ → Y = ◊.
-#X #Y #f * -X -Y //
+fact liftsv_inv_nil1_aux: ∀f,X,Y. ⬆*[f] X ≡ Y → X = ◊ → Y = ◊.
+#f #X #Y * -X -Y //
#T1s #T2s #T1 #T2 #_ #_ #H destruct
qed-.
(* Basic_2A1: includes: liftv_inv_nil1 *)
-lemma liftsv_inv_nil1: ∀Y,f. ⬆*[f] ◊ ≡ Y → Y = ◊.
+lemma liftsv_inv_nil1: ∀f,Y. ⬆*[f] ◊ ≡ Y → Y = ◊.
/2 width=5 by liftsv_inv_nil1_aux/ qed-.
-fact liftsv_inv_cons1_aux: ∀X,Y. ∀f:rtmap. ⬆*[f] X ≡ Y →
+fact liftsv_inv_cons1_aux: ∀f:rtmap. ∀X,Y. ⬆*[f] X ≡ Y →
∀T1,T1s. X = T1 @ T1s →
∃∃T2,T2s. ⬆*[f] T1 ≡ T2 & ⬆*[f] T1s ≡ T2s &
Y = T2 @ T2s.
-#X #Y #f * -X -Y
+#f #X #Y * -X -Y
[ #U1 #U1s #H destruct
| #T1s #T2s #T1 #T2 #HT12 #HT12s #U1 #U1s #H destruct /2 width=5 by ex3_2_intro/
]
qed-.
(* Basic_2A1: includes: liftv_inv_cons1 *)
-lemma liftsv_inv_cons1: ∀T1,T1s,Y. ∀f:rtmap. ⬆*[f] T1 @ T1s ≡ Y →
+lemma liftsv_inv_cons1: ∀f:rtmap. ∀T1,T1s,Y. ⬆*[f] T1 @ T1s ≡ Y →
∃∃T2,T2s. ⬆*[f] T1 ≡ T2 & ⬆*[f] T1s ≡ T2s &
Y = T2 @ T2s.
/2 width=3 by liftsv_inv_cons1_aux/ qed-.
-fact liftsv_inv_nil2_aux: ∀X,Y,f. ⬆*[f] X ≡ Y → Y = ◊ → X = ◊.
-#X #Y #f * -X -Y //
+fact liftsv_inv_nil2_aux: ∀f,X,Y. ⬆*[f] X ≡ Y → Y = ◊ → X = ◊.
+#f #X #Y * -X -Y //
#T1s #T2s #T1 #T2 #_ #_ #H destruct
qed-.
-lemma liftsv_inv_nil2: ∀X,f. ⬆*[f] X ≡ ◊ → X = ◊.
+lemma liftsv_inv_nil2: ∀f,X. ⬆*[f] X ≡ ◊ → X = ◊.
/2 width=5 by liftsv_inv_nil2_aux/ qed-.
-fact liftsv_inv_cons2_aux: ∀X,Y. ∀f:rtmap. ⬆*[f] X ≡ Y →
+fact liftsv_inv_cons2_aux: ∀f:rtmap. ∀X,Y. ⬆*[f] X ≡ Y →
∀T2,T2s. Y = T2 @ T2s →
∃∃T1,T1s. ⬆*[f] T1 ≡ T2 & ⬆*[f] T1s ≡ T2s &
X = T1 @ T1s.
-#X #Y #f * -X -Y
+#f #X #Y * -X -Y
[ #U2 #U2s #H destruct
| #T1s #T2s #T1 #T2 #HT12 #HT12s #U2 #U2s #H destruct /2 width=5 by ex3_2_intro/
]
qed-.
-lemma liftsv_inv_cons2: ∀X,T2,T2s. ∀f:rtmap. ⬆*[f] X ≡ T2 @ T2s →
+lemma liftsv_inv_cons2: ∀f:rtmap. ∀X,T2,T2s. ⬆*[f] X ≡ T2 @ T2s →
∃∃T1,T1s. ⬆*[f] T1 ≡ T2 & ⬆*[f] T1s ≡ T2s &
X = T1 @ T1s.
/2 width=3 by liftsv_inv_cons2_aux/ qed-.
(* Basic_1: was: lifts1_flat (left to right) *)
-lemma lifts_inv_applv1: ∀V1s,U1,T2. ∀f:rtmap. ⬆*[f] Ⓐ V1s.U1 ≡ T2 →
+lemma lifts_inv_applv1: ∀f:rtmap. ∀V1s,U1,T2. ⬆*[f] Ⓐ V1s.U1 ≡ T2 →
∃∃V2s,U2. ⬆*[f] V1s ≡ V2s & ⬆*[f] U1 ≡ U2 &
T2 = Ⓐ V2s.U2.
-#V1s elim V1s -V1s
+#f #V1s elim V1s -V1s
[ /3 width=5 by ex3_2_intro, liftsv_nil/
-| #V1 #V1s #IHV1s #T1 #X #f #H elim (lifts_inv_flat1 … H) -H
+| #V1 #V1s #IHV1s #T1 #X #H elim (lifts_inv_flat1 … H) -H
#V2 #Y #HV12 #HY #H destruct elim (IHV1s … HY) -IHV1s -HY
- #V2c #T2 #HV12s #HT12 #H destruct /3 width=5 by ex3_2_intro, liftsv_cons/
+ #V2s #T2 #HV12s #HT12 #H destruct /3 width=5 by ex3_2_intro, liftsv_cons/
]
qed-.
-lemma lifts_inv_applv2: ∀V2s,U2,T1. ∀f:rtmap. ⬆*[f] T1 ≡ Ⓐ V2s.U2 →
+lemma lifts_inv_applv2: ∀f:rtmap. ∀V2s,U2,T1. ⬆*[f] T1 ≡ Ⓐ V2s.U2 →
∃∃V1s,U1. ⬆*[f] V1s ≡ V2s & ⬆*[f] U1 ≡ U2 &
T1 = Ⓐ V1s.U1.
-#V2s elim V2s -V2s
+#f #V2s elim V2s -V2s
[ /3 width=5 by ex3_2_intro, liftsv_nil/
-| #V2 #V2s #IHV2s #T2 #X #f #H elim (lifts_inv_flat2 … H) -H
+| #V2 #V2s #IHV2s #T2 #X #H elim (lifts_inv_flat2 … H) -H
#V1 #Y #HV12 #HY #H destruct elim (IHV2s … HY) -IHV2s -HY
#V1s #T1 #HV12s #HT12 #H destruct /3 width=5 by ex3_2_intro, liftsv_cons/
]
qed-.
(* Basic_1: was: lifts1_flat (right to left) *)
-lemma lifts_applv: ∀V1s,V2s. ∀f:rtmap. ⬆*[f] V1s ≡ V2s →
+lemma lifts_applv: ∀f:rtmap. ∀V1s,V2s. ⬆*[f] V1s ≡ V2s →
∀T1,T2. ⬆*[f] T1 ≡ T2 →
⬆*[f] Ⓐ V1s.T1 ≡ Ⓐ V2s.T2.
-#V1s #V2s #f #H elim H -V1s -V2s /3 width=1 by lifts_flat/
+#f #V1s #V2s #H elim H -V1s -V2s /3 width=1 by lifts_flat/
qed.
(* Basic_1: removed theorems 2: lifts1_nil lifts1_cons *)
(* Forward lemmas with weight for terms *************************************)
(* Basic_2A1: includes: lift_fwd_tw *)
-lemma lifts_fwd_tw: ∀T1,T2,f. ⬆*[f] T1 ≡ T2 → ♯{T1} = ♯{T2}.
-#T1 #T2 #f #H elim H -T1 -T2 -f normalize //
+lemma lifts_fwd_tw: ∀f,T1,T2. ⬆*[f] T1 ≡ T2 → ♯{T1} = ♯{T2}.
+#f #T1 #T2 #H elim H -f -T1 -T2 normalize //
qed-.
lemma lreq_eq_repl_fwd: ∀L1,L2. eq_repl_fwd … (λf. L1 ≡[f] L2).
/2 width=3 by lexs_eq_repl_fwd/ qed-.
-lemma sle_lreq_trans: ∀L1,L2,f2. L1 ≡[f2] L2 →
+lemma sle_lreq_trans: ∀f2,L1,L2. L1 ≡[f2] L2 →
∀f1. f1 ⊆ f2 → L1 ≡[f1] L2.
/2 width=3 by sle_lexs_trans/ qed-.
(* Basic inversion lemmas ***************************************************)
(* Basic_2A1: includes: lreq_inv_atom1 *)
-lemma lreq_inv_atom1: ∀Y,f. ⋆ ≡[f] Y → Y = ⋆.
+lemma lreq_inv_atom1: ∀f,Y. ⋆ ≡[f] Y → Y = ⋆.
/2 width=4 by lexs_inv_atom1/ qed-.
(* Basic_2A1: includes: lreq_inv_pair1 *)
-lemma lreq_inv_next1: ∀J,K1,Y,W1,g. K1.ⓑ{J}W1 ≡[⫯g] Y →
+lemma lreq_inv_next1: ∀g,J,K1,Y,W1. K1.ⓑ{J}W1 ≡[⫯g] Y →
∃∃K2. K1 ≡[g] K2 & Y = K2.ⓑ{J}W1.
-#J #K1 #Y #W1 #g #H elim (lexs_inv_next1 … H) -H /2 width=3 by ex2_intro/
+#g #J #K1 #Y #W1 #H elim (lexs_inv_next1 … H) -H /2 width=3 by ex2_intro/
qed-.
(* Basic_2A1: includes: lreq_inv_zero1 lreq_inv_succ1 *)
-lemma lreq_inv_push1: ∀J,K1,Y,W1,g. K1.ⓑ{J}W1 ≡[↑g] Y →
+lemma lreq_inv_push1: ∀g,J,K1,Y,W1. K1.ⓑ{J}W1 ≡[↑g] Y →
∃∃K2,W2. K1 ≡[g] K2 & Y = K2.ⓑ{J}W2.
-#J #K1 #Y #W1 #g #H elim (lexs_inv_push1 … H) -H /2 width=4 by ex2_2_intro/ qed-.
+#g #J #K1 #Y #W1 #H elim (lexs_inv_push1 … H) -H /2 width=4 by ex2_2_intro/ qed-.
(* Basic_2A1: includes: lreq_inv_atom2 *)
-lemma lreq_inv_atom2: ∀X,f. X ≡[f] ⋆ → X = ⋆.
+lemma lreq_inv_atom2: ∀f,X. X ≡[f] ⋆ → X = ⋆.
/2 width=4 by lexs_inv_atom2/ qed-.
(* Basic_2A1: includes: lreq_inv_pair2 *)
-lemma lreq_inv_next2: ∀J,X,K2,W2,g. X ≡[⫯g] K2.ⓑ{J}W2 →
+lemma lreq_inv_next2: ∀g,J,X,K2,W2. X ≡[⫯g] K2.ⓑ{J}W2 →
∃∃K1. K1 ≡[g] K2 & X = K1.ⓑ{J}W2.
-#J #X #K2 #W2 #g #H elim (lexs_inv_next2 … H) -H /2 width=3 by ex2_intro/ qed-.
+#g #J #X #K2 #W2 #H elim (lexs_inv_next2 … H) -H /2 width=3 by ex2_intro/ qed-.
(* Basic_2A1: includes: lreq_inv_zero2 lreq_inv_succ2 *)
-lemma lreq_inv_push2: ∀J,X,K2,W2,g. X ≡[↑g] K2.ⓑ{J}W2 →
+lemma lreq_inv_push2: ∀g,J,X,K2,W2. X ≡[↑g] K2.ⓑ{J}W2 →
∃∃K1,W1. K1 ≡[g] K2 & X = K1.ⓑ{J}W1.
-#J #X #K2 #W2 #g #H elim (lexs_inv_push2 … H) -H /2 width=4 by ex2_2_intro/ qed-.
+#g #J #X #K2 #W2 #H elim (lexs_inv_push2 … H) -H /2 width=4 by ex2_2_intro/ qed-.
(* Basic_2A1: includes: lreq_inv_pair *)
-lemma lreq_inv_next: ∀I1,I2,L1,L2,V1,V2,f.
+lemma lreq_inv_next: ∀f,I1,I2,L1,L2,V1,V2.
L1.ⓑ{I1}V1 ≡[⫯f] (L2.ⓑ{I2}V2) →
∧∧ L1 ≡[f] L2 & V1 = V2 & I1 = I2.
/2 width=1 by lexs_inv_next/ qed-.
(* Basic_2A1: includes: lreq_inv_succ *)
-lemma lreq_inv_push: ∀I1,I2,L1,L2,V1,V2,f.
+lemma lreq_inv_push: ∀f,I1,I2,L1,L2,V1,V2.
L1.ⓑ{I1}V1 ≡[↑f] (L2.ⓑ{I2}V2) →
L1 ≡[f] L2 ∧ I1 = I2.
-#I1 #I2 #L1 #L2 #V1 #V2 #f #H elim (lexs_inv_push … H) -H /2 width=1 by conj/
+#f #I1 #I2 #L1 #L2 #V1 #V2 #H elim (lexs_inv_push … H) -H /2 width=1 by conj/
qed-.
-lemma lreq_inv_tl: ∀I,L1,L2,V,f. L1 ≡[⫱f] L2 → L1.ⓑ{I}V ≡[f] L2.ⓑ{I}V.
+lemma lreq_inv_tl: ∀f,I,L1,L2,V. L1 ≡[⫱f] L2 → L1.ⓑ{I}V ≡[f] L2.ⓑ{I}V.
/2 width=1 by lexs_inv_tl/ qed-.
(* Basic_2A1: removed theorems 5:
(* Forward lemmas with length for local environments ************************)
(* Basic_2A1: includes: lreq_fwd_length *)
-lemma lreq_fwd_length: ∀L1,L2,f. L1 ≡[f] L2 → |L1| = |L2|.
+lemma lreq_fwd_length: ∀f,L1,L2. L1 ≡[f] L2 → |L1| = |L2|.
/2 width=4 by lexs_fwd_length/ qed-.
theorem lreq_canc_dx: ∀f. right_cancellable … (lreq f).
/3 width=3 by lexs_canc_dx, lreq_trans, lreq_sym/ qed-.
-theorem lreq_join: ∀L1,L2,f1. L1 ≡[f1] L2 → ∀f2. L1 ≡[f2] L2 →
+theorem lreq_join: ∀f1,L1,L2. L1 ≡[f1] L2 → ∀f2. L1 ≡[f2] L2 →
∀f. f1 ⋓ f2 ≡ f → L1 ≡[f] L2.
/2 width=5 by lexs_join/ qed-.
-theorem lreq_meet: ∀L1,L2,f1. L1 ≡[f1] L2 → ∀f2. L1 ≡[f2] L2 →
+theorem lreq_meet: ∀f1,L1,L2. L1 ≡[f1] L2 → ∀f2. L1 ≡[f2] L2 →
∀f. f1 ⋒ f2 ≡ f → L1 ≡[f] L2.
/2 width=5 by lexs_meet/ qed-.
(* Basic_1: was just: pr3_pr2_pr2_t *)
(* Basic_1: includes: pr3_pr0_pr2_t *)
-lemma lpr_cpr_trans: ∀G. c_r_transitive … (cpr G) (λ_. lpr G).
+lemma lpr_cpr_trans: ∀G. b_c_transitive … (cpr G) (λ_. lpr G).
#G #L2 #T1 #T2 #HT12 elim HT12 -G -L2 -T1 -T2
[ /2 width=3 by/
| #G #L2 #K2 #V0 #V2 #W2 #i #HLK2 #_ #HVW2 #IHV02 #L1 #HL12
(* Advanced properties ******************************************************)
(* Basic_1: was only: pr3_pr2_pr3_t pr3_wcpr0_t *)
-lemma lpr_cprs_trans: ∀G. c_rs_transitive … (cpr G) (λ_. lpr G).
-#G @c_r_trans_LTC1 /2 width=3 by lpr_cpr_trans/ (**) (* full auto fails *)
+lemma lpr_cprs_trans: ∀G. b_rs_transitive … (cpr G) (λ_. lpr G).
+#G @b_c_trans_LTC1 /2 width=3 by lpr_cpr_trans/ (**) (* full auto fails *)
qed-.
(* Basic_1: was: pr3_strip *)
@(cpxs_strap1 … (⋆(iter d2 ℕ (next h) s)))
[ /3 width=3 by lt_to_le/
| @(cpx_st … (d1-d2-1)) <plus_minus_k_k
- /2 width=1 by deg_iter, monotonic_le_minus_r/
+ /2 width=1 by deg_iter, monotonic_le_minus_c/
]
qed.
(* Properties on sn extended parallel reduction for local environments ******)
-lemma lpx_cpx_trans: ∀h,o,G. c_r_transitive … (cpx h o G) (λ_.lpx h o G).
+lemma lpx_cpx_trans: ∀h,o,G. b_c_transitive … (cpx h o G) (λ_.lpx h o G).
#h #o #G #L2 #T1 #T2 #HT12 elim HT12 -G -L2 -T1 -T2
[ /2 width=3 by/
| /3 width=2 by cpx_cpxs, cpx_st/
(* Advanced properties ******************************************************)
-lemma lpx_cpxs_trans: ∀h,o,G. c_rs_transitive … (cpx h o G) (λ_.lpx h o G).
-#h #o #G @c_r_trans_LTC1 /2 width=3 by lpx_cpx_trans/ (**) (* full auto fails *)
+lemma lpx_cpxs_trans: ∀h,o,G. b_rs_transitive … (cpx h o G) (λ_.lpx h o G).
+#h #o #G @b_c_trans_LTC1 /2 width=3 by lpx_cpx_trans/ (**) (* full auto fails *)
qed-.
lemma cpxs_bind2_dx: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, o] V2 →
| #G #L #K #V1 #V2 #W2 #i #d2 #HLK #_ #HVW2 #IHV12 #d1 #H #Hd21
elim (da_inv_lref … H) -H * #K0 #V0 [| #d0 ] #HLK0
lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct
- #HV1 #H destruct lapply (le_plus_to_le_r … Hd21) -Hd21
+ #HV1 #H destruct lapply (le_plus_to_le_c … Hd21) -Hd21
/3 width=7 by cpxs_delta/
| /4 width=3 by cpxs_bind_dx, da_inv_bind/
| /4 width=3 by cpxs_flat_dx, da_inv_flat/
qed-.
(* Basic_1: was just: pr3_iso_appls_abbr *)
-lemma cpxs_fwd_theta_vector: ∀h,o,G,L,V1c,V2c. ⬆[0, 1] V1c ≡ V2c →
- ∀a,V,T,U. ⦃G, L⦄ ⊢ ⒶV1c.ⓓ{a}V.T ➡*[h, o] U →
- ⒶV1c. ⓓ{a}V. T ≂ U ∨ ⦃G, L⦄ ⊢ ⓓ{a}V.ⒶV2c.T ➡*[h, o] U.
-#h #o #G #L #V1c #V2c * -V1c -V2c /3 width=1 by or_intror/
-#V1c #V2c #V1a #V2a #HV12a #HV12c #a
+lemma cpxs_fwd_theta_vector: ∀h,o,G,L,V1b,V2b. ⬆[0, 1] V1b ≡ V2b →
+ ∀a,V,T,U. ⦃G, L⦄ ⊢ ⒶV1b.ⓓ{a}V.T ➡*[h, o] U →
+ ⒶV1b. ⓓ{a}V. T ≂ U ∨ ⦃G, L⦄ ⊢ ⓓ{a}V.ⒶV2b.T ➡*[h, o] U.
+#h #o #G #L #V1b #V2b * -V1b -V2b /3 width=1 by or_intror/
+#V1b #V2b #V1a #V2a #HV12a #HV12b #a
generalize in match HV12a; -HV12a
generalize in match V2a; -V2a
generalize in match V1a; -V1a
-elim HV12c -V1c -V2c /2 width=1 by cpxs_fwd_theta/
-#V1c #V2c #V1b #V2b #HV12b #_ #IHV12c #V1a #V2a #HV12a #V #T #U #H
+elim HV12b -V1b -V2b /2 width=1 by cpxs_fwd_theta/
+#V1b #V2b #V1b #V2b #HV12b #_ #IHV12b #V1a #V2a #HV12a #V #T #U #H
elim (cpxs_inv_appl1 … H) -H *
-[ -IHV12c -HV12a -HV12b #V0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or_introl/
+[ -IHV12b -HV12a -HV12b #V0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or_introl/
| #b #W0 #T0 #HT0 #HU
- elim (IHV12c … HV12b … HT0) -IHV12c -HT0 #HT0
+ elim (IHV12b … HV12b … HT0) -IHV12b -HT0 #HT0
[ -HV12a -HV12b -HU
elim (tsts_inv_pair1 … HT0) #V1 #T1 #H destruct
- | @or_intror -V1c (**) (* explicit constructor *)
+ | @or_intror -V1b (**) (* explicit constructor *)
@(cpxs_trans … HU) -U
elim (cpxs_inv_abbr1 … HT0) -HT0 *
[ -HV12a -HV12b #V1 #T1 #_ #_ #H destruct
| -V1b #X #HT1 #H #H0 destruct
elim (lift_inv_bind1 … H) -H #W1 #T1 #HW01 #HT01 #H destruct
- @(cpxs_trans … (+ⓓV.ⓐV2a.ⓛ{b}W1.T1)) [ /3 width=1 by cpxs_flat_dx, cpxs_bind_dx/ ] -T -V2b -V2c
+ @(cpxs_trans … (+ⓓV.ⓐV2a.ⓛ{b}W1.T1)) [ /3 width=1 by cpxs_flat_dx, cpxs_bind_dx/ ] -T -V2b -V2b
@(cpxs_strap2 … (ⓐV1a.ⓛ{b}W0.T0))
/4 width=7 by cpxs_beta_dx, cpx_zeta, lift_bind, lift_flat/
]
]
| #b #V0a #Va #V0 #T0 #HV10a #HV0a #HT0 #HU
- elim (IHV12c … HV12b … HT0) -HV12b -IHV12c -HT0 #HT0
+ elim (IHV12b … HV12b … HT0) -HV12b -IHV12b -HT0 #HT0
[ -HV12a -HV10a -HV0a -HU
elim (tsts_inv_pair1 … HT0) #V1 #T1 #H destruct
- | @or_intror -V1c -V1b (**) (* explicit constructor *)
+ | @or_intror -V1b -V1b (**) (* explicit constructor *)
@(cpxs_trans … HU) -U
elim (cpxs_inv_abbr1 … HT0) -HT0 *
[ #V1 #T1 #HV1 #HT1 #H destruct
| #X #HT1 #H #H0 destruct
elim (lift_inv_bind1 … H) -H #V1 #T1 #HW01 #HT01 #H destruct
lapply (cpxs_lift … HV10a (L.ⓓV0) (Ⓕ) … HV12a … HV0a) -V0a [ /2 width=1 by drop_drop/ ] #HV2a
- @(cpxs_trans … (+ⓓV.ⓐV2a.ⓓ{b}V1.T1)) [ /3 width=1 by cpxs_flat_dx, cpxs_bind_dx/ ] -T -V2b -V2c
+ @(cpxs_trans … (+ⓓV.ⓐV2a.ⓓ{b}V1.T1)) [ /3 width=1 by cpxs_flat_dx, cpxs_bind_dx/ ] -T -V2b -V2b
@(cpxs_strap2 … (ⓐV1a.ⓓ{b}V0.T0)) [ /4 width=7 by cpx_zeta, lift_bind, lift_flat/ ] -V -V1 -T1
@(cpxs_strap2 … (ⓓ{b}V0.ⓐV2a.T0)) /3 width=3 by cpxs_pair_sn, cpxs_bind_dx, cpr_cpx, cpr_theta/
]
(* Relocation properties ****************************************************)
(* Basic_1: was just: sn3_lift *)
-lemma csx_lift: ∀h,o,G,L2,L1,T1,c,l,k. ⦃G, L1⦄ ⊢ ⬊*[h, o] T1 →
- ∀T2. ⬇[c, l, k] L2 ≡ L1 → ⬆[l, k] T1 ≡ T2 → ⦃G, L2⦄ ⊢ ⬊*[h, o] T2.
-#h #o #G #L2 #L1 #T1 #c #l #k #H elim H -T1 #T1 #_ #IHT1 #T2 #HL21 #HT12
+lemma csx_lift: ∀h,o,G,L2,L1,T1,b,l,k. ⦃G, L1⦄ ⊢ ⬊*[h, o] T1 →
+ ∀T2. ⬇[b, l, k] L2 ≡ L1 → ⬆[l, k] T1 ≡ T2 → ⦃G, L2⦄ ⊢ ⬊*[h, o] T2.
+#h #o #G #L2 #L1 #T1 #b #l #k #H elim H -T1 #T1 #_ #IHT1 #T2 #HL21 #HT12
@csx_intro #T #HLT2 #HT2
elim (cpx_inv_lift1 … HLT2 … HL21 … HT12) -HLT2 #T0 #HT0 #HLT10
@(IHT1 … HLT10) // -L1 -L2 #H destruct
qed.
(* Basic_1: was just: sn3_gen_lift *)
-lemma csx_inv_lift: ∀h,o,G,L2,L1,T1,c,l,k. ⦃G, L1⦄ ⊢ ⬊*[h, o] T1 →
- ∀T2. ⬇[c, l, k] L1 ≡ L2 → ⬆[l, k] T2 ≡ T1 → ⦃G, L2⦄ ⊢ ⬊*[h, o] T2.
-#h #o #G #L2 #L1 #T1 #c #l #k #H elim H -T1 #T1 #_ #IHT1 #T2 #HL12 #HT21
+lemma csx_inv_lift: ∀h,o,G,L2,L1,T1,b,l,k. ⦃G, L1⦄ ⊢ ⬊*[h, o] T1 →
+ ∀T2. ⬇[b, l, k] L1 ≡ L2 → ⬆[l, k] T2 ≡ T1 → ⦃G, L2⦄ ⊢ ⬊*[h, o] T2.
+#h #o #G #L2 #L1 #T1 #b #l #k #H elim H -T1 #T1 #_ #IHT1 #T2 #HL12 #HT21
@csx_intro #T #HLT2 #HT2
elim (lift_total T l k) #T0 #HT0
lapply (cpx_lift … HLT2 … HL12 … HT21 … HT0) -HLT2 #HLT10
qed.
(* Basic_1: was just: sn3_appls_abbr *)
-lemma csx_applv_theta: ∀h,o,a,G,L,V1c,V2c. ⬆[0, 1] V1c ≡ V2c →
- ∀V,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⓓ{a}V.ⒶV2c.T →
- ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶV1c.ⓓ{a}V.T.
-#h #o #a #G #L #V1c #V2c * -V1c -V2c /2 width=1 by/
-#V1c #V2c #V1 #V2 #HV12 #H
+lemma csx_applv_theta: ∀h,o,a,G,L,V1b,V2b. ⬆[0, 1] V1b ≡ V2b →
+ ∀V,T. ⦃G, L⦄ ⊢ ⬊*[h, o] ⓓ{a}V.ⒶV2b.T →
+ ⦃G, L⦄ ⊢ ⬊*[h, o] ⒶV1b.ⓓ{a}V.T.
+#h #o #a #G #L #V1b #V2b * -V1b -V2b /2 width=1 by/
+#V1b #V2b #V1 #V2 #HV12 #H
generalize in match HV12; -HV12 generalize in match V2; -V2 generalize in match V1; -V1
-elim H -V1c -V2c /2 width=3 by csx_appl_theta/
-#V1c #V2c #V1 #V2 #HV12 #HV12c #IHV12c #W1 #W2 #HW12 #V #T #H
+elim H -V1b -V2b /2 width=3 by csx_appl_theta/
+#V1b #V2b #V1 #V2 #HV12 #HV12b #IHV12b #W1 #W2 #HW12 #V #T #H
lapply (csx_appl_theta … HW12 … H) -H -HW12 #H
lapply (csx_fwd_pair_sn … H) #HW1
lapply (csx_fwd_flat_dx … H) #H1
-@csx_appl_simple_tsts /2 width=3 by simple_flat/ -IHV12c -HW1 -H1 #X #H1 #H2
-elim (cpxs_fwd_theta_vector … (V2@V2c) … H1) -H1 /2 width=1 by liftv_cons/ -HV12c -HV12
+@csx_appl_simple_tsts /2 width=3 by simple_flat/ -IHV12b -HW1 -H1 #X #H1 #H2
+elim (cpxs_fwd_theta_vector … (V2@V2b) … H1) -H1 /2 width=1 by liftv_cons/ -HV12b -HV12
[ -H #H elim H2 -H2 //
| -H2 /3 width=5 by csx_cpxs_trans, cpxs_flat_dx/
]
[ /3 width=1 by csx_applv_cnx/
|2,3,6: /2 width=1 by csx_applv_beta, csx_applv_sort, csx_applv_cast/
| /2 width=7 by csx_applv_delta/
-| #G #L #V1c #V2c #HV12c #a #V #T #H #HV
- @(csx_applv_theta … HV12c) -HV12c
+| #G #L #V1b #V2b #HV12b #a #V #T #H #HV
+ @(csx_applv_theta … HV12b) -HV12b
@csx_abbr //
]
qed.
[ #G #L #s #L0 #cs #HL0 #X #H #L2 #HL20
>(lifts_inv_sort1 … H) -H
lapply (acr_gcr … H1RP H2RP (⓪)) #HAtom
- lapply (c4 … HAtom G L2 (◊)) /2 width=1 by/
+ lapply (b4 … HAtom G L2 (◊)) /2 width=1 by/
| #I #G #L1 #K1 #V1 #B #i #HLK1 #HKV1B #IHB #L0 #cs #HL01 #X #H #L2 #HL20
lapply (acr_gcr … H1RP H2RP B) #HB
elim (lifts_inv_lref1 … H) -H #i1 #Hi1 #H destruct
[ #K2 #HK20 #H destruct
elim (lift_total V0 0 (i0 +1)) #V #HV0
elim (lifts_lift_trans … Hi0 … Hcs0 … HV10 … HV0) -HV10 #V2 #HV12 #HV2
- lapply (c5 … HB ? G ? ? (◊) … HV0 HLK2) /3 width=7 by drops_cons, lifts_cons/ (* Note: uses IHB HL20 V2 HV0 *)
+ lapply (b5 … HB ? G ? ? (◊) … HV0 HLK2) /3 width=7 by drops_cons, lifts_cons/ (* Note: uses IHB HL20 V2 HV0 *)
| -HLK1 -IHB -HL01 -HL20 -HK1b -Hi0 -Hcs0
#K2 #V2 #A2 #HKV2A #H1KV0A #H2KV0A #_ #H1 #H2 destruct
lapply (drop_fwd_drop2 … HLK2) #HLK2b
lapply (aaa_mono … H2KV0A … HKV0B) #H destruct -H2KV0A -HKV0B
elim (lift_total V0 0 (i0 +1)) #V3 #HV03
elim (lift_total V2 0 (i0 +1)) #V #HV2
- lapply (c5 … HB ? G ? ? (◊) … (ⓝV3.V) … HLK2) /2 width=1 by lift_flat/
- lapply (c7 … HB G L2 (◊)) /3 width=7 by gcr_lift/
+ lapply (b5 … HB ? G ? ? (◊) … (ⓝV3.V) … HLK2) /2 width=1 by lift_flat/
+ lapply (b7 … HB G L2 (◊)) /3 width=7 by gcr_lift/
]
| #a #G #L #V #T #B #A #_ #_ #IHB #IHA #L0 #cs #HL0 #X #H #L2 #HL20
elim (lifts_inv_bind1 … H) -H #V0 #T0 #HV0 #HT0 #H destruct
lapply (acr_gcr … H1RP H2RP A) #HA
lapply (acr_gcr … H1RP H2RP B) #HB
- lapply (c1 … HB) -HB #HB
- lapply (c6 … HA G L2 (◊) (◊)) /4 width=5 by lsubc_pair, drops_skip, liftv_nil/
+ lapply (b1 … HB) -HB #HB
+ lapply (b6 … HA G L2 (◊) (◊)) /4 width=5 by lsubc_pair, drops_skip, liftv_nil/
| #a #G #L #W #T #B #A #HLWB #_ #IHB #IHA #L0 #cs #HL0 #X #H #L2 #HL02
elim (lifts_inv_bind1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct
@(acr_abst … H1RP H2RP) /2 width=5 by/
| #G #L #V #T #A #_ #_ #IH1A #IH2A #L0 #cs #HL0 #X #H #L2 #HL20
elim (lifts_inv_flat1 … H) -H #V0 #T0 #HV0 #HT0 #H destruct
lapply (acr_gcr … H1RP H2RP A) #HA
- lapply (c7 … HA G L2 (◊)) /3 width=5 by/
+ lapply (b7 … HA G L2 (◊)) /3 width=5 by/
]
qed.
∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → RP G L T.
#RR #RS #RP #H1RP #H2RP #G #L #T #A #HT
lapply (acr_gcr … H1RP H2RP A) #HA
-@(c1 … HA) /2 width=4 by acr_aaa/
+@(b1 … HA) /2 width=4 by acr_aaa/
qed.
⬇[i] L ≡ K.ⓑ{I}V1 → C G L (ⒶVs.#i).
definition S6 ≝ λRP,C:candidate.
- ∀G,L,V1c,V2c. ⬆[0, 1] V1c ≡ V2c →
- ∀a,V,T. C G (L.ⓓV) (ⒶV2c.T) → RP G L V → C G L (ⒶV1c.ⓓ{a}V.T).
+ ∀G,L,V1b,V2b. ⬆[0, 1] V1b ≡ V2b →
+ ∀a,V,T. C G (L.ⓓV) (ⒶV2b.T) → RP G L V → C G L (ⒶV1b.ⓓ{a}V.T).
definition S7 ≝ λC:candidate.
∀G,L,Vs,T,W. C G L (ⒶVs.T) → C G L (ⒶVs.W) → C G L (ⒶVs.ⓝW.T).
(* requirements for the generic reducibility candidate *)
record gcr (RR:relation4 genv lenv term term) (RS:relation term) (RP,C:candidate) : Prop ≝
-{ c1: S1 RP C;
- c2: S2 RR RS RP C;
- c3: S3 C;
- c4: S4 RP C;
- c5: S5 C;
- c6: S6 RP C;
- c7: S7 C
+{ b1: S1 RP C;
+ b2: S2 RR RS RP C;
+ b3: S3 C;
+ b4: S4 RP C;
+ b5: S5 C;
+ b6: S6 RP C;
+ b7: S7 C
}.
(* the functional construction for candidates *)
[ #G #L #T #H
elim (cp1 … H1RP G L) #s #HK
lapply (H L (⋆s) T (◊) ? ? ?) -H //
- [ lapply (c2 … IHB G L (◊) … HK) //
- | /3 width=6 by c1, cp3/
+ [ lapply (b2 … IHB G L (◊) … HK) //
+ | /3 width=6 by b1, cp3/
]
| #G #L #Vs #HVs #T #H1T #H2T #L0 #V0 #X #cs #HL0 #H #HB
- elim (lifts_inv_applv1 … H) -H #V0c #T0 #HV0c #HT0 #H destruct
- lapply (c1 … IHB … HB) #HV0
- @(c2 … IHA … (V0 @ V0c))
+ elim (lifts_inv_applv1 … H) -H #V0b #T0 #HV0b #HT0 #H destruct
+ lapply (b1 … IHB … HB) #HV0
+ @(b2 … IHA … (V0 @ V0b))
/3 width=14 by gcp2_lifts_all, gcp2_lifts, gcp0_lifts, lifts_simple_dx, conj/
| #a #G #L #Vs #U #T #W #HA #L0 #V0 #X #cs #HL0 #H #HB
- elim (lifts_inv_applv1 … H) -H #V0c #Y #HV0c #HY #H destruct
+ elim (lifts_inv_applv1 … H) -H #V0b #Y #HV0b #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
- @(c3 … IHA … (V0 @ V0c)) /5 width=6 by lifts_applv, lifts_flat, lifts_bind/
+ @(b3 … IHA … (V0 @ V0b)) /5 width=6 by lifts_applv, lifts_flat, lifts_bind/
| #G #L #Vs #HVs #s #L0 #V0 #X #cs #HL0 #H #HB
- elim (lifts_inv_applv1 … H) -H #V0c #Y #HV0c #HY #H destruct
+ elim (lifts_inv_applv1 … H) -H #V0b #Y #HV0b #HY #H destruct
>(lifts_inv_sort1 … HY) -Y
- lapply (c1 … IHB … HB) #HV0
- @(c4 … IHA … (V0 @ V0c)) /3 width=7 by gcp2_lifts_all, conj/
+ lapply (b1 … IHB … HB) #HV0
+ @(b4 … IHA … (V0 @ V0b)) /3 width=7 by gcp2_lifts_all, conj/
| #I #G #L #K #Vs #V1 #V2 #i #HA #HV12 #HLK #L0 #V0 #X #cs #HL0 #H #HB
- elim (lifts_inv_applv1 … H) -H #V0c #Y #HV0c #HY #H destruct
+ elim (lifts_inv_applv1 … H) -H #V0b #Y #HV0b #HY #H destruct
elim (lifts_inv_lref1 … HY) -HY #i0 #Hi0 #H destruct
elim (drops_drop_trans … HL0 … HLK) #X #cs0 #i1 #HL02 #H #Hi1 #Hcs0
>(at_mono … Hi1 … Hi0) in HL02; -i1 #HL02
elim (lift_total W1 0 (i0 + 1)) #W2 #HW12
elim (lifts_lift_trans … Hcs0 … HVW1 … HW12) // -Hcs0 -Hi0 #V3 #HV13 #HVW2
>(lift_mono … HV13 … HV12) in HVW2; -V3 #HVW2
- @(c5 … IHA … (V0 @ V0c) … HW12 HL02) /3 width=5 by lifts_applv/
-| #G #L #V1c #V2c #HV12c #a #V #T #HA #HV #L0 #V10 #X #cs #HL0 #H #HB
- elim (lifts_inv_applv1 … H) -H #V10c #Y #HV10c #HY #H destruct
+ @(b5 … IHA … (V0 @ V0b) … HW12 HL02) /3 width=5 by lifts_applv/
+| #G #L #V1b #V2b #HV12b #a #V #T #HA #HV #L0 #V10 #X #cs #HL0 #H #HB
+ elim (lifts_inv_applv1 … H) -H #V10b #Y #HV10b #HY #H destruct
elim (lifts_inv_bind1 … HY) -HY #V0 #T0 #HV0 #HT0 #H destruct
elim (lift_total V10 0 1) #V20 #HV120
- elim (liftv_total 0 1 V10c) #V20c #HV120c
- @(c6 … IHA … (V10 @ V10c) (V20 @ V20c)) /3 width=7 by gcp2_lifts, liftv_cons/
+ elim (liftv_total 0 1 V10b) #V20b #HV120b
+ @(b6 … IHA … (V10 @ V10b) (V20 @ V20b)) /3 width=7 by gcp2_lifts, liftv_cons/
@(HA … (cs + 1)) /2 width=2 by drops_skip/
[ @lifts_applv //
- elim (liftsv_liftv_trans_le … HV10c … HV120c) -V10c #V10c #HV10c #HV120c
- >(liftv_mono … HV12c … HV10c) -V1c //
+ elim (liftsv_liftv_trans_le … HV10b … HV120b) -V10b #V10b #HV10b #HV120b
+ >(liftv_mono … HV12b … HV10b) -V1b //
| @(gcr_lift … H1RP … HB … HV120) /2 width=2 by drop_drop/
]
| #G #L #Vs #T #W #HA #HW #L0 #V0 #X #cs #HL0 #H #HB
- elim (lifts_inv_applv1 … H) -H #V0c #Y #HV0c #HY #H destruct
+ elim (lifts_inv_applv1 … H) -H #V0b #Y #HV0b #HY #H destruct
elim (lifts_inv_flat1 … HY) -HY #W0 #T0 #HW0 #HT0 #H destruct
- @(c7 … IHA … (V0 @ V0c)) /3 width=5 by lifts_applv/
+ @(b7 … IHA … (V0 @ V0b)) /3 width=5 by lifts_applv/
]
qed.
lapply (acr_gcr … H1RP H2RP B) #HCB
elim (lifts_inv_bind1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct
lapply (gcr_lifts … H1RP … HL0 … HW0 HW) -HW #HW0
-lapply (c3 … HCA … a G L0 (◊)) #H @H -H
-lapply (c6 … HCA G L0 (◊) (◊) ?) // #H @H -H
+lapply (b3 … HCA … a G L0 (◊)) #H @H -H
+lapply (b6 … HCA G L0 (◊) (◊) ?) // #H @H -H
[ @(HA … HL0) //
-| lapply (c1 … HCB) -HCB #HCB
- lapply (c7 … H2RP G L0 (◊)) /3 width=1 by/
+| lapply (b1 … HCB) -HCB #HCB
+ lapply (b7 … H2RP G L0 (◊)) /3 width=1 by/
]
qed.
(* Properties on context-sensitive parallel computation for terms ***********)
-lemma lprs_cpr_trans: ∀G. c_r_transitive … (cpr G) (λ_. lprs G).
-/3 width=5 by c_r_trans_LTC2, lpr_cprs_trans/ qed-.
+lemma lprs_cpr_trans: ∀G. b_c_transitive … (cpr G) (λ_. lprs G).
+/3 width=5 by b_c_trans_LTC2, lpr_cprs_trans/ qed-.
(* Basic_1: was just: pr3_pr3_pr3_t *)
(* Note: alternative proof /3 width=5 by s_r_trans_LTC1, lprs_cpr_trans/ *)
-lemma lprs_cprs_trans: ∀G. c_rs_transitive … (cpr G) (λ_. lprs G).
-#G @c_r_to_c_rs_trans @c_r_trans_LTC2
-@c_rs_trans_TC1 /2 width=3 by lpr_cprs_trans/ (**) (* full auto too slow *)
+lemma lprs_cprs_trans: ∀G. b_rs_transitive … (cpr G) (λ_. lprs G).
+#G @b_c_to_b_rs_trans @b_c_trans_LTC2
+@b_rs_trans_TC1 /2 width=3 by lpr_cprs_trans/ (**) (* full auto too slow *)
qed-.
lemma lprs_cprs_conf_dx: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡* T1 →
(* Properties on context-sensitive extended parallel computation for terms **)
-lemma lpxs_cpx_trans: ∀h,o,G. c_r_transitive … (cpx h o G) (λ_.lpxs h o G).
-/3 width=5 by c_r_trans_LTC2, lpx_cpxs_trans/ qed-.
+lemma lpxs_cpx_trans: ∀h,o,G. b_c_transitive … (cpx h o G) (λ_.lpxs h o G).
+/3 width=5 by b_c_trans_LTC2, lpx_cpxs_trans/ qed-.
(* Note: alternative proof: /3 width=5 by s_r_trans_TC1, lpxs_cpx_trans/ *)
-lemma lpxs_cpxs_trans: ∀h,o,G. c_rs_transitive … (cpx h o G) (λ_.lpxs h o G).
-#h #o #G @c_r_to_c_rs_trans @c_r_trans_LTC2
-@c_rs_trans_TC1 /2 width=3 by lpx_cpxs_trans/ (**) (* full auto too slow *)
+lemma lpxs_cpxs_trans: ∀h,o,G. b_rs_transitive … (cpx h o G) (λ_.lpxs h o G).
+#h #o #G @b_c_to_b_rs_trans @b_c_trans_LTC2
+@b_rs_trans_TC1 /2 width=3 by lpx_cpxs_trans/ (**) (* full auto too slow *)
qed-.
lemma cpxs_bind2: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡*[h, o] V2 →
(* Basic_1: was: csubc_drop_conf_O *)
(* Note: the constant 0 can not be generalized *)
-lemma lsubc_drop_O1_trans: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → ∀K2,c,k. ⬇[c, 0, k] L2 ≡ K2 →
- ∃∃K1. ⬇[c, 0, k] L1 ≡ K1 & G ⊢ K1 ⫃[RP] K2.
+lemma lsubc_drop_O1_trans: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → ∀K2,b,k. ⬇[b, 0, k] L2 ≡ K2 →
+ ∃∃K1. ⬇[b, 0, k] L1 ≡ K1 & G ⊢ K1 ⫃[RP] K2.
#RP #G #L1 #L2 #H elim H -L1 -L2
-[ #X #c #k #H elim (drop_inv_atom1 … H) -H /4 width=3 by drop_atom, ex2_intro/
-| #I #L1 #L2 #V #_ #IHL12 #X #c #k #H
+[ #X #b #k #H elim (drop_inv_atom1 … H) -H /4 width=3 by drop_atom, ex2_intro/
+| #I #L1 #L2 #V #_ #IHL12 #X #b #k #H
elim (drop_inv_O1_pair1 … H) -H * #Hm #H destruct
- [ elim (IHL12 L2 c 0) -IHL12 // #X #H <(drop_inv_O2 … H) -H
+ [ elim (IHL12 L2 b 0) -IHL12 // #X #H <(drop_inv_O2 … H) -H
/3 width=3 by lsubc_pair, drop_pair, ex2_intro/
| elim (IHL12 … H) -L2 /3 width=3 by drop_drop_lt, ex2_intro/
]
-| #L1 #L2 #V #W #A #HV #H1W #H2W #_ #IHL12 #X #c #k #H
+| #L1 #L2 #V #W #A #HV #H1W #H2W #_ #IHL12 #X #b #k #H
elim (drop_inv_O1_pair1 … H) -H * #Hm #H destruct
- [ elim (IHL12 L2 c 0) -IHL12 // #X #H <(drop_inv_O2 … H) -H
+ [ elim (IHL12 L2 b 0) -IHL12 // #X #H <(drop_inv_O2 … H) -H
/3 width=8 by lsubc_beta, drop_pair, ex2_intro/
| elim (IHL12 … H) -L2 /3 width=3 by drop_drop_lt, ex2_intro/
]
(* Relocation properties ****************************************************)
lemma scpds_lift: ∀h,o,G,d. d_liftable (scpds h o d G).
-#h #o #G #d2 #K #T1 #T2 * #T #d1 #Hd21 #Hd1 #HT1 #HT2 #L #c #l #k
+#h #o #G #d2 #K #T1 #T2 * #T #d1 #Hd21 #Hd1 #HT1 #HT2 #L #b #l #k
elim (lift_total T l k)
/3 width=15 by cprs_lift, da_lift, lstas_lift, ex4_2_intro/
qed.
lemma scpds_inv_lift1: ∀h,o,G,d. d_deliftable_sn (scpds h o d G).
-#h #o #G #d2 #L #U1 #U2 * #U #d1 #Hd21 #Hd1 #HU1 #HU2 #K #c #l #k #HLK #T1 #HTU1
+#h #o #G #d2 #L #U1 #U2 * #U #d1 #Hd21 #Hd1 #HU1 #HU2 #K #b #l #k #HLK #T1 #HTU1
lapply (da_inv_lift … Hd1 … HLK … HTU1) -Hd1 #Hd1
elim (lstas_inv_lift1 … HU1 … HLK … HTU1) -U1 #T #HTU #HT1
elim (cprs_inv_lift1 … HU2 … HLK … HTU) -U -L
#h #o #G #L #T1 #T #T2 #d1 #d2 #d #Hd21 #HTd1 #HT1 * #T0 #d0 #Hd0 #HTd0 #HT0 #HT02
lapply (lstas_da_conf … HT1 … HTd1) #HTd12
lapply (da_mono … HTd12 … HTd0) -HTd12 -HTd0 #H destruct
-lapply (le_minus_to_plus_r … Hd21 Hd0) -Hd21 -Hd0
+lapply (le_minus_to_plus_c … Hd21 Hd0) -Hd21 -Hd0
/3 width=7 by lstas_trans, ex4_2_intro/
qed-.
inductive cpg (h): rtc → relation4 genv lenv term term ≝
| cpg_atom : ∀I,G,L. cpg h (𝟘𝟘) G L (⓪{I}) (⓪{I})
| cpg_ess : ∀G,L,s. cpg h (𝟘𝟙) G L (⋆s) (⋆(next h s))
-| cpg_delta: ∀r,G,L,V1,V2,W2. cpg h r G L V1 V2 →
- ⬆*[1] V2 ≡ W2 → cpg h (↓r) G (L.ⓓV1) (#0) W2
-| cpg_ell : ∀r,G,L,V1,V2,W2. cpg h r G L V1 V2 →
- ⬆*[1] V2 ≡ W2 → cpg h ((↓r)+𝟘𝟙) G (L.ⓛV1) (#0) W2
-| cpt_lref : ∀r,I,G,L,V,T,U,i. cpg h r G L (#i) T →
- ⬆*[1] T ≡ U → cpg h r G (L.ⓑ{I}V) (#⫯i) U
-| cpg_bind : ∀rV,rT,p,I,G,L,V1,V2,T1,T2.
- cpg h rV G L V1 V2 → cpg h rT G (L.ⓑ{I}V1) T1 T2 →
- cpg h ((↓rV)+rT) G L (ⓑ{p,I}V1.T1) (ⓑ{p,I}V2.T2)
-| cpg_flat : ∀rV,rT,I,G,L,V1,V2,T1,T2.
- cpg h rV G L V1 V2 → cpg h rT G L T1 T2 →
- cpg h ((↓rV)+rT) G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2)
-| cpg_zeta : ∀r,G,L,V,T1,T,T2. cpg h r G (L.ⓓV) T1 T →
- ⬆*[1] T2 ≡ T → cpg h ((↓r)+𝟙𝟘) G L (+ⓓV.T1) T2
-| cpg_eps : ∀r,G,L,V,T1,T2. cpg h r G L T1 T2 → cpg h ((↓r)+𝟙𝟘) G L (ⓝV.T1) T2
-| cpg_ee : ∀r,G,L,V1,V2,T. cpg h r G L V1 V2 → cpg h ((↓r)+𝟘𝟙) G L (ⓝV1.T) V2
-| cpg_beta : ∀rV,rW,rT,p,G,L,V1,V2,W1,W2,T1,T2.
- cpg h rV G L V1 V2 → cpg h rW G L W1 W2 → cpg h rT G (L.ⓛW1) T1 T2 →
- cpg h ((↓rV)+(↓rW)+(↓rT)+𝟙𝟘) G L (ⓐV1.ⓛ{p}W1.T1) (ⓓ{p}ⓝW2.V2.T2)
-| cpg_theta: ∀rV,rW,rT,p,G,L,V1,V,V2,W1,W2,T1,T2.
- cpg h rV G L V1 V → ⬆*[1] V ≡ V2 → cpg h rW G L W1 W2 →
- cpg h rT G (L.ⓓW1) T1 T2 →
- cpg h ((↓rV)+(↓rW)+(↓rT)+𝟙𝟘) G L (ⓐV1.ⓓ{p}W1.T1) (ⓓ{p}W2.ⓐV2.T2)
+| cpg_delta: ∀c,G,L,V1,V2,W2. cpg h c G L V1 V2 →
+ ⬆*[1] V2 ≡ W2 → cpg h (↓c) G (L.ⓓV1) (#0) W2
+| cpg_ell : ∀c,G,L,V1,V2,W2. cpg h c G L V1 V2 →
+ ⬆*[1] V2 ≡ W2 → cpg h ((↓c)+𝟘𝟙) G (L.ⓛV1) (#0) W2
+| cpt_lref : ∀c,I,G,L,V,T,U,i. cpg h c G L (#i) T →
+ ⬆*[1] T ≡ U → cpg h c G (L.ⓑ{I}V) (#⫯i) U
+| cpg_bind : ∀cV,cT,p,I,G,L,V1,V2,T1,T2.
+ cpg h cV G L V1 V2 → cpg h cT G (L.ⓑ{I}V1) T1 T2 →
+ cpg h ((↓cV)+cT) G L (ⓑ{p,I}V1.T1) (ⓑ{p,I}V2.T2)
+| cpg_flat : ∀cV,cT,I,G,L,V1,V2,T1,T2.
+ cpg h cV G L V1 V2 → cpg h cT G L T1 T2 →
+ cpg h ((↓cV)+cT) G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2)
+| cpg_zeta : ∀c,G,L,V,T1,T,T2. cpg h c G (L.ⓓV) T1 T →
+ ⬆*[1] T2 ≡ T → cpg h ((↓c)+𝟙𝟘) G L (+ⓓV.T1) T2
+| cpg_eps : ∀c,G,L,V,T1,T2. cpg h c G L T1 T2 → cpg h ((↓c)+𝟙𝟘) G L (ⓝV.T1) T2
+| cpg_ee : ∀c,G,L,V1,V2,T. cpg h c G L V1 V2 → cpg h ((↓c)+𝟘𝟙) G L (ⓝV1.T) V2
+| cpg_beta : ∀cV,cW,cT,p,G,L,V1,V2,W1,W2,T1,T2.
+ cpg h cV G L V1 V2 → cpg h cW G L W1 W2 → cpg h cT G (L.ⓛW1) T1 T2 →
+ cpg h ((↓cV)+(↓cW)+(↓cT)+𝟙𝟘) G L (ⓐV1.ⓛ{p}W1.T1) (ⓓ{p}ⓝW2.V2.T2)
+| cpg_theta: ∀cV,cW,cT,p,G,L,V1,V,V2,W1,W2,T1,T2.
+ cpg h cV G L V1 V → ⬆*[1] V ≡ V2 → cpg h cW G L W1 W2 →
+ cpg h cT G (L.ⓓW1) T1 T2 →
+ cpg h ((↓cV)+(↓cW)+(↓cT)+𝟙𝟘) G L (ⓐV1.ⓓ{p}W1.T1) (ⓓ{p}W2.ⓐV2.T2)
.
interpretation
"context-sensitive generic parallel rt-transition (term)"
- 'PRed h r G L T1 T2 = (cpg h r G L T1 T2).
+ 'PRed c h G L T1 T2 = (cpg h c G L T1 T2).
(* Basic properties *********************************************************)
(* Note: this is "∀h,g,L. reflexive … (cpg h (𝟘𝟘) L)" *)
-lemma cpg_refl: ∀h,G,T,L. ⦃G, L⦄ ⊢ T ➡[h, 𝟘𝟘] T.
+lemma cpg_refl: ∀h,G,T,L. ⦃G, L⦄ ⊢ T ➡[𝟘𝟘, h] T.
#h #G #T elim T -T // * /2 width=1 by cpg_bind, cpg_flat/
qed.
-lemma cpg_pair_sn: ∀h,r,I,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, r] V2 →
- ∀T. ⦃G, L⦄ ⊢ ②{I}V1.T ➡[h, ↓r] ②{I}V2.T.
-#h #r * /2 width=1 by cpg_bind, cpg_flat/
+lemma cpg_pair_sn: ∀c,h,I,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[c, h] V2 →
+ ∀T. ⦃G, L⦄ ⊢ ②{I}V1.T ➡[↓c, h] ②{I}V2.T.
+#c #h * /2 width=1 by cpg_bind, cpg_flat/
qed.
(* Basic inversion lemmas ***************************************************)
-fact cpg_inv_atom1_aux: ∀h,r,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, r] T2 → ∀J. T1 = ⓪{J} →
+fact cpg_inv_atom1_aux: ∀c,h,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[c, h] T2 → ∀J. T1 = ⓪{J} →
∨∨ T2 = ⓪{J}
| ∃∃s. J = Sort s & T2 = ⋆(next h s)
- | ∃∃rV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[h, rV] V2 & ⬆*[1] V2 ≡ T2 &
- L = K.ⓓV1 & J = LRef 0 & r = ↓rV
- | ∃∃rV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[h, rV] V2 & ⬆*[1] V2 ≡ T2 &
- L = K.ⓛV1 & J = LRef 0 & r = (↓rV)+𝟘𝟙
- | ∃∃I,K,V,T,i. ⦃G, K⦄ ⊢ #i ➡[h, r] T & ⬆*[1] T ≡ T2 &
+ | ∃∃cV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[cV, h] V2 & ⬆*[1] V2 ≡ T2 &
+ L = K.ⓓV1 & J = LRef 0 & c = ↓cV
+ | ∃∃cV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[cV, h] V2 & ⬆*[1] V2 ≡ T2 &
+ L = K.ⓛV1 & J = LRef 0 & c = (↓cV)+𝟘𝟙
+ | ∃∃I,K,V,T,i. ⦃G, K⦄ ⊢ #i ➡[c, h] T & ⬆*[1] T ≡ T2 &
L = K.ⓑ{I}V & J = LRef (⫯i).
-#h #r #G #L #T1 #T2 * -r -G -L -T1 -T2
+#c #h #G #L #T1 #T2 * -c -G -L -T1 -T2
[ #I #G #L #J #H destruct /2 width=1 by or5_intro0/
| #G #L #s #J #H destruct /3 width=3 by or5_intro1, ex2_intro/
-| #r #G #L #V1 #V2 #W2 #HV12 #VW2 #J #H destruct /3 width=8 by or5_intro2, ex5_4_intro/
-| #r #G #L #V1 #V2 #W2 #HV12 #VW2 #J #H destruct /3 width=8 by or5_intro3, ex5_4_intro/
-| #r #I #G #L #V #T #U #i #HT #HTU #J #H destruct /3 width=9 by or5_intro4, ex4_5_intro/
-| #rV #rT #p #I #G #L #V1 #V2 #T1 #T2 #_ #_ #J #H destruct
-| #rV #rT #I #G #L #V1 #V2 #T1 #T2 #_ #_ #J #H destruct
-| #r #G #L #V #T1 #T #T2 #_ #_ #J #H destruct
-| #r #G #L #V #T1 #T2 #_ #J #H destruct
-| #r #G #L #V1 #V2 #T #_ #J #H destruct
-| #rV #rW #rT #p #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #J #H destruct
-| #rV #rW #rT #p #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #J #H destruct
+| #c #G #L #V1 #V2 #W2 #HV12 #VW2 #J #H destruct /3 width=8 by or5_intro2, ex5_4_intro/
+| #c #G #L #V1 #V2 #W2 #HV12 #VW2 #J #H destruct /3 width=8 by or5_intro3, ex5_4_intro/
+| #c #I #G #L #V #T #U #i #HT #HTU #J #H destruct /3 width=9 by or5_intro4, ex4_5_intro/
+| #cV #cT #p #I #G #L #V1 #V2 #T1 #T2 #_ #_ #J #H destruct
+| #cV #cT #I #G #L #V1 #V2 #T1 #T2 #_ #_ #J #H destruct
+| #c #G #L #V #T1 #T #T2 #_ #_ #J #H destruct
+| #c #G #L #V #T1 #T2 #_ #J #H destruct
+| #c #G #L #V1 #V2 #T #_ #J #H destruct
+| #cV #cW #cT #p #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #J #H destruct
+| #cV #cW #cT #p #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #J #H destruct
]
qed-.
-lemma cpg_inv_atom1: ∀h,r,J,G,L,T2. ⦃G, L⦄ ⊢ ⓪{J} ➡[h, r] T2 →
+lemma cpg_inv_atom1: ∀c,h,J,G,L,T2. ⦃G, L⦄ ⊢ ⓪{J} ➡[c, h] T2 →
∨∨ T2 = ⓪{J}
| ∃∃s. J = Sort s & T2 = ⋆(next h s)
- | ∃∃rV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[h, rV] V2 & ⬆*[1] V2 ≡ T2 &
- L = K.ⓓV1 & J = LRef 0 & r = ↓rV
- | ∃∃rV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[h, rV] V2 & ⬆*[1] V2 ≡ T2 &
- L = K.ⓛV1 & J = LRef 0 & r = (↓rV)+𝟘𝟙
- | ∃∃I,K,V,T,i. ⦃G, K⦄ ⊢ #i ➡[h, r] T & ⬆*[1] T ≡ T2 &
+ | ∃∃cV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[cV, h] V2 & ⬆*[1] V2 ≡ T2 &
+ L = K.ⓓV1 & J = LRef 0 & c = ↓cV
+ | ∃∃cV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[cV, h] V2 & ⬆*[1] V2 ≡ T2 &
+ L = K.ⓛV1 & J = LRef 0 & c = (↓cV)+𝟘𝟙
+ | ∃∃I,K,V,T,i. ⦃G, K⦄ ⊢ #i ➡[c, h] T & ⬆*[1] T ≡ T2 &
L = K.ⓑ{I}V & J = LRef (⫯i).
/2 width=3 by cpg_inv_atom1_aux/ qed-.
-lemma cpg_inv_sort1: ∀h,r,G,L,T2,s. ⦃G, L⦄ ⊢ ⋆s ➡[h, r] T2 →
+lemma cpg_inv_sort1: ∀c,h,G,L,T2,s. ⦃G, L⦄ ⊢ ⋆s ➡[c, h] T2 →
T2 = ⋆s ∨ T2 = ⋆(next h s).
-#h #r #G #L #T2 #s #H
+#c #h #G #L #T2 #s #H
elim (cpg_inv_atom1 … H) -H /2 width=1 by or_introl/ *
[ #s0 #H destruct /2 width=1 by or_intror/
-|2,3: #rV #K #V1 #V2 #_ #_ #_ #H destruct
+|2,3: #cV #K #V1 #V2 #_ #_ #_ #H destruct
| #I #K #V1 #V2 #i #_ #_ #_ #H destruct
]
qed-.
-lemma cpg_inv_zero1: ∀h,r,G,L,T2. ⦃G, L⦄ ⊢ #0 ➡[h, r] T2 →
+lemma cpg_inv_zero1: ∀c,h,G,L,T2. ⦃G, L⦄ ⊢ #0 ➡[c, h] T2 →
∨∨ T2 = #0
- | ∃∃rV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[h, rV] V2 & ⬆*[1] V2 ≡ T2 &
- L = K.ⓓV1 & r = ↓rV
- | ∃∃rV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[h, rV] V2 & ⬆*[1] V2 ≡ T2 &
- L = K.ⓛV1 & r = (↓rV)+𝟘𝟙.
-#h #r #G #L #T2 #H
+ | ∃∃cV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[cV, h] V2 & ⬆*[1] V2 ≡ T2 &
+ L = K.ⓓV1 & c = ↓cV
+ | ∃∃cV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[cV, h] V2 & ⬆*[1] V2 ≡ T2 &
+ L = K.ⓛV1 & c = (↓cV)+𝟘𝟙.
+#c #h #G #L #T2 #H
elim (cpg_inv_atom1 … H) -H /2 width=1 by or3_intro0/ *
[ #s #H destruct
-|2,3: #rV #K #V1 #V2 #HV12 #HVT2 #H1 #_ #H2 destruct /3 width=8 by or3_intro1, or3_intro2, ex4_4_intro/
+|2,3: #cV #K #V1 #V2 #HV12 #HVT2 #H1 #_ #H2 destruct /3 width=8 by or3_intro1, or3_intro2, ex4_4_intro/
| #I #K #V1 #V2 #i #_ #_ #_ #H destruct
]
qed-.
-lemma cpg_inv_lref1: ∀h,r,G,L,T2,i. ⦃G, L⦄ ⊢ #⫯i ➡[h, r] T2 →
+lemma cpg_inv_lref1: ∀c,h,G,L,T2,i. ⦃G, L⦄ ⊢ #⫯i ➡[c, h] T2 →
(T2 = #⫯i) ∨
- ∃∃I,K,V,T. ⦃G, K⦄ ⊢ #i ➡[h, r] T & ⬆*[1] T ≡ T2 & L = K.ⓑ{I}V.
-#h #r #G #L #T2 #i #H
+ ∃∃I,K,V,T. ⦃G, K⦄ ⊢ #i ➡[c, h] T & ⬆*[1] T ≡ T2 & L = K.ⓑ{I}V.
+#c #h #G #L #T2 #i #H
elim (cpg_inv_atom1 … H) -H /2 width=1 by or_introl/ *
[ #s #H destruct
-|2,3: #rV #K #V1 #V2 #_ #_ #_ #H destruct
+|2,3: #cV #K #V1 #V2 #_ #_ #_ #H destruct
| #I #K #V1 #V2 #j #HV2 #HVT2 #H1 #H2 destruct /3 width=7 by ex3_4_intro, or_intror/
]
qed-.
-lemma cpg_inv_gref1: ∀h,r,G,L,T2,l. ⦃G, L⦄ ⊢ §l ➡[h, r] T2 → T2 = §l.
-#h #r #G #L #T2 #l #H
+lemma cpg_inv_gref1: ∀c,h,G,L,T2,l. ⦃G, L⦄ ⊢ §l ➡[c, h] T2 → T2 = §l.
+#c #h #G #L #T2 #l #H
elim (cpg_inv_atom1 … H) -H // *
[ #s #H destruct
-|2,3: #rV #K #V1 #V2 #_ #_ #_ #H destruct
+|2,3: #cV #K #V1 #V2 #_ #_ #_ #H destruct
| #I #K #V1 #V2 #i #_ #_ #_ #H destruct
]
qed-.
-fact cpg_inv_bind1_aux: ∀h,r,G,L,U,U2. ⦃G, L⦄ ⊢ U ➡[h, r] U2 →
+fact cpg_inv_bind1_aux: ∀c,h,G,L,U,U2. ⦃G, L⦄ ⊢ U ➡[c, h] U2 →
∀p,J,V1,U1. U = ⓑ{p,J}V1.U1 → (
- ∃∃rV,rT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V2 & ⦃G, L.ⓑ{J}V1⦄ ⊢ U1 ➡[h, rT] T2 &
- U2 = ⓑ{p,J}V2.T2 & r = (↓rV)+rT
+ ∃∃cV,cT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[cV, h] V2 & ⦃G, L.ⓑ{J}V1⦄ ⊢ U1 ➡[cT, h] T2 &
+ U2 = ⓑ{p,J}V2.T2 & c = (↓cV)+cT
) ∨
- ∃∃rT,T. ⦃G, L.ⓓV1⦄ ⊢ U1 ➡[h, rT] T & ⬆*[1] U2 ≡ T &
- p = true & J = Abbr & r = (↓rT)+𝟙𝟘.
-#h #r #G #L #U #U2 * -r -G -L -U -U2
+ ∃∃cT,T. ⦃G, L.ⓓV1⦄ ⊢ U1 ➡[cT, h] T & ⬆*[1] U2 ≡ T &
+ p = true & J = Abbr & c = (↓cT)+𝟙𝟘.
+#c #h #G #L #U #U2 * -c -G -L -U -U2
[ #I #G #L #q #J #W #U1 #H destruct
| #G #L #s #q #J #W #U1 #H destruct
-| #r #G #L #V1 #V2 #W2 #_ #_ #q #J #W #U1 #H destruct
-| #r #G #L #V1 #V2 #W2 #_ #_ #q #J #W #U1 #H destruct
-| #r #I #G #L #V #T #U #i #_ #_ #q #J #W #U1 #H destruct
-| #rv #rT #p #I #G #L #V1 #V2 #T1 #T2 #HV12 #HT12 #q #J #W #U1 #H destruct /3 width=8 by ex4_4_intro, or_introl/
-| #rv #rT #I #G #L #V1 #V2 #T1 #T2 #_ #_ #q #J #W #U1 #H destruct
-| #r #G #L #V #T1 #T #T2 #HT1 #HT2 #q #J #W #U1 #H destruct /3 width=5 by ex5_2_intro, or_intror/
-| #r #G #L #V #T1 #T2 #_ #q #J #W #U1 #H destruct
-| #r #G #L #V1 #V2 #T #_ #q #J #W #U1 #H destruct
-| #rv #rW #rT #p #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #q #J #W #U1 #H destruct
-| #rv #rW #rT #p #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #q #J #W #U1 #H destruct
+| #c #G #L #V1 #V2 #W2 #_ #_ #q #J #W #U1 #H destruct
+| #c #G #L #V1 #V2 #W2 #_ #_ #q #J #W #U1 #H destruct
+| #c #I #G #L #V #T #U #i #_ #_ #q #J #W #U1 #H destruct
+| #rv #cT #p #I #G #L #V1 #V2 #T1 #T2 #HV12 #HT12 #q #J #W #U1 #H destruct /3 width=8 by ex4_4_intro, or_introl/
+| #rv #cT #I #G #L #V1 #V2 #T1 #T2 #_ #_ #q #J #W #U1 #H destruct
+| #c #G #L #V #T1 #T #T2 #HT1 #HT2 #q #J #W #U1 #H destruct /3 width=5 by ex5_2_intro, or_intror/
+| #c #G #L #V #T1 #T2 #_ #q #J #W #U1 #H destruct
+| #c #G #L #V1 #V2 #T #_ #q #J #W #U1 #H destruct
+| #rv #cW #cT #p #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #q #J #W #U1 #H destruct
+| #rv #cW #cT #p #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #q #J #W #U1 #H destruct
]
qed-.
-lemma cpg_inv_bind1: ∀h,r,p,I,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓑ{p,I}V1.T1 ➡[h, r] U2 → (
- ∃∃rV,rT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V2 & ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ➡[h, rT] T2 &
- U2 = ⓑ{p,I}V2.T2 & r = (↓rV)+rT
+lemma cpg_inv_bind1: ∀c,h,p,I,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓑ{p,I}V1.T1 ➡[c, h] U2 → (
+ ∃∃cV,cT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[cV, h] V2 & ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ➡[cT, h] T2 &
+ U2 = ⓑ{p,I}V2.T2 & c = (↓cV)+cT
) ∨
- ∃∃rT,T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[h, rT] T & ⬆*[1] U2 ≡ T &
- p = true & I = Abbr & r = (↓rT)+𝟙𝟘.
+ ∃∃cT,T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[cT, h] T & ⬆*[1] U2 ≡ T &
+ p = true & I = Abbr & c = (↓cT)+𝟙𝟘.
/2 width=3 by cpg_inv_bind1_aux/ qed-.
-lemma cpg_inv_abbr1: ∀h,r,p,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{p}V1.T1 ➡[h, r] U2 → (
- ∃∃rV,rT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V2 & ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[h, rT] T2 &
- U2 = ⓓ{p}V2.T2 & r = (↓rV)+rT
+lemma cpg_inv_abbr1: ∀c,h,p,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{p}V1.T1 ➡[c, h] U2 → (
+ ∃∃cV,cT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[cV, h] V2 & ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[cT, h] T2 &
+ U2 = ⓓ{p}V2.T2 & c = (↓cV)+cT
) ∨
- ∃∃rT,T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[h, rT] T & ⬆*[1] U2 ≡ T &
- p = true & r = (↓rT)+𝟙𝟘.
-#h #r #p #G #L #V1 #T1 #U2 #H elim (cpg_inv_bind1 … H) -H *
+ ∃∃cT,T. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡[cT, h] T & ⬆*[1] U2 ≡ T &
+ p = true & c = (↓cT)+𝟙𝟘.
+#c #h #p #G #L #V1 #T1 #U2 #H elim (cpg_inv_bind1 … H) -H *
/3 width=8 by ex4_4_intro, ex4_2_intro, or_introl, or_intror/
qed-.
-lemma cpg_inv_abst1: ∀h,r,p,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{p}V1.T1 ➡[h, r] U2 →
- ∃∃rV,rT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 ➡[h, rT] T2 &
- U2 = ⓛ{p} V2. T2 & r = (↓rV)+rT.
-#h #r #p #G #L #V1 #T1 #U2 #H elim (cpg_inv_bind1 … H) -H *
+lemma cpg_inv_abst1: ∀c,h,p,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{p}V1.T1 ➡[c, h] U2 →
+ ∃∃cV,cT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[cV, h] V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 ➡[cT, h] T2 &
+ U2 = ⓛ{p} V2. T2 & c = (↓cV)+cT.
+#c #h #p #G #L #V1 #T1 #U2 #H elim (cpg_inv_bind1 … H) -H *
[ /3 width=8 by ex4_4_intro/
-| #r #T #_ #_ #_ #H destruct
+| #c #T #_ #_ #_ #H destruct
]
qed-.
-fact cpg_inv_flat1_aux: ∀h,r,G,L,U,U2. ⦃G, L⦄ ⊢ U ➡[h, r] U2 →
+fact cpg_inv_flat1_aux: ∀c,h,G,L,U,U2. ⦃G, L⦄ ⊢ U ➡[c, h] U2 →
∀J,V1,U1. U = ⓕ{J}V1.U1 →
- ∨∨ ∃∃rV,rT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, rT] T2 &
- U2 = ⓕ{J}V2.T2 & r = (↓rV)+rT
- | ∃∃rT. ⦃G, L⦄ ⊢ U1 ➡[h, rT] U2 & J = Cast & r = (↓rT)+𝟙𝟘
- | ∃∃rV. ⦃G, L⦄ ⊢ V1 ➡[h, rV] U2 & J = Cast & r = (↓rV)+𝟘𝟙
- | ∃∃rV,rW,rT,p,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V2 & ⦃G, L⦄ ⊢ W1 ➡[h, rW] W2 & ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[h, rT] T2 &
- J = Appl & U1 = ⓛ{p}W1.T1 & U2 = ⓓ{p}ⓝW2.V2.T2 & r = (↓rV)+(↓rW)+(↓rT)+𝟙𝟘
- | ∃∃rV,rW,rT,p,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V & ⬆*[1] V ≡ V2 & ⦃G, L⦄ ⊢ W1 ➡[h, rW] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[h, rT] T2 &
- J = Appl & U1 = ⓓ{p}W1.T1 & U2 = ⓓ{p}W2.ⓐV2.T2 & r = (↓rV)+(↓rW)+(↓rT)+𝟙𝟘.
-#h #r #G #L #U #U2 * -r -G -L -U -U2
+ ∨∨ ∃∃cV,cT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[cV, h] V2 & ⦃G, L⦄ ⊢ U1 ➡[cT, h] T2 &
+ U2 = ⓕ{J}V2.T2 & c = (↓cV)+cT
+ | ∃∃cT. ⦃G, L⦄ ⊢ U1 ➡[cT, h] U2 & J = Cast & c = (↓cT)+𝟙𝟘
+ | ∃∃cV. ⦃G, L⦄ ⊢ V1 ➡[cV, h] U2 & J = Cast & c = (↓cV)+𝟘𝟙
+ | ∃∃cV,cW,cT,p,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[cV, h] V2 & ⦃G, L⦄ ⊢ W1 ➡[cW, h] W2 & ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[cT, h] T2 &
+ J = Appl & U1 = ⓛ{p}W1.T1 & U2 = ⓓ{p}ⓝW2.V2.T2 & c = (↓cV)+(↓cW)+(↓cT)+𝟙𝟘
+ | ∃∃cV,cW,cT,p,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[cV, h] V & ⬆*[1] V ≡ V2 & ⦃G, L⦄ ⊢ W1 ➡[cW, h] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[cT, h] T2 &
+ J = Appl & U1 = ⓓ{p}W1.T1 & U2 = ⓓ{p}W2.ⓐV2.T2 & c = (↓cV)+(↓cW)+(↓cT)+𝟙𝟘.
+#c #h #G #L #U #U2 * -c -G -L -U -U2
[ #I #G #L #J #W #U1 #H destruct
| #G #L #s #J #W #U1 #H destruct
-| #r #G #L #V1 #V2 #W2 #_ #_ #J #W #U1 #H destruct
-| #r #G #L #V1 #V2 #W2 #_ #_ #J #W #U1 #H destruct
-| #r #I #G #L #V #T #U #i #_ #_ #J #W #U1 #H destruct
-| #rv #rT #p #I #G #L #V1 #V2 #T1 #T2 #_ #_ #J #W #U1 #H destruct
-| #rv #rT #I #G #L #V1 #V2 #T1 #T2 #HV12 #HT12 #J #W #U1 #H destruct /3 width=8 by or5_intro0, ex4_4_intro/
-| #r #G #L #V #T1 #T #T2 #_ #_ #J #W #U1 #H destruct
-| #r #G #L #V #T1 #T2 #HT12 #J #W #U1 #H destruct /3 width=3 by or5_intro1, ex3_intro/
-| #r #G #L #V1 #V2 #T #HV12 #J #W #U1 #H destruct /3 width=3 by or5_intro2, ex3_intro/
-| #rv #rW #rT #p #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HW12 #HT12 #J #W #U1 #H destruct /3 width=15 by or5_intro3, ex7_9_intro/
-| #rv #rW #rT #p #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 #HW12 #HT12 #J #W #U1 #H destruct /3 width=17 by or5_intro4, ex8_10_intro/
+| #c #G #L #V1 #V2 #W2 #_ #_ #J #W #U1 #H destruct
+| #c #G #L #V1 #V2 #W2 #_ #_ #J #W #U1 #H destruct
+| #c #I #G #L #V #T #U #i #_ #_ #J #W #U1 #H destruct
+| #rv #cT #p #I #G #L #V1 #V2 #T1 #T2 #_ #_ #J #W #U1 #H destruct
+| #rv #cT #I #G #L #V1 #V2 #T1 #T2 #HV12 #HT12 #J #W #U1 #H destruct /3 width=8 by or5_intro0, ex4_4_intro/
+| #c #G #L #V #T1 #T #T2 #_ #_ #J #W #U1 #H destruct
+| #c #G #L #V #T1 #T2 #HT12 #J #W #U1 #H destruct /3 width=3 by or5_intro1, ex3_intro/
+| #c #G #L #V1 #V2 #T #HV12 #J #W #U1 #H destruct /3 width=3 by or5_intro2, ex3_intro/
+| #rv #cW #cT #p #G #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HW12 #HT12 #J #W #U1 #H destruct /3 width=15 by or5_intro3, ex7_9_intro/
+| #rv #cW #cT #p #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 #HW12 #HT12 #J #W #U1 #H destruct /3 width=17 by or5_intro4, ex8_10_intro/
]
qed-.
-lemma cpg_inv_flat1: ∀h,r,I,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓕ{I}V1.U1 ➡[h, r] U2 →
- ∨∨ ∃∃rV,rT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, rT] T2 &
- U2 = ⓕ{I}V2.T2 & r = (↓rV)+rT
- | ∃∃rT. ⦃G, L⦄ ⊢ U1 ➡[h, rT] U2 & I = Cast & r = (↓rT)+𝟙𝟘
- | ∃∃rV. ⦃G, L⦄ ⊢ V1 ➡[h, rV] U2 & I = Cast & r = (↓rV)+𝟘𝟙
- | ∃∃rV,rW,rT,p,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V2 & ⦃G, L⦄ ⊢ W1 ➡[h, rW] W2 & ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[h, rT] T2 &
- I = Appl & U1 = ⓛ{p}W1.T1 & U2 = ⓓ{p}ⓝW2.V2.T2 & r = (↓rV)+(↓rW)+(↓rT)+𝟙𝟘
- | ∃∃rV,rW,rT,p,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V & ⬆*[1] V ≡ V2 & ⦃G, L⦄ ⊢ W1 ➡[h, rW] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[h, rT] T2 &
- I = Appl & U1 = ⓓ{p}W1.T1 & U2 = ⓓ{p}W2.ⓐV2.T2 & r = (↓rV)+(↓rW)+(↓rT)+𝟙𝟘.
+lemma cpg_inv_flat1: ∀c,h,I,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓕ{I}V1.U1 ➡[c, h] U2 →
+ ∨∨ ∃∃cV,cT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[cV, h] V2 & ⦃G, L⦄ ⊢ U1 ➡[cT, h] T2 &
+ U2 = ⓕ{I}V2.T2 & c = (↓cV)+cT
+ | ∃∃cT. ⦃G, L⦄ ⊢ U1 ➡[cT, h] U2 & I = Cast & c = (↓cT)+𝟙𝟘
+ | ∃∃cV. ⦃G, L⦄ ⊢ V1 ➡[cV, h] U2 & I = Cast & c = (↓cV)+𝟘𝟙
+ | ∃∃cV,cW,cT,p,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[cV, h] V2 & ⦃G, L⦄ ⊢ W1 ➡[cW, h] W2 & ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[cT, h] T2 &
+ I = Appl & U1 = ⓛ{p}W1.T1 & U2 = ⓓ{p}ⓝW2.V2.T2 & c = (↓cV)+(↓cW)+(↓cT)+𝟙𝟘
+ | ∃∃cV,cW,cT,p,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[cV, h] V & ⬆*[1] V ≡ V2 & ⦃G, L⦄ ⊢ W1 ➡[cW, h] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[cT, h] T2 &
+ I = Appl & U1 = ⓓ{p}W1.T1 & U2 = ⓓ{p}W2.ⓐV2.T2 & c = (↓cV)+(↓cW)+(↓cT)+𝟙𝟘.
/2 width=3 by cpg_inv_flat1_aux/ qed-.
-lemma cpg_inv_appl1: ∀h,r,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓐV1.U1 ➡[h, r] U2 →
- ∨∨ ∃∃rV,rT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, rT] T2 &
- U2 = ⓐV2.T2 & r = (↓rV)+rT
- | ∃∃rV,rW,rT,p,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V2 & ⦃G, L⦄ ⊢ W1 ➡[h, rW] W2 & ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[h, rT] T2 &
- U1 = ⓛ{p}W1.T1 & U2 = ⓓ{p}ⓝW2.V2.T2 & r = (↓rV)+(↓rW)+(↓rT)+𝟙𝟘
- | ∃∃rV,rW,rT,p,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V & ⬆*[1] V ≡ V2 & ⦃G, L⦄ ⊢ W1 ➡[h, rW] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[h, rT] T2 &
- U1 = ⓓ{p}W1.T1 & U2 = ⓓ{p}W2.ⓐV2.T2 & r = (↓rV)+(↓rW)+(↓rT)+𝟙𝟘.
-#h #r #G #L #V1 #U1 #U2 #H elim (cpg_inv_flat1 … H) -H *
+lemma cpg_inv_appl1: ∀c,h,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓐV1.U1 ➡[c, h] U2 →
+ ∨∨ ∃∃cV,cT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[cV, h] V2 & ⦃G, L⦄ ⊢ U1 ➡[cT, h] T2 &
+ U2 = ⓐV2.T2 & c = (↓cV)+cT
+ | ∃∃cV,cW,cT,p,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[cV, h] V2 & ⦃G, L⦄ ⊢ W1 ➡[cW, h] W2 & ⦃G, L.ⓛW1⦄ ⊢ T1 ➡[cT, h] T2 &
+ U1 = ⓛ{p}W1.T1 & U2 = ⓓ{p}ⓝW2.V2.T2 & c = (↓cV)+(↓cW)+(↓cT)+𝟙𝟘
+ | ∃∃cV,cW,cT,p,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[cV, h] V & ⬆*[1] V ≡ V2 & ⦃G, L⦄ ⊢ W1 ➡[cW, h] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ➡[cT, h] T2 &
+ U1 = ⓓ{p}W1.T1 & U2 = ⓓ{p}W2.ⓐV2.T2 & c = (↓cV)+(↓cW)+(↓cT)+𝟙𝟘.
+#c #h #G #L #V1 #U1 #U2 #H elim (cpg_inv_flat1 … H) -H *
[ /3 width=8 by or3_intro0, ex4_4_intro/
-|2,3: #r #_ #H destruct
+|2,3: #c #_ #H destruct
| /3 width=15 by or3_intro1, ex6_9_intro/
| /3 width=17 by or3_intro2, ex7_10_intro/
]
qed-.
-lemma cpg_inv_cast1: ∀h,r,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓝV1.U1 ➡[h, r] U2 →
- ∨∨ ∃∃rV,rT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, rV] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, rT] T2 &
- U2 = ⓝV2.T2 & r = (↓rV)+rT
- | ∃∃rT. ⦃G, L⦄ ⊢ U1 ➡[h, rT] U2 & r = (↓rT)+𝟙𝟘
- | ∃∃rV. ⦃G, L⦄ ⊢ V1 ➡[h, rV] U2 & r = (↓rV)+𝟘𝟙.
-#h #r #G #L #V1 #U1 #U2 #H elim (cpg_inv_flat1 … H) -H *
+lemma cpg_inv_cast1: ∀c,h,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓝV1.U1 ➡[c, h] U2 →
+ ∨∨ ∃∃cV,cT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[cV, h] V2 & ⦃G, L⦄ ⊢ U1 ➡[cT, h] T2 &
+ U2 = ⓝV2.T2 & c = (↓cV)+cT
+ | ∃∃cT. ⦃G, L⦄ ⊢ U1 ➡[cT, h] U2 & c = (↓cT)+𝟙𝟘
+ | ∃∃cV. ⦃G, L⦄ ⊢ V1 ➡[cV, h] U2 & c = (↓cV)+𝟘𝟙.
+#c #h #G #L #V1 #U1 #U2 #H elim (cpg_inv_flat1 … H) -H *
[ /3 width=8 by or3_intro0, ex4_4_intro/
|2,3: /3 width=3 by or3_intro1, or3_intro2, ex2_intro/
-| #rv #rW #rT #p #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #H destruct
-| #rv #rW #rT #p #V #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #H destruct
+| #rv #cW #cT #p #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #H destruct
+| #rv #cW #cT #p #V #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #H destruct
]
qed-.
(* Basic forward lemmas *****************************************************)
-lemma cpg_fwd_bind1_minus: ∀h,r,I,G,L,V1,T1,T. ⦃G, L⦄ ⊢ -ⓑ{I}V1.T1 ➡[h, r] T → ∀b.
- ∃∃V2,T2. ⦃G, L⦄ ⊢ ⓑ{b,I}V1.T1 ➡[h, r] ⓑ{b,I}V2.T2 &
+lemma cpg_fwd_bind1_minus: ∀c,h,I,G,L,V1,T1,T. ⦃G, L⦄ ⊢ -ⓑ{I}V1.T1 ➡[c, h] T → ∀b.
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ ⓑ{b,I}V1.T1 ➡[c, h] ⓑ{b,I}V2.T2 &
T = -ⓑ{I}V2.T2.
-#h #r #I #G #L #V1 #T1 #T #H #b elim (cpg_inv_bind1 … H) -H *
-[ #rV #rT #V2 #T2 #HV12 #HT12 #H1 #H2 destruct /3 width=4 by cpg_bind, ex2_2_intro/
-| #r #T2 #_ #_ #H destruct
+#c #h #I #G #L #V1 #T1 #T #H #b elim (cpg_inv_bind1 … H) -H *
+[ #cV #cT #V2 #T2 #HV12 #HT12 #H1 #H2 destruct /3 width=4 by cpg_bind, ex2_2_intro/
+| #c #T2 #_ #_ #H destruct
]
qed-.
(* Properties with generic slicing for local environments *******************)
-lemma cpg_delift: ∀h,r,I,G,K,V,T1,L,l. ⬇*[i] L ≡ (K.ⓑ{I}V) →
+lemma cpg_delift: ∀c,h,I,G,K,V,T1,L,i. ⬇*[i] L ≡ (K.ⓑ{I}V) →
∃∃T2,T. ⦃G, L⦄ ⊢ T1 ➡[h, 𝟘𝟘] T2 & ⬆*[↑1] T ≡ T2.
-#h #r #I #G #K #V #T1 elim T1 -T1
+#h #c #I #G #K #V #T1 elim T1 -T1
[ * #i #L #l /2 width=4 by cpg_atom, lift_sort, lift_gref, ex2_2_intro/
elim (lt_or_eq_or_gt i l) #Hil [1,3: /4 width=4 by cpg_atom, lift_lref_ge_minus, lift_lref_lt, ylt_inj, yle_inj, ex2_2_intro/ ]
destruct
]
qed-.
-lemma cpg_inv_lref1: ∀h,r,G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡[h, r] T2 →
+lemma cpg_inv_lref1: ∀h,c,G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡[h, c] T2 →
T2 = #i ∨
- ∃∃I,K,V,V2. ⬇[i] L ≡ K. ⓑ{I}V & ⦃G, K⦄ ⊢ V ➡[h, r] V2 &
+ ∃∃I,K,V,V2. ⬇[i] L ≡ K. ⓑ{I}V & ⦃G, K⦄ ⊢ V ➡[h, c] V2 &
⬆[O, i+1] V2 ≡ T2.
-#h #r #G #L #T2 #i #H
+#h #c #G #L #T2 #i #H
elim (cpg_inv_atom1 … H) -H /2 width=1 by or_introl/ *
[ #s #d #_ #_ #H destruct
| #I #K #V #V2 #j #HLK #HV2 #HVT2 #H destruct /3 width=7 by ex3_4_intro, or_intror/
(* Properties with length for local environments ****************************)
-lemma cpg_inv_lref1_ge: ∀h,r,G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡[h, r] T2 → |L| ≤ i → T2 = #i.
-#h #r #G #L #T2 #i #H elim (cpg_inv_lref1 … H) -H // *
+lemma cpg_inv_lref1_ge: ∀h,c,G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡[h, c] T2 → |L| ≤ i → T2 = #i.
+#h #c #G #L #T2 #i #H elim (cpg_inv_lref1 … H) -H // *
#I #K #V1 #V2 #HLK #_ #_ #HL -h -G -V2 lapply (drop_fwd_length_lt2 … HLK) -K -I -V1
#H elim (lt_refl_false i) /2 width=3 by lt_to_le_to_lt/
qed-.
(* Properties with restricted refinement for local environments *************)
-lemma lsubr_cpg_trans: ∀h,r,G. lsub_trans … (cpg h r G) lsubr.
-#h #r #G #L1 #T1 #T2 #H elim H -r -G -L1 -T1 -T2
+lemma lsubr_cpg_trans: ∀h,c,G. lsub_trans … (cpg h c G) lsubr.
+#h #c #G #L1 #T1 #T2 #H elim H -c -G -L1 -T1 -T2
[ //
| /2 width=2 by cpg_st/
-| #r #G #L1 #V1 #V2 #W2 #_ #HVW2 #IH #X #H
+| #c #G #L1 #V1 #V2 #W2 #_ #HVW2 #IH #X #H
elim (lsubr_inv_abbr2 … H) -H #L2 #HL21 #H destruct
/3 width=3 by cpg_delta/
-| #r #G #L1 #V1 #V2 #W2 #_ #HVW2 #IH #X #H
+| #c #G #L1 #V1 #V2 #W2 #_ #HVW2 #IH #X #H
elim (lsubr_inv_abst2 … H) -H * #L2 [2: #V ] #HL21 #H destruct
/4 width=3 by cpg_delta, cpg_lt, cpg_ct/
-| #r #I1 #G #L1 #V1 #T1 #U1 #i #_ #HTU1 #IH #X #H
+| #c #I1 #G #L1 #V1 #T1 #U1 #i #_ #HTU1 #IH #X #H
elim (lsubr_fwd_pair2 … H) -H #I2 #L2 #V2 #HL21 #H destruct
/3 width=3 by cpt_lref/
|6,11: /4 width=1 by cpg_bind, cpg_beta, lsubr_pair/
(* Properties with generic slicing for local environments *******************)
(* Note: the main property of simple terms *)
-lemma cpg_inv_appl1_simple: ∀h,r,G,L,V1,T1,U. ⦃G, L⦄ ⊢ ⓐV1.T1 ➡[h, r] U → 𝐒⦃T1⦄ →
- ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, r] V2 & ⦃G, L⦄ ⊢ T1 ➡[h, r] T2 &
+lemma cpg_inv_appl1_simple: ∀h,c,G,L,V1,T1,U. ⦃G, L⦄ ⊢ ⓐV1.T1 ➡[h, c] U → 𝐒⦃T1⦄ →
+ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, c] V2 & ⦃G, L⦄ ⊢ T1 ➡[h, c] T2 &
U = ⓐV2.T2.
-#h #r #G #L #V1 #T1 #U #H #HT1
+#h #c #G #L #V1 #T1 #U #H #HT1
elim (cpg_inv_appl1 … H) -H *
[ /2 width=5 by ex3_2_intro/
| #a #V2 #W1 #W2 #U1 #U2 #_ #_ #_ #H #_ destruct
(* Basic_1: includes: pr0_lift pr2_lift *)
lemma cpr_lift: ∀G. d_liftable (cpr G).
#G #K #T1 #T2 #H elim H -G -K -T1 -T2
-[ #I #G #K #L #c #l #k #_ #U1 #H1 #U2 #H2
+[ #I #G #K #L #b #l #k #_ #U1 #H1 #U2 #H2
>(lift_mono … H1 … H2) -H1 -H2 //
-| #G #K #KV #V #V2 #W2 #i #HKV #HV2 #HVW2 #IHV2 #L #c #l #k #HLK #U1 #H #U2 #HWU2
+| #G #K #KV #V #V2 #W2 #i #HKV #HV2 #HVW2 #IHV2 #L #b #l #k #HLK #U1 #H #U2 #HWU2
elim (lift_inv_lref1 … H) * #Hil #H destruct
[ elim (lift_trans_ge … HVW2 … HWU2) -W2 /2 width=1 by ylt_fwd_le_succ1/ #W2 #HVW2 #HWU2
elim (drop_trans_le … HLK … HKV) -K /2 width=2 by ylt_fwd_le/ #X #HLK #H
| lapply (lift_trans_be … HVW2 … HWU2 ? ?) -W2 /2 width=1 by yle_succ_dx/ >plus_plus_comm_23 #HVU2
lapply (drop_trans_ge_comm … HLK … HKV ?) -K // -Hil /3 width=6 by cpr_delta, drop_inv_gen/
]
-| #a #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #c #l #k #HLK #U1 #H1 #U2 #H2
+| #a #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #b #l #k #HLK #U1 #H1 #U2 #H2
elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 destruct
elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /4 width=6 by cpr_bind, drop_skip/
-| #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #c #l #k #HLK #U1 #H1 #U2 #H2
+| #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #b #l #k #HLK #U1 #H1 #U2 #H2
elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 destruct
elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /3 width=6 by cpr_flat/
-| #G #K #V #T1 #T #T2 #_ #HT2 #IHT1 #L #c #l #k #HLK #U1 #H #U2 #HTU2
+| #G #K #V #T1 #T #T2 #_ #HT2 #IHT1 #L #b #l #k #HLK #U1 #H #U2 #HTU2
elim (lift_inv_bind1 … H) -H #VV1 #TT1 #HVV1 #HTT1 #H destruct
elim (lift_conf_O1 … HTU2 … HT2) -T2 /4 width=6 by cpr_zeta, drop_skip/
-| #G #K #V #T1 #T2 #_ #IHT12 #L #c #l #k #HLK #U1 #H #U2 #HTU2
+| #G #K #V #T1 #T2 #_ #IHT12 #L #b #l #k #HLK #U1 #H #U2 #HTU2
elim (lift_inv_flat1 … H) -H #VV1 #TT1 #HVV1 #HTT1 #H destruct /3 width=6 by cpr_eps/
-| #a #G #K #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #L #c #l #k #HLK #X1 #HX1 #X2 #HX2
+| #a #G #K #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #L #b #l #k #HLK #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 #X #T3 #HX #HT23 #HX2 destruct
elim (lift_inv_flat1 … HX) -HX #W3 #V3 #HW23 #HV23 #HX destruct /4 width=6 by cpr_beta, drop_skip/
-| #a #G #K #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #L #c #l #k #HLK #X1 #HX1 #X2 #HX2
+| #a #G #K #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #L #b #l #k #HLK #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
(* Basic_1: includes: pr0_gen_lift pr2_gen_lift *)
lemma cpr_inv_lift1: ∀G. d_deliftable_sn (cpr G).
#G #L #U1 #U2 #H elim H -G -L -U1 -U2
-[ * #i #G #L #K #c #l #k #_ #T1 #H
+[ * #i #G #L #K #b #l #k #_ #T1 #H
[ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by ex2_intro/
| elim (lift_inv_lref2 … H) -H * #Hil #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/
| lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by ex2_intro/
]
-| #G #L #LV #V #V2 #W2 #i #HLV #HV2 #HVW2 #IHV2 #K #c #l #k #HLK #T1 #H
+| #G #L #LV #V #V2 #W2 #i #HLV #HV2 #HVW2 #IHV2 #K #b #l #k #HLK #T1 #H
elim (lift_inv_lref2 … H) -H * #Hil #H destruct
[ elim (drop_conf_lt … HLK … HLV) -L // #L #U #HKL #HLV #HUV
elim (IHV2 … HLV … HUV) -V #U2 #HUV2 #HU2
#V1 #HV1 >plus_minus // <minus_minus /2 width=1 by le_S/ <minus_n_n <plus_n_O
/3 width=8 by cpr_delta, ex2_intro/
]
-| #a #I #G #L #V1 #V2 #U1 #U2 #_ #_ #IHV12 #IHU12 #K #c #l #k #HLK #X #H
+| #a #I #G #L #V1 #V2 #U1 #U2 #_ #_ #IHV12 #IHU12 #K #b #l #k #HLK #X #H
elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
elim (IHV12 … HLK … HWV1) -IHV12 #W2 #HW12 #HWV2
elim (IHU12 … HTU1) -IHU12 -HTU1 /3 width=6 by cpr_bind, drop_skip, lift_bind, ex2_intro/
-| #I #G #L #V1 #V2 #U1 #U2 #_ #_ #IHV12 #IHU12 #K #c #l #k #HLK #X #H
+| #I #G #L #V1 #V2 #U1 #U2 #_ #_ #IHV12 #IHU12 #K #b #l #k #HLK #X #H
elim (lift_inv_flat2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
elim (IHV12 … HLK … HWV1) -V1
elim (IHU12 … HLK … HTU1) -U1 -HLK /3 width=6 by cpr_flat, lift_flat, ex2_intro/
-| #G #L #V #U1 #U #U2 #_ #HU2 #IHU1 #K #c #l #k #HLK #X #H
+| #G #L #V #U1 #U #U2 #_ #HU2 #IHU1 #K #b #l #k #HLK #X #H
elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
- elim (IHU1 (K.ⓓW1) c … HTU1) /2 width=1 by drop_skip/ -L -U1 #T #HTU #HT1
+ elim (IHU1 (K.ⓓW1) b … HTU1) /2 width=1 by drop_skip/ -L -U1 #T #HTU #HT1
elim (lift_div_le … HU2 … HTU) -U /3 width=6 by cpr_zeta, ex2_intro/
-| #G #L #V #U1 #U2 #_ #IHU12 #K #c #l #k #HLK #X #H
+| #G #L #V #U1 #U2 #_ #IHU12 #K #b #l #k #HLK #X #H
elim (lift_inv_flat2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
elim (IHU12 … HLK … HTU1) -L -U1 /3 width=3 by cpr_eps, ex2_intro/
-| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #K #c #l #k #HLK #X #HX
+| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #K #b #l #k #HLK #X #HX
elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct
elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct
elim (IHV12 … HLK … HV01) -V1
- elim (IHT12 (K.ⓛW0) c … HT01) -T1 /2 width=1 by drop_skip/
+ elim (IHT12 (K.ⓛW0) b … HT01) -T1 /2 width=1 by drop_skip/
elim (IHW12 … HLK … HW01) -W1 /4 width=7 by cpr_beta, lift_flat, lift_bind, ex2_intro/
-| #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #K #c #l #k #HLK #X #HX
+| #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #K #b #l #k #HLK #X #HX
elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct
elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct
elim (IHV1 … HLK … HV01) -V1 #V3 #HV3 #HV03
- elim (IHT12 (K.ⓓW0) c … HT01) -T1 /2 width=1 by drop_skip/ #T3 #HT32 #HT03
+ elim (IHT12 (K.ⓓW0) b … HT01) -T1 /2 width=1 by drop_skip/ #T3 #HT32 #HT03
elim (IHW12 … HLK … HW01) -W1
elim (lift_trans_le … HV3 … HV2) -V
/4 width=9 by cpr_theta, lift_flat, lift_bind, ex2_intro/
(* Properties on lazy sn pointwise extensions *******************************)
lemma cpr_llpx_sn_conf: ∀R. (∀L. reflexive … (R L)) → d_liftable R → d_deliftable_sn R →
- ∀G. c_r_confluent1 … (cpr G) (llpx_sn R 0).
+ ∀G. b_c_confluent1 … (cpr G) (llpx_sn R 0).
#R #H1R #H2R #H3R #G #Ls #T1 #T2 #H elim H -G -Ls -T1 -T2
[ //
-| #G #Ls #Ks #V1c #V2c #W2c #i #HLKs #_ #HVW2c #IHV12c #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
- #Kd #V1l #HLKd #HV1c #HV1sd
+| #G #Ls #Ks #V1b #V2b #W2b #i #HLKs #_ #HVW2b #IHV12b #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
+ #Kd #V1l #HLKd #HV1b #HV1sd
lapply (drop_fwd_drop2 … HLKs) -HLKs #HLKs
lapply (drop_fwd_drop2 … HLKd) -HLKd #HLKd
- @(llpx_sn_lift_le … HLKs HLKd … HVW2c) -HLKs -HLKd -HVW2c /2 width=1 by/ (**) (* full auto too slow *)
+ @(llpx_sn_lift_le … HLKs HLKd … HVW2b) -HLKs -HLKd -HVW2b /2 width=1 by/ (**) (* full auto too slow *)
| #a #I #G #Ls #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #Ld #H elim (llpx_sn_inv_bind_O … H) -H
/4 width=5 by llpx_sn_bind_repl_SO, llpx_sn_bind/
| #I #G #Ls #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
lemma cpx_lift: ∀h,o,G. d_liftable (cpx h o G).
#h #o #G #K #T1 #T2 #H elim H -G -K -T1 -T2
-[ #I #G #K #L #c #l #k #_ #U1 #H1 #U2 #H2
+[ #I #G #K #L #b #l #k #_ #U1 #H1 #U2 #H2
>(lift_mono … H1 … H2) -H1 -H2 //
-| #G #K #s #d #Hkd #L #c #l #k #_ #U1 #H1 #U2 #H2
+| #G #K #s #d #Hkd #L #b #l #k #_ #U1 #H1 #U2 #H2
>(lift_inv_sort1 … H1) -U1
>(lift_inv_sort1 … H2) -U2 /2 width=2 by cpx_st/
-| #I #G #K #KV #V #V2 #W2 #i #HKV #HV2 #HVW2 #IHV2 #L #c #l #k #HLK #U1 #H #U2 #HWU2
+| #I #G #K #KV #V #V2 #W2 #i #HKV #HV2 #HVW2 #IHV2 #L #b #l #k #HLK #U1 #H #U2 #HWU2
elim (lift_inv_lref1 … H) * #Hil #H destruct
[ elim (lift_trans_ge … HVW2 … HWU2) -W2 /2 width=1 by ylt_fwd_le_succ1/ #W2 #HVW2 #HWU2
elim (drop_trans_le … HLK … HKV) -K /2 width=2 by ylt_fwd_le/ #X #HLK #H
| lapply (lift_trans_be … HVW2 … HWU2 ? ?) -W2 /2 width=1 by yle_succ_dx/ >plus_plus_comm_23 #HVU2
lapply (drop_trans_ge_comm … HLK … HKV ?) -K /3 width=7 by cpx_delta, drop_inv_gen/
]
-| #a #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #c #l #k #HLK #U1 #H1 #U2 #H2
+| #a #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #b #l #k #HLK #U1 #H1 #U2 #H2
elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 destruct
elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /4 width=6 by cpx_bind, drop_skip/
-| #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #c #l #k #HLK #U1 #H1 #U2 #H2
+| #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #b #l #k #HLK #U1 #H1 #U2 #H2
elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 destruct
elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /3 width=6 by cpx_flat/
-| #G #K #V #T1 #T #T2 #_ #HT2 #IHT1 #L #c #l #k #HLK #U1 #H #U2 #HTU2
+| #G #K #V #T1 #T #T2 #_ #HT2 #IHT1 #L #b #l #k #HLK #U1 #H #U2 #HTU2
elim (lift_inv_bind1 … H) -H #VV1 #TT1 #HVV1 #HTT1 #H destruct
elim (lift_conf_O1 … HTU2 … HT2) -T2 /4 width=6 by cpx_zeta, drop_skip/
-| #G #K #V #T1 #T2 #_ #IHT12 #L #c #l #k #HLK #U1 #H #U2 #HTU2
+| #G #K #V #T1 #T2 #_ #IHT12 #L #b #l #k #HLK #U1 #H #U2 #HTU2
elim (lift_inv_flat1 … H) -H #VV1 #TT1 #HVV1 #HTT1 #H destruct /3 width=6 by cpx_eps/
-| #G #K #V1 #V2 #T #_ #IHV12 #L #c #l #k #HLK #U1 #H #U2 #HVU2
+| #G #K #V1 #V2 #T #_ #IHV12 #L #b #l #k #HLK #U1 #H #U2 #HVU2
elim (lift_inv_flat1 … H) -H #VV1 #TT1 #HVV1 #HTT1 #H destruct /3 width=6 by cpx_ct/
-| #a #G #K #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #L #c #l #k #HLK #X1 #HX1 #X2 #HX2
+| #a #G #K #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #L #b #l #k #HLK #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 #X #T3 #HX #HT23 #HX2 destruct
elim (lift_inv_flat1 … HX) -HX #W3 #V3 #HW23 #HV23 #HX destruct /4 width=6 by cpx_beta, drop_skip/
-| #a #G #K #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #L #c #l #k #HLK #X1 #HX1 #X2 #HX2
+| #a #G #K #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #L #b #l #k #HLK #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
lemma cpx_inv_lift1: ∀h,o,G. d_deliftable_sn (cpx h o G).
#h #o #G #L #U1 #U2 #H elim H -G -L -U1 -U2
-[ * #i #G #L #K #c #l #k #_ #T1 #H
+[ * #i #G #L #K #b #l #k #_ #T1 #H
[ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by cpx_atom, lift_sort, ex2_intro/
| elim (lift_inv_lref2 … H) -H * #Hil #H destruct /3 width=3 by cpx_atom, lift_lref_ge_minus, lift_lref_lt, ex2_intro/
| lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by cpx_atom, lift_gref, ex2_intro/
]
-| #G #L #s #d #Hkd #K #c #l #k #_ #T1 #H
+| #G #L #s #d #Hkd #K #b #l #k #_ #T1 #H
lapply (lift_inv_sort2 … H) -H #H destruct /3 width=3 by cpx_st, lift_sort, ex2_intro/
-| #I #G #L #LV #V #V2 #W2 #i #HLV #HV2 #HVW2 #IHV2 #K #c #l #k #HLK #T1 #H
+| #I #G #L #LV #V #V2 #W2 #i #HLV #HV2 #HVW2 #IHV2 #K #b #l #k #HLK #T1 #H
elim (lift_inv_lref2 … H) -H * #Hil #H destruct
[ elim (drop_conf_lt … HLK … HLV) -L // #L #U #HKL #HLV #HUV
elim (IHV2 … HLV … HUV) -V #U2 #HUV2 #HU2
elim (lift_split … HVW2 l (i - k + 1)) -HVW2 /3 width=1 by yle_succ, yle_pred_sn, le_S_S/ -Hil -Hlim
#V1 #HV1 >plus_minus // <minus_minus /2 width=1 by le_S/ <minus_n_n <plus_n_O /3 width=9 by cpx_delta, ex2_intro/
]
-| #a #I #G #L #V1 #V2 #U1 #U2 #_ #_ #IHV12 #IHU12 #K #c #l #k #HLK #X #H
+| #a #I #G #L #V1 #V2 #U1 #U2 #_ #_ #IHV12 #IHU12 #K #b #l #k #HLK #X #H
elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
elim (IHV12 … HLK … HWV1) -IHV12 #W2 #HW12 #HWV2
elim (IHU12 … HTU1) -IHU12 -HTU1 /3 width=6 by cpx_bind, drop_skip, lift_bind, ex2_intro/
-| #I #G #L #V1 #V2 #U1 #U2 #_ #_ #IHV12 #IHU12 #K #c #l #k #HLK #X #H
+| #I #G #L #V1 #V2 #U1 #U2 #_ #_ #IHV12 #IHU12 #K #b #l #k #HLK #X #H
elim (lift_inv_flat2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
elim (IHV12 … HLK … HWV1) -V1
elim (IHU12 … HLK … HTU1) -U1 -HLK /3 width=5 by cpx_flat, lift_flat, ex2_intro/
-| #G #L #V #U1 #U #U2 #_ #HU2 #IHU1 #K #c #l #k #HLK #X #H
+| #G #L #V #U1 #U #U2 #_ #HU2 #IHU1 #K #b #l #k #HLK #X #H
elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
- elim (IHU1 (K.ⓓW1) c … HTU1) /2 width=1 by drop_skip/ -L -U1 #T #HTU #HT1
+ elim (IHU1 (K.ⓓW1) b … HTU1) /2 width=1 by drop_skip/ -L -U1 #T #HTU #HT1
elim (lift_div_le … HU2 … HTU) -U /3 width=5 by cpx_zeta, ex2_intro/
-| #G #L #V #U1 #U2 #_ #IHU12 #K #c #l #k #HLK #X #H
+| #G #L #V #U1 #U2 #_ #IHU12 #K #b #l #k #HLK #X #H
elim (lift_inv_flat2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
elim (IHU12 … HLK … HTU1) -L -U1 /3 width=3 by cpx_eps, ex2_intro/
-| #G #L #V1 #V2 #U1 #_ #IHV12 #K #c #l #k #HLK #X #H
+| #G #L #V1 #V2 #U1 #_ #IHV12 #K #b #l #k #HLK #X #H
elim (lift_inv_flat2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct
elim (IHV12 … HLK … HWV1) -L -V1 /3 width=3 by cpx_ct, ex2_intro/
-| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #K #c #l #k #HLK #X #HX
+| #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #K #b #l #k #HLK #X #HX
elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct
elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct
elim (IHV12 … HLK … HV01) -V1 #V3 #HV32 #HV03
- elim (IHT12 (K.ⓛW0) c … HT01) -T1 /2 width=1 by drop_skip/ #T3 #HT32 #HT03
+ elim (IHT12 (K.ⓛW0) b … HT01) -T1 /2 width=1 by drop_skip/ #T3 #HT32 #HT03
elim (IHW12 … HLK … HW01) -W1
/4 width=7 by cpx_beta, lift_bind, lift_flat, ex2_intro/
-| #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #K #c #l #k #HLK #X #HX
+| #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #K #b #l #k #HLK #X #HX
elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct
elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct
elim (IHV1 … HLK … HV01) -V1 #V3 #HV3 #HV03
- elim (IHT12 (K.ⓓW0) c … HT01) -T1 /2 width=1 by drop_skip/ #T3 #HT32 #HT03
+ elim (IHT12 (K.ⓓW0) b … HT01) -T1 /2 width=1 by drop_skip/ #T3 #HT32 #HT03
elim (IHW12 … HLK … HW01) -W1 #W3 #HW32 #HW03
elim (lift_trans_le … HV3 … HV2) -V
/4 width=9 by cpx_theta, lift_bind, lift_flat, ex2_intro/
∀L1. L2 ≡[T1, 0] L1 → ⦃G, L1⦄ ⊢ T1 ➡[h, o] T2.
/3 width=3 by lleq_cpx_trans, lleq_sym/ qed-.
-lemma cpx_lleq_conf_sn: ∀h,o,G. c_r_confluent1 … (cpx h o G) (lleq 0).
+lemma cpx_lleq_conf_sn: ∀h,o,G. b_c_confluent1 … (cpx h o G) (lleq 0).
/3 width=6 by cpx_llpx_sn_conf, lift_mono, ex2_intro/ qed-.
lemma cpx_lleq_conf_dx: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ➡[h, o] T2 →
(* Note: lemma 1000 *)
lemma cpx_llpx_sn_conf: ∀R. (∀L. reflexive … (R L)) → d_liftable R → d_deliftable_sn R →
- ∀h,o,G. c_r_confluent1 … (cpx h o G) (llpx_sn R 0).
+ ∀h,o,G. b_c_confluent1 … (cpx h o G) (llpx_sn R 0).
#R #H1R #H2R #H3R #h #o #G #Ls #T1 #T2 #H elim H -G -Ls -T1 -T2
[ //
| /3 width=4 by llpx_sn_fwd_length, llpx_sn_sort/
-| #I #G #Ls #Ks #V1c #V2c #W2c #i #HLKs #_ #HVW2c #IHV12c #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
- #Kd #V1l #HLKd #HV1c #HV1sd
+| #I #G #Ls #Ks #V1b #V2b #W2b #i #HLKs #_ #HVW2b #IHV12b #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H
+ #Kd #V1l #HLKd #HV1b #HV1sd
lapply (drop_fwd_drop2 … HLKs) -HLKs #HLKs
lapply (drop_fwd_drop2 … HLKd) -HLKd #HLKd
- @(llpx_sn_lift_le … HLKs HLKd … HVW2c) -HLKs -HLKd -HVW2c /2 width=1 by/ (**) (* full auto too slow *)
+ @(llpx_sn_lift_le … HLKs HLKd … HVW2b) -HLKs -HLKd -HVW2b /2 width=1 by/ (**) (* full auto too slow *)
| #a #I #G #Ls #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #Ld #H elim (llpx_sn_inv_bind_O … H) -H
/4 width=5 by llpx_sn_bind_repl_SO, llpx_sn_bind/
| #I #G #Ls #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H
(* Properties with generic slicing for local environments *******************)
(* Basic_2A1: includes: aaa_lift *)
-lemma aaa_lifts: ∀G,L1,T1,A. ⦃G, L1⦄ ⊢ T1 ⁝ A → ∀L2,c,f. ⬇*[c, f] L2 ≡ L1 →
+lemma aaa_lifts: ∀G,L1,T1,A. ⦃G, L1⦄ ⊢ T1 ⁝ A → ∀b,f,L2. ⬇*[b, f] L2 ≡ L1 →
∀T2. ⬆*[f] T1 ≡ T2 → ⦃G, L2⦄ ⊢ T2 ⁝ A.
@fqup_wf_ind_eq #G0 #L0 #T0 #IH #G #L1 * *
-[ #s #HG #HL #HT #A #H #L2 #c #f #HL21 #X #HX -IH -c
+[ #s #HG #HL #HT #A #H #b #f #L2 #HL21 #X #HX -b -IH
lapply (aaa_inv_sort … H) -H #H destruct
>(lifts_inv_sort1 … HX) -HX //
-| #i1 #HG #HL #HT #A #H #L2 #c #f #HL21 #X #HX
+| #i1 #HG #HL #HT #A #H #b #f #L2 #HL21 #X #HX
elim (aaa_inv_lref_gen … H) -H #J #K1 #V1 #HLK1 #HA
elim (lifts_inv_lref1 … HX) -HX #i2 #Hf #H destruct
lapply (drops_trans … HL21 … HLK1 ??) -HL21 [1,2: // ] #H
lapply (drops_inv_tls_at … Hf … HY) -HY #HY -Hf
elim (drops_inv_skip2 … HY) -HY #K2 #V2 #HK21 #HV12 #H destruct
/4 width=12 by aaa_lref_gen, fqup_lref, drops_inv_gen/
-| #l #HG #HL #HT #A #H #L2 #c #f #HL21 #X #HX -IH -c -f
+| #l #HG #HL #HT #A #H #b #f #L2 #HL21 #X #HX -b -f -IH
elim (aaa_inv_gref … H)
-| #p * #V1 #T1 #HG #HL #HT #A #H #L2 #c #f #HL21 #X #HX
+| #p * #V1 #T1 #HG #HL #HT #A #H #b #f #L2 #HL21 #X #HX
[ elim (aaa_inv_abbr … H) -H #B #HB #HA
elim (lifts_inv_bind1 … HX) -HX #V2 #T2 #HV12 #HT12 #H destruct
/4 width=9 by aaa_abbr, drops_skip/
elim (lifts_inv_bind1 … HX) -HX #V2 #T2 #HV12 #HT12 #H destruct
/4 width=8 by aaa_abst, drops_skip/
]
-| * #V1 #T1 #HG #HL #HT #A #H #L2 #c #f #HL21 #X #HX
+| * #V1 #T1 #HG #HL #HT #A #H #b #f #L2 #HL21 #X #HX
[ elim (aaa_inv_appl … H) -H #B #HB #HA
elim (lifts_inv_flat1 … HX) -HX #V2 #T2 #HV12 #HT12 #H destruct
/3 width=10 by aaa_appl/
(* Inversion lemmas with generic slicing for local environments *************)
(* Basic_2A1: includes: aaa_inv_lift *)
-lemma aaa_inv_lifts: ∀G,L2,T2,A. ⦃G, L2⦄ ⊢ T2 ⁝ A → ∀L1,c,f. ⬇*[c, f] L2 ≡ L1 →
+lemma aaa_inv_lifts: ∀G,L2,T2,A. ⦃G, L2⦄ ⊢ T2 ⁝ A → ∀b,f,L1. ⬇*[b, f] L2 ≡ L1 →
∀T1. ⬆*[f] T1 ≡ T2 → ⦃G, L1⦄ ⊢ T1 ⁝ A.
@fqup_wf_ind_eq #G0 #L0 #T0 #IH #G #L2 * *
-[ #s #HG #HL #HT #A #H #L1 #c #f #HL21 #X #HX -IH -c
+[ #s #HG #HL #HT #A #H #b #f #L1 #HL21 #X #HX -b -IH
lapply (aaa_inv_sort … H) -H #H destruct
>(lifts_inv_sort2 … HX) -HX //
-| #i2 #HG #HL #HT #A #H #L1 #c #f #HL21 #X #HX
+| #i2 #HG #HL #HT #A #H #b #f #L1 #HL21 #X #HX
elim (aaa_inv_lref_gen … H) -H #J #K2 #V2 #HLK2 #HA
elim (lifts_inv_lref2 … HX) -HX #i1 #Hf #H destruct
lapply (drops_split_div … HL21 (𝐔❴i1❵) ???) -HL21 [4: * |*: // ] #Y #HLK1 #HY
lapply (drops_inv_tls_at … Hf … HY) -HY #HY -Hf
elim (drops_inv_skip1 … HY) -HY #K1 #V1 #HK21 #HV12 #H destruct
/4 width=12 by aaa_lref_gen, fqup_lref, drops_inv_F/
-| #l #HG #HL #HT #A #H #L1 #c #f #HL21 #X #HX -IH -c -f
+| #l #HG #HL #HT #A #H #b #f #L1 #HL21 #X #HX -IH -b -f
elim (aaa_inv_gref … H)
-| #p * #V2 #T2 #HG #HL #HT #A #H #L1 #c #f #HL21 #X #HX
+| #p * #V2 #T2 #HG #HL #HT #A #H #b #f #L1 #HL21 #X #HX
[ elim (aaa_inv_abbr … H) -H #B #HB #HA
elim (lifts_inv_bind2 … HX) -HX #V1 #T1 #HV12 #HT12 #H destruct
/4 width=9 by aaa_abbr, drops_skip/
elim (lifts_inv_bind2 … HX) -HX #V1 #T1 #HV12 #HT12 #H destruct
/4 width=8 by aaa_abst, drops_skip/
]
-| * #V2 #T2 #HG #HL #HT #A #H #L1 #c #f #HL21 #X #HX
+| * #V2 #T2 #HG #HL #HT #A #H #b #f #L1 #HL21 #X #HX
[ elim (aaa_inv_appl … H) -H #B #HB #HA
elim (lifts_inv_flat2 … HX) -HX #V1 #T1 #HV12 #HT12 #H destruct
/3 width=10 by aaa_appl/
(* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
inductive frees: relation3 lenv term rtmap ≝
-| frees_atom: ∀I,f. 𝐈⦃f⦄ → frees (⋆) (⓪{I}) f
-| frees_sort: ∀I,L,V,s,f. frees L (⋆s) f →
+| frees_atom: ∀f,I. 𝐈⦃f⦄ → frees (⋆) (⓪{I}) f
+| frees_sort: ∀f,I,L,V,s. frees L (⋆s) f →
frees (L.ⓑ{I}V) (⋆s) (↑f)
-| frees_zero: ∀I,L,V,f. frees L V f →
+| frees_zero: ∀f,I,L,V. frees L V f →
frees (L.ⓑ{I}V) (#0) (⫯f)
-| frees_lref: ∀I,L,V,i,f. frees L (#i) f →
+| frees_lref: ∀f,I,L,V,i. frees L (#i) f →
frees (L.ⓑ{I}V) (#⫯i) (↑f)
-| frees_gref: ∀I,L,V,p,f. frees L (§p) f →
- frees (L.ⓑ{I}V) (§p) (↑f)
-| frees_bind: ∀I,L,V,T,a,f1,f2,f. frees L V f1 → frees (L.ⓑ{I}V) T f2 →
- f1 ⋓ ⫱f2 ≡ f → frees L (ⓑ{a,I}V.T) f
-| frees_flat: ∀I,L,V,T,f1,f2,f. frees L V f1 → frees L T f2 →
+| frees_gref: ∀f,I,L,V,l. frees L (§l) f →
+ frees (L.ⓑ{I}V) (§l) (↑f)
+| frees_bind: ∀f1,f2,f,p,I,L,V,T. frees L V f1 → frees (L.ⓑ{I}V) T f2 →
+ f1 ⋓ ⫱f2 ≡ f → frees L (ⓑ{p,I}V.T) f
+| frees_flat: ∀f1,f2,f,I,L,V,T. frees L V f1 → frees L T f2 →
f1 ⋓ f2 ≡ f → frees L (ⓕ{I}V.T) f
.
(* Basic inversion lemmas ***************************************************)
-fact frees_inv_atom_aux: ∀L,X,f. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀J. L = ⋆ → X = ⓪{J} → 𝐈⦃f⦄.
-#L #X #f #H elim H -L -X -f /3 width=3 by isid_push/
-[5,6: #I #L #V #T [ #p ] #f1 #f2 #f #_ #_ #_ #_ #_ #J #_ #H destruct
-|*: #I #L #V [1,3,4: #x ] #f #_ #_ #J #H destruct
+fact frees_inv_atom_aux: ∀f,L,X. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀J. L = ⋆ → X = ⓪{J} → 𝐈⦃f⦄.
+#f #L #X #H elim H -f -L -X /3 width=3 by isid_push/
+[5,6: #f1 #f2 #f [ #p ] #I #L #V #T #_ #_ #_ #_ #_ #J #_ #H destruct
+|*: #f #I #L #V [1,3,4: #x ] #_ #_ #J #H destruct
]
qed-.
-lemma frees_inv_atom: ∀I,f. ⋆ ⊢ 𝐅*⦃⓪{I}⦄ ≡ f → 𝐈⦃f⦄.
+lemma frees_inv_atom: ∀f,I. ⋆ ⊢ 𝐅*⦃⓪{I}⦄ ≡ f → 𝐈⦃f⦄.
/2 width=6 by frees_inv_atom_aux/ qed-.
-fact frees_inv_sort_aux: ∀L,X,f. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀x. X = ⋆x → 𝐈⦃f⦄.
-#L #X #f #H elim H -L -X -f /3 width=3 by isid_push/
-[ #_ #L #V #f #_ #_ #x #H destruct
-| #_ #L #_ #i #f #_ #_ #x #H destruct
-| #I #L #V #T #p #f1 #f2 #f #_ #_ #_ #_ #_ #x #H destruct
-| #I #L #V #T #f1 #f2 #f #_ #_ #_ #_ #_ #x #H destruct
+fact frees_inv_sort_aux: ∀f,L,X. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀x. X = ⋆x → 𝐈⦃f⦄.
+#L #X #f #H elim H -f -L -X /3 width=3 by isid_push/
+[ #f #_ #L #V #_ #_ #x #H destruct
+| #f #_ #L #_ #i #_ #_ #x #H destruct
+| #f1 #f2 #f #p #I #L #V #T #_ #_ #_ #_ #_ #x #H destruct
+| #f1 #f2 #f #I #L #V #T #_ #_ #_ #_ #_ #x #H destruct
]
qed-.
-lemma frees_inv_sort: ∀L,s,f. L ⊢ 𝐅*⦃⋆s⦄ ≡ f → 𝐈⦃f⦄.
+lemma frees_inv_sort: ∀f,L,s. L ⊢ 𝐅*⦃⋆s⦄ ≡ f → 𝐈⦃f⦄.
/2 width=5 by frees_inv_sort_aux/ qed-.
-fact frees_inv_gref_aux: ∀L,X,f. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀x. X = §x → 𝐈⦃f⦄.
-#L #X #f #H elim H -L -X -f /3 width=3 by isid_push/
-[ #_ #L #V #f #_ #_ #x #H destruct
-| #_ #L #_ #i #f #_ #_ #x #H destruct
-| #I #L #V #T #p #f1 #f2 #f #_ #_ #_ #_ #_ #x #H destruct
-| #I #L #V #T #f1 #f2 #f #_ #_ #_ #_ #_ #x #H destruct
+fact frees_inv_gref_aux: ∀f,L,X. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀x. X = §x → 𝐈⦃f⦄.
+#f #L #X #H elim H -f -L -X /3 width=3 by isid_push/
+[ #f #_ #L #V #_ #_ #x #H destruct
+| #f #_ #L #_ #i #_ #_ #x #H destruct
+| #f1 #f2 #f #p #I #L #V #T #_ #_ #_ #_ #_ #x #H destruct
+| #f1 #f2 #f #I #L #V #T #_ #_ #_ #_ #_ #x #H destruct
]
qed-.
-lemma frees_inv_gref: ∀L,l,f. L ⊢ 𝐅*⦃§l⦄ ≡ f → 𝐈⦃f⦄.
+lemma frees_inv_gref: ∀f,L,l. L ⊢ 𝐅*⦃§l⦄ ≡ f → 𝐈⦃f⦄.
/2 width=5 by frees_inv_gref_aux/ qed-.
-fact frees_inv_zero_aux: ∀L,X,f. L ⊢ 𝐅*⦃X⦄ ≡ f → X = #0 →
+fact frees_inv_zero_aux: ∀f,L,X. L ⊢ 𝐅*⦃X⦄ ≡ f → X = #0 →
(L = ⋆ ∧ 𝐈⦃f⦄) ∨
- ∃∃I,K,V,g. K ⊢ 𝐅*⦃V⦄ ≡ g & L = K.ⓑ{I}V & f = ⫯g.
-#L #X #f * -L -X -f
+ ∃∃g,I,K,V. K ⊢ 𝐅*⦃V⦄ ≡ g & L = K.ⓑ{I}V & f = ⫯g.
+#f #L #X * -f -L -X
[ /3 width=1 by or_introl, conj/
-| #I #L #V #s #f #_ #H destruct
+| #f #I #L #V #s #_ #H destruct
| /3 width=7 by ex3_4_intro, or_intror/
-| #I #L #V #i #f #_ #H destruct
-| #I #L #V #l #f #_ #H destruct
-| #I #L #V #T #p #f1 #f2 #f #_ #_ #_ #H destruct
-| #I #L #V #T #f1 #f2 #f #_ #_ #_ #H destruct
+| #f #I #L #V #i #_ #H destruct
+| #f #I #L #V #l #_ #H destruct
+| #f1 #f2 #f #p #I #L #V #T #_ #_ #_ #H destruct
+| #f1 #f2 #f #I #L #V #T #_ #_ #_ #H destruct
]
qed-.
-lemma frees_inv_zero: ∀L,f. L ⊢ 𝐅*⦃#0⦄ ≡ f →
+lemma frees_inv_zero: ∀f,L. L ⊢ 𝐅*⦃#0⦄ ≡ f →
(L = ⋆ ∧ 𝐈⦃f⦄) ∨
- ∃∃I,K,V,g. K ⊢ 𝐅*⦃V⦄ ≡ g & L = K.ⓑ{I}V & f = ⫯g.
+ ∃∃g,I,K,V. K ⊢ 𝐅*⦃V⦄ ≡ g & L = K.ⓑ{I}V & f = ⫯g.
/2 width=3 by frees_inv_zero_aux/ qed-.
-fact frees_inv_lref_aux: ∀L,X,f. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀j. X = #(⫯j) →
+fact frees_inv_lref_aux: ∀f,L,X. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀j. X = #(⫯j) →
(L = ⋆ ∧ 𝐈⦃f⦄) ∨
- ∃∃I,K,V,g. K ⊢ 𝐅*⦃#j⦄ ≡ g & L = K.ⓑ{I}V & f = ↑g.
-#L #X #f * -L -X -f
+ ∃∃g,I,K,V. K ⊢ 𝐅*⦃#j⦄ ≡ g & L = K.ⓑ{I}V & f = ↑g.
+#f #L #X * -f -L -X
[ /3 width=1 by or_introl, conj/
-| #I #L #V #s #f #_ #j #H destruct
-| #I #L #V #f #_ #j #H destruct
-| #I #L #V #i #f #Ht #j #H destruct /3 width=7 by ex3_4_intro, or_intror/
-| #I #L #V #l #f #_ #j #H destruct
-| #I #L #V #T #p #f1 #f2 #f #_ #_ #_ #j #H destruct
-| #I #L #V #T #f1 #f2 #f #_ #_ #_ #j #H destruct
+| #f #I #L #V #s #_ #j #H destruct
+| #f #I #L #V #_ #j #H destruct
+| #f #I #L #V #i #Hf #j #H destruct /3 width=7 by ex3_4_intro, or_intror/
+| #f #I #L #V #l #_ #j #H destruct
+| #f1 #f2 #f #p #I #L #V #T #_ #_ #_ #j #H destruct
+| #f1 #f2 #f #I #L #V #T #_ #_ #_ #j #H destruct
]
qed-.
-lemma frees_inv_lref: ∀L,i,f. L ⊢ 𝐅*⦃#(⫯i)⦄ ≡ f →
+lemma frees_inv_lref: ∀f,L,i. L ⊢ 𝐅*⦃#(⫯i)⦄ ≡ f →
(L = ⋆ ∧ 𝐈⦃f⦄) ∨
- ∃∃I,K,V,g. K ⊢ 𝐅*⦃#i⦄ ≡ g & L = K.ⓑ{I}V & f = ↑g.
+ ∃∃g,I,K,V. K ⊢ 𝐅*⦃#i⦄ ≡ g & L = K.ⓑ{I}V & f = ↑g.
/2 width=3 by frees_inv_lref_aux/ qed-.
-fact frees_inv_bind_aux: ∀L,X,f. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀I,V,T,a. X = ⓑ{a,I}V.T →
+fact frees_inv_bind_aux: ∀f,L,X. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀p,I,V,T. X = ⓑ{p,I}V.T →
∃∃f1,f2. L ⊢ 𝐅*⦃V⦄ ≡ f1 & L.ⓑ{I}V ⊢ 𝐅*⦃T⦄ ≡ f2 & f1 ⋓ ⫱f2 ≡ f.
-#L #X #f * -L -X -f
-[ #I #f #_ #J #W #U #b #H destruct
-| #I #L #V #s #f #_ #J #W #U #b #H destruct
-| #I #L #V #f #_ #J #W #U #b #H destruct
-| #I #L #V #i #f #_ #J #W #U #b #H destruct
-| #I #L #V #l #f #_ #J #W #U #b #H destruct
-| #I #L #V #T #p #f1 #f2 #f #HV #HT #Hf #J #W #U #b #H destruct /2 width=5 by ex3_2_intro/
-| #I #L #V #T #f1 #f2 #f #_ #_ #_ #J #W #U #b #H destruct
+#f #L #X * -f -L -X
+[ #f #I #_ #q #J #W #U #H destruct
+| #f #I #L #V #s #_ #q #J #W #U #H destruct
+| #f #I #L #V #_ #q #J #W #U #H destruct
+| #f #I #L #V #i #_ #q #J #W #U #H destruct
+| #f #I #L #V #l #_ #q #J #W #U #H destruct
+| #f1 #f2 #f #p #I #L #V #T #HV #HT #Hf #q #J #W #U #H destruct /2 width=5 by ex3_2_intro/
+| #f1 #f2 #f #I #L #V #T #_ #_ #_ #q #J #W #U #H destruct
]
qed-.
-lemma frees_inv_bind: ∀I,L,V,T,a,f. L ⊢ 𝐅*⦃ⓑ{a,I}V.T⦄ ≡ f →
+lemma frees_inv_bind: ∀f,p,I,L,V,T. L ⊢ 𝐅*⦃ⓑ{p,I}V.T⦄ ≡ f →
∃∃f1,f2. L ⊢ 𝐅*⦃V⦄ ≡ f1 & L.ⓑ{I}V ⊢ 𝐅*⦃T⦄ ≡ f2 & f1 ⋓ ⫱f2 ≡ f.
/2 width=4 by frees_inv_bind_aux/ qed-.
-fact frees_inv_flat_aux: ∀L,X,f. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀I,V,T. X = ⓕ{I}V.T →
+fact frees_inv_flat_aux: ∀f,L,X. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀I,V,T. X = ⓕ{I}V.T →
∃∃f1,f2. L ⊢ 𝐅*⦃V⦄ ≡ f1 & L ⊢ 𝐅*⦃T⦄ ≡ f2 & f1 ⋓ f2 ≡ f.
-#L #X #f * -L -X -f
-[ #I #f #_ #J #W #U #H destruct
-| #I #L #V #s #f #_ #J #W #U #H destruct
-| #I #L #V #f #_ #J #W #U #H destruct
-| #I #L #V #i #f #_ #J #W #U #H destruct
-| #I #L #V #l #f #_ #J #W #U #H destruct
-| #I #L #V #T #p #f1 #f2 #f #_ #_ #_ #J #W #U #H destruct
-| #I #L #V #T #f1 #f2 #f #HV #HT #Hf #J #W #U #H destruct /2 width=5 by ex3_2_intro/
+#f #L #X * -f -L -X
+[ #f #I #_ #J #W #U #H destruct
+| #f #I #L #V #s #_ #J #W #U #H destruct
+| #f #I #L #V #_ #J #W #U #H destruct
+| #f #I #L #V #i #_ #J #W #U #H destruct
+| #f #I #L #V #l #_ #J #W #U #H destruct
+| #f1 #f2 #f #p #I #L #V #T #_ #_ #_ #J #W #U #H destruct
+| #f1 #f2 #f #I #L #V #T #HV #HT #Hf #J #W #U #H destruct /2 width=5 by ex3_2_intro/
]
qed-.
-lemma frees_inv_flat: ∀I,L,V,T,f. L ⊢ 𝐅*⦃ⓕ{I}V.T⦄ ≡ f →
+lemma frees_inv_flat: ∀f,I,L,V,T. L ⊢ 𝐅*⦃ⓕ{I}V.T⦄ ≡ f →
∃∃f1,f2. L ⊢ 𝐅*⦃V⦄ ≡ f1 & L ⊢ 𝐅*⦃T⦄ ≡ f2 & f1 ⋓ f2 ≡ f.
/2 width=4 by frees_inv_flat_aux/ qed-.
(* Basic forward lemmas ****************************************************)
-lemma frees_fwd_isfin: ∀L,T,f. L ⊢ 𝐅*⦃T⦄ ≡ f → 𝐅⦃f⦄.
-#L #T #f #H elim H -L -T -f
+lemma frees_fwd_isfin: ∀f,L,T. L ⊢ 𝐅*⦃T⦄ ≡ f → 𝐅⦃f⦄.
+#f #L #T #H elim H -f -L -T
/3 width=5 by sor_isfin, isfin_isid, isfin_tl, isfin_push, isfin_next/
qed-.
(* Basic properties ********************************************************)
lemma frees_eq_repl_back: ∀L,T. eq_repl_back … (λf. L ⊢ 𝐅*⦃T⦄ ≡ f).
-#L #T #f1 #H elim H -L -T -f1
+#L #T #f1 #H elim H -f1 -L -T
[ /3 width=3 by frees_atom, isid_eq_repl_back/
-| #I #L #V #s #f1 #_ #IH #f2 #Hf12
+| #f1 #I #L #V #s #_ #IH #f2 #Hf12
elim (eq_inv_px … Hf12) -Hf12 /3 width=3 by frees_sort/
-| #I #L #V #f1 #_ #IH #f2 #Hf12
+| #f1 #I #L #V #_ #IH #f2 #Hf12
elim (eq_inv_nx … Hf12) -Hf12 /3 width=3 by frees_zero/
-| #I #L #V #i #f1 #_ #IH #f2 #Hf12
+| #f1 #I #L #V #i #_ #IH #f2 #Hf12
elim (eq_inv_px … Hf12) -Hf12 /3 width=3 by frees_lref/
-| #I #L #V #l #f1 #_ #IH #f2 #Hf12
+| #f1 #I #L #V #l #_ #IH #f2 #Hf12
elim (eq_inv_px … Hf12) -Hf12 /3 width=3 by frees_gref/
| /3 width=7 by frees_bind, sor_eq_repl_back3/
| /3 width=7 by frees_flat, sor_eq_repl_back3/
#L #T @eq_repl_sym /2 width=3 by frees_eq_repl_back/
qed-.
-lemma frees_sort_gen: ∀L,s,f. 𝐈⦃f⦄ → L ⊢ 𝐅*⦃⋆s⦄ ≡ f.
-#L elim L -L
+lemma frees_sort_gen: ∀f,L,s. 𝐈⦃f⦄ → L ⊢ 𝐅*⦃⋆s⦄ ≡ f.
+#f #L elim L -L
/4 width=3 by frees_eq_repl_back, frees_sort, frees_atom, eq_push_inv_isid/
qed.
-lemma frees_gref_gen: ∀L,p,f. 𝐈⦃f⦄ → L ⊢ 𝐅*⦃§p⦄ ≡ f.
-#L elim L -L
+lemma frees_gref_gen: ∀f,L,p. 𝐈⦃f⦄ → L ⊢ 𝐅*⦃§p⦄ ≡ f.
+#f #L elim L -L
/4 width=3 by frees_eq_repl_back, frees_gref, frees_atom, eq_push_inv_isid/
qed.
(* Main inversion lemmas ****************************************************)
-theorem frees_mono: ∀L,T,f1. L ⊢ 𝐅*⦃T⦄ ≡ f1 → ∀f2. L ⊢ 𝐅*⦃T⦄ ≡ f2 → f1 ≗ f2.
-#L #T #f1 #H elim H -L -T -f1
+theorem frees_mono: ∀f1,L,T. L ⊢ 𝐅*⦃T⦄ ≡ f1 → ∀f2. L ⊢ 𝐅*⦃T⦄ ≡ f2 → f1 ≗ f2.
+#f1 #L #T #H elim H -f1 -L -T
[ /3 width=2 by frees_inv_atom, isid_inv_eq_repl/
| /4 width=5 by frees_inv_sort, eq_push_inv_isid, isid_inv_eq_repl, eq_trans/
-| #I #L #V #f1 #_ #IH #x #H elim (frees_inv_zero … H) -H *
+| #f1 #I #L #V #_ #IH #x #H elim (frees_inv_zero … H) -H *
[ #H destruct
- | #Z #Y #X #f2 #Hf2 #H1 #H2 destruct /3 width=5 by eq_next/
+ | #f2 #Z #Y #X #Hf2 #H1 #H2 destruct /3 width=5 by eq_next/
]
-| #I #L #V #i #f1 #_ #IH #x #H elim (frees_inv_lref … H) -H *
+| #f1 #I #L #V #i #_ #IH #x #H elim (frees_inv_lref … H) -H *
[ #H destruct
- | #Z #Y #X #f2 #Hf2 #H1 #H2 destruct /3 width=5 by eq_push/
+ | #f2 #Z #Y #X #Hf2 #H1 #H2 destruct /3 width=5 by eq_push/
]
| /4 width=5 by frees_inv_gref, eq_push_inv_isid, isid_inv_eq_repl, eq_trans/
-| #I #L #V #T #p #f1 #f2 #f #_ #_ #Hf #IHV #IHT #g #H elim (frees_inv_bind … H) -H
+| #f1 #f2 #f #p #I #L #V #T #_ #_ #Hf #IHV #IHT #g #H elim (frees_inv_bind … H) -H
#g1 #g2 #HV #HT #Hg @(sor_mono … Hf) -Hf
/5 width=3 by sor_eq_repl_fwd2, sor_eq_repl_fwd1, tl_eq_repl/ (**) (* full auto too slow *)
-| #I #L #V #T #f1 #f2 #f #_ #_ #Hf #IHV #IHT #g #H elim (frees_inv_flat … H) -H
+| #f1 #f2 #f #I #L #V #T #_ #_ #Hf #IHV #IHT #g #H elim (frees_inv_flat … H) -H
#g1 #g2 #HV #HT #Hg @(sor_mono … Hf) -Hf
/4 width=3 by sor_eq_repl_fwd2, sor_eq_repl_fwd1/ (**) (* full auto too slow *)
]
(* Properties with ranged equivalence for local environments ****************)
-lemma frees_lreq_conf: ∀L1,T,f. L1 ⊢ 𝐅*⦃T⦄ ≡ f → ∀L2. L1 ≡[f] L2 → L2 ⊢ 𝐅*⦃T⦄ ≡ f.
-#L1 #T #f #H elim H -L1 -T -f
-[ #I #f #Hf #X #H lapply (lreq_inv_atom1 … H) -H
+lemma frees_lreq_conf: ∀f,L1,T. L1 ⊢ 𝐅*⦃T⦄ ≡ f → ∀L2. L1 ≡[f] L2 → L2 ⊢ 𝐅*⦃T⦄ ≡ f.
+#f #L1 #T #H elim H -f -L1 -T
+[ #f #I #Hf #X #H lapply (lreq_inv_atom1 … H) -H
#H destruct /2 width=1 by frees_atom/
-| #I #L1 #V1 #s #f #_ #IH #X #H elim (lreq_inv_push1 … H) -H
+| #f #I #L1 #V1 #s #_ #IH #X #H elim (lreq_inv_push1 … H) -H
/3 width=1 by frees_sort/
-| #I #L1 #V1 #f #_ #IH #X #H elim (lreq_inv_next1 … H) -H
+| #f #I #L1 #V1 #_ #IH #X #H elim (lreq_inv_next1 … H) -H
/3 width=1 by frees_zero/
-| #I #L1 #V1 #i #f #_ #IH #X #H elim (lreq_inv_push1 … H) -H
+| #f #I #L1 #V1 #i #_ #IH #X #H elim (lreq_inv_push1 … H) -H
/3 width=1 by frees_lref/
-| #I #L1 #V1 #l #f #_ #IH #X #H elim (lreq_inv_push1 … H) -H
+| #f #I #L1 #V1 #l #_ #IH #X #H elim (lreq_inv_push1 … H) -H
/3 width=1 by frees_gref/
| /6 width=5 by frees_bind, lreq_inv_tl, sle_lreq_trans, sor_inv_sle_dx, sor_inv_sle_sn/
| /5 width=5 by frees_flat, sle_lreq_trans, sor_inv_sle_dx, sor_inv_sle_sn/
]
qed-.
-lemma lreq_frees_trans: ∀L1,T,f. L1 ⊢ 𝐅*⦃T⦄ ≡ f → ∀L2. L2 ≡[f] L1 → L2 ⊢ 𝐅*⦃T⦄ ≡ f.
+lemma lreq_frees_trans: ∀f,L1,T. L1 ⊢ 𝐅*⦃T⦄ ≡ f → ∀L2. L2 ≡[f] L1 → L2 ⊢ 𝐅*⦃T⦄ ≡ f.
/3 width=3 by frees_lreq_conf, lreq_sym/ qed-.
/3 width=8 by ex3_5_intro, or_introl, or_intror, conj/
qed-.
-lemma lfeq_inv_bind: ∀I,L1,L2,V,T,p. L1 ≡[ⓑ{p,I}V.T] L2 →
+lemma lfeq_inv_bind: ∀p,I,L1,L2,V,T. L1 ≡[ⓑ{p,I}V.T] L2 →
L1 ≡[V] L2 ∧ L1.ⓑ{I}V ≡[T] L2.ⓑ{I}V.
-#I #L1 #L2 #V #T #p #H elim (lfxs_inv_bind … H) -H /2 width=3 by conj/
+#p #I #L1 #L2 #V #T #H elim (lfxs_inv_bind … H) -H /2 width=3 by conj/
qed-.
lemma lfeq_inv_flat: ∀I,L1,L2,V,T. L1 ≡[ⓕ{I}V.T] L2 →
(* Basic forward lemmas *****************************************************)
-lemma lfeq_fwd_bind_sn: ∀I,L1,L2,V,T,p. L1 ≡[ⓑ{p,I}V.T] L2 → L1 ≡[V] L2.
+lemma lfeq_fwd_bind_sn: ∀p,I,L1,L2,V,T. L1 ≡[ⓑ{p,I}V.T] L2 → L1 ≡[V] L2.
/2 width=4 by lfxs_fwd_bind_sn/ qed-.
-lemma lfeq_fwd_bind_dx: ∀I,L1,L2,V,T,p.
+lemma lfeq_fwd_bind_dx: ∀p,I,L1,L2,V,T.
L1 ≡[ⓑ{p,I}V.T] L2 → L1.ⓑ{I}V ≡[T] L2.ⓑ{I}V.
/2 width=2 by lfxs_fwd_bind_dx/ qed-.
(* Main properties **********************************************************)
-theorem lfeq_bind: ∀I,L1,L2,V1,V2,T,p.
+theorem lfeq_bind: ∀p,I,L1,L2,V1,V2,T.
L1 ≡[V1] L2 → L1.ⓑ{I}V1 ≡[T] L2.ⓑ{I}V2 →
L1 ≡[ⓑ{p,I}V1.T] L2.
/2 width=2 by lfxs_bind/ qed.
(* Properties with ranged equivalence for local environments ****************)
-lemma lreq_lfeq: ∀L1,L2,T,f. L1 ⊢ 𝐅*⦃T⦄ ≡ f → L1 ≡[f] L2 → L1 ≡[T] L2.
+lemma lreq_lfeq: ∀f,L1,L2,T. L1 ⊢ 𝐅*⦃T⦄ ≡ f → L1 ≡[f] L2 → L1 ≡[T] L2.
/2 width=3 by ex2_intro/ qed.
(* Advanced properties ******************************************************)
]
qed-.
-lemma lfxs_inv_bind: ∀R,I,L1,L2,V1,V2,T,p. L1 ⦻*[R, ⓑ{p,I}V1.T] L2 → R L1 V1 V2 →
+lemma lfxs_inv_bind: ∀R,p,I,L1,L2,V1,V2,T. L1 ⦻*[R, ⓑ{p,I}V1.T] L2 → R L1 V1 V2 →
L1 ⦻*[R, V1] L2 ∧ L1.ⓑ{I}V1 ⦻*[R, T] L2.ⓑ{I}V2.
-#R #I #L1 #L2 #V1 #V2 #T #p * #f #Hf #HL #HV elim (frees_inv_bind … Hf) -Hf
+#R #p #I #L1 #L2 #V1 #V2 #T * #f #Hf #HL #HV elim (frees_inv_bind … Hf) -Hf
/6 width=6 by sle_lexs_trans, lexs_inv_tl, sor_inv_sle_dx, sor_inv_sle_sn, ex2_intro, conj/
qed-.
(* Basic forward lemmas *****************************************************)
-lemma lfxs_fwd_bind_sn: ∀R,I,L1,L2,V,T,p. L1 ⦻*[R, ⓑ{p,I}V.T] L2 → L1 ⦻*[R, V] L2.
-#R #I #L1 #L2 #V #T #p * #f #Hf #HL elim (frees_inv_bind … Hf) -Hf
+lemma lfxs_fwd_bind_sn: ∀R,p,I,L1,L2,V,T. L1 ⦻*[R, ⓑ{p,I}V.T] L2 → L1 ⦻*[R, V] L2.
+#R #p #I #L1 #L2 #V #T * #f #Hf #HL elim (frees_inv_bind … Hf) -Hf
/4 width=6 by sle_lexs_trans, sor_inv_sle_sn, ex2_intro/
qed-.
-lemma lfxs_fwd_bind_dx: ∀R,I,L1,L2,V1,V2,T,p. L1 ⦻*[R, ⓑ{p,I}V1.T] L2 →
+lemma lfxs_fwd_bind_dx: ∀R,p,I,L1,L2,V1,V2,T. L1 ⦻*[R, ⓑ{p,I}V1.T] L2 →
R L1 V1 V2 → L1.ⓑ{I}V1 ⦻*[R, T] L2.ⓑ{I}V2.
-#R #I #L1 #L2 #V1 #V2 #T #p #H #HV elim (lfxs_inv_bind … H HV) -H -HV //
+#R #p #I #L1 #L2 #V1 #V2 #T #H #HV elim (lfxs_inv_bind … H HV) -H -HV //
qed-.
lemma lfxs_fwd_flat_sn: ∀R,I,L1,L2,V,T. L1 ⦻*[R, ⓕ{I}V.T] L2 → L1 ⦻*[R, V] L2.
(* Main properties ******************************************************)
-theorem lfxs_bind: ∀R,I,L1,L2,V1,V2,T,p.
+theorem lfxs_bind: ∀R,p,I,L1,L2,V1,V2,T.
L1 ⦻*[R, V1] L2 → L1.ⓑ{I}V1 ⦻*[R, T] L2.ⓑ{I}V2 →
L1 ⦻*[R, ⓑ{p,I}V1.T] L2.
-#R #I #L1 #L2 #V1 #V2 #T #p * #f1 #HV #Hf1 * #f2 #HT #Hf2
+#R #p #I #L1 #L2 #V1 #V2 #T * #f1 #HV #Hf1 * #f2 #HT #Hf2
elim (lexs_fwd_pair … Hf2) -Hf2 #Hf2 #_ elim (sor_isfin_ex f1 (⫱f2))
/3 width=7 by frees_fwd_isfin, frees_bind, lexs_join, isfin_tl, ex2_intro/
qed.
(* Note: the premise 𝐔⦃f⦄ cannot be removed *)
(* Basic_2A1: includes: lsuba_drop_O1_conf *)
lemma lsuba_drops_conf_isuni: ∀G,L1,L2. G ⊢ L1 ⫃⁝ L2 →
- ∀K1,c,f. 𝐔⦃f⦄ → ⬇*[c, f] L1 ≡ K1 →
- ∃∃K2. G ⊢ K1 ⫃⁝ K2 & ⬇*[c, f] L2 ≡ K2.
+ ∀b,f,K1. 𝐔⦃f⦄ → ⬇*[b, f] L1 ≡ K1 →
+ ∃∃K2. G ⊢ K1 ⫃⁝ K2 & ⬇*[b, f] L2 ≡ K2.
#G #L1 #L2 #H elim H -L1 -L2
[ /2 width=3 by ex2_intro/
-| #I #L1 #L2 #V #HL12 #IH #K1 #c #f #Hf #H
+| #I #L1 #L2 #V #HL12 #IH #b #f #K1 #Hf #H
elim (drops_inv_pair1_isuni … Hf H) -Hf -H *
[ #Hf #H destruct -IH
/3 width=3 by lsuba_pair, drops_refl, ex2_intro/
| #g #Hg #HLK1 #H destruct -HL12
elim (IH … Hg HLK1) -L1 -Hg /3 width=3 by drops_drop, ex2_intro/
]
-| #L1 #L2 #W #V #A #HV #HW #HL12 #IH #K1 #c #f #Hf #H
+| #L1 #L2 #W #V #A #HV #HW #HL12 #IH #b #f #K1 #Hf #H
elim (drops_inv_pair1_isuni … Hf H) -Hf -H *
[ #Hf #H destruct -IH
/3 width=3 by drops_refl, lsuba_beta, ex2_intro/
(* Note: the premise 𝐔⦃f⦄ cannot be removed *)
(* Basic_2A1: includes: lsuba_drop_O1_trans *)
lemma lsuba_drop_O1_trans: ∀G,L1,L2. G ⊢ L1 ⫃⁝ L2 →
- ∀K2,c,f. 𝐔⦃f⦄ → ⬇*[c, f] L2 ≡ K2 →
- ∃∃K1. G ⊢ K1 ⫃⁝ K2 & ⬇*[c, f] L1 ≡ K1.
+ ∀b,f,K2. 𝐔⦃f⦄ → ⬇*[b, f] L2 ≡ K2 →
+ ∃∃K1. G ⊢ K1 ⫃⁝ K2 & ⬇*[b, f] L1 ≡ K1.
#G #L1 #L2 #H elim H -L1 -L2
[ /2 width=3 by ex2_intro/
-| #I #L1 #L2 #V #HL12 #IH #K2 #c #f #Hf #H
+| #I #L1 #L2 #V #HL12 #IH #b #f #K2 #Hf #H
elim (drops_inv_pair1_isuni … Hf H) -Hf -H *
[ #Hf #H destruct -IH
/3 width=3 by lsuba_pair, drops_refl, ex2_intro/
| #g #Hg #HLK2 #H destruct -HL12
elim (IH … Hg HLK2) -L2 -Hg /3 width=3 by drops_drop, ex2_intro/
]
-| #L1 #L2 #W #V #A #HV #HW #HL12 #IH #K2 #c #f #Hf #H
+| #L1 #L2 #W #V #A #HV #HW #HL12 #IH #b #f #K2 #Hf #H
elim (drops_inv_pair1_isuni … Hf H) -Hf -H *
[ #Hf #H destruct -IH
/3 width=3 by drops_refl, lsuba_beta, ex2_intro/
(* Basic_2A1: includes: lsubr_fwd_drop2_pair *)
lemma lsubr_fwd_drops2_pair: ∀L1,L2. L1 ⫃ L2 →
- ∀I,K2,W,c,f. 𝐔⦃f⦄ → ⬇*[c, f] L2 ≡ K2.ⓑ{I}W →
- (∃∃K1. K1 ⫃ K2 & ⬇*[c, f] L1 ≡ K1.ⓑ{I}W) ∨
- ∃∃K1,V. K1 ⫃ K2 & ⬇*[c, f] L1 ≡ K1.ⓓⓝW.V & I = Abst.
+ ∀b,f,I,K2,W. 𝐔⦃f⦄ → ⬇*[b, f] L2 ≡ K2.ⓑ{I}W →
+ (∃∃K1. K1 ⫃ K2 & ⬇*[b, f] L1 ≡ K1.ⓑ{I}W) ∨
+ ∃∃K1,V. K1 ⫃ K2 & ⬇*[b, f] L1 ≡ K1.ⓓⓝW.V & I = Abst.
#L1 #L2 #H elim H -L1 -L2
-[ #L #I #K2 #W #c #f #_ #H
+[ #L #b #f #I #K2 #W #_ #H
elim (drops_inv_atom1 … H) -H #H destruct
-| #J #L1 #L2 #V #HL12 #IH #I #K2 #W #c #f #Hf #H
+| #J #L1 #L2 #V #HL12 #IH #b #f #I #K2 #W #Hf #H
elim (drops_inv_pair1_isuni … Hf H) -Hf -H *
[ #Hf #H destruct -IH
/4 width=4 by drops_refl, ex2_intro, or_introl/
elim (IH … Hg HLK2) -IH -Hg -HLK2 *
/4 width=4 by drops_drop, ex3_2_intro, ex2_intro, or_introl, or_intror/
]
-| #L1 #L2 #V1 #V2 #HL12 #IH #I #K2 #W #c #f #Hf #H
+| #L1 #L2 #V1 #V2 #HL12 #IH #b #f #I #K2 #W #Hf #H
elim (drops_inv_pair1_isuni … Hf H) -Hf -H *
[ #Hf #H destruct -IH
/4 width=4 by drops_refl, ex3_2_intro, or_intror/
(* Basic_2A1: includes: lsubr_fwd_drop2_abbr *)
lemma lsubr_fwd_drops2_abbr: ∀L1,L2. L1 ⫃ L2 →
- ∀K2,V,c,f. 𝐔⦃f⦄ → ⬇*[c, f] L2 ≡ K2.ⓓV →
- ∃∃K1. K1 ⫃ K2 & ⬇*[c, f] L1 ≡ K1.ⓓV.
-#L1 #L2 #HL12 #K2 #V #c #f #Hf #HLK2
+ ∀b,f,K2,V. 𝐔⦃f⦄ → ⬇*[b, f] L2 ≡ K2.ⓓV →
+ ∃∃K1. K1 ⫃ K2 & ⬇*[b, f] L1 ≡ K1.ⓓV.
+#L1 #L2 #HL12 #b #f #K2 #V #Hf #HLK2
elim (lsubr_fwd_drops2_pair … HL12 … Hf HLK2) -L2 -Hf // *
#K1 #W #_ #_ #H destruct
qed-.