]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpg_drops.ma
- notational change for cpg and cpx
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_transition / cpg_drops.ma
index 4ccf66fdfa32ea49c65e90627b36a7a72d5b6578..779ea690b035a6c6b06e4298d1577f747f1be200 100644 (file)
@@ -17,12 +17,12 @@ include "basic_2/s_computation/fqup_weight.ma".
 include "basic_2/s_computation/fqup_drops.ma".
 include "basic_2/rt_transition/cpg.ma".
 
-(* CONTEXT-SENSITIVE GENERIC PARALLEL RT-TRANSITION FOR TERMS ***************)
+(* COUNTED CONTEXT-SENSITIVE PARALLEL RT-TRANSITION FOR TERMS ***************)
 
 (* Advanced properties ******************************************************)
 
-lemma cpg_delta_drops: â\88\80c,h,G,K,V,V2,i,L,T2. â¬\87*[i] L â\89¡ K.â\93\93V â\86\92 â¦\83G, Kâ¦\84 â\8a¢ V â\9e¡[c, h] V2 →
-                       â¬\86*[⫯i] V2 â\89¡ T2 â\86\92  â¦\83G, Lâ¦\84 â\8a¢ #i â\9e¡[â\86\93c, h] T2.
+lemma cpg_delta_drops: â\88\80c,h,G,K,V,V2,i,L,T2. â¬\87*[i] L â\89¡ K.â\93\93V â\86\92 â¦\83G, Kâ¦\84 â\8a¢ V â¬\88[c, h] V2 →
+                       â¬\86*[⫯i] V2 â\89¡ T2 â\86\92  â¦\83G, Lâ¦\84 â\8a¢ #i â¬\88[c, h] T2.
 #c #h #G #K #V #V2 #i elim i -i
 [ #L #T2 #HLK lapply (drops_fwd_isid … HLK ?) // #H destruct /3 width=3 by cpg_delta/
 | #i #IH #L0 #T0 #H0 #HV2 #HVT2
@@ -31,8 +31,8 @@ lemma cpg_delta_drops: ∀c,h,G,K,V,V2,i,L,T2. ⬇*[i] L ≡ K.ⓓV → ⦃G, K
 ]
 qed.
 
