]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma
partial commit in the relocation component to move some files ...
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / relocation / drops.ma
index a38d5c7a56d765a68f573e1b5066dc193d9424a7..27281ee17925880e47abbce5166ec0dcfed00331 100644 (file)
@@ -48,9 +48,9 @@ definition d_deliftable2_sn: predicate (lenv → relation term) ≝
                              ∃∃T2. ⬆*[f] T2 ≡ U2 & R K T1 T2.
 
 definition dropable_sn: predicate (rtmap → relation lenv) ≝
-                        λR. ∀L1,K1,c,f. ⬇*[c, f] L1 ≡ K1 → ∀L2,u2. R u2 L1 L2 →
-                        ∀u1. f ⊚ u1 ≡ u2 →
-                        ∃∃K2. R u1 K1 K2 & ⬇*[c, f] L2 ≡ K2.
+                        λR. ∀L1,K1,c,f. ⬇*[c, f] L1 ≡ K1 → ∀L2,f2. R f2 L1 L2 →
+                        ∀f1. f ⊚ f1 ≡ f2 →
+                        ∃∃K2. R f1 K1 K2 & ⬇*[c, f] L2 ≡ K2.
 
 (* Basic inversion lemmas ***************************************************)
 
@@ -117,18 +117,18 @@ lemma drops_inv_skip2: ∀I,X,K2,V2,c,f. ⬇*[c, ↑f] X ≡ K2.ⓑ{I}V2 →
 
 (* Basic properties *********************************************************)
 
-lemma drops_eq_repl_back: ∀L1,L2,c. eq_stream_repl_back … (λf. ⬇*[c, f] L1 ≡ L2).
+lemma drops_eq_repl_back: ∀L1,L2,c. eq_repl_back … (λf. ⬇*[c, f] L1 ≡ L2).
 #L1 #L2 #c #f1 #H elim H -L1 -L2 -f1
 [ /4 width=3 by drops_atom, isid_eq_repl_back/
-| #I #L1 #L2 #V #f1 #_ #IH #f2 #H elim (next_inv_sn … H) -H
-  /3 width=1 by drops_drop/
-| #I #L1 #L2 #V1 #v2 #f1 #_ #HV #IH #f2 #H elim (push_inv_sn … H) -H
+| #I #L1 #L2 #V #f1 #_ #IH #f2 #H elim (eq_inv_nx … H) -H
+  /3 width=3 by drops_drop/
+| #I #L1 #L2 #V1 #v2 #f1 #_ #HV #IH #f2 #H elim (eq_inv_px … H) -H
   /3 width=3 by drops_skip, lifts_eq_repl_back/
 ]
 qed-.
 
-lemma drops_eq_repl_fwd: ∀L1,L2,c. eq_stream_repl_fwd … (λf. ⬇*[c, f] L1 ≡ L2).
-#L1 #L2 #c @eq_stream_repl_sym /2 width=3 by drops_eq_repl_back/ (**) (* full auto fails *)
+lemma drops_eq_repl_fwd: ∀L1,L2,c. eq_repl_fwd … (λf. ⬇*[c, f] L1 ≡ L2).
+#L1 #L2 #c @eq_repl_sym /2 width=3 by drops_eq_repl_back/ (**) (* full auto fails *)
 qed-.
 
 (* Basic_2A1: includes: drop_refl *)
@@ -163,7 +163,7 @@ fact drops_fwd_drop2_aux: ∀X,Y,c,f2. ⬇*[c, f2] X ≡ Y → ∀I,K,V. Y = K.
 | #I #L1 #L2 #V #f2 #_ #IHL #J #K #W #H elim (IHL … H) -IHL
   /3 width=7 by after_next, ex3_2_intro, drops_drop/
 | #I #L1 #L2 #V1 #V2 #f2 #HL #_ #_ #J #K #W #H destruct
-  lapply (isid_after_dx 𝐈𝐝 f2 ?) /3 width=9 by after_push, ex3_2_intro, drops_drop/
+  lapply (isid_after_dx 𝐈𝐝 … f2) /3 width=9 by after_push, ex3_2_intro, drops_drop/
 ]
 qed-.
 
@@ -176,16 +176,16 @@ lemma drops_fwd_drop2: ∀I,X,K,V,c,f2. ⬇*[c, f2] X ≡ K.ⓑ{I}V →
 lemma drops_after_fwd_drop2: ∀I,X,K,V,c,f2. ⬇*[c, f2] X ≡ K.ⓑ{I}V →
                              ∀f1,f. 𝐈⦃f1⦄ → f2 ⊚ ⫯f1 ≡ f → ⬇*[c, f] X ≡ K.
 #I #X #K #V #c #f2 #H #f1 #f #Hf1 #Hf elim (drops_fwd_drop2 … H) -H
-#g1 #g #Hg1 #Hg #HK lapply (after_mono … Hg Hf ??) -Hg -Hf
-/3 width=3 by drops_eq_repl_back, isid_inv_eq_repl, next_eq_repl/
+#g1 #g #Hg1 #Hg #HK lapply (after_mono_eq … Hg … Hf ??) -Hg -Hf
+/3 width=5 by drops_eq_repl_back, isid_inv_eq_repl, eq_next/
 qed-.
 
 (* Basic_1: includes: drop_gen_refl *)
 (* Basic_2A1: includes: drop_inv_O2 *)
 lemma drops_fwd_isid: ∀L1,L2,c,f. ⬇*[c, f] L1 ≡ L2 → 𝐈⦃f⦄ → L1 = L2.
 #L1 #L2 #c #f #H elim H -L1 -L2 -f //
-[ #I #L1 #L2 #V #f #_ #_ #H elim (isid_inv_next … H)
-| /5 width=3 by isid_inv_push, lifts_fwd_isid, eq_f3, sym_eq/
+[ #I #L1 #L2 #V #f #_ #_ #H elim (isid_inv_next … H) //
+| /5 width=5 by isid_inv_push, lifts_fwd_isid, eq_f3, sym_eq/
 ]
 qed-.