]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/reduction/cnx.ma
some resuls on pointwise extensions (all of them are now in the
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / reduction / cnx.ma
index 61e6e6cc6d40f76f515069980c1dc9ffba443f11..983e2ac364eeaae2238a8091c351bb93922d7d50 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/notation/relations/normal_5.ma".
+include "basic_2/notation/relations/prednormal_5.ma".
 include "basic_2/reduction/cnr.ma".
 include "basic_2/reduction/cpx.ma".
 
-(* CONTEXT-SENSITIVE EXTENDED NORMAL TERMS **********************************)
+(* NORMAL TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION ********************)
 
 definition cnx: ∀h. sd h → relation3 genv lenv term ≝
                 λh,g,G,L. NF … (cpx h g G L) (eq …).
 
 interpretation
-   "context-sensitive extended normality (term)"
-   'Normal h g L T = (cnx h g L T).
+   "normality for context-sensitive extended reduction (term)"
+   'PRedNormal h g L T = (cnx h g L T).
 
 (* Basic inversion lemmas ***************************************************)
 
@@ -106,6 +106,16 @@ lemma cnx_sort_iter: ∀h,g,G,L,k,l. deg h g k l → ⦃G, L⦄ ⊢ ➡[h, g] 
 lapply (deg_iter … l Hkl) -Hkl <minus_n_n /2 width=6 by cnx_sort/
 qed.
 
+lemma cnx_lref_free: ∀h,g,G,L,i. |L| ≤ i → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃#i⦄.
+#h #g #G #L #i #Hi #X #H elim (cpx_inv_lref1 … H) -H // *
+#I #K #V1 #V2 #HLK lapply (ldrop_fwd_length_lt2 … HLK) -HLK
+#H elim (lt_refl_false i) /2 width=3 by lt_to_le_to_lt/
+qed.
+
+lemma cnx_lref_atom: ∀h,g,G,L,i. ⇩[i] L ≡ ⋆ → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃#i⦄.
+#h #g #G #L #i #HL @cnx_lref_free >(ldrop_fwd_length … HL) -HL //
+qed.
+
 lemma cnx_abst: ∀h,g,a,G,L,W,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃W⦄ → ⦃G, L.ⓛW⦄ ⊢ ➡[h, g] 𝐍⦃T⦄ →
                 ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃ⓛ{a}W.T⦄.
 #h #g #a #G #L #W #T #HW #HT #X #H