-lemma cpg_ell_drops: â\88\80c,h,G,K,V,V2,i,L,T2. â¬\87*[i] L â\89¡ K.â\93\9bV â\86\92 â¦\83G, Kâ¦\84 â\8a¢ V â\9e¡[c, h] V2 →
-                     â¬\86*[⫯i] V2 â\89¡ T2 â\86\92  â¦\83G, Lâ¦\84 â\8a¢ #i â\9e¡[(↓c)+𝟘𝟙, h] T2.
+lemma cpg_ell_drops: â\88\80c,h,G,K,V,V2,i,L,T2. â¬\87*[i] L â\89¡ K.â\93\9bV â\86\92 â¦\83G, Kâ¦\84 â\8a¢ V â¬\88[c, h] V2 →
+                     â¬\86*[⫯i] V2 â\89¡ T2 â\86\92  â¦\83G, Lâ¦\84 â\8a¢ #i â¬\88[(↓c)+𝟘𝟙, h] T2.
 #c #h #G #K #V #V2 #i elim i -i
 [ #L #T2 #HLK lapply (drops_fwd_isid … HLK ?) // #H destruct /3 width=3 by cpg_ell/
 | #i #IH #L0 #T0 #H0 #HV2 #HVT2
@@ -43,11 +43,11 @@ qed.
 
 (* Advanced inversion lemmas ************************************************)
 
-lemma cpg_inv_lref1_drops: â\88\80c,h,G,i,L,T2. â¦\83G, Lâ¦\84 â\8a¢ #i â\9e¡[c, h] T2 →
+lemma cpg_inv_lref1_drops: â\88\80c,h,G,i,L,T2. â¦\83G, Lâ¦\84 â\8a¢ #i â¬\88[c, h] T2 →
                            ∨∨ T2 = #i ∧ c = 𝟘𝟘
-                            | â\88\83â\88\83cV,K,V,V2. â¬\87*[i] L â\89¡ K.â\93\93V & â¦\83G, Kâ¦\84 â\8a¢ V â\9e¡[cV, h] V2 &
-                                           ⬆*[⫯i] V2 ≡ T2 & c = cV
-                            | â\88\83â\88\83cV,K,V,V2. â¬\87*[i] L â\89¡ K.â\93\9bV & â¦\83G, Kâ¦\84 â\8a¢ V â\9e¡[cV, h] V2 &
+                            | â\88\83â\88\83cV,K,V,V2. â¬\87*[i] L â\89¡ K.â\93\93V & â¦\83G, Kâ¦\84 â\8a¢ V â¬\88[cV, h] V2 &
+                                           ⬆*[⫯i] V2 ≡ T2 & c = cV
+                            | â\88\83â\88\83cV,K,V,V2. â¬\87*[i] L â\89¡ K.â\93\9bV & â¦\83G, Kâ¦\84 â\8a¢ V â¬\88[cV, h] V2 &
                                            ⬆*[⫯i] V2 ≡ T2 & c = (↓cV) + 𝟘𝟙.
 #c #h #G #i elim i -i
 [ #L #T2 #H elim (cpg_inv_zero1 … H) -H * /3 width=1 by or3_intro0, conj/
@@ -61,6 +61,23 @@ lemma cpg_inv_lref1_drops: ∀c,h,G,i,L,T2. ⦃G, L⦄ ⊢ #i ➡[c, h] T2 →
 ]
 qed-.
 
+lemma cpg_inv_atom1_drops: ∀c,h,I,G,L,T2. ⦃G, L⦄ ⊢ ⓪{I} ⬈[c, h] T2 →
+                           ∨∨ T2 = ⓪{I} ∧ c = 𝟘𝟘
+                            | ∃∃s. T2 = ⋆(next h s) & I = Sort s & c = 𝟘𝟙
+                            | ∃∃cV,i,K,V,V2. ⬇*[i] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V ⬈[cV, h] V2 &
+                                             ⬆*[⫯i] V2 ≡ T2 & I = LRef i & c = cV
+                            | ∃∃cV,i,K,V,V2. ⬇*[i] L ≡ K.ⓛV & ⦃G, K⦄ ⊢ V ⬈[cV, h] V2 &
+                                             ⬆*[⫯i] V2 ≡ T2 & I = LRef i & c = (↓cV) + 𝟘𝟙.
+#c #h * #n #G #L #T2 #H
+[ elim (cpg_inv_sort1 … H) -H *
+  /3 width=3 by or4_intro0, or4_intro1, ex3_intro, conj/
+| elim (cpg_inv_lref1_drops … H) -H *
+  /3 width=10 by or4_intro0, or4_intro2, or4_intro3, ex5_5_intro, conj/
+| elim (cpg_inv_gref1 … H) -H
+  /3 width=1 by or4_intro0, conj/
+]
+qed-.
+
 (* Properties with generic slicing for local environments *******************)
 
 lemma cpg_lifts: ∀c,h,G. d_liftable2 (cpg h c G).
@@ -135,7 +152,7 @@ qed-.
 
 (* Inversion lemmas with generic slicing for local environments *************)
 
-lemma cpg_inv_lift1: ∀c,h,G. d_deliftable2_sn (cpg h c G).
+lemma cpg_inv_lifts1: ∀c,h,G. d_deliftable2_sn (cpg h c G).
 #c #h #G #L #U generalize in match c; -c
 @(fqup_wf_ind_eq … G L U) -G -L -U #G0 #L0 #U0 #IH #G #L * *
 [ #s #HG #HL #HU #c #X2 #H2 #b #f #K #HLK #X1 #H1 destruct -IH