]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/dynamic/cnv_drops.ma
update in ground_2, static_2, basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / dynamic / cnv_drops.ma
index c979465c31d3ce4d730f2c36967a96cb7009e057..654a694b2cbc9193fdc53d5533be25e32030f30b 100644 (file)
 include "basic_2/rt_computation/cpms_drops.ma".
 include "basic_2/dynamic/cnv.ma".
 
-(* CONTEXT_SENSITIVE NATIVE VALIDITY FOR TERMS ******************************)
+(* CONTEXT-SENSITIVE NATIVE VALIDITY FOR TERMS ******************************)
 
 (* Advanced dproperties *****************************************************)
 
 (* Basic_2A1: uses: snv_lref *)
-lemma cnv_lref_drops (a) (h) (G): ∀I,K,V,i,L. ⦃G, K⦄ ⊢ V ![a, h] →
-                                  ⬇*[i] L ≘ K.ⓑ{I}V →  ⦃G, L⦄ ⊢ #i ![a, h].
-#a #h #G #I #K #V #i elim i -i
+lemma cnv_lref_drops (h) (a) (G):
+      ∀I,K,V,i,L. ❪G,K❫ ⊢ V ![h,a] →
+      ⇩[i] L ≘ K.ⓑ[I]V → ❪G,L❫ ⊢ #i ![h,a].
+#h #a #G #I #K #V #i elim i -i
 [ #L #HV #H
   lapply (drops_fwd_isid … H ?) -H // #H destruct
   /2 width=1 by cnv_zero/
@@ -35,10 +36,10 @@ qed.
 (* Advanced inversion lemmas ************************************************)
 
 (* Basic_2A1: uses: snv_inv_lref *)
-lemma cnv_inv_lref_drops (a) (h) (G):
-                         ∀i,L. ⦃G, L⦄ ⊢ #i ![a, h] →
-                         ∃∃I,K,V. ⬇*[i] L ≘ K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ![a, h].
-#a #h #G #i elim i -i
+lemma cnv_inv_lref_drops (h) (a) (G):
+      ∀i,L. ❪G,L❫ ⊢ #i ![h,a] →
+      ∃∃I,K,V. ⇩[i] L ≘ K.ⓑ[I]V & ❪G,K❫ ⊢ V ![h,a].
+#h #a #G #i elim i -i
 [ #L #H
   elim (cnv_inv_zero … H) -H #I #K #V #HV #H destruct
   /3 width=5 by drops_refl, ex2_3_intro/
@@ -49,11 +50,35 @@ lemma cnv_inv_lref_drops (a) (h) (G):
 ]
 qed-.
 
+lemma cnv_inv_lref_pair (h) (a) (G):
+      ∀i,L. ❪G,L❫ ⊢ #i ![h,a] →
+      ∀I,K,V. ⇩[i] L ≘ K.ⓑ[I]V → ❪G,K❫ ⊢ V ![h,a].
+#h #a #G #i #L #H #I #K #V #HLK
+elim (cnv_inv_lref_drops … H) -H #Z #Y #X #HLY #HX
+lapply (drops_mono … HLY … HLK) -L #H destruct //
+qed-.
+
+lemma cnv_inv_lref_atom (h) (a) (b) (G):
+      ∀i,L. ❪G,L❫ ⊢ #i ![h,a] → ⇩*[b,𝐔❨i❩] L ≘ ⋆ → ⊥.
+#h #a #b #G #i #L #H #Hi
+elim (cnv_inv_lref_drops … H) -H #Z #Y #X #HLY #_
+lapply (drops_gen b … HLY) -HLY #HLY
+lapply (drops_mono … HLY … Hi) -L #H destruct
+qed-.
+
+lemma cnv_inv_lref_unit (h) (a) (G):
+      ∀i,L. ❪G,L❫ ⊢ #i ![h,a] →
+      ∀I,K. ⇩[i] L ≘ K.ⓤ[I] → ⊥.
+#h #a #G #i #L #H #I #K #HLK
+elim (cnv_inv_lref_drops … H) -H #Z #Y #X #HLY #_
+lapply (drops_mono … HLY … HLK) -L #H destruct
+qed-.
+
 (* Properties with generic slicing for local environments *******************)
 
 (* Basic_2A1: uses: snv_lift *)
-lemma cnv_lifts (a) (h): ∀G. d_liftable1 (cnv a h G).
-#a #h #G #K #T
+lemma cnv_lifts (h) (a): ∀G. d_liftable1 (cnv h a G).
+#h #a #G #K #T
 @(fqup_wf_ind_eq (Ⓣ) … G K T) -G -K -T #G0 #K0 #T0 #IH #G #K * * [|||| * ]
 [ #s #HG #HK #HT #_ #b #f #L #_ #X #H2 destruct
   >(lifts_inv_sort1 … H2) -X -K -f //
@@ -91,8 +116,8 @@ qed-.
 (* Inversion lemmas with generic slicing for local environments *************)
 
 (* Basic_2A1: uses: snv_inv_lift *)
-lemma cnv_inv_lifts (a) (h): ∀G. d_deliftable1 (cnv a h G).
-#a #h #G #L #U
+lemma cnv_inv_lifts (h) (a): ∀G. d_deliftable1 (cnv h a G).
+#h #a #G #L #U
 @(fqup_wf_ind_eq (Ⓣ) … G L U) -G -L -U #G0 #L0 #U0 #IH #G #L * * [|||| * ]
 [ #s #HG #HL #HU #H1 #b #f #K #HLK #X #H2 destruct
   >(lifts_inv_sort2 … H2) -X -L -f //
@@ -100,7 +125,7 @@ lemma cnv_inv_lifts (a) (h): ∀G. d_deliftable1 (cnv a h G).
   elim (cnv_inv_lref_drops … H1) -H1 #I0 #L0 #W #HL0 #HW
   elim (lifts_inv_lref2 … H2) -H2 #i #Hf #H destruct
 (**) (* this should be a lemma *)
-  lapply (drops_split_div â\80¦ HLK (ð\9d\90\94â\9d´iâ\9dµ) ???) -HLK [4,8: * |*: // ] #Y0 #HK #HLY0
+  lapply (drops_split_div â\80¦ HLK (ð\9d\90\94â\9d¨iâ\9d©) ???) -HLK [4,8: * |*: // ] #Y0 #HK #HLY0
   lapply (drops_conf … HL0 … HLY0 ??) -HLY0 [3,6: |*: /2 width=6 by after_uni_dx/ ] #HLY0
   lapply (drops_tls_at … Hf … HLY0) -HLY0 #HLY0
   elim (drops_inv_skip1 … HLY0) -HLY0 #Z #K0 #HLK0 #HZ #H destruct