}
]
*)
+(*
class "orange"
[ { "functional" * } {
[ { "reduction and type machine" * } {
]
}
]
+*)
}
class "top" { * }
--- /dev/null
+(* A Basic_A2 lemma we do not need so far *)
+lemma lsubf_sle_div: ∀f,f2,L1,L2. ⦃L1, f⦄ ⫃ 𝐅* ⦃L2, f2⦄ →
+ ∀f1. f1 ⊆ f2 → ⦃L1, f⦄ ⫃ 𝐅* ⦃L2, f1⦄.
+#f #f2 #L1 #L2 #H elim H -f -f2 -L1 -L2
+/4 width=3 by lsubf_beta, lsubf_bind, lsubf_atom, sle_tl, sle_trans/
+qed-.
include "basic_2/notation/relations/rdropstar_3.ma".
include "basic_2/notation/relations/rdropstar_4.ma".
include "basic_2/relocation/lreq.ma".
-include "basic_2/relocation/lifts.ma".
+include "basic_2/relocation/lifts_bind.ma".
(* GENERIC SLICING FOR LOCAL ENVIRONMENTS ***********************************)
*)
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 (â\86\91f) (L1.â\93\91{I}V1) (L2.â\93\91{I}V2)
+| drops_drop: ∀f,I,L1,L2. drops b f L1 L2 → drops b (⫯f) (L1.ⓘ{I}) L2
+| drops_skip: ∀f,I1,I2,L1,L2.
+ drops b f L1 L2 → ⬆*[f] I2 ≡ I1 →
+ drops b (â\86\91f) (L1.â\93\98{I1}) (L2.â\93\98{I2})
.
interpretation "uniform slicing (local environment)"
λR. ∀L,U. R L U → ∀b,f,K. ⬇*[b, f] L ≡ K →
∀T. ⬆*[f] T ≡ U → R K T.
-definition d_liftable2_sn: predicate (lenv → relation term) ≝
- λ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 → ∀b,f,K. ⬇*[b, f] L ≡ K →
- ∀T1. ⬆*[f] T1 ≡ U1 →
- ∃∃T2. ⬆*[f] T2 ≡ U2 & R K T1 T2.
-
-definition d_liftable2_bi: predicate (lenv → relation term) ≝
- λ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_bi: predicate (lenv → relation term) ≝
- λ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 d_liftable2_sn: ∀C:Type[0]. ∀S:rtmap → relation C.
+ predicate (lenv → relation C) ≝
+ λC,S,R. ∀K,T1,T2. R K T1 T2 → ∀b,f,L. ⬇*[b, f] L ≡ K →
+ ∀U1. S f T1 U1 →
+ ∃∃U2. S f T2 U2 & R L U1 U2.
+
+definition d_deliftable2_sn: ∀C:Type[0]. ∀S:rtmap → relation C.
+ predicate (lenv → relation C) ≝
+ λC,S,R. ∀L,U1,U2. R L U1 U2 → ∀b,f,K. ⬇*[b, f] L ≡ K →
+ ∀T1. S f T1 U1 →
+ ∃∃T2. S f T2 U2 & R K T1 T2.
+
+definition d_liftable2_bi: ∀C:Type[0]. ∀S:rtmap → relation C.
+ predicate (lenv → relation C) ≝
+ λC,S,R. ∀K,T1,T2. R K T1 T2 → ∀b,f,L. ⬇*[b, f] L ≡ K →
+ ∀U1. S f T1 U1 →
+ ∀U2. S f T2 U2 → R L U1 U2.
+
+definition d_deliftable2_bi: ∀C:Type[0]. ∀S:rtmap → relation C.
+ predicate (lenv → relation C) ≝
+ λC,S,R. ∀L,U1,U2. R L U1 U2 → ∀b,f,K. ⬇*[b, f] L ≡ K →
+ ∀T1. S f T1 U1 →
+ ∀T2. S f T2 U2 → R K T1 T2.
definition co_dropable_sn: predicate (rtmap → relation lenv) ≝
λR. ∀b,f,L1,K1. ⬇*[b, f] L1 ≡ K1 → 𝐔⦃f⦄ →
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/
-| #f1 #I #L1 #L2 #V #_ #IH #f2 #H elim (eq_inv_nx … H) -H
+| #f1 #I #L1 #L2 #_ #IH #f2 #H elim (eq_inv_nx … H) -H
/3 width=3 by drops_drop/
-| #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/
+| #f1 #I1 #I2 #L1 #L2 #_ #HI #IH #f2 #H elim (eq_inv_px … H) -H
+ /3 width=3 by drops_skip, liftsb_eq_repl_back/
]
qed-.
* /2 width=1 by drops_TF/
qed-.
+lemma d_liftable2_sn_bi: ∀C,S. (∀f,c. is_mono … (S f c)) →
+ ∀R. d_liftable2_sn C S R → d_liftable2_bi C S R.
+#C #S #HS #R #HR #K #T1 #T2 #HT12 #b #f #L #HLK #U1 #HTU1 #U2 #HTU2
+elim (HR … HT12 … HLK … HTU1) -HR -K -T1 #X #HTX #HUX
+<(HS … HTX … HTU2) -T2 -U2 -b -f //
+qed-.
+
+lemma d_deliftable2_sn_bi: ∀C,S. (∀f. is_inj2 … (S f)) →
+ ∀R. d_deliftable2_sn C S R → d_deliftable2_bi C S R.
+#C #S #HS #R #HR #L #U1 #U2 #HU12 #b #f #K #HLK #T1 #HTU1 #T2 #HTU2
+elim (HR … HU12 … HLK … HTU1) -HR -L -U1 #X #HUX #HTX
+<(HS … HUX … HTU2) -U2 -T2 -b -f //
+qed-.
+
(* Basic inversion lemmas ***************************************************)
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/
-| #f #I #L1 #L2 #V #_ #H destruct
-| #f #I #L1 #L2 #V1 #V2 #_ #_ #H destruct
+| #f #I #L1 #L2 #_ #H destruct
+| #f #I1 #I2 #L1 #L2 #_ #_ #H destruct
]
qed-.
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: ∀b,f,X,Y. ⬇*[b, f] X ≡ Y → ∀g,I,K,V. X = K.ⓑ{I}V → f = ⫯g →
+fact drops_inv_drop1_aux: ∀b,f,X,Y. ⬇*[b, f] X ≡ Y → ∀g,I,K. X = K.ⓘ{I} → 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)
+[ #f #Hf #g #J #K #H destruct
+| #f #I #L1 #L2 #HL #g #J #K #H1 #H2 <(injective_next … H2) -g destruct //
+| #f #I1 #I2 #L1 #L2 #_ #_ #g #J #K #_ #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: ∀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-.
+lemma drops_inv_drop1: ∀b,f,I,K,Y. ⬇*[b, ⫯f] K.ⓘ{I} ≡ Y → ⬇*[b, f] K ≡ Y.
+/2 width=6 by drops_inv_drop1_aux/ qed-.
-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.
+fact drops_inv_skip1_aux: ∀b,f,X,Y. ⬇*[b, f] X ≡ Y → ∀g,I1,K1. X = K1.ⓘ{I1} → f = ↑g →
+ ∃∃I2,K2. ⬇*[b, g] K1 ≡ K2 & ⬆*[g] I2 ≡ I1 & Y = K2.ⓘ{I2}.
#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
+[ #f #Hf #g #J1 #K1 #H destruct
+| #f #I #L1 #L2 #_ #g #J1 #K1 #_ #H2 elim (discr_next_push … H2)
+| #f #I1 #I2 #L1 #L2 #HL #HI #g #J1 #K1 #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: ∀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.
+lemma drops_inv_skip1: ∀b,f,I1,K1,Y. ⬇*[b, ↑f] K1.ⓘ{I1} ≡ Y →
+ ∃∃I2,K2. ⬇*[b, f] K1 ≡ K2 & ⬆*[f] I2 ≡ I1 & Y = K2.ⓘ{I2}.
/2 width=5 by drops_inv_skip1_aux/ qed-.
-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.
+fact drops_inv_skip2_aux: ∀b,f,X,Y. ⬇*[b, f] X ≡ Y → ∀g,I2,K2. Y = K2.ⓘ{I2} → f = ↑g →
+ ∃∃I1,K1. ⬇*[b, g] K1 ≡ K2 & ⬆*[g] I2 ≡ I1 & X = K1.ⓘ{I1}.
#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
+[ #f #Hf #g #J2 #K2 #H destruct
+| #f #I #L1 #L2 #_ #g #J2 #K2 #_ #H2 elim (discr_next_push … H2)
+| #f #I1 #I2 #L1 #L2 #HL #HV #g #J2 #K2 #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: ∀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.
+lemma drops_inv_skip2: ∀b,f,I2,X,K2. ⬇*[b, ↑f] X ≡ K2.ⓘ{I2} →
+ ∃∃I1,K1. ⬇*[b, f] K1 ≡ K2 & ⬆*[f] I2 ≡ I1 & X = K1.ⓘ{I1}.
/2 width=5 by drops_inv_skip2_aux/ qed-.
(* Basic forward lemmas *****************************************************)
-fact drops_fwd_drop2_aux: ∀b,f2,X,Y. ⬇*[b, f2] X ≡ Y → ∀I,K,V. Y = K.ⓑ{I}V →
+fact drops_fwd_drop2_aux: ∀b,f2,X,Y. ⬇*[b, f2] X ≡ Y → ∀I,K. Y = K.ⓘ{I} →
∃∃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
+[ #f2 #Hf2 #J #K #H destruct
+| #f2 #I #L1 #L2 #_ #IHL #J #K #H elim (IHL … H) -IHL
/3 width=7 by after_next, ex3_2_intro, drops_drop/
-| #f2 #I #L1 #L2 #V1 #V2 #HL #_ #_ #J #K #W #H destruct
+| #f2 #I1 #I2 #L1 #L2 #HL #_ #_ #J #K #H destruct
lapply (after_isid_dx 𝐈𝐝 … f2) /3 width=9 by after_push, ex3_2_intro, drops_drop/
]
qed-.
-lemma drops_fwd_drop2: ∀b,f2,I,X,K,V. ⬇*[b, f2] X ≡ K.ⓑ{I}V →
+lemma drops_fwd_drop2: ∀b,f2,I,X,K. ⬇*[b, f2] X ≡ K.ⓘ{I} →
∃∃f1,f. 𝐈⦃f1⦄ & f2 ⊚ ⫯f1 ≡ f & ⬇*[b, f] X ≡ K.
-/2 width=5 by drops_fwd_drop2_aux/ qed-.
+/2 width=4 by drops_fwd_drop2_aux/ qed-.
(* Properties with test for identity ****************************************)
(* Basic_2A1: includes: drop_refl *)
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/
+#L #I #IHL #f #Hf elim (isid_inv_gen … Hf) -Hf
+/3 width=1 by drops_skip, liftsb_refl/
qed.
(* Forward lemmas test for identity *****************************************)
(* Basic_2A1: includes: drop_inv_O2 *)
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/
+[ #f #I #L1 #L2 #_ #_ #H elim (isid_inv_next … H) //
+| /5 width=5 by isid_inv_push, liftsb_fwd_isid, eq_f2, sym_eq/
]
qed-.
-lemma drops_after_fwd_drop2: ∀b,f2,I,X,K,V. ⬇*[b, f2] X ≡ K.ⓑ{I}V →
+lemma drops_after_fwd_drop2: ∀b,f2,I,X,K. ⬇*[b, f2] X ≡ K.ⓘ{I} →
∀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
+#b #f2 #I #X #K #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-.
lemma drops_isuni_ex: ∀f. 𝐔⦃f⦄ → ∀L. ∃K. ⬇*[Ⓕ, f] L ≡ K.
#f #H elim H -f /4 width=2 by drops_refl, drops_TF, ex_intro/
-#f #_ #g #H #IH * /2 width=2 by ex_intro/
-#L #I #V destruct
-elim (IH L) -IH /3 width=2 by drops_drop, ex_intro/
+#f #_ #g #H #IH destruct * /2 width=2 by ex_intro/
+#L #I elim (IH L) -IH /3 width=2 by drops_drop, ex_intro/
qed-.
(* Inversion lemmas with test for uniformity ********************************)
lemma drops_inv_isuni: ∀f,L1,L2. ⬇*[Ⓣ, f] L1 ≡ L2 → 𝐔⦃f⦄ →
(𝐈⦃f⦄ ∧ L1 = L2) ∨
- ∃∃g,I,K,V. ⬇*[Ⓣ, g] K ≡ L2 & 𝐔⦃g⦄ & L1 = K.ⓑ{I}V & f = ⫯g.
+ ∃∃g,I,K. ⬇*[Ⓣ, g] K ≡ L2 & 𝐔⦃g⦄ & L1 = K.ⓘ{I} & 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/
+| /4 width=7 by isuni_inv_next, ex4_3_intro, or_intror/
+| /7 width=6 by drops_fwd_isid, liftsb_fwd_isid, isuni_inv_push, isid_push, or_introl, conj, eq_f2, sym_eq/
]
qed-.
(* Basic_2A1: was: drop_inv_O1_pair1 *)
-lemma drops_inv_pair1_isuni: ∀b,f,I,K,L2,V. 𝐔⦃f⦄ → ⬇*[b, f] K.ⓑ{I}V ≡ L2 →
- (ð\9d\90\88â¦\83fâ¦\84 â\88§ L2 = K.â\93\91{I}V) ∨
+lemma drops_inv_bind1_isuni: ∀b,f,I,K,L2. 𝐔⦃f⦄ → ⬇*[b, f] K.ⓘ{I} ≡ L2 →
+ (ð\9d\90\88â¦\83fâ¦\84 â\88§ L2 = K.â\93\98{I}) ∨
∃∃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
+#b #f #I #K #L2 #Hf #H elim (isuni_split … Hf) -Hf * #g #Hg #H0 destruct
+[ lapply (drops_inv_skip1 … H) -H * #Z #Y #HY #HZ #H destruct
+ <(drops_fwd_isid … HY Hg) -Y >(liftsb_fwd_isid … HZ Hg) -Z
/4 width=3 by isid_push, or_introl, conj/
| lapply (drops_inv_drop1 … H) -H /3 width=4 by ex3_intro, or_intror/
]
qed-.
(* Basic_2A1: was: drop_inv_O1_pair2 *)
-lemma drops_inv_pair2_isuni: ∀b,f,I,K,V,L1. 𝐔⦃f⦄ → ⬇*[b, f] L1 ≡ K.ⓑ{I}V →
- (ð\9d\90\88â¦\83fâ¦\84 â\88§ L1 = K.â\93\91{I}V) ∨
- ∃∃g,I1,K1,V1. 𝐔⦃g⦄ & ⬇*[b, g] K1 ≡ K.ⓑ{I}V & L1 = K1.ⓑ{I1}V1 & f = ⫯g.
-#b #f #I #K #V *
+lemma drops_inv_bind2_isuni: ∀b,f,I,K,L1. 𝐔⦃f⦄ → ⬇*[b, f] L1 ≡ K.ⓘ{I} →
+ (ð\9d\90\88â¦\83fâ¦\84 â\88§ L1 = K.â\93\98{I}) ∨
+ ∃∃g,I1,K1. 𝐔⦃g⦄ & ⬇*[b, g] K1 ≡ K.ⓘ{I} & L1 = K1.ⓘ{I1} & f = ⫯g.
+#b #f #I #K *
[ #Hf #H elim (drops_inv_atom1 … H) -H #H destruct
-| #L1 #I1 #V1 #Hf #H elim (drops_inv_pair1_isuni … Hf H) -Hf -H *
+| #L1 #I1 #Hf #H elim (drops_inv_bind1_isuni … Hf H) -Hf -H *
[ #Hf #H destruct /3 width=1 by or_introl, conj/
- | /3 width=8 by ex4_4_intro, or_intror/
+ | /3 width=7 by ex4_3_intro, or_intror/
]
]
qed-.
-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 *
+lemma drops_inv_bind2_isuni_next: ∀b,f,I,K,L1. 𝐔⦃f⦄ → ⬇*[b, ⫯f] L1 ≡ K.ⓘ{I} →
+ ∃∃I1,K1. ⬇*[b, f] K1 ≡ K.ⓘ{I} & L1 = K1.ⓘ{I1}.
+#b #f #I #K #L1 #Hf #H elim (drops_inv_bind2_isuni … H) -H /2 width=3 by isuni_next/ -Hf *
[ #H elim (isid_inv_next … H) -H //
-| /2 width=5 by ex2_3_intro/
+| /2 width=4 by ex2_2_intro/
]
qed-.
fact drops_inv_TF_aux: ∀f,L1,L2. ⬇*[Ⓕ, f] L1 ≡ L2 → 𝐔⦃f⦄ →
- ∀I,K,V. L2 = K.ⓑ{I}V →
- ⬇*[Ⓣ, f] L1 ≡ K.ⓑ{I}V.
+ ∀I,K. L2 = K.ⓘ{I} → ⬇*[Ⓣ, f] L1 ≡ K.ⓘ{I}.
#f #L1 #L2 #H elim H -f -L1 -L2
-[ #f #_ #_ #J #K #W #H destruct
-| #f #I #L1 #L2 #V #_ #IH #Hf #J #K #W #H destruct
+[ #f #_ #_ #J #K #H destruct
+| #f #I #L1 #L2 #_ #IH #Hf #J #K #H destruct
/4 width=3 by drops_drop, isuni_inv_next/
-| #f #I #L1 #L2 #V1 #V2 #HL12 #HV21 #_ #Hf #J #K #W #H destruct
+| #f #I1 #I2 #L1 #L2 #HL12 #HI21 #_ #Hf #J #K #H destruct
lapply (isuni_inv_push … Hf ??) -Hf [1,2: // ] #Hf
- <(drops_fwd_isid … HL12) -K // <(lifts_fwd_isid … HV21) -V1
+ <(drops_fwd_isid … HL12) -K // <(liftsb_fwd_isid … HI21) -I1
/3 width=3 by drops_refl, isid_push/
]
qed-.
(* Basic_2A1: includes: drop_inv_FT *)
-lemma drops_inv_TF: ∀f,I,L,K,V. ⬇*[Ⓕ, f] L ≡ K.ⓑ{I}V → 𝐔⦃f⦄ →
- ⬇*[Ⓣ, f] L ≡ K.ⓑ{I}V.
+lemma drops_inv_TF: ∀f,I,L,K. ⬇*[Ⓕ, f] L ≡ K.ⓘ{I} → 𝐔⦃f⦄ → ⬇*[Ⓣ, f] L ≡ K.ⓘ{I}.
/2 width=3 by drops_inv_TF_aux/ qed-.
(* Basic_2A1: includes: drop_inv_gen *)
-lemma drops_inv_gen: ∀b,f,I,L,K,V. ⬇*[b, f] L ≡ K.ⓑ{I}V → 𝐔⦃f⦄ →
- ⬇*[Ⓣ, f] L ≡ K.ⓑ{I}V.
+lemma drops_inv_gen: ∀b,f,I,L,K. ⬇*[b, f] L ≡ K.ⓘ{I} → 𝐔⦃f⦄ → ⬇*[Ⓣ, f] L ≡ K.ⓘ{I}.
* /2 width=1 by drops_inv_TF/
qed-.
(* Basic_2A1: includes: drop_inv_T *)
-lemma drops_inv_F: ∀b,f,I,L,K,V. ⬇*[Ⓕ, f] L ≡ K.ⓑ{I}V → 𝐔⦃f⦄ →
- ⬇*[b, f] L ≡ K.ⓑ{I}V.
+lemma drops_inv_F: ∀b,f,I,L,K. ⬇*[Ⓕ, f] L ≡ K.ⓘ{I} → 𝐔⦃f⦄ → ⬇*[b, f] L ≡ K.ⓘ{I}.
* /2 width=1 by drops_inv_TF/
qed-.
(* Basic_1: was: drop_S *)
(* Basic_2A1: was: drop_fwd_drop2 *)
-lemma drops_isuni_fwd_drop2: ∀b,f,I,X,K,V. 𝐔⦃f⦄ → ⬇*[b, f] X ≡ K.ⓑ{I}V → ⬇*[b, ⫯f] X ≡ K.
+lemma drops_isuni_fwd_drop2: ∀b,f,I,X,K. 𝐔⦃f⦄ → ⬇*[b, f] X ≡ K.ⓘ{I} → ⬇*[b, ⫯f] X ≡ K.
/3 width=7 by drops_after_fwd_drop2, after_isid_isuni/ qed-.
(* Inversion lemmas with uniform relocations ********************************)
∃∃n,f1. ⬇*[b, 𝐔❴n❵] L ≡ ⋆ & 𝐔❴n❵ ⊚ f1 ≡ f.
#b #L elim L -L
[ /3 width=4 by drops_atom, after_isid_sn, ex2_2_intro/
-| #L #I #V #IH #f #H elim (pn_split f) * #g #H0 destruct
- [ elim (drops_inv_skip1 … H) -H #K #W #_ #_ #H destruct
+| #L #I #IH #f #H elim (pn_split f) * #g #H0 destruct
+ [ elim (drops_inv_skip1 … H) -H #J #K #_ #_ #H destruct
| lapply (drops_inv_drop1 … H) -H #HL
elim (IH … HL) -IH -HL /3 width=8 by drops_drop, after_next, ex2_2_intro/
]
]
qed-.
-lemma drops_inv_succ: ∀l,L1,L2. ⬇*[⫯l] L1 ≡ L2 →
- ∃∃I,K,V. ⬇*[l] K ≡ L2 & L1 = K.ⓑ{I}V.
-#l #L1 #L2 #H elim (drops_inv_isuni … H) -H // *
+lemma drops_inv_succ: ∀L1,L2,i. ⬇*[⫯i] L1 ≡ L2 →
+ ∃∃I,K. ⬇*[i] K ≡ L2 & L1 = K.ⓘ{I}.
+#L1 #L2 #i #H elim (drops_inv_isuni … H) -H // *
[ #H elim (isid_inv_next … H) -H //
-| /2 width=5 by ex2_3_intro/
+| /2 width=4 by ex2_2_intro/
]
qed-.
(* Properties with uniform relocations **************************************)
-lemma drops_F_uni: ∀L,i. ⬇*[Ⓕ, 𝐔❴i❵] L ≡ ⋆ ∨ ∃∃I,K,V. ⬇*[i] L ≡ K.ⓑ{I}V.
+lemma drops_F_uni: ∀L,i. ⬇*[Ⓕ, 𝐔❴i❵] L ≡ ⋆ ∨ ∃∃I,K. ⬇*[i] L ≡ K.ⓘ{I}.
#L elim L -L /2 width=1 by or_introl/
-#L #I #V #IH * /4 width=4 by drops_refl, ex1_3_intro, or_intror/
+#L #I #IH * /4 width=3 by drops_refl, ex1_2_intro, or_intror/
#i elim (IH i) -IH /3 width=1 by drops_drop, or_introl/
-* /4 width=4 by drops_drop, ex1_3_intro, or_intror/
-qed-.
+* /4 width=3 by drops_drop, ex1_2_intro, or_intror/
+qed-.
(* Basic_2A1: includes: drop_split *)
lemma drops_split_trans: ∀b,f,L1,L2. ⬇*[b, f] L1 ≡ L2 → ∀f1,f2. f1 ⊚ f2 ≡ f → 𝐔⦃f1⦄ →
[ #f #H0f #f1 #f2 #Hf #Hf1 @(ex2_intro … (⋆)) @drops_atom
#H lapply (H0f H) -b
#H elim (after_inv_isid3 … Hf H) -f //
-| #f #I #L1 #L2 #V #HL12 #IHL12 #f1 #f2 #Hf #Hf1 elim (after_inv_xxn … Hf) -Hf [1,3: * |*: // ]
+| #f #I #L1 #L2 #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
- /4 width=5 by drops_drop, drops_skip, lifts_refl, isuni_isid, ex2_intro/
+ /4 width=5 by drops_drop, drops_skip, liftsb_refl, isuni_isid, ex2_intro/
| #g1 #Hf #H destruct elim (IHL12 … Hf) -f
/3 width=5 by ex2_intro, drops_drop, isuni_inv_next/
]
-| #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
+| #f #I1 #I2 #L1 #L2 #_ #HI21 #IHL12 #f1 #f2 #Hf #Hf1 elim (after_inv_xxp … Hf) -Hf [2,3: // ]
+ #g1 #g2 #Hf #H1 #H2 destruct elim (liftsb_split_trans … HI21 … Hf) -HI21
elim (IHL12 … Hf) -f /3 width=5 by ex2_intro, drops_skip, isuni_fwd_push/
]
qed-.
∃∃L2. ⬇*[Ⓕ, f2] L ≡ L2 & ⬇*[Ⓕ, f] L1 ≡ L2.
#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: // ]
+| #f1 #I #L1 #L #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/
-| #f1 #I #L1 #L #V1 #V #HL1 #HV1 #IH #f2 #f #Hf #Hf2
+| #f1 #I1 #I #L1 #L #HL1 #HI1 #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
lapply (after_isid_inv_dx … Hg … Hg2) -Hg #Hg
- /5 width=7 by drops_eq_repl_back, drops_F, drops_refl, drops_skip, lifts_eq_repl_back, isid_push, ex2_intro/
- | lapply (isuni_inv_next … Hf2 ??) -Hf2 [1,2: // ] #Hg2 -HL1 -HV1
+ /5 width=7 by drops_eq_repl_back, drops_F, drops_refl, drops_skip, liftsb_eq_repl_back, isid_push, ex2_intro/
+ | lapply (isuni_inv_next … Hf2 ??) -Hf2 [1,2: // ] #Hg2 -HL1 -HI1
elim (IH … Hg) -f1 /3 width=3 by drops_drop, ex2_intro/
]
]
⬇*[b,↑⫱*[⫯i2]f] L1 ≡ L2.
/3 width=3 by drops_eq_repl_fwd, at_inv_tls/ qed-.
-lemma drops_split_trans_pair2: ∀b,f,I,L,K0,V. ⬇*[b, f] L ≡ K0.ⓑ{I}V → ∀n. @⦃O, f⦄ ≡ n →
- ∃∃K,W. ⬇*[n]L ≡ K.ⓑ{I}W & ⬇*[b, ⫱*[⫯n]f] K ≡ K0 & ⬆*[⫱*[⫯n]f] V ≡ W.
-#b #f #I #L #K0 #V #H #n #Hf
+lemma drops_split_trans_bind2: ∀b,f,I,L,K0. ⬇*[b, f] L ≡ K0.ⓘ{I} → ∀i. @⦃O, f⦄ ≡ i →
+ ∃∃J,K. ⬇*[i]L ≡ K.ⓘ{J} & ⬇*[b, ⫱*[⫯i]f] K ≡ K0 & ⬆*[⫱*[⫯i]f] I ≡ J.
+#b #f #I #L #K0 #H #i #Hf
elim (drops_split_trans … H) -H [ |5: @(after_uni_dx … Hf) |2,3: skip ] /2 width=1 by after_isid_dx/ #Y #HLY #H
lapply (drops_tls_at … Hf … H) -H #H
-elim (drops_inv_skip2 … H) -H #K #W #HK0 #HVW #H destruct
+elim (drops_inv_skip2 … H) -H #J #K #HK0 #HIJ #H destruct
/3 width=5 by drops_inv_gen, ex3_2_intro/
qed-.
+(* Properties with context-sensitive equivalence for terms ******************)
+
+lemma ceq_lift_sn: d_liftable2_sn … liftsb ceq.
+/2 width=3 by ex2_intro/ qed-.
+
+lemma ceq_inv_lift_sn: d_deliftable2_sn … liftsb ceq.
+/2 width=3 by ex2_intro/ qed-.
+
+(* Note: d_deliftable2_sn cfull does not hold *)
+lemma cfull_lift_sn: d_liftable2_sn … liftsb cfull.
+#K #I1 #I2 #_ #b #f #L #_ #J1 #_ -K -I1 -b
+elim (liftsb_total I2 f) /2 width=3 by ex2_intro/
+qed-.
+
(* Basic_2A1: removed theorems 12:
drops_inv_nil drops_inv_cons d1_liftable_liftables
drop_refl_atom_O2 drop_inv_pair1
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/relocation/drops.ma".
-
-(* GENERIC SLICING FOR LOCAL ENVIRONMENTS ***********************************)
-
-(* Properties with context-sensitive equivalence for terms ******************)
-
-lemma ceq_lift_sn: d_liftable2_sn ceq.
-/2 width=3 by ex2_intro/ qed-.
-
-lemma ceq_inv_lift_sn: d_deliftable2_sn ceq.
-/2 width=3 by ex2_intro/ qed-.
-
-(* Note: d_deliftable2_sn cfull does not hold *)
-lemma cfull_lift_sn: d_liftable2_sn cfull.
-#K #T1 #T2 #_ #b #f #L #_ #U1 #_ -K -T1 -b
-elim (lifts_total T2 f) /2 width=3 by ex2_intro/
-qed-.
(* *)
(**************************************************************************)
-include "basic_2/relocation/lifts_lifts.ma".
+include "basic_2/relocation/lifts_lifts_bind.ma".
include "basic_2/relocation/drops_weight.ma".
(* GENERIC SLICING FOR LOCAL ENVIRONMENTS ***********************************)
-(* Main properties on generic relocation ************************************)
-
-lemma d_liftable2_sn_bi: ∀R. d_liftable2_sn R → d_liftable2_bi R.
-#R #HR #K #T1 #T2 #HT12 #b #f #L #HLK #U1 #HTU1 #U2 #HTU2
-elim (HR … HT12 … HLK … HTU1) -HR -K -T1 #X #HTX #HUX
-<(lifts_mono … HTX … HTU2) -T2 -U2 -b -f //
-qed-.
-
-lemma d_deliftable2_sn_bi: ∀R. d_deliftable2_sn R → d_deliftable2_bi R.
-#R #HR #L #U1 #U2 #HU12 #b #f #K #HLK #T1 #HTU1 #T2 #HTU2
-elim (HR … HU12 … HLK … HTU1) -HR -L -U1 #X #HUX #HTX
-<(lifts_inj … HUX … HTU2) -U2 -T2 -b -f //
-qed-.
-
(* Main properties **********************************************************)
(* Basic_2A1: includes: drop_conf_ge drop_conf_be drop_conf_le *)
[ #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/
-| #f1 #I #K1 #K #V1 #_ #IH #b2 #f #L2 #HL2 #f2 #Hf elim (after_inv_nxx … Hf) -Hf [2,3: // ]
+| #f1 #I1 #K1 #K #_ #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/
-| #f1 #I #K1 #K #V1 #V #_ #HV1 #IH #b2 #f #L2 #HL2 #f2 #Hf elim (after_inv_pxx … Hf) -Hf [1,3: * |*:// ]
+| #f1 #I1 #I #K1 #K #_ #HI1 #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_div3/
+ [ elim (drops_inv_skip1 … HL2) -HL2 /3 width=6 by drops_skip, liftsb_div3/
| /4 width=3 by drops_inv_drop1, drops_drop/
]
]
#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/
-| #f1 #I #K1 #K #V1 #_ #IH #b2 #f2 #L2 #HL2 #f #Hf elim (after_inv_nxx … Hf) -Hf
+| #f1 #I1 #K1 #K #_ #IH #b2 #f2 #L2 #HL2 #f #Hf elim (after_inv_nxx … Hf) -Hf
/3 width=3 by drops_drop/
-| #f1 #I #K1 #K #V1 #V #_ #HV1 #IH #b2 #f2 #L2 #HL2 #f #Hf elim (after_inv_pxx … Hf) -Hf [1,3: * |*: // ]
+| #f1 #I1 #I #K1 #K #_ #HI1 #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/
+ [ elim (drops_inv_skip1 … HL2) -HL2 /3 width=6 by drops_skip, liftsb_trans/
| /4 width=3 by drops_inv_drop1, drops_drop/
]
]
#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/
-| #f1 #I #L #K #V #Hf1 #IH #f2 elim (pn_split f2) *
+| #f1 #I #L #K #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)
+ #H destruct elim (drops_inv_x_bind_xy … Hf1)
| /4 width=5 by drops_inv_drop1, isuni_inv_next, eq_next/
]
-| #f1 #I #L #K #V #W #Hf1 #_ #IH #f2 elim (pn_split f2) *
+| #f1 #I1 #I2 #L #K #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/
| lapply (drops_inv_drop1 … Hf2) -Hf2 -IH -HU2 #Hg2
lapply (drops_fwd_isid … Hf1 ?) -Hf1 /2 width=3 by isuni_inv_push/ -HU1
- #H destruct elim (drops_inv_x_pair_xy … Hg2)
+ #H destruct elim (drops_inv_x_bind_xy … Hg2)
]
]
qed-.
(* Basic_2A1: includes: drop_conf_lt *)
lemma drops_conf_skip1: ∀b2,f,L,L2. ⬇*[b2, f] L ≡ L2 →
- ∀b1,f1,I,K1,V1. ⬇*[b1, f1] L ≡ K1.ⓑ{I}V1 →
+ ∀b1,f1,I1,K1. ⬇*[b1, f1] L ≡ K1.ⓘ{I1} →
∀f2. f1 ⊚ ↑f2 ≡ f →
- ∃∃K2,V2. L2 = K2.ⓑ{I}V2 &
- ⬇*[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
+ ∃∃I2,K2. L2 = K2.ⓘ{I2} &
+ ⬇*[b2, f2] K1 ≡ K2 & ⬆*[f2] I2 ≡ I1.
+#b2 #f #L #L2 #H2 #b1 #f1 #I1 #K1 #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: ∀b1,f1,L1,L. ⬇*[b1, f1] L1 ≡ L →
- ∀b2,f2,I,K2,V2. ⬇*[b2, f2] L ≡ K2.ⓑ{I}V2 →
+ ∀b2,f2,I2,K2. ⬇*[b2, f2] L ≡ K2.ⓘ{I2} →
∀f. f1 ⊚ f2 ≡ ↑f →
- ∃∃K1,V1. L1 = K1.ⓑ{I}V1 &
- ⬇*[b1∧b2, f] K1 ≡ K2 & ⬆*[f] V2 ≡ V1.
-#b1 #f1 #L1 #L #H1 #b2 #f2 #I #K2 #V2 #H2 #f #Hf
+ ∃∃I1,K1. L1 = K1.ⓘ{I1} &
+ ⬇*[b1∧b2, f] K1 ≡ K2 & ⬆*[f] I2 ≡ I1.
+#b1 #f1 #L1 #L #H1 #b2 #f2 #I2 #K2 #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: ∀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.
-#f1 #f2 #I1 #I2 #L #K #V1 #V2 #Hf1 #Hf2 #HU1 #HU2
+lemma drops_conf_div_bind: ∀f1,f2,I1,I2,L,K.
+ ⬇*[Ⓣ, f1] L ≡ K.ⓘ{I1} → ⬇*[Ⓣ, f2] L ≡ K.ⓘ{I2} →
+ 𝐔⦃f1⦄ → 𝐔⦃f2⦄ → f1 ≗ f2 ∧ I1 = I2.
+#f1 #f2 #I1 #I2 #L #K #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
lapply (eq_inv_nn … H ????) -H [5: |*: // ] #H12
lapply (drops_eq_repl_back … Hf1 … H12) -Hf1 #H0
lapply (drops_mono … H0 … Hf2) -L #H
-destruct /2 width=1 by and3_intro/
+destruct /2 width=1 by conj/
qed-.
-lemma drops_inv_uni: ∀L,i. ⬇*[Ⓕ, 𝐔❴i❵] L ≡ ⋆ → ∀I,K,V. ⬇*[i] L ≡ K.ⓑ{I}V → ⊥.
-#L #i #H1 #I #K #V #H2
+lemma drops_inv_uni: ∀L,i. ⬇*[Ⓕ, 𝐔❴i❵] L ≡ ⋆ → ∀I,K. ⬇*[i] L ≡ K.ⓘ{I} → ⊥.
+#L #i #H1 #I #K #H2
lapply (drops_F … H2) -H2 #H2
lapply (drops_mono … H2 … H1) -L -i #H destruct
qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/syntax/lenv_ext2.ma".
+include "basic_2/relocation/drops.ma".
+
+(* GENERIC SLICING FOR LOCAL ENVIRONMENTS ***********************************)
+
+(* Properties with the extension to binders of a context-sensitive relation *)
+
+lemma cext2_d_liftable2_sn: ∀R. d_liftable2_sn … lifts R →
+ d_liftable2_sn … liftsb (cext2 R).
+#R #HR #K #I1 #I2 * -I1 -I2 #I [| #T1 #T2 #HT12 ]
+#b #f #L #HLK #Z1 #H
+[ lapply (liftsb_inv_unit_sn … H)
+| lapply (liftsb_inv_pair_sn … H) * #U1 #HTU1
+] -H #H destruct /3 width=3 by ext2_unit, ex2_intro/
+elim (HR … HT12 … HLK … HTU1) -HR -b -K -T1 /3 width=3 by ext2_pair, ex2_intro/
+qed-.
#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 //
-| #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
+| #f #I1 #L1 #K1 #_ #IH #X2 #K2 #HX #H elim (length_inv_succ_sn … H) -H
+ #I2 #L2 #H12 #H destruct lapply (drops_inv_drop1 … HX) -HX
#HLK2 @(IH … HLK2 H12) (**) (* auto fails *)
-| #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 >length_pair >length_pair /2 width=1 by/ (**) (* full auto fails *)
+| #f #I1 #I2 #L1 #K1 #_ #_ #IH #X2 #Y2 #HX #H elim (length_inv_succ_sn … H) -H
+ #I2 #L2 #H12 #H destruct elim (drops_inv_skip1 … HX) -HX
+ #I2 #K2 #HLK2 #_ #H destruct
+ lapply (IH … HLK2 H12) -f >length_bind >length_bind /2 width=1 by/ (**) (* full auto fails *)
]
qed-.
∃∃n. 𝐂⦃f⦄ ≡ n & |L1| = |L2| + n.
#f #L1 #L2 #H elim H -f -L1 -L2
[ /4 width=3 by fcla_isid, ex2_intro/
-| #f #I #L1 #L2 #V #_ * >length_pair /3 width=3 by fcla_next, ex2_intro, eq_f/
-| #f #I #L1 #L2 #V1 #V2 #_ #_ * >length_pair >length_pair /3 width=3 by fcla_push, ex2_intro/
+| #f #I #L1 #L2 #_ * >length_bind /3 width=3 by fcla_next, ex2_intro, eq_f/
+| #f #I1 #I2 #L1 #L2 #_ #_ * >length_bind >length_bind /3 width=3 by fcla_push, ex2_intro/
]
qed-.
#k #Hm #H <(fcla_mono … Hm … Hn) -f //
qed-.
-lemma drops_fwd_fcla_lt2: ∀f,L1,I2,K2,V2. ⬇*[Ⓣ, f] L1 ≡ K2.ⓑ{I2}V2 →
+lemma drops_fwd_fcla_lt2: ∀f,L1,I2,K2. ⬇*[Ⓣ, f] L1 ≡ K2.ⓘ{I2} →
∃∃n. 𝐂⦃f⦄ ≡ n & n < |L1|.
-#f #L1 #I2 #K2 #V2 #H elim (drops_fwd_fcla … H) -H
+#f #L1 #I2 #K2 #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: ∀f,L1,I2,K2,V2,n.
- â¬\87*[â\93\89, f] L1 â\89¡ K2.â\93\91{I2}V2 → 𝐂⦃f⦄ ≡ n →
+lemma drops_fcla_fwd_lt2: ∀f,L1,I2,K2,n.
+ â¬\87*[â\93\89, f] L1 â\89¡ K2.â\93\98{I2} → 𝐂⦃f⦄ ≡ n →
n < |L1|.
-#f #L1 #I2 #K2 #V2 #n #H #Hn elim (drops_fwd_fcla_lt2 … H) -H
+#f #L1 #I2 #K2 #n #H #Hn elim (drops_fwd_fcla_lt2 … H) -H
#k #Hm #H <(fcla_mono … Hm … Hn) -f //
qed-.
#RN #RP #b #f #L1 #K1 #H elim H -f -L1 -K1
[ #f #Hf #_ #f2 #X #H #f1 #Hf2 >(lexs_inv_atom1 … H) -X
/4 width=3 by lexs_atom, drops_atom, ex2_intro/
-| #f #I #L1 #K1 #V1 #_ #IH #Hf #f2 #X #H #f1 #Hf2
+| #f #I1 #L1 #K1 #_ #IH #Hf #f2 #X #H #f1 #Hf2
elim (coafter_inv_nxx … Hf2) -Hf2 [2,3: // ] #g2 #Hg2 #H2 destruct
- elim (lexs_inv_push1 … H) -H #L2 #V2 #HL12 #HV12 #H destruct
+ elim (lexs_inv_push1 … H) -H #I2 #L2 #HL12 #HI12 #H destruct
elim (IH … HL12 … Hg2) -g2
/3 width=3 by isuni_inv_next, drops_drop, ex2_intro/
-| #f #I #L1 #K1 #V1 #W1 #HLK #HWV #IH #Hf #f2 #X #H #f1 #Hf2
+| #f #I1 #J1 #L1 #K1 #HLK #HJI1 #IH #Hf #f2 #X #H #f1 #Hf2
lapply (isuni_inv_push … Hf ??) -Hf [3: |*: // ] #Hf
lapply (drops_fwd_isid … HLK … Hf) -HLK #H0 destruct
- lapply (lifts_fwd_isid … HWV … Hf) -HWV #H0 destruct
+ lapply (liftsb_fwd_isid … HJI1 … Hf) -HJI1 #H0 destruct
elim (coafter_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 (lexs_inv_push1 … H) | elim (lexs_inv_next1 … H) ] -H #I2 #L2 #HL12 #HI12 #H destruct
elim (IH … HL12 … Hg2) -g2 -IH /2 width=1 by isuni_isid/ #K2 #HK12 #HLK2
lapply (drops_fwd_isid … HLK2 … Hf) -HLK2 #H0 destruct
/4 width=3 by drops_refl, lexs_next, lexs_push, isid_push, ex2_intro/
qed-.
(* Basic_2A1: includes: lpx_sn_liftable_dedropable *)
-lemma lexs_liftable_co_dedropable_sn: ∀RN,RP. (∀L. reflexive ? (RN L)) → (∀L. reflexive ? (RP L)) →
- d_liftable2_sn RN → d_liftable2_sn RP → co_dedropable_sn (lexs RN RP).
+lemma lexs_liftable_co_dedropable_sn: ∀RN,RP. (∀L. reflexive … (RN L)) → (∀L. reflexive … (RP L)) →
+ d_liftable2_sn … liftsb RN → d_liftable2_sn … liftsb RP → co_dedropable_sn (lexs RN RP).
#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/
-| #f #I #L1 #K1 #V1 #_ #IHLK1 #K2 #f1 #HK12 #f2 #Hf2
+| #f #I1 #L1 #K1 #_ #IHLK1 #K2 #f1 #HK12 #f2 #Hf2
elim (coafter_inv_nxx … Hf2) -Hf2 [2,3: // ] #g2 #Hg2 #H destruct
elim (IHLK1 … HK12 … Hg2) -K1
/3 width=6 by drops_drop, lexs_next, lexs_push, ex3_intro/
-| #f #I #L1 #K1 #V1 #W1 #HLK1 #HWV1 #IHLK1 #X #f1 #H #f2 #Hf2
+| #f #I1 #J1 #L1 #K1 #HLK1 #HJI1 #IHLK1 #X #f1 #H #f2 #Hf2
elim (coafter_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
+ [ elim (lexs_inv_push1 … H) | elim (lexs_inv_next1 … H) ] -H #J2 #K2 #HK12 #HJ12 #H destruct
+ [ elim (H2RP … HJ12 … HLK1 … HJI1) | elim (H2RN … HJ12 … HLK1 … HJI1) ] -J1
elim (IHLK1 … HK12 … Hg2) -K1
/3 width=6 by drops_skip, lexs_next, lexs_push, ex3_intro/
]
#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/
-| #f #I #L2 #K2 #V2 #_ #IH #Hf #f2 #X #HX #f1 #Hf2
+| #f #I2 #L2 #K2 #_ #IH #Hf #f2 #X #HX #f1 #Hf2
elim (coafter_inv_nxx … Hf2) -Hf2 [2,3: // ] #g2 #Hg2 #H destruct
- elim (lexs_inv_push2 … HX) -HX #L1 #V1 #HL12 #HV12 #H destruct
- elim (IH … HL12 … Hg2) -L2 -V2 -g2
+ elim (lexs_inv_push2 … HX) -HX #I1 #L1 #HL12 #HI12 #H destruct
+ elim (IH … HL12 … Hg2) -L2 -I2 -g2
/3 width=3 by drops_drop, isuni_inv_next, ex2_intro/
-| #f #I #L2 #K2 #V2 #W2 #_ #HWV2 #IH #Hf #f2 #X #HX #f1 #Hf2
+| #f #I2 #J2 #L2 #K2 #_ #HJI2 #IH #Hf #f2 #X #HX #f1 #Hf2
elim (coafter_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 (lexs_inv_push2 … HX) | elim (lexs_inv_next2 … HX) ] -HX #I1 #L1 #HL12 #HI12 #H destruct
elim (IH … HL12 … Hg2) -L2 -g2 /2 width=3 by isuni_fwd_push/ #K1 #HLK1 #HK12
lapply (isuni_inv_push … Hf ??) -Hf [3,6: |*: // ] #Hf
- lapply (lifts_fwd_isid … HWV2 … Hf) #H destruct -HWV2
+ lapply (liftsb_fwd_isid … HJI2 … Hf) #H destruct -HJI2
lapply (drops_fwd_isid … HLK1 … Hf) #H destruct -HLK1
/4 width=5 by lexs_next, lexs_push, drops_refl, isid_push, ex2_intro/
]
(* Basic_2A1: includes: lpx_sn_drop_conf *) (**)
lemma lexs_drops_conf_next: ∀RN,RP.
∀f2,L1,L2. L1 ⪤*[RN, RP, f2] L2 →
- ∀b,f,I,K1,V1. ⬇*[b,f] L1 ≡ K1.ⓑ{I}V1 → 𝐔⦃f⦄ →
+ ∀b,f,I1,K1. ⬇*[b, f] L1 ≡ K1.ⓘ{I1} → 𝐔⦃f⦄ →
∀f1. f ~⊚ ⫯f1 ≡ f2 →
- ∃∃K2,V2. ⬇*[b,f] L2 ≡ K2.ⓑ{I}V2 & K1 ⪤*[RN, RP, f1] K2 & RN K1 V1 V2.
-#RN #RP #f2 #L1 #L2 #HL12 #b #f #I #K1 #V1 #HLK1 #Hf #f1 #Hf2
+ ∃∃I2,K2. ⬇*[b, f] L2 ≡ K2.ⓘ{I2} & K1 ⪤*[RN, RP, f1] K2 & RN K1 I1 I2.
+#RN #RP #f2 #L1 #L2 #HL12 #b #f #I1 #K1 #HLK1 #Hf #f1 #Hf2
elim (lexs_co_dropable_sn … HLK1 … Hf … HL12 … Hf2) -L1 -f2 -Hf
#X #HX #HLK2 elim (lexs_inv_next1 … HX) -HX
-#K2 #V2 #HK12 #HV12 #H destruct /2 width=5 by ex3_2_intro/
+#I2 #K2 #HK12 #HI12 #H destruct /2 width=5 by ex3_2_intro/
qed-.
lemma lexs_drops_conf_push: ∀RN,RP.
∀f2,L1,L2. L1 ⪤*[RN, RP, f2] L2 →
- ∀b,f,I,K1,V1. ⬇*[b,f] L1 ≡ K1.ⓑ{I}V1 → 𝐔⦃f⦄ →
+ ∀b,f,I1,K1. ⬇*[b, f] L1 ≡ K1.ⓘ{I1} → 𝐔⦃f⦄ →
∀f1. f ~⊚ ↑f1 ≡ f2 →
- ∃∃K2,V2. ⬇*[b,f] L2 ≡ K2.ⓑ{I}V2 & K1 ⪤*[RN, RP, f1] K2 & RP K1 V1 V2.
-#RN #RP #f2 #L1 #L2 #HL12 #b #f #I #K1 #V1 #HLK1 #Hf #f1 #Hf2
+ ∃∃I2,K2. ⬇*[b, f] L2 ≡ K2.ⓘ{I2} & K1 ⪤*[RN, RP, f1] K2 & RP K1 I1 I2.
+#RN #RP #f2 #L1 #L2 #HL12 #b #f #I1 #K1 #HLK1 #Hf #f1 #Hf2
elim (lexs_co_dropable_sn … HLK1 … Hf … HL12 … Hf2) -L1 -f2 -Hf
#X #HX #HLK2 elim (lexs_inv_push1 … HX) -HX
-#K2 #V2 #HK12 #HV12 #H destruct /2 width=5 by ex3_2_intro/
+#I2 #K2 #HK12 #HI12 #H destruct /2 width=5 by ex3_2_intro/
qed-.
(* Basic_2A1: includes: lpx_sn_drop_trans *)
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⦄ →
+ ∀b,f,I2,K2. ⬇*[b, f] L2 ≡ K2.ⓘ{I2} → 𝐔⦃f⦄ →
∀f1. f ~⊚ ⫯f1 ≡ f2 →
- ∃∃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_co_dropable_dx … HL12 … HLK1 … Hf … Hf2) -L2 -f2 -Hf
+ ∃∃I1,K1. ⬇*[b, f] L1 ≡ K1.ⓘ{I1} & K1 ⪤*[RN, RP, f1] K2 & RN K1 I1 I2.
+#RN #RP #f2 #L1 #L2 #HL12 #b #f #I2 #K2 #HLK2 #Hf #f1 #Hf2
+elim (lexs_co_dropable_dx … HL12 … HLK2 … 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/
+#I1 #K1 #HK12 #HI12 #H destruct /2 width=5 by ex3_2_intro/
qed-.
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⦄ →
+ ∀b,f,I2,K2. ⬇*[b, f] L2 ≡ K2.ⓘ{I2} → 𝐔⦃f⦄ →
∀f1. f ~⊚ ↑f1 ≡ f2 →
- ∃∃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_co_dropable_dx … HL12 … HLK1 … Hf … Hf2) -L2 -f2 -Hf
+ ∃∃I1,K1. ⬇*[b, f] L1 ≡ K1.ⓘ{I1} & K1 ⪤*[RN, RP, f1] K2 & RP K1 I1 I2.
+#RN #RP #f2 #L1 #L2 #HL12 #b #f #I2 #K2 #HLK2 #Hf #f1 #Hf2
+elim (lexs_co_dropable_dx … HL12 … HLK2 … 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/
+#I1 #K1 #HK12 #HI12 #H destruct /2 width=5 by ex3_2_intro/
qed-.
lemma drops_lexs_trans_next: ∀RN,RP. (∀L. reflexive ? (RN L)) → (∀L. reflexive ? (RP L)) →
- d_liftable2_sn RN → d_liftable2_sn RP →
+ d_liftable2_sn … liftsb RN → d_liftable2_sn … liftsb RP →
∀f1,K1,K2. K1 ⪤*[RN, RP, f1] K2 →
- ∀b,f,I,L1,V1. ⬇*[b,f] L1.ⓑ{I}V1 ≡ K1 →
+ ∀b,f,I1,L1. ⬇*[b, f] L1.ⓘ{I1} ≡ K1 →
∀f2. f ~⊚ f1 ≡ ⫯f2 →
- ∃∃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
+ ∃∃I2,L2. ⬇*[b, f] L2.ⓘ{I2} ≡ K2 & L1 ⪤*[RN, RP, f2] L2 & RN L1 I1 I2 & L1.ⓘ{I1} ≡[f] L2.ⓘ{I2}.
+#RN #RP #H1RN #H1RP #H2RN #H2RP #f1 #K1 #K2 #HK12 #b #f #I1 #L1 #HLK1 #f2 #Hf2
elim (lexs_liftable_co_dedropable_sn … 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/
+#I2 #L2 #H2L12 #HI12 #H destruct /2 width=6 by ex4_2_intro/
qed-.
lemma drops_lexs_trans_push: ∀RN,RP. (∀L. reflexive ? (RN L)) → (∀L. reflexive ? (RP L)) →
- d_liftable2_sn RN → d_liftable2_sn RP →
+ d_liftable2_sn … liftsb RN → d_liftable2_sn … liftsb RP →
∀f1,K1,K2. K1 ⪤*[RN, RP, f1] K2 →
- ∀b,f,I,L1,V1. ⬇*[b,f] L1.ⓑ{I}V1 ≡ K1 →
+ ∀b,f,I1,L1. ⬇*[b, f] L1.ⓘ{I1} ≡ K1 →
∀f2. f ~⊚ f1 ≡ ↑f2 →
- ∃∃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
+ ∃∃I2,L2. ⬇*[b, f] L2.ⓘ{I2} ≡ K2 & L1 ⪤*[RN, RP, f2] L2 & RP L1 I1 I2 & L1.ⓘ{I1} ≡[f] L2.ⓘ{I2}.
+#RN #RP #H1RN #H1RP #H2RN #H2RP #f1 #K1 #K2 #HK12 #b #f #I1 #L1 #HLK1 #f2 #Hf2
elim (lexs_liftable_co_dedropable_sn … 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/
+#I2 #L2 #H2L12 #HI12 #H destruct /2 width=6 by ex4_2_intro/
qed-.
lemma drops_atom2_lexs_conf: ∀RN,RP,b,f1,L1. ⬇*[b, f1] L1 ≡ ⋆ → 𝐔⦃f1⦄ →
(* *)
(**************************************************************************)
-include "basic_2/relocation/drops_ceq.ma".
include "basic_2/relocation/drops_lexs.ma".
(* GENERIC SLICING FOR LOCAL ENVIRONMENTS ***********************************)
(* Basic_2A1: includes: lreq_drop_trans_be *)
lemma lreq_drops_trans_next: ∀f2,L1,L2. L1 ≡[f2] L2 →
- ∀b,f,I,K2,V. ⬇*[b,f] L2 ≡ K2.ⓑ{I}V → 𝐔⦃f⦄ →
+ ∀b,f,I,K2. ⬇*[b, f] L2 ≡ K2.ⓘ{I} → 𝐔⦃f⦄ →
∀f1. f ~⊚ ⫯f1 ≡ f2 →
- ∃∃K1. ⬇*[b,f] L1 ≡ K1.ⓑ{I}V & K1 ≡[f1] K2.
-#f2 #L1 #L2 #HL12 #b #f #I #K1 #V #HLK1 #Hf #f1 #Hf2
+ ∃∃K1. ⬇*[b, f] L1 ≡ K1.ⓘ{I} & K1 ≡[f1] K2.
+#f2 #L1 #L2 #HL12 #b #f #I #K1 #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: ∀f2,L1,L2. L1 ≡[f2] L2 →
- ∀b,f,I,K1,V. ⬇*[b,f] L1 ≡ K1.ⓑ{I}V → 𝐔⦃f⦄ →
+ ∀b,f,I,K1. ⬇*[b, f] L1 ≡ K1.ⓘ{I} → 𝐔⦃f⦄ →
∀f1. f ~⊚ ⫯f1 ≡ f2 →
- ∃∃K2. ⬇*[b,f] L2 ≡ K2.ⓑ{I}V & K1 ≡[f1] K2.
-#f2 #L1 #L2 #HL12 #b #f #I #K1 #V #HLK1 #Hf #f1 #Hf2
+ ∃∃K2. ⬇*[b, f] L2 ≡ K2.ⓘ{I} & K1 ≡[f1] K2.
+#f2 #L1 #L2 #HL12 #b #f #I #K1 #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: ∀f1,K1,K2. K1 ≡[f1] K2 →
- ∀b,f,I,L1,V. ⬇*[b,f] L1.ⓑ{I}V ≡ K1 →
+ ∀b,f,I,L1. ⬇*[b, f] L1.ⓘ{I} ≡ K1 →
∀f2. f ~⊚ f1 ≡ ⫯f2 →
- ∃∃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
+ ∃∃L2. ⬇*[b, f] L2.ⓘ{I} ≡ K2 & L1 ≡[f2] L2 & L1.ⓘ{I} ≡[f] L2.ⓘ{I}.
+#f1 #K1 #K2 #HK12 #b #f #I #L1 #HLK1 #f2 #Hf2
elim (drops_lexs_trans_next … HK12 … HLK1 … Hf2) -f1 -K1
/2 width=6 by cfull_lift_sn, ceq_lift_sn, ex3_intro/ qed-.
(* Properties with reflexive and transitive closure *************************)
(* Basic_2A1: was: d_liftable_LTC *)
-lemma d2_liftable_sn_LTC: ∀R. d_liftable2_sn R → d_liftable2_sn (LTC … R).
-#R #HR #K #T1 #T2 #H elim H -T2
+lemma d2_liftable_sn_LTC: ∀C,S,R. d_liftable2_sn C S R → d_liftable2_sn C S (LTC … R).
+#C #S #R #HR #K #T1 #T2 #H elim H -T2
[ #T2 #HT12 #b #f #L #HLK #U1 #HTU1
elim (HR … HT12 … HLK … HTU1) /3 width=3 by inj, ex2_intro/
| #T #T2 #_ #HT2 #IHT1 #b #f #L #HLK #U1 #HTU1
qed-.
(* 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
+lemma d2_deliftable_sn_LTC: ∀C,S,R. d_deliftable2_sn C S R → d_deliftable2_sn C S (LTC … R).
+#C #S #R #HR #L #U1 #U2 #H elim H -U2
[ #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 #b #f #K #HLK #T1 #HTU1
qed-.
(* Basic_2A1: was: d_liftable_llstar *)
-lemma d2_liftable_sn_llstar: ∀R. d_liftable2_sn R → ∀d. d_liftable2_sn (llstar … R d).
-#R #HR #d #K #T1 #T2 #H @(lstar_ind_r … d T2 H) -d -T2
+lemma d2_liftable_sn_llstar: ∀C,S,R. d_liftable2_sn C S R → ∀d. d_liftable2_sn C S (llstar … R d).
+#C #S #R #HR #d #K #T1 #T2 #H @(lstar_ind_r … d T2 H) -d -T2
[ #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
qed-.
(* 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 #l #L #U1 #U2 #H @(lstar_ind_r … l U2 H) -l -U2
+lemma d2_deliftable_sn_llstar: ∀C,S,R. d_deliftable2_sn C S R →
+ ∀d. d_deliftable2_sn C S (llstar … R d).
+#C #S #R #HR #l #L #U1 #U2 #H @(lstar_ind_r … l U2 H) -l -U2
[ /2 width=3 by lstar_O, ex2_intro/
| #l #U #U2 #_ #HU2 #IHU1 #b #f #K #HLK #T1 #HTU1
elim (IHU1 … HLK … HTU1) -IHU1 -U1 #T #HTU #HT1
(**************************************************************************)
include "basic_2/syntax/cl_restricted_weight.ma".
-include "basic_2/relocation/lifts_weight.ma".
+include "basic_2/relocation/lifts_weight_bind.ma".
include "basic_2/relocation/drops.ma".
(* GENERIC SLICING FOR LOCAL ENVIRONMENTS ***********************************)
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/
-| #f #I #L1 #L2 #V1 #V2 #_ #HV21 #IHL12 normalize
- >(lifts_fwd_tw … HV21) -HV21 /2 width=1 by monotonic_le_plus_l/
+| #f #I1 #I2 #L1 #L2 #_ #HI21 #IHL12 normalize
+ >(liftsb_fwd_bw … HI21) -HI21 /2 width=1 by monotonic_le_plus_l/
]
qed-.
#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/
-| #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/
+| #f #I1 #I2 #L1 #L2 #_ #HI21 #IHL12 #H normalize in ⊢ (?%%);
+ >(liftsb_fwd_bw … HI21) -I2 /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: ∀b,f,I,L,K,V. ⬇*[b, f] L ≡ K.ⓑ{I}V → ∀T. ♯{K, V} < ♯{L, T}.
+lemma drops_bind2_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/
+normalize in ⊢ (%→?→?%%); /3 width=3 by le_to_lt_to_lt, monotonic_lt_plus_r/
qed-.
(* Advanced inversion lemma *************************************************)
-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
+lemma drops_inv_x_bind_xy: ∀b,f,I,L. ⬇*[b, f] L ≡ L.ⓘ{I} → ⊥.
+#b #f #I #L #H lapply (drops_fwd_lw … H) -b -f
/2 width=4 by lt_le_false/ (**) (* full auto is a bit slow: 19s *)
qed-.
(* GENERIC ENTRYWISE EXTENSION OF CONTEXT-SENSITIVE REALTIONS FOR TERMS *****)
(* Basic_2A1: includes: lpx_sn_atom lpx_sn_pair *)
-inductive lexs (RN,RP:relation3 lenv term term): rtmap → relation lenv ≝
+inductive lexs (RN,RP:relation3 lenv bind bind): rtmap → relation lenv ≝
| lexs_atom: ∀f. lexs RN RP f (⋆) (⋆)
-| lexs_next: ∀f,I,L1,L2,V1,V2.
- lexs RN RP f L1 L2 → RN L1 V1 V2 →
- lexs RN RP (⫯f) (L1.â\93\91{I}V1) (L2.â\93\91{I}V2)
-| lexs_push: ∀f,I,L1,L2,V1,V2.
- lexs RN RP f L1 L2 → RP L1 V1 V2 →
- lexs RN RP (â\86\91f) (L1.â\93\91{I}V1) (L2.â\93\91{I}V2)
+| lexs_next: ∀f,I1,I2,L1,L2.
+ lexs RN RP f L1 L2 → RN L1 I1 I2 →
+ lexs RN RP (⫯f) (L1.â\93\98{I1}) (L2.â\93\98{I2})
+| lexs_push: ∀f,I1,I2,L1,L2.
+ lexs RN RP f L1 L2 → RP L1 I1 I2 →
+ lexs RN RP (â\86\91f) (L1.â\93\98{I1}) (L2.â\93\98{I2})
.
interpretation "generic entrywise extension (local environment)"
'RelationStar RN RP f L1 L2 = (lexs RN RP f L1 L2).
-definition R_pw_confluent2_lexs: relation3 lenv term term → relation3 lenv term term →
- relation3 lenv term term → relation3 lenv term term →
- relation3 lenv term term → relation3 lenv term term →
- relation3 rtmap lenv term ≝
+definition R_pw_confluent2_lexs: relation3 lenv bind bind → relation3 lenv bind bind →
+ relation3 lenv bind bind → relation3 lenv bind bind →
+ relation3 lenv bind bind → relation3 lenv bind bind →
+ relation3 rtmap lenv bind ≝
λR1,R2,RN1,RP1,RN2,RP2,f,L0,T0.
∀T1. R1 L0 T0 T1 → ∀T2. R2 L0 T0 T2 →
∀L1. L0 ⪤*[RN1, RP1, f] L1 → ∀L2. L0 ⪤*[RN2, RP2, f] L2 →
∃∃T. R2 L1 T1 T & R1 L2 T2 T.
-definition lexs_transitive: relation5 (relation3 lenv term term)
- (relation3 lenv term term) … ≝
+definition lexs_transitive: relation5 (relation3 lenv bind bind)
+ (relation3 lenv bind bind) … ≝
λR1,R2,R3,RN,RP.
∀f,L1,T1,T. R1 L1 T1 T → ∀L2. L1 ⪤*[RN, RP, f] L2 →
∀T2. R2 L2 T T2 → R3 L1 T1 T2.
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
+#f #I1 #I2 #L1 #L2 #_ #_ #H destruct
qed-.
(* Basic_2A1: includes lpx_sn_inv_atom1 *)
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,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.
+fact lexs_inv_next1_aux: ∀RN,RP,f,X,Y. X ⪤*[RN, RP, f] Y → ∀g,J1,K1. X = K1.ⓘ{J1} → f = ⫯g →
+ ∃∃J2,K2. K1 ⪤*[RN, RP, g] K2 & RN K1 J1 J2 & Y = K2.ⓘ{J2}.
#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
+[ #f #g #J1 #K1 #H destruct
+| #f #I1 #I2 #L1 #L2 #HL #HI #g #J1 #K1 #H1 #H2 <(injective_next … H2) -g destruct
/2 width=5 by ex3_2_intro/
-| #f #I #L1 #L2 #V1 #V2 #_ #_ #g #J #K1 #W1 #_ #H elim (discr_push_next … H)
+| #f #I1 #I2 #L1 #L2 #_ #_ #g #J1 #K1 #_ #H elim (discr_push_next … H)
]
qed-.
(* Basic_2A1: includes lpx_sn_inv_pair1 *)
-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.
+lemma lexs_inv_next1: ∀RN,RP,g,J1,K1,Y. K1.ⓘ{J1} ⪤*[RN, RP, ⫯g] Y →
+ ∃∃J2,K2. K1 ⪤*[RN, RP, g] K2 & RN K1 J1 J2 & Y = K2.ⓘ{J2}.
/2 width=7 by lexs_inv_next1_aux/ qed-.
-
-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.
+fact lexs_inv_push1_aux: ∀RN,RP,f,X,Y. X ⪤*[RN, RP, f] Y → ∀g,J1,K1. X = K1.ⓘ{J1} → f = ↑g →
+ ∃∃J2,K2. K1 ⪤*[RN, RP, g] K2 & RP K1 J1 J2 & Y = K2.ⓘ{J2}.
#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
+[ #f #g #J1 #K1 #H destruct
+| #f #I1 #I2 #L1 #L2 #_ #_ #g #J1 #K1 #_ #H elim (discr_next_push … H)
+| #f #I1 #I2 #L1 #L2 #HL #HI #g #J1 #K1 #H1 #H2 <(injective_push … H2) -g destruct
/2 width=5 by ex3_2_intro/
]
qed-.
-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.
+lemma lexs_inv_push1: ∀RN,RP,g,J1,K1,Y. K1.ⓘ{J1} ⪤*[RN, RP, ↑g] Y →
+ ∃∃J2,K2. K1 ⪤*[RN, RP, g] K2 & RP K1 J1 J2 & Y = K2.ⓘ{J2}.
/2 width=7 by lexs_inv_push1_aux/ qed-.
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
+#f #I1 #I2 #L1 #L2 #_ #_ #H destruct
qed-.
(* Basic_2A1: includes lpx_sn_inv_atom2 *)
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,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.
+fact lexs_inv_next2_aux: ∀RN,RP,f,X,Y. X ⪤*[RN, RP, f] Y → ∀g,J2,K2. Y = K2.ⓘ{J2} → f = ⫯g →
+ ∃∃J1,K1. K1 ⪤*[RN, RP, g] K2 & RN K1 J1 J2 & X = K1.ⓘ{J1}.
#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
+[ #f #g #J2 #K2 #H destruct
+| #f #I1 #I2 #L1 #L2 #HL #HI #g #J2 #K2 #H1 #H2 <(injective_next … H2) -g destruct
/2 width=5 by ex3_2_intro/
-| #f #I #L1 #L2 #V1 #V2 #_ #_ #g #J #K2 #W2 #_ #H elim (discr_push_next … H)
+| #f #I1 #I2 #L1 #L2 #_ #_ #g #J2 #K2 #_ #H elim (discr_push_next … H)
]
qed-.
(* Basic_2A1: includes lpx_sn_inv_pair2 *)
-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.
+lemma lexs_inv_next2: ∀RN,RP,g,J2,X,K2. X ⪤*[RN, RP, ⫯g] K2.ⓘ{J2} →
+ ∃∃J1,K1. K1 ⪤*[RN, RP, g] K2 & RN K1 J1 J2 & X = K1.ⓘ{J1}.
/2 width=7 by lexs_inv_next2_aux/ qed-.
-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.
+fact lexs_inv_push2_aux: ∀RN,RP,f,X,Y. X ⪤*[RN, RP, f] Y → ∀g,J2,K2. Y = K2.ⓘ{J2} → f = ↑g →
+ ∃∃J1,K1. K1 ⪤*[RN, RP, g] K2 & RP K1 J1 J2 & X = K1.ⓘ{J1}.
#RN #RP #f #X #Y * -f -X -Y
-[ #f #J #K2 #W2 #g #H 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
+[ #f #J2 #K2 #g #H destruct
+| #f #I1 #I2 #L1 #L2 #_ #_ #g #J2 #K2 #_ #H elim (discr_next_push … H)
+| #f #I1 #I2 #L1 #L2 #HL #HI #g #J2 #K2 #H1 #H2 <(injective_push … H2) -g destruct
/2 width=5 by ex3_2_intro/
]
qed-.
-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.
+lemma lexs_inv_push2: ∀RN,RP,g,J2,X,K2. X ⪤*[RN, RP, ↑g] K2.ⓘ{J2} →
+ ∃∃J1,K1. K1 ⪤*[RN, RP, g] K2 & RP K1 J1 J2 & X = K1.ⓘ{J1}.
/2 width=7 by lexs_inv_push2_aux/ qed-.
(* Basic_2A1: includes lpx_sn_inv_pair *)
-lemma lexs_inv_next: ∀RN,RP,f,I1,I2,L1,L2,V1,V2.
- L1.â\93\91{I1}V1 ⪤*[RN, RP, ⫯f] (L2.â\93\91{I2}V2) →
- ∧∧ L1 ⪤*[RN, RP, f] L2 & RN L1 V1 V2 & I1 = I2.
-#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/
+lemma lexs_inv_next: ∀RN,RP,f,I1,I2,L1,L2.
+ L1.â\93\98{I1} ⪤*[RN, RP, ⫯f] L2.â\93\98{I2} →
+ L1 ⪤*[RN, RP, f] L2 ∧ RN L1 I1 I2.
+#RN #RP #f #I1 #I2 #L1 #L2 #H elim (lexs_inv_next1 … H) -H
+#I0 #L0 #HL10 #HI10 #H destruct /2 width=1 by conj/
qed-.
-lemma lexs_inv_push: ∀RN,RP,f,I1,I2,L1,L2,V1,V2.
- L1.â\93\91{I1}V1 ⪤*[RN, RP, â\86\91f] (L2.â\93\91{I2}V2) →
- ∧∧ L1 ⪤*[RN, RP, f] L2 & RP L1 V1 V2 & I1 = I2.
-#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/
+lemma lexs_inv_push: ∀RN,RP,f,I1,I2,L1,L2.
+ L1.â\93\98{I1} ⪤*[RN, RP, â\86\91f] L2.â\93\98{I2} →
+ L1 ⪤*[RN, RP, f] L2 ∧ RP L1 I1 I2.
+#RN #RP #f #I1 #I2 #L1 #L2 #H elim (lexs_inv_push1 … H) -H
+#I0 #L0 #HL10 #HI10 #H destruct /2 width=1 by conj/
qed-.
-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.â\93\91{I}V1 ⪤*[RN, RP, f] L2.â\93\91{I}V2.
-#RN #RP #f #I #L2 #L2 #V1 #V2 elim (pn_split f) *
+lemma lexs_inv_tl: ∀RN,RP,f,I1,I2,L1,L2. L1 ⪤*[RN, RP, ⫱f] L2 →
+ RN L1 I1 I2 → RP L1 I1 I2 →
+ L1.â\93\98{I1} ⪤*[RN, RP, f] L2.â\93\98{I2}.
+#RN #RP #f #I1 #I2 #L2 #L2 elim (pn_split f) *
/2 width=1 by lexs_next, lexs_push/
qed-.
(* Basic forward lemmas *****************************************************)
-lemma lexs_fwd_pair: ∀RN,RP,f,I1,I2,L1,L2,V1,V2.
- L1.â\93\91{I1}V1 ⪤*[RN, RP, f] L2.â\93\91{I2}V2 →
- L1 ⪤*[RN, RP, ⫱f] L2 ∧ I1 = I2.
-#RN #RP #f #I1 #I2 #L2 #L2 #V1 #V2 #Hf
+lemma lexs_fwd_bind: ∀RN,RP,f,I1,I2,L1,L2.
+ L1.â\93\98{I1} ⪤*[RN, RP, f] L2.â\93\98{I2} →
+ L1 ⪤*[RN, RP, ⫱f] L2.
+#RN #RP #f #I1 #I2 #L1 #L2 #Hf
elim (pn_split f) * #g #H destruct
-[ elim (lexs_inv_push … Hf) | elim (lexs_inv_next … Hf) ] -Hf
-/2 width=1 by conj/
+[ elim (lexs_inv_push … Hf) | elim (lexs_inv_next … Hf) ] -Hf //
qed-.
(* 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 -f1 -L1 -L2 //
-#f1 #I #L1 #L2 #V1 #v2 #_ #HV #IH #f2 #H
+#f1 #I1 #I2 #L1 #L2 #_ #HI #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/
]
#RN #RP #L1 #L2 @eq_repl_sym /2 width=3 by lexs_eq_repl_back/ (**) (* full auto fails *)
qed-.
-(* Basic_2A1: includes: lpx_sn_refl *)
+(* Basic_2A1: uses: lpx_sn_refl *)
lemma lexs_refl: ∀RN,RP.
(∀L. reflexive … (RN L)) →
(∀L. reflexive … (RP L)) →
∀f.reflexive … (lexs RN RP f).
#RN #RP #HRN #HRP #f #L generalize in match f; -f elim L -L //
-#L #I #V #IH * * /2 width=1 by lexs_next, lexs_push/
+#L #I #IH #f elim (pn_split f) *
+#g #H destruct /2 width=1 by lexs_next, lexs_push/
qed.
lemma lexs_sym: ∀RN,RP.
- (∀L1,L2,T1,T2. RN L1 T1 T2 → RN L2 T2 T1) →
- (∀L1,L2,T1,T2. RP L1 T1 T2 → RP L2 T2 T1) →
+ (∀L1,L2,I1,I2. RN L1 I1 I2 → RN L2 I2 I1) →
+ (∀L1,L2,I1,I2. RP L1 I1 I2 → RP L2 I2 I1) →
∀f. symmetric … (lexs RN RP f).
#RN #RP #HRN #HRP #f #L1 #L2 #H elim H -L1 -L2 -f
/3 width=2 by lexs_next, lexs_push/
qed-.
-lemma lexs_pair_repl: ∀RN,RP,f,I,L1,L2,V1,V2.
- L1.ⓑ{I}V1 ⪤*[RN, RP, f] L2.ⓑ{I}V2 →
- ∀W1,W2. RN L1 W1 W2 → RP L1 W1 W2 →
- L1.ⓑ{I}W1 ⪤*[RN, RP, f] L2.ⓑ{I}W2.
-#RN #RP #f #I #L1 #L2 #V1 #V2 #HL12 #W1 #W2 #HN #HP
-elim (lexs_fwd_pair … HL12) -HL12 /2 width=1 by lexs_inv_tl/
-qed-.
+lemma lexs_pair_repl: ∀RN,RP,f,I1,I2,L1,L2.
+ L1.ⓘ{I1} ⪤*[RN, RP, f] L2.ⓘ{I2} →
+ ∀J1,J2. RN L1 J1 J2 → RP L1 J1 J2 →
+ L1.ⓘ{J1} ⪤*[RN, RP, f] L2.ⓘ{J2}.
+/3 width=3 by lexs_inv_tl, lexs_fwd_bind/ qed-.
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,I1,I2. RN1 L1 I1 I2 → RN2 L1 I1 I2) →
+ (∀L1,I1,I2. RP1 L1 I1 I2 → RP2 L1 I1 I2) →
∀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-.
lemma lexs_co_isid: ∀RN1,RP1,RN2,RP2.
- (∀L1,T1,T2. RP1 L1 T1 T2 → RP2 L1 T1 T2) →
+ (∀L1,I1,I2. RP1 L1 I1 I2 → RP2 L1 I1 I2) →
∀f,L1,L2. L1 ⪤*[RN1, RP1, f] L2 → 𝐈⦃f⦄ →
L1 ⪤*[RN2, RP2, f] L2.
#RN1 #RP1 #RN2 #RP2 #HR #f #L1 #L2 #H elim H -f -L1 -L2 //
-#f #I #K1 #K2 #V1 #V2 #_ #HV12 #IH #H
+#f #I1 #I2 #K1 #K2 #_ #HI12 #IH #H
[ elim (isid_inv_next … H) -H //
| /4 width=3 by lexs_push, isid_inv_push/
-]
+]
qed-.
-lemma sle_lexs_trans: ∀RN,RP. (∀L,T1,T2. RN L T1 T2 → RP L T1 T2) →
+lemma sle_lexs_trans: ∀RN,RP. (∀L,I1,I2. RN L I1 I2 → RP L I1 I2) →
∀f2,L1,L2. L1 ⪤*[RN, RP, f2] L2 →
∀f1. f1 ⊆ f2 → L1 ⪤*[RN, RP, f1] L2.
#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/
-| elim (sle_inv_xp … H) -H [2,3: // ]
+#f2 #I1 #I2 #L1 #L2 #_ #HI12 #IH #f1 #H12
+[ elim (pn_split f1) * ]
+[ /4 width=5 by lexs_push, sle_inv_pn/
+| /4 width=5 by lexs_next, sle_inv_nn/
+| elim (sle_inv_xp … H12) -H12 [2,3: // ]
#g1 #H #H1 destruct /3 width=5 by lexs_push/
]
qed-.
-lemma sle_lexs_conf: ∀RN,RP. (∀L,T1,T2. RP L T1 T2 → RN L T1 T2) →
+lemma sle_lexs_conf: ∀RN,RP. (∀L,I1,I2. RP L I1 I2 → RN L I1 I2) →
∀f1,L1,L2. L1 ⪤*[RN, RP, f1] L2 →
∀f2. f1 ⊆ f2 → L1 ⪤*[RN, RP, f2] L2.
#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/
-| elim (sle_inv_nx … H) -H [2,3: // ]
+#f1 #I1 #I2 #L1 #L2 #_ #HI12 #IH #f2 #H12
+[2: elim (pn_split f2) * ]
+[ /4 width=5 by lexs_push, sle_inv_pp/
+| /4 width=5 by lexs_next, sle_inv_pn/
+| elim (sle_inv_nx … H12) -H12 [2,3: // ]
#g2 #H #H2 destruct /3 width=5 by lexs_next/
]
qed-.
∃∃L. L1 ⪤*[R1, RP, g] L & L ⪤*[R2, cfull, f] L2.
#R1 #R2 #RP #HR1 #HR2 #f #L1 #L2 #H elim H -f -L1 -L2
[ /2 width=3 by lexs_atom, ex2_intro/ ]
-#f #I #L1 #L2 #V1 #V2 #_ #HV12 #IH #y #H
+#f #I1 #I2 #L1 #L2 #_ #HI12 #IH #y #H
[ elim (sle_inv_nx … H ??) -H [ |*: // ] #g #Hfg #H destruct
elim (IH … Hfg) -IH -Hfg /3 width=5 by lexs_next, ex2_intro/
| elim (sle_inv_px … H ??) -H [1,3: * |*: // ] #g #Hfg #H destruct
qed-.
lemma lexs_dec: ∀RN,RP.
- (∀L,T1,T2. Decidable (RN L T1 T2)) →
- (∀L,T1,T2. Decidable (RP L T1 T2)) →
+ (∀L,I1,I2. Decidable (RN L I1 I2)) →
+ (∀L,I1,I2. Decidable (RP L I1 I2)) →
∀L1,L2,f. Decidable (L1 ⪤*[RN, RP, f] L2).
-#RN #RP #HRN #HRP #L1 elim L1 -L1 [ * | #L1 #I1 #V1 #IH * ]
+#RN #RP #HRN #HRP #L1 elim L1 -L1 [ * | #L1 #I1 #IH * ]
[ /2 width=1 by lexs_atom, or_introl/
-| #L2 #I2 #V2 #f @or_intror #H
+| #L2 #I2 #f @or_intror #H
lapply (lexs_inv_atom1 … H) -H #H destruct
| #f @or_intror #H
lapply (lexs_inv_atom2 … H) -H #H destruct
-| #L2 #I2 #V2 #f elim (eq_bind2_dec I1 I2) #H destruct
- [2: @or_intror #H elim (lexs_fwd_pair … H) -H /2 width=1 by/ ]
- elim (IH L2 (⫱f)) -IH #HL12
- [2: @or_intror #H elim (lexs_fwd_pair … H) -H /2 width=1 by/ ]
+| #L2 #I2 #f elim (IH L2 (⫱f)) -IH #HL12
+ [2: /4 width=3 by lexs_fwd_bind, or_intror/ ]
elim (pn_split f) * #g #H destruct
- [ elim (HRP L1 V1 V2) | elim (HRN L1 V1 V2) ] -HRP -HRN #HV12
+ [ elim (HRP L1 I1 I2) | elim (HRN L1 I1 I2) ] -HRP -HRN #HV12
[1,3: /3 width=1 by lexs_push, lexs_next, or_introl/ ]
@or_intror #H
[ elim (lexs_inv_push … H) | elim (lexs_inv_next … H) ] -H
(* Basic_2A1: includes: lpx_sn_fwd_length *)
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 //
-#f #I #L1 #L2 #V1 #V2 >length_pair >length_pair //
+#f #I1 #I2 #L1 #L2 >length_bind >length_bind //
qed-.
L1 ⪤*[RN, RP, f] L2.
#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 //
-| #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/
-| #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/
+| #f #I1 #I #K1 #K #HK1 #HI1 #IH #L2 #H elim (lexs_inv_next1 … H) -H
+ #I2 #K2 #HK2 #HI2 #H destruct /4 width=6 by lexs_next/
+| #f #I1 #I #K1 #K #HK1 #HI1 #IH #L2 #H elim (lexs_inv_push1 … H) -H
+ #I2 #K2 #HK2 #HI2 #H destruct /4 width=6 by lexs_push/
]
qed-.
(* Basic_2A1: includes: lpx_sn_conf *)
theorem lexs_conf (RN1) (RP1) (RN2) (RP2):
∀L,f.
- (∀g,I,K,V,n. ⬇*[n] L ≡ K.ⓑ{I}V → ⫯g = ⫱*[n] f → R_pw_confluent2_lexs RN1 RN2 RN1 RP1 RN2 RP2 g K V) →
- (∀g,I,K,V,n. ⬇*[n] L ≡ K.ⓑ{I}V → ↑g = ⫱*[n] f → R_pw_confluent2_lexs RP1 RP2 RN1 RP1 RN2 RP2 g K V) →
+ (∀g,I,K,n. ⬇*[n] L ≡ K.ⓘ{I} → ⫯g = ⫱*[n] f → R_pw_confluent2_lexs RN1 RN2 RN1 RP1 RN2 RP2 g K I) →
+ (∀g,I,K,n. ⬇*[n] L ≡ K.ⓘ{I} → ↑g = ⫱*[n] f → R_pw_confluent2_lexs RP1 RP2 RN1 RP1 RN2 RP2 g K I) →
pw_confluent2 … (lexs RN1 RP1 f) (lexs RN2 RP2 f) L.
#RN1 #RP1 #RN2 #RP2 #L elim L -L
[ #f #_ #_ #L1 #H1 #L2 #H2 >(lexs_inv_atom1 … H1) >(lexs_inv_atom1 … H2) -H2 -H1
/2 width=3 by lexs_atom, ex2_intro/
-| #L #I #V #IH #f elim (pn_split f) * #g #H destruct
+| #L #I0 #IH #f elim (pn_split f) * #g #H destruct
#HN #HP #Y1 #H1 #Y2 #H2
- [ elim (lexs_inv_push1 … H1) -H1 #L1 #V1 #HL1 #HV1 #H destruct
- elim (lexs_inv_push1 … H2) -H2 #L2 #V2 #HL2 #HV2 #H destruct
- elim (HP … 0 … HV1 … HV2 … HL1 … HL2) -HV1 -HV2 /2 width=2 by drops_refl/ #V #HV1 #HV2
+ [ elim (lexs_inv_push1 … H1) -H1 #I1 #L1 #HL1 #HI01 #H destruct
+ elim (lexs_inv_push1 … H2) -H2 #I2 #L2 #HL2 #HI02 #H destruct
+ elim (HP … 0 … HI01 … HI02 … HL1 … HL2) -HI01 -HI02 /2 width=2 by drops_refl/ #I #HI1 #HI2
elim (IH … HL1 … HL2) -IH -HL1 -HL2 /3 width=5 by drops_drop, lexs_push, ex2_intro/
- | elim (lexs_inv_next1 … H1) -H1 #L1 #V1 #HL1 #HV1 #H destruct
- elim (lexs_inv_next1 … H2) -H2 #L2 #V2 #HL2 #HV2 #H destruct
- elim (HN … 0 … HV1 … HV2 … HL1 … HL2) -HV1 -HV2 /2 width=2 by drops_refl/ #V #HV1 #HV2
+ | elim (lexs_inv_next1 … H1) -H1 #I1 #L1 #HL1 #HI01 #H destruct
+ elim (lexs_inv_next1 … H2) -H2 #I2 #L2 #HL2 #HI02 #H destruct
+ elim (HN … 0 … HI01 … HI02 … HL1 … HL2) -HI01 -HI02 /2 width=2 by drops_refl/ #I #HI1 #HI2
elim (IH … HL1 … HL2) -IH -HL1 -HL2 /3 width=5 by drops_drop, lexs_next, ex2_intro/
]
]
∀f2. L1 ⪤*[RN, RP, f2] L2 →
∀f. f1 ⋒ f2 ≡ f → L1 ⪤*[RN, RP, f] L2.
#RN #RP #L1 #L2 #f1 #H elim H -f1 -L1 -L2 //
-#f1 #I #L1 #L2 #V1 #V2 #_ #HV12 #IH #f2 #H #f #Hf
+#f1 #I1 #I2 #L1 #L2 #_ #HI12 #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)
∀f2. L1 ⪤*[RN, RP, f2] L2 →
∀f. f1 ⋓ f2 ≡ f → L1 ⪤*[RN, RP, f] L2.
#RN #RP #L1 #L2 #f1 #H elim H -f1 -L1 -L2 //
-#f1 #I #L1 #L2 #V1 #V2 #_ #HV12 #IH #f2 #H #f #Hf
+#f1 #I1 #I2 #L1 #L2 #_ #HI12 #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)
(* *)
(**************************************************************************)
-include "basic_2/syntax/bind_ext2.ma".
+include "basic_2/syntax/ext2.ma".
include "basic_2/relocation/lifts.ma".
(* GENERIC RELOCATION FOR BINDERS *******************************************)
(**************************************************************************)
include "basic_2/notation/relations/lazyeq_3.ma".
+include "basic_2/syntax/ext2.ma".
include "basic_2/relocation/lexs.ma".
(* RANGED EQUIVALENCE FOR LOCAL ENVIRONMENTS ********************************)
(* Basic_2A1: includes: lreq_refl *)
lemma lreq_refl: ∀f. reflexive … (lreq f).
-/2 width=1 by lexs_refl/ qed.
+/3 width=1 by lexs_refl, ext2_refl/ qed.
(* Basic_2A1: includes: lreq_sym *)
lemma lreq_sym: ∀f. symmetric … (lreq f).
-#f #L1 #L2 #H elim H -L1 -L2 -f
-/2 width=1 by lexs_next, lexs_push/
-qed-.
+/3 width=1 by lexs_sym, ext2_sym/ qed-.
(* Basic inversion lemmas ***************************************************)
/2 width=4 by lexs_inv_atom1/ qed-.
(* Basic_2A1: includes: lreq_inv_pair1 *)
-lemma lreq_inv_next1: ∀g,J,K1,Y,W1. K1.ⓑ{J}W1 ≡[⫯g] Y →
- â\88\83â\88\83K2. K1 â\89¡[g] K2 & Y = K2.â\93\91{J}W1.
-#g #J #K1 #Y #W1 #H elim (lexs_inv_next1 … H) -H /2 width=3 by ex2_intro/
+lemma lreq_inv_next1: ∀g,J,K1,Y. K1.ⓘ{J} ≡[⫯g] Y →
+ â\88\83â\88\83K2. K1 â\89¡[g] K2 & Y = K2.â\93\98{J}.
+#g #J #K1 #Y #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: ∀g,J,K1,Y,W1. K1.ⓑ{J}W1 ≡[↑g] Y →
- ∃∃K2,W2. K1 ≡[g] K2 & Y = K2.ⓑ{J}W2.
-#g #J #K1 #Y #W1 #H elim (lexs_inv_push1 … H) -H /2 width=4 by ex2_2_intro/ qed-.
+lemma lreq_inv_push1: ∀g,J1,K1,Y. K1.ⓘ{J1} ≡[↑g] Y →
+ ∃∃J2,K2. K1 ≡[g] K2 & Y = K2.ⓘ{J2}.
+#g #J1 #K1 #Y #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: ∀f,X. X ≡[f] ⋆ → X = ⋆.
/2 width=4 by lexs_inv_atom2/ qed-.
(* Basic_2A1: includes: lreq_inv_pair2 *)
-lemma lreq_inv_next2: ∀g,J,X,K2,W2. X ≡[⫯g] K2.ⓑ{J}W2 →
- ∃∃K1. K1 ≡[g] K2 & X = K1.ⓑ{J}W2.
-#g #J #X #K2 #W2 #H elim (lexs_inv_next2 … H) -H /2 width=3 by ex2_intro/ qed-.
+lemma lreq_inv_next2: ∀g,J,X,K2. X ≡[⫯g] K2.ⓘ{J} →
+ ∃∃K1. K1 ≡[g] K2 & X = K1.ⓘ{J}.
+#g #J #X #K2 #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: ∀g,J,X,K2,W2. X ≡[↑g] K2.ⓑ{J}W2 →
- ∃∃K1,W1. K1 ≡[g] K2 & X = K1.ⓑ{J}W1.
-#g #J #X #K2 #W2 #H elim (lexs_inv_push2 … H) -H /2 width=4 by ex2_2_intro/ qed-.
+lemma lreq_inv_push2: ∀g,J2,X,K2. X ≡[↑g] K2.ⓘ{J2} →
+ ∃∃J1,K1. K1 ≡[g] K2 & X = K1.ⓘ{J1}.
+#g #J2 #X #K2 #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: ∀f,I1,I2,L1,L2,V1,V2.
- L1.ⓑ{I1}V1 ≡[⫯f] (L2.ⓑ{I2}V2) →
- ∧∧ L1 ≡[f] L2 & V1 = V2 & I1 = I2.
+lemma lreq_inv_next: ∀f,I1,I2,L1,L2. L1.ⓘ{I1} ≡[⫯f] L2.ⓘ{I2} →
+ L1 ≡[f] L2 ∧ I1 = I2.
/2 width=1 by lexs_inv_next/ qed-.
(* Basic_2A1: includes: lreq_inv_succ *)
-lemma lreq_inv_push: ∀f,I1,I2,L1,L2,V1,V2.
- L1.ⓑ{I1}V1 ≡[↑f] (L2.ⓑ{I2}V2) →
- L1 ≡[f] L2 ∧ I1 = I2.
-#f #I1 #I2 #L1 #L2 #V1 #V2 #H elim (lexs_inv_push … H) -H /2 width=1 by conj/
+lemma lreq_inv_push: ∀f,I1,I2,L1,L2. L1.ⓘ{I1} ≡[↑f] L2.ⓘ{I2} → L1 ≡[f] L2.
+#f #I1 #I2 #L1 #L2 #H elim (lexs_inv_push … H) -H /2 width=1 by conj/
qed-.
-lemma lreq_inv_tl: ∀f,I,L1,L2,V. L1 ≡[⫱f] L2 → L1.ⓑ{I}V ≡[f] L2.ⓑ{I}V.
+lemma lreq_inv_tl: ∀f,I,L1,L2. L1 ≡[⫱f] L2 → L1.ⓘ{I} ≡[f] L2.ⓘ{I}.
/2 width=1 by lexs_inv_tl/ qed-.
(* Basic_2A1: removed theorems 5:
include "ground_2/lib/star.ma".
include "basic_2/notation/relations/suptermplus_6.ma".
+include "basic_2/notation/relations/suptermplus_7.ma".
include "basic_2/s_transition/fqu.ma".
(* PLUS-ITERATED SUPCLOSURE *************************************************)
-definition fqup: tri_relation genv lenv term ≝ tri_TC … fqu.
+definition fqup: bool → tri_relation genv lenv term ≝
+ λb. tri_TC … (fqu b).
+
+interpretation "extended plus-iterated structural successor (closure)"
+ 'SupTermPlus b G1 L1 T1 G2 L2 T2 = (fqup b G1 L1 T1 G2 L2 T2).
interpretation "plus-iterated structural successor (closure)"
- 'SupTermPlus G1 L1 T1 G2 L2 T2 = (fqup G1 L1 T1 G2 L2 T2).
+ 'SupTermPlus G1 L1 T1 G2 L2 T2 = (fqup true G1 L1 T1 G2 L2 T2).
(* Basic properties *********************************************************)
-lemma fqu_fqup: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄.
+lemma fqu_fqup: ∀b,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐[b] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ⊐+[b] ⦃G2, L2, T2⦄.
/2 width=1 by tri_inj/ qed.
-lemma fqup_strap1: ∀G1,G,G2,L1,L,L2,T1,T,T2.
- ⦃G1, L1, T1⦄ ⊐+ ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐ ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄.
+lemma fqup_strap1: ∀b,G1,G,G2,L1,L,L2,T1,T,T2.
+ ⦃G1, L1, T1⦄ ⊐+[b] ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐[b] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ⊐+[b] ⦃G2, L2, T2⦄.
/2 width=5 by tri_step/ qed.
-lemma fqup_strap2: ∀G1,G,G2,L1,L,L2,T1,T,T2.
- ⦃G1, L1, T1⦄ ⊐ ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄.
+lemma fqup_strap2: ∀b,G1,G,G2,L1,L,L2,T1,T,T2.
+ ⦃G1, L1, T1⦄ ⊐[b] ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐+[b] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ⊐+[b] ⦃G2, L2, T2⦄.
/2 width=5 by tri_TC_strap/ qed.
-lemma fqup_pair_sn: ∀I,G,L,V,T. ⦃G, L, ②{I}V.T⦄ ⊐+ ⦃G, L, V⦄.
+lemma fqup_pair_sn: ∀b,I,G,L,V,T. ⦃G, L, ②{I}V.T⦄ ⊐+[b] ⦃G, L, V⦄.
/2 width=1 by fqu_pair_sn, fqu_fqup/ qed.
-lemma fqup_bind_dx: ∀a,I,G,L,V,T. ⦃G, L, ⓑ{a,I}V.T⦄ ⊐+ ⦃G, L.ⓑ{I}V, T⦄.
+lemma fqup_bind_dx: ∀b,p,I,G,L,V,T. ⦃G, L, ⓑ{p,I}V.T⦄ ⊐+[b] ⦃G, L.ⓑ{I}V, T⦄.
/2 width=1 by fqu_bind_dx, fqu_fqup/ qed.
-lemma fqup_flat_dx: ∀I,G,L,V,T. ⦃G, L, ⓕ{I}V.T⦄ ⊐+ ⦃G, L, T⦄.
+lemma fqup_clear: ∀p,I,G,L,V,T. ⦃G, L, ⓑ{p,I}V.T⦄ ⊐+[Ⓕ] ⦃G, L.ⓧ, T⦄.
+/3 width=1 by fqu_clear, fqu_fqup/ qed.
+
+lemma fqup_flat_dx: ∀b,I,G,L,V,T. ⦃G, L, ⓕ{I}V.T⦄ ⊐+[b] ⦃G, L, T⦄.
/2 width=1 by fqu_flat_dx, fqu_fqup/ qed.
-lemma fqup_flat_dx_pair_sn: ∀I1,I2,G,L,V1,V2,T. ⦃G, L, ⓕ{I1}V1.②{I2}V2.T⦄ ⊐+ ⦃G, L, V2⦄.
+lemma fqup_flat_dx_pair_sn: ∀b,I1,I2,G,L,V1,V2,T. ⦃G, L, ⓕ{I1}V1.②{I2}V2.T⦄ ⊐+[b] ⦃G, L, V2⦄.
/2 width=5 by fqu_pair_sn, fqup_strap1/ qed.
-lemma fqup_bind_dx_flat_dx: ∀a,G,I1,I2,L,V1,V2,T. ⦃G, L, ⓑ{a,I1}V1.ⓕ{I2}V2.T⦄ ⊐+ ⦃G, L.ⓑ{I1}V1, T⦄.
+lemma fqup_bind_dx_flat_dx: ∀b,p,G,I1,I2,L,V1,V2,T. ⦃G, L, ⓑ{p,I1}V1.ⓕ{I2}V2.T⦄ ⊐+[b] ⦃G, L.ⓑ{I1}V1, T⦄.
/2 width=5 by fqu_flat_dx, fqup_strap1/ qed.
-lemma fqup_flat_dx_bind_dx: ∀a,I1,I2,G,L,V1,V2,T. ⦃G, L, ⓕ{I1}V1.ⓑ{a,I2}V2.T⦄ ⊐+ ⦃G, L.ⓑ{I2}V2, T⦄.
+lemma fqup_flat_dx_bind_dx: ∀b,p,I1,I2,G,L,V1,V2,T. ⦃G, L, ⓕ{I1}V1.ⓑ{p,I2}V2.T⦄ ⊐+[b] ⦃G, L.ⓑ{I2}V2, T⦄.
/2 width=5 by fqu_bind_dx, fqup_strap1/ qed.
(* Basic eliminators ********************************************************)
-lemma fqup_ind: ∀G1,L1,T1. ∀R:relation3 ….
- (∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → R G2 L2 T2) →
- (∀G,G2,L,L2,T,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐ ⦃G2, L2, T2⦄ → R G L T → R G2 L2 T2) →
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → R G2 L2 T2.
-#G1 #L1 #T1 #R #IH1 #IH2 #G2 #L2 #T2 #H
+lemma fqup_ind: ∀b,G1,L1,T1. ∀R:relation3 ….
+ (∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐[b] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
+ (∀G,G2,L,L2,T,T2. ⦃G1, L1, T1⦄ ⊐+[b] ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐[b] ⦃G2, L2, T2⦄ → R G L T → R G2 L2 T2) →
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐+[b] ⦃G2, L2, T2⦄ → R G2 L2 T2.
+#b #G1 #L1 #T1 #R #IH1 #IH2 #G2 #L2 #T2 #H
@(tri_TC_ind … IH1 IH2 G2 L2 T2 H)
qed-.
-lemma fqup_ind_dx: ∀G2,L2,T2. ∀R:relation3 ….
- (∀G1,L1,T1. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → R G1 L1 T1) →
- (∀G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ⊐ ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐+ ⦃G2, L2, T2⦄ → R G L T → R G1 L1 T1) →
- ∀G1,L1,T1. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → R G1 L1 T1.
-#G2 #L2 #T2 #R #IH1 #IH2 #G1 #L1 #T1 #H
+lemma fqup_ind_dx: ∀b,G2,L2,T2. ∀R:relation3 ….
+ (∀G1,L1,T1. ⦃G1, L1, T1⦄ ⊐[b] ⦃G2, L2, T2⦄ → R G1 L1 T1) →
+ (∀G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ⊐[b] ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐+[b] ⦃G2, L2, T2⦄ → R G L T → R G1 L1 T1) →
+ ∀G1,L1,T1. ⦃G1, L1, T1⦄ ⊐+[b] ⦃G2, L2, T2⦄ → R G1 L1 T1.
+#b #G2 #L2 #T2 #R #IH1 #IH2 #G1 #L1 #T1 #H
@(tri_TC_ind_dx … IH1 IH2 G1 L1 T1 H)
qed-.
(* Properties with generic slicing for local environments *******************)
-lemma fqup_drops_succ: ∀G,K,T,l,L,U. ⬇*[⫯l] L ≡ K → ⬆*[⫯l] T ≡ U →
- ⦃G, L, U⦄ ⊐+ ⦃G, K, T⦄.
-#G #K #T #l elim l -l
+lemma fqup_drops_succ: ∀b,G,K,T,i,L,U. ⬇*[⫯i] L ≡ K → ⬆*[⫯i] T ≡ U →
+ ⦃G, L, U⦄ ⊐+[b] ⦃G, K, T⦄.
+#b #G #K #T #i elim i -i
[ #L #U #HLK #HTU elim (drops_inv_succ … HLK) -HLK
- #I #Y #V #HY #H destruct <(drops_fwd_isid … HY) -K
- /3 width=1 by fqu_fqup, fqu_drop/
+ #I #Y #HY #H destruct <(drops_fwd_isid … HY) -K //
+ /3 width=2 by fqu_fqup, fqu_drop/
| #l #IH #L #U #HLK #HTU elim (drops_inv_succ … HLK) -HLK
- #I #Y #V #HY #H destruct
+ #I #Y #HY #H destruct
elim (lifts_split_trans … HTU … (𝐔❴⫯l❵) (𝐔❴1❵)) -HTU
/4 width=5 by fqup_strap2, fqu_drop/
]
qed.
-lemma fqup_drops_strap1: ∀G1,G2,L1,K1,K2,T1,T2,U1,l. ⬇*[l] L1 ≡ K1 → ⬆*[l] T1 ≡ U1 →
- ⦃G1, K1, T1⦄ ⊐ ⦃G2, K2, T2⦄ → ⦃G1, L1, U1⦄ ⊐+ ⦃G2, K2, T2⦄.
-#G1 #G2 #L1 #K1 #K2 #T1 #T2 #U1 *
+lemma fqup_drops_strap1: ∀b,G1,G2,L1,K1,K2,T1,T2,U1,i. ⬇*[i] L1 ≡ K1 → ⬆*[i] T1 ≡ U1 →
+ ⦃G1, K1, T1⦄ ⊐[b] ⦃G2, K2, T2⦄ → ⦃G1, L1, U1⦄ ⊐+[b] ⦃G2, K2, T2⦄.
+#b #G1 #G2 #L1 #K1 #K2 #T1 #T2 #U1 *
[ #HLK1 #HTU1 #HT12
>(drops_fwd_isid … HLK1) -L1 //
<(lifts_fwd_isid … HTU1) -U1 /2 width=1 by fqu_fqup/
]
qed-.
-lemma fqup_lref: ∀I,G,L,K,V,i. ⬇*[i] L ≡ K.ⓑ{I}V → ⦃G, L, #i⦄ ⊐+ ⦃G, K, V⦄.
+lemma fqup_lref: ∀b,I,G,L,K,V,i. ⬇*[i] L ≡ K.ⓑ{I}V → ⦃G, L, #i⦄ ⊐+[b] ⦃G, K, V⦄.
/2 width=6 by fqup_drops_strap1/ qed.
(* Main properties **********************************************************)
-theorem fqup_trans: tri_transitive … fqup.
+theorem fqup_trans: ∀b. tri_transitive … (fqup b).
/2 width=5 by tri_TC_transitive/ qed-.
(* Forward lemmas with weight for closures **********************************)
-lemma fqup_fwd_fw: ∀G1,G2,L1,L2,T1,T2.
- â¦\83G1, L1, T1â¦\84 â\8a\90+ â¦\83G2, L2, T2â¦\84 â\86\92 â\99¯{G2, L2, T2} < â\99¯{G1, L1, T1}.
-#G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
+lemma fqup_fwd_fw: ∀b,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+[b] ⦃G2, L2, T2⦄ →
+ ♯{G2, L2, T2} < ♯{G1, L1, T1}.
+#b #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
/3 width=3 by fqu_fwd_fw, transitive_lt/
qed-.
(* Advanced eliminators *****************************************************)
-lemma fqup_wf_ind: ∀R:relation3 …. (
- ∀G1,L1,T1. (∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → R G2 L2 T2) →
+lemma fqup_wf_ind: ∀b. ∀R:relation3 …. (
+ ∀G1,L1,T1. (∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐+[b] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
R G1 L1 T1
) → ∀G1,L1,T1. R G1 L1 T1.
-#R #HR @(f3_ind … fw) #x #IHx #G1 #L1 #T1 #H destruct /4 width=1 by fqup_fwd_fw/
+#b #R #HR @(f3_ind … fw) #x #IHx #G1 #L1 #T1 #H destruct
+/4 width=2 by fqup_fwd_fw/
qed-.
-lemma fqup_wf_ind_eq: ∀R:relation3 …. (
- ∀G1,L1,T1. (∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → R G2 L2 T2) →
+lemma fqup_wf_ind_eq: ∀b. ∀R:relation3 …. (
+ ∀G1,L1,T1. (∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐+[b] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
∀G2,L2,T2. G1 = G2 → L1 = L2 → T1 = T2 → R G2 L2 T2
) → ∀G1,L1,T1. R G1 L1 T1.
-#R #HR @(f3_ind … fw) #x #IHx #G1 #L1 #T1 #H destruct /4 width=7 by fqup_fwd_fw/
+#b #R #HR @(f3_ind … fw) #x #IHx #G1 #L1 #T1 #H destruct
+/4 width=7 by fqup_fwd_fw/
qed-.
include "ground_2/lib/star.ma".
include "basic_2/notation/relations/suptermstar_6.ma".
+include "basic_2/notation/relations/suptermstar_7.ma".
include "basic_2/s_transition/fquq.ma".
(* STAR-ITERATED SUPCLOSURE *************************************************)
-definition fqus: tri_relation genv lenv term ≝ tri_TC … fquq.
+definition fqus: bool → tri_relation genv lenv term ≝
+ λb. tri_TC … (fquq b).
+
+interpretation "extended star-iterated structural successor (closure)"
+ 'SupTermStar b G1 L1 T1 G2 L2 T2 = (fqus b G1 L1 T1 G2 L2 T2).
interpretation "star-iterated structural successor (closure)"
- 'SupTermStar G1 L1 T1 G2 L2 T2 = (fqus G1 L1 T1 G2 L2 T2).
+ 'SupTermStar G1 L1 T1 G2 L2 T2 = (fqus true G1 L1 T1 G2 L2 T2).
(* Basic eliminators ********************************************************)
-lemma fqus_ind: ∀G1,L1,T1. ∀R:relation3 …. R G1 L1 T1 →
- (∀G,G2,L,L2,T,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐⸮ ⦃G2, L2, T2⦄ → R G L T → R G2 L2 T2) →
- ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ → R G2 L2 T2.
-#G1 #L1 #T1 #R #IH1 #IH2 #G2 #L2 #T2 #H
+lemma fqus_ind: ∀b,G1,L1,T1. ∀R:relation3 …. R G1 L1 T1 →
+ (∀G,G2,L,L2,T,T2. ⦃G1, L1, T1⦄ ⊐*[b] ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐⸮[b] ⦃G2, L2, T2⦄ → R G L T → R G2 L2 T2) →
+ ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐*[b] ⦃G2, L2, T2⦄ → R G2 L2 T2.
+#b #G1 #L1 #T1 #R #IH1 #IH2 #G2 #L2 #T2 #H
@(tri_TC_star_ind … IH1 IH2 G2 L2 T2 H) //
qed-.
-lemma fqus_ind_dx: ∀G2,L2,T2. ∀R:relation3 …. R G2 L2 T2 →
- (∀G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐* ⦃G2, L2, T2⦄ → R G L T → R G1 L1 T1) →
- ∀G1,L1,T1. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ → R G1 L1 T1.
-#G2 #L2 #T2 #R #IH1 #IH2 #G1 #L1 #T1 #H
+lemma fqus_ind_dx: ∀b,G2,L2,T2. ∀R:relation3 …. R G2 L2 T2 →
+ (∀G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ⊐⸮[b] ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐*[b] ⦃G2, L2, T2⦄ → R G L T → R G1 L1 T1) →
+ ∀G1,L1,T1. ⦃G1, L1, T1⦄ ⊐*[b] ⦃G2, L2, T2⦄ → R G1 L1 T1.
+#b #G2 #L2 #T2 #R #IH1 #IH2 #G1 #L1 #T1 #H
@(tri_TC_star_ind_dx … IH1 IH2 G1 L1 T1 H) //
qed-.
(* Basic properties *********************************************************)
-lemma fqus_refl: tri_reflexive … fqus.
+lemma fqus_refl: ∀b. tri_reflexive … (fqus b).
/2 width=1 by tri_inj/ qed.
-lemma fquq_fqus: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄.
+lemma fquq_fqus: ∀b,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮[b] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ⊐*[b] ⦃G2, L2, T2⦄.
/2 width=1 by tri_inj/ qed.
-lemma fqus_strap1: ∀G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐⸮ ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄.
+lemma fqus_strap1: ∀b,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐*[b] ⦃G, L, T⦄ →
+ ⦃G, L, T⦄ ⊐⸮[b] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐*[b] ⦃G2, L2, T2⦄.
/2 width=5 by tri_step/ qed-.
-lemma fqus_strap2: ∀G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐* ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄.
+lemma fqus_strap2: ∀b,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐⸮[b] ⦃G, L, T⦄ →
+ ⦃G, L, T⦄ ⊐*[b] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐*[b] ⦃G2, L2, T2⦄.
/2 width=5 by tri_TC_strap/ qed-.
(* Basic inversion lemmas ***************************************************)
-lemma fqus_inv_fqu_sn: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+lemma fqus_inv_fqu_sn: ∀b,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐*[b] ⦃G2, L2, T2⦄ →
(∧∧ G1 = G2 & L1 = L2 & T1 = T2) ∨
- ∃∃G,L,T. ⦃G1, L1, T1⦄ ⊐ ⦃G, L, T⦄ & ⦃G, L, T⦄ ⊐* ⦃G2, L2, T2⦄.
-#G1 #G2 #L1 #L2 #T1 #T2 #H12 @(fqus_ind_dx … H12) -G1 -L1 -T1 /3 width=1 by and3_intro, or_introl/
+ ∃∃G,L,T. ⦃G1, L1, T1⦄ ⊐[b] ⦃G, L, T⦄ & ⦃G, L, T⦄ ⊐*[b] ⦃G2, L2, T2⦄.
+#b #G1 #G2 #L1 #L2 #T1 #T2 #H12 @(fqus_ind_dx … H12) -G1 -L1 -T1 /3 width=1 by and3_intro, or_introl/
#G1 #G #L1 #L #T1 #T * /3 width=5 by ex2_3_intro, or_intror/
* #HG #HL #HT #_ destruct //
qed-.
-lemma fqus_inv_sort1: ∀G1,G2,L1,L2,T2,s. ⦃G1, L1, ⋆s⦄ ⊐* ⦃G2, L2, T2⦄ →
+lemma fqus_inv_sort1: ∀b,G1,G2,L1,L2,T2,s. ⦃G1, L1, ⋆s⦄ ⊐*[b] ⦃G2, L2, T2⦄ →
(∧∧ G1 = G2 & L1 = L2 & ⋆s = T2) ∨
- ∃∃J,L,V. ⦃G1, L, ⋆s⦄ ⊐* ⦃G2, L2, T2⦄ & L1 = L.ⓑ{J}V.
-#G1 #G2 #L1 #L2 #T2 #s #H elim (fqus_inv_fqu_sn … H) -H * /3 width=1 by and3_intro, or_introl/
-#G #L #T #H elim (fqu_inv_sort1 … H) -H /3 width=5 by ex2_3_intro, or_intror/
+ ∃∃J,L. ⦃G1, L, ⋆s⦄ ⊐*[b] ⦃G2, L2, T2⦄ & L1 = L.ⓘ{J}.
+#b #G1 #G2 #L1 #L2 #T2 #s #H elim (fqus_inv_fqu_sn … H) -H * /3 width=1 by and3_intro, or_introl/
+#G #L #T #H elim (fqu_inv_sort1 … H) -H /3 width=4 by ex2_2_intro, or_intror/
qed-.
-lemma fqus_inv_lref1: ∀G1,G2,L1,L2,T2,i. ⦃G1, L1, #i⦄ ⊐* ⦃G2, L2, T2⦄ →
+lemma fqus_inv_lref1: ∀b,G1,G2,L1,L2,T2,i. ⦃G1, L1, #i⦄ ⊐*[b] ⦃G2, L2, T2⦄ →
∨∨ ∧∧ G1 = G2 & L1 = L2 & #i = T2
- | ∃∃J,L,V. ⦃G1, L, V⦄ ⊐* ⦃G2, L2, T2⦄ & L1 = L.ⓑ{J}V & i = 0
- | ∃∃J,L,V,j. ⦃G1, L, #j⦄ ⊐* ⦃G2, L2, T2⦄ & L1 = L.ⓑ{J}V & i = ⫯j.
-#G1 #G2 #L1 #L2 #T2 #i #H elim (fqus_inv_fqu_sn … H) -H * /3 width=1 by and3_intro, or3_intro0/
+ | ∃∃J,L,V. ⦃G1, L, V⦄ ⊐*[b] ⦃G2, L2, T2⦄ & L1 = L.ⓑ{J}V & i = 0
+ | ∃∃J,L,j. ⦃G1, L, #j⦄ ⊐*[b] ⦃G2, L2, T2⦄ & L1 = L.ⓘ{J} & i = ⫯j.
+#b #G1 #G2 #L1 #L2 #T2 #i #H elim (fqus_inv_fqu_sn … H) -H * /3 width=1 by and3_intro, or3_intro0/
#G #L #T #H elim (fqu_inv_lref1 … H) -H * /3 width=7 by or3_intro1, or3_intro2, ex3_4_intro, ex3_3_intro/
qed-.
-lemma fqus_inv_gref1: ∀G1,G2,L1,L2,T2,l. ⦃G1, L1, §l⦄ ⊐* ⦃G2, L2, T2⦄ →
+lemma fqus_inv_gref1: ∀b,G1,G2,L1,L2,T2,l. ⦃G1, L1, §l⦄ ⊐*[b] ⦃G2, L2, T2⦄ →
(∧∧ G1 = G2 & L1 = L2 & §l = T2) ∨
- ∃∃J,L,V. ⦃G1, L, §l⦄ ⊐* ⦃G2, L2, T2⦄ & L1 = L.ⓑ{J}V.
-#G1 #G2 #L1 #L2 #T2 #l #H elim (fqus_inv_fqu_sn … H) -H * /3 width=1 by and3_intro, or_introl/
-#G #L #T #H elim (fqu_inv_gref1 … H) -H /3 width=5 by ex2_3_intro, or_intror/
+ ∃∃J,L. ⦃G1, L, §l⦄ ⊐*[b] ⦃G2, L2, T2⦄ & L1 = L.ⓘ{J}.
+#b #G1 #G2 #L1 #L2 #T2 #l #H elim (fqus_inv_fqu_sn … H) -H * /3 width=1 by and3_intro, or_introl/
+#G #L #T #H elim (fqu_inv_gref1 … H) -H /3 width=4 by ex2_2_intro, or_intror/
qed-.
-lemma fqus_inv_bind1: ∀p,I,G1,G2,L1,L2,V1,T1,T2. ⦃G1, L1, ⓑ{p,I}V1.T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+lemma fqus_inv_bind1: ∀b,p,I,G1,G2,L1,L2,V1,T1,T2. ⦃G1, L1, ⓑ{p,I}V1.T1⦄ ⊐*[b] ⦃G2, L2, T2⦄ →
∨∨ ∧∧ G1 = G2 & L1 = L2 & ⓑ{p,I}V1.T1 = T2
- | ⦃G1, L1, V1⦄ ⊐* ⦃G2, L2, T2⦄
- | ⦃G1, L1.ⓑ{I}V1, T1⦄ ⊐* ⦃G2, L2, T2⦄
- | ∃∃J,L,V,T. ⦃G1, L, T⦄ ⊐* ⦃G2, L2, T2⦄ & ⬆*[1] T ≡ ⓑ{p,I}V1.T1 & L1 = L.ⓑ{J}V.
-#p #I #G1 #G2 #L1 #L2 #V1 #T1 #T2 #H elim (fqus_inv_fqu_sn … H) -H * /3 width=1 by and3_intro, or4_intro0/
+ | ⦃G1, L1, V1⦄ ⊐*[b] ⦃G2, L2, T2⦄
+ | ⦃G1, L1.ⓑ{I}V1, T1⦄ ⊐*[b] ⦃G2, L2, T2⦄
+ | ⦃G1, L1.ⓧ, T1⦄ ⊐*[b] ⦃G2, L2, T2⦄ ∧ b = Ⓕ
+ | ∃∃J,L,T. ⦃G1, L, T⦄ ⊐*[b] ⦃G2, L2, T2⦄ & ⬆*[1] T ≡ ⓑ{p,I}V1.T1 & L1 = L.ⓘ{J}.
+#b #p #I #G1 #G2 #L1 #L2 #V1 #T1 #T2 #H elim (fqus_inv_fqu_sn … H) -H * /3 width=1 by and3_intro, or5_intro0/
#G #L #T #H elim (fqu_inv_bind1 … H) -H *
-[3: #J #V ] #H1 #H2 #H3 #H destruct
-/3 width=7 by or4_intro1, or4_intro2, or4_intro3, ex3_4_intro/
+[4: #J ] #H1 #H2 #H3 [4: #Hb ] #H destruct
+/3 width=6 by or5_intro1, or5_intro2, or5_intro3, or5_intro4, ex3_3_intro, conj/
+qed-.
+
+
+lemma fqus_inv_bind1_true: ∀p,I,G1,G2,L1,L2,V1,T1,T2. ⦃G1, L1, ⓑ{p,I}V1.T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+ ∨∨ ∧∧ G1 = G2 & L1 = L2 & ⓑ{p,I}V1.T1 = T2
+ | ⦃G1, L1, V1⦄ ⊐* ⦃G2, L2, T2⦄
+ | ⦃G1, L1.ⓑ{I}V1, T1⦄ ⊐* ⦃G2, L2, T2⦄
+ | ∃∃J,L,T. ⦃G1, L, T⦄ ⊐* ⦃G2, L2, T2⦄ & ⬆*[1] T ≡ ⓑ{p,I}V1.T1 & L1 = L.ⓘ{J}.
+#p #I #G1 #G2 #L1 #L2 #V1 #T1 #T2 #H elim (fqus_inv_bind1 … H) -H [1,4: * ]
+/3 width=1 by and3_intro, or4_intro0, or4_intro1, or4_intro2, or4_intro3, ex3_3_intro/
+#_ #H destruct
qed-.
-lemma fqus_inv_flat1: ∀I,G1,G2,L1,L2,V1,T1,T2. ⦃G1, L1, ⓕ{I}V1.T1⦄ ⊐* ⦃G2, L2, T2⦄ →
+lemma fqus_inv_flat1: ∀b,I,G1,G2,L1,L2,V1,T1,T2. ⦃G1, L1, ⓕ{I}V1.T1⦄ ⊐*[b] ⦃G2, L2, T2⦄ →
∨∨ ∧∧ G1 = G2 & L1 = L2 & ⓕ{I}V1.T1 = T2
- | ⦃G1, L1, V1⦄ ⊐* ⦃G2, L2, T2⦄
- | ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄
- | ∃∃J,L,V,T. ⦃G1, L, T⦄ ⊐* ⦃G2, L2, T2⦄ & ⬆*[1] T ≡ ⓕ{I}V1.T1 & L1 = L.ⓑ{J}V.
-#I #G1 #G2 #L1 #L2 #V1 #T1 #T2 #H elim (fqus_inv_fqu_sn … H) -H * /3 width=1 by and3_intro, or4_intro0/
+ | ⦃G1, L1, V1⦄ ⊐*[b] ⦃G2, L2, T2⦄
+ | ⦃G1, L1, T1⦄ ⊐*[b] ⦃G2, L2, T2⦄
+ | ∃∃J,L,T. ⦃G1, L, T⦄ ⊐*[b] ⦃G2, L2, T2⦄ & ⬆*[1] T ≡ ⓕ{I}V1.T1 & L1 = L.ⓘ{J}.
+#b #I #G1 #G2 #L1 #L2 #V1 #T1 #T2 #H elim (fqus_inv_fqu_sn … H) -H * /3 width=1 by and3_intro, or4_intro0/
#G #L #T #H elim (fqu_inv_flat1 … H) -H *
-[3: #J #V ] #H1 #H2 #H3 #H destruct
-/3 width=7 by or4_intro1, or4_intro2, or4_intro3, ex3_4_intro/
+[3: #J ] #H1 #H2 #H3 #H destruct
+/3 width=6 by or4_intro1, or4_intro2, or4_intro3, ex3_3_intro/
qed-.
(* Advanced inversion lemmas ************************************************)
-lemma fqus_inv_atom1: ∀I,G1,G2,L2,T2. ⦃G1, ⋆, ⓪{I}⦄ ⊐* ⦃G2, L2, T2⦄ →
+lemma fqus_inv_atom1: ∀b,I,G1,G2,L2,T2. ⦃G1, ⋆, ⓪{I}⦄ ⊐*[b] ⦃G2, L2, T2⦄ →
∧∧ G1 = G2 & ⋆ = L2 & ⓪{I} = T2.
-#I #G1 #G2 #L2 #T2 #H elim (fqus_inv_fqu_sn … H) -H * /2 width=1 by and3_intro/
+#b #I #G1 #G2 #L2 #T2 #H elim (fqus_inv_fqu_sn … H) -H * /2 width=1 by and3_intro/
#G #L #T #H elim (fqu_inv_atom1 … H)
qed-.
-lemma fqus_inv_sort1_pair: ∀I,G1,G2,L1,L2,V1,T2,s. ⦃G1, L1.ⓑ{I}V1, ⋆s⦄ ⊐* ⦃G2, L2, T2⦄ →
- (â\88§â\88§ G1 = G2 & L1.â\93\91{I}V1 = L2 & â\8b\86s = T2) â\88¨ â¦\83G1, L1, â\8b\86sâ¦\84 â\8a\90* ⦃G2, L2, T2⦄.
-#I #G1 #G2 #L1 #L2 #V #T2 #s #H elim (fqus_inv_fqu_sn … H) -H * /3 width=1 by and3_intro, or_introl/
-#G #L #T #H elim (fqu_inv_sort1_pair … H) -H
+lemma fqus_inv_sort1_bind: ∀b,I,G1,G2,L1,L2,T2,s. ⦃G1, L1.ⓘ{I}, ⋆s⦄ ⊐*[b] ⦃G2, L2, T2⦄ →
+ (â\88§â\88§ G1 = G2 & L1.â\93\98{I} = L2 & â\8b\86s = T2) â\88¨ â¦\83G1, L1, â\8b\86sâ¦\84 â\8a\90*[b] ⦃G2, L2, T2⦄.
+#b #I #G1 #G2 #L1 #L2 #T2 #s #H elim (fqus_inv_fqu_sn … H) -H * /3 width=1 by and3_intro, or_introl/
+#G #L #T #H elim (fqu_inv_sort1_bind … H) -H
#H1 #H2 #H3 #H destruct /2 width=1 by or_intror/
qed-.
-lemma fqus_inv_zero1_pair: ∀I,G1,G2,L1,L2,V1,T2. ⦃G1, L1.ⓑ{I}V1, #0⦄ ⊐* ⦃G2, L2, T2⦄ →
- (∧∧ G1 = G2 & L1.ⓑ{I}V1 = L2 & #0 = T2) ∨ ⦃G1, L1, V1⦄ ⊐* ⦃G2, L2, T2⦄.
-#I #G1 #G2 #L1 #L2 #V1 #T2 #H elim (fqus_inv_fqu_sn … H) -H * /3 width=1 by and3_intro, or_introl/
+lemma fqus_inv_zero1_pair: ∀b,I,G1,G2,L1,L2,V1,T2. ⦃G1, L1.ⓑ{I}V1, #0⦄ ⊐*[b] ⦃G2, L2, T2⦄ →
+ (∧∧ G1 = G2 & L1.ⓑ{I}V1 = L2 & #0 = T2) ∨ ⦃G1, L1, V1⦄ ⊐*[b] ⦃G2, L2, T2⦄.
+#b #I #G1 #G2 #L1 #L2 #V1 #T2 #H elim (fqus_inv_fqu_sn … H) -H * /3 width=1 by and3_intro, or_introl/
#G #L #T #H elim (fqu_inv_zero1_pair … H) -H
#H1 #H2 #H3 #H destruct /2 width=1 by or_intror/
qed-.
-lemma fqus_inv_lref1_pair: ∀I,G1,G2,L1,L2,V1,T2,i. ⦃G1, L1.ⓑ{I}V1, #⫯i⦄ ⊐* ⦃G2, L2, T2⦄ →
- (â\88§â\88§ G1 = G2 & L1.â\93\91{I}V1 = L2 & #(⫯i) = T2) â\88¨ â¦\83G1, L1, #iâ¦\84 â\8a\90* ⦃G2, L2, T2⦄.
-#I #G1 #G2 #L1 #L2 #V #T2 #i #H elim (fqus_inv_fqu_sn … H) -H * /3 width=1 by and3_intro, or_introl/
-#G #L #T #H elim (fqu_inv_lref1_pair … H) -H
+lemma fqus_inv_lref1_bind: ∀b,I,G1,G2,L1,L2,T2,i. ⦃G1, L1.ⓘ{I}, #⫯i⦄ ⊐*[b] ⦃G2, L2, T2⦄ →
+ (â\88§â\88§ G1 = G2 & L1.â\93\98{I} = L2 & #(⫯i) = T2) â\88¨ â¦\83G1, L1, #iâ¦\84 â\8a\90*[b] ⦃G2, L2, T2⦄.
+#b #I #G1 #G2 #L1 #L2 #T2 #i #H elim (fqus_inv_fqu_sn … H) -H * /3 width=1 by and3_intro, or_introl/
+#G #L #T #H elim (fqu_inv_lref1_bind … H) -H
#H1 #H2 #H3 #H destruct /2 width=1 by or_intror/
qed-.
-lemma fqus_inv_gref1_pair: ∀I,G1,G2,L1,L2,V1,T2,l. ⦃G1, L1.ⓑ{I}V1, §l⦄ ⊐* ⦃G2, L2, T2⦄ →
- (â\88§â\88§ G1 = G2 & L1.â\93\91{I}V1 = L2 & §l = T2) â\88¨ â¦\83G1, L1, §lâ¦\84 â\8a\90* ⦃G2, L2, T2⦄.
-#I #G1 #G2 #L1 #L2 #V #T2 #l #H elim (fqus_inv_fqu_sn … H) -H * /3 width=1 by and3_intro, or_introl/
-#G #L #T #H elim (fqu_inv_gref1_pair … H) -H
+lemma fqus_inv_gref1_bind: ∀b,I,G1,G2,L1,L2,T2,l. ⦃G1, L1.ⓘ{I}, §l⦄ ⊐*[b] ⦃G2, L2, T2⦄ →
+ (â\88§â\88§ G1 = G2 & L1.â\93\98{I} = L2 & §l = T2) â\88¨ â¦\83G1, L1, §lâ¦\84 â\8a\90*[b] ⦃G2, L2, T2⦄.
+#b #I #G1 #G2 #L1 #L2 #T2 #l #H elim (fqus_inv_fqu_sn … H) -H * /3 width=1 by and3_intro, or_introl/
+#G #L #T #H elim (fqu_inv_gref1_bind … H) -H
#H1 #H2 #H3 #H destruct /2 width=1 by or_intror/
qed-.
(* Properties with generic slicing for local environments *******************)
-lemma fqus_drops: ∀G,L,K,T,U,l. ⬇*[l] L ≡ K → ⬆*[l] T ≡ U →
- ⦃G, L, U⦄ ⊐* ⦃G, K, T⦄.
-#G #L #K #T #U * /3 width=3 by fqup_drops_succ, fqup_fqus/
+lemma fqus_drops: ∀b,G,L,K,T,U,i. ⬇*[i] L ≡ K → ⬆*[i] T ≡ U →
+ ⦃G, L, U⦄ ⊐*[b] ⦃G, K, T⦄.
+#b #G #L #K #T #U * /3 width=3 by fqup_drops_succ, fqup_fqus/
#HLK #HTU <(lifts_fwd_isid … HTU) -U // <(drops_fwd_isid … HLK) -K //
qed.
(* Alternative definition with plus-iterated supclosure *********************)
-lemma fqup_fqus: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄.
-#G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
+lemma fqup_fqus: ∀b,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+[b] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐*[b] ⦃G2, L2, T2⦄.
+#b #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2
/3 width=5 by fqus_strap1, fquq_fqus, fqu_fquq/
qed.
(* Basic_2A1: was: fqus_inv_gen *)
-lemma fqus_inv_fqup: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ ∨ (∧∧ G1 = G2 & L1 = L2 & T1 = T2).
-#G1 #G2 #L1 #L2 #T1 #T2 #H @(fqus_ind … H) -G2 -L2 -T2 //
+lemma fqus_inv_fqup: ∀b,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐*[b] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ⊐+[b] ⦃G2, L2, T2⦄ ∨ (∧∧ G1 = G2 & L1 = L2 & T1 = T2).
+#b #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqus_ind … H) -G2 -L2 -T2 //
#G #G2 #L #L2 #T #T2 #_ *
[ #H2 * /3 width=5 by fqup_strap1, or_introl/
* /3 width=1 by fqu_fqup, or_introl/
-| * #HG #HL #HT destruct * /2 width=1 by or_introl/
- * /2 width=1 by or_intror/
+| * #HG #HL #HT destruct //
]
qed-.
(* Advanced properties ******************************************************)
-lemma fqus_strap1_fqu: ∀G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐ ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄.
-#G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H2 elim (fqus_inv_fqup … H1) -H1
+lemma fqus_strap1_fqu: ∀b,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐*[b] ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐[b] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ⊐+[b] ⦃G2, L2, T2⦄.
+#b #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H2 elim (fqus_inv_fqup … H1) -H1
[ /2 width=5 by fqup_strap1/
| * /2 width=1 by fqu_fqup/
]
qed-.
-lemma fqus_strap2_fqu: ∀G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐* ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄.
-#G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H2 elim (fqus_inv_fqup … H2) -H2
+lemma fqus_strap2_fqu: ∀b,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐[b] ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐*[b] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ⊐+[b] ⦃G2, L2, T2⦄.
+#b #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H2 elim (fqus_inv_fqup … H2) -H2
[ /2 width=5 by fqup_strap2/
| * /2 width=1 by fqu_fqup/
]
qed-.
-lemma fqus_fqup_trans: ∀G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄.
-#G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H2 @(fqup_ind … H2) -H2 -G2 -L2 -T2
+lemma fqus_fqup_trans: ∀b,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐*[b] ⦃G, L, T⦄ → ⦃G, L, T⦄ ⊐+[b] ⦃G2, L2, T2⦄ →
+ ⦃G1, L1, T1⦄ ⊐+[b] ⦃G2, L2, T2⦄.
+#b #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H2 @(fqup_ind … H2) -H2 -G2 -L2 -T2
/2 width=5 by fqus_strap1_fqu, fqup_strap1/
qed-.
-lemma fqup_fqus_trans: ∀G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G, L, T⦄ →
- ⦃G, L, T⦄ ⊐* ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄.
-#G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 @(fqup_ind_dx … H1) -H1 -G1 -L1 -T1
+lemma fqup_fqus_trans: ∀b,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊐+[b] ⦃G, L, T⦄ →
+ ⦃G, L, T⦄ ⊐*[b] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐+[b] ⦃G2, L2, T2⦄.
+#b #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 @(fqup_ind_dx … H1) -H1 -G1 -L1 -T1
/3 width=5 by fqus_strap2_fqu, fqup_strap2/
qed-.
(* Advanced inversion lemmas for plus-iterated supclosure *******************)
-lemma fqup_inv_step_sn: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ →
- ∃∃G,L,T. ⦃G1, L1, T1⦄ ⊐ ⦃G, L, T⦄ & ⦃G, L, T⦄ ⊐* ⦃G2, L2, T2⦄.
-#G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind_dx … H) -G1 -L1 -T1 /2 width=5 by ex2_3_intro/
+lemma fqup_inv_step_sn: ∀b,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+[b] ⦃G2, L2, T2⦄ →
+ ∃∃G,L,T. ⦃G1, L1, T1⦄ ⊐[b] ⦃G, L, T⦄ & ⦃G, L, T⦄ ⊐*[b] ⦃G2, L2, T2⦄.
+#b #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind_dx … H) -G1 -L1 -T1 /2 width=5 by ex2_3_intro/
#G1 #G #L1 #L #T1 #T #H1 #_ * /4 width=9 by fqus_strap2, fqu_fquq, ex2_3_intro/
qed-.
(* Main properties **********************************************************)
-theorem fqus_trans: tri_transitive … fqus.
+theorem fqus_trans: ∀b. tri_transitive … (fqus b).
/2 width=5 by tri_TC_transitive/ qed-.
(* Forward lemmas with weight for closures **********************************)
-lemma fqus_fwd_fw: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ → ♯{G2, L2, T2} ≤ ♯{G1, L1, T1}.
-#G1 #G2 #L1 #L2 #T1 #T2 #H @(fqus_ind … H) -L2 -T2
+lemma fqus_fwd_fw: ∀b,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐*[b] ⦃G2, L2, T2⦄ →
+ ♯{G2, L2, T2} ≤ ♯{G1, L1, T1}.
+#b #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqus_ind … H) -L2 -T2
/3 width=3 by fquq_fwd_fw, transitive_le/
qed-.
(* Advanced inversion lemmas ************************************************)
-lemma fqus_inv_refl_atom3: ∀I,G,L,X. ⦃G, L, ⓪{I}⦄ ⊐* ⦃G, L, X⦄ → ⓪{I} = X.
-#I #G #L #X #H elim (fqus_inv_fqu_sn … H) -H * //
+lemma fqus_inv_refl_atom3: ∀b,I,G,L,X. ⦃G, L, ⓪{I}⦄ ⊐*[b] ⦃G, L, X⦄ → ⓪{I} = X.
+#b #I #G #L #X #H elim (fqus_inv_fqu_sn … H) -H * //
#G0 #L0 #T0 #H1 #H2 lapply (fqu_fwd_fw … H1) lapply (fqus_fwd_fw … H2) -H2 -H1
#H2 #H1 lapply (le_to_lt_to_lt … H2 H1) -G0 -L0 -T0
#H elim (lt_le_false … H) -H /2 width=1 by monotonic_le_plus_r/
(**************************************************************************)
include "basic_2/notation/relations/supterm_6.ma".
+include "basic_2/notation/relations/supterm_7.ma".
include "basic_2/syntax/lenv.ma".
include "basic_2/syntax/genv.ma".
include "basic_2/relocation/lifts.ma".
fqu_cpx_trans requires fqu_drop for all terms
frees_fqus_drops requires fqu_drop restricted on atoms
*)
-inductive fqu: tri_relation genv lenv term ≝
-| fqu_lref_O : ∀I,G,L,V. fqu G (L.ⓑ{I}V) (#0) G L V
-| fqu_pair_sn: ∀I,G,L,V,T. fqu G L (②{I}V.T) G L V
-| fqu_bind_dx: ∀p,I,G,L,V,T. fqu G L (ⓑ{p,I}V.T) G (L.ⓑ{I}V) T
-| fqu_flat_dx: ∀I,G,L,V,T. fqu G L (ⓕ{I}V.T) G L T
-| fqu_drop : ∀I,G,L,V,T,U. ⬆*[1] T ≡ U → fqu G (L.ⓑ{I}V) U G L T
+inductive fqu (b:bool): tri_relation genv lenv term ≝
+| fqu_lref_O : ∀I,G,L,V. fqu b G (L.ⓑ{I}V) (#0) G L V
+| fqu_pair_sn: ∀I,G,L,V,T. fqu b G L (②{I}V.T) G L V
+| fqu_bind_dx: ∀p,I,G,L,V,T. fqu b G L (ⓑ{p,I}V.T) G (L.ⓑ{I}V) T
+| fqu_clear : ∀p,I,G,L,V,T. b = Ⓕ → fqu b G L (ⓑ{p,I}V.T) G (L.ⓧ) T
+| fqu_flat_dx: ∀I,G,L,V,T. fqu b G L (ⓕ{I}V.T) G L T
+| fqu_drop : ∀I,G,L,T,U. ⬆*[1] T ≡ U → fqu b G (L.ⓘ{I}) U G L T
.
+interpretation
+ "extended structural successor (closure)"
+ 'SupTerm b G1 L1 T1 G2 L2 T2 = (fqu b G1 L1 T1 G2 L2 T2).
+
interpretation
"structural successor (closure)"
- 'SupTerm G1 L1 T1 G2 L2 T2 = (fqu G1 L1 T1 G2 L2 T2).
+ 'SupTerm G1 L1 T1 G2 L2 T2 = (fqu true G1 L1 T1 G2 L2 T2).
(* Basic properties *********************************************************)
-lemma fqu_lref_S: ∀I,G,L,V,i. ⦃G, L.ⓑ{I}V, #⫯i⦄ ⊐ ⦃G, L, #i⦄.
+lemma fqu_lref_S: ∀b,I,G,L,V,i. ⦃G, L.ⓑ{I}V, #⫯i⦄ ⊐[b] ⦃G, L, #i⦄.
/2 width=1 by fqu_drop/ qed.
(* Basic inversion lemmas ***************************************************)
-fact fqu_inv_sort1_aux: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+fact fqu_inv_sort1_aux: ∀b,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐[b] ⦃G2, L2, T2⦄ →
∀s. T1 = ⋆s →
- ∃∃J,V. G1 = G2 & L1 = L2.ⓑ{J}V & T2 = ⋆s.
-#G1 #G2 #L1 #L2 #T1 #T2 * -G1 -G2 -L1 -L2 -T1 -T2
+ ∃∃J. G1 = G2 & L1 = L2.ⓘ{J} & T2 = ⋆s.
+#b #G1 #G2 #L1 #L2 #T1 #T2 * -G1 -G2 -L1 -L2 -T1 -T2
[ #I #G #L #T #s #H destruct
| #I #G #L #V #T #s #H destruct
| #p #I #G #L #V #T #s #H destruct
+| #p #I #G #L #V #T #_ #s #H destruct
| #I #G #L #V #T #s #H destruct
-| #I #G #L #V #T #U #HI12 #s #H destruct
- lapply (lifts_inv_sort2 … HI12) -HI12 /2 width=3 by ex3_2_intro/
+| #I #G #L #T #U #HI12 #s #H destruct
+ lapply (lifts_inv_sort2 … HI12) -HI12 /2 width=2 by ex3_intro/
]
qed-.
-lemma fqu_inv_sort1: ∀G1,G2,L1,L2,T2,s. ⦃G1, L1, ⋆s⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∃∃J,V. G1 = G2 & L1 = L2.ⓑ{J}V & T2 = ⋆s.
-/2 width=3 by fqu_inv_sort1_aux/ qed-.
+lemma fqu_inv_sort1: ∀b,G1,G2,L1,L2,T2,s. ⦃G1, L1, ⋆s⦄ ⊐[b] ⦃G2, L2, T2⦄ →
+ ∃∃J. G1 = G2 & L1 = L2.ⓘ{J} & T2 = ⋆s.
+/2 width=4 by fqu_inv_sort1_aux/ qed-.
-fact fqu_inv_lref1_aux: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+fact fqu_inv_lref1_aux: ∀b,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐[b] ⦃G2, L2, T2⦄ →
∀i. T1 = #i →
(∃∃J,V. G1 = G2 & L1 = L2.ⓑ{J}V & T2 = V & i = 0) ∨
- ∃∃J,V,j. G1 = G2 & L1 = L2.ⓑ{J}V & T2 = #j & i = ⫯j.
-#G1 #G2 #L1 #L2 #T1 #T2 * -G1 -G2 -L1 -L2 -T1 -T2
+ ∃∃J,j. G1 = G2 & L1 = L2.ⓘ{J} & T2 = #j & i = ⫯j.
+#b #G1 #G2 #L1 #L2 #T1 #T2 * -G1 -G2 -L1 -L2 -T1 -T2
[ #I #G #L #T #i #H destruct /3 width=4 by ex4_2_intro, or_introl/
| #I #G #L #V #T #i #H destruct
| #p #I #G #L #V #T #i #H destruct
+| #p #I #G #L #V #T #_ #i #H destruct
| #I #G #L #V #T #i #H destruct
-| #I #G #L #V #T #U #HI12 #i #H destruct
- elim (lifts_inv_lref2_uni … HI12) -HI12 /3 width=3 by ex4_3_intro, or_intror/
+| #I #G #L #T #U #HI12 #i #H destruct
+ elim (lifts_inv_lref2_uni … HI12) -HI12 /3 width=3 by ex4_2_intro, or_intror/
]
qed-.
-lemma fqu_inv_lref1: ∀G1,G2,L1,L2,T2,i. ⦃G1, L1, #i⦄ ⊐ ⦃G2, L2, T2⦄ →
+lemma fqu_inv_lref1: ∀b,G1,G2,L1,L2,T2,i. ⦃G1, L1, #i⦄ ⊐[b] ⦃G2, L2, T2⦄ →
(∃∃J,V. G1 = G2 & L1 = L2.ⓑ{J}V & T2 = V & i = 0) ∨
- ∃∃J,V,j. G1 = G2 & L1 = L2.ⓑ{J}V & T2 = #j & i = ⫯j.
-/2 width=3 by fqu_inv_lref1_aux/ qed-.
+ ∃∃J,j. G1 = G2 & L1 = L2.ⓘ{J} & T2 = #j & i = ⫯j.
+/2 width=4 by fqu_inv_lref1_aux/ qed-.
-fact fqu_inv_gref1_aux: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+fact fqu_inv_gref1_aux: ∀b,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐[b] ⦃G2, L2, T2⦄ →
∀l. T1 = §l →
- ∃∃J,V. G1 = G2 & L1 = L2.ⓑ{J}V & T2 = §l.
-#G1 #G2 #L1 #L2 #T1 #T2 * -G1 -G2 -L1 -L2 -T1 -T2
+ ∃∃J. G1 = G2 & L1 = L2.ⓘ{J} & T2 = §l.
+#b #G1 #G2 #L1 #L2 #T1 #T2 * -G1 -G2 -L1 -L2 -T1 -T2
[ #I #G #L #T #l #H destruct
| #I #G #L #V #T #l #H destruct
| #p #I #G #L #V #T #l #H destruct
+| #p #I #G #L #V #T #_ #l #H destruct
| #I #G #L #V #T #s #H destruct
-| #I #G #L #V #T #U #HI12 #l #H destruct
- lapply (lifts_inv_gref2 … HI12) -HI12 /2 width=3 by ex3_2_intro/
+| #I #G #L #T #U #HI12 #l #H destruct
+ lapply (lifts_inv_gref2 … HI12) -HI12 /2 width=3 by ex3_intro/
]
qed-.
-lemma fqu_inv_gref1: ∀G1,G2,L1,L2,T2,l. ⦃G1, L1, §l⦄ ⊐ ⦃G2, L2, T2⦄ →
- ∃∃J,V. G1 = G2 & L1 = L2.ⓑ{J}V & T2 = §l.
-/2 width=3 by fqu_inv_gref1_aux/ qed-.
+lemma fqu_inv_gref1: ∀b,G1,G2,L1,L2,T2,l. ⦃G1, L1, §l⦄ ⊐[b] ⦃G2, L2, T2⦄ →
+ ∃∃J. G1 = G2 & L1 = L2.ⓘ{J} & T2 = §l.
+/2 width=4 by fqu_inv_gref1_aux/ qed-.
-fact fqu_inv_bind1_aux: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+fact fqu_inv_bind1_aux: ∀b,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐[b] ⦃G2, L2, T2⦄ →
∀p,I,V1,U1. T1 = ⓑ{p,I}V1.U1 →
∨∨ ∧∧ G1 = G2 & L1 = L2 & V1 = T2
| ∧∧ G1 = G2 & L1.ⓑ{I}V1 = L2 & U1 = T2
- | ∃∃J,V. G1 = G2 & L1 = L2.ⓑ{J}V & ⬆*[1] T2 ≡ ⓑ{p,I}V1.U1.
-#G1 #G2 #L1 #L2 #T1 #T2 * -G1 -G2 -L1 -L2 -T1 -T2
+ | ∧∧ G1 = G2 & L1.ⓧ = L2 & U1 = T2 & b = Ⓕ
+ | ∃∃J. G1 = G2 & L1 = L2.ⓘ{J} & ⬆*[1] T2 ≡ ⓑ{p,I}V1.U1.
+#b #G1 #G2 #L1 #L2 #T1 #T2 * -G1 -G2 -L1 -L2 -T1 -T2
[ #I #G #L #T #q #J #V0 #U0 #H destruct
-| #I #G #L #V #T #q #J #V0 #U0 #H destruct /3 width=1 by and3_intro, or3_intro0/
-| #p #I #G #L #V #T #q #J #V0 #U0 #H destruct /3 width=1 by and3_intro, or3_intro1/
+| #I #G #L #V #T #q #J #V0 #U0 #H destruct /3 width=1 by and3_intro, or4_intro0/
+| #p #I #G #L #V #T #q #J #V0 #U0 #H destruct /3 width=1 by and3_intro, or4_intro1/
+| #p #I #G #L #V #T #Hb #q #J #V0 #U0 #H destruct /3 width=1 by and4_intro, or4_intro2/
| #I #G #L #V #T #q #J #V0 #U0 #H destruct
-| #I #G #L #V #T #U #HTU #q #J #V0 #U0 #H destruct /3 width=3 by or3_intro2, ex3_2_intro/
+| #I #G #L #T #U #HTU #q #J #V0 #U0 #H destruct /3 width=2 by or4_intro3, ex3_intro/
]
qed-.
-lemma fqu_inv_bind1: ∀p,I,G1,G2,L1,L2,V1,U1,T2. ⦃G1, L1, ⓑ{p,I}V1.U1⦄ ⊐ ⦃G2, L2, T2⦄ →
+lemma fqu_inv_bind1: ∀b,p,I,G1,G2,L1,L2,V1,U1,T2. ⦃G1, L1, ⓑ{p,I}V1.U1⦄ ⊐[b] ⦃G2, L2, T2⦄ →
∨∨ ∧∧ G1 = G2 & L1 = L2 & V1 = T2
| ∧∧ G1 = G2 & L1.ⓑ{I}V1 = L2 & U1 = T2
- | ∃∃J,V. G1 = G2 & L1 = L2.ⓑ{J}V & ⬆*[1] T2 ≡ ⓑ{p,I}V1.U1.
+ | ∧∧ G1 = G2 & L1.ⓧ = L2 & U1 = T2 & b = Ⓕ
+ | ∃∃J. G1 = G2 & L1 = L2.ⓘ{J} & ⬆*[1] T2 ≡ ⓑ{p,I}V1.U1.
/2 width=4 by fqu_inv_bind1_aux/ qed-.
-fact fqu_inv_flat1_aux: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+lemma fqu_inv_bind1_true: ∀p,I,G1,G2,L1,L2,V1,U1,T2. ⦃G1, L1, ⓑ{p,I}V1.U1⦄ ⊐ ⦃G2, L2, T2⦄ →
+ ∨∨ ∧∧ G1 = G2 & L1 = L2 & V1 = T2
+ | ∧∧ G1 = G2 & L1.ⓑ{I}V1 = L2 & U1 = T2
+ | ∃∃J. G1 = G2 & L1 = L2.ⓘ{J} & ⬆*[1] T2 ≡ ⓑ{p,I}V1.U1.
+#p #I #G1 #G2 #L1 #L2 #V1 #U1 #T2 #H elim (fqu_inv_bind1 … H) -H
+/3 width=1 by or3_intro0, or3_intro1, or3_intro2/
+* #_ #_ #_ #H destruct
+qed-.
+
+fact fqu_inv_flat1_aux: ∀b,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐[b] ⦃G2, L2, T2⦄ →
∀I,V1,U1. T1 = ⓕ{I}V1.U1 →
∨∨ ∧∧ G1 = G2 & L1 = L2 & V1 = T2
| ∧∧ G1 = G2 & L1 = L2 & U1 = T2
- | ∃∃J,V. G1 = G2 & L1 = L2.ⓑ{J}V & ⬆*[1] T2 ≡ ⓕ{I}V1.U1.
-#G1 #G2 #L1 #L2 #T1 #T2 * -G1 -G2 -L1 -L2 -T1 -T2
+ | ∃∃J. G1 = G2 & L1 = L2.ⓘ{J} & ⬆*[1] T2 ≡ ⓕ{I}V1.U1.
+#b #G1 #G2 #L1 #L2 #T1 #T2 * -G1 -G2 -L1 -L2 -T1 -T2
[ #I #G #L #T #J #V0 #U0 #H destruct
| #I #G #L #V #T #J #V0 #U0 #H destruct /3 width=1 by and3_intro, or3_intro0/
| #p #I #G #L #V #T #J #V0 #U0 #H destruct
+| #p #I #G #L #V #T #_ #J #V0 #U0 #H destruct
| #I #G #L #V #T #J #V0 #U0 #H destruct /3 width=1 by and3_intro, or3_intro1/
-| #I #G #L #V #T #U #HTU #J #V0 #U0 #H destruct /3 width=3 by or3_intro2, ex3_2_intro/
+| #I #G #L #T #U #HTU #J #V0 #U0 #H destruct /3 width=2 by or3_intro2, ex3_intro/
]
qed-.
-lemma fqu_inv_flat1: ∀I,G1,G2,L1,L2,V1,U1,T2. ⦃G1, L1, ⓕ{I}V1.U1⦄ ⊐ ⦃G2, L2, T2⦄ →
+lemma fqu_inv_flat1: ∀b,I,G1,G2,L1,L2,V1,U1,T2. ⦃G1, L1, ⓕ{I}V1.U1⦄ ⊐[b] ⦃G2, L2, T2⦄ →
∨∨ ∧∧ G1 = G2 & L1 = L2 & V1 = T2
| ∧∧ G1 = G2 & L1 = L2 & U1 = T2
- | ∃∃J,V. G1 = G2 & L1 = L2.ⓑ{J}V & ⬆*[1] T2 ≡ ⓕ{I}V1.U1.
+ | ∃∃J. G1 = G2 & L1 = L2.ⓘ{J} & ⬆*[1] T2 ≡ ⓕ{I}V1.U1.
/2 width=4 by fqu_inv_flat1_aux/ qed-.
(* Advanced inversion lemmas ************************************************)
-lemma fqu_inv_atom1: ∀I,G1,G2,L2,T2. ⦃G1, ⋆, ⓪{I}⦄ ⊐ ⦃G2, L2, T2⦄ → ⊥.
-* #x #G1 #G2 #L2 #T2 #H
+lemma fqu_inv_atom1: ∀b,I,G1,G2,L2,T2. ⦃G1, ⋆, ⓪{I}⦄ ⊐[b] ⦃G2, L2, T2⦄ → ⊥.
+#b * #x #G1 #G2 #L2 #T2 #H
[ elim (fqu_inv_sort1 … H) | elim (fqu_inv_lref1 … H) * | elim (fqu_inv_gref1 … H) ] -H
-#I #V [3: #i ] #_ #H destruct
+#I [2: #V |3: #i ] #_ #H destruct
qed-.
-lemma fqu_inv_sort1_pair: ∀I,G1,G2,K,L2,V,T2,s. ⦃G1, K.ⓑ{I}V, ⋆s⦄ ⊐ ⦃G2, L2, T2⦄ →
+lemma fqu_inv_sort1_bind: ∀b,I,G1,G2,K,L2,T2,s. ⦃G1, K.ⓘ{I}, ⋆s⦄ ⊐[b] ⦃G2, L2, T2⦄ →
∧∧ G1 = G2 & L2 = K & T2 = ⋆s.
-#I #G1 #G2 #K #L2 #V #T2 #s #H elim (fqu_inv_sort1 … H) -H
-#Z #X #H1 #H2 #H3 destruct /2 width=1 by and3_intro/
+#b #I #G1 #G2 #K #L2 #T2 #s #H elim (fqu_inv_sort1 … H) -H
+#Z #X #H1 #H2 destruct /2 width=1 by and3_intro/
qed-.
-lemma fqu_inv_zero1_pair: ∀I,G1,G2,K,L2,V,T2. ⦃G1, K.ⓑ{I}V, #0⦄ ⊐ ⦃G2, L2, T2⦄ →
+lemma fqu_inv_zero1_pair: ∀b,I,G1,G2,K,L2,V,T2. ⦃G1, K.ⓑ{I}V, #0⦄ ⊐[b] ⦃G2, L2, T2⦄ →
∧∧ G1 = G2 & L2 = K & T2 = V.
-#I #G1 #G2 #K #L2 #V #T2 #H elim (fqu_inv_lref1 … H) -H *
-#Z #X [2: #x ] #H1 #H2 #H3 #H4 destruct /2 width=1 by and3_intro/
+#b #I #G1 #G2 #K #L2 #V #T2 #H elim (fqu_inv_lref1 … H) -H *
+#Z #X #H1 #H2 #H3 #H4 destruct /2 width=1 by and3_intro/
qed-.
-lemma fqu_inv_lref1_pair: ∀I,G1,G2,K,L2,V,T2,i. ⦃G1, K.ⓑ{I}V, #(⫯i)⦄ ⊐ ⦃G2, L2, T2⦄ →
+lemma fqu_inv_lref1_bind: ∀b,I,G1,G2,K,L2,T2,i. ⦃G1, K.ⓘ{I}, #(⫯i)⦄ ⊐[b] ⦃G2, L2, T2⦄ →
∧∧ G1 = G2 & L2 = K & T2 = #i.
-#I #G1 #G2 #K #L2 #V #T2 #i #H elim (fqu_inv_lref1 … H) -H *
-#Z #X [2: #x ] #H1 #H2 #H3 #H4 destruct /2 width=1 by and3_intro/
+#b #I #G1 #G2 #K #L2 #T2 #i #H elim (fqu_inv_lref1 … H) -H *
+#Z #X #H1 #H2 #H3 #H4 destruct /2 width=1 by and3_intro/
qed-.
-lemma fqu_inv_gref1_pair: ∀I,G1,G2,K,L2,V,T2,l. ⦃G1, K.ⓑ{I}V, §l⦄ ⊐ ⦃G2, L2, T2⦄ →
+lemma fqu_inv_gref1_bind: ∀b,I,G1,G2,K,L2,T2,l. ⦃G1, K.ⓘ{I}, §l⦄ ⊐[b] ⦃G2, L2, T2⦄ →
∧∧ G1 = G2 & L2 = K & T2 = §l.
-#I #G1 #G2 #K #L2 #V #T2 #l #H elim (fqu_inv_gref1 … H) -H
-#Z #X #H1 #H2 #H3 destruct /2 width=1 by and3_intro/
+#b #I #G1 #G2 #K #L2 #T2 #l #H elim (fqu_inv_gref1 … H) -H
+#Z #H1 #H2 #H3 destruct /2 width=1 by and3_intro/
qed-.
(* Basic_2A1: removed theorems 3:
(* Forward lemmas with length for local environments ************************)
-fact fqu_fwd_length_lref1_aux: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+fact fqu_fwd_length_lref1_aux: ∀b,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐[b] ⦃G2, L2, T2⦄ →
∀i. T1 = #i → |L2| < |L1|.
-#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 // [2: #p]
-#I #G #L #V #T #j #H destruct
+#b #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 // [2,3: #p]
+#I #G #L #V #T [2: #_ ] #j #H destruct
qed-.
-lemma fqu_fwd_length_lref1: ∀G1,G2,L1,L2,T2,i. ⦃G1, L1, #i⦄ ⊐ ⦃G2, L2, T2⦄ →
+lemma fqu_fwd_length_lref1: ∀b,G1,G2,L1,L2,T2,i. ⦃G1, L1, #i⦄ ⊐[b] ⦃G2, L2, T2⦄ →
|L2| < |L1|.
-/2 width=7 by fqu_fwd_length_lref1_aux/
+/2 width=8 by fqu_fwd_length_lref1_aux/
qed-.
(* Inversion lemmas with length for local environments **********************)
-fact fqu_inv_eq_aux: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ →
+fact fqu_inv_eq_aux: ∀b,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐[b] ⦃G2, L2, T2⦄ →
G1 = G2 → |L1| = |L2| → T1 = T2 → ⊥.
-#G1 #G2 #L1 #L2 #T1 #T2 * -G1 -G2 -L1 -L2 -T1 -T2
+#b #G1 #G2 #L1 #L2 #T1 #T2 * -G1 -G2 -L1 -L2 -T1 -T2
[1: #I #G #L #V #_ #H elim (succ_inv_refl_sn … H)
-|5: #I #G #L #V #T #U #_ #_ #H elim (succ_inv_refl_sn … H)
+|6: #I #G #L #T #U #_ #_ #H elim (succ_inv_refl_sn … H)
]
/2 width=4 by discr_tpair_xy_y, discr_tpair_xy_x/
qed-.
-lemma fqu_inv_eq: ∀G,L1,L2,T. ⦃G, L1, T⦄ ⊐ ⦃G, L2, T⦄ → |L1| = |L2| → ⊥.
-#G #L1 #L2 #T #H #H0 @(fqu_inv_eq_aux … H … H0) // (**) (* full auto fails *)
+lemma fqu_inv_eq: ∀b,G,L1,L2,T. ⦃G, L1, T⦄ ⊐[b] ⦃G, L2, T⦄ → |L1| = |L2| → ⊥.
+#b #G #L1 #L2 #T #H #H0 @(fqu_inv_eq_aux … H … H0) // (**) (* full auto fails *)
qed-.
(* Forward lemmas with weight for closures **********************************)
-lemma fqu_fwd_fw: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → ♯{G2, L2, T2} < ♯{G1, L1, T1}.
-#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 //
-#I #I1 #I2 #G #L #V #HI12 normalize in ⊢ (?%%); -I
-<(lifts_fwd_tw … HI12) -I1 /3 width=1 by monotonic_lt_plus_r, monotonic_lt_plus_l/
+lemma fqu_fwd_fw: ∀b,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐[b] ⦃G2, L2, T2⦄ →
+ ♯{G2, L2, T2} < ♯{G1, L1, T1}.
+#b #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 //
+#I #I1 #I2 #G #L #HI12 normalize in ⊢ (?%%); -I1
+<(lifts_fwd_tw … HI12) /3 width=1 by monotonic_lt_plus_r, monotonic_lt_plus_l/
qed-.
(* Advanced eliminators *****************************************************)
-lemma fqu_wf_ind: ∀R:relation3 …. (
- ∀G1,L1,T1. (∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → R G2 L2 T2) →
+lemma fqu_wf_ind: ∀b. ∀R:relation3 …. (
+ ∀G1,L1,T1. (∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐[b] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
R G1 L1 T1
) → ∀G1,L1,T1. R G1 L1 T1.
-#R #HR @(f3_ind … fw) #x #IHx #G1 #L1 #T1 #H destruct /4 width=1 by fqu_fwd_fw/
+#b #R #HR @(f3_ind … fw) #x #IHx #G1 #L1 #T1 #H destruct /4 width=2 by fqu_fwd_fw/
qed-.
(**************************************************************************)
include "basic_2/notation/relations/suptermopt_6.ma".
+include "basic_2/notation/relations/suptermopt_7.ma".
include "basic_2/s_transition/fqu.ma".
(* OPTIONAL SUPCLOSURE ******************************************************)
(* Basic_2A1: was: fquqa *)
(* Basic_2A1: includes: fquq_inv_gen *)
-definition fquq: tri_relation genv lenv term ≝ tri_RC … fqu.
+definition fquq: bool → tri_relation genv lenv term ≝
+ λb. tri_RC … (fqu b).
+
+interpretation
+ "extended optional structural successor (closure)"
+ 'SupTermOpt b G1 L1 T1 G2 L2 T2 = (fquq b G1 L1 T1 G2 L2 T2).
interpretation
"optional structural successor (closure)"
- 'SupTermOpt G1 L1 T1 G2 L2 T2 = (fquq G1 L1 T1 G2 L2 T2).
+ 'SupTermOpt G1 L1 T1 G2 L2 T2 = (fquq true G1 L1 T1 G2 L2 T2).
(* Basic properties *********************************************************)
(* Basic_2A1: includes: fquqa_refl *)
-lemma fquq_refl: tri_reflexive … fquq.
+lemma fquq_refl: ∀b. tri_reflexive … (fquq b).
// qed.
-lemma fqu_fquq: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄.
+lemma fqu_fquq: ∀b,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐[b] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐⸮[b] ⦃G2, L2, T2⦄.
/2 width=1 by or_introl/ qed.
(* Basic_2A1: removed theorems 8:
(* Forward lemmas with length for local environments ************************)
-lemma fquq_fwd_length_lref1: ∀G1,G2,L1,L2,T2,i. ⦃G1, L1, #i⦄ ⊐⸮ ⦃G2, L2, T2⦄ → |L2| ≤ |L1|.
-#G1 #G2 #L1 #L2 #T2 #i #H elim H -H [2: * ]
-/3 width=5 by fqu_fwd_length_lref1, lt_to_le/
+lemma fquq_fwd_length_lref1: ∀b,G1,G2,L1,L2,T2,i. ⦃G1, L1, #i⦄ ⊐⸮[b] ⦃G2, L2, T2⦄ →
+ |L2| ≤ |L1|.
+#b #G1 #G2 #L1 #L2 #T2 #i #H elim H -H [2: * ]
+/3 width=6 by fqu_fwd_length_lref1, lt_to_le/
qed-.
(* Forward lemmas with weight for closures **********************************)
-lemma fquq_fwd_fw: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → ♯{G2, L2, T2} ≤ ♯{G1, L1, T1}.
-#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -H [2: * ]
-/3 width=1 by fqu_fwd_fw, lt_to_le/
+lemma fquq_fwd_fw: ∀b,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮[b] ⦃G2, L2, T2⦄ →
+ ♯{G2, L2, T2} ≤ ♯{G1, L1, T1}.
+#b #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -H [2: * ]
+/3 width=2 by fqu_fwd_fw, lt_to_le/
qed-.
inductive aaa: relation4 genv lenv term aarity ≝
| aaa_sort: ∀G,L,s. aaa G L (⋆s) (⓪)
| aaa_zero: ∀I,G,L,V,B. aaa G L V B → aaa G (L.ⓑ{I}V) (#0) B
-| aaa_lref: ∀I,G,L,V,A,i. aaa G L (#i) A → aaa G (L.ⓑ{I}V) (#⫯i) A
+| aaa_lref: ∀I,G,L,A,i. aaa G L (#i) A → aaa G (L.ⓘ{I}) (#⫯i) A
| aaa_abbr: ∀p,G,L,V,T,B,A.
aaa G L V B → aaa G (L.ⓓV) T A → aaa G L (ⓓ{p}V.T) A
| aaa_abst: ∀p,G,L,V,T,B,A.
fact aaa_inv_sort_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀s. T = ⋆s → A = ⓪.
#G #L #T #A * -G -L -T -A //
[ #I #G #L #V #B #_ #s #H destruct
-| #I #G #L #V #A #i #_ #s #H destruct
+| #I #G #L #A #i #_ #s #H destruct
| #p #G #L #V #T #B #A #_ #_ #s #H destruct
| #p #G #L #V #T #B #A #_ #_ #s #H destruct
| #G #L #V #T #B #A #_ #_ #s #H destruct
∃∃I,K,V. L = K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ⁝ A.
#G #L #T #A * -G -L -T -A /2 width=5 by ex2_3_intro/
[ #G #L #s #H destruct
-| #I #G #L #V #A #i #_ #H destruct
+| #I #G #L #A #i #_ #H destruct
| #p #G #L #V #T #B #A #_ #_ #H destruct
| #p #G #L #V #T #B #A #_ #_ #H destruct
| #G #L #V #T #B #A #_ #_ #H destruct
/2 width=3 by aaa_inv_zero_aux/ qed-.
fact aaa_inv_lref_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀i. T = #(⫯i) →
- ∃∃I,K,V. L = K.ⓑ{I}V & ⦃G, K⦄ ⊢ #i ⁝ A.
+ ∃∃I,K. L = K.ⓘ{I} & ⦃G, K⦄ ⊢ #i ⁝ A.
#G #L #T #A * -G -L -T -A
[ #G #L #s #j #H destruct
| #I #G #L #V #B #_ #j #H destruct
-| #I #G #L #V #A #i #HA #j #H destruct /2 width=5 by ex2_3_intro/
+| #I #G #L #A #i #HA #j #H destruct /2 width=4 by ex2_2_intro/
| #p #G #L #V #T #B #A #_ #_ #j #H destruct
| #p #G #L #V #T #B #A #_ #_ #j #H destruct
| #G #L #V #T #B #A #_ #_ #j #H destruct
qed-.
lemma aaa_inv_lref: ∀G,L,A,i. ⦃G, L⦄ ⊢ #⫯i ⁝ A →
- ∃∃I,K,V. L = K.ⓑ{I}V & ⦃G, K⦄ ⊢ #i ⁝ A.
+ ∃∃I,K. L = K.ⓘ{I} & ⦃G, K⦄ ⊢ #i ⁝ A.
/2 width=3 by aaa_inv_lref_aux/ qed-.
fact aaa_inv_gref_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀l. T = §l → ⊥.
#G #L #T #A * -G -L -T -A
[ #G #L #s #k #H destruct
| #I #G #L #V #B #_ #k #H destruct
-| #I #G #L #V #A #i #_ #k #H destruct
+| #I #G #L #A #i #_ #k #H destruct
| #p #G #L #V #T #B #A #_ #_ #k #H destruct
| #p #G #L #V #T #B #A #_ #_ #k #H destruct
| #G #L #V #T #B #A #_ #_ #k #H destruct
#G #L #T #A * -G -L -T -A
[ #G #L #s #q #W #U #H destruct
| #I #G #L #V #B #_ #q #W #U #H destruct
-| #I #G #L #V #A #i #_ #q #W #U #H destruct
+| #I #G #L #A #i #_ #q #W #U #H destruct
| #p #G #L #V #T #B #A #HV #HT #q #W #U #H destruct /2 width=2 by ex2_intro/
| #p #G #L #V #T #B #A #_ #_ #q #W #U #H destruct
| #G #L #V #T #B #A #_ #_ #q #W #U #H destruct
#G #L #T #A * -G -L -T -A
[ #G #L #s #q #W #U #H destruct
| #I #G #L #V #B #_ #q #W #U #H destruct
-| #I #G #L #V #A #i #_ #q #W #U #H destruct
+| #I #G #L #A #i #_ #q #W #U #H destruct
| #p #G #L #V #T #B #A #_ #_ #q #W #U #H destruct
| #p #G #L #V #T #B #A #HV #HT #q #W #U #H destruct /2 width=5 by ex3_2_intro/
| #G #L #V #T #B #A #_ #_ #q #W #U #H destruct
#G #L #T #A * -G -L -T -A
[ #G #L #s #W #U #H destruct
| #I #G #L #V #B #_ #W #U #H destruct
-| #I #G #L #V #A #i #_ #W #U #H destruct
+| #I #G #L #A #i #_ #W #U #H destruct
| #p #G #L #V #T #B #A #_ #_ #W #U #H destruct
| #p #G #L #V #T #B #A #_ #_ #W #U #H destruct
| #G #L #V #T #B #A #HV #HT #W #U #H destruct /2 width=3 by ex2_intro/
∃∃B. ⦃G, L⦄ ⊢ V ⁝ B & ⦃G, L⦄ ⊢ T ⁝ ②B.A.
/2 width=3 by aaa_inv_appl_aux/ qed-.
-fact aaa_inv_cast_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀W,U. T = ⓝW. U →
+fact aaa_inv_cast_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀W,U. T = ⓝW.U →
⦃G, L⦄ ⊢ W ⁝ A ∧ ⦃G, L⦄ ⊢ U ⁝ A.
#G #L #T #A * -G -L -T -A
[ #G #L #s #W #U #H destruct
| #I #G #L #V #B #_ #W #U #H destruct
-| #I #G #L #V #A #i #_ #W #U #H destruct
+| #I #G #L #A #i #_ #W #U #H destruct
| #p #G #L #V #T #B #A #_ #_ #W #U #H destruct
| #p #G #L #V #T #B #A #_ #_ #W #U #H destruct
| #G #L #V #T #B #A #_ #_ #W #U #H destruct
]
qed-.
-lemma aaa_inv_cast: ∀G,L,W,T,A. ⦃G, L⦄ ⊢ ⓝW. T ⁝ A →
+lemma aaa_inv_cast: ∀G,L,W,T,A. ⦃G, L⦄ ⊢ ⓝW.T ⁝ A →
⦃G, L⦄ ⊢ W ⁝ A ∧ ⦃G, L⦄ ⊢ T ⁝ A.
/2 width=3 by aaa_inv_cast_aux/ qed-.
[ #G #L #s #A2 #H >(aaa_inv_sort … H) -H //
| #I1 #G #L #V1 #B #_ #IH #A2 #H
elim (aaa_inv_zero … H) -H #I2 #K2 #V2 #H #HA2 destruct /2 width=1 by/
-| #I1 #G #L #V1 #B #i #_ #IH #A2 #H
- elim (aaa_inv_lref … H) -H #I2 #K2 #V2 #H #HA2 destruct /2 width=1 by/
+| #I1 #G #L #B #i #_ #IH #A2 #H
+ elim (aaa_inv_lref … H) -H #I2 #K2 #H #HA2 destruct /2 width=1 by/
| #p #G #L #V #T #B1 #A1 #_ #_ #_ #IH #A2 #H
elim (aaa_inv_abbr … H) -H /2 width=1 by/
| #p #G #L #V1 #T1 #B1 #A1 #_ #_ #IHB1 #IHA1 #X #H
#I #G #K #V #B #i elim i -i
[ #L #H lapply (drops_fwd_isid … H ?) -H //
#H destruct /2 width=1 by aaa_zero/
-| #i #IH #L <uni_succ #H #HB lapply (drops_inv_pair2_isuni_next … H) -H // *
- #Z #Y #X #HY #H destruct /3 width=1 by aaa_lref/
+| #i #IH #L <uni_succ #H #HB lapply (drops_inv_bind2_isuni_next … H) -H // *
+ #Z #Y #HY #H destruct /3 width=1 by aaa_lref/
]
qed.
#G #A #i elim i -i
[ #L #H elim (aaa_inv_zero … H) -H /3 width=5 by drops_refl, ex2_3_intro/
| #i #IH #L #H elim (aaa_inv_lref … H) -H
- #I #K #V #H #HA destruct elim (IH … HA) -IH -HA /3 width=5 by drops_drop, ex2_3_intro/
+ #I #K #H #HA destruct elim (IH … HA) -IH -HA /3 width=5 by drops_drop, ex2_3_intro/
]
qed-.
(* Note: it should use drops_split_trans_pair2 *)
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 * *
+@(fqup_wf_ind_eq (Ⓣ)) #G0 #L0 #T0 #IH #G #L1 * *
[ #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 //
lapply (drops_trans … HL21 … HLK1 ??) -HL21 [1,2: // ] #H
elim (drops_split_trans … H) -H [ |*: /2 width=6 by after_uni_dx/ ] #Y #HLK2 #HY
lapply (drops_tls_at … Hf … HY) -HY #HY -Hf
- elim (drops_inv_skip2 … HY) -HY #K2 #V2 #HK21 #HV12 #H destruct
+ elim (drops_inv_skip2 … HY) -HY #Z #K2 #HK21 #HZ #H destruct
+ elim (liftsb_inv_pair_sn … HZ) -HZ #V2 #HV12 #H destruct
/4 width=12 by aaa_lref_drops, fqup_lref, drops_inv_gen/
| #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 #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/
+ /5 width=9 by aaa_abbr, drops_skip, ext2_pair/
| elim (aaa_inv_abst … H) -H #B #A0 #HB #HA #H0
- elim (lifts_inv_bind1 … HX) -HX #V2 #T2 #HV12 #HT12 #H destruct
- /4 width=8 by aaa_abst, drops_skip/
+ elim (lifts_inv_bind1 … HX) -HX #V2 #T2 #HV12 #HT12 #H destruct
+ /5 width=8 by aaa_abst, drops_skip, ext2_pair/
]
| * #V1 #T1 #HG #HL #HT #A #H #b #f #L2 #HL21 #X #HX
[ elim (aaa_inv_appl … H) -H #B #HB #HA
(* Basic_2A1: includes: aaa_inv_lift *)
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 * *
+@(fqup_wf_ind_eq (Ⓣ)) #G0 #L0 #T0 #IH #G #L2 * *
[ #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 //
lapply (drops_split_div … HL21 (𝐔❴i1❵) ???) -HL21 [4: * |*: // ] #Y #HLK1 #HY
lapply (drops_conf … HLK2 … HY ??) -HY [1,2: /2 width=6 by after_uni_dx/ ] #HY
lapply (drops_tls_at … Hf … HY) -HY #HY -Hf
- elim (drops_inv_skip1 … HY) -HY #K1 #V1 #HK21 #HV12 #H destruct
+ elim (drops_inv_skip1 … HY) -HY #Z #K1 #HK21 #HZ #H destruct
+ elim (liftsb_inv_pair_dx … HZ) -HZ #V1 #HV12 #H destruct
/4 width=12 by aaa_lref_drops, fqup_lref, drops_inv_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 #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/
+ /5 width=9 by aaa_abbr, drops_skip, ext2_pair/
| elim (aaa_inv_abst … H) -H #B #A0 #HB #HA #H0
elim (lifts_inv_bind2 … HX) -HX #V1 #T1 #HV12 #HT12 #H destruct
- /4 width=8 by aaa_abst, drops_skip/
+ /5 width=8 by aaa_abst, drops_skip, ext2_pair/
]
| * #V2 #T2 #HG #HL #HT #A #H #b #f #L1 #HL21 #X #HX
[ elim (aaa_inv_appl … H) -H #B #HB #HA
[ elim (aaa_inv_abbr … H)
| elim (aaa_inv_abst … H)
] -H /2 width=2 by ex_intro/
+| #p #I #G #L #V #T #H destruct
| * #G #L #V #T #X #H
[ elim (aaa_inv_appl … H)
| elim (aaa_inv_cast … H)
[ #G #L1 #s1 #X #H1 elim (tdeq_inv_sort1 … H1) -H1 //
| #I #G #L1 #V1 #B #_ #IH #X #H1 >(tdeq_inv_lref1 … H1) -H1
#Y #H2 elim (lfdeq_inv_zero_pair_sn … H2) -H2
- #L2 #V2 #HL12 #HV12 #H2 destruct /3 width=1 by aaa_zero/
-| #I #G #L1 #V1 #A #i #_ #IH #X #H1 >(tdeq_inv_lref1 … H1) -H1
- #Y #H2 elim (lfdeq_inv_lref_pair_sn … H2) -H2
- #L2 #V2 #HL12 #H2 destruct /3 width=1 by aaa_lref/
+ #L2 #V2 #HL12 #HV12 #H destruct /3 width=1 by aaa_zero/
+| #I #G #L1 #A #i #_ #IH #X #H1 >(tdeq_inv_lref1 … H1) -H1
+ #Y #H2 elim (lfdeq_inv_lref_bind_sn … H2) -H2
+ #J #L2 #HL12 #H destruct /3 width=1 by aaa_lref/
| #p #G #L1 #V1 #T1 #B #A #_ #_ #IHV #IHT #X #H1 elim (tdeq_inv_pair1 … H1) -H1
- #V2 #T2 #HV12 #HT12 #H1 #L2 #H2 elim (lfdeq_inv_bind … H2) -H2
- /4 width=2 by aaa_abbr, lfdeq_pair_repl_dx/
+ #V2 #T2 #HV12 #HT12 #H #L2 #H2 elim (lfdeq_inv_bind … H2) -H2 destruct
+ /5 width=2 by aaa_abbr, lfdeq_bind_repl_dx, ext2_pair/
| #p #G #L1 #V1 #T1 #B #A #_ #_ #IHV #IHT #X #H1 elim (tdeq_inv_pair1 … H1) -H1
- #V2 #T2 #HV12 #HT12 #H1 #L2 #H2 elim (lfdeq_inv_bind … H2) -H2
- /4 width=2 by aaa_abst, lfdeq_pair_repl_dx/
+ #V2 #T2 #HV12 #HT12 #H #L2 #H2 elim (lfdeq_inv_bind … H2) -H2 destruct
+ /5 width=2 by aaa_abst, lfdeq_bind_repl_dx, ext2_pair/
| #G #L1 #V1 #T1 #B #A #_ #_ #IHV #IHT #X #H1 elim (tdeq_inv_pair1 … H1) -H1
- #V2 #T2 #HV12 #HT12 #H1 #L2 #H2 elim (lfdeq_inv_flat … H2) -H2
+ #V2 #T2 #HV12 #HT12 #H #L2 #H2 elim (lfdeq_inv_flat … H2) -H2 destruct
/3 width=3 by aaa_appl/
| #G #L1 #V1 #T1 #A #_ #_ #IHV #IHT #X #H1 elim (tdeq_inv_pair1 … H1) -H1
- #V2 #T2 #HV12 #HT12 #H1 #L2 #H2 elim (lfdeq_inv_flat … H2) -H2
+ #V2 #T2 #HV12 #HT12 #H #L2 #H2 elim (lfdeq_inv_flat … H2) -H2 destruct
/3 width=1 by aaa_cast/
]
qed-.
(* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
inductive frees: relation3 lenv term rtmap ≝
-| 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: ∀f,I,L,V. frees L V f →
+| frees_sort: ∀f,L,s. 𝐈⦃f⦄ → frees L (⋆s) f
+| frees_atom: ∀f,i. 𝐈⦃f⦄ → frees (⋆) (#i) (↑*[i]⫯f)
+| frees_pair: ∀f,I,L,V. frees L V f →
frees (L.ⓑ{I}V) (#0) (⫯f)
-| frees_lref: ∀f,I,L,V,i. frees L (#i) f →
- frees (L.ⓑ{I}V) (#⫯i) (↑f)
-| frees_gref: ∀f,I,L,V,l. frees L (§l) f →
- frees (L.ⓑ{I}V) (§l) (↑f)
+| frees_unit: ∀f,I,L. 𝐈⦃f⦄ → frees (L.ⓤ{I}) (#0) (⫯f)
+| frees_lref: ∀f,I,L,i. frees L (#i) f →
+ frees (L.ⓘ{I}) (#⫯i) (↑f)
+| frees_gref: ∀f,L,l. 𝐈⦃f⦄ → frees L (§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 →
interpretation
"context-sensitive free variables (term)"
- 'FreeStar L T t = (frees L T t).
+ 'FreeStar L T f = (frees L T f).
(* Basic inversion lemmas ***************************************************)
-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: ∀f,I. ⋆ ⊢ 𝐅*⦃⓪{I}⦄ ≡ f → 𝐈⦃f⦄.
-/2 width=6 by frees_inv_atom_aux/ qed-.
-
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
+#L #X #f #H elim H -f -L -X //
+[ #f #i #_ #x #H destruct
+| #f #_ #L #V #_ #_ #x #H destruct
+| #f #_ #L #_ #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
]
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: ∀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
+fact frees_inv_atom_aux: ∀f,L,X. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀i. L = ⋆ → X = #i →
+ ∃∃g. 𝐈⦃g⦄ & f = ↑*[i]⫯g.
+#f #L #X #H elim H -f -L -X
+[ #f #L #s #_ #j #_ #H destruct
+| #f #i #Hf #j #_ #H destruct /2 width=3 by ex2_intro/
+| #f #I #L #V #_ #_ #j #H destruct
+| #f #I #L #_ #j #H destruct
+| #f #I #L #i #_ #_ #j #H destruct
+| #f #L #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_gref: ∀f,L,l. L ⊢ 𝐅*⦃§l⦄ ≡ f → 𝐈⦃f⦄.
-/2 width=5 by frees_inv_gref_aux/ qed-.
+lemma frees_inv_atom: ∀f,i. ⋆ ⊢ 𝐅*⦃#i⦄ ≡ f → ∃∃g. 𝐈⦃g⦄ & f = ↑*[i]⫯g.
+/2 width=5 by frees_inv_atom_aux/ qed-.
-fact frees_inv_zero_aux: ∀f,L,X. L ⊢ 𝐅*⦃X⦄ ≡ f → X = #0 →
- (L = ⋆ ∧ 𝐈⦃f⦄) ∨
- ∃∃g,I,K,V. K ⊢ 𝐅*⦃V⦄ ≡ g & L = K.ⓑ{I}V & f = ⫯g.
+fact frees_inv_pair_aux: ∀f,L,X. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀I,K,V. L = K.ⓑ{I}V → X = #0 →
+ ∃∃g. K ⊢ 𝐅*⦃V⦄ ≡ g & f = ⫯g.
#f #L #X * -f -L -X
-[ /3 width=1 by or_introl, conj/
-| #f #I #L #V #s #_ #H destruct
-| /3 width=7 by ex3_4_intro, or_intror/
-| #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
+[ #f #L #s #_ #Z #Y #X #_ #H destruct
+| #f #i #_ #Z #Y #X #H destruct
+| #f #I #L #V #Hf #Z #Y #X #H #_ destruct /2 width=3 by ex2_intro/
+| #f #I #L #_ #Z #Y #X #H destruct
+| #f #I #L #i #_ #Z #Y #X #_ #H destruct
+| #f #L #l #_ #Z #Y #X #_ #H destruct
+| #f1 #f2 #f #p #I #L #V #T #_ #_ #_ #Z #Y #X #_ #H destruct
+| #f1 #f2 #f #I #L #V #T #_ #_ #_ #Z #Y #X #_ #H destruct
]
qed-.
-lemma frees_inv_zero: ∀f,L. L ⊢ 𝐅*⦃#0⦄ ≡ f →
- (L = ⋆ ∧ 𝐈⦃f⦄) ∨
- ∃∃g,I,K,V. K ⊢ 𝐅*⦃V⦄ ≡ g & L = K.ⓑ{I}V & f = ⫯g.
-/2 width=3 by frees_inv_zero_aux/ qed-.
+lemma frees_inv_pair: ∀f,I,K,V. K.ⓑ{I}V ⊢ 𝐅*⦃#0⦄ ≡ f → ∃∃g. K ⊢ 𝐅*⦃V⦄ ≡ g & f = ⫯g.
+/2 width=6 by frees_inv_pair_aux/ qed-.
-fact frees_inv_lref_aux: ∀f,L,X. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀j. X = #(⫯j) →
- (L = ⋆ ∧ 𝐈⦃f⦄) ∨
- ∃∃g,I,K,V. K ⊢ 𝐅*⦃#j⦄ ≡ g & L = K.ⓑ{I}V & f = ↑g.
+fact frees_inv_unit_aux: ∀f,L,X. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀I,K. L = K.ⓤ{I} → X = #0 →
+ ∃∃g. 𝐈⦃g⦄ & f = ⫯g.
#f #L #X * -f -L -X
-[ /3 width=1 by or_introl, conj/
-| #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
+[ #f #L #s #_ #Z #Y #_ #H destruct
+| #f #i #_ #Z #Y #H destruct
+| #f #I #L #V #_ #Z #Y #H destruct
+| #f #I #L #Hf #Z #Y #H destruct /2 width=3 by ex2_intro/
+| #f #I #L #i #_ #Z #Y #_ #H destruct
+| #f #L #l #_ #Z #Y #_ #H destruct
+| #f1 #f2 #f #p #I #L #V #T #_ #_ #_ #Z #Y #_ #H destruct
+| #f1 #f2 #f #I #L #V #T #_ #_ #_ #Z #Y #_ #H destruct
+]
+qed-.
+
+lemma frees_inv_unit: ∀f,I,K. K.ⓤ{I} ⊢ 𝐅*⦃#0⦄ ≡ f → ∃∃g. 𝐈⦃g⦄ & f = ⫯g.
+/2 width=7 by frees_inv_unit_aux/ qed-.
+
+fact frees_inv_lref_aux: ∀f,L,X. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀I,K,j. L = K.ⓘ{I} → X = #(⫯j) →
+ ∃∃g. K ⊢ 𝐅*⦃#j⦄ ≡ g & f = ↑g.
+#f #L #X * -f -L -X
+[ #f #L #s #_ #Z #Y #j #_ #H destruct
+| #f #i #_ #Z #Y #j #H destruct
+| #f #I #L #V #_ #Z #Y #j #_ #H destruct
+| #f #I #L #_ #Z #Y #j #_ #H destruct
+| #f #I #L #i #Hf #Z #Y #j #H1 #H2 destruct /2 width=3 by ex2_intro/
+| #f #L #l #_ #Z #Y #j #_ #H destruct
+| #f1 #f2 #f #p #I #L #V #T #_ #_ #_ #Z #Y #j #_ #H destruct
+| #f1 #f2 #f #I #L #V #T #_ #_ #_ #Z #Y #j #_ #H destruct
+]
+qed-.
+
+lemma frees_inv_lref: ∀f,I,K,i. K.ⓘ{I} ⊢ 𝐅*⦃#(⫯i)⦄ ≡ f →
+ ∃∃g. K ⊢ 𝐅*⦃#i⦄ ≡ g & f = ↑g.
+/2 width=6 by frees_inv_lref_aux/ qed-.
+
+fact frees_inv_gref_aux: ∀f,L,X. L ⊢ 𝐅*⦃X⦄ ≡ f → ∀x. X = §x → 𝐈⦃f⦄.
+#f #L #X #H elim H -f -L -X //
+[ #f #i #_ #x #H destruct
+| #f #_ #L #V #_ #_ #x #H destruct
+| #f #_ #L #_ #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_lref: ∀f,L,i. L ⊢ 𝐅*⦃#(⫯i)⦄ ≡ f →
- (L = ⋆ ∧ 𝐈⦃f⦄) ∨
- ∃∃g,I,K,V. K ⊢ 𝐅*⦃#i⦄ ≡ g & L = K.ⓑ{I}V & f = ↑g.
-/2 width=3 by frees_inv_lref_aux/ qed-.
+lemma frees_inv_gref: ∀f,L,l. L ⊢ 𝐅*⦃§l⦄ ≡ f → 𝐈⦃f⦄.
+/2 width=5 by frees_inv_gref_aux/ qed-.
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.
#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 #L #s #_ #q #J #W #U #H destruct
+| #f #i #_ #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
+| #f #I #L #_ #q #J #W #U #H destruct
+| #f #I #L #i #_ #q #J #W #U #H destruct
+| #f #L #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
]
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.
#f #L #X * -f -L -X
-[ #f #I #_ #J #W #U #H destruct
-| #f #I #L #V #s #_ #J #W #U #H destruct
+[ #f #L #s #_ #J #W #U #H destruct
+| #f #i #_ #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
+| #f #I #L #_ #J #W #U #H destruct
+| #f #I #L #i #_ #J #W #U #H destruct
+| #f #L #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/
]
/2 width=4 by frees_inv_flat_aux/ qed-.
(* Advanced inversion lemmas ***********************************************)
-
+(*
lemma frees_inv_zero_pair: ∀f,I,K,V. K.ⓑ{I}V ⊢ 𝐅*⦃#0⦄ ≡ f →
∃∃g. K ⊢ 𝐅*⦃V⦄ ≡ g & f = ⫯g.
#f #I #K #V #H elim (frees_inv_zero … H) -H *
[ #H destruct
-| #g #Z #Y #X #Hg #H1 #H2 destruct /3 width=3 by ex2_intro/
+| #g #Z #Y #X #Hg #H1 #H2 destruct /3 width=3 by ex2_intro/
+| #g #Z #Y #_ #H destruct
]
qed-.
-lemma frees_inv_lref_pair: ∀f,I,K,V,i. K.ⓑ{I}V ⊢ 𝐅*⦃#(⫯i)⦄ ≡ f →
- ∃∃g. K ⊢ 𝐅*⦃#i⦄ ≡ g & f = ↑g.
-#f #I #K #V #i #H elim (frees_inv_lref … H) -H *
+lemma frees_inv_zero_unit: ∀f,I,K. K.ⓤ{I} ⊢ 𝐅*⦃#0⦄ ≡ f → ∃∃g. 𝐈⦃g⦄ & f = ⫯g.
+#f #I #K #H elim (frees_inv_zero … H) -H *
[ #H destruct
-| #g #Z #Y #X #Hg #H1 #H2 destruct /3 width=3 by ex2_intro/
+| #g #Z #Y #X #_ #H destruct
+| /2 width=3 by ex2_intro/
]
qed-.
-(* Basic forward lemmas ****************************************************)
-
-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/
+lemma frees_inv_lref_bind: ∀f,I,K,i. K.ⓘ{I} ⊢ 𝐅*⦃#(⫯i)⦄ ≡ f →
+ ∃∃g. K ⊢ 𝐅*⦃#i⦄ ≡ g & f = ↑g.
+#f #I #K #i #H elim (frees_inv_lref … H) -H *
+[ #H destruct
+| #g #Z #Y #Hg #H1 #H2 destruct /3 width=3 by ex2_intro/
+]
qed-.
-
+*)
(* Basic properties ********************************************************)
lemma frees_eq_repl_back: ∀L,T. eq_repl_back … (λf. L ⊢ 𝐅*⦃T⦄ ≡ f).
#L #T #f1 #H elim H -f1 -L -T
-[ /3 width=3 by frees_atom, isid_eq_repl_back/
-| #f1 #I #L #V #s #_ #IH #f2 #Hf12
- elim (eq_inv_px … Hf12) -Hf12 /3 width=3 by frees_sort/
-| #f1 #I #L #V #_ #IH #f2 #Hf12
- elim (eq_inv_nx … Hf12) -Hf12 /3 width=3 by frees_zero/
-| #f1 #I #L #V #i #_ #IH #f2 #Hf12
- elim (eq_inv_px … Hf12) -Hf12 /3 width=3 by frees_lref/
-| #f1 #I #L #V #l #_ #IH #f2 #Hf12
- elim (eq_inv_px … Hf12) -Hf12 /3 width=3 by frees_gref/
+[ /3 width=3 by frees_sort, isid_eq_repl_back/
+| #f1 #i #Hf1 #g2 #H
+ elim (eq_inv_pushs_sn … H) -H #g #Hg #H destruct
+ elim (eq_inv_nx … Hg) -Hg
+ /3 width=3 by frees_atom, isid_eq_repl_back/
+| #f1 #I #L #V #_ #IH #g2 #H
+ elim (eq_inv_nx … H) -H
+ /3 width=3 by frees_pair/
+| #f1 #I #L #Hf1 #g2 #H
+ elim (eq_inv_nx … H) -H
+ /3 width=3 by frees_unit, isid_eq_repl_back/
+| #f1 #I #L #i #_ #IH #g2 #H
+ elim (eq_inv_px … H) -H /3 width=3 by frees_lref/
+| /3 width=3 by frees_gref, isid_eq_repl_back/
| /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: ∀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.
+(* Forward lemmas with test for finite colength *****************************)
-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.
+lemma frees_fwd_isfin: ∀f,L,T. L ⊢ 𝐅*⦃T⦄ ≡ f → 𝐅⦃f⦄.
+#f #L #T #H elim H -f -L -T
+/4 width=5 by sor_isfin, isfin_isid, isfin_tl, isfin_pushs, isfin_push, isfin_next/
+qed-.
(* Basic_2A1: removed theorems 30:
frees_eq frees_be frees_inv
include "ground_2/relocation/nstream_coafter.ma".
include "basic_2/relocation/drops_drops.ma".
include "basic_2/static/frees_fqup.ma".
-include "basic_2/static/frees_frees.ma".
(* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
(* Advanced properties ******************************************************)
-lemma frees_lref_atom: ∀b,L,i. ⬇*[b, 𝐔❴i❵] L ≡ ⋆ →
- ∀f. 𝐈⦃f⦄ → L ⊢ 𝐅*⦃#i⦄ ≡ f.
+lemma frees_atom_drops: ∀b,L,i. ⬇*[b, 𝐔❴i❵] L ≡ ⋆ →
+ ∀f. 𝐈⦃f⦄ → L ⊢ 𝐅*⦃#i⦄ ≡ ↑*[i]⫯f.
#b #L elim L -L /2 width=1 by frees_atom/
-#L #I #V #IH *
+#L #I #IH *
[ #H lapply (drops_fwd_isid … H ?) -H // #H destruct
-| /5 width=3 by frees_eq_repl_back, frees_lref, drops_inv_drop1, eq_push_inv_isid/
+| /4 width=3 by frees_lref, drops_inv_drop1/
]
qed.
-lemma frees_lref_pair: ∀f,K,V. K ⊢ 𝐅*⦃V⦄ ≡ f →
- ∀i,I,L. ⬇*[i] L ≡ K.ⓑ{I}V → L ⊢ 𝐅*⦃#i⦄ ≡ ↑*[i] ⫯f.
+lemma frees_pair_drops: ∀f,K,V. K ⊢ 𝐅*⦃V⦄ ≡ f →
+ ∀i,I,L. ⬇*[i] L ≡ K.ⓑ{I}V → L ⊢ 𝐅*⦃#i⦄ ≡ ↑*[i] ⫯f.
#f #K #V #Hf #i elim i -i
-[ #I #L #H lapply (drops_fwd_isid … H ?) -H /2 width=1 by frees_zero/
+[ #I #L #H lapply (drops_fwd_isid … H ?) -H /2 width=1 by frees_pair/
| #i #IH #I #L #H elim (drops_inv_succ … H) -H /3 width=2 by frees_lref/
]
qed.
+lemma frees_unit_drops: ∀f. 𝐈⦃f⦄ → ∀I,K,i,L. ⬇*[i] L ≡ K.ⓤ{I} →
+ L ⊢ 𝐅*⦃#i⦄ ≡ ↑*[i] ⫯f.
+#f #Hf #I #K #i elim i -i
+[ #L #H lapply (drops_fwd_isid … H ?) -H /2 width=1 by frees_unit/
+| #i #IH #Y #H elim (drops_inv_succ … H) -H
+ #J #L #HLK #H destruct /3 width=1 by frees_lref/
+]
+qed.
+(*
lemma frees_sort_pushs: ∀f,K,s. K ⊢ 𝐅*⦃⋆s⦄ ≡ f →
∀i,L. ⬇*[i] L ≡ K → L ⊢ 𝐅*⦃⋆s⦄ ≡ ↑*[i] f.
#f #K #s #Hf #i elim i -i
| #i #IH #L #H elim (drops_inv_succ … H) -H /3 width=1 by frees_sort/
]
qed.
-
+*)
lemma frees_lref_pushs: ∀f,K,j. K ⊢ 𝐅*⦃#j⦄ ≡ f →
∀i,L. ⬇*[i] L ≡ K → L ⊢ 𝐅*⦃#(i+j)⦄ ≡ ↑*[i] f.
#f #K #j #Hf #i elim i -i
[ #L #H lapply (drops_fwd_isid … H ?) -H //
| #i #IH #L #H elim (drops_inv_succ … H) -H
- #I #Y #V #HYK #H destruct /3 width=1 by frees_lref/
+ #I #Y #HYK #H destruct /3 width=1 by frees_lref/
]
qed.
-
+(*
lemma frees_gref_pushs: ∀f,K,l. K ⊢ 𝐅*⦃§l⦄ ≡ f →
∀i,L. ⬇*[i] L ≡ K → L ⊢ 𝐅*⦃§l⦄ ≡ ↑*[i] f.
#f #K #l #Hf #i elim i -i
| #i #IH #L #H elim (drops_inv_succ … H) -H /3 width=1 by frees_gref/
]
qed.
-
+*)
(* Advanced inversion lemmas ************************************************)
-lemma frees_inv_lref_drops: ∀i,f,L. L ⊢ 𝐅*⦃#i⦄ ≡ f →
- (⬇*[Ⓕ, 𝐔❴i❵] L ≡ ⋆ ∧ 𝐈⦃f⦄) ∨
- ∃∃g,I,K,V. K ⊢ 𝐅*⦃V⦄ ≡ g &
- ⬇*[i] L ≡ K.ⓑ{I}V & f = ↑*[i] ⫯g.
-#i elim i -i
-[ #f #L #H elim (frees_inv_zero … H) -H *
- /4 width=7 by ex3_4_intro, or_introl, or_intror, conj, drops_refl/
-| #i #IH #f #L #H elim (frees_inv_lref … H) -H * /3 width=1 by or_introl, conj/
- #g #I #K #V #Hg #H1 #H2 destruct
- elim (IH … Hg) -IH -Hg *
- [ /4 width=3 by or_introl, conj, isid_push, drops_drop/
- | /4 width=7 by drops_drop, ex3_4_intro, or_intror/
+lemma frees_inv_lref_drops: ∀L,i,f. L ⊢ 𝐅*⦃#i⦄ ≡ f →
+ ∨∨ ∃∃g. ⬇*[Ⓕ, 𝐔❴i❵] L ≡ ⋆ & 𝐈⦃g⦄ & f = ↑*[i] ⫯g
+ | ∃∃g,I,K,V. K ⊢ 𝐅*⦃V⦄ ≡ g &
+ ⬇*[i] L ≡ K.ⓑ{I}V & f = ↑*[i] ⫯g
+ | ∃∃g,I,K. ⬇*[i] L ≡ K.ⓤ{I} & f = ↑*[i] ⫯g.
+#L elim L -L
+[ #i #g | #L #I #IH * [ #g cases I -I [ #I | #I #V ] -IH | #i #g ] ] #H
+[ elim (frees_inv_atom … H) -H #f #Hf #H destruct
+ /3 width=3 by or3_intro0, ex3_intro/
+| elim (frees_inv_unit … H) -H #f #Hf #H destruct
+ /4 width=3 by drops_refl, or3_intro2, ex2_3_intro/
+| elim (frees_inv_pair … H) -H #f #Hf #H destruct
+ /4 width=7 by drops_refl, or3_intro1, ex3_4_intro/
+| elim (frees_inv_lref … H) -H #f #Hf #H destruct
+ elim (IH … Hf) -IH -Hf *
+ [ /4 width=3 by drops_drop, or3_intro0, ex3_intro/
+ | /4 width=7 by drops_drop, or3_intro1, ex3_4_intro/
+ | /4 width=3 by drops_drop, or3_intro2, ex2_3_intro/
]
]
qed-.
∀f,L. ⬇*[b, f] L ≡ K → ∀U. ⬆*[f] T ≡ U →
∀f2. f ~⊚ f1 ≡ f2 → L ⊢ 𝐅*⦃U⦄ ≡ f2.
#b #f1 #K #T #H lapply (frees_fwd_isfin … H) elim H -f1 -K -T
-[ #f1 #I #Hf1 #_ #f #L #H1 #U #H2 #f2 #H3
+[ #f1 #K #s #Hf1 #_ #f #L #HLK #U #H2 #f2 #H3
lapply (coafter_isid_inv_dx … H3 … Hf1) -f1 #Hf2
- elim (lifts_inv_atom1 … H2) -H2 *
- /2 width=1 by frees_sort_gen, frees_gref_gen/
- #i #j #Hij #H #H0 destruct
+ >(lifts_inv_sort1 … H2) -U /2 width=1 by frees_sort/
+| #f1 #i #Hf1 #_ #f #L #H1 #U #H2 #f2 #H3
+ elim (lifts_inv_lref1 … H2) -H2 #j #Hij #H destruct
+ elim (coafter_fwd_xnx_pushs … Hij H3) -H3 #g2 #Hg2 #H2 destruct
+ lapply (coafter_isid_inv_dx … Hg2 … Hf1) -f1 #Hf2
elim (drops_inv_atom2 … H1) -H1 #n #g #H1 #Hf
elim (after_at_fwd … Hij … Hf) -f #x #_ #Hj -g -i
lapply (at_inv_uni … Hj) -Hj #H destruct
- /3 width=8 by frees_lref_atom, drops_trans/
-| #f1 #I #K #V #s #_ #IH #Hf1 #f #L #H1 #U #H2 #f2 #H3
- lapply (isfin_inv_push … Hf1 ??) -Hf1 [3: |*: // ] #Hf1
- lapply (lifts_inv_sort1 … H2) -H2 #H destruct
- elim (drops_split_trans_pair2 … H1) -H1 [ |*: // ] #Y #W #HLY #HYK #_
- elim (coafter_fwd_xpx_pushs … 0 … H3) [ |*: // ] #g2 #H2 destruct
- lapply (coafter_tls_succ … H3 ??) -H3 [3: |*: // ] #H3
- lapply (IH … HYK … H3) -IH -H3 -HYK [1,3: // | skip ] #Hf2 #H destruct
- /3 width=5 by drops_isuni_fwd_drop2, frees_sort_pushs/
+ /3 width=8 by frees_atom_drops, drops_trans/
| #f1 #I #K #V #_ #IH #Hf1 #f #L #H1 #U #H2 #f2 #H3
lapply (isfin_inv_next … Hf1 ??) -Hf1 [3: |*: // ] #Hf1
lapply (lifts_inv_lref1 … H2) -H2 * #j #Hf #H destruct
- elim (drops_split_trans_pair2 … H1) -H1 [ |*: // ] #Y #W #HLY #HYK #HVW
- elim (coafter_fwd_xnx_pushs … 0 … H3) [ |*: // ] #g2 #H2 destruct
- lapply (coafter_tls_succ … H3 ??) -H3 [3: |*: // ]
- <tls_S in ⊢ (???%→?); <tl_next_rew #H3 #H destruct
+ elim (drops_split_trans_bind2 … H1) -H1 [ |*: // ] #Z #Y #HLY #HYK #H
+ elim (liftsb_inv_pair_sn … H) -H #W #HVW #H destruct
+ elim (coafter_fwd_xnx_pushs … Hf H3) -H3 #g2 #H3 #H2 destruct
lapply (IH … HYK … HVW … H3) -IH -H3 -HYK -HVW //
- /2 width=5 by frees_lref_pair/
-| #f1 #I #K #V #i #_ #IH #Hf1 #f #L #H1 #U #H2 #f2 #H3
+ /2 width=5 by frees_pair_drops/
+| #f1 #I #K #Hf1 #_ #f #L #H1 #U #H2 #f2 #H3
+ lapply (lifts_inv_lref1 … H2) -H2 * #j #Hf #H destruct
+ elim (coafter_fwd_xnx_pushs … Hf H3) -H3 #g2 #H3 #H2 destruct
+ lapply (coafter_isid_inv_dx … H3 … Hf1) -f1 #Hg2
+ elim (drops_split_trans_bind2 … H1 … Hf) -H1 -Hf #Z #Y #HLY #_ #H
+ lapply (liftsb_inv_unit_sn … H) -H #H destruct
+ /2 width=3 by frees_unit_drops/
+| #f1 #I #K #i #_ #IH #Hf1 #f #L #H1 #U #H2 #f2 #H3
lapply (isfin_inv_push … Hf1 ??) -Hf1 [3: |*: // ] #Hf1
lapply (lifts_inv_lref1 … H2) -H2 * #x #Hf #H destruct
elim (at_inv_nxx … Hf) -Hf [ |*: // ] #j #Hf #H destruct
- elim (drops_split_trans_pair2 … H1) -H1 [ |*: // ] #Y #W #HLY #HYK #_
- elim (coafter_fwd_xpx_pushs … 0 … H3) [ |*: // ] #g2 #H2 destruct
- lapply (coafter_tls_succ … H3 ??) -H3 [3: |*: // ] #H3 #H destruct
+ elim (drops_split_trans_bind2 … H1) -H1 [ |*: // ] #Z #Y #HLY #HYK #_
+ elim (coafter_fwd_xpx_pushs … 0 … H3) [ |*: // ] #g2 #H3 #H2 destruct
lapply (drops_isuni_fwd_drop2 … HLY) -HLY // #HLY
lapply (IH … HYK … H3) -IH -H3 -HYK [4: |*: /2 width=2 by lifts_lref/ ]
>plus_S1 /2 width=3 by frees_lref_pushs/ (**) (* full auto fails *)
-| #f1 #I #K #V #l #_ #IH #Hf1 #f #L #H1 #U #H2 #f2 #H3
- lapply (isfin_inv_push … Hf1 ??) -Hf1 [3: |*: // ] #Hf1
- lapply (lifts_inv_gref1 … H2) -H2 #H destruct
- elim (drops_split_trans_pair2 … H1) -H1 [ |*: // ] #Y #W #HLY #HYK #_
- elim (coafter_fwd_xpx_pushs … 0 … H3) [ |*: // ] #g2 #H2 destruct
- lapply (coafter_tls_succ … H3 ??) -H3 [3: |*: // ] #H3 #H destruct
- lapply (IH … HYK … H3) -IH -H3 -HYK [1,3: // | skip ]
- /3 width=5 by drops_isuni_fwd_drop2, frees_gref_pushs/
+| #f1 #K #l #Hf1 #_ #f #L #HLK #U #H2 #f2 #H3
+ lapply (coafter_isid_inv_dx … H3 … Hf1) -f1 #Hf2
+ >(lifts_inv_gref1 … H2) -U /2 width=1 by frees_gref/
| #f1V #f1T #f1 #p #I #K #V #T #_ #_ #H1f1 #IHV #IHT #H2f1 #f #L #H1 #Y #H2 #f2 #H3
elim (sor_inv_isfin3 … H1f1) // #Hf1V #H
lapply (isfin_inv_tl … H) -H
elim (lifts_inv_bind1 … H2) -H2 #W #U #HVW #HTU #H destruct
elim (coafter_sor … H3 … H1f1) /2 width=5 by coafter_isfin2_fwd/ -H3 -H1f1 #f2V #f2T #Hf2V #H
- elim (coafter_inv_tl1 … H) -H /4 width=5 by frees_bind, drops_skip/
+ elim (coafter_inv_tl1 … H) -H
+ /5 width=5 by frees_bind, drops_skip, ext2_pair/
| #f1V #f1T #f1 #I #K #V #T #_ #_ #H1f1 #IHV #IHT #H2f1 #f #L #H1 #Y #H2 #f2 #H3
elim (sor_inv_isfin3 … H1f1) //
elim (lifts_inv_flat1 … H2) -H2 #W #U #HVW #HTU #H destruct
/3 width=7 by frees_eq_repl_back, coafter_inj/
qed-.
-lemma frees_inv_drops: ∀f2,L,U. L ⊢ 𝐅*⦃U⦄ ≡ f2 →
- ∀f,K. ⬇*[Ⓣ, f] L ≡ K → ∀f1. f ~⊚ f1 ≡ f2 →
- ∃∃T. K ⊢ 𝐅*⦃T⦄ ≡ f1 & ⬆*[f] T ≡ U.
-#f2 #L #U #H lapply (frees_fwd_isfin … H) elim H -f2 -L -U
-[ #f2 #I #Hf2 #_ #f #K #H1 #f1 #H2
- lapply (coafter_fwd_isid2 … H2 ??) -H2 // -Hf2 #Hf1
- elim (drops_inv_atom1 … H1) -H1 #H #Hf destruct
- /4 width=3 by frees_atom, lifts_refl, ex2_intro/
-| #f2 #I #L #W #s #_ #IH #Hf2 #f #Y #H1 #f1 #H2
- lapply (isfin_inv_push … Hf2 ??) -Hf2 [3: |*: // ] #Hf2
- elim (coafter_inv_xxp … H2) -H2 [1,3: * |*: // ]
- [ #g #g1 #Hf2 #H #H0 destruct
- elim (drops_inv_skip1 … H1) -H1 #K #V #HLK #_ #H destruct
- | #g #Hf2 #H destruct
- lapply (drops_inv_drop1 … H1) -H1 #HLK
+lemma frees_inv_drops_next: ∀f1,L1,T1. L1 ⊢ 𝐅*⦃T1⦄ ≡ f1 →
+ ∀I2,L2,V2,n. ⬇*[n] L1 ≡ L2.ⓑ{I2}V2 →
+ ∀g1. ⫯g1 = ⫱*[n] f1 →
+ ∃∃g2. L2 ⊢ 𝐅*⦃V2⦄ ≡ g2 & g2 ⊆ g1.
+#f1 #L1 #T1 #H elim H -f1 -L1 -T1
+[ #f1 #L1 #s #Hf1 #I2 #L2 #V2 #n #_ #g1 #H1 -I2 -L1 -s
+ lapply (isid_tls n … Hf1) -Hf1 <H1 -f1 #Hf1
+ elim (isid_inv_next … Hf1) -Hf1 //
+| #f1 #i #_ #I2 #L2 #V2 #n #H
+ elim (drops_inv_atom1 … H) -H #H destruct
+| #f1 #I1 #L1 #V1 #Hf1 #IH #I2 #L2 #V2 *
+ [ -IH #HL12 lapply (drops_fwd_isid … HL12 ?) -HL12 //
+ #H destruct #g1 #Hgf1 >(injective_next … Hgf1) -g1
+ /2 width=3 by ex2_intro/
+ | -Hf1 #n #HL12 lapply (drops_inv_drop1 … HL12) -HL12
+ #HL12 #g1 <tls_xn <tl_next_rew #Hgf1 elim (IH … HL12 … Hgf1) -IH -HL12 -Hgf1
+ /2 width=3 by ex2_intro/
+ ]
+| #f1 #I1 #L1 #Hf1 #I2 #L2 #V2 *
+ [ #HL12 lapply (drops_fwd_isid … HL12 ?) -HL12 // #H destruct
+ | #n #_ #g1 #Hgf1 elim (isid_inv_next … Hgf1) -Hgf1 <tls_xn /2 width=1 by isid_tls/
+ ]
+| #f1 #I1 #L1 #i #_ #IH #I2 #L2 #V2 *
+ [ -IH #_ #g1 #Hgf1 elim (discr_next_push … Hgf1)
+ | #n #HL12 lapply (drops_inv_drop1 … HL12) -HL12
+ #HL12 #g1 <tls_xn #Hgf1 elim (IH … HL12 … Hgf1) -IH -HL12 -Hgf1
+ /2 width=3 by ex2_intro/
]
- elim (IH … HLK … Hf2) -L // -f2 #X #Hg1 #HX
- lapply (lifts_inv_sort2 … HX) -HX #H destruct
- /3 width=3 by frees_sort, lifts_sort, ex2_intro/
-| #f2 #I #L #W #_ #IH #Hf2 #f #Y #H1 #f1 #H2
- lapply (isfin_inv_next … Hf2 ??) -Hf2 [3: |*: // ] #Hf2
- elim (coafter_inv_xxn … H2) -H2 [ |*: // ] #g #g1 #Hf2 #H0 #H destruct
- elim (drops_inv_skip1 … H1) -H1 #K #V #HLK #HVW #H destruct
- elim (IH … HLK … Hf2) -L // -f2 #X #Hg1 #HX
- lapply (lifts_inj … HX … HVW) -W #H destruct
- /3 width=3 by frees_zero, lifts_lref, ex2_intro/
-| #f2 #I #L #W #j #_ #IH #Hf2 #f #Y #H1 #f1 #H2
- lapply (isfin_inv_push … Hf2 ??) -Hf2 [3: |*: // ] #Hf2
- elim (coafter_inv_xxp … H2) -H2 [1,3: * |*: // ]
- [ #g #g1 #Hf2 #H #H0 destruct
- elim (drops_inv_skip1 … H1) -H1 #K #V #HLK #_ #H destruct
- | #g #Hf2 #H destruct
- lapply (drops_inv_drop1 … H1) -H1 #HLK
+| #f1 #L1 #l #Hf1 #I2 #L2 #V2 #n #_ #g1 #H1 -I2 -L1 -l
+ lapply (isid_tls n … Hf1) -Hf1 <H1 -f1 #Hf1
+ elim (isid_inv_next … Hf1) -Hf1 //
+| #fV1 #fT1 #f1 #p #I1 #L1 #V1 #T1 #_ #_ #Hf1 #IHV1 #IHT1 #I2 #L2 #V2 #n #HL12 #g1 #Hgf1
+ lapply (sor_tls … Hf1 n) -Hf1 <Hgf1 -Hgf1 #Hf1
+ elim (sor_xxn_tl … Hf1) [1,2: * |*: // ] -Hf1
+ #gV1 #gT1 #Hg1
+ [ -IHT1 #H1 #_ elim (IHV1 … HL12 … H1) -IHV1 -HL12 -H1
+ /3 width=6 by sor_inv_sle_sn_trans, ex2_intro/
+ | -IHV1 #_ >tls_xn #H2 elim (IHT1 … H2) -IHT1 -H2
+ /3 width=6 by drops_drop, sor_inv_sle_dx_trans, ex2_intro/
]
- elim (IH … HLK … Hf2) -L // -f2 #X #Hg1 #HX
- elim (lifts_inv_lref2 … HX) -HX #i #Hij #H destruct
- /4 width=7 by frees_lref, lifts_lref, at_S1, at_next, ex2_intro/
-| #f2 #I #L #W #l #_ #IH #Hf2 #f #Y #H1 #f1 #H2
- lapply (isfin_inv_push … Hf2 ??) -Hf2 [3: |*: // ] #Hf2
- elim (coafter_inv_xxp … H2) -H2 [1,3: * |*: // ]
- [ #g #g1 #Hf2 #H #H0 destruct
- elim (drops_inv_skip1 … H1) -H1 #K #V #HLK #_ #H destruct
- | #g #Hf2 #H destruct
- lapply (drops_inv_drop1 … H1) -H1 #HLK
+| #fV1 #fT1 #f1 #I1 #L1 #V1 #T1 #_ #_ #Hf1 #IHV1 #IHT1 #I2 #L2 #V2 #n #HL12 #g1 #Hgf1
+ lapply (sor_tls … Hf1 n) -Hf1 <Hgf1 -Hgf1 #Hf1
+ elim (sor_xxn_tl … Hf1) [1,2: * |*: // ] -Hf1
+ #gV1 #gT1 #Hg1
+ [ -IHT1 #H1 #_ elim (IHV1 … HL12 … H1) -IHV1 -HL12 -H1
+ /3 width=6 by sor_inv_sle_sn_trans, ex2_intro/
+ | -IHV1 #_ #H2 elim (IHT1 … HL12 … H2) -IHT1 -HL12 -H2
+ /3 width=6 by sor_inv_sle_dx_trans, ex2_intro/
]
- elim (IH … HLK … Hf2) -L // -f2 #X #Hg1 #HX
- lapply (lifts_inv_gref2 … HX) -HX #H destruct
- /3 width=3 by frees_gref, lifts_gref, ex2_intro/
-| #f2W #f2U #f2 #p #I #L #W #U #_ #_ #H1f2 #IHW #IHU #H2f2 #f #K #H1 #f1 #H2
- elim (sor_inv_isfin3 … H1f2) // #H1f2W #H
- lapply (isfin_inv_tl … H) -H #H1f2U
- elim (coafter_inv_sor … H2 … H1f2) -H2 -H1f2 // #f1W #f1U #H2f2W #H #Hf1
- elim (coafter_inv_tl0 … H) -H #g1 #H2f2U #H destruct
- elim (IHW … H1 … H2f2W) -IHW -H2f2W // -H1f2W #V #Hf1W #HVW
- elim (IHU … H2f2U) -IHU -H2f2U
- /3 width=5 by frees_bind, drops_skip, lifts_bind, ex2_intro/
-| #f2W #f2U #f2 #I #L #W #U #_ #_ #H1f2 #IHW #IHU #H2f2 #f #K #H1 #f1 #H2
- elim (sor_inv_isfin3 … H1f2) // #H1f2W #H1f2U
- elim (coafter_inv_sor … H2 … H1f2) -H2 -H1f2 // #f1W #f1U #H2f2W #H2f2U #Hf1
- elim (IHW … H1 … H2f2W) -IHW -H2f2W // -H1f2W
- elim (IHU … H1 … H2f2U) -L -H2f2U
- /3 width=5 by frees_flat, lifts_flat, ex2_intro/
]
qed-.
(* *)
(**************************************************************************)
-include "basic_2/relocation/drops.ma".
include "basic_2/s_computation/fqup_weight.ma".
-include "basic_2/static/frees.ma".
+include "basic_2/static/lsubf_lsubr.ma".
(* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
(* Note: this replaces lemma 1400 concluding the "big tree" theorem *)
lemma frees_total: ∀L,T. ∃f. L ⊢ 𝐅*⦃T⦄ ≡ f.
-#L #T @(fqup_wf_ind_eq … (⋆) L T) -L -T
-#G0 #L0 #T0 #IH #G #L *
-[ cases L -L /3 width=2 by frees_atom, ex_intro/
- #L #I #V *
- [ #s #HG #HL #HT destruct
- elim (IH G L (⋆s)) -IH /3 width=2 by frees_sort_gen, fqu_fqup, fqu_drop, lifts_sort, ex_intro/
- | * [2: #i ] #HG #HL #HT destruct
- [ elim (IH G L (#i)) -IH /3 width=2 by frees_lref, fqu_fqup, ex_intro/
- | elim (IH G L V) -IH /3 width=2 by frees_zero, fqu_fqup, fqu_lref_O, ex_intro/
+#L #T @(fqup_wf_ind_eq (Ⓣ) … (⋆) L T) -L -T
+#G0 #L0 #T0 #IH #G #L * *
+[ /3 width=2 by frees_sort, ex_intro/
+| cases L -L /3 width=2 by frees_atom, ex_intro/
+ #L #I *
+ [ cases I -I #I [2: #V ] #HG #HL #HT destruct
+ [ elim (IH G L V) -IH
+ /3 width=2 by frees_pair, fqu_fqup, fqu_lref_O, ex_intro/
+ | -IH /3 width=2 by frees_unit, ex_intro/
]
- | #l #HG #HL #HT destruct
- elim (IH G L (§l)) -IH /3 width=2 by frees_gref_gen, fqu_fqup, fqu_drop, lifts_gref, ex_intro/
- ]
-| * [ #p ] #I #V #T #HG #HL #HT destruct elim (IH G L V) // #f1 #HV
- [ elim (IH G (L.ⓑ{I}V) T) -IH // #f2 #HT
- elim (sor_isfin_ex f1 (⫱f2))
- /3 width=6 by frees_fwd_isfin, frees_bind, isfin_tl, ex_intro/
- | elim (IH G L T) -IH // #f2 #HT
- elim (sor_isfin_ex f1 f2)
- /3 width=6 by frees_fwd_isfin, frees_flat, ex_intro/
+ | #i #HG #HL #HT destruct
+ elim (IH G L (#i)) -IH
+ /3 width=2 by frees_lref, fqu_fqup, fqu_drop, ex_intro/
]
+| /3 width=2 by frees_gref, ex_intro/
+| #p #I #V #T #HG #HL #HT destruct
+ elim (IH G L V) // #f1 #HV
+ elim (IH G (L.ⓑ{I}V) T) -IH // #f2 #HT
+ elim (sor_isfin_ex f1 (⫱f2))
+ /3 width=6 by frees_fwd_isfin, frees_bind, isfin_tl, ex_intro/
+| #I #V #T #HG #HL #HT destruct
+ elim (IH G L V) // #f1 #HV
+ elim (IH G L T) -IH // #f2 #HT
+ elim (sor_isfin_ex f1 f2)
+ /3 width=6 by frees_fwd_isfin, frees_flat, ex_intro/
]
qed-.
-(* Properties with plus-iterated supclosure *********************************)
+(* Advanced main properties *************************************************)
-lemma frees_drops_next: ∀f1,L1,T1. L1 ⊢ 𝐅*⦃T1⦄ ≡ f1 →
- ∀I2,L2,V2,n. ⬇*[n] L1 ≡ L2.ⓑ{I2}V2 →
- ∀g1. ⫯g1 = ⫱*[n] f1 →
- ∃∃g2. L2 ⊢ 𝐅*⦃V2⦄ ≡ g2 & g2 ⊆ g1.
-#f1 #L1 #T1 #H elim H -f1 -L1 -T1
-[ #f1 #I1 #Hf1 #I2 #L2 #V2 #n #HL12
- elim (drops_inv_atom1 … HL12) -HL12 #H destruct
-| #f1 #I1 #L1 #V1 #s #_ #IH #I2 #L2 #V2 *
- [ -IH #_ #g1 #Hgf1 elim (discr_next_push … Hgf1)
- | #n #HL12 lapply (drops_inv_drop1 … HL12) -HL12
- #HL12 #g1 <tls_xn #Hgf1 elim (IH … HL12 … Hgf1) -IH -HL12 -Hgf1
- /2 width=3 by ex2_intro/
- ]
-| #f1 #I1 #L1 #V1 #Hf1 #IH #I2 #L2 #V2 *
- [ -IH #HL12 lapply (drops_fwd_isid … HL12 ?) -HL12 //
- #H destruct #g1 #Hgf1 >(injective_next … Hgf1) -g1
- /3 width=3 by sle_refl, ex2_intro/
- | -Hf1 #n #HL12 lapply (drops_inv_drop1 … HL12) -HL12
- #HL12 #g1 <tls_xn <tl_next_rew #Hgf1 elim (IH … HL12 … Hgf1) -IH -HL12 -Hgf1
- /2 width=3 by ex2_intro/
- ]
-| #f1 #I1 #L1 #V1 #i #_ #IH #I2 #L2 #V2 *
- [ -IH #_ #g1 #Hgf1 elim (discr_next_push … Hgf1)
- | #n #HL12 lapply (drops_inv_drop1 … HL12) -HL12
- #HL12 #g1 <tls_xn #Hgf1 elim (IH … HL12 … Hgf1) -IH -HL12 -Hgf1
- /2 width=3 by ex2_intro/
- ]
-| #f1 #I1 #L1 #V1 #l #_ #IH #I2 #L2 #V2 *
- [ -IH #_ #g1 #Hgf1 elim (discr_next_push … Hgf1)
- | #n #HL12 lapply (drops_inv_drop1 … HL12) -HL12
- #HL12 #g1 <tls_xn #Hgf1 elim (IH … HL12 … Hgf1) -IH -HL12 -Hgf1
- /2 width=3 by ex2_intro/
- ]
-| #fV1 #fT1 #f1 #p #I1 #L1 #V1 #T1 #_ #_ #Hf1 #IHV1 #IHT1 #I2 #L2 #V2 #n #HL12 #g1 #Hgf1
- lapply (sor_tls … Hf1 n) -Hf1 <Hgf1 -Hgf1 #Hf1
- elim (sor_xxn_tl … Hf1) [1,2: * |*: // ] -Hf1
- #gV1 #gT1 #Hg1
- [ -IHT1 #H1 #_ elim (IHV1 … HL12 … H1) -IHV1 -HL12 -H1
- /3 width=6 by sor_inv_sle_sn_trans, ex2_intro/
- | -IHV1 #_ >tls_xn #H2 elim (IHT1 … H2) -IHT1 -H2
- /3 width=6 by drops_drop, sor_inv_sle_dx_trans, ex2_intro/
- ]
-| #fV1 #fT1 #f1 #I1 #L1 #V1 #T1 #_ #_ #Hf1 #IHV1 #IHT1 #I2 #L2 #V2 #n #HL12 #g1 #Hgf1
- lapply (sor_tls … Hf1 n) -Hf1 <Hgf1 -Hgf1 #Hf1
- elim (sor_xxn_tl … Hf1) [1,2: * |*: // ] -Hf1
- #gV1 #gT1 #Hg1
- [ -IHT1 #H1 #_ elim (IHV1 … HL12 … H1) -IHV1 -HL12 -H1
- /3 width=6 by sor_inv_sle_sn_trans, ex2_intro/
- | -IHV1 #_ #H2 elim (IHT1 … HL12 … H2) -IHT1 -HL12 -H2
- /3 width=6 by sor_inv_sle_dx_trans, ex2_intro/
- ]
+theorem frees_bind_void: ∀f1,L,V. L ⊢ 𝐅*⦃V⦄ ≡ f1 → ∀f2,T. L.ⓧ ⊢ 𝐅*⦃T⦄ ≡ f2 →
+ ∀f. f1 ⋓ ⫱f2 ≡ f → ∀p,I. L ⊢ 𝐅*⦃ⓑ{p,I}V.T⦄ ≡ f.
+#f1 #L #V #Hf1 #f2 #T #Hf2 #f #Hf #p #I
+elim (frees_total (L.ⓑ{I}V) T) #f0 #Hf0
+lapply (lsubr_lsubf … Hf2 … Hf0) -Hf2 /2 width=5 by lsubr_unit/ #H02
+elim (pn_split f2) * #g2 #H destruct
+[ elim (lsubf_inv_push2 … H02) -H02 #g0 #Z #Y #H02 #H0 #H destruct
+ lapply (lsubf_inv_refl … H02) -H02 #H02
+ lapply (sor_eq_repl_fwd2 … Hf … H02) -g2 #Hf
+ /2 width=5 by frees_bind/
+| elim (lsubf_inv_unit2 … H02) -H02 * [ #g0 #Y #_ #_ #H destruct ]
+ #z1 #g0 #z #Z #Y #X #H02 #Hz1 #Hz #H0 #H destruct
+ lapply (lsubf_inv_refl … H02) -H02 #H02
+ lapply (frees_mono … Hz1 … Hf1) -Hz1 #H1
+ lapply (sor_eq_repl_back1 … Hz … H02) -g0 #Hz
+ lapply (sor_eq_repl_back2 … Hz … H1) -z1 #Hz
+ lapply (sor_sym … Hz) -Hz #Hz
+ lapply (sor_mono … f Hz ?) // -Hz #H
+ lapply (sor_inv_sle_sn … Hf) -Hf #Hf
+ lapply (frees_eq_repl_back … Hf0 (⫯f) ?) /2 width=5 by eq_next/ -z #Hf0
+ @(frees_bind … Hf1 Hf0) -Hf1 -Hf0 (**) (* constructor needed *)
+ /2 width=1 by sor_sle_dx/
+]
+qed-.
+
+(* Advanced inversion lemmas ************************************************)
+
+lemma frees_inv_bind_void: ∀f,p,I,L,V,T. L ⊢ 𝐅*⦃ⓑ{p,I}V.T⦄ ≡ f →
+ ∃∃f1,f2. L ⊢ 𝐅*⦃V⦄ ≡ f1 & L.ⓧ ⊢ 𝐅*⦃T⦄ ≡ f2 & f1 ⋓ ⫱f2 ≡ f.
+#f #p #I #L #V #T #H
+elim (frees_inv_bind … H) -H #f1 #f2 #Hf1 #Hf2 #Hf
+elim (frees_total (L.ⓧ) T) #f0 #Hf0
+lapply (lsubr_lsubf … Hf0 … Hf2) -Hf2 /2 width=5 by lsubr_unit/ #H20
+elim (pn_split f0) * #g0 #H destruct
+[ elim (lsubf_inv_push2 … H20) -H20 #g2 #I #Y #H20 #H2 #H destruct
+ lapply (lsubf_inv_refl … H20) -H20 #H20
+ lapply (sor_eq_repl_back2 … Hf … H20) -g2 #Hf
+ /2 width=5 by ex3_2_intro/
+| elim (lsubf_inv_unit2 … H20) -H20 * [ #g2 #Y #_ #_ #H destruct ]
+ #z1 #z0 #g2 #Z #Y #X #H20 #Hz1 #Hg2 #H2 #H destruct
+ lapply (lsubf_inv_refl … H20) -H20 #H0
+ lapply (frees_mono … Hz1 … Hf1) -Hz1 #H1
+ lapply (sor_eq_repl_back1 … Hg2 … H0) -z0 #Hg2
+ lapply (sor_eq_repl_back2 … Hg2 … H1) -z1 #Hg2
+ @(ex3_2_intro … Hf1 Hf0) -Hf1 -Hf0 (**) (* constructor needed *)
+ /2 width=3 by sor_trans2_idem/
]
qed-.
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/
-| #f1 #I #L #V #_ #IH #x #H elim (frees_inv_zero … H) -H *
- [ #H destruct
- | #f2 #Z #Y #X #Hf2 #H1 #H2 destruct /3 width=5 by eq_next/
- ]
-| #f1 #I #L #V #i #_ #IH #x #H elim (frees_inv_lref … H) -H *
- [ #H destruct
- | #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/
-| #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
+[ /3 width=3 by frees_inv_sort, isid_inv_eq_repl/
+| #f1 #i #Hf1 #g2 #H
+ elim (frees_inv_atom … H) -H #f2 #Hf2 #H destruct
+ /4 width=5 by isid_inv_eq_repl, pushs_eq_repl, eq_next/
+| #f1 #I #L #V #_ #IH #g2 #H elim (frees_inv_pair … H) -H
+ #f2 #Hf2 #H destruct /3 width=5 by eq_next/
+| #f1 #I #L #Hf1 #g2 #H elim (frees_inv_unit … H) -H
+ #f2 #Hf2 #H destruct /3 width=5 by isid_inv_eq_repl, eq_next/
+| #f1 #I #L #i #_ #IH #g2 #H elim (frees_inv_lref … H) -H
+ #f2 #Hf2 #H destruct /3 width=5 by eq_push/
+| /3 width=3 by frees_inv_gref, isid_inv_eq_repl/
+| #f1V #f1T #f1 #p #I #L #V #T #_ #_ #Hf1 #IHV #IHT #f2 #H elim (frees_inv_bind … H) -H
+ #f2V #f2T #HV #HT #Hf2 @(sor_mono … Hf1) -Hf1
/5 width=3 by sor_eq_repl_fwd2, sor_eq_repl_fwd1, tl_eq_repl/ (**) (* full auto too slow *)
-| #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
+| #f1V #f1T #f1 #I #L #V #T #_ #_ #Hf1 #IHV #IHT #f2 #H elim (frees_inv_flat … H) -H
+ #f2V #f2T #HV #HT #Hf2 @(sor_mono … Hf1) -Hf1
/4 width=3 by sor_eq_repl_fwd2, sor_eq_repl_fwd1/ (**) (* full auto too slow *)
]
qed-.
∀G,L1,T,A. ⦃G, L1⦄ ⊢ T ⁝ A → ∀b,f,L0. ⬇*[b, f] L0 ≡ L1 →
∀T0. ⬆*[f] T ≡ T0 → ∀L2. G ⊢ L2 ⫃[RP] L0 →
⦃G, L2, T0⦄ ϵ[RP] 〚A〛.
-#RR #RS #RP #H1RP #H2RP #G #L1 #T @(fqup_wf_ind_eq … G L1 T) -G -L1 -T
+#RR #RS #RP #H1RP #H2RP #G #L1 #T @(fqup_wf_ind_eq (Ⓣ) … G L1 T) -G -L1 -T
#Z #Y #X #IH #G #L1 * [ * | * [ #p ] * ]
[ #s #HG #HL #HT #A #HA #b #f #L0 #HL01 #X0 #H0 #L2 #HL20 destruct -IH
lapply (aaa_inv_sort … HA) -HA #H destruct
lapply (drops_trans … HL01 … HLK1 ??) -HL01 [3: |*: // ] #H
elim (drops_split_trans … H) -H [ |*: /2 width=6 by after_uni_dx/ ] #Y #HLK0 #HY
lapply (drops_tls_at … Hf … HY) -Hf -HY #HY
- elim (drops_inv_skip2 … HY) -HY #K0 #V0 #HK01 #HV10 #H destruct
+ elim (drops_inv_skip2 … HY) -HY #Z #K0 #HK01 #HZ #H destruct
+ elim (liftsb_inv_pair_sn … HZ) -HZ #V0 #HV10 #H destruct
elim (lifts_total V0 (𝐔❴⫯j❵)) #V #HV0
elim (lsubc_drops_trans_isuni … HL20 … HLK0) -HL20 -HLK0 // #Y #HLK2 #H
- elim (lsubc_inv_pair2 … H) -H *
+ elim (lsubc_inv_bind2 … H) -H *
[ #K2 #HK20 #H destruct
lapply (drops_isuni_fwd_drop2 … HLK2) // #HLK2b
lapply (s5 … HA ? G ? ? (◊) … HV0 ?) -HA
/4 width=11 by acr_lifts, fqup_lref, drops_inv_gen/
- | #K2 #V2 #B #HKV2 #HK2V0 #HKV0B #_ #H1 #H2 destruct -IH -HLK1
+ | #K2 #V2 #W2 #B #HKV2 #HK2V0 #HKV0B #_ #H1 #H2 destruct -IH -HLK1
lapply (drops_isuni_fwd_drop2 … HLK2) // #HLK2b
lapply (aaa_lifts … HKV1 … HK01 … HV10) -HKV1 -HK01 -HV10 #HKV0A
lapply (aaa_mono … HKV0B … HKV0A) #H destruct -HKV0B -HKV0A
elim (lifts_total V2 (𝐔❴⫯j❵)) #V3 #HV23
- lapply (s5 … HA … G … (◊) … (ⓝV0.V2) (ⓝV.V3) ????)
+ lapply (s5 … HA … G … (◊) … (ⓝW2.V2) (ⓝV.V3) ????)
[3: |*: /2 width=9 by drops_inv_gen, lifts_flat/ ] -HLK2
lapply (s7 … HA G L2 (◊)) -HA /3 width=7 by acr_lifts/
]
lapply (acr_gcr … H1RP H2RP A) #HA
lapply (acr_gcr … H1RP H2RP B) #HB
lapply (s1 … HB) -HB #HB
- lapply (s6 … HA G L2 (◊) (◊)) /4 width=10 by lsubc_pair, liftsv_nil, drops_skip/
+ lapply (s6 … HA G L2 (◊) (◊)) /5 width=10 by lsubc_bind, liftsv_nil, drops_skip, ext2_pair/
| #W #T #HG #HL #HT #Z0 #HA #b #f #L0 #HL01 #X0 #H0 #L2 #HL20 destruct
elim (aaa_inv_abst … HA) -HA #B #A #HW #HT #H destruct
elim (lifts_inv_bind1 … H0) -H0 #W0 #T0 #HW0 #HT0 #H destruct
lapply (aaa_lifts … HW … (f3∘f) L2 … W3 ?) -HW
[4: |*: /2 width=8 by drops_trans, lifts_trans/ ] #HW3
@(IH … ((↑f3)∘↑f) … (L2. ⓛW3)) -IH
- /3 width=12 by lsubc_beta, drops_trans, drops_skip, lifts_trans/
+ /4 width=12 by lsubc_beta, drops_trans, drops_skip, lifts_trans, ext2_pair/
| #V #T #HG #HL #HT #A #HA #b #f #L0 #HL01 #X0 #H0 #L2 #HL20 destruct
elim (aaa_inv_appl … HA) -HA #B #HV #HT
elim (lifts_inv_flat1 … H0) -H0 #V0 #T0 #HV0 #HT0 #H destruct
lapply (s1 … IHB … HB) #HV0
@(s2 … IHA … (V0@V0s)) /3 width=13 by cp0, gcp2_all, lifts_simple_dx, conj/
| #p #G #L #Vs #U #T #W #HA #f #L0 #V0 #X #HL0 #H #HB
- elim (lifts_inv_applv1 … H) -H #V0s #Y #HV0s #HY #H destruct
- elim (lifts_inv_flat1 … HY) -HY #U0 #X #HU0 #HX #H destruct
+ elim (lifts_inv_applv1 … H) -H #V0s #X0 #HV0s #H0 #H destruct
+ elim (lifts_inv_flat1 … H0) -H0 #U0 #X #HU0 #HX #H destruct
elim (lifts_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT0 #H destruct
@(s3 … IHA … (V0@V0s)) /5 width=6 by lifts_applv, lifts_flat, lifts_bind/
| #G #L #Vs #HVs #s #f #L0 #V0 #X #HL0 #H #HB
- elim (lifts_inv_applv1 … H) -H #V0s #Y #HV0s #HY #H destruct
- >(lifts_inv_sort1 … HY) -Y
+ elim (lifts_inv_applv1 … H) -H #V0s #X0 #HV0s #H0 #H destruct
+ >(lifts_inv_sort1 … H0) -X0
lapply (s1 … IHB … HB) #HV0
@(s4 … IHA … (V0@V0s)) /3 width=7 by gcp2_all, conj/
| #I #G #L #K #Vs #V1 #V2 #i #HA #HV12 #HLK #f #L0 #V0 #X #HL0 #H #HB
- elim (lifts_inv_applv1 … H) -H #V0s #Y #HV0s #HY #H destruct
- elim (lifts_inv_lref1 … HY) -HY #j #Hf #H destruct
+ elim (lifts_inv_applv1 … H) -H #V0s #X0 #HV0s #H0 #H destruct
+ elim (lifts_inv_lref1 … H0) -H0 #j #Hf #H destruct
lapply (drops_trans … HL0 … HLK ??) [3: |*: // ] -HLK #H
elim (drops_split_trans … H) -H [ |*: /2 width=6 by after_uni_dx/ ] #Y #HLK0 #HY
lapply (drops_tls_at … Hf … HY) -HY #HY
- elim (drops_inv_skip2 … HY) -HY #K0 #W1 #_ #HVW1 #H destruct
+ elim (drops_inv_skip2 … HY) -HY #Z #K0 #HK0 #HZ #H destruct
+ elim (liftsb_inv_pair_sn … HZ) -HZ #W1 #HVW1 #H destruct
elim (lifts_total W1 (𝐔❴⫯j❵)) #W2 #HW12
lapply (lifts_trans … HVW1 … HW12 ??) -HVW1 [3: |*: // ] #H
lapply (lifts_conf … HV12 … H f ?) -V1 [ /2 width=3 by after_uni_succ_sn/ ] #HVW2
@(s5 … IHA … (V0@V0s) … HW12) /3 width=4 by drops_inv_gen, lifts_applv/
| #G #L #V1s #V2s #HV12s #p #V #T #HA #HV #f #L0 #V10 #X #HL0 #H #HB
- elim (lifts_inv_applv1 … H) -H #V10s #Y #HV10s #HY #H destruct
- elim (lifts_inv_bind1 … HY) -HY #V0 #T0 #HV0 #HT0 #H destruct
+ elim (lifts_inv_applv1 … H) -H #V10s #X0 #HV10s #H0 #H destruct
+ elim (lifts_inv_bind1 … H0) -H0 #V0 #T0 #HV0 #HT0 #H destruct
elim (lifts_total V10 (𝐔❴1❵)) #V20 #HV120
elim (liftsv_total (𝐔❴1❵) V10s) #V20s #HV120s
@(s6 … IHA … (V10@V10s) (V20@V20s)) /3 width=7 by cp2, liftsv_cons/
- @(HA … (↑f)) /2 width=2 by drops_skip/
+ @(HA … (↑f)) /3 width=2 by drops_skip, ext2_pair/
[ @lifts_applv //
lapply (liftsv_trans … HV10s … HV120s ??) -V10s [3: |*: // ] #H
elim (liftsv_split_trans … H (𝐔❴1❵) (↑f)) /2 width=1 by after_uni_one_sn/ #V10s #HV10s #HV120s
| @(acr_lifts … H1RP … HB … HV120) /3 width=2 by drops_refl, drops_drop/
]
| #G #L #Vs #T #W #HA #HW #f #L0 #V0 #X #HL0 #H #HB
- elim (lifts_inv_applv1 … H) -H #V0s #Y #HV0s #HY #H destruct
- elim (lifts_inv_flat1 … HY) -HY #W0 #T0 #HW0 #HT0 #H destruct
+ elim (lifts_inv_applv1 … H) -H #V0s #X0 #HV0s #H0 #H destruct
+ elim (lifts_inv_flat1 … H0) -H0 #W0 #T0 #HW0 #HT0 #H destruct
@(s7 … IHA … (V0@V0s)) /3 width=5 by lifts_applv/
]
qed.
(**************************************************************************)
include "basic_2/notation/relations/lazyeq_5.ma".
-include "basic_2/syntax/tdeq.ma".
+include "basic_2/syntax/tdeq_ext.ma".
include "basic_2/static/lfxs.ma".
(* DEGREE-BASED EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES ******)
interpretation
"degree-based ranged equivalence (local environment)"
- 'LazyEq h o f L1 L2 = (lexs (cdeq h o) cfull f L1 L2).
+ 'LazyEq h o f L1 L2 = (lexs (cdeq_ext h o) cfull f L1 L2).
(*
definition lfdeq_transitive: predicate (relation3 lenv term term) ≝
λR. ∀L2,T1,T2. R L2 T1 T2 → ∀L1. L1 ≡[h, o, T1] L2 → R L1 T1 T2.
lemma frees_tdeq_conf_lexs: ∀h,o,f,L1,T1. L1 ⊢ 𝐅*⦃T1⦄ ≡ f → ∀T2. T1 ≡[h, o] T2 →
∀L2. L1 ≡[h, o, f] L2 → L2 ⊢ 𝐅*⦃T2⦄ ≡ f.
#h #o #f #L1 #T1 #H elim H -f -L1 -T1
-[ #f #I1 #Hf #X #H1 elim (tdeq_fwd_atom1 … H1) -H1
- #I2 #H1 #Y #H2 lapply (lexs_inv_atom1 … H2) -H2
- #H2 destruct /2 width=1 by frees_atom/
-| #f #I #L1 #V1 #s1 #_ #IH #X #H1 elim (tdeq_inv_sort1 … H1) -H1
- #s2 #d #Hs1 #Hs2 #H1 #Y #H2 elim (lexs_inv_push1 … H2) -H2
- #L2 #V2 #HL12 #_ #H2 destruct /4 width=3 by frees_sort, tdeq_sort/
-| #f #I #L1 #V1 #_ #IH #X #H1 >(tdeq_inv_lref1 … H1) -H1
- #Y #H2 elim (lexs_inv_next1 … H2) -H2
- #L2 #V2 #HL12 #HV12 #H2 destruct /3 width=1 by frees_zero/
-| #f #I #L1 #V1 #i #_ #IH #X #H1 >(tdeq_inv_lref1 … H1) -H1
- #Y #H2 elim (lexs_inv_push1 … H2) -H2
- #L2 #V2 #HL12 #_ #H2 destruct /3 width=1 by frees_lref/
-| #f #I #L1 #V1 #l #_ #IH #X #H1 >(tdeq_inv_gref1 … H1) -H1
- #Y #H2 elim (lexs_inv_push1 … H2) -H2
- #L2 #V2 #HL12 #_ #H2 destruct /3 width=1 by frees_gref/
-| #f1V #f1T #f1 #p #I #L1 #V1 #T1 #_ #_ #Hf1 #IHV #IHT #X #H1 elim (tdeq_inv_pair1 … H1) -H1
- #V2 #T2 #HV12 #HT12 #H1 #L2 #HL12 destruct
- /6 width=5 by frees_bind, lexs_inv_tl, sle_lexs_trans, sor_inv_sle_dx, sor_inv_sle_sn/
-| #f1V #f1T #f1 #I #L1 #V1 #T1 #_ #_ #Hf1 #IHV #IHT #X #H1 elim (tdeq_inv_pair1 … H1) -H1
- #V2 #T2 #HV12 #HT12 #H1 #L2 #HL12 destruct
+[ #f #L1 #s1 #Hf #X #H1 #L2 #_
+ elim (tdeq_inv_sort1 … H1) -H1 #s2 #d #_ #_ #H destruct
+ /2 width=3 by frees_sort/
+| #f #i #Hf #X #H1
+ >(tdeq_inv_lref1 … H1) -X #Y #H2
+ >(lexs_inv_atom1 … H2) -Y
+ /2 width=1 by frees_atom/
+| #f #I #L1 #V1 #_ #IH #X #H1
+ >(tdeq_inv_lref1 … H1) -X #Y #H2
+ elim (lexs_inv_next1 … H2) -H2 #Z #L2 #HL12 #HZ #H destruct
+ elim (ext2_inv_pair_sn … HZ) -HZ #V2 #HV12 #H destruct
+ /3 width=1 by frees_pair/
+| #f #I #L1 #Hf #X #H1
+ >(tdeq_inv_lref1 … H1) -X #Y #H2
+ elim (lexs_inv_next1 … H2) -H2 #Z #L2 #_ #HZ #H destruct
+ >(ext2_inv_unit_sn … HZ) -Z /2 width=1 by frees_unit/
+| #f #I #L1 #i #_ #IH #X #H1
+ >(tdeq_inv_lref1 … H1) -X #Y #H2
+ elim (lexs_inv_push1 … H2) -H2 #J #L2 #HL12 #_ #H destruct
+ /3 width=1 by frees_lref/
+| #f #L1 #l #Hf #X #H1 #L2 #_
+ >(tdeq_inv_gref1 … H1) -X /2 width=1 by frees_gref/
+| #f1V #f1T #f1 #p #I #L1 #V1 #T1 #_ #_ #Hf1 #IHV #IHT #X #H1
+ elim (tdeq_inv_pair1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H1 #L2 #HL12 destruct
+ /6 width=5 by frees_bind, lexs_inv_tl, ext2_pair, sle_lexs_trans, sor_inv_sle_dx, sor_inv_sle_sn/
+| #f1V #f1T #f1 #I #L1 #V1 #T1 #_ #_ #Hf1 #IHV #IHT #X #H1
+ elim (tdeq_inv_pair1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H1 #L2 #HL12 destruct
/5 width=5 by frees_flat, sle_lexs_trans, sor_inv_sle_dx, sor_inv_sle_sn/
]
qed-.
lemma frees_tdeq_conf: ∀h,o,f,L,T1. L ⊢ 𝐅*⦃T1⦄ ≡ f →
∀T2. T1 ≡[h, o] T2 → L ⊢ 𝐅*⦃T2⦄ ≡ f.
-/3 width=7 by frees_tdeq_conf_lexs, lexs_refl/ qed-.
+/4 width=7 by frees_tdeq_conf_lexs, lexs_refl, ext2_refl/ qed-.
-lemma frees_lfdeq_conf_lexs: ∀h,o. lexs_frees_confluent (cdeq h o) cfull.
-/3 width=7 by frees_tdeq_conf_lexs, sle_refl, ex2_intro/ qed-.
+lemma frees_lexs_conf: ∀h,o,f,L1,T. L1 ⊢ 𝐅*⦃T⦄ ≡ f →
+ ∀L2. L1 ≡[h, o, f] L2 → L2 ⊢ 𝐅*⦃T⦄ ≡ f.
+/2 width=7 by frees_tdeq_conf_lexs, tdeq_refl/ qed-.
+
+lemma frees_lfdeq_conf_lexs: ∀h,o. lexs_frees_confluent (cdeq_ext h o) cfull.
+/3 width=7 by frees_tdeq_conf_lexs, ex2_intro/ qed-.
lemma tdeq_lfdeq_conf_sn: ∀h,o. s_r_confluent1 … (cdeq h o) (lfdeq h o).
#h #o #L1 #T1 #T2 #HT12 #L2 *
(* Basic_2A1: uses: lleq_sym *)
lemma lfdeq_sym: ∀h,o,T. symmetric … (lfdeq h o T).
#h #o #T #L1 #L2 *
-/4 width=7 by frees_tdeq_conf_lexs, lfxs_sym, tdeq_sym, sle_refl, ex2_intro/
+/4 width=7 by frees_tdeq_conf_lexs, lfxs_sym, tdeq_sym, ex2_intro/
qed-.
lemma lfdeq_atom: ∀h,o,I. ⋆ ≡[h, o, ⓪{I}] ⋆.
L1 ≡[h, o, ⋆s] L2 → L1.ⓑ{I}V1 ≡[h, o, ⋆s] L2.ⓑ{I}V2.
/2 width=1 by lfxs_sort/ qed.
-lemma lfdeq_zero: ∀h,o,I,L1,L2,V.
+lemma lfdeq_pair: ∀h,o,I,L1,L2,V.
L1 ≡[h, o, V] L2 → L1.ⓑ{I}V ≡[h, o, #0] L2.ⓑ{I}V.
-/2 width=1 by lfxs_zero/ qed.
+/2 width=1 by lfxs_pair/ qed.
+
+lemma lfdeq_unit: ∀h,o,f,I,L1,L2. 𝐈⦃f⦄ → L1 ⪤*[cdeq_ext h o, cfull, f] L2 →
+ L1.ⓤ{I} ≡[h, o, #0] L2.ⓤ{I}.
+/2 width=3 by lfxs_unit/ qed.
lemma lfdeq_lref: ∀h,o,I,L1,L2,V1,V2,i.
L1 ≡[h, o, #i] L2 → L1.ⓑ{I}V1 ≡[h, o, #⫯i] L2.ⓑ{I}V2.
L1 ≡[h, o, §l] L2 → L1.ⓑ{I}V1 ≡[h, o, §l] L2.ⓑ{I}V2.
/2 width=1 by lfxs_gref/ qed.
-lemma lfdeq_pair_repl_dx: ∀h,o,I,L1,L2.∀T:term.∀V,V1.
- L1.â\93\91{I}V â\89¡[h, o, T] L2.â\93\91{I}V1 →
- ∀V2. V ≡[h, o] V2 →
- L1.â\93\91{I}V â\89¡[h, o, T] L2.â\93\91{I}V2.
-/2 width=2 by lfxs_pair_repl_dx/ qed-.
+lemma lfdeq_bind_repl_dx: ∀h,o,I,I1,L1,L2.∀T:term.
+ L1.â\93\98{I} â\89¡[h, o, T] L2.â\93\98{I1} →
+ ∀I2. I ≡[h, o] I2 →
+ L1.â\93\98{I} â\89¡[h, o, T] L2.â\93\98{I2}.
+/2 width=2 by lfxs_bind_repl_dx/ qed-.
(* Basic inversion lemmas ***************************************************)
/2 width=3 by lfxs_inv_atom_dx/ qed-.
lemma lfdeq_inv_zero: ∀h,o,Y1,Y2. Y1 ≡[h, o, #0] Y2 →
- (Y1 = ⋆ ∧ Y2 = ⋆) ∨
- ∃∃I,L1,L2,V1,V2. L1 ≡[h, o, V1] L2 & V1 ≡[h, o] V2 &
- Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2.
+ ∨∨ Y1 = ⋆ ∧ Y2 = ⋆
+ | ∃∃I,L1,L2,V1,V2. L1 ≡[h, o, V1] L2 & V1 ≡[h, o] V2 &
+ Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2
+ | ∃∃f,I,L1,L2. 𝐈⦃f⦄ & L1 ⪤*[cdeq_ext h o, cfull, f] L2 &
+ Y1 = L1.ⓤ{I} & Y2 = L2.ⓤ{I}.
#h #o #Y1 #Y2 #H elim (lfxs_inv_zero … H) -H *
-/3 width=9 by ex4_5_intro, or_introl, or_intror, conj/
+/3 width=9 by or3_intro0, or3_intro1, or3_intro2, ex4_5_intro, ex4_4_intro, conj/
qed-.
lemma lfdeq_inv_lref: ∀h,o,Y1,Y2,i. Y1 ≡[h, o, #⫯i] Y2 →
- (Y1 = ⋆ ∧ Y2 = ⋆) ∨
- ∃∃I,L1,L2,V1,V2. L1 ≡[h, o, #i] L2 &
- Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2.
+ (Y1 = ⋆ ∧ Y2 = ⋆) ∨
+ ∃∃I1,I2,L1,L2. L1 ≡[h, o, #i] L2 &
+ Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}.
/2 width=1 by lfxs_inv_lref/ qed-.
(* Basic_2A1: uses: lleq_inv_bind lleq_inv_bind_O *)
lemma lfdeq_inv_zero_pair_sn: ∀h,o,I,Y2,L1,V1. L1.ⓑ{I}V1 ≡[h, o, #0] Y2 →
∃∃L2,V2. L1 ≡[h, o, V1] L2 & V1 ≡[h, o] V2 & Y2 = L2.ⓑ{I}V2.
-#h #o #I #Y2 #L1 #V1 #H elim (lfxs_inv_zero_pair_sn … H) -H /2 width=5 by ex3_2_intro/
-qed-.
+/2 width=1 by lfxs_inv_zero_pair_sn/ qed-.
lemma lfdeq_inv_zero_pair_dx: ∀h,o,I,Y1,L2,V2. Y1 ≡[h, o, #0] L2.ⓑ{I}V2 →
∃∃L1,V1. L1 ≡[h, o, V1] L2 & V1 ≡[h, o] V2 & Y1 = L1.ⓑ{I}V1.
-#h #o #I #Y1 #L2 #V2 #H elim (lfxs_inv_zero_pair_dx … H) -H
-#L1 #V1 #HL12 #HV12 #H destruct /2 width=5 by ex3_2_intro/
-qed-.
+/2 width=1 by lfxs_inv_zero_pair_dx/ qed-.
-lemma lfdeq_inv_lref_pair_sn: ∀h,o,I,Y2,L1,V1,i. L1.ⓑ{I}V1 ≡[h, o, #⫯i] Y2 →
- ∃∃L2,V2. L1 ≡[h, o, #i] L2 & Y2 = L2.ⓑ{I}V2.
-/2 width=2 by lfxs_inv_lref_pair_sn/ qed-.
+lemma lfdeq_inv_lref_bind_sn: ∀h,o,I1,Y2,L1,i. L1.ⓘ{I1} ≡[h, o, #⫯i] Y2 →
+ ∃∃I2,L2. L1 ≡[h, o, #i] L2 & Y2 = L2.ⓘ{I2}.
+/2 width=2 by lfxs_inv_lref_bind_sn/ qed-.
-lemma lfdeq_inv_lref_pair_dx: ∀h,o,I,Y1,L2,V2,i. Y1 ≡[h, o, #⫯i] L2.ⓑ{I}V2 →
- ∃∃L1,V1. L1 ≡[h, o, #i] L2 & Y1 = L1.ⓑ{I}V1.
-/2 width=2 by lfxs_inv_lref_pair_dx/ qed-.
+lemma lfdeq_inv_lref_bind_dx: ∀h,o,I2,Y1,L2,i. Y1 ≡[h, o, #⫯i] L2.ⓘ{I2} →
+ ∃∃I1,L1. L1 ≡[h, o, #i] L2 & Y1 = L1.ⓘ{I1}.
+/2 width=2 by lfxs_inv_lref_bind_dx/ qed-.
(* Basic forward lemmas *****************************************************)
lemma lfdeq_fwd_flat_dx: ∀h,o,I,L1,L2,V,T. L1 ≡[h, o, ⓕ{I}V.T] L2 → L1 ≡[h, o, T] L2.
/2 width=3 by lfxs_fwd_flat_dx/ qed-.
-lemma lfdeq_fwd_dx: ∀h,o,I,L1,K2,V2. ∀T:term. L1 ≡[h, o, T] K2.ⓑ{I}V2 →
- ∃∃K1,V1. L1 = K1.ⓑ{I}V1.
+lemma lfdeq_fwd_dx: ∀h,o,I2,L1,K2. ∀T:term. L1 ≡[h, o, T] K2.ⓘ{I2} →
+ ∃∃I1,K1. L1 = K1.ⓘ{I1}.
/2 width=5 by lfxs_fwd_dx/ qed-.
(* Basic_2A1: removed theorems 10:
include "basic_2/relocation/lifts_tdeq.ma".
include "basic_2/static/lfxs_drops.ma".
-include "basic_2/static/lfdeq_fqup.ma".
+include "basic_2/static/lfdeq.ma".
(* DEGREE-BASED EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES ******)
∀T. ⬆*[i] T ≡ U → K1 ≡[h, o, T] K2.
/2 width=8 by lfxs_inv_lifts_bi/ qed-.
-lemma lfdeq_inv_lref_sn: ∀h,o,L1,L2,i. L1 ≡[h, o, #i] L2 → ∀I,K1,V1. ⬇*[i] L1 ≡ K1.ⓑ{I}V1 →
- ∃∃K2,V2. ⬇*[i] L2 ≡ K2.ⓑ{I}V2 & K1 ≡[h, o, V1] K2 & V1 ≡[h, o] V2.
-/2 width=3 by lfxs_inv_lref_sn/ qed-.
+lemma lfdeq_inv_lref_pair_sn: ∀h,o,L1,L2,i. L1 ≡[h, o, #i] L2 → ∀I,K1,V1. ⬇*[i] L1 ≡ K1.ⓑ{I}V1 →
+ ∃∃K2,V2. ⬇*[i] L2 ≡ K2.ⓑ{I}V2 & K1 ≡[h, o, V1] K2 & V1 ≡[h, o] V2.
+/2 width=3 by lfxs_inv_lref_pair_sn/ qed-.
-lemma lfdeq_inv_lref_dx: ∀h,o,L1,L2,i. L1 ≡[h, o, #i] L2 → ∀I,K2,V2. ⬇*[i] L2 ≡ K2.ⓑ{I}V2 →
- ∃∃K1,V1. ⬇*[i] L1 ≡ K1.ⓑ{I}V1 & K1 ≡[h, o, V1] K2 & V1 ≡[h, o] V2.
-/2 width=3 by lfxs_inv_lref_dx/ qed-.
+lemma lfdeq_inv_lref_pair_dx: ∀h,o,L1,L2,i. L1 ≡[h, o, #i] L2 → ∀I,K2,V2. ⬇*[i] L2 ≡ K2.ⓑ{I}V2 →
+ ∃∃K1,V1. ⬇*[i] L1 ≡ K1.ⓑ{I}V1 & K1 ≡[h, o, V1] K2 & V1 ≡[h, o] V2.
+/2 width=3 by lfxs_inv_lref_pair_dx/ qed-.
lemma lfdeq_pair: ∀h,o,V1,V2. V1 ≡[h, o] V2 →
∀I,L. ∀T:term. L.ⓑ{I}V1 ≡[h, o, T] L.ⓑ{I}V2.
/2 width=1 by lfxs_pair/ qed.
+
+(* Advanced inversion lemmas ************************************************)
+
+lemma lfdeq_inv_bind_void: ∀h,o,p,I,L1,L2,V,T. L1 ≡[h, o, ⓑ{p,I}V.T] L2 →
+ L1 ≡[h, o, V] L2 ∧ L1.ⓧ ≡[h, o, T] L2.ⓧ.
+/2 width=3 by lfxs_inv_bind_void/ qed-.
+
+(* Advanced forward lemmas **************************************************)
+
+lemma lfdeq_fwd_bind_dx_void: ∀h,o,p,I,L1,L2,V,T.
+ L1 ≡[h, o, ⓑ{p,I}V.T] L2 → L1.ⓧ ≡[h, o, T] L2.ⓧ.
+/2 width=4 by lfxs_fwd_bind_dx_void/ qed-.
include "basic_2/s_computation/fqus_fqup.ma".
include "basic_2/static/lfdeq_drops.ma".
+include "basic_2/static/lfdeq_fqup.ma".
include "basic_2/static/lfdeq_lfdeq.ma".
(* DEGREE-BASED EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES ******)
(* Properties with supclosure ***********************************************)
-lemma fqu_tdeq_conf: ∀h,o,G1,G2,L1,L2,U1,T1. ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, T1⦄ →
+lemma fqu_tdeq_conf: ∀h,o,b,G1,G2,L1,L2,U1,T1. ⦃G1, L1, U1⦄ ⊐[b] ⦃G2, L2, T1⦄ →
∀U2. U1 ≡[h, o] U2 →
- ∃∃L,T2. ⦃G1, L1, U2⦄ ⊐ ⦃G2, L, T2⦄ & L2 ≡[h, o, T1] L & T1 ≡[h, o] T2.
-#h #o #G1 #G2 #L1 #L2 #U1 #T1 #H elim H -G1 -G2 -L1 -L2 -U1 -T1
-[ #I #G #L #W #X #H >(tdeq_inv_lref1 … H) -H
+ ∃∃L,T2. ⦃G1, L1, U2⦄ ⊐[b] ⦃G2, L, T2⦄ & L2 ≡[h, o, T1] L & T1 ≡[h, o] T2.
+#h #o #b #G1 #G2 #L1 #L2 #U1 #T1 #H elim H -G1 -G2 -L1 -L2 -U1 -T1
+[ #I #G #L #W #X #H >(tdeq_inv_lref1 … H) -X
/2 width=5 by fqu_lref_O, ex3_2_intro/
| #I #G #L #W1 #U1 #X #H
elim (tdeq_inv_pair1 … H) -H #W2 #U2 #HW12 #_ #H destruct
| #p #I #G #L #W1 #U1 #X #H
elim (tdeq_inv_pair1 … H) -H #W2 #U2 #HW12 #HU12 #H destruct
/3 width=5 by lfdeq_pair, fqu_bind_dx, ex3_2_intro/
+| #p #I #G #L #W1 #U1 #Hb #X #H
+ elim (tdeq_inv_pair1 … H) -H #W2 #U2 #HW12 #HU12 #H destruct
+ /3 width=5 by fqu_clear, ex3_2_intro/
| #I #G #L #W1 #U1 #X #H
elim (tdeq_inv_pair1 … H) -H #W2 #U2 #_ #HU12 #H destruct
/2 width=5 by fqu_flat_dx, ex3_2_intro/
-| #I #G #L #W #T1 #U1 #HTU1 #U2 #HU12
- elim (tdeq_inv_lifts_sn … HU12 … HTU1) -U1 #T2 #HTU2 #HT12
+| #I #G #L #T1 #U1 #HTU1 #U2 #HU12
+ elim (tdeq_inv_lifts_sn … HU12 … HTU1) -U1
/3 width=5 by fqu_drop, ex3_2_intro/
]
qed-.
-lemma tdeq_fqu_trans: ∀h,o,G1,G2,L1,L2,U1,T1. ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, T1⦄ →
+lemma tdeq_fqu_trans: ∀h,o,b,G1,G2,L1,L2,U1,T1. ⦃G1, L1, U1⦄ ⊐[b] ⦃G2, L2, T1⦄ →
∀U2. U2 ≡[h, o] U1 →
- ∃∃L,T2. ⦃G1, L1, U2⦄ ⊐ ⦃G2, L, T2⦄ & T2 ≡[h, o] T1 & L ≡[h, o, T1] L2.
-#h #o #G1 #G2 #L1 #L2 #U1 #T1 #H12 #U2 #HU21
+ ∃∃L,T2. ⦃G1, L1, U2⦄ ⊐[b] ⦃G2, L, T2⦄ & T2 ≡[h, o] T1 & L ≡[h, o, T1] L2.
+#h #o #b #G1 #G2 #L1 #L2 #U1 #T1 #H12 #U2 #HU21
elim (fqu_tdeq_conf … o … H12 U2) -H12
/3 width=5 by lfdeq_sym, tdeq_sym, ex3_2_intro/
qed-.
(* Basic_2A1: was just: lleq_fqu_trans *)
-lemma lfdeq_fqu_trans: ∀h,o,G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊐ ⦃G2, K2, U⦄ →
+lemma lfdeq_fqu_trans: ∀h,o,b,G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊐[b] ⦃G2, K2, U⦄ →
∀L1. L1 ≡[h, o, T] L2 →
- ∃∃K1,U0. ⦃G1, L1, T⦄ ⊐ ⦃G2, K1, U0⦄ & U0 ≡[h, o] U & K1 ≡[h, o, U] K2.
-#h #o #G1 #G2 #L2 #K2 #T #U #H elim H -G1 -G2 -L2 -K2 -T -U
+ ∃∃K1,U0. ⦃G1, L1, T⦄ ⊐[b] ⦃G2, K1, U0⦄ & U0 ≡[h, o] U & K1 ≡[h, o, U] K2.
+#h #o #b #G1 #G2 #L2 #K2 #T #U #H elim H -G1 -G2 -L2 -K2 -T -U
[ #I #G #L2 #V2 #L1 #H elim (lfdeq_inv_zero_pair_dx … H) -H
#K1 #V1 #HV1 #HV12 #H destruct
/3 width=7 by tdeq_lfdeq_conf_sn, fqu_lref_O, ex3_2_intro/
| elim (lfdeq_inv_flat … H)
] -H
/2 width=5 by fqu_pair_sn, ex3_2_intro/
-| #a #I #G #L2 #V #T #L1 #H elim (lfdeq_inv_bind … H) -H
+| #p #I #G #L2 #V #T #L1 #H elim (lfdeq_inv_bind … H) -H
/2 width=5 by fqu_bind_dx, ex3_2_intro/
+| #p #I #G #L2 #V #T #Hb #L1 #H elim (lfdeq_inv_bind_void … H) -H
+ /3 width=5 by fqu_clear, ex3_2_intro/
| #I #G #L2 #V #T #L1 #H elim (lfdeq_inv_flat … H) -H
/2 width=5 by fqu_flat_dx, ex3_2_intro/
-| #I #G #L2 #V2 #T #U #HTU #Y #HU
+| #I #G #L2 #T #U #HTU #Y #HU
elim (lfdeq_fwd_dx … HU) #L1 #V1 #H destruct
/5 width=12 by lfdeq_inv_lifts_bi, fqu_drop, drops_refl, drops_drop, ex3_2_intro/
]
qed-.
(* Basic_2A1: was just: lleq_fquq_trans *)
-lemma lfdeq_fquq_trans: ∀h,o,G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊐⸮ ⦃G2, K2, U⦄ →
+lemma lfdeq_fquq_trans: ∀h,o,b,G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊐⸮[b] ⦃G2, K2, U⦄ →
∀L1. L1 ≡[h, o, T] L2 →
- ∃∃K1,U0. ⦃G1, L1, T⦄ ⊐⸮ ⦃G2, K1, U0⦄ & U0 ≡[h, o] U & K1 ≡[h, o, U] K2.
-#h #o #G1 #G2 #L2 #K2 #T #U #H elim H -H
+ ∃∃K1,U0. ⦃G1, L1, T⦄ ⊐⸮[b] ⦃G2, K1, U0⦄ & U0 ≡[h, o] U & K1 ≡[h, o, U] K2.
+#h #o #b #G1 #G2 #L2 #K2 #T #U #H elim H -H
[ #H #L1 #HL12 elim (lfdeq_fqu_trans … H … HL12) -L2 /3 width=5 by fqu_fquq, ex3_2_intro/
| * #HG #HL #HT destruct /2 width=5 by ex3_2_intro/
]
qed-.
(* Basic_2A1: was just: lleq_fqup_trans *)
-lemma lfdeq_fqup_trans: ∀h,o,G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊐+ ⦃G2, K2, U⦄ →
+lemma lfdeq_fqup_trans: ∀h,o,b,G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊐+[b] ⦃G2, K2, U⦄ →
∀L1. L1 ≡[h, o, T] L2 →
- ∃∃K1,U0. ⦃G1, L1, T⦄ ⊐+ ⦃G2, K1, U0⦄ & U0 ≡[h, o] U & K1 ≡[h, o, U] K2.
-#h #o #G1 #G2 #L2 #K2 #T #U #H @(fqup_ind … H) -G2 -K2 -U
+ ∃∃K1,U0. ⦃G1, L1, T⦄ ⊐+[b] ⦃G2, K1, U0⦄ & U0 ≡[h, o] U & K1 ≡[h, o, U] K2.
+#h #o #b #G1 #G2 #L2 #K2 #T #U #H @(fqup_ind … H) -G2 -K2 -U
[ #G2 #K2 #U #HTU #L1 #HL12 elim (lfdeq_fqu_trans … HTU … HL12) -L2
/3 width=5 by fqu_fqup, ex3_2_intro/
| #G #G2 #K #K2 #U #U2 #_ #HU2 #IHTU #L1 #HL12
qed-.
(* Basic_2A1: was just: lleq_fqus_trans *)
-lemma lfdeq_fqus_trans: ∀h,o,G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊐* ⦃G2, K2, U⦄ →
+lemma lfdeq_fqus_trans: ∀h,o,b,G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊐*[b] ⦃G2, K2, U⦄ →
∀L1. L1 ≡[h, o, T] L2 →
- ∃∃K1,U0. ⦃G1, L1, T⦄ ⊐* ⦃G2, K1, U0⦄ & U0 ≡[h, o] U & K1 ≡[h, o, U] K2.
-#h #o #G1 #G2 #L2 #K2 #T #U #H #L1 #HL12 elim(fqus_inv_fqup … H) -H
+ ∃∃K1,U0. ⦃G1, L1, T⦄ ⊐*[b] ⦃G2, K1, U0⦄ & U0 ≡[h, o] U & K1 ≡[h, o, U] K2.
+#h #o #b #G1 #G2 #L2 #K2 #T #U #H #L1 #HL12 elim(fqus_inv_fqup … H) -H
[ #H elim (lfdeq_fqup_trans … H … HL12) -L2 /3 width=5 by fqup_fqus, ex3_2_intro/
| * #HG #HL #HT destruct /2 width=5 by ex3_2_intro/
]
(* *)
(**************************************************************************)
+include "basic_2/syntax/ext2_ext2.ma".
include "basic_2/syntax/tdeq_tdeq.ma".
include "basic_2/static/lfxs_lfxs.ma".
include "basic_2/static/lfdeq.ma".
L1 ≡[h, o, ⓕ{I}V.T] L2.
/2 width=1 by lfxs_flat/ qed.
+theorem lfdeq_bind_void: ∀h,o,p,I,L1,L2,V,T.
+ L1 ≡[h, o, V] L2 → L1.ⓧ ≡[h, o, T] L2.ⓧ →
+ L1 ≡[h, o, ⓑ{p,I}V.T] L2.
+/2 width=1 by lfxs_bind_void/ qed.
+
(* Basic_2A1: uses: lleq_trans *)
theorem lfdeq_trans: ∀h,o,T. Transitive … (lfdeq h o T).
#h #o #T #L1 #L * #f1 #Hf1 #HL1 #L2 * #f2 #Hf2 #HL2
lapply (frees_tdeq_conf_lexs … Hf1 T … HL1) // #H0
lapply (frees_mono … Hf2 … H0) -Hf2 -H0
-/4 width=7 by lexs_trans, lexs_eq_repl_back, tdeq_trans, ex2_intro/
+/5 width=7 by lexs_trans, lexs_eq_repl_back, tdeq_trans, ext2_trans, ex2_intro/
qed-.
(* Basic_2A1: uses: lleq_canc_sn *)
lemma lfdneq_inv_flat: ∀h,o,I,L1,L2,V,T. (L1 ≡[h, o, ⓕ{I}V.T] L2 → ⊥) →
(L1 ≡[h, o, V] L2 → ⊥) ∨ (L1 ≡[h, o, T] L2 → ⊥).
/3 width=2 by lfnxs_inv_flat, tdeq_dec/ qed-.
+
+lemma lfdneq_inv_bind_void: ∀h,o,p,I,L1,L2,V,T. (L1 ≡[h, o, ⓑ{p,I}V.T] L2 → ⊥) →
+ (L1 ≡[h, o, V] L2 → ⊥) ∨ (L1.ⓧ ≡[h, o, T] L2.ⓧ → ⊥).
+/3 width=3 by lfnxs_inv_bind_void, tdeq_dec/ qed-.
include "ground_2/relocation/rtmap_id.ma".
include "basic_2/notation/relations/relationstar_4.ma".
+include "basic_2/syntax/lenv_ext2.ma".
include "basic_2/relocation/lexs.ma".
include "basic_2/static/frees.ma".
(* GENERIC EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ****)
definition lfxs (R) (T): relation lenv ≝
- λL1,L2. ∃∃f. L1 ⊢ 𝐅*⦃T⦄ ≡ f & L1 ⪤*[R, cfull, f] L2.
+ λL1,L2. ∃∃f. L1 ⊢ 𝐅*⦃T⦄ ≡ f & L1 ⪤*[cext2 R, cfull, f] L2.
interpretation "generic extension on referred entries (local environment)"
'RelationStar R T L1 L2 = (lfxs R T L1 L2).
-definition R_frees_confluent: predicate (relation3 lenv term term) ≝
+definition R_frees_confluent: predicate (relation3 …) ≝
λRN.
∀f1,L,T1. L ⊢ 𝐅*⦃T1⦄ ≡ f1 → ∀T2. RN L T1 T2 →
∃∃f2. L ⊢ 𝐅*⦃T2⦄ ≡ f2 & f2 ⊆ f1.
-definition lexs_frees_confluent: relation (relation3 lenv term term) ≝
+definition lexs_frees_confluent: relation (relation3 …) ≝
λRN,RP.
∀f1,L1,T. L1 ⊢ 𝐅*⦃T⦄ ≡ f1 →
∀L2. L1 ⪤*[RN, RP, f1] L2 →
(* Basic properties *********************************************************)
lemma lfxs_atom: ∀R,I. ⋆ ⪤*[R, ⓪{I}] ⋆.
-/3 width=3 by lexs_atom, frees_atom, ex2_intro/
+#R * /3 width=3 by frees_sort, frees_atom, frees_gref, lexs_atom, ex2_intro/
qed.
(* Basic_2A1: uses: llpx_sn_sort *)
-lemma lfxs_sort: ∀R,I,L1,L2,V1,V2,s.
- L1 ⪤*[R, ⋆s] L2 → L1.ⓑ{I}V1 ⪤*[R, ⋆s] L2.ⓑ{I}V2.
-#R #I #L1 #L2 #V1 #V2 #s * /3 width=3 by lexs_push, frees_sort, ex2_intro/
+lemma lfxs_sort: ∀R,I1,I2,L1,L2,s.
+ L1 ⪤*[R, ⋆s] L2 → L1.ⓘ{I1} ⪤*[R, ⋆s] L2.ⓘ{I2}.
+#R #I1 #I2 #L1 #L2 #s * #f #Hf #H12
+lapply (frees_inv_sort … Hf) -Hf
+/4 width=3 by frees_sort, lexs_push, isid_push, ex2_intro/
qed.
-lemma lfxs_zero: ∀R,I,L1,L2,V1,V2. L1 ⪤*[R, V1] L2 →
+lemma lfxs_pair: ∀R,I,L1,L2,V1,V2. L1 ⪤*[R, V1] L2 →
R L1 V1 V2 → L1.ⓑ{I}V1 ⪤*[R, #0] L2.ⓑ{I}V2.
-#R #I #L1 #L2 #V1 #V2 * /3 width=3 by lexs_next, frees_zero, ex2_intro/
+#R #I1 #I2 #L1 #L2 #V1 *
+/4 width=3 by ext2_pair, frees_pair, lexs_next, ex2_intro/
qed.
-lemma lfxs_lref: ∀R,I,L1,L2,V1,V2,i.
- L1 ⪤*[R, #i] L2 → L1.ⓑ{I}V1 ⪤*[R, #⫯i] L2.ⓑ{I}V2.
-#R #I #L1 #L2 #V1 #V2 #i * /3 width=3 by lexs_push, frees_lref, ex2_intro/
+lemma lfxs_unit: ∀R,f,I,L1,L2. 𝐈⦃f⦄ → L1 ⪤*[cext2 R, cfull, f] L2 →
+ L1.ⓤ{I} ⪤*[R, #0] L2.ⓤ{I}.
+/4 width=3 by frees_unit, lexs_next, ext2_unit, ex2_intro/ qed.
+
+lemma lfxs_lref: ∀R,I1,I2,L1,L2,i.
+ L1 ⪤*[R, #i] L2 → L1.ⓘ{I1} ⪤*[R, #⫯i] L2.ⓘ{I2}.
+#R #I1 #I2 #L1 #L2 #i * /3 width=3 by lexs_push, frees_lref, ex2_intro/
qed.
(* Basic_2A1: uses: llpx_sn_gref *)
-lemma lfxs_gref: ∀R,I,L1,L2,V1,V2,l.
- L1 ⪤*[R, §l] L2 → L1.ⓑ{I}V1 ⪤*[R, §l] L2.ⓑ{I}V2.
-#R #I #L1 #L2 #V1 #V2 #l * /3 width=3 by lexs_push, frees_gref, ex2_intro/
+lemma lfxs_gref: ∀R,I1,I2,L1,L2,l.
+ L1 ⪤*[R, §l] L2 → L1.ⓘ{I1} ⪤*[R, §l] L2.ⓘ{I2}.
+#R #I1 #I2 #L1 #L2 #l * #f #Hf #H12
+lapply (frees_inv_gref … Hf) -Hf
+/4 width=3 by frees_gref, lexs_push, isid_push, ex2_intro/
qed.
-lemma lfxs_pair_repl_dx: ∀R,I,L1,L2,T,V,V1.
- L1.â\93\91{I}V ⪤*[R, T] L2.â\93\91{I}V1 →
- ∀V2. R L1 V V2 →
- L1.â\93\91{I}V ⪤*[R, T] L2.â\93\91{I}V2.
-#R #I #L1 #L2 #T #V #V1 * #f #Hf #HL12 #V2 #HR
+lemma lfxs_bind_repl_dx: ∀R,I,I1,L1,L2,T.
+ L1.â\93\98{I} ⪤*[R, T] L2.â\93\98{I1} →
+ ∀I2. cext2 R L1 I I2 →
+ L1.â\93\98{I} ⪤*[R, T] L2.â\93\98{I2}.
+#R #I #I1 #L1 #L2 #T * #f #Hf #HL12 #I2 #HR
/3 width=5 by lexs_pair_repl, ex2_intro/
qed-.
-lemma lfxs_sym: ∀R. lexs_frees_confluent R cfull →
+lemma lfxs_sym: ∀R. lexs_frees_confluent (cext2 R) cfull →
(∀L1,L2,T1,T2. R L1 T1 T2 → R L2 T2 T1) →
∀T. symmetric … (lfxs R T).
#R #H1R #H2R #T #L1 #L2 * #f1 #Hf1 #HL12 elim (H1R … Hf1 … HL12) -Hf1
-/4 width=5 by sle_lexs_trans, lexs_sym, ex2_intro/
+/5 width=5 by sle_lexs_trans, lexs_sym, cext2_sym, ex2_intro/
qed-.
(* Basic_2A1: uses: llpx_sn_co *)
lemma lfxs_co: ∀R1,R2. (∀L,T1,T2. R1 L T1 T2 → R2 L T1 T2) →
∀L1,L2,T. L1 ⪤*[R1, T] L2 → L1 ⪤*[R2, T] L2.
-#R1 #R2 #HR #L1 #L2 #T * /4 width=7 by lexs_co, ex2_intro/
+#R1 #R2 #HR #L1 #L2 #T * /5 width=7 by lexs_co, cext2_co, ex2_intro/
qed-.
lemma lfxs_isid: ∀R1,R2,L1,L2,T1,T2.
qed-.
lemma lfxs_inv_sort: ∀R,Y1,Y2,s. Y1 ⪤*[R, ⋆s] Y2 →
- (Y1 = ⋆ ∧ Y2 = ⋆) ∨
- ∃∃I,L1,L2,V1,V2. L1 ⪤*[R, ⋆s] L2 &
- Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2.
-#R * [ | #Y1 #I #V1 ] #Y2 #s * #f #H1 #H2
+ ∨∨ Y1 = ⋆ ∧ Y2 = ⋆
+ | ∃∃I1,I2,L1,L2. L1 ⪤*[R, ⋆s] L2 &
+ Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}.
+#R * [ | #Y1 #I1 ] #Y2 #s * #f #H1 #H2
[ lapply (lexs_inv_atom1 … H2) -H2 /3 width=1 by or_introl, conj/
| lapply (frees_inv_sort … H1) -H1 #Hf
elim (isid_inv_gen … Hf) -Hf #g #Hg #H destruct
- elim (lexs_inv_push1 … H2) -H2 #L2 #V2 #H12 #_ #H destruct
- /5 width=8 by frees_sort_gen, ex3_5_intro, ex2_intro, or_intror/
+ elim (lexs_inv_push1 … H2) -H2 #I2 #L2 #H12 #_ #H destruct
+ /5 width=7 by frees_sort, ex3_4_intro, ex2_intro, or_intror/
]
qed-.
lemma lfxs_inv_zero: ∀R,Y1,Y2. Y1 ⪤*[R, #0] Y2 →
- (Y1 = ⋆ ∧ Y2 = ⋆) ∨
- ∃∃I,L1,L2,V1,V2. L1 ⪤*[R, V1] L2 & R L1 V1 V2 &
- Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2.
-#R #Y1 #Y2 * #f #H1 #H2 elim (frees_inv_zero … H1) -H1 *
-[ #H #_ lapply (lexs_inv_atom1_aux … H2 H) -H2 /3 width=1 by or_introl, conj/
-| #I1 #L1 #V1 #g #HV1 #HY1 #Hg elim (lexs_inv_next1_aux … H2 … HY1 Hg) -H2 -Hg
- /4 width=9 by ex4_5_intro, ex2_intro, or_intror/
+ ∨∨ Y1 = ⋆ ∧ Y2 = ⋆
+ | ∃∃I,L1,L2,V1,V2. L1 ⪤*[R, V1] L2 & R L1 V1 V2 &
+ Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2
+ | ∃∃f,I,L1,L2. 𝐈⦃f⦄ & L1 ⪤*[cext2 R, cfull, f] L2 &
+ Y1 = L1.ⓤ{I} & Y2 = L2.ⓤ{I}.
+#R * [ | #Y1 * #I1 [ | #X ] ] #Y2 * #f #H1 #H2
+[ lapply (lexs_inv_atom1 … H2) -H2 /3 width=1 by or3_intro0, conj/
+| elim (frees_inv_unit … H1) -H1 #g #HX #H destruct
+ elim (lexs_inv_next1 … H2) -H2 #I2 #L2 #HL12 #H #H2 destruct
+ >(ext2_inv_unit_sn … H) -H /3 width=8 by or3_intro2, ex4_4_intro/
+| elim (frees_inv_pair … H1) -H1 #g #Hg #H destruct
+ elim (lexs_inv_next1 … H2) -H2 #Z2 #L2 #HL12 #H
+ elim (ext2_inv_pair_sn … H) -H
+ /4 width=9 by or3_intro1, ex4_5_intro, ex2_intro/
]
qed-.
lemma lfxs_inv_lref: ∀R,Y1,Y2,i. Y1 ⪤*[R, #⫯i] Y2 →
- (Y1 = ⋆ ∧ Y2 = ⋆) ∨
- ∃∃I,L1,L2,V1,V2. L1 ⪤*[R, #i] L2 &
- Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2.
-#R #Y1 #Y2 #i * #f #H1 #H2 elim (frees_inv_lref … H1) -H1 *
-[ #H #_ lapply (lexs_inv_atom1_aux … H2 H) -H2 /3 width=1 by or_introl, conj/
-| #I1 #L1 #V1 #g #HV1 #HY1 #Hg elim (lexs_inv_push1_aux … H2 … HY1 Hg) -H2 -Hg
- /4 width=8 by ex3_5_intro, ex2_intro, or_intror/
+ ∨∨ Y1 = ⋆ ∧ Y2 = ⋆
+ | ∃∃I1,I2,L1,L2. L1 ⪤*[R, #i] L2 &
+ Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}.
+#R * [ | #Y1 #I1 ] #Y2 #i * #f #H1 #H2
+[ lapply (lexs_inv_atom1 … H2) -H2 /3 width=1 by or_introl, conj/
+| elim (frees_inv_lref … H1) -H1 #g #Hg #H destruct
+ elim (lexs_inv_push1 … H2) -H2
+ /4 width=7 by ex3_4_intro, ex2_intro, or_intror/
]
qed-.
lemma lfxs_inv_gref: ∀R,Y1,Y2,l. Y1 ⪤*[R, §l] Y2 →
- (Y1 = ⋆ ∧ Y2 = ⋆) ∨
- ∃∃I,L1,L2,V1,V2. L1 ⪤*[R, §l] L2 &
- Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2.
-#R * [ | #Y1 #I #V1 ] #Y2 #l * #f #H1 #H2
+ ∨∨ Y1 = ⋆ ∧ Y2 = ⋆
+ | ∃∃I1,I2,L1,L2. L1 ⪤*[R, §l] L2 &
+ Y1 = L1.ⓘ{I1} & Y2 = L2.ⓘ{I2}.
+#R * [ | #Y1 #I1 ] #Y2 #l * #f #H1 #H2
[ lapply (lexs_inv_atom1 … H2) -H2 /3 width=1 by or_introl, conj/
| lapply (frees_inv_gref … H1) -H1 #Hf
elim (isid_inv_gen … Hf) -Hf #g #Hg #H destruct
- elim (lexs_inv_push1 … H2) -H2 #L2 #V2 #H12 #_ #H destruct
- /5 width=8 by frees_gref_gen, ex3_5_intro, ex2_intro, or_intror/
+ elim (lexs_inv_push1 … H2) -H2 #I2 #L2 #H12 #_ #H destruct
+ /5 width=7 by frees_gref, ex3_4_intro, ex2_intro, or_intror/
]
qed-.
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 #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/
+/6 width=6 by sle_lexs_trans, lexs_inv_tl, ext2_pair, sor_inv_sle_dx, sor_inv_sle_sn, ex2_intro, conj/
qed-.
(* Basic_2A1: uses: llpx_sn_inv_flat *)
(* Advanced inversion lemmas ************************************************)
-lemma lfxs_inv_sort_pair_sn: ∀R,I,Y2,L1,V1,s. L1.ⓑ{I}V1 ⪤*[R, ⋆s] Y2 →
- ∃∃L2,V2. L1 ⪤*[R, ⋆s] L2 & Y2 = L2.ⓑ{I}V2.
-#R #I #Y2 #L1 #V1 #s #H elim (lfxs_inv_sort … H) -H *
+lemma lfxs_inv_sort_bind_sn: ∀R,I1,K1,L2,s. K1.ⓘ{I1} ⪤*[R, ⋆s] L2 →
+ ∃∃I2,K2. K1 ⪤*[R, ⋆s] K2 & L2 = K2.ⓘ{I2}.
+#R #I1 #K1 #L2 #s #H elim (lfxs_inv_sort … H) -H *
[ #H destruct
-| #J #Y1 #L2 #X1 #V2 #Hs #H1 #H2 destruct /2 width=4 by ex2_2_intro/
+| #Z1 #I2 #Y1 #K2 #Hs #H1 #H2 destruct /2 width=4 by ex2_2_intro/
]
qed-.
-lemma lfxs_inv_sort_pair_dx: ∀R,I,Y1,L2,V2,s. Y1 ⪤*[R, ⋆s] L2.ⓑ{I}V2 →
- ∃∃L1,V1. L1 ⪤*[R, ⋆s] L2 & Y1 = L1.ⓑ{I}V1.
-#R #I #Y1 #L2 #V2 #s #H elim (lfxs_inv_sort … H) -H *
+lemma lfxs_inv_sort_bind_dx: ∀R,I2,K2,L1,s. L1 ⪤*[R, ⋆s] K2.ⓘ{I2} →
+ ∃∃I1,K1. K1 ⪤*[R, ⋆s] K2 & L1 = K1.ⓘ{I1}.
+#R #I2 #K2 #L1 #s #H elim (lfxs_inv_sort … H) -H *
[ #_ #H destruct
-| #J #L1 #Y2 #V1 #X2 #Hs #H1 #H2 destruct /2 width=4 by ex2_2_intro/
+| #I1 #Z2 #K1 #Y2 #Hs #H1 #H2 destruct /2 width=4 by ex2_2_intro/
]
qed-.
-lemma lfxs_inv_zero_pair_sn: ∀R,I,Y2,L1,V1. L1.ⓑ{I}V1 ⪤*[R, #0] Y2 →
- ∃∃L2,V2. L1 ⪤*[R, V1] L2 & R L1 V1 V2 &
- Y2 = L2.ⓑ{I}V2.
-#R #I #Y2 #L1 #V1 #H elim (lfxs_inv_zero … H) -H *
+lemma lfxs_inv_zero_pair_sn: ∀R,I,L2,K1,V1. K1.ⓑ{I}V1 ⪤*[R, #0] L2 →
+ ∃∃K2,V2. K1 ⪤*[R, V1] K2 & R K1 V1 V2 &
+ L2 = K2.ⓑ{I}V2.
+#R #I #L2 #K1 #V1 #H elim (lfxs_inv_zero … H) -H *
[ #H destruct
-| #J #Y1 #L2 #X1 #V2 #HV1 #HV12 #H1 #H2 destruct
+| #Z #Y1 #K2 #X1 #V2 #HK12 #HV12 #H1 #H2 destruct
/2 width=5 by ex3_2_intro/
+| #f #Z #Y1 #Y2 #_ #_ #H destruct
]
qed-.
-lemma lfxs_inv_zero_pair_dx: ∀R,I,Y1,L2,V2. Y1 ⪤*[R, #0] L2.ⓑ{I}V2 →
- ∃∃L1,V1. L1 ⪤*[R, V1] L2 & R L1 V1 V2 &
- Y1 = L1.ⓑ{I}V1.
-#R #I #Y1 #L2 #V2 #H elim (lfxs_inv_zero … H) -H *
+lemma lfxs_inv_zero_pair_dx: ∀R,I,L1,K2,V2. L1 ⪤*[R, #0] K2.ⓑ{I}V2 →
+ ∃∃K1,V1. K1 ⪤*[R, V1] K2 & R K1 V1 V2 &
+ L1 = K1.ⓑ{I}V1.
+#R #I #L1 #K2 #V2 #H elim (lfxs_inv_zero … H) -H *
[ #_ #H destruct
-| #J #L1 #Y2 #V1 #X2 #HV1 #HV12 #H1 #H2 destruct
+| #Z #K1 #Y2 #V1 #X2 #HK12 #HV12 #H1 #H2 destruct
/2 width=5 by ex3_2_intro/
+| #f #Z #Y1 #Y2 #_ #_ #_ #H destruct
+]
+qed-.
+
+lemma lfxs_inv_zero_unit_sn: ∀R,I,K1,L2. K1.ⓤ{I} ⪤*[R, #0] L2 →
+ ∃∃f,K2. 𝐈⦃f⦄ & K1 ⪤*[cext2 R, cfull, f] K2 &
+ L2 = K2.ⓤ{I}.
+#R #I #K1 #L2 #H elim (lfxs_inv_zero … H) -H *
+[ #H destruct
+| #Z #Y1 #Y2 #X1 #X2 #_ #_ #H destruct
+| #f #Z #Y1 #K2 #Hf #HK12 #H1 #H2 destruct /2 width=5 by ex3_2_intro/
+]
+qed-.
+
+lemma lfxs_inv_zero_unit_dx: ∀R,I,L1,K2. L1 ⪤*[R, #0] K2.ⓤ{I} →
+ ∃∃f,K1. 𝐈⦃f⦄ & K1 ⪤*[cext2 R, cfull, f] K2 &
+ L1 = K1.ⓤ{I}.
+#R #I #L1 #K2 #H elim (lfxs_inv_zero … H) -H *
+[ #_ #H destruct
+| #Z #Y1 #Y2 #X1 #X2 #_ #_ #_ #H destruct
+| #f #Z #K1 #Y2 #Hf #HK12 #H1 #H2 destruct /2 width=5 by ex3_2_intro/
]
qed-.
-lemma lfxs_inv_lref_pair_sn: ∀R,I,Y2,L1,V1,i. L1.ⓑ{I}V1 ⪤*[R, #⫯i] Y2 →
- ∃∃L2,V2. L1 ⪤*[R, #i] L2 & Y2 = L2.ⓑ{I}V2.
-#R #I #Y2 #L1 #V1 #i #H elim (lfxs_inv_lref … H) -H *
+lemma lfxs_inv_lref_bind_sn: ∀R,I1,K1,L2,i. K1.ⓘ{I1} ⪤*[R, #⫯i] L2 →
+ ∃∃I2,K2. K1 ⪤*[R, #i] K2 & L2 = K2.ⓘ{I2}.
+#R #I1 #K1 #L2 #i #H elim (lfxs_inv_lref … H) -H *
[ #H destruct
-| #J #Y1 #L2 #X1 #V2 #Hi #H1 #H2 destruct /2 width=4 by ex2_2_intro/
+| #Z1 #I2 #Y1 #K2 #Hi #H1 #H2 destruct /2 width=4 by ex2_2_intro/
]
qed-.
-lemma lfxs_inv_lref_pair_dx: ∀R,I,Y1,L2,V2,i. Y1 ⪤*[R, #⫯i] L2.ⓑ{I}V2 →
- ∃∃L1,V1. L1 ⪤*[R, #i] L2 & Y1 = L1.ⓑ{I}V1.
-#R #I #Y1 #L2 #V2 #i #H elim (lfxs_inv_lref … H) -H *
+lemma lfxs_inv_lref_bind_dx: ∀R,I2,K2,L1,i. L1 ⪤*[R, #⫯i] K2.ⓘ{I2} →
+ ∃∃I1,K1. K1 ⪤*[R, #i] K2 & L1 = K1.ⓘ{I1}.
+#R #I2 #K2 #L1 #i #H elim (lfxs_inv_lref … H) -H *
[ #_ #H destruct
-| #J #L1 #Y2 #V1 #X2 #Hi #H1 #H2 destruct /2 width=4 by ex2_2_intro/
+| #I1 #Z2 #K1 #Y2 #Hi #H1 #H2 destruct /2 width=4 by ex2_2_intro/
]
qed-.
-lemma lfxs_inv_gref_pair_sn: ∀R,I,Y2,L1,V1,l. L1.ⓑ{I}V1 ⪤*[R, §l] Y2 →
- ∃∃L2,V2. L1 ⪤*[R, §l] L2 & Y2 = L2.ⓑ{I}V2.
-#R #I #Y2 #L1 #V1 #l #H elim (lfxs_inv_gref … H) -H *
+lemma lfxs_inv_gref_bind_sn: ∀R,I1,K1,L2,l. K1.ⓘ{I1} ⪤*[R, §l] L2 →
+ ∃∃I2,K2. K1 ⪤*[R, §l] K2 & L2 = K2.ⓘ{I2}.
+#R #I1 #K1 #L2 #l #H elim (lfxs_inv_gref … H) -H *
[ #H destruct
-| #J #Y1 #L2 #X1 #V2 #Hl #H1 #H2 destruct /2 width=4 by ex2_2_intro/
+| #Z1 #I2 #Y1 #K2 #Hl #H1 #H2 destruct /2 width=4 by ex2_2_intro/
]
qed-.
-lemma lfxs_inv_gref_pair_dx: ∀R,I,Y1,L2,V2,l. Y1 ⪤*[R, §l] L2.ⓑ{I}V2 →
- ∃∃L1,V1. L1 ⪤*[R, §l] L2 & Y1 = L1.ⓑ{I}V1.
-#R #I #Y1 #L2 #V2 #l #H elim (lfxs_inv_gref … H) -H *
+lemma lfxs_inv_gref_bind_dx: ∀R,I2,K2,L1,l. L1 ⪤*[R, §l] K2.ⓘ{I2} →
+ ∃∃I1,K1. K1 ⪤*[R, §l] K2 & L1 = K1.ⓘ{I1}.
+#R #I2 #K2 #L1 #l #H elim (lfxs_inv_gref … H) -H *
[ #_ #H destruct
-| #J #L1 #Y2 #V1 #X2 #Hl #H1 #H2 destruct /2 width=4 by ex2_2_intro/
+| #I1 #Z2 #K1 #Y2 #Hl #H1 #H2 destruct /2 width=4 by ex2_2_intro/
]
qed-.
#R #I #L1 #L2 #V #T #H elim (lfxs_inv_flat … H) -H //
qed-.
-lemma lfxs_fwd_dx: ∀R,I,L1,K2,T,V2. L1 ⪤*[R, T] K2.ⓑ{I}V2 →
- ∃∃K1,V1. L1 = K1.ⓑ{I}V1.
-#R #I #L1 #K2 #T #V2 * #f elim (pn_split f) * #g #Hg #_ #Hf destruct
-[ elim (lexs_inv_push2 … Hf) | elim (lexs_inv_next2 … Hf) ] -Hf #K1 #V1 #_ #_ #H destruct
+lemma lfxs_fwd_dx: ∀R,I2,L1,K2,T. L1 ⪤*[R, T] K2.ⓘ{I2} →
+ ∃∃I1,K1. L1 = K1.ⓘ{I1}.
+#R #I2 #L1 #K2 #T * #f elim (pn_split f) * #g #Hg #_ #Hf destruct
+[ elim (lexs_inv_push2 … Hf) | elim (lexs_inv_next2 … Hf) ] -Hf #I1 #K1 #_ #_ #H destruct
/2 width=3 by ex1_2_intro/
qed-.
(* *)
(**************************************************************************)
-include "basic_2/relocation/drops_ceq.ma".
include "basic_2/relocation/drops_lexs.ma".
+include "basic_2/relocation/drops_ext2.ma".
include "basic_2/static/frees_drops.ma".
include "basic_2/static/lfxs.ma".
(* Basic_2A1: includes: llpx_sn_lift_le llpx_sn_lift_ge *)
lemma lfxs_liftable_dedropable_sn: ∀R. (∀L. reflexive ? (R L)) →
- d_liftable2_sn R → dedropable_sn R.
+ d_liftable2_sn … lifts R → dedropable_sn R.
#R #H1R #H2R #b #f #L1 #K1 #HLK1 #K2 #T * #f1 #Hf1 #HK12 #U #HTU
elim (frees_total L1 U) #f2 #Hf2
lapply (frees_fwd_coafter … Hf2 … HLK1 … HTU … Hf1) -HTU #Hf
-elim (lexs_liftable_co_dedropable_sn … H1R … H2R … HLK1 … HK12 … Hf) -f1 -K1
-/3 width=6 by cfull_lift_sn, ex3_intro, ex2_intro/
+elim (lexs_liftable_co_dedropable_sn … HLK1 … HK12 … Hf) -f1 -K1
+/3 width=6 by cext2_d_liftable2_sn, cfull_lift_sn, ext2_refl, ex3_intro, ex2_intro/
qed-.
(* Inversion lemmas with generic slicing for local environments *************)
lapply (drops_mono … HY … HLK2) -L2 -i #H destruct //
qed-.
-lemma lfxs_inv_lref_sn: ∀R,L1,L2,i. L1 ⪤*[R, #i] L2 → ∀I,K1,V1. ⬇*[i] L1 ≡ K1.ⓑ{I}V1 →
- ∃∃K2,V2. ⬇*[i] L2 ≡ K2.ⓑ{I}V2 & K1 ⪤*[R, V1] K2 & R K1 V1 V2.
+lemma lfxs_inv_lref_pair_sn: ∀R,L1,L2,i. L1 ⪤*[R, #i] L2 → ∀I,K1,V1. ⬇*[i] L1 ≡ K1.ⓑ{I}V1 →
+ ∃∃K2,V2. ⬇*[i] L2 ≡ K2.ⓑ{I}V2 & K1 ⪤*[R, V1] K2 & R K1 V1 V2.
#R #L1 #L2 #i #HL12 #I #K1 #V1 #HLK1 elim (lfxs_dropable_sn … HLK1 … HL12 (#0)) -HLK1 -HL12 //
#Y #HY #HLK2 elim (lfxs_inv_zero_pair_sn … HY) -HY
#K2 #V2 #HK12 #HV12 #H destruct /2 width=5 by ex3_2_intro/
qed-.
-lemma lfxs_inv_lref_dx: ∀R,L1,L2,i. L1 ⪤*[R, #i] L2 → ∀I,K2,V2. ⬇*[i] L2 ≡ K2.ⓑ{I}V2 →
- ∃∃K1,V1. ⬇*[i] L1 ≡ K1.ⓑ{I}V1 & K1 ⪤*[R, V1] K2 & R K1 V1 V2.
+lemma lfxs_inv_lref_pair_dx: ∀R,L1,L2,i. L1 ⪤*[R, #i] L2 → ∀I,K2,V2. ⬇*[i] L2 ≡ K2.ⓑ{I}V2 →
+ ∃∃K1,V1. ⬇*[i] L1 ≡ K1.ⓑ{I}V1 & K1 ⪤*[R, V1] K2 & R K1 V1 V2.
#R #L1 #L2 #i #HL12 #I #K2 #V2 #HLK2 elim (lfxs_dropable_dx … HL12 … HLK2 … (#0)) -HLK2 -HL12 //
#Y #HLK1 #HY elim (lfxs_inv_zero_pair_dx … HY) -HY
#K1 #V1 #HK12 #HV12 #H destruct /2 width=5 by ex3_2_intro/
qed-.
+
+lemma lfxs_inv_lref_unit_sn: ∀R,L1,L2,i. L1 ⪤*[R, #i] L2 → ∀I,K1. ⬇*[i] L1 ≡ K1.ⓤ{I} →
+ ∃∃f,K2. ⬇*[i] L2 ≡ K2.ⓤ{I} & K1 ⪤*[cext2 R, cfull, f] K2 & 𝐈⦃f⦄.
+#R #L1 #L2 #i #HL12 #I #K1 #HLK1 elim (lfxs_dropable_sn … HLK1 … HL12 (#0)) -HLK1 -HL12 //
+#Y #HY #HLK2 elim (lfxs_inv_zero_unit_sn … HY) -HY
+#f #K2 #Hf #HK12 #H destruct /2 width=5 by ex3_2_intro/
+qed-.
+
+lemma lfxs_inv_lref_unit_dx: ∀R,L1,L2,i. L1 ⪤*[R, #i] L2 → ∀I,K2. ⬇*[i] L2 ≡ K2.ⓤ{I} →
+ ∃∃f,K1. ⬇*[i] L1 ≡ K1.ⓤ{I} & K1 ⪤*[cext2 R, cfull, f] K2 & 𝐈⦃f⦄.
+#R #L1 #L2 #i #HL12 #I #K2 #HLK2 elim (lfxs_dropable_dx … HL12 … HLK2 … (#0)) -HLK2 -HL12 //
+#Y #HLK1 #HY elim (lfxs_inv_zero_unit_dx … HY) -HY
+#f #K2 #Hf #HK12 #H destruct /2 width=5 by ex3_2_intro/
+qed-.
(* Basic_2A1: uses: llpx_sn_refl *)
lemma lfxs_refl: ∀R. (∀L. reflexive … (R L)) → ∀L,T. L ⪤*[R, T] L.
-#R #HR #L #T elim (frees_total L T) /3 width=3 by lexs_refl, ex2_intro/
+#R #HR #L #T elim (frees_total L T)
+/4 width=3 by lexs_refl, ext2_refl, ex2_intro/
qed.
lemma lfxs_pair: ∀R. (∀L. reflexive … (R L)) →
#R #HR #L #V1 #V2 #HV12 #I #T
elim (frees_total (L.ⓑ{I}V1) T) #f #Hf
elim (pn_split f) * #g #H destruct
-/4 width=3 by lexs_refl, lexs_next, lexs_push, ex2_intro/
+/5 width=3 by lexs_refl, lexs_next, lexs_push, ext2_refl, ext2_pair, ex2_intro/
qed.
+
+(* Advanced inversion lemmas ************************************************)
+
+lemma lfxs_inv_bind_void: ∀R,p,I,L1,L2,V,T. L1 ⪤*[R, ⓑ{p,I}V.T] L2 →
+ L1 ⪤*[R, V] L2 ∧ L1.ⓧ ⪤*[R, T] L2.ⓧ.
+#R #p #I #L1 #L2 #V #T * #f #Hf #HL elim (frees_inv_bind_void … Hf) -Hf
+/6 width=6 by sle_lexs_trans, lexs_inv_tl, sor_inv_sle_dx, sor_inv_sle_sn, ex2_intro, conj/
+qed-.
+
+(* Advanced forward lemmas **************************************************)
+
+lemma lfxs_fwd_bind_dx_void: ∀R,p,I,L1,L2,V,T. L1 ⪤*[R, ⓑ{p,I}V.T] L2 →
+ L1.ⓧ ⪤*[R, T] L2.ⓧ.
+#R #p #I #L1 #L2 #V #T #H elim (lfxs_inv_bind_void … H) -H //
+qed-.
(**************************************************************************)
include "basic_2/relocation/lexs_lexs.ma".
-include "basic_2/static/frees_fqup.ma".
-include "basic_2/static/frees_frees.ma".
+include "basic_2/static/frees_drops.ma".
include "basic_2/static/lfxs.ma".
(* GENERIC EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ****)
(* Advanced properties ******************************************************)
lemma lfxs_inv_frees: ∀R,L1,L2,T. L1 ⪤*[R, T] L2 →
- ∀f. L1 ⊢ 𝐅*⦃T⦄ ≡ f → L1 ⪤*[R, cfull, f] L2.
+ ∀f. L1 ⊢ 𝐅*⦃T⦄ ≡ f → L1 ⪤*[cext2 R, cfull, f] L2.
#R #L1 #L2 #T * /3 width=6 by frees_mono, lexs_eq_repl_back/
qed-.
∀L1,L2,T. Decidable (L1 ⪤*[R, T] L2).
#R #HR #L1 #L2 #T
elim (frees_total L1 T) #f #Hf
-elim (lexs_dec R cfull HR … L1 L2 f)
-/4 width=3 by lfxs_inv_frees, cfull_dec, ex2_intro, or_intror, or_introl/
+elim (lexs_dec (cext2 R) cfull … L1 L2 f)
+/4 width=3 by lfxs_inv_frees, cfull_dec, ext2_dec, ex2_intro, or_intror, or_introl/
qed-.
lemma lfxs_pair_sn_split: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) →
- lexs_frees_confluent … R1 cfull →
+ lexs_frees_confluent … (cext2 R1) cfull →
∀L1,L2,V. L1 ⪤*[R1, V] L2 → ∀I,T.
∃∃L. L1 ⪤*[R1, ②{I}V.T] L & L ⪤*[R2, V] L2.
#R1 #R2 #HR1 #HR2 #HR #L1 #L2 #V * #f #Hf #HL12 * [ #p ] #I #T
lapply(frees_mono … H … Hf) -H #H1
lapply (sor_eq_repl_back1 … Hy … H1) -y1 #Hy
lapply (sor_inv_sle_sn … Hy) -y2 #Hfg
-elim (lexs_sle_split … HR1 HR2 … HL12 … Hfg) -HL12 #L #HL1 #HL2
+elim (lexs_sle_split (cext2 R1) (cext2 R2) … HL12 … Hfg) -HL12 /2 width=1 by ext2_refl/ #L #HL1 #HL2
lapply (sle_lexs_trans … HL1 … Hfg) // #H
elim (HR … Hf … H) -HR -Hf -H
/4 width=7 by sle_lexs_trans, ex2_intro/
qed-.
lemma lfxs_flat_dx_split: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) →
- lexs_frees_confluent … R1 cfull →
+ lexs_frees_confluent … (cext2 R1) cfull →
∀L1,L2,T. L1 ⪤*[R1, T] L2 → ∀I,V.
∃∃L. L1 ⪤*[R1, ⓕ{I}V.T] L & L ⪤*[R2, T] L2.
#R1 #R2 #HR1 #HR2 #HR #L1 #L2 #T * #f #Hf #HL12 #I #V
lapply(frees_mono … H … Hf) -H #H2
lapply (sor_eq_repl_back2 … Hy … H2) -y2 #Hy
lapply (sor_inv_sle_dx … Hy) -y1 #Hfg
-elim (lexs_sle_split … HR1 HR2 … HL12 … Hfg) -HL12 #L #HL1 #HL2
+elim (lexs_sle_split (cext2 R1) (cext2 R2) … HL12 … Hfg) -HL12 /2 width=1 by ext2_refl/ #L #HL1 #HL2
lapply (sle_lexs_trans … HL1 … Hfg) // #H
elim (HR … Hf … H) -HR -Hf -H
/4 width=7 by sle_lexs_trans, ex2_intro/
qed-.
lemma lfxs_bind_dx_split: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) →
- lexs_frees_confluent … R1 cfull →
+ lexs_frees_confluent … (cext2 R1) cfull →
∀I,L1,L2,V1,T. L1.ⓑ{I}V1 ⪤*[R1, T] L2 → ∀p.
∃∃L,V. L1 ⪤*[R1, ⓑ{p,I}V1.T] L & L.ⓑ{I}V ⪤*[R2, T] L2 & R1 L1 V1 V.
#R1 #R2 #HR1 #HR2 #HR #I #L1 #L2 #V1 #T * #f #Hf #HL12 #p
lapply (sor_eq_repl_back2 … Hy … H2) -y2 #Hy
lapply (sor_inv_sle_dx … Hy) -y1 #Hfg
lapply (sle_inv_tl_sn … Hfg) -Hfg #Hfg
-elim (lexs_sle_split … HR1 HR2 … HL12 … Hfg) -HL12 #Y #H #HL2
+elim (lexs_sle_split (cext2 R1) (cext2 R2) … HL12 … Hfg) -HL12 /2 width=1 by ext2_refl/ #Y #H #HL2
lapply (sle_lexs_trans … H … Hfg) // #H0
-elim (lexs_inv_next1 … H) -H #L #V #HL1 #HV0 #H destruct
+elim (lexs_inv_next1 … H) -H #Z #L #HL1 #H
+elim (ext2_inv_pair_sn … H) -H #V #HV #H1 #H2 destruct
elim (HR … Hf … H0) -HR -Hf -H0
/4 width=7 by sle_lexs_trans, ex3_2_intro, ex2_intro/
qed-.
L1 ⪤*[R, V1] L2 → L1.ⓑ{I}V1 ⪤*[R, T] L2.ⓑ{I}V2 →
L1 ⪤*[R, ⓑ{p,I}V1.T] L2.
#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))
+lapply (lexs_fwd_bind … 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.
/3 width=7 by frees_fwd_isfin, frees_flat, lexs_join, ex2_intro/
qed.
+theorem lfxs_bind_void: ∀R,p,I,L1,L2,V,T.
+ L1 ⪤*[R, V] L2 → L1.ⓧ ⪤*[R, T] L2.ⓧ →
+ L1 ⪤*[R, ⓑ{p,I}V.T] L2.
+#R #p #I #L1 #L2 #V #T * #f1 #HV #Hf1 * #f2 #HT #Hf2
+lapply (lexs_fwd_bind … Hf2) -Hf2 #Hf2 elim (sor_isfin_ex f1 (⫱f2))
+/3 width=7 by frees_fwd_isfin, frees_bind_void, lexs_join, isfin_tl, ex2_intro/
+qed.
+
theorem lfxs_conf: ∀R1,R2.
- lexs_frees_confluent R1 cfull →
- lexs_frees_confluent R2 cfull →
+ lexs_frees_confluent (cext2 R1) cfull →
+ lexs_frees_confluent (cext2 R2) cfull →
R_confluent2_lfxs R1 R2 R1 R2 →
∀T. confluent2 … (lfxs R1 T) (lfxs R2 T).
#R1 #R2 #HR1 #HR2 #HR12 #T #L0 #L1 * #f1 #Hf1 #HL01 #L2 * #f #Hf #HL02
lapply (sle_lexs_trans … HL1 … H1) // -HL1 -H1 #HL1
lapply (sle_lexs_trans … HL2 … H2) // -HL2 -H2 #HL2
/3 width=5 by ex2_intro/
-| #g #I #K0 #V0 #n #HLK0 #Hgf #V1 #HV01 #V2 #HV02 #K1 #HK01 #K2 #HK02
- elim (frees_drops_next … Hf … HLK0 … Hgf) -Hf -HLK0 -Hgf #g0 #Hg0 #H0
- lapply (sle_lexs_trans … HK01 … H0) // -HK01 #HK01
- lapply (sle_lexs_trans … HK02 … H0) // -HK02 #HK02
- elim (HR12 … HV01 … HV02 K1 … K2) /2 width=3 by ex2_intro/
+| #g * #I0 [2: #V0 ] #K0 #n #HLK0 #Hgf #Z1 #H1 #Z2 #H2 #K1 #HK01 #K2 #HK02
+ [ elim (ext2_inv_pair_sn … H1) -H1 #V1 #HV01 #H destruct
+ elim (ext2_inv_pair_sn … H2) -H2 #V2 #HV02 #H destruct
+ elim (frees_inv_drops_next … Hf … HLK0 … Hgf) -Hf -HLK0 -Hgf #g0 #Hg0 #H0
+ lapply (sle_lexs_trans … HK01 … H0) // -HK01 #HK01
+ lapply (sle_lexs_trans … HK02 … H0) // -HK02 #HK02
+ elim (HR12 … HV01 … HV02 K1 … K2) /3 width=3 by ext2_pair, ex2_intro/
+ | lapply (ext2_inv_unit_sn … H1) -H1 #H destruct
+ lapply (ext2_inv_unit_sn … H2) -H2 #H destruct
+ /3 width=3 by ext2_unit, ex2_intro/
+ ]
]
qed-.
#R #HR #I #L1 #L2 #V #T #H elim (lfxs_dec … HR L1 L2 V)
/4 width=1 by lfxs_flat, or_intror, or_introl/
qed-.
+
+lemma lfnxs_inv_bind_void: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) →
+ ∀p,I,L1,L2,V,T. (L1 ⪤*[R, ⓑ{p,I}V.T] L2 → ⊥) →
+ (L1 ⪤*[R, V] L2 → ⊥) ∨ (L1.ⓧ ⪤*[R, T] L2.ⓧ → ⊥).
+#R #HR #p #I #L1 #L2 #V #T #H elim (lfxs_dec … HR L1 L2 V)
+/4 width=2 by lfxs_bind_void, or_intror, or_introl/
+qed-.
inductive lsuba (G:genv): relation lenv ≝
| lsuba_atom: lsuba G (⋆) (⋆)
-| lsuba_pair: ∀I,L1,L2,V. lsuba G L1 L2 → lsuba G (L1.ⓑ{I}V) (L2.ⓑ{I}V)
+| lsuba_bind: ∀I,L1,L2. lsuba G L1 L2 → lsuba G (L1.ⓘ{I}) (L2.ⓘ{I})
| lsuba_beta: ∀L1,L2,W,V,A. ⦃G, L1⦄ ⊢ ⓝW.V ⁝ A → ⦃G, L2⦄ ⊢ W ⁝ A →
lsuba G L1 L2 → lsuba G (L1.ⓓⓝW.V) (L2.ⓛW)
.
fact lsuba_inv_atom1_aux: ∀G,L1,L2. G ⊢ L1 ⫃⁝ L2 → L1 = ⋆ → L2 = ⋆.
#G #L1 #L2 * -L1 -L2
[ //
-| #I #L1 #L2 #V #_ #H destruct
+| #I #L1 #L2 #_ #H destruct
| #L1 #L2 #W #V #A #_ #_ #_ #H destruct
]
qed-.
lemma lsuba_inv_atom1: ∀G,L2. G ⊢ ⋆ ⫃⁝ L2 → L2 = ⋆.
/2 width=4 by lsuba_inv_atom1_aux/ qed-.
-fact lsuba_inv_pair1_aux: ∀G,L1,L2. G ⊢ L1 ⫃⁝ L2 → ∀I,K1,X. L1 = K1.ⓑ{I}X →
- (â\88\83â\88\83K2. G â\8a¢ K1 â«\83â\81\9d K2 & L2 = K2.â\93\91{I}X) ∨
+fact lsuba_inv_bind1_aux: ∀G,L1,L2. G ⊢ L1 ⫃⁝ L2 → ∀I,K1. L1 = K1.ⓘ{I} →
+ (â\88\83â\88\83K2. G â\8a¢ K1 â«\83â\81\9d K2 & L2 = K2.â\93\98{I}) ∨
∃∃K2,W,V,A. ⦃G, K1⦄ ⊢ ⓝW.V ⁝ A & ⦃G, K2⦄ ⊢ W ⁝ A &
- G ⊢ K1 ⫃⁝ K2 & I = Abbr & L2 = K2.ⓛW & X = ⓝW.V.
+ G ⊢ K1 ⫃⁝ K2 & I = BPair Abbr (ⓝW.V) & L2 = K2.ⓛW.
#G #L1 #L2 * -L1 -L2
-[ #J #K1 #X #H destruct
-| #I #L1 #L2 #V #HL12 #J #K1 #X #H destruct /3 width=3 by ex2_intro, or_introl/
-| #L1 #L2 #W #V #A #HV #HW #HL12 #J #K1 #X #H destruct /3 width=9 by or_intror, ex6_4_intro/
+[ #J #K1 #H destruct
+| #I #L1 #L2 #HL12 #J #K1 #H destruct /3 width=3 by ex2_intro, or_introl/
+| #L1 #L2 #W #V #A #HV #HW #HL12 #J #K1 #H destruct /3 width=9 by ex5_4_intro, or_intror/
]
qed-.
-lemma lsuba_inv_pair1: ∀I,G,K1,L2,X. G ⊢ K1.ⓑ{I}X ⫃⁝ L2 →
- (â\88\83â\88\83K2. G â\8a¢ K1 â«\83â\81\9d K2 & L2 = K2.â\93\91{I}X) ∨
+lemma lsuba_inv_bind1: ∀I,G,K1,L2. G ⊢ K1.ⓘ{I} ⫃⁝ L2 →
+ (â\88\83â\88\83K2. G â\8a¢ K1 â«\83â\81\9d K2 & L2 = K2.â\93\98{I}) ∨
∃∃K2,W,V,A. ⦃G, K1⦄ ⊢ ⓝW.V ⁝ A & ⦃G, K2⦄ ⊢ W ⁝ A & G ⊢ K1 ⫃⁝ K2 &
- I = Abbr & L2 = K2.ⓛW & X = ⓝW.V.
-/2 width=3 by lsuba_inv_pair1_aux/ qed-.
+ I = BPair Abbr (ⓝW.V) & L2 = K2.ⓛW.
+/2 width=3 by lsuba_inv_bind1_aux/ qed-.
fact lsuba_inv_atom2_aux: ∀G,L1,L2. G ⊢ L1 ⫃⁝ L2 → L2 = ⋆ → L1 = ⋆.
#G #L1 #L2 * -L1 -L2
[ //
-| #I #L1 #L2 #V #_ #H destruct
+| #I #L1 #L2 #_ #H destruct
| #L1 #L2 #W #V #A #_ #_ #_ #H destruct
]
qed-.
lemma lsubc_inv_atom2: ∀G,L1. G ⊢ L1 ⫃⁝ ⋆ → L1 = ⋆.
/2 width=4 by lsuba_inv_atom2_aux/ qed-.
-fact lsuba_inv_pair2_aux: ∀G,L1,L2. G ⊢ L1 ⫃⁝ L2 → ∀I,K2,W. L2 = K2.ⓑ{I}W →
- (â\88\83â\88\83K1. G â\8a¢ K1 â«\83â\81\9d K2 & L1 = K1.â\93\91{I}W) ∨
- ∃∃K1,V,A. ⦃G, K1⦄ ⊢ ⓝW.V ⁝ A & ⦃G, K2⦄ ⊢ W ⁝ A &
- G ⊢ K1 ⫃⁝ K2 & I = Abst & L1 = K1.ⓓⓝW.V.
+fact lsuba_inv_bind2_aux: ∀G,L1,L2. G ⊢ L1 ⫃⁝ L2 → ∀I,K2. L2 = K2.ⓘ{I} →
+ (â\88\83â\88\83K1. G â\8a¢ K1 â«\83â\81\9d K2 & L1 = K1.â\93\98{I}) ∨
+ ∃∃K1,V,W, A. ⦃G, K1⦄ ⊢ ⓝW.V ⁝ A & ⦃G, K2⦄ ⊢ W ⁝ A &
+ G ⊢ K1 ⫃⁝ K2 & I = BPair Abst W & L1 = K1.ⓓⓝW.V.
#G #L1 #L2 * -L1 -L2
-[ #J #K2 #U #H destruct
-| #I #L1 #L2 #V #HL12 #J #K2 #U #H destruct /3 width=3 by ex2_intro, or_introl/
-| #L1 #L2 #W #V #A #HV #HW #HL12 #J #K2 #U #H destruct /3 width=7 by or_intror, ex5_3_intro/
+[ #J #K2 #H destruct
+| #I #L1 #L2 #HL12 #J #K2 #H destruct /3 width=3 by ex2_intro, or_introl/
+| #L1 #L2 #W #V #A #HV #HW #HL12 #J #K2 #H destruct /3 width=9 by ex5_4_intro, or_intror/
]
qed-.
-lemma lsuba_inv_pair2: ∀I,G,L1,K2,W. G ⊢ L1 ⫃⁝ K2.ⓑ{I}W →
- (â\88\83â\88\83K1. G â\8a¢ K1 â«\83â\81\9d K2 & L1 = K1.â\93\91{I}W) ∨
- ∃∃K1,V,A. ⦃G, K1⦄ ⊢ ⓝW.V ⁝ A & ⦃G, K2⦄ ⊢ W ⁝ A & G ⊢ K1 ⫃⁝ K2 &
- I = Abst & L1 = K1.ⓓⓝW.V.
-/2 width=3 by lsuba_inv_pair2_aux/ qed-.
+lemma lsuba_inv_bind2: ∀I,G,L1,K2. G ⊢ L1 ⫃⁝ K2.ⓘ{I} →
+ (â\88\83â\88\83K1. G â\8a¢ K1 â«\83â\81\9d K2 & L1 = K1.â\93\98{I}) ∨
+ ∃∃K1,V,W,A. ⦃G, K1⦄ ⊢ ⓝW.V ⁝ A & ⦃G, K2⦄ ⊢ W ⁝ A & G ⊢ K1 ⫃⁝ K2 &
+ I = BPair Abst W & L1 = K1.ⓓⓝW.V.
+/2 width=3 by lsuba_inv_bind2_aux/ qed-.
(* Basic properties *********************************************************)
lemma lsuba_refl: ∀G,L. G ⊢ L ⫃⁝ L.
-#G #L elim L -L /2 width=1 by lsuba_atom, lsuba_pair/
+#G #L elim L -L /2 width=1 by lsuba_atom, lsuba_bind/
qed.
#G #L1 #V #A #H elim H -G -L1 -V -A
[ //
| #I #G #L1 #V #A #HA #IH #L2 #H
- elim (lsuba_inv_pair1 … H) -H * /3 width=1 by aaa_zero/
- #L0 #W0 #V0 #A0 #HV0 #HW0 #HL10 #H1 #H2 #H3 destruct
+ elim (lsuba_inv_bind1 … H) -H * /3 width=1 by aaa_zero/
+ #L0 #W0 #V0 #A0 #HV0 #HW0 #HL10 #H1 #H2 destruct
lapply (aaa_mono … HV0 … HA) #H destruct -V0 -L1 /2 width=1 by aaa_zero/
-| #I #G #K1 #V #A #i #_ #IH #L2 #H
- elim (lsuba_inv_pair1 … H) -H * /3 width=1 by aaa_lref/
-| /4 width=2 by lsuba_pair, aaa_abbr/
-| /4 width=1 by lsuba_pair, aaa_abst/
+| #I #G #K1 #A #i #_ #IH #L2 #H
+ elim (lsuba_inv_bind1 … H) -H * /3 width=1 by aaa_lref/
+| /4 width=2 by lsuba_bind, aaa_abbr/
+| /4 width=1 by lsuba_bind, aaa_abst/
| /3 width=3 by aaa_appl/
| /3 width=1 by aaa_cast/
]
#G #L2 #V #A #H elim H -G -L2 -V -A
[ //
| #I #G #L2 #V #A #HA #IH #L1 #H
- elim (lsuba_inv_pair2 … H) -H * /3 width=1 by aaa_zero/
- #L0 #V0 #A0 #HV0 #HV #HL02 #H1 #H2 destruct
- lapply (aaa_mono … HV … HA) #H destruct -L2 /2 width=1 by aaa_zero/
-| #I #G #K2 #V #A #i #_ #IH #L1 #H
- elim (lsuba_inv_pair2 … H) -H * /3 width=1 by aaa_lref/
-| /4 width=2 by lsuba_pair, aaa_abbr/
-| /4 width=1 by lsuba_pair, aaa_abst/
+ elim (lsuba_inv_bind2 … H) -H * /3 width=1 by aaa_zero/
+ #L0 #V0 #W0 #A0 #HV0 #HW0 #HL02 #H1 #H2 destruct
+ lapply (aaa_mono … HW0 … HA) #H destruct -L2 /2 width=1 by aaa_zero/
+| #I #G #K2 #A #i #_ #IH #L1 #H
+ elim (lsuba_inv_bind2 … H) -H * /3 width=1 by aaa_lref/
+| /4 width=2 by lsuba_bind, aaa_abbr/
+| /4 width=1 by lsuba_bind, aaa_abst/
| /3 width=3 by aaa_appl/
| /3 width=1 by aaa_cast/
]
∃∃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 #b #f #K1 #Hf #H
- elim (drops_inv_pair1_isuni … Hf H) -Hf -H *
+| #I #L1 #L2 #HL12 #IH #b #f #K1 #Hf #H
+ elim (drops_inv_bind1_isuni … Hf H) -Hf -H *
[ #Hf #H destruct -IH
- /3 width=3 by lsuba_pair, drops_refl, ex2_intro/
+ /3 width=3 by lsuba_bind, 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 #b #f #K1 #Hf #H
- elim (drops_inv_pair1_isuni … Hf H) -Hf -H *
+ elim (drops_inv_bind1_isuni … Hf H) -Hf -H *
[ #Hf #H destruct -IH
/3 width=3 by drops_refl, lsuba_beta, ex2_intro/
| #g #Hg #HLK1 #H destruct -HL12
∃∃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 #b #f #K2 #Hf #H
- elim (drops_inv_pair1_isuni … Hf H) -Hf -H *
+| #I #L1 #L2 #HL12 #IH #b #f #K2 #Hf #H
+ elim (drops_inv_bind1_isuni … Hf H) -Hf -H *
[ #Hf #H destruct -IH
- /3 width=3 by lsuba_pair, drops_refl, ex2_intro/
+ /3 width=3 by lsuba_bind, 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 #b #f #K2 #Hf #H
- elim (drops_inv_pair1_isuni … Hf H) -Hf -H *
+ elim (drops_inv_bind1_isuni … Hf H) -Hf -H *
[ #Hf #H destruct -IH
/3 width=3 by drops_refl, lsuba_beta, ex2_intro/
| #g #Hg #HLK2 #H destruct -HL12
theorem lsuba_trans: ∀G. Transitive … (lsuba G).
#G #L1 #L #H elim H -L1 -L
[ #X #H >(lsuba_inv_atom1 … H) -H //
-| #I #L1 #L #Y #HL1 #IHL1 #X #H
- elim (lsuba_inv_pair1 … H) -H * #L2
- [ #HL2 #H destruct /3 width=1 by lsuba_pair/
- | #W #V #A #HV #HW #HL2 #H1 #H2 #H3 destruct
+| #I #L1 #L #HL1 #IHL1 #Y #H
+ elim (lsuba_inv_bind1 … H) -H * #L2
+ [ #HL2 #H destruct /3 width=1 by lsuba_bind/
+ | #W #V #A #HV #HW #HL2 #H1 #H2 destruct
/3 width=3 by lsuba_beta, lsuba_aaa_trans/
]
-| #L1 #L #W #V #A #HV #HW #HL1 #IHL1 #X #H
- elim (lsuba_inv_pair1 … H) -H * #L2
+| #L1 #L #W #V #A #HV #HW #HL1 #IHL1 #Y #H
+ elim (lsuba_inv_bind1 … H) -H * #L2
[ #HL2 #H destruct /3 width=5 by lsuba_beta, lsuba_aaa_conf/
| #W0 #V0 #A0 #_ #_ #_ #H destruct
]
(* Forward lemmas with restricted refinement for local environments *********)
lemma lsuba_fwd_lsubr: ∀G,L1,L2. G ⊢ L1 ⫃⁝ L2 → L1 ⫃ L2.
-#G #L1 #L2 #H elim H -L1 -L2 /2 width=1 by lsubr_pair, lsubr_beta/
+#G #L1 #L2 #H elim H -L1 -L2 /2 width=1 by lsubr_bind, lsubr_beta/
qed-.
inductive lsubc (RP) (G): relation lenv ≝
| lsubc_atom: lsubc RP G (⋆) (⋆)
-| lsubc_pair: ∀I,L1,L2,V. lsubc RP G L1 L2 → lsubc RP G (L1.ⓑ{I}V) (L2.ⓑ{I}V)
+| lsubc_bind: ∀I,L1,L2. lsubc RP G L1 L2 → lsubc RP G (L1.ⓘ{I}) (L2.ⓘ{I})
| lsubc_beta: ∀L1,L2,V,W,A. ⦃G, L1, V⦄ ϵ[RP] 〚A〛 → ⦃G, L1, W⦄ ϵ[RP] 〚A〛 → ⦃G, L2⦄ ⊢ W ⁝ A →
lsubc RP G L1 L2 → lsubc RP G (L1. ⓓⓝW.V) (L2.ⓛW)
.
fact lsubc_inv_atom1_aux: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → L1 = ⋆ → L2 = ⋆.
#RP #G #L1 #L2 * -L1 -L2
[ //
-| #I #L1 #L2 #V #_ #H destruct
+| #I #L1 #L2 #_ #H destruct
| #L1 #L2 #V #W #A #_ #_ #_ #_ #H destruct
]
qed-.
lemma lsubc_inv_atom1: ∀RP,G,L2. G ⊢ ⋆ ⫃[RP] L2 → L2 = ⋆.
/2 width=5 by lsubc_inv_atom1_aux/ qed-.
-fact lsubc_inv_pair1_aux: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → ∀I,K1,X. L1 = K1.ⓑ{I}X →
- (â\88\83â\88\83K2. G â\8a¢ K1 â«\83[RP] K2 & L2 = K2.â\93\91{I}X) ∨
+fact lsubc_inv_bind1_aux: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → ∀I,K1. L1 = K1.ⓘ{I} →
+ (â\88\83â\88\83K2. G â\8a¢ K1 â«\83[RP] K2 & L2 = K2.â\93\98{I}) ∨
∃∃K2,V,W,A. ⦃G, K1, V⦄ ϵ[RP] 〚A〛 & ⦃G, K1, W⦄ ϵ[RP] 〚A〛 & ⦃G, K2⦄ ⊢ W ⁝ A &
G ⊢ K1 ⫃[RP] K2 &
- L2 = K2. ⓛW & X = ⓝW.V & I = Abbr.
+ L2 = K2. ⓛW & I = BPair Abbr (ⓝW.V).
#RP #G #L1 #L2 * -L1 -L2
-[ #I #K1 #V #H destruct
-| #J #L1 #L2 #V #HL12 #I #K1 #W #H destruct /3 width=3 by ex2_intro, or_introl/
-| #L1 #L2 #V1 #W2 #A #HV1 #H1W2 #H2W2 #HL12 #I #K1 #V #H destruct /3 width=10 by ex7_4_intro, or_intror/
+[ #I #K1 #H destruct
+| #J #L1 #L2 #HL12 #I #K1 #H destruct /3 width=3 by ex2_intro, or_introl/
+| #L1 #L2 #V1 #W2 #A #HV1 #H1W2 #H2W2 #HL12 #I #K1 #H destruct
+ /3 width=10 by ex6_4_intro, or_intror/
]
qed-.
(* Basic_1: was: csubc_gen_head_r *)
-lemma lsubc_inv_pair1: ∀RP,I,G,K1,L2,X. G ⊢ K1.ⓑ{I}X ⫃[RP] L2 →
- (â\88\83â\88\83K2. G â\8a¢ K1 â«\83[RP] K2 & L2 = K2.â\93\91{I}X) ∨
+lemma lsubc_inv_bind1: ∀RP,I,G,K1,L2. G ⊢ K1.ⓘ{I} ⫃[RP] L2 →
+ (â\88\83â\88\83K2. G â\8a¢ K1 â«\83[RP] K2 & L2 = K2.â\93\98{I}) ∨
∃∃K2,V,W,A. ⦃G, K1, V⦄ ϵ[RP] 〚A〛 & ⦃G, K1, W⦄ ϵ[RP] 〚A〛 & ⦃G, K2⦄ ⊢ W ⁝ A &
G ⊢ K1 ⫃[RP] K2 &
- L2 = K2.ⓛW & X = ⓝW.V & I = Abbr.
-/2 width=3 by lsubc_inv_pair1_aux/ qed-.
+ L2 = K2.ⓛW & I = BPair Abbr (ⓝW.V).
+/2 width=3 by lsubc_inv_bind1_aux/ qed-.
fact lsubc_inv_atom2_aux: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → L2 = ⋆ → L1 = ⋆.
#RP #G #L1 #L2 * -L1 -L2
[ //
-| #I #L1 #L2 #V #_ #H destruct
+| #I #L1 #L2 #_ #H destruct
| #L1 #L2 #V #W #A #_ #_ #_ #_ #H destruct
]
qed-.
lemma lsubc_inv_atom2: ∀RP,G,L1. G ⊢ L1 ⫃[RP] ⋆ → L1 = ⋆.
/2 width=5 by lsubc_inv_atom2_aux/ qed-.
-fact lsubc_inv_pair2_aux: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → ∀I,K2,W. L2 = K2.ⓑ{I} W →
- (â\88\83â\88\83K1. G â\8a¢ K1 â«\83[RP] K2 & L1 = K1. â\93\91{I} W) ∨
- ∃∃K1,V,A. ⦃G, K1, V⦄ ϵ[RP] 〚A〛 & ⦃G, K1, W⦄ ϵ[RP] 〚A〛 & ⦃G, K2⦄ ⊢ W ⁝ A &
- G ⊢ K1 ⫃[RP] K2 &
- L1 = K1.ⓓⓝW.V & I = Abst.
+fact lsubc_inv_bind2_aux: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → ∀I,K2. L2 = K2.ⓘ{I} →
+ (â\88\83â\88\83K1. G â\8a¢ K1 â«\83[RP] K2 & L1 = K1. â\93\98{I}) ∨
+ ∃∃K1,V,W,A. ⦃G, K1, V⦄ ϵ[RP] 〚A〛 & ⦃G, K1, W⦄ ϵ[RP] 〚A〛 & ⦃G, K2⦄ ⊢ W ⁝ A &
+ G ⊢ K1 ⫃[RP] K2 &
+ L1 = K1.ⓓⓝW.V & I = BPair Abst W.
#RP #G #L1 #L2 * -L1 -L2
-[ #I #K2 #W #H destruct
-| #J #L1 #L2 #V #HL12 #I #K2 #W #H destruct /3 width=3 by ex2_intro, or_introl/
-| #L1 #L2 #V1 #W2 #A #HV1 #H1W2 #H2W2 #HL12 #I #K2 #W #H destruct /3 width=8 by ex6_3_intro, or_intror/
+[ #I #K2 #H destruct
+| #J #L1 #L2 #HL12 #I #K2 #H destruct /3 width=3 by ex2_intro, or_introl/
+| #L1 #L2 #V1 #W2 #A #HV1 #H1W2 #H2W2 #HL12 #I #K2 #H destruct
+ /3 width=10 by ex6_4_intro, or_intror/
]
qed-.
(* Basic_1: was just: csubc_gen_head_l *)
-lemma lsubc_inv_pair2: ∀RP,I,G,L1,K2,W. G ⊢ L1 ⫃[RP] K2.ⓑ{I} W →
- (â\88\83â\88\83K1. G â\8a¢ K1 â«\83[RP] K2 & L1 = K1.â\93\91{I} W) ∨
- ∃∃K1,V,A. ⦃G, K1, V⦄ ϵ[RP] 〚A〛 & ⦃G, K1, W⦄ ϵ[RP] 〚A〛 & ⦃G, K2⦄ ⊢ W ⁝ A &
- G ⊢ K1 ⫃[RP] K2 &
- L1 = K1.ⓓⓝW.V & I = Abst.
-/2 width=3 by lsubc_inv_pair2_aux/ qed-.
+lemma lsubc_inv_bind2: ∀RP,I,G,L1,K2. G ⊢ L1 ⫃[RP] K2.ⓘ{I} →
+ (â\88\83â\88\83K1. G â\8a¢ K1 â«\83[RP] K2 & L1 = K1.â\93\98{I}) ∨
+ ∃∃K1,V,W,A. ⦃G, K1, V⦄ ϵ[RP] 〚A〛 & ⦃G, K1, W⦄ ϵ[RP] 〚A〛 & ⦃G, K2⦄ ⊢ W ⁝ A &
+ G ⊢ K1 ⫃[RP] K2 &
+ L1 = K1.ⓓⓝW.V & I = BPair Abst W.
+/2 width=3 by lsubc_inv_bind2_aux/ qed-.
(* Basic properties *********************************************************)
(* Basic_1: was just: csubc_refl *)
lemma lsubc_refl: ∀RP,G,L. G ⊢ L ⫃[RP] L.
-#RP #G #L elim L -L /2 width=1 by lsubc_pair/
+#RP #G #L elim L -L /2 width=1 by lsubc_bind/
qed.
(* Basic_1: removed theorems 3:
∃∃K1. ⬇*[b, f] L1 ≡ K1 & G ⊢ K1 ⫃[RP] K2.
#RP #G #L1 #L2 #H elim H -L1 -L2
[ /2 width=3 by ex2_intro/
-| #I #L1 #L2 #V #HL12 #IH #b #f #K2 #Hf #H
- elim (drops_inv_pair1_isuni … Hf H) -Hf -H *
+| #I #L1 #L2 #HL12 #IH #b #f #K2 #Hf #H
+ elim (drops_inv_bind1_isuni … Hf H) -Hf -H *
[ #Hf #H destruct -IH
- /3 width=3 by lsubc_pair, drops_refl, ex2_intro/
+ /3 width=3 by lsubc_bind, 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 #V #W #A #HV #H1W #H2W #HL12 #IH #b #f #K2 #Hf #H
- elim (drops_inv_pair1_isuni … Hf H) -Hf -H *
+ elim (drops_inv_bind1_isuni … Hf H) -Hf -H *
[ #Hf #H destruct -IH
/3 width=8 by drops_refl, lsubc_beta, ex2_intro/
| #g #Hg #HLK2 #H destruct -HL12
#RR #RS #RP #HR #b #f #G #L1 #K1 #H elim H -f -L1 -K1
[ #f #Hf #Y #H lapply (lsubc_inv_atom1 … H) -H
#H destruct /4 width=3 by lsubc_atom, drops_atom, ex2_intro/
-| #f #I #L1 #K1 #V1 #_ #IH #K2 #HK12 elim (IH … HK12) -K1
- /3 width=5 by lsubc_pair, drops_drop, ex2_intro/
-| #f #I #L1 #K1 #V1 #V2 #HLK1 #HV21 #IH #X #H elim (lsubc_inv_pair1 … H) -H *
+| #f #I #L1 #K1 #_ #IH #K2 #HK12 elim (IH … HK12) -K1
+ /3 width=5 by lsubc_bind, drops_drop, ex2_intro/
+| #f #Z #I #L1 #K1 #HLK1 #HZ #IH #Y #H elim (lsubc_inv_bind1 … H) -H *
[ #K2 #HK12 #H destruct -HLK1
- elim (IH … HK12) -K1 /3 width=5 by lsubc_pair, drops_skip, ex2_intro/
- | #K2 #V #W2 #A #HV #H1W2 #H2W2 #HK12 #H1 #H2 #H3 destruct
+ elim (IH … HK12) -K1 /3 width=5 by lsubc_bind, drops_skip, ex2_intro/
+ | #K2 #V2 #W2 #A #HV2 #H1W2 #H2W2 #HK12 #H1 #H2 destruct
+ elim (liftsb_inv_pair_sn … HZ) -HZ #V1 #HV21 #H destruct
elim (lifts_inv_flat1 … HV21) -HV21 #W3 #V3 #HW23 #HV3 #H destruct
elim (IH … HK12) -IH -HK12 #K #HL1K #HK2
- lapply (acr_lifts … HR … HV … HLK1 … HV3) -HV
+ lapply (acr_lifts … HR … HV2 … HLK1 … HV3) -HV2
lapply (acr_lifts … HR … H1W2 … HLK1 … HW23) -H1W2
- /4 width=10 by lsubc_beta, aaa_lifts, drops_skip, ex2_intro/
+ /4 width=10 by lsubc_beta, aaa_lifts, drops_skip, ext2_pair, ex2_intro/
]
]
qed-.
lemma lsuba_lsubc: ∀RR,RS,RP. gcp RR RS RP → gcr RR RS RP RP →
∀G,L1,L2. G ⊢ L1 ⫃⁝ L2 → G ⊢ L1 ⫃[RP] L2.
-#RR #RS #RP #H1RP #H2RP #G #L1 #L2 #H elim H -L1 -L2 /2 width=1 by lsubc_pair/
+#RR #RS #RP #H1RP #H2RP #G #L1 #L2 #H elim H -L1 -L2 /2 width=1 by lsubc_bind/
#L1 #L2 #W #V #A #H elim (aaa_inv_cast … H) -H /3 width=4 by acr_aaa, lsubc_beta/
qed.
(* Forward lemmas with restricted refinement for local environments *********)
lemma lsubc_fwd_lsubr: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 → L1 ⫃ L2.
-#RP #G #L1 #L2 #H elim H -L1 -L2 /2 width=1 by lsubr_pair, lsubr_beta/
+#RP #G #L1 #L2 #H elim H -L1 -L2 /2 width=1 by lsubr_bind, lsubr_beta/
qed-.
(* RESTRICTED REFINEMENT FOR CONTEXT-SENSITIVE FREE VARIABLES ***************)
inductive lsubf: relation4 lenv rtmap lenv rtmap ≝
-| lsubf_atom: ∀f1,f2. f2 ⊆ f1 → lsubf (⋆) f1 (⋆) f2
-| lsubf_pair: ∀f1,f2,I,L1,L2,V. f2 ⊆ f1 → lsubf L1 (⫱f1) L2 (⫱f2) →
- lsubf (L1.ⓑ{I}V) f1 (L2.ⓑ{I}V) f2
-| lsubf_beta: ∀f,f1,f2,L1,L2,W,V. L1 ⊢ 𝐅*⦃V⦄ ≡ f → f ⊆ ⫱f1 → f2 ⊆ f1 →
- lsubf L1 (⫱f1) L2 (⫱f2) → lsubf (L1.ⓓⓝW.V) f1 (L2.ⓛW) f2
+| lsubf_atom: ∀f1,f2. f1 ≗ f2 → lsubf (⋆) f1 (⋆) f2
+| lsubf_push: ∀f1,f2,I1,I2,L1,L2. lsubf L1 (f1) L2 (f2) →
+ lsubf (L1.ⓘ{I1}) (↑f1) (L2.ⓘ{I2}) (↑f2)
+| lsubf_bind: ∀f1,f2,I,L1,L2. lsubf L1 f1 L2 f2 →
+ lsubf (L1.ⓘ{I}) (⫯f1) (L2.ⓘ{I}) (⫯f2)
+| lsubf_beta: ∀f,f0,f1,f2,L1,L2,W,V. L1 ⊢ 𝐅*⦃V⦄ ≡ f → f0 ⋓ f ≡ f1 →
+ lsubf L1 f0 L2 f2 → lsubf (L1.ⓓⓝW.V) (⫯f1) (L2.ⓛW) (⫯f2)
+| lsubf_unit: ∀f,f0,f1,f2,I1,I2,L1,L2,V. L1 ⊢ 𝐅*⦃V⦄ ≡ f → f0 ⋓ f ≡ f1 →
+ lsubf L1 f0 L2 f2 → lsubf (L1.ⓑ{I1}V) (⫯f1) (L2.ⓤ{I2}) (⫯f2)
.
interpretation
(* Basic inversion lemmas ***************************************************)
fact lsubf_inv_atom1_aux: ∀f1,f2,L1,L2. ⦃L1, f1⦄ ⫃𝐅* ⦃L2, f2⦄ → L1 = ⋆ →
- L2 = ⋆ ∧ f2 ⊆ f1.
+ f1 ≗ f2 ∧ L2 = ⋆.
#f1 #f2 #L1 #L2 * -f1 -f2 -L1 -L2
[ /2 width=1 by conj/
-| #f1 #f2 #I #L1 #L2 #V #_ #_ #H destruct
-| #f #f1 #f2 #L1 #L2 #W #V #_ #_ #_ #_ #H destruct
+| #f1 #f2 #I1 #I2 #L1 #L2 #_ #H destruct
+| #f1 #f2 #I #L1 #L2 #_ #H destruct
+| #f #f0 #f1 #f2 #L1 #L2 #W #V #_ #_ #_ #H destruct
+| #f #f0 #f1 #f2 #I1 #I2 #L1 #L2 #V #_ #_ #_ #H destruct
]
qed-.
-lemma lsubf_inv_atom1: ∀f1,f2,L2. ⦃⋆, f1⦄ ⫃𝐅* ⦃L2, f2⦄ → L2 = ⋆ ∧ f2 ⊆ f1.
+lemma lsubf_inv_atom1: ∀f1,f2,L2. ⦃⋆, f1⦄ ⫃𝐅* ⦃L2, f2⦄ → f1 ≗ f2 ∧ L2 = ⋆.
/2 width=3 by lsubf_inv_atom1_aux/ qed-.
+fact lsubf_inv_push1_aux: ∀f1,f2,L1,L2. ⦃L1, f1⦄ ⫃𝐅* ⦃L2, f2⦄ →
+ ∀g1,I1,K1. f1 = ↑g1 → L1 = K1.ⓘ{I1} →
+ ∃∃g2,I2,K2. ⦃K1, g1⦄ ⫃𝐅* ⦃K2, g2⦄ & f2 = ↑g2 & L2 = K2.ⓘ{I2}.
+#f1 #f2 #L1 #L2 * -f1 -f2 -L1 -L2
+[ #f1 #f2 #_ #g1 #J1 #K1 #_ #H destruct
+| #f1 #f2 #I1 #I2 #L1 #L2 #H12 #g1 #J1 #K1 #H1 #H2 destruct
+ <(injective_push … H1) -g1 /2 width=6 by ex3_3_intro/
+| #f1 #f2 #I #L1 #L2 #_ #g1 #J1 #K1 #H elim (discr_next_push … H)
+| #f #f0 #f1 #f2 #L1 #L2 #W #V #_ #_ #_ #g1 #J1 #K1 #H elim (discr_next_push … H)
+| #f #f0 #f1 #f2 #I1 #I2 #L1 #L2 #V #_ #_ #_ #g1 #J1 #K1 #H elim (discr_next_push … H)
+]
+qed-.
+
+lemma lsubf_inv_push1: ∀g1,f2,I1,K1,L2. ⦃K1.ⓘ{I1}, ↑g1⦄ ⫃𝐅* ⦃L2, f2⦄ →
+ ∃∃g2,I2,K2. ⦃K1, g1⦄ ⫃𝐅* ⦃K2, g2⦄ & f2 = ↑g2 & L2 = K2.ⓘ{I2}.
+/2 width=6 by lsubf_inv_push1_aux/ qed-.
+
fact lsubf_inv_pair1_aux: ∀f1,f2,L1,L2. ⦃L1, f1⦄ ⫃𝐅* ⦃L2, f2⦄ →
- ∀I,K1,X. L1 = K1.ⓑ{I}X →
- (∃∃K2. f2 ⊆ f1 & ⦃K1, ⫱f1⦄ ⫃𝐅* ⦃K2, ⫱f2⦄ & L2 = K2.ⓑ{I}X) ∨
- ∃∃f,K2,W,V. K1 ⊢ 𝐅*⦃V⦄ ≡ f & f ⊆ ⫱f1 &
- f2 ⊆ f1 & ⦃K1, ⫱f1⦄ ⫃𝐅* ⦃K2, ⫱f2⦄ & I = Abbr & L2 = K2.ⓛW & X = ⓝW.V.
+ ∀g1,I,K1,X. f1 = ⫯g1 → L1 = K1.ⓑ{I}X →
+ ∨∨ ∃∃g2,K2. ⦃K1, g1⦄ ⫃𝐅* ⦃K2, g2⦄ & f2 = ⫯g2 & L2 = K2.ⓑ{I}X
+ | ∃∃g,g0,g2,K2,W,V. ⦃K1, g0⦄ ⫃𝐅* ⦃K2, g2⦄ &
+ K1 ⊢ 𝐅*⦃V⦄ ≡ g & g0 ⋓ g ≡ g1 & f2 = ⫯g2 &
+ I = Abbr & X = ⓝW.V & L2 = K2.ⓛW
+ | ∃∃g,g0,g2,J,K2. ⦃K1, g0⦄ ⫃𝐅* ⦃K2, g2⦄ &
+ K1 ⊢ 𝐅*⦃X⦄ ≡ g & g0 ⋓ g ≡ g1 & f2 = ⫯g2 &
+ L2 = K2.ⓤ{J}.
#f1 #f2 #L1 #L2 * -f1 -f2 -L1 -L2
-[ #f1 #f2 #_ #J #K1 #X #H destruct
-| #f1 #f2 #I #L1 #L2 #V #H21 #HL12 #J #K1 #X #H destruct
- /3 width=3 by ex3_intro, or_introl/
-| #f #f1 #f2 #L1 #L2 #W #V #Hf #Hf1 #H21 #HL12 #J #K1 #X #H destruct
- /3 width=11 by ex7_4_intro, or_intror/
+[ #f1 #f2 #_ #g1 #J #K1 #X #_ #H destruct
+| #f1 #f2 #I1 #I2 #L1 #L2 #H12 #g1 #J #K1 #X #H elim (discr_push_next … H)
+| #f1 #f2 #I #L1 #L2 #H12 #g1 #J #K1 #X #H1 #H2 destruct
+ <(injective_next … H1) -g1 /3 width=5 by or3_intro0, ex3_2_intro/
+| #f #f0 #f1 #f2 #L1 #L2 #W #V #Hf #Hf1 #H12 #g1 #J #K1 #X #H1 #H2 destruct
+ <(injective_next … H1) -g1 /3 width=12 by or3_intro1, ex7_6_intro/
+| #f #f0 #f1 #f2 #I1 #I2 #L1 #L2 #V #Hf #Hf1 #H12 #g1 #J #K1 #X #H1 #H2 destruct
+ <(injective_next … H1) -g1 /3 width=10 by or3_intro2, ex5_5_intro/
]
qed-.
-lemma lsubf_inv_pair1: ∀f1,f2,I,K1,L2,X. ⦃K1.ⓑ{I}X, f1⦄ ⫃𝐅* ⦃L2, f2⦄ →
- (∃∃K2. f2 ⊆ f1 & ⦃K1, ⫱f1⦄ ⫃𝐅* ⦃K2, ⫱f2⦄ & L2 = K2.ⓑ{I}X) ∨
- ∃∃f,K2,W,V. K1 ⊢ 𝐅*⦃V⦄ ≡ f & f ⊆ ⫱f1 &
- f2 ⊆ f1 & ⦃K1, ⫱f1⦄ ⫃𝐅* ⦃K2, ⫱f2⦄ & I = Abbr & L2 = K2.ⓛW & X = ⓝW.V.
-/2 width=3 by lsubf_inv_pair1_aux/ qed-.
+lemma lsubf_inv_pair1: ∀g1,f2,I,K1,L2,X. ⦃K1.ⓑ{I}X, ⫯g1⦄ ⫃𝐅* ⦃L2, f2⦄ →
+ ∨∨ ∃∃g2,K2. ⦃K1, g1⦄ ⫃𝐅* ⦃K2, g2⦄ & f2 = ⫯g2 & L2 = K2.ⓑ{I}X
+ | ∃∃g,g0,g2,K2,W,V. ⦃K1, g0⦄ ⫃𝐅* ⦃K2, g2⦄ &
+ K1 ⊢ 𝐅*⦃V⦄ ≡ g & g0 ⋓ g ≡ g1 & f2 = ⫯g2 &
+ I = Abbr & X = ⓝW.V & L2 = K2.ⓛW
+ | ∃∃g,g0,g2,J,K2. ⦃K1, g0⦄ ⫃𝐅* ⦃K2, g2⦄ &
+ K1 ⊢ 𝐅*⦃X⦄ ≡ g & g0 ⋓ g ≡ g1 & f2 = ⫯g2 &
+ L2 = K2.ⓤ{J}.
+/2 width=5 by lsubf_inv_pair1_aux/ qed-.
+
+fact lsubf_inv_unit1_aux: ∀f1,f2,L1,L2. ⦃L1, f1⦄ ⫃𝐅* ⦃L2, f2⦄ →
+ ∀g1,I,K1. f1 = ⫯g1 → L1 = K1.ⓤ{I} →
+ ∃∃g2,K2. ⦃K1, g1⦄ ⫃𝐅* ⦃K2, g2⦄ & f2 = ⫯g2 & L2 = K2.ⓤ{I}.
+#f1 #f2 #L1 #L2 * -f1 -f2 -L1 -L2
+[ #f1 #f2 #_ #g1 #J #K1 #_ #H destruct
+| #f1 #f2 #I1 #I2 #L1 #L2 #H12 #g1 #J #K1 #H elim (discr_push_next … H)
+| #f1 #f2 #I #L1 #L2 #H12 #g1 #J #K1 #H1 #H2 destruct
+ <(injective_next … H1) -g1 /2 width=5 by ex3_2_intro/
+| #f #f0 #f1 #f2 #L1 #L2 #W #V #_ #_ #_ #g1 #J #K1 #_ #H destruct
+| #f #f0 #f1 #f2 #I1 #I2 #L1 #L2 #V #_ #_ #_ #g1 #J #K1 #_ #H destruct
+]
+qed-.
+
+lemma lsubf_inv_unit1: ∀g1,f2,I,K1,L2. ⦃K1.ⓤ{I}, ⫯g1⦄ ⫃𝐅* ⦃L2, f2⦄ →
+ ∃∃g2,K2. ⦃K1, g1⦄ ⫃𝐅* ⦃K2, g2⦄ & f2 = ⫯g2 & L2 = K2.ⓤ{I}.
+/2 width=5 by lsubf_inv_unit1_aux/ qed-.
fact lsubf_inv_atom2_aux: ∀f1,f2,L1,L2. ⦃L1, f1⦄ ⫃𝐅* ⦃L2, f2⦄ → L2 = ⋆ →
- L1 = ⋆ ∧ f2 ⊆ f1.
+ f1 ≗ f2 ∧ L1 = ⋆.
#f1 #f2 #L1 #L2 * -f1 -f2 -L1 -L2
[ /2 width=1 by conj/
-| #f1 #f2 #I #L1 #L2 #V #_ #_ #H destruct
-| #f #f1 #f2 #L1 #L2 #W #V #_ #_ #_ #_ #H destruct
+| #f1 #f2 #I1 #I2 #L1 #L2 #_ #H destruct
+| #f1 #f2 #I #L1 #L2 #_ #H destruct
+| #f #f0 #f1 #f2 #L1 #L2 #W #V #_ #_ #_ #H destruct
+| #f #f0 #f1 #f2 #I1 #I2 #L1 #L2 #V #_ #_ #_ #H destruct
]
qed-.
-lemma lsubf_inv_atom2: ∀f1,f2,L1. ⦃L1, f1⦄ ⫃𝐅* ⦃⋆, f2⦄ → L1 = ⋆ ∧ f2 ⊆ f1.
+lemma lsubf_inv_atom2: ∀f1,f2,L1. ⦃L1, f1⦄ ⫃𝐅* ⦃⋆, f2⦄ → f1 ≗ f2 ∧ L1 = ⋆.
/2 width=3 by lsubf_inv_atom2_aux/ qed-.
+fact lsubf_inv_push2_aux: ∀f1,f2,L1,L2. ⦃L1, f1⦄ ⫃𝐅* ⦃L2, f2⦄ →
+ ∀g2,I2,K2. f2 = ↑g2 → L2 = K2.ⓘ{I2} →
+ ∃∃g1,I1,K1. ⦃K1, g1⦄ ⫃𝐅* ⦃K2, g2⦄ & f1 = ↑g1 & L1 = K1.ⓘ{I1}.
+#f1 #f2 #L1 #L2 * -f1 -f2 -L1 -L2
+[ #f1 #f2 #_ #g2 #J2 #K2 #_ #H destruct
+| #f1 #f2 #I1 #I2 #L1 #L2 #H12 #g2 #J2 #K2 #H1 #H2 destruct
+ <(injective_push … H1) -g2 /2 width=6 by ex3_3_intro/
+| #f1 #f2 #I #L1 #L2 #_ #g2 #J2 #K2 #H elim (discr_next_push … H)
+| #f #f0 #f1 #f2 #L1 #L2 #W #V #_ #_ #_ #g2 #J2 #K2 #H elim (discr_next_push … H)
+| #f #f0 #f1 #f2 #I1 #I2 #L1 #L2 #V #_ #_ #_ #g2 #J2 #K2 #H elim (discr_next_push … H)
+]
+qed-.
+
+lemma lsubf_inv_push2: ∀f1,g2,I2,L1,K2. ⦃L1, f1⦄ ⫃𝐅* ⦃K2.ⓘ{I2}, ↑g2⦄ →
+ ∃∃g1,I1,K1. ⦃K1, g1⦄ ⫃𝐅* ⦃K2, g2⦄ & f1 = ↑g1 & L1 = K1.ⓘ{I1}.
+/2 width=6 by lsubf_inv_push2_aux/ qed-.
+
fact lsubf_inv_pair2_aux: ∀f1,f2,L1,L2. ⦃L1, f1⦄ ⫃𝐅* ⦃L2, f2⦄ →
- ∀I,K2,W. L2 = K2.ⓑ{I}W →
- (∃∃K1.f2 ⊆ f1 & ⦃K1, ⫱f1⦄ ⫃𝐅* ⦃K2, ⫱f2⦄ & L1 = K1.ⓑ{I}W) ∨
- ∃∃f,K1,V. K1 ⊢ 𝐅*⦃V⦄ ≡ f & f ⊆ ⫱f1 &
- f2 ⊆ f1 & ⦃K1, ⫱f1⦄ ⫃𝐅* ⦃K2, ⫱f2⦄ & I = Abst & L1 = K1.ⓓⓝW.V.
+ ∀g2,I,K2,W. f2 = ⫯g2 → L2 = K2.ⓑ{I}W →
+ ∨∨ ∃∃g1,K1. ⦃K1, g1⦄ ⫃𝐅* ⦃K2, g2⦄ & f1 = ⫯g1 & L1 = K1.ⓑ{I}W
+ | ∃∃g,g0,g1,K1,V. ⦃K1, g0⦄ ⫃𝐅* ⦃K2, g2⦄ &
+ K1 ⊢ 𝐅*⦃V⦄ ≡ g & g0 ⋓ g ≡ g1 & f1 = ⫯g1 &
+ I = Abst & L1 = K1.ⓓⓝW.V.
#f1 #f2 #L1 #L2 * -f1 -f2 -L1 -L2
-[ #f1 #f2 #_ #J #K2 #X #H destruct
-| #f1 #f2 #I #L1 #L2 #V #H21 #HL12 #J #K2 #X #H destruct
- /3 width=3 by ex3_intro, or_introl/
-| #f #f1 #f2 #L1 #L2 #W #V #Hf #Hf1 #H21 #HL12 #J #K2 #X #H destruct
- /3 width=7 by ex6_3_intro, or_intror/
+[ #f1 #f2 #_ #g2 #J #K2 #X #_ #H destruct
+| #f1 #f2 #I1 #I2 #L1 #L2 #H12 #g2 #J #K2 #X #H elim (discr_push_next … H)
+| #f1 #f2 #I #L1 #L2 #H12 #g2 #J #K2 #X #H1 #H2 destruct
+ <(injective_next … H1) -g2 /3 width=5 by ex3_2_intro, or_introl/
+| #f #f0 #f1 #f2 #L1 #L2 #W #V #Hf #Hf1 #H12 #g2 #J #K2 #X #H1 #H2 destruct
+ <(injective_next … H1) -g2 /3 width=10 by ex6_5_intro, or_intror/
+| #f #f0 #f1 #f2 #I1 #I2 #L1 #L2 #V #_ #_ #_ #g2 #J #K2 #X #_ #H destruct
]
qed-.
-lemma lsubf_inv_pair2: ∀f1,f2,I,L1,K2,W. ⦃L1, f1⦄ ⫃𝐅* ⦃K2.ⓑ{I}W, f2⦄ →
- (∃∃K1.f2 ⊆ f1 & ⦃K1, ⫱f1⦄ ⫃𝐅* ⦃K2, ⫱f2⦄ & L1 = K1.ⓑ{I}W) ∨
- ∃∃f,K1,V. K1 ⊢ 𝐅*⦃V⦄ ≡ f & f ⊆ ⫱f1 &
- f2 ⊆ f1 & ⦃K1, ⫱f1⦄ ⫃𝐅* ⦃K2, ⫱f2⦄ & I = Abst & L1 = K1.ⓓⓝW.V.
+lemma lsubf_inv_pair2: ∀f1,g2,I,L1,K2,W. ⦃L1, f1⦄ ⫃𝐅* ⦃K2.ⓑ{I}W, ⫯g2⦄ →
+ ∨∨ ∃∃g1,K1. ⦃K1, g1⦄ ⫃𝐅* ⦃K2, g2⦄ & f1 = ⫯g1 & L1 = K1.ⓑ{I}W
+ | ∃∃g,g0,g1,K1,V. ⦃K1, g0⦄ ⫃𝐅* ⦃K2, g2⦄ &
+ K1 ⊢ 𝐅*⦃V⦄ ≡ g & g0 ⋓ g ≡ g1 & f1 = ⫯g1 &
+ I = Abst & L1 = K1.ⓓⓝW.V.
/2 width=5 by lsubf_inv_pair2_aux/ qed-.
+fact lsubf_inv_unit2_aux: ∀f1,f2,L1,L2. ⦃L1, f1⦄ ⫃𝐅* ⦃L2, f2⦄ →
+ ∀g2,I,K2. f2 = ⫯g2 → L2 = K2.ⓤ{I} →
+ ∨∨ ∃∃g1,K1. ⦃K1, g1⦄ ⫃𝐅* ⦃K2, g2⦄ & f1 = ⫯g1 & L1 = K1.ⓤ{I}
+ | ∃∃g,g0,g1,J,K1,V. ⦃K1, g0⦄ ⫃𝐅* ⦃K2, g2⦄ &
+ K1 ⊢ 𝐅*⦃V⦄ ≡ g & g0 ⋓ g ≡ g1 & f1 = ⫯g1 &
+ L1 = K1.ⓑ{J}V.
+#f1 #f2 #L1 #L2 * -f1 -f2 -L1 -L2
+[ #f1 #f2 #_ #g2 #J #K2 #_ #H destruct
+| #f1 #f2 #I1 #I2 #L1 #L2 #H12 #g2 #J #K2 #H elim (discr_push_next … H)
+| #f1 #f2 #I #L1 #L2 #H12 #g2 #J #K2 #H1 #H2 destruct
+ <(injective_next … H1) -g2 /3 width=5 by ex3_2_intro, or_introl/
+| #f #f0 #f1 #f2 #L1 #L2 #W #V #_ #_ #_ #g2 #J #K2 #_ #H destruct
+| #f #f0 #f1 #f2 #I1 #I2 #L1 #L2 #V #Hf #Hf1 #H12 #g2 #J #K2 #H1 #H2 destruct
+ <(injective_next … H1) -g2 /3 width=11 by ex5_6_intro, or_intror/
+]
+qed-.
+
+lemma lsubf_inv_unit2: ∀f1,g2,I,L1,K2. ⦃L1, f1⦄ ⫃𝐅* ⦃K2.ⓤ{I}, ⫯g2⦄ →
+ ∨∨ ∃∃g1,K1. ⦃K1, g1⦄ ⫃𝐅* ⦃K2, g2⦄ & f1 = ⫯g1 & L1 = K1.ⓤ{I}
+ | ∃∃g,g0,g1,J,K1,V. ⦃K1, g0⦄ ⫃𝐅* ⦃K2, g2⦄ &
+ K1 ⊢ 𝐅*⦃V⦄ ≡ g & g0 ⋓ g ≡ g1 & f1 = ⫯g1 &
+ L1 = K1.ⓑ{J}V.
+/2 width=5 by lsubf_inv_unit2_aux/ qed-.
+
+(* Advanced inversion lemmas ************************************************)
+
+lemma lsubf_inv_atom: ∀f1,f2. ⦃⋆, f1⦄ ⫃𝐅* ⦃⋆, f2⦄ → f1 ≗ f2.
+#f1 #f2 #H elim (lsubf_inv_atom1 … H) -H //
+qed-.
+
+lemma lsubf_inv_push_sn: ∀g1,f2,I1,I2,K1,K2. ⦃K1.ⓘ{I1}, ↑g1⦄ ⫃𝐅* ⦃K2.ⓘ{I2}, f2⦄ →
+ ∃∃g2. ⦃K1, g1⦄ ⫃𝐅* ⦃K2, g2⦄ & f2 = ↑g2.
+#g1 #f2 #I #K1 #K2 #X #H elim (lsubf_inv_push1 … H) -H
+#g2 #I #Y #H0 #H2 #H destruct /2 width=3 by ex2_intro/
+qed-.
+
+lemma lsubf_inv_bind_sn: ∀g1,f2,I,K1,K2. ⦃K1.ⓘ{I}, ⫯g1⦄ ⫃𝐅* ⦃K2.ⓘ{I}, f2⦄ →
+ ∃∃g2. ⦃K1, g1⦄ ⫃𝐅* ⦃K2, g2⦄ & f2 = ⫯g2.
+#g1 #f2 * #I [2: #X ] #K1 #K2 #H
+[ elim (lsubf_inv_pair1 … H) -H *
+ [ #z2 #Y2 #H2 #H #H0 destruct /2 width=3 by ex2_intro/
+ | #z #z0 #z2 #Y2 #W #V #_ #_ #_ #_ #H0 #_ #H destruct
+ | #z #z0 #z2 #Z2 #Y2 #_ #_ #_ #_ #H destruct
+ ]
+| elim (lsubf_inv_unit1 … H) -H
+ #z2 #Y2 #H2 #H #H0 destruct /2 width=3 by ex2_intro/
+]
+qed-.
+
+lemma lsubf_inv_beta_sn: ∀g1,f2,K1,K2,V,W. ⦃K1.ⓓⓝW.V, ⫯g1⦄ ⫃𝐅* ⦃K2.ⓛW, f2⦄ →
+ ∃∃g,g0,g2. ⦃K1, g0⦄ ⫃𝐅* ⦃K2, g2⦄ & K1 ⊢ 𝐅*⦃V⦄ ≡ g & g0 ⋓ g ≡ g1 & f2 = ⫯g2.
+#g1 #f2 #K1 #K2 #V #W #H elim (lsubf_inv_pair1 … H) -H *
+[ #z2 #Y2 #_ #_ #H destruct
+| #z #z0 #z2 #Y2 #X0 #X #H02 #Hz #Hg1 #H #_ #H0 #H1 destruct
+ /2 width=7 by ex4_3_intro/
+| #z #z0 #z2 #Z2 #Y2 #_ #_ #_ #_ #H destruct
+]
+qed-.
+
+lemma lsubf_inv_unit_sn: ∀g1,f2,I,J,K1,K2,V. ⦃K1.ⓑ{I}V, ⫯g1⦄ ⫃𝐅* ⦃K2.ⓤ{J}, f2⦄ →
+ ∃∃g,g0,g2. ⦃K1, g0⦄ ⫃𝐅* ⦃K2, g2⦄ & K1 ⊢ 𝐅*⦃V⦄ ≡ g & g0 ⋓ g ≡ g1 & f2 = ⫯g2.
+#g1 #f2 #I #J #K1 #K2 #V #H elim (lsubf_inv_pair1 … H) -H *
+[ #z2 #Y2 #_ #_ #H destruct
+| #z #z0 #z2 #Y2 #X0 #X #_ #_ #_ #_ #_ #_ #H destruct
+| #z #z0 #z2 #Z2 #Y2 #H02 #Hz #Hg1 #H0 #H1 destruct
+ /2 width=7 by ex4_3_intro/
+]
+qed-.
+
+lemma lsubf_inv_refl: ∀L,f1,f2. ⦃L,f1⦄ ⫃𝐅* ⦃L,f2⦄ → f1 ≗ f2.
+#L elim L -L /2 width=1 by lsubf_inv_atom/
+#L #I #IH #f1 #f2 #H12
+elim (pn_split f1) * #g1 #H destruct
+[ elim (lsubf_inv_push_sn … H12) | elim (lsubf_inv_bind_sn … H12) ] -H12
+#g2 #H12 #H destruct /3 width=5 by eq_next, eq_push/
+qed-.
+
(* Basic forward lemmas *****************************************************)
+lemma lsubf_fwd_bind_tl: ∀f1,f2,I,L1,L2.
+ ⦃L1.ⓘ{I}, f1⦄ ⫃𝐅* ⦃L2.ⓘ{I}, f2⦄ → ⦃L1, ⫱f1⦄ ⫃𝐅* ⦃L2, ⫱f2⦄.
+#f1 #f2 #I #L1 #L2 #H
+elim (pn_split f1) * #g1 #H0 destruct
+[ elim (lsubf_inv_push_sn … H) | elim (lsubf_inv_bind_sn … H) ] -H
+#g2 #H12 #H destruct //
+qed-.
+
+lemma lsubf_fwd_isid_dx: ∀f1,f2,L1,L2. ⦃L1, f1⦄ ⫃𝐅* ⦃L2, f2⦄ → 𝐈⦃f2⦄ → 𝐈⦃f1⦄.
+#f1 #f2 #L1 #L2 #H elim H -f1 -f2 -L1 -L2
+[ /2 width=3 by isid_eq_repl_fwd/
+| /4 width=3 by isid_inv_push, isid_push/
+| #f1 #f2 #I #L1 #L2 #_ #_ #H elim (isid_inv_next … H) -H //
+| #f #f0 #f1 #f2 #L1 #L2 #W #V #_ #_ #_ #_ #H elim (isid_inv_next … H) -H //
+| #f #f0 #f1 #f2 #I1 #I2 #L1 #L2 #V #_ #_ #_ #_ #H elim (isid_inv_next … H) -H //
+]
+qed-.
+
+lemma lsubf_fwd_isid_sn: ∀f1,f2,L1,L2. ⦃L1, f1⦄ ⫃𝐅* ⦃L2, f2⦄ → 𝐈⦃f1⦄ → 𝐈⦃f2⦄.
+#f1 #f2 #L1 #L2 #H elim H -f1 -f2 -L1 -L2
+[ /2 width=3 by isid_eq_repl_back/
+| /4 width=3 by isid_inv_push, isid_push/
+| #f1 #f2 #I #L1 #L2 #_ #_ #H elim (isid_inv_next … H) -H //
+| #f #f0 #f1 #f2 #L1 #L2 #W #V #_ #_ #_ #_ #H elim (isid_inv_next … H) -H //
+| #f #f0 #f1 #f2 #I1 #I2 #L1 #L2 #V #_ #_ #_ #_ #H elim (isid_inv_next … H) -H //
+]
+qed-.
+
lemma lsubf_fwd_sle: ∀f1,f2,L1,L2. ⦃L1, f1⦄ ⫃𝐅* ⦃L2, f2⦄ → f2 ⊆ f1.
-#f1 #f2 #L1 #L2 * //
+#f1 #f2 #L1 #L2 #H elim H -f1 -f2 -L1 -L2
+/3 width=5 by sor_inv_sle_sn_trans, sle_next, sle_push, sle_refl_eq, eq_sym/
qed-.
(* Basic properties *********************************************************)
-lemma lsubf_pair_nn: ∀f1,f2,L1,L2. ⦃L1, f1⦄ ⫃𝐅* ⦃L2, f2⦄ →
- ∀I,V. ⦃L1.ⓑ{I}V, ⫯f1⦄ ⫃𝐅* ⦃L2.ⓑ{I}V, ⫯f2⦄.
-/4 width=5 by lsubf_fwd_sle, lsubf_pair, sle_next/ qed.
+axiom lsubf_eq_repl_back1: ∀f2,L1,L2. eq_repl_back … (λf1. ⦃L1, f1⦄ ⫃𝐅* ⦃L2, f2⦄).
+
+lemma lsubf_eq_repl_fwd1: ∀f2,L1,L2. eq_repl_fwd … (λf1. ⦃L1, f1⦄ ⫃𝐅* ⦃L2, f2⦄).
+#f2 #L1 #L2 @eq_repl_sym /2 width=3 by lsubf_eq_repl_back1/
+qed-.
+
+axiom lsubf_eq_repl_back2: ∀f1,L1,L2. eq_repl_back … (λf2. ⦃L1, f1⦄ ⫃𝐅* ⦃L2, f2⦄).
+
+lemma lsubf_eq_repl_fwd2: ∀f1,L1,L2. eq_repl_fwd … (λf2. ⦃L1, f1⦄ ⫃𝐅* ⦃L2, f2⦄).
+#f1 #L1 #L2 @eq_repl_sym /2 width=3 by lsubf_eq_repl_back2/
+qed-.
-lemma lsubf_refl: ∀L,f1,f2. f2 ⊆ f1 → ⦃L, f1⦄ ⫃𝐅* ⦃L, f2⦄.
-#L elim L -L /4 width=1 by lsubf_atom, lsubf_pair, sle_tl/
+lemma lsubf_refl: bi_reflexive … lsubf.
+#L elim L -L /2 width=1 by lsubf_atom, eq_refl/
+#L #I #IH #f elim (pn_split f) * #g #H destruct
+/2 width=1 by lsubf_push, lsubf_bind/
qed.
-lemma lsubf_sle_div: ∀f,f2,L1,L2. ⦃L1, f⦄ ⫃ 𝐅* ⦃L2, f2⦄ →
- ∀f1. f1 ⊆ f2 → ⦃L1, f⦄ ⫃ 𝐅* ⦃L2, f1⦄.
-#f #f2 #L1 #L2 #H elim H -f -f2 -L1 -L2
-/4 width=3 by lsubf_beta, lsubf_pair, lsubf_atom, sle_tl, sle_trans/
+lemma lsubf_refl_eq: ∀f1,f2,L. f1 ≗ f2 → ⦃L, f1⦄ ⫃𝐅* ⦃L, f2⦄.
+/2 width=3 by lsubf_eq_repl_back2/ qed.
+
+lemma lsubf_tl_dx: ∀g1,f2,I,L1,L2. ⦃L1, g1⦄ ⫃𝐅* ⦃L2, ⫱f2⦄ →
+ ∃∃f1. ⦃L1.ⓘ{I}, f1⦄ ⫃𝐅* ⦃L2.ⓘ{I}, f2⦄ & g1 = ⫱f1.
+#g1 #f2 #I #L1 #L2 #H
+elim (pn_split f2) * #g2 #H2 destruct
+@ex2_intro [1,2,4,5: /2 width=2 by lsubf_push, lsubf_bind/ ] // (**) (* constructor needed *)
qed-.
+
include "basic_2/static/lsubf.ma".
+axiom lsubf_inv_sor_dx: ∀f1,f2,L1,L2. ⦃L1, f1⦄ ⫃𝐅* ⦃L2, f2⦄ →
+ ∀x2,y2. x2⋓y2 ≡ f2 →
+ ∃∃x1,y1. ⦃L1, x1⦄ ⫃𝐅* ⦃L2, x2⦄ & ⦃L1, y1⦄ ⫃𝐅* ⦃L2, y2⦄ & x1⋓y1 ≡ f1.
+
(* RESTRICTED REFINEMENT FOR CONTEXT-SENSITIVE FREE VARIABLES ***************)
(* Properties with context-sensitive free variables *************************)
-lemma lsubf_frees_trans: ∀f2,L2,T. L2 ⊢ 𝐅*⦃T⦄ ≡ f2 → ∀f,L1. ⦃L1, f⦄ ⫃𝐅* ⦃L2, f2⦄ →
- â\88\83â\88\83f1. L1 â\8a¢ ð\9d\90\85*â¦\83Tâ¦\84 â\89¡ f1 & f1 â\8a\86 f.
+lemma lsubf_frees_trans: ∀f2,L2,T. L2 ⊢ 𝐅*⦃T⦄ ≡ f2 →
+ â\88\80f1,L1. â¦\83L1, f1â¦\84 â«\83ð\9d\90\85* â¦\83L2, f2â¦\84 â\86\92 L1 â\8a¢ ð\9d\90\85*â¦\83Tâ¦\84 â\89¡ f1.
#f2 #L2 #T #H elim H -f2 -L2 -T
-[ #f2 #I #Hf2 #f #L1 #H elim (lsubf_inv_atom2 … H) -H
- #H #_ destruct /3 width=3 by frees_atom, sle_isid_sn, ex2_intro/
-| #f2 #I #K2 #W #s #_ #IH #f #L1 #H elim (lsubf_inv_pair2 … H) -H *
- [ #K1 #_ #H12 #H | #g #K1 #V #Hg #Hf #_ #H12 #H1 #H2 ]
- destruct elim (IH … H12) -K2
- /3 width=3 by frees_sort, sle_inv_tl_dx, ex2_intro/
-| #f2 #I #K2 #W #_ #IH #f #L1 #H elim (lsubf_inv_pair2 … H) -H *
- [ #K1 #H elim (sle_inv_nx … H ??) -H [ <tl_next_rew |*: // ]
- #g2 #_ #H1 #H12 #H2 destruct elim (IH … H12) -K2
- /3 width=7 by frees_zero, sle_next, ex2_intro/
- | #g #K1 #V #Hg #Hf #H elim (sle_inv_nx … H ??) -H [ <tl_next_rew |*: // ]
- #g2 #_ #H1 #H12 #H2 #H3 destruct elim (IH … H12) -K2
- #f1 #Hf1 elim (sor_isfin_ex … f1 g ??)
- /4 width=7 by frees_fwd_isfin, frees_flat, frees_zero, sor_inv_sle, sle_next, ex2_intro/
+[ /3 width=5 by lsubf_fwd_isid_dx, frees_sort/
+| #f2 #i #Hf2 #g1 #Y1 #H
+ elim (lsubf_inv_atom2 … H) -H #Hg1 #H destruct
+ elim (eq_inv_pushs_dx … Hg1) -Hg1 #g #Hg #H destruct
+ elim (eq_inv_xn … Hg) -Hg
+ /3 width=3 by frees_atom, isid_eq_repl_fwd/
+| #f2 #I #K2 #W #_ #IH #g1 #Y1 #H elim (lsubf_inv_pair2 … H) -H *
+ [ #f1 #K1 #H12 #H1 #H2 destruct /3 width=1 by frees_pair/
+ | #f #f0 #f1 #K1 #V #H12 #Hf #Hf1 #H1 #H2 #H3 destruct
+ /4 width=5 by frees_pair, frees_flat/
+ ]
+| #f2 #I #L2 #Hf2 #g1 #Y1 #H elim (lsubf_inv_unit2 … H) -H *
+ [ #f1 #L1 #H12 #H1 #H2 destruct
+ /3 width=5 by lsubf_fwd_isid_dx, frees_unit/
+ | #f #f0 #f1 #J #L1 #V #H12 #Hf #Hf1 #H1 #H2 destruct
+ /5 width=9 by lsubf_fwd_isid_dx, frees_eq_repl_back, frees_pair, sor_isid_inv_sn/
]
-| #f2 #I #K2 #W #i #_ #IH #f #L1 #H elim (lsubf_inv_pair2 … H) -H *
- [ #K1 #_ #H12 #H | #g #K1 #V #Hg #Hf #_ #H12 #H1 #H2 ]
- destruct elim (IH … H12) -K2
- /3 width=3 by frees_lref, sle_inv_tl_dx, ex2_intro/
-| #f2 #I #K2 #W #l #_ #IH #f #L1 #H elim (lsubf_inv_pair2 … H) -H *
- [ #K1 #_ #H12 #H | #g #K1 #V #Hg #Hf #_ #H12 #H1 #H2 ]
- destruct elim (IH … H12) -K2
- /3 width=3 by frees_gref, sle_inv_tl_dx, ex2_intro/
-| #f2V #f2T #f2 #p #I #L2 #V #T #_ #_ #Hf2 #IHV #IHT #f #L1 #H12
- elim (IHV f L1) -IHV [2: /3 width=4 by lsubf_sle_div, sor_inv_sle_sn/ ]
- elim (IHT (⫯f) (L1.ⓑ{I}V)) -IHT [2: /4 width=4 by lsubf_sle_div, lsubf_pair_nn, sor_inv_sle_dx, sor_inv_tl_dx/ ]
- -f2V -f2T -f2 -L2 #f1T #HT #Hf1T #f1V #HV #Hf1V elim (sor_isfin_ex … f1V (⫱f1T) ??)
- /4 width=9 by frees_fwd_isfin, frees_bind, sor_inv_sle, sle_xn_tl, isfin_tl, ex2_intro/
-| #f2V #f2T #f2 #I #L2 #V #T #_ #_ #Hf2 #IHV #IHT #f #L1 #H12
- elim (IHV f L1) -IHV [2: /3 width=4 by lsubf_sle_div, sor_inv_sle_sn/ ]
- elim (IHT f L1) -IHT [2: /3 width=4 by lsubf_sle_div, sor_inv_sle_dx/ ]
- -f2V -f2T -f2 -L2 #f1T #HT #Hf1T #f1V #HV #Hf1V elim (sor_isfin_ex … f1V f1T ??)
- /3 width=7 by frees_fwd_isfin, frees_flat, sor_inv_sle, ex2_intro/
+| #f2 #I #L2 #i #_ #IH #g1 #L1 #H elim (lsubf_inv_push2 … H) -H
+ /3 width=1 by frees_lref/
+| /3 width=5 by lsubf_fwd_isid_dx, frees_gref/
+| #f2V #f2T #f2 #p #I #L2 #V #T #_ #_ #Hf2 #IHV #IHT #f1 #L1 #H12
+ elim (lsubf_inv_sor_dx … H12 … Hf2) -f2 #f1V #g1T #HV #HT #Hf1
+ elim (lsubf_tl_dx … (BPair I V) … HT) -HT #f1T #HT #H destruct
+ /3 width=5 by frees_bind/
+| #f2V #f2T #f2 #I #L2 #V #T #_ #_ #Hf2 #IHV #IHT #f1 #L1 #H12
+ elim (lsubf_inv_sor_dx … H12 … Hf2) -f2 /3 width=5 by frees_flat/
]
qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "ground_2/relocation/nstream_sor.ma".
+include "basic_2/static/frees_frees.ma".
+include "basic_2/static/lsubf.ma".
+
+(* RESTRICTED REFINEMENT FOR CONTEXT-SENSITIVE FREE VARIABLES ***************)
+
+(* Main properties **********************************************************)
+
+theorem lsubf_sor: ∀K,L,g1,f1. ⦃K, g1⦄ ⫃𝐅* ⦃L, f1⦄ →
+ ∀g2,f2. ⦃K, g2⦄ ⫃𝐅* ⦃L, f2⦄ →
+ ∀g. g1 ⋓ g2 ≡ g → ∀f. f1 ⋓ f2 ≡ f → ⦃K, g⦄ ⫃𝐅* ⦃L, f⦄.
+#K elim K -K
+[ #L #g1 #f1 #H1 #g2 #f2 #H2 #g #Hg #f #Hf
+ elim (lsubf_inv_atom1 … H1) -H1 #H1 #H destruct
+ lapply (lsubf_inv_atom … H2) -H2 #H2
+ /5 width=4 by lsubf_atom, sor_mono, sor_eq_repl_back2, sor_eq_repl_back1/
+| #K #J #IH #L #g1 #f1 #H1 #g2 #f2 #H2 #g #Hg #f #Hf
+ elim (pn_split g1) * #y1 #H destruct
+ elim (pn_split g2) * #y2 #H destruct
+ [ elim (sor_inv_ppx … Hg) -Hg [|*: // ] #y #Hy #H destruct
+ elim (lsubf_inv_push1 … H1) -H1 #x1 #Z1 #Y1 #H1 #H #H0 destruct
+ elim (lsubf_inv_push_sn … H2) -H2 #x2 #H2 #H destruct
+ elim (sor_inv_ppx … Hf) -Hf [|*: // ] #x #Hx #H destruct
+ /3 width=8 by lsubf_push/
+ | elim (sor_inv_pnx … Hg) -Hg [|*: // ] #y #Hy #H destruct
+ elim (lsubf_inv_push1 … H1) -H1 #x1 #Z1 #Y1 #H1 #H #H0 destruct
+ generalize in match H2; -H2 cases J -J #J [| #V ] #H2
+ [ elim (lsubf_inv_unit1 … H2) -H2 #x2 #Y2 #H2 #H #H0 destruct
+ | elim (lsubf_inv_pair1 … H2) -H2 *
+ [ #x2 #Z2 #H2 #H #H0 destruct
+ | #y3 #y4 #x2 #Y2 #W #U #H2 #Hy3 #Hy2 #H #H0 #H3 #H4 destruct
+ | #y3 #y4 #x2 #Z2 #Y2 #H2 #Hy3 #Hy2 #H #H0 destruct
+ ]
+ ]
+ elim (sor_inv_pnx … Hf) -Hf [1,6,11,16:|*: // ] #x #Hx #H destruct
+ /3 width=12 by lsubf_unit, lsubf_beta, lsubf_bind, sor_trans2/
+ | elim (sor_inv_npx … Hg) -Hg [|*: // ] #y #Hy #H destruct
+ elim (lsubf_inv_push1 … H2) -H2 #x2 #Z2 #Y2 #H2 #H #H0 destruct
+ generalize in match H1; -H1 cases J -J #J [| #V ] #H1
+ [ elim (lsubf_inv_unit1 … H1) -H1 #x1 #Y1 #H1 #H #H0 destruct
+ | elim (lsubf_inv_pair1 … H1) -H1 *
+ [ #x1 #Z1 #H1 #H #H0 destruct
+ | #y3 #y4 #x1 #Y1 #W #U #H1 #Hy3 #Hy1 #H #H0 #H3 #H4 destruct
+ | #y3 #y4 #x1 #Z1 #Y1 #H1 #Hy3 #Hy1 #H #H0 destruct
+ ]
+ ]
+ elim (sor_inv_npx … Hf) -Hf [1,6,11,16:|*: // ] #x #Hx #H destruct
+ /3 width=12 by lsubf_unit, lsubf_beta, lsubf_bind, sor_trans1_sym/
+ | elim (sor_inv_nnx … Hg) -Hg [|*: // ] #y #Hy #H destruct
+ generalize in match H2; generalize in match H1; -H1 -H2 cases J -J #J [| #V ] #H1 #H2
+ [ elim (lsubf_inv_unit1 … H1) -H1 #x1 #Y1 #H1 #H #H0 destruct
+ elim (lsubf_inv_bind_sn … H2) -H2 #x2 #H2 #H destruct
+ | elim (lsubf_inv_pair1 … H1) -H1 *
+ [ #x1 #Z1 #H1 #H #H0 destruct
+ elim (lsubf_inv_bind_sn … H2) -H2 #x2 #H2 #H destruct
+ | #y3 #y4 #x1 #Y1 #W #U #H1 #Hy3 #Hy1 #H #H0 #H3 #H4 destruct
+ elim (lsubf_inv_beta_sn … H2) -H2 #z3 #z4 #x2 #H2 #Hz3 #Hy2 #H destruct
+ lapply (frees_mono … Hz3 … Hy3) -Hz3 #H3
+ lapply (sor_eq_repl_back2 … Hy2 … H3) -z3 #Hy2
+ | #y3 #y4 #x1 #Z1 #Y1 #H1 #Hy3 #Hy1 #H #H0 destruct
+ elim (lsubf_inv_unit_sn … H2) -H2 #z3 #z4 #x2 #H2 #Hz3 #Hy2 #H destruct
+ lapply (frees_mono … Hz3 … Hy3) -Hz3 #H3
+ lapply (sor_eq_repl_back2 … Hy2 … H3) -z3 #Hy2
+ ]
+ ]
+ elim (sor_inv_nnx … Hf) -Hf [1,6,11,16:|*: // ] #x #Hx #H destruct
+ /3 width=12 by lsubf_unit, lsubf_beta, lsubf_bind, sor_distr_dx/
+ ]
+]
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/static/lsubr.ma".
+include "basic_2/static/lsubf_lsubf.ma".
+
+(* RESTRICTED REFINEMENT FOR CONTEXT-SENSITIVE FREE VARIABLES ***************)
+
+(* Forward lemmas with restricted refinement for local environments *********)
+
+lemma lsubf_fwd_lsubr_isdiv: ∀f1,f2,L1,L2. ⦃L1, f1⦄ ⫃𝐅* ⦃L2, f2⦄ →
+ 𝛀⦃f1⦄ → 𝛀⦃f2⦄ → L1 ⫃ L2.
+#f1 #f2 #L1 #L2 #H elim H -f1 -f2 -L1 -L2
+/4 width=3 by lsubr_bind, isdiv_inv_next/
+[ #f1 #f2 #I1 #I2 #L1 #L2 #_ #_ #H
+ elim (isdiv_inv_push … H) //
+| /5 width=5 by lsubf_fwd_sle, lsubr_beta, sle_inv_isdiv_sn, isdiv_inv_next/
+| /5 width=5 by lsubf_fwd_sle, lsubr_unit, sle_inv_isdiv_sn, isdiv_inv_next/
+]
+qed-.
+
+(* Properties with restricted refinement for local environments *************)
+
+lemma lsubr_lsubf_isid: ∀L1,L2. L1 ⫃ L2 →
+ ∀f1,f2. 𝐈⦃f1⦄ → 𝐈⦃f2⦄ → ⦃L1, f1⦄ ⫃𝐅* ⦃L2, f2⦄.
+#L1 #L2 #H elim H -L1 -L2
+[ /3 width=1 by lsubf_atom, isid_inv_eq_repl/
+| #I #L1 #L2 | #L1 #L2 #V #W | #I1 #I2 #L1 #L2 #V
+]
+#_ #IH #f1 #f2 #Hf1 #Hf2
+elim (isid_inv_gen … Hf1) -Hf1 #g1 #Hg1 #H destruct
+elim (isid_inv_gen … Hf2) -Hf2 #g2 #Hg2 #H destruct
+/3 width=1 by lsubf_push/
+qed.
+
+lemma lsubr_lsubf: ∀f2,L2,T. L2 ⊢ 𝐅*⦃T⦄ ≡ f2 → ∀L1. L1 ⫃ L2 →
+ ∀f1. L1 ⊢ 𝐅*⦃T⦄ ≡ f1 → ⦃L1, f1⦄ ⫃𝐅* ⦃L2, f2⦄.
+#f2 #L2 #T #H elim H -f2 -L2 -T
+[ #f2 #L2 #s #Hf2 #L1 #HL12 #f1 #Hf1
+ lapply (frees_inv_sort … Hf1) -Hf1 /2 width=1 by lsubr_lsubf_isid/
+| #f2 #i #Hf2 #Y1 #HY1
+ >(lsubr_inv_atom2 … HY1) -Y1 #g1 #Hg1
+ elim (frees_inv_atom … Hg1) -Hg1 #f1 #Hf1 #H destruct
+ /5 width=5 by lsubf_atom, isid_inv_eq_repl, pushs_eq_repl, eq_next/
+| #f2 #Z #L2 #W #_ #IH #Y1 #HY1 #g1 #Hg1 elim (lsubr_inv_pair2 … HY1) -HY1 *
+ [ #L1 #HL12 #H destruct
+ elim (frees_inv_pair … Hg1) -Hg1 #f1 #Hf1 #H destruct
+ /3 width=1 by lsubf_bind/
+ | #L1 #V #HL12 #H1 #H2 destruct
+ elim (frees_inv_pair … Hg1) -Hg1 #f1 #Hf1 #H destruct
+ elim (frees_inv_flat … Hf1) -Hf1 /3 width=5 by lsubf_beta/
+ ]
+| #f2 #I2 #L2 #Hf2 #Y1 #HY1 #g1 #Hg1 elim (lsubr_inv_unit2 … HY1) -HY1 *
+ [ #L1 #HL12 #H destruct
+ elim (frees_inv_unit … Hg1) -Hg1 #f1 #Hf1 #H destruct
+ /3 width=1 by lsubf_bind, lsubr_lsubf_isid/
+ | #I #L1 #V #HL12 #H destruct
+ elim (frees_inv_pair … Hg1) -Hg1 #f1 #Hf1 #H destruct
+ /3 width=5 by lsubf_unit, sor_isid_sn, lsubr_lsubf_isid/
+ ]
+| #f2 #I2 #L2 #i #_ #IH #Y1 #HY1 #g1 #Hg1
+ elim (lsubr_fwd_bind2 … HY1) -HY1 #I1 #L1 #HL12 #H destruct
+ elim (frees_inv_lref … Hg1) -Hg1 #f1 #Hf1 #H destruct
+ /3 width=1 by lsubf_push/
+| #f2 #L2 #l #Hf2 #L1 #HL12 #f1 #Hf1
+ lapply (frees_inv_gref … Hf1) -Hf1 /2 width=1 by lsubr_lsubf_isid/
+| #f2V #f2T #f2 #p #I #L2 #V #T #_ #_ #Hf2 #IHV #IHT #L1 #HL12 #f1 #Hf1
+ elim (frees_inv_bind … Hf1) -Hf1 #f1V #f1T #Hf1V #Hf1T #Hf1
+ /5 width=8 by lsubf_sor, lsubf_fwd_bind_tl, lsubr_bind/
+| #f2V #f2T #f2 #I #L2 #V #T #_ #_ #Hf2 #IHV #IHT #L1 #HL12 #f1 #Hf1
+ elim (frees_inv_flat … Hf1) -Hf1 #f1V #f1T #Hf1V #Hf1T #Hf1
+ /3 width=8 by lsubf_sor/
+]
+qed.
(* RESTRICTED REFINEMENT FOR LOCAL ENVIRONMENTS *****************************)
+(* Basic_2A1: just tpr_cpr and tprs_cprs require the extended lsubr_atom *)
+(* Basic_2A1: includes: lsubr_pair *)
inductive lsubr: relation lenv ≝
-| lsubr_atom: ∀L. lsubr L (⋆)
-| lsubr_pair: ∀I,L1,L2,V. lsubr L1 L2 → lsubr (L1.ⓑ{I}V) (L2.ⓑ{I}V)
+| lsubr_atom: lsubr (⋆) (⋆)
+| lsubr_bind: ∀I,L1,L2. lsubr L1 L2 → lsubr (L1.ⓘ{I}) (L2.ⓘ{I})
| lsubr_beta: ∀L1,L2,V,W. lsubr L1 L2 → lsubr (L1.ⓓⓝW.V) (L2.ⓛW)
+| lsubr_unit: ∀I1,I2,L1,L2,V. lsubr L1 L2 → lsubr (L1.ⓑ{I1}V) (L2.ⓤ{I2})
.
interpretation
(* Basic properties *********************************************************)
lemma lsubr_refl: ∀L. L ⫃ L.
-#L elim L -L /2 width=1 by lsubr_atom, lsubr_pair/
+#L elim L -L /2 width=1 by lsubr_atom, lsubr_bind/
qed.
(* Basic inversion lemmas ***************************************************)
fact lsubr_inv_atom1_aux: ∀L1,L2. L1 ⫃ L2 → L1 = ⋆ → L2 = ⋆.
#L1 #L2 * -L1 -L2 //
-[ #I #L1 #L2 #V #_ #H destruct
+[ #I #L1 #L2 #_ #H destruct
| #L1 #L2 #V #W #_ #H destruct
+| #I1 #I2 #L1 #L2 #V #_ #H destruct
]
qed-.
lemma lsubr_inv_atom1: ∀L2. ⋆ ⫃ L2 → L2 = ⋆.
/2 width=3 by lsubr_inv_atom1_aux/ qed-.
-fact lsubr_inv_abst1_aux: ∀L1,L2. L1 ⫃ L2 → ∀K1,W. L1 = K1.ⓛW →
- L2 = ⋆ ∨ ∃∃K2. K1 ⫃ K2 & L2 = K2.ⓛW.
+fact lsubr_inv_bind1_aux: ∀L1,L2. L1 ⫃ L2 → ∀I,K1. L1 = K1.ⓘ{I} →
+ ∨∨ ∃∃K2. K1 ⫃ K2 & L2 = K2.ⓘ{I}
+ | ∃∃K2,V,W. K1 ⫃ K2 & L2 = K2.ⓛW &
+ I = BPair Abbr (ⓝW.V)
+ | ∃∃J1,J2,K2,V. K1 ⫃ K2 & L2 = K2.ⓤ{J2} &
+ I = BPair J1 V.
#L1 #L2 * -L1 -L2
-[ #L #K1 #W #H destruct /2 width=1 by or_introl/
-| #I #L1 #L2 #V #HL12 #K1 #W #H destruct /3 width=3 by ex2_intro, or_intror/
-| #L1 #L2 #V1 #V2 #_ #K1 #W #H destruct
+[ #J #K1 #H destruct
+| #I #L1 #L2 #HL12 #J #K1 #H destruct /3 width=3 by or3_intro0, ex2_intro/
+| #L1 #L2 #V #W #HL12 #J #K1 #H destruct /3 width=6 by or3_intro1, ex3_3_intro/
+| #I1 #I2 #L1 #L2 #V #HL12 #J #K1 #H destruct /3 width=4 by or3_intro2, ex3_4_intro/
]
qed-.
-lemma lsubr_inv_abst1: ∀K1,L2,W. K1.ⓛW ⫃ L2 →
- L2 = ⋆ ∨ ∃∃K2. K1 ⫃ K2 & L2 = K2.ⓛW.
-/2 width=3 by lsubr_inv_abst1_aux/ qed-.
+(* Basic_2A1: uses: lsubr_inv_pair1 *)
+lemma lsubr_inv_bind1: ∀I,K1,L2. K1.ⓘ{I} ⫃ L2 →
+ ∨∨ ∃∃K2. K1 ⫃ K2 & L2 = K2.ⓘ{I}
+ | ∃∃K2,V,W. K1 ⫃ K2 & L2 = K2.ⓛW &
+ I = BPair Abbr (ⓝW.V)
+ | ∃∃J1,J2,K2,V. K1 ⫃ K2 & L2 = K2.ⓤ{J2} &
+ I = BPair J1 V.
+/2 width=3 by lsubr_inv_bind1_aux/ qed-.
+
+fact lsubr_inv_atom2_aux: ∀L1,L2. L1 ⫃ L2 → L2 = ⋆ → L1 = ⋆.
+#L1 #L2 * -L1 -L2 //
+[ #I #L1 #L2 #_ #H destruct
+| #L1 #L2 #V #W #_ #H destruct
+| #I1 #I2 #L1 #L2 #V #_ #H destruct
+]
+qed-.
-fact lsubr_inv_pair2_aux: ∀L1,L2. L1 ⫃ L2 → ∀I,K2,W. L2 = K2.ⓑ{I}W →
- (∃∃K1. K1 ⫃ K2 & L1 = K1.ⓑ{I}W) ∨
- ∃∃K1,V. K1 ⫃ K2 & L1 = K1.ⓓⓝW.V & I = Abst.
+lemma lsubr_inv_atom2: ∀L1. L1 ⫃ ⋆ → L1 = ⋆.
+/2 width=3 by lsubr_inv_atom2_aux/ qed-.
+
+fact lsubr_inv_bind2_aux: ∀L1,L2. L1 ⫃ L2 → ∀I,K2. L2 = K2.ⓘ{I} →
+ ∨∨ ∃∃K1. K1 ⫃ K2 & L1 = K1.ⓘ{I}
+ | ∃∃K1,W,V. K1 ⫃ K2 & L1 = K1.ⓓⓝW.V & I = BPair Abst W
+ | ∃∃J1,J2,K1,V. K1 ⫃ K2 & L1 = K1.ⓑ{J1}V & I = BUnit J2.
#L1 #L2 * -L1 -L2
-[ #L #J #K2 #W #H destruct
-| #I #L1 #L2 #V #HL12 #J #K2 #W #H destruct /3 width=3 by ex2_intro, or_introl/
-| #L1 #L2 #V1 #V2 #HL12 #J #K2 #W #H destruct /3 width=4 by ex3_2_intro, or_intror/
+[ #J #K2 #H destruct
+| #I #L1 #L2 #HL12 #J #K2 #H destruct /3 width=3 by ex2_intro, or3_intro0/
+| #L1 #L2 #V1 #V2 #HL12 #J #K2 #H destruct /3 width=6 by ex3_3_intro, or3_intro1/
+| #I1 #I2 #L1 #L2 #V #HL12 #J #K2 #H destruct /3 width=5 by ex3_4_intro, or3_intro2/
]
qed-.
-lemma lsubr_inv_pair2: ∀I,L1,K2,W. L1 ⫃ K2.ⓑ{I}W →
- (∃∃K1. K1 ⫃ K2 & L1 = K1.ⓑ{I}W) ∨
- ∃∃K1,V1. K1 ⫃ K2 & L1 = K1.ⓓⓝW.V1 & I = Abst.
-/2 width=3 by lsubr_inv_pair2_aux/ qed-.
+lemma lsubr_inv_bind2: ∀I,L1,K2. L1 ⫃ K2.ⓘ{I} →
+ ∨∨ ∃∃K1. K1 ⫃ K2 & L1 = K1.ⓘ{I}
+ | ∃∃K1,W,V. K1 ⫃ K2 & L1 = K1.ⓓⓝW.V & I = BPair Abst W
+ | ∃∃J1,J2,K1,V. K1 ⫃ K2 & L1 = K1.ⓑ{J1}V & I = BUnit J2.
+/2 width=3 by lsubr_inv_bind2_aux/ qed-.
(* Advanced inversion lemmas ************************************************)
+lemma lsubr_inv_abst1: ∀K1,L2,W. K1.ⓛW ⫃ L2 →
+ ∨∨ ∃∃K2. K1 ⫃ K2 & L2 = K2.ⓛW
+ | ∃∃I2,K2. K1 ⫃ K2 & L2 = K2.ⓤ{I2}.
+#K1 #L2 #W #H elim (lsubr_inv_bind1 … H) -H *
+/3 width=4 by ex2_2_intro, ex2_intro, or_introl, or_intror/
+#K2 #V2 #W2 #_ #_ #H destruct
+qed-.
+
+lemma lsubr_inv_unit1: ∀I,K1,L2. K1.ⓤ{I} ⫃ L2 →
+ ∃∃K2. K1 ⫃ K2 & L2 = K2.ⓤ{I}.
+#I #K1 #L2 #H elim (lsubr_inv_bind1 … H) -H *
+[ #K2 #HK12 #H destruct /2 width=3 by ex2_intro/
+| #K2 #V #W #_ #_ #H destruct
+| #I1 #I2 #K2 #V #_ #_ #H destruct
+]
+qed-.
+
+lemma lsubr_inv_pair2: ∀I,L1,K2,W. L1 ⫃ K2.ⓑ{I}W →
+ ∨∨ ∃∃K1. K1 ⫃ K2 & L1 = K1.ⓑ{I}W
+ | ∃∃K1,V. K1 ⫃ K2 & L1 = K1.ⓓⓝW.V & I = Abst.
+#I #L1 #K2 #W #H elim (lsubr_inv_bind2 … H) -H *
+[ /3 width=3 by ex2_intro, or_introl/
+| #K2 #X #V #HK12 #H1 #H2 destruct /3 width=4 by ex3_2_intro, or_intror/
+| #I1 #I1 #K2 #V #_ #_ #H destruct
+]
+qed-.
+
lemma lsubr_inv_abbr2: ∀L1,K2,V. L1 ⫃ K2.ⓓV →
∃∃K1. K1 ⫃ K2 & L1 = K1.ⓓV.
#L1 #K2 #V #H elim (lsubr_inv_pair2 … H) -H *
-[ #K1 #HK12 #H destruct /2 width=3 by ex2_intro/
-| #K1 #V1 #_ #_ #H destruct
+[ /2 width=3 by ex2_intro/
+| #K1 #X #_ #_ #H destruct
]
qed-.
lemma lsubr_inv_abst2: ∀L1,K2,W. L1 ⫃ K2.ⓛW →
- (∃∃K1. K1 ⫃ K2 & L1 = K1.ⓛW) ∨
- ∃∃K1,V. K1 ⫃ K2 & L1 = K1.ⓓⓝW.V.
+ ∨∨ ∃∃K1. K1 ⫃ K2 & L1 = K1.ⓛW
+ | ∃∃K1,V. K1 ⫃ K2 & L1 = K1.ⓓⓝW.V.
#L1 #K2 #W #H elim (lsubr_inv_pair2 … H) -H *
-[ #K1 #HK12 #H destruct /3 width=3 by ex2_intro, or_introl/
-| #K1 #V1 #HK12 #H #_ destruct /3 width=4 by ex2_2_intro, or_intror/
+/3 width=4 by ex2_2_intro, ex2_intro, or_introl, or_intror/
+qed-.
+
+lemma lsubr_inv_unit2: ∀I,L1,K2. L1 ⫃ K2.ⓤ{I} →
+ ∨∨ ∃∃K1. K1 ⫃ K2 & L1 = K1.ⓤ{I}
+ | ∃∃J,K1,V. K1 ⫃ K2 & L1 = K1.ⓑ{J}V.
+#I #L1 #K2 #H elim (lsubr_inv_bind2 … H) -H *
+[ /3 width=3 by ex2_intro, or_introl/
+| #K1 #W #V #_ #_ #H destruct
+| #I1 #I2 #K1 #V #HK12 #H1 #H2 destruct /3 width=5 by ex2_3_intro, or_intror/
]
qed-.
(* Basic forward lemmas *****************************************************)
-lemma lsubr_fwd_pair2: ∀I2,L1,K2,V2. L1 ⫃ K2.ⓑ{I2}V2 →
- ∃∃I1,K1,V1. K1 ⫃ K2 & L1 = K1.ⓑ{I1}V1.
-#I2 #L1 #K2 #V2 #H elim (lsubr_inv_pair2 … H) -H *
-[ #K1 #HK12 #H destruct /3 width=5 by ex2_3_intro/
-| #K1 #V1 #HK12 #H1 #H2 destruct /3 width=5 by ex2_3_intro/
+lemma lsubr_fwd_bind1: ∀I1,K1,L2. K1.ⓘ{I1} ⫃ L2 →
+ ∃∃I2,K2. K1 ⫃ K2 & L2 = K2.ⓘ{I2}.
+#I1 #K1 #L2 #H elim (lsubr_inv_bind1 … H) -H *
+[ #K2 #HK12 #H destruct /3 width=4 by ex2_2_intro/
+| #K2 #W1 #V1 #HK12 #H1 #H2 destruct /3 width=4 by ex2_2_intro/
+| #I1 #I2 #K2 #V1 #HK12 #H1 #H2 destruct /3 width=4 by ex2_2_intro/
+]
+qed-.
+
+lemma lsubr_fwd_bind2: ∀I2,L1,K2. L1 ⫃ K2.ⓘ{I2} →
+ ∃∃I1,K1. K1 ⫃ K2 & L1 = K1.ⓘ{I1}.
+#I2 #L1 #K2 #H elim (lsubr_inv_bind2 … H) -H *
+[ #K1 #HK12 #H destruct /3 width=4 by ex2_2_intro/
+| #K1 #W1 #V1 #HK12 #H1 #H2 destruct /3 width=4 by ex2_2_intro/
+| #I1 #I2 #K1 #V1 #HK12 #H1 #H2 destruct /3 width=4 by ex2_2_intro/
]
qed-.
(* Forward lemmas with generic slicing for local environments ***************)
(* Basic_2A1: includes: lsubr_fwd_drop2_pair *)
-lemma lsubr_fwd_drops2_pair: ∀L1,L2. L1 ⫃ L2 →
- ∀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.
+lemma lsubr_fwd_drops2_bind: ∀L1,L2. L1 ⫃ L2 →
+ ∀b,f,I,K2. 𝐔⦃f⦄ → ⬇*[b, f] L2 ≡ K2.ⓘ{I} →
+ ∨∨ ∃∃K1. K1 ⫃ K2 & ⬇*[b, f] L1 ≡ K1.ⓘ{I}
+ | ∃∃K1,W,V. K1 ⫃ K2 & ⬇*[b, f] L1 ≡ K1.ⓓⓝW.V & I = BPair Abst W
+ | ∃∃J1,J2,K1,V. K1 ⫃ K2 & ⬇*[b, f] L1 ≡ K1.ⓑ{J1}V & I = BUnit J2.
#L1 #L2 #H elim H -L1 -L2
-[ #L #b #f #I #K2 #W #_ #H
+[ #b #f #I #K2 #_ #H
elim (drops_inv_atom1 … H) -H #H destruct
-| #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/
- | #g #Hg #HLK2 #H destruct -HL12
- 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 #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/
- | #g #Hg #HLK2 #H destruct -HL12
- elim (IH … Hg HLK2) -IH -Hg -HLK2 *
- /4 width=4 by drops_drop, ex3_2_intro, ex2_intro, or_introl, or_intror/
- ]
+| #J #L1 #L2 | #L1 #L2 #V #W | #I1 #I2 #L1 #L2 #V1
+]
+#HL12 #IH #b #f #I #K2 #Hf #H
+elim (drops_inv_bind1_isuni … Hf H) -Hf -H *
+[1,3,5: #Hf #H destruct -IH
+ /4 width=6 by drops_refl, or3_intro0, or3_intro1, or3_intro2, ex3_4_intro, ex3_3_intro, ex2_intro/
+|2,4,6: #g #Hg #HLK2 #H destruct -HL12
+ elim (IH … Hg HLK2) -IH -Hg -HLK2 *
+ /4 width=6 by drops_drop, or3_intro0, or3_intro1, or3_intro2, ex3_4_intro, ex3_3_intro, ex2_intro/
]
qed-.
∀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
+elim (lsubr_fwd_drops2_bind … HL12 … Hf HLK2) -L2 -Hf // *
+[ #K1 #W #V #_ #_ #H destruct
+| #I1 #I2 #K1 #V #_ #_ #H destruct
+]
qed-.
(* Forward lemmas with length for local environments ************************)
-lemma lsubr_fwd_length: ∀L1,L2. L1 ⫃ L2 → |L2| ≤ |L1|.
-#L1 #L2 #H elim H -L1 -L2 /2 width=1 by O, le_S_S/
+lemma lsubr_fwd_length: ∀L1,L2. L1 ⫃ L2 → |L2| = |L1|.
+#L1 #L2 #H elim H -L1 -L2 normalize //
qed-.
(* RESTRICTED REFINEMENT FOR LOCAL ENVIRONMENTS *****************************)
-(* Auxiliary inversion lemmas ***********************************************)
-
-fact lsubr_inv_pair1_aux: ∀L1,L2. L1 ⫃ L2 → ∀I,K1,X. L1 = K1.ⓑ{I}X →
- ∨∨ L2 = ⋆
- | ∃∃K2. K1 ⫃ K2 & L2 = K2.ⓑ{I}X
- | ∃∃K2,V,W. K1 ⫃ K2 & L2 = K2.ⓛW &
- I = Abbr & X = ⓝW.V.
-#L1 #L2 * -L1 -L2
-[ #L #J #K1 #X #H destruct /2 width=1 by or3_intro0/
-| #I #L1 #L2 #V #HL12 #J #K1 #X #H destruct /3 width=3 by or3_intro1, ex2_intro/
-| #L1 #L2 #V #W #HL12 #J #K1 #X #H destruct /3 width=6 by or3_intro2, ex4_3_intro/
-]
-qed-.
-
-lemma lsubr_inv_pair1: ∀I,K1,L2,X. K1.ⓑ{I}X ⫃ L2 →
- ∨∨ L2 = ⋆
- | ∃∃K2. K1 ⫃ K2 & L2 = K2.ⓑ{I}X
- | ∃∃K2,V,W. K1 ⫃ K2 & L2 = K2.ⓛW &
- I = Abbr & X = ⓝW.V.
-/2 width=3 by lsubr_inv_pair1_aux/ qed-.
-
(* Main properties **********************************************************)
theorem lsubr_trans: Transitive … lsubr.
-#L1 #L #H elim H -L1 -L
-[ #L1 #X #H
- lapply (lsubr_inv_atom1 … H) -H //
-| #I #L1 #L #V #_ #IHL1 #X #H
- elim (lsubr_inv_pair1 … H) -H // *
- #L2 [2: #V2 #W2 ] #HL2 #H1 [ #H2 #H3 ] destruct /3 width=1 by lsubr_pair, lsubr_beta/
-| #L1 #L #V1 #W #_ #IHL1 #X #H
- elim (lsubr_inv_abst1 … H) -H // *
- #L2 #HL2 #H destruct /3 width=1 by lsubr_beta/
+#L1 #L #H elim H -L1 -L //
+[ #I #L1 #L #_ #IH #X #H elim (lsubr_inv_bind1 … H) -H *
+ [ #L2 #HL2 #H | #L2 #V #W #HL2 #H1 #H2 | #I1 #I2 #L2 #V #Hl2 #H1 #H2 ]
+ destruct /3 width=1 by lsubr_bind, lsubr_beta, lsubr_unit/
+| #L1 #L #V #W #_ #IH #X #H elim (lsubr_inv_abst1 … H) -H *
+ [ #L2 #HL2 #H | #I #L2 #HL2 #H ]
+ destruct /3 width=1 by lsubr_beta, lsubr_unit/
+| #I1 #I2 #L1 #L #V #_ #IH #X #H elim (lsubr_inv_unit1 … H) -H
+ /4 width=1 by lsubr_unit/
]
qed-.
(**************************************************************************)
include "ground_2/notation/functions/append_2.ma".
+include "basic_2/notation/functions/snitem_2.ma".
include "basic_2/notation/functions/snbind1_2.ma".
include "basic_2/notation/functions/snbind2_3.ma".
include "basic_2/notation/functions/snvoid_1.ma".
(* APPEND FOR LOCAL ENVIRONMENTS ********************************************)
rec definition append L K on K ≝ match K with
-[ LAtom ⇒ L
-| LPair K I V ⇒ (append L K).ⓑ{I}V
+[ LAtom ⇒ L
+| LBind K I ⇒ (append L K).ⓘ{I}
].
interpretation "append (local environment)" 'Append L1 L2 = (append L1 L2).
-(*
+
+interpretation "local environment tail binding construction (generic)"
+ 'SnItem I L = (append (LBind LAtom I) L).
+
interpretation "local environment tail binding construction (unary)"
- 'SnBind1 I L = (append (LUnit LAtom I) L).
-*)
+ 'SnBind1 I L = (append (LBind LAtom (BUnit I)) L).
+
interpretation "local environment tail binding construction (binary)"
- 'SnBind2 I T L = (append (LPair LAtom I T) L).
-(*
+ 'SnBind2 I T L = (append (LBind LAtom (BPair I T)) L).
+
interpretation "tail exclusion (local environment)"
- 'SnVoid L = (append (LUnit LAtom Void) L).
-*)
+ 'SnVoid L = (append (LBind LAtom (BUnit Void)) L).
+
interpretation "tail abbreviation (local environment)"
- 'SnAbbr T L = (append (LPair LAtom Abbr T) L).
+ 'SnAbbr T L = (append (LBind LAtom (BPair Abbr T)) L).
interpretation "tail abstraction (local environment)"
- 'SnAbst L T = (append (LPair LAtom Abst T) L).
+ 'SnAbst L T = (append (LBind LAtom (BPair Abst T)) L).
definition d_appendable_sn: predicate (lenv→relation term) ≝ λR.
∀K,T1,T2. R K T1 T2 → ∀L. R (L@@K) T1 T2.
lemma append_atom: ∀L. L @@ ⋆ = L.
// qed.
-lemma append_pair: ∀I,L,K,V. L@@(K.ⓑ{I}V) = (L@@K).ⓑ{I}V.
+(* Basic_2A1: uses: append_pair *)
+lemma append_bind: ∀I,L,K. L@@(K.ⓘ{I}) = (L@@K).ⓘ{I}.
// qed.
lemma append_atom_sn: ∀L. ⋆@@L = L.
#L elim L -L //
-#L #I #V >append_pair //
+#L #I >append_bind //
qed.
lemma append_assoc: associative … append.
#L1 #L2 #L3 elim L3 -L3 //
qed.
-lemma append_shift: ∀L,K,I,V. L@@(ⓑ{I}V.K) = (L.ⓑ{I}V)@@K.
-#L #K #I #V <append_assoc //
+lemma append_shift: ∀L,K,I. L@@(ⓘ{I}.K) = (L.ⓘ{I})@@K.
+#L #K #I <append_assoc //
qed.
(* Basic inversion lemmas ***************************************************)
lemma append_inj_sn: ∀K,L1,L2. L1@@K = L2@@K → L1 = L2.
#K elim K -K //
-#K #I #V #IH #L1 #L2 >append_pair #H
-elim (destruct_lpair_lpair_aux … H) -H /2 width=1 by/ (**) (* destruct lemma needed *)
+#K #I #IH #L1 #L2 >append_bind #H
+elim (destruct_lbind_lbind_aux … H) -H /2 width=1 by/ (**) (* destruct lemma needed *)
qed-.
(* Basic_1: uses: chead_ctail *)
(* Basic_2A1: uses: lpair_ltail *)
-lemma lenv_case_tail: ∀L. L = ⋆ ∨ ∃∃K,I,V. L = ⓑ{I}V.K.
+lemma lenv_case_tail: ∀L. L = ⋆ ∨ ∃∃K,I. L = ⓘ{I}.K.
#L elim L -L /2 width=1 by or_introl/
-#L #I #V * [2: * ] /3 width=4 by ex1_3_intro, or_intror/
+#L #I * [2: * ] /3 width=3 by ex1_2_intro, or_intror/
qed-.
lemma append_length: ∀L1,L2. |L1 @@ L2| = |L1| + |L2|.
#L1 #L2 elim L2 -L2 //
-#L2 #I #V2 >append_pair >length_pair >length_pair //
+#L2 #I >append_bind >length_bind >length_bind //
qed.
-lemma ltail_length: ∀I,L,V. |ⓑ{I}V.L| = ⫯|L|.
-#I #L #V >append_length //
+lemma ltail_length: ∀I,L. |ⓘ{I}.L| = ⫯|L|.
+#I #L >append_length //
qed.
(* Advanced inversion lemmas on length for local environments ***************)
(* Basic_2A1: was: length_inv_pos_dx_ltail *)
-lemma length_inv_succ_dx_ltail: ∀L,l. |L| = ⫯l →
- ∃∃I,K,V. |K| = l & L = ⓑ{I}V.K.
-#Y #l #H elim (length_inv_succ_dx … H) -H #I #L #V #Hl #HLK destruct
-elim (lenv_case_tail … L) [2: * #K #J #W ]
-#H destruct /2 width=5 by ex2_3_intro/
+lemma length_inv_succ_dx_ltail: ∀L,n. |L| = ⫯n →
+ ∃∃I,K. |K| = n & L = ⓘ{I}.K.
+#Y #n #H elim (length_inv_succ_dx … H) -H #I #L #Hn #HLK destruct
+elim (lenv_case_tail … L) [2: * #K #J ]
+#H destruct /2 width=4 by ex2_2_intro/
qed-.
(* Basic_2A1: was: length_inv_pos_sn_ltail *)
-lemma length_inv_succ_sn_ltail: ∀L,l. ⫯l = |L| →
- ∃∃I,K,V. l = |K| & L = ⓑ{I}V.K.
-#Y #l #H elim (length_inv_succ_sn … H) -H #I #L #V #Hl #HLK destruct
-elim (lenv_case_tail … L) [2: * #K #J #W ]
-#H destruct /2 width=5 by ex2_3_intro/
+lemma length_inv_succ_sn_ltail: ∀L,n. ⫯n = |L| →
+ ∃∃I,K. n = |K| & L = ⓘ{I}.K.
+#Y #n #H elim (length_inv_succ_sn … H) -H #I #L #Hn #HLK destruct
+elim (lenv_case_tail … L) [2: * #K #J ]
+#H destruct /2 width=4 by ex2_2_intro/
qed-.
(* Inversion lemmas with length for local environments **********************)
L1 = L2 ∧ K1 = K2.
#K1 elim K1 -K1
[ * /2 width=1 by conj/
- #K2 #I2 #V2 #L1 #L2 #_ >length_atom >length_pair
+ #K2 #I2 #L1 #L2 #_ >length_atom >length_bind
#H destruct
-| #K1 #I1 #V1 #IH *
- [ #L1 #L2 #_ >length_atom >length_pair
+| #K1 #I1 #IH *
+ [ #L1 #L2 #_ >length_atom >length_bind
#H destruct
- | #K2 #I2 #V2 #L1 #L2 #H1 >length_pair >length_pair #H2
- elim (destruct_lpair_lpair_aux … H1) -H1 #H1 #H3 #H4 destruct (**) (* destruct lemma needed *)
+ | #K2 #I2 #L1 #L2 #H1 >length_bind >length_bind #H2
+ elim (destruct_lbind_lbind_aux … H1) -H1 #H1 #H3 destruct (**) (* destruct lemma needed *)
elim (IH … H1) -IH -H1 /3 width=4 by conj/
]
]
L1 = L2 ∧ K1 = K2.
#K1 elim K1 -K1
[ * /2 width=1 by conj/
- #K2 #I2 #V2 #L1 #L2 >append_atom >append_pair #H destruct
- >length_pair >append_length >plus_n_Sm
+ #K2 #I2 #L1 #L2 >append_atom >append_bind #H destruct
+ >length_bind >append_length >plus_n_Sm
#H elim (plus_xSy_x_false … H)
-| #K1 #I1 #V1 #IH *
- [ #L1 #L2 >append_pair >append_atom #H destruct
- >length_pair >append_length >plus_n_Sm #H
+| #K1 #I1 #IH *
+ [ #L1 #L2 >append_bind >append_atom #H destruct
+ >length_bind >append_length >plus_n_Sm #H
lapply (discr_plus_x_xy … H) -H #H destruct
- | #K2 #I2 #V2 #L1 #L2 >append_pair >append_pair #H1 #H2
- elim (destruct_lpair_lpair_aux … H1) -H1 #H1 #H3 #H4 destruct (**) (* destruct lemma needed *)
+ | #K2 #I2 #L1 #L2 >append_bind >append_bind #H1 #H2
+ elim (destruct_lbind_lbind_aux … H1) -H1 #H1 #H3 destruct (**) (* destruct lemma needed *)
elim (IH … H1) -IH -H1 /2 width=1 by conj/
]
]
(* Basic_1: was: c_tail_ind *)
(* Basic_2A1: was: lenv_ind_alt *)
lemma lenv_ind_tail: ∀R:predicate lenv.
- R (⋆) → (∀I,L,T. R L → R (ⓑ{I}T.L)) → ∀L. R L.
+ R (⋆) → (∀I,L. R L → R (ⓘ{I}.L)) → ∀L. R L.
#R #IH1 #IH2 #L @(f_ind … length … L) -L #x #IHx * //
-#L #I #V -IH1 #H destruct
-elim (lenv_case_tail … L) [2: * #K #J #W ]
+#L #I -IH1 #H destruct
+elim (lenv_case_tail … L) [2: * #K #J ]
#H destruct /3 width=1 by/
qed-.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "basic_2/syntax/bind.ma".
-
-(* EXTENSION TO BINDERS OF A RELATION FOR TERMS *****************************)
-
-inductive ext2 (R:relation term): relation bind ≝
-| ext2_unit: ∀I. ext2 R (BUnit I) (BUnit I)
-| ext2_pair: ∀I,V1,V2. R V1 V2 → ext2 R (BPair I V1) (BPair I V2)
-.
-
-(* Basic_inversion lemmas **************************************************)
-
-fact ext2_inv_unit_sn_aux: ∀R,Z1,Z2. ext2 R Z1 Z2 →
- ∀I. Z1 = BUnit I → Z2 = BUnit I.
-#R #Z1 #Z2 * -Z1 -Z2 #I [2: #V1 #V2 #_ ]
-#J #H destruct //
-qed-.
-
-lemma ext2_inv_unit_sn: ∀R,I,Z2. ext2 R (BUnit I) Z2 → Z2 = BUnit I.
-/2 width=4 by ext2_inv_unit_sn_aux/ qed-.
-
-fact ext2_inv_pair_sn_aux: ∀R,Z1,Z2. ext2 R Z1 Z2 →
- ∀I,V1. Z1 = BPair I V1 →
- ∃∃V2. R V1 V2 & Z2 = BPair I V2.
-#R #Z1 #Z2 * -Z1 -Z2 #I [2: #V1 #V2 #HV12 ]
-#J #W1 #H destruct /2 width=3 by ex2_intro/
-qed-.
-
-lemma ext2_inv_pair_sn: ∀R,Z2,I,V1. ext2 R (BPair I V1) Z2 →
- ∃∃V2. R V1 V2 & Z2 = BPair I V2.
-/2 width=3 by ext2_inv_pair_sn_aux/ qed-.
-
-fact ext2_inv_unit_dx_aux: ∀R,Z1,Z2. ext2 R Z1 Z2 →
- ∀I. Z2 = BUnit I → Z1 = BUnit I.
-#R #Z1 #Z2 * -Z1 -Z2 #I [2: #V1 #V2 #_ ]
-#J #H destruct //
-qed-.
-
-lemma ext2_inv_unit_dx: ∀R,I,Z1. ext2 R Z1 (BUnit I) → Z1 = BUnit I.
-/2 width=4 by ext2_inv_unit_dx_aux/ qed-.
-
-fact ext2_inv_pair_dx_aux: ∀R,Z1,Z2. ext2 R Z1 Z2 →
- ∀I,V2. Z2 = BPair I V2 →
- ∃∃V1. R V1 V2 & Z1 = BPair I V1.
-#R #Z1 #Z2 * -Z1 -Z2 #I [2: #V1 #V2 #HV12 ]
-#J #W2 #H destruct /2 width=3 by ex2_intro/
-qed-.
-
-lemma ext2_inv_pair_dx: ∀R,Z1,I,V2. ext2 R Z1 (BPair I V2) →
- ∃∃V1. R V1 V2 & Z1 = BPair I V1.
-/2 width=3 by ext2_inv_pair_dx_aux/ qed-.
-
-(* Advanced inversion lemmas ***********************************************)
-
-lemma ext2_inv_unit: ∀R,I1,I2. ext2 R (BUnit I1) (BUnit I2) → I1 = I2.
-#R #I1 #I2 #H lapply (ext2_inv_unit_sn … H) -H
-#H destruct //
-qed-.
-
-lemma ext2_inv_pair: ∀R,I1,I2,V1,V2. ext2 R (BPair I1 V1) (BPair I2 V2) →
- I1 = I2 ∧ R V1 V2.
-#R #I1 #I2 #V1 #V2 #H elim (ext2_inv_pair_sn … H) -H
-#V #HV #H destruct /2 width=1 by conj/
-qed-.
-
-(* Basic properties ********************************************************)
-
-lemma ext2_refl: ∀R. reflexive … R → reflexive … (ext2 R).
-#R #HR * /2 width=1 by ext2_pair/
-qed.
-
-lemma ext2_sym: ∀R. symmetric … R → symmetric … (ext2 R).
-#R #HR #T1 #T2 * /3 width=1 by ext2_unit, ext2_pair/
-qed-.
-
-lemma ext2_dec: ∀R. (∀T1,T2. Decidable (R T1 T2)) →
- ∀I1,I2. Decidable (ext2 R I1 I2).
-#R #HR * #I1 [2: #T1 ] * #I2 [2,4: #T2 ]
-[ elim (eq_bind2_dec I1 I2) #HI12 destruct
- [ elim (HR T1 T2) -HR #HT12 /3 width=1 by ext2_pair, or_introl/ ]
- @or_intror #H elim (ext2_inv_pair … H) -H /2 width=1 by/
-| @or_intror #H lapply (ext2_inv_unit_sn … H) -H
- #H destruct
-| @or_intror #H lapply (ext2_inv_unit_dx … H) -H
- #H destruct
-| elim (eq_bind1_dec I1 I2) #HI12 destruct
- /4 width=2 by ext2_inv_unit, ext2_unit, or_intror, or_introl/
-]
-qed-.
lemma rfw_shift: ∀p,I,K,V,T. ♯{K.ⓑ{I}V, T} < ♯{K, ⓑ{p,I}V.T}.
normalize /2 width=1 by monotonic_le_plus_r/
qed.
-(*
+
lemma rfw_clear: ∀p,I1,I2,K,V,T. ♯{K.ⓤ{I1}, T} < ♯{K, ⓑ{p,I2}V.T}.
normalize /4 width=1 by monotonic_le_plus_r, le_S_S/
qed.
-*)
+
lemma rfw_tpair_sn: ∀I,L,V,T. ♯{L, V} < ♯{L, ②{I}V.T}.
normalize in ⊢ (?→?→?→?→?%%); //
qed.
lemma fw_shift: ∀p,I,G,K,V,T. ♯{G, K.ⓑ{I}V, T} < ♯{G, K, ⓑ{p,I}V.T}.
normalize /2 width=1 by monotonic_le_plus_r/
qed.
-(*
+
lemma fw_clear: ∀p,I1,I2,G,K,V,T. ♯{G, K.ⓤ{I1}, T} < ♯{G, K, ⓑ{p,I2}V.T}.
normalize /4 width=1 by monotonic_le_plus_r, le_S_S/
qed.
-*)
+
lemma fw_tpair_sn: ∀I,G,L,V,T. ♯{G, L, V} < ♯{G, L, ②{I}V.T}.
normalize in ⊢ (?→?→?→?→?→?%%); //
qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/syntax/bind.ma".
+
+(* EXTENSION TO BINDERS OF A RELATION FOR TERMS *****************************)
+
+inductive ext2 (R:relation term): relation bind ≝
+| ext2_unit: ∀I. ext2 R (BUnit I) (BUnit I)
+| ext2_pair: ∀I,V1,V2. R V1 V2 → ext2 R (BPair I V1) (BPair I V2)
+.
+
+(* Basic_inversion lemmas **************************************************)
+
+fact ext2_inv_unit_sn_aux: ∀R,Z1,Z2. ext2 R Z1 Z2 →
+ ∀I. Z1 = BUnit I → Z2 = BUnit I.
+#R #Z1 #Z2 * -Z1 -Z2 #I [2: #V1 #V2 #_ ]
+#J #H destruct //
+qed-.
+
+lemma ext2_inv_unit_sn: ∀R,I,Z2. ext2 R (BUnit I) Z2 → Z2 = BUnit I.
+/2 width=4 by ext2_inv_unit_sn_aux/ qed-.
+
+fact ext2_inv_pair_sn_aux: ∀R,Z1,Z2. ext2 R Z1 Z2 →
+ ∀I,V1. Z1 = BPair I V1 →
+ ∃∃V2. R V1 V2 & Z2 = BPair I V2.
+#R #Z1 #Z2 * -Z1 -Z2 #I [2: #V1 #V2 #HV12 ]
+#J #W1 #H destruct /2 width=3 by ex2_intro/
+qed-.
+
+lemma ext2_inv_pair_sn: ∀R,Z2,I,V1. ext2 R (BPair I V1) Z2 →
+ ∃∃V2. R V1 V2 & Z2 = BPair I V2.
+/2 width=3 by ext2_inv_pair_sn_aux/ qed-.
+
+fact ext2_inv_unit_dx_aux: ∀R,Z1,Z2. ext2 R Z1 Z2 →
+ ∀I. Z2 = BUnit I → Z1 = BUnit I.
+#R #Z1 #Z2 * -Z1 -Z2 #I [2: #V1 #V2 #_ ]
+#J #H destruct //
+qed-.
+
+lemma ext2_inv_unit_dx: ∀R,I,Z1. ext2 R Z1 (BUnit I) → Z1 = BUnit I.
+/2 width=4 by ext2_inv_unit_dx_aux/ qed-.
+
+fact ext2_inv_pair_dx_aux: ∀R,Z1,Z2. ext2 R Z1 Z2 →
+ ∀I,V2. Z2 = BPair I V2 →
+ ∃∃V1. R V1 V2 & Z1 = BPair I V1.
+#R #Z1 #Z2 * -Z1 -Z2 #I [2: #V1 #V2 #HV12 ]
+#J #W2 #H destruct /2 width=3 by ex2_intro/
+qed-.
+
+lemma ext2_inv_pair_dx: ∀R,Z1,I,V2. ext2 R Z1 (BPair I V2) →
+ ∃∃V1. R V1 V2 & Z1 = BPair I V1.
+/2 width=3 by ext2_inv_pair_dx_aux/ qed-.
+
+(* Advanced inversion lemmas ***********************************************)
+
+lemma ext2_inv_unit: ∀R,I1,I2. ext2 R (BUnit I1) (BUnit I2) → I1 = I2.
+#R #I1 #I2 #H lapply (ext2_inv_unit_sn … H) -H
+#H destruct //
+qed-.
+
+lemma ext2_inv_pair: ∀R,I1,I2,V1,V2. ext2 R (BPair I1 V1) (BPair I2 V2) →
+ I1 = I2 ∧ R V1 V2.
+#R #I1 #I2 #V1 #V2 #H elim (ext2_inv_pair_sn … H) -H
+#V #HV #H destruct /2 width=1 by conj/
+qed-.
+
+(* Basic properties ********************************************************)
+
+lemma ext2_refl: ∀R. reflexive … R → reflexive … (ext2 R).
+#R #HR * /2 width=1 by ext2_pair/
+qed.
+
+lemma ext2_sym: ∀R. symmetric … R → symmetric … (ext2 R).
+#R #HR #T1 #T2 * /3 width=1 by ext2_unit, ext2_pair/
+qed-.
+
+lemma ext2_dec: ∀R. (∀T1,T2. Decidable (R T1 T2)) →
+ ∀I1,I2. Decidable (ext2 R I1 I2).
+#R #HR * #I1 [2: #T1 ] * #I2 [2,4: #T2 ]
+[ elim (eq_bind2_dec I1 I2) #HI12 destruct
+ [ elim (HR T1 T2) -HR #HT12 /3 width=1 by ext2_pair, or_introl/ ]
+ @or_intror #H elim (ext2_inv_pair … H) -H /2 width=1 by/
+| @or_intror #H lapply (ext2_inv_unit_sn … H) -H
+ #H destruct
+| @or_intror #H lapply (ext2_inv_unit_dx … H) -H
+ #H destruct
+| elim (eq_bind1_dec I1 I2) #HI12 destruct
+ /4 width=2 by ext2_inv_unit, ext2_unit, or_intror, or_introl/
+]
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/syntax/ext2.ma".
+
+(* EXTENSION TO BINDERS OF A RELATION FOR TERMS *****************************)
+
+(* Main properties **********************************************************)
+
+theorem ext2_trans: ∀R. Transitive … R → Transitive … (ext2 R).
+#R #HR #I1 #I #H elim H -I1 -I
+[ #I1 #J #H >(ext2_inv_unit_sn … H) -J /2 width=1 by ext2_unit/
+| #I1 #V1 #V #HV1 #J #H elim (ext2_inv_pair_sn … H) -H
+ #V2 #HV2 #H destruct /3 width=3 by ext2_pair/
+]
+qed-.
| GRef: nat → item0 (* reference by position: starting at 0 *)
.
-(* binary binding items *)
+(* unary binding items *)
inductive bind1: Type[0] ≝
| Void: bind1 (* exclusion *)
.
(* local environments *)
inductive lenv: Type[0] ≝
-| LAtom: lenv (* empty *)
-| LPair: lenv → bind2 → term → lenv (* binary binding construction *)
+| LAtom: lenv (* empty *)
+| LBind: lenv → bind → lenv (* binding construction *)
.
interpretation "sort (local environment)"
'Star = LAtom.
-(*
+
+interpretation "local environment binding construction (generic)"
+ 'DxItem L I = (LBind L I).
+
interpretation "local environment binding construction (unary)"
- 'DxBind1 L I = (LUnit L I).
-*)
+ 'DxBind1 L I = (LBind L (BUnit I)).
+
interpretation "local environment binding construction (binary)"
- 'DxBind2 L I T = (LPair L I T).
-(*
+ 'DxBind2 L I T = (LBind L (BPair I T)).
+
interpretation "void (local environment)"
- 'DxVoid L = (LPair L Void).
-*)
+ 'DxVoid L = (LBind L (BUnit Void)).
+
interpretation "abbreviation (local environment)"
- 'DxAbbr L T = (LPair L Abbr T).
+ 'DxAbbr L T = (LBind L (BPair Abbr T)).
interpretation "abstraction (local environment)"
- 'DxAbst L T = (LPair L Abst T).
+ 'DxAbst L T = (LBind L (BPair Abst T)).
-definition ceq: relation3 lenv term term ≝ λL,T1,T2. T1 = T2.
+definition cfull: relation3 lenv bind bind ≝ λL,I1,I2. ⊤.
-definition cfull: relation3 lenv term term ≝ λL,T1,T2. ⊤.
+definition ceq: relation3 lenv bind bind ≝ λL. eq ….
(* Basic properties *********************************************************)
lemma eq_lenv_dec: ∀L1,L2:lenv. Decidable (L1 = L2).
-#L1 elim L1 -L1 [| #L1 #I1 #V1 #IHL1 ] * [2,4: #L2 #I2 #V2 ]
-[1,4: @or_intror #H destruct
-| elim (eq_bind2_dec I1 I2) #HI
- [ elim (eq_term_dec V1 V2) #HV
- [ elim (IHL1 L2) -IHL1 /2 width=1 by or_introl/ #HL ]
- ]
- @or_intror #H destruct /2 width=1 by/
-| /2 width=1 by or_introl/
+#L1 elim L1 -L1 [| #L1 #I1 #IHL1 ] * [2,4: #L2 #I2 ]
+[3: /2 width=1 by or_introl/
+|2: elim (eq_bind_dec I1 I2) #HI
+ [ elim (IHL1 L2) -IHL1 #HL /2 width=1 by or_introl/ ]
]
+@or_intror #H destruct /2 width=1 by/
qed-.
lemma cfull_dec: ∀L,T1,T2. Decidable (cfull L T1 T2).
(* Basic inversion lemmas ***************************************************)
-fact destruct_lpair_lpair_aux: ∀I1,I2,L1,L2,V1,V2. L1.ⓑ{I1}V1 = L2.ⓑ{I2}V2 →
- ∧∧L1 = L2 & I1 = I2 & V1 = V2.
-#I1 #I2 #L1 #L2 #V1 #V2 #H destruct /2 width=1 by and3_intro/
+fact destruct_lbind_lbind_aux: ∀I1,I2,L1,L2. L1.ⓘ{I1} = L2.ⓘ{I2} →
+ L1 = L2 ∧ I1 = I2.
+#I1 #I2 #L1 #L2 #H destruct /2 width=1 by conj/
qed-.
-lemma discr_lpair_x_xy: ∀I,V,L. L = L.ⓑ{I}V → ⊥.
-#I #V #L elim L -L
+(* Basic_2A1: uses: discr_lpair_x_xy *)
+lemma discr_lbind_x_xy: ∀I,L. L = L.ⓘ{I} → ⊥.
+#I #L elim L -L
[ #H destruct
-| #L #J #W #IHL #H
- elim (destruct_lpair_lpair_aux … H) -H #H1 #H2 #H3 destruct /2 width=1 by/ (**) (* destruct lemma needed *)
+| #L #J #IHL #H elim (destruct_lbind_lbind_aux … H) -H (**) (* destruct lemma needed *)
+ #H1 #H2 destruct /2 width=1 by/
]
qed-.
-lemma discr_lpair_xy_x: ∀I,V,L. L.ⓑ{I}V = L → ⊥.
-/2 width=4 by discr_lpair_x_xy/ qed-.
+(* Basic_2A1: uses: discr_lpair_xy_x *)
+lemma discr_lbind_xy_x: ∀I,L. L.ⓘ{I} = L → ⊥.
+/2 width=4 by discr_lbind_x_xy/ qed-.
(* *)
(**************************************************************************)
-include "basic_2/syntax/bind_ext2.ma".
+include "basic_2/syntax/ext2.ma".
include "basic_2/syntax/lenv.ma".
(* EXTENSION TO BINDERS OF A CONTEXT-SENSITIVE RELATION FOR TERMS ***********)
(* LENGTH OF A LOCAL ENVIRONMENT ********************************************)
rec definition length L ≝ match L with
-[ LAtom ⇒ 0
-| LPair L _ _ ⇒ ⫯(length L)
+[ LAtom ⇒ 0
+| LBind L _ ⇒ ⫯(length L)
].
interpretation "length (local environment)" 'card L = (length L).
lemma length_atom: |⋆| = 0.
// qed.
-lemma length_pair: ∀I,L,V. |L.ⓑ{I}V| = ⫯|L|.
+(* Basic_2A1: uses: length_pair *)
+lemma length_bind: ∀I,L. |L.ⓘ{I}| = ⫯|L|.
// qed.
(* Basic inversion lemmas ***************************************************)
lemma length_inv_zero_dx: ∀L. |L| = 0 → L = ⋆.
-* // #L #I #V >length_pair
+* // #L #I >length_bind
#H destruct
qed-.
(* Basic_2A1: was: length_inv_pos_dx *)
lemma length_inv_succ_dx: ∀n,L. |L| = ⫯n →
- ∃∃I,K,V. |K| = n & L = K. ⓑ{I}V.
-#n * [ >length_atom #H destruct ]
-#L #I #V >length_pair /3 width=5 by ex2_3_intro, injective_S/
+ ∃∃I,K. |K| = n & L = K. ⓘ{I}.
+#n *
+[ >length_atom #H destruct
+| #L #I >length_bind /3 width=4 by ex2_2_intro, injective_S/
+]
qed-.
(* Basic_2A1: was: length_inv_pos_sn *)
lemma length_inv_succ_sn: ∀n,L. ⫯n = |L| →
- ∃∃I,K,V. n = |K| & L = K. ⓑ{I}V.
-#l #L #H lapply (sym_eq ??? H) -H
-#H elim (length_inv_succ_dx … H) -H /2 width=5 by ex2_3_intro/
+ ∃∃I,K. n = |K| & L = K. ⓘ{I}.
+#n #L #H lapply (sym_eq ??? H) -H
+#H elim (length_inv_succ_dx … H) -H /2 width=4 by ex2_2_intro/
qed-.
(* Basic_2A1: removed theorems 1: length_inj *)
(* WEIGHT OF A LOCAL ENVIRONMENT ********************************************)
rec definition lw L ≝ match L with
-[ LAtom ⇒ 0
-| LPair L _ V ⇒ lw L + ♯{V}
+[ LAtom ⇒ 0
+| LBind L I ⇒ lw L + ♯{I}
].
interpretation "weight (local environment)" 'Weight L = (lw L).
(* Basic properties *********************************************************)
-lemma lw_pair: ∀I,L,V. ♯{L} < ♯{L.ⓑ{I}V}.
+(* Basic_2A1: uses: lw_pair *)
+lemma lw_bind: ∀I,L. ♯{L} < ♯{L.ⓘ{I}}.
normalize /2 width=1 by monotonic_le_plus_r/ qed.
(* Basic_1: removed theorems 4: clt_cong clt_head clt_thead clt_wf_ind *)
include "basic_2/notation/relations/lazyeq_4.ma".
include "basic_2/syntax/item_sd.ma".
-include "basic_2/syntax/lenv.ma".
+include "basic_2/syntax/term.ma".
(* DEGREE-BASED EQUIVALENCE ON TERMS ****************************************)
"degree-based equivalence (terms)"
'LazyEq h o T1 T2 = (tdeq h o T1 T2).
-definition cdeq: ∀h. sd h → relation3 lenv term term ≝
- λh,o,L. tdeq h o.
-
(* Basic inversion lemmas ***************************************************)
fact tdeq_inv_sort1_aux: ∀h,o,X,Y. X ≡[h, o] Y → ∀s1. X = ⋆s1 →
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/syntax/tdeq.ma".
+include "basic_2/syntax/lenv_ext2.ma".
+
+(* EXTENDED DEGREE-BASED EQUIVALENCE ****************************************)
+
+definition cdeq: ∀h. sd h → relation3 lenv term term ≝
+ λh,o,L. tdeq h o.
+
+definition cdeq_ext: ∀h. sd h → relation3 lenv bind bind ≝
+ λh,o. cext2 (cdeq h o).
+
+interpretation
+ "degree-based equivalence (binders)"
+ 'LazyEq h o I1 I2 = (ext2 (tdeq h o) I1 I2).
<subsection name="A2">Stage "A2": "Extending the Applicability Condition"</subsection>
+ <news class="alpha" date="2017 October 16.">
+ Exclusion binder in local environments.
+ Ported components: syntax, relocation, s_transition, s_computation, static.
+ </news>
<news class="alpha" date="2017 April 16.">
Strong rt-normalization
for simply typed terms
}
]
*)
+(*
class "blue"
[ { "conversion" * } {
[ { "context-sensitive r-conversion" * } {
]
}
]
+*)
class "green"
[ { "static typing" * } {
[ { "generic reducibility" * } {
}
]
[ { "context-sensitive free variables" * } {
- [ "lsubf ( ⦃?,?⦄ ⫃𝐅* ⦃?,?⦄ )" "lsubf_frees" * ]
+ [ "lsubf ( ⦃?,?⦄ ⫃𝐅* ⦃?,?⦄ )" "lsubf_lsubr" + "lsubf_frees" + "lsubf_lsubf" * ]
[ "frees ( ? ⊢ 𝐅*⦃?⦄ ≡ ? )" "frees_drops" + "frees_fqup" + "frees_frees" * ]
}
]
class "grass"
[ { "s-computation" * } {
[ { "iterated structural successor for closures" * } {
- [ "fqus ( ⦃?,?,?⦄ ⊐* ⦃?,?,?⦄ )" "fqus_weight" + "fqus_drops" + "fqus_fqup" + "fqus_fqus" * ]
- [ "fqup ( ⦃?,?,?⦄ ⊐+ ⦃?,?,?⦄ )" "fqup_weight" + "fqup_drops" + "fqup_fqup" * ]
+ [ "fqus ( ⦃?,?,?⦄ ⊐*[?] ⦃?,?,?⦄ ) ( ⦃?,?,?⦄ ⊐* ⦃?,?,?⦄ )" "fqus_weight" + "fqus_drops" + "fqus_fqup" + "fqus_fqus" * ]
+ [ "fqup ( ⦃?,?,?⦄ ⊐+[?] ⦃?,?,?⦄ ) ( ⦃?,?,?⦄ ⊐+ ⦃?,?,?⦄ )" "fqup_weight" + "fqup_drops" + "fqup_fqup" * ]
}
]
}
class "yellow"
[ { "s-transition" * } {
[ { "structural successor for closures" * } {
- [ "fquq ( ⦃?,?,?⦄ ⊐⸮ ⦃?,?,?⦄ )" "fquq_length" + "fquq_weight" * ]
- [ "fqu ( ⦃?,?,?⦄ ⊐ ⦃?,?,?⦄ )" "fqu_length" + "fqu_weight" * ]
+ [ "fquq ( ⦃?,?,?⦄ ⊐⸮[?] ⦃?,?,?⦄ ) ( ⦃?,?,?⦄ ⊐⸮ ⦃?,?,?⦄ )" "fquq_length" + "fquq_weight" * ]
+ [ "fqu ( ⦃?,?,?⦄ ⊐[?] ⦃?,?,?⦄ ) ( ⦃?,?,?⦄ ⊐ ⦃?,?,?⦄ )" "fqu_length" + "fqu_weight" * ]
}
]
}
[ { "relocation" * } {
[ { "generic slicing for local environments" * } {
[ "drops_vector ( ⬇*[?,?] ? ≡ ? ) ( ⬇*[?] ? ≡ ? )" * ]
- [ "drops ( ⬇*[?,?] ? ≡ ? ) ( ⬇*[?] ? ≡ ? )" "drops_lstar" + "drops_weight" + "drops_length" + "drops_ceq" + "drops_lexs" + "drops_lreq" + "drops_drops" * ]
+ [ "drops ( ⬇*[?,?] ? ≡ ? ) ( ⬇*[?] ? ≡ ? )" "drops_lstar" + "drops_weight" + "drops_length" + "drops_ext2" + "drops_lexs" + "drops_lreq" + "drops_drops" * ]
}
]
[ { "generic relocation" * } {
}
]
[ { "degree-based equivalence for terms" * } {
- [ "deq ( ? ≡[?,?] ? ) " "deq_deq" * ]
+ [ "tdeq ( ? ≡[?,?] ? ) " "tdeq_ext" + "tdeq_tdeq" * ]
}
]
[ { "closures" * } {
}
]
[ { "binders for local environments" * } {
- [ "bind_ext2" * ]
+ [ "ext2" "ext2_ext2" * ]
[ "bind" "bind_weight" * ]
}
]
(* Basic properties *********************************************************)
-corec lemma sle_refl: ∀f1,f2. f1 ≗ f2 → f1 ⊆ f2.
-#f1 #f2 * -f1 -f2
-#f1 #f2 #g1 #g2 #H12 #H1 #H2
-[ @(sle_push … H1 H2) | @(sle_next … H1 H2) ] -H1 -H2 /2 width=1 by/
+axiom sle_eq_repl_back1: ∀f2. eq_repl_back … (λf1. f1 ⊆ f2).
+
+lemma sle_eq_repl_fwd1: ∀f2. eq_repl_fwd … (λf1. f1 ⊆ f2).
+#f2 @eq_repl_sym /2 width=3 by sle_eq_repl_back1/
+qed-.
+
+axiom sle_eq_repl_back2: ∀f1. eq_repl_back … (λf2. f1 ⊆ f2).
+
+lemma sle_eq_repl_fwd2: ∀f1. eq_repl_fwd … (λf2. f1 ⊆ f2).
+#f1 @eq_repl_sym /2 width=3 by sle_eq_repl_back2/
+qed-.
+
+corec lemma sle_refl: ∀f. f ⊆ f.
+#f cases (pn_split f) * #g #H
+[ @(sle_push … H H) | @(sle_next … H H) ] -H //
qed.
+lemma sle_refl_eq: ∀f1,f2. f1 ≗ f2 → f1 ⊆ f2.
+/2 width=3 by sle_eq_repl_back2/ qed.
+
(* Basic inversion lemmas ***************************************************)
lemma sle_inv_xp: ∀g1,g2. g1 ⊆ g2 → ∀f2. ↑f2 = g2 →
∀f2, f3. f2 ⋓ f3 ≡ f0 →
∀f. f1 ⋓ f2 ≡ f → f ⋓ f3 ≡ f4.
+lemma sor_trans1_sym: ∀f0,f1,f2,f3,f4,f.
+ f0⋓f4 ≡ f1 → f1⋓f2 ≡ f → f0⋓f2 ≡ f3 → f3⋓f4 ≡ f.
+/4 width=6 by sor_sym, sor_trans1/ qed-.
+
+corec theorem sor_trans2_idem: ∀f0,f1,f2. f0 ⋓ f1 ≡ f2 →
+ ∀f. f1 ⋓ f2 ≡ f → f1 ⋓ f0 ≡ f.
+#f0 #f1 #f2 * -f0 -f1 -f2
+#f0 #f1 #f2 #g0 #g1 #g2 #Hf2 #H0 #H1 #H2 #g #Hg
+[ cases (sor_inv_ppx … Hg … H1 H2)
+| cases (sor_inv_pnx … Hg … H1 H2)
+| cases (sor_inv_nnx … Hg … H1 H2)
+| cases (sor_inv_nnx … Hg … H1 H2)
+] -g2 #f #Hf #H
+/3 width=7 by sor_nn, sor_np, sor_pn, sor_pp/
+qed-.
+
corec theorem sor_distr_dx: ∀f1,f2,f. f1 ⋓ f2 ≡ f → ∀g1,g2,g. g1 ⋓ g2 ≡ g →
∀g0. g1 ⋓ g0 ≡ f1 → g2 ⋓ g0 ≡ f2 → g ⋓ g0 ≡ f.
#f1 #f2 #f cases (pn_split f) * #x #Hx #Hf #g1 #g2 #g #Hg #g0 #Hf1 #Hf2
]
]
qed-.
+