]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_conversion/cpce_drops.ma
backport of WIP on \lambda\delta to matita 0.99.3
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_conversion / cpce_drops.ma
index 6e8afb63c7f156abf0a495e4716680811697f5e0..985f9f8f4e7cc98d1aa0c857de43e775449a9f8c 100644 (file)
@@ -13,6 +13,7 @@
 (**************************************************************************)
 
 include "static_2/relocation/drops.ma".
+include "static_2/relocation/lifts_lifts.ma".
 include "basic_2/rt_conversion/cpce.ma".
 
 (* CONTEXT-SENSITIVE PARALLEL ETA-CONVERSION FOR TERMS **********************)
@@ -44,3 +45,46 @@ lemma cpce_zero_drops (h) (G):
   /5 width=8 by cpce_lref, drops_drop/
 ]
 qed.
+
+(* Inversion lemmas with uniform slicing for local environments *************)
+
+lemma cpce_inv_lref_sn_drops (h) (G) (i) (L):
+      ∀X2. ⦃G,L⦄ ⊢ #i ⬌η[h] X2 →
+      ∀I,K. ⇩*[i] L ≘ K.ⓘ{I} →
+      ∨∨ ∧∧ ∀n,p,W,V,U. I = BPair Abst W → ⦃G,K⦄ ⊢ W ➡*[n,h] ⓛ{p}V.U → ⊥ & #i = X2
+       | ∃∃n,p,W,V1,V2,W2,U. ⦃G,K⦄ ⊢ W ➡*[n,h] ⓛ{p}V1.U & ⦃G,K⦄ ⊢ V1 ⬌η[h] V2
+                           & ⇧*[↑i] V2 ≘ W2 & I = BPair Abst W & +ⓛW2.ⓐ#0.#(↑i) = X2.
+#h #G #i elim i -i
+[ #L #X2 #HX2 #I #K #HLK
+  lapply (drops_fwd_isid … HLK ?) -HLK [ // ] #H destruct
+  /2 width=1 by cpce_inv_zero_sn/
+| #i #IH #L0 #X0 #HX0 #J #K #H0
+  elim (drops_inv_succ … H0) -H0 #I #L #HLK #H destruct
+  elim (cpce_inv_lref_sn … HX0) -HX0 #X2 #HX2 #HX20
+  elim (IH … HX2 … HLK) -IH -I -L *
+  [ #HJ #H destruct
+    lapply (lifts_inv_lref1_uni … HX20) -HX20 #H destruct
+    /4 width=7 by or_introl, conj/
+  | #n #p #W #V1 #V2 #W2 #U #HWU #HV12 #HVW2 #H1 #H2 destruct
+    elim (lifts_inv_bind1 … HX20) -HX20 #X2 #X #HWX2 #HX #H destruct
+    elim (lifts_inv_flat1 … HX) -HX #X0 #X1 #H0 #H1 #H destruct
+    lapply (lifts_inv_push_zero_sn … H0) -H0 #H destruct
+    elim (lifts_inv_push_succ_sn … H1) -H1 #j #Hj #H destruct
+    lapply (lifts_inv_lref1_uni … Hj) -Hj #H destruct
+    /4 width=12 by lifts_trans_uni, ex5_7_intro, or_intror/
+  ]
+]
+qed-.
+
+lemma cpce_inv_zero_sn_drops (h) (G) (i) (L):
+      ∀X2. ⦃G,L⦄ ⊢ #i ⬌η[h] X2 →
+      ∀I,K. ⇩*[i] L ≘ K.ⓘ{I} →
+      (∀n,p,W,V,U. I = BPair Abst W → ⦃G,K⦄ ⊢ W ➡*[n,h] ⓛ{p}V.U → ⊥) →
+      #i = X2.
+#h #G #i #L #X2 #HX2 #I #K #HLK #HI
+elim (cpce_inv_lref_sn_drops … HX2 … HLK) -L *
+[ #_ #H //
+| #n #p #W #V1 #V2 #W2 #U #HWU #_ #_ #H destruct
+  elim (HI … HWU) -n -p -K -X2 -V1 -V2 -W2 -U -i //
+]
+qed-.