]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/static_2/s_computation/fqup_drops.ma
more additions and corrections for the article
[helm.git] / matita / matita / contribs / lambdadelta / static_2 / s_computation / fqup_drops.ma
index 37bd15ada9016789cd7ce90baff4b2e18bd8030a..c21d46c6026be363201aa91140a85b6e272829f8 100644 (file)
@@ -20,7 +20,7 @@ include "static_2/s_computation/fqup.ma".
 (* Properties with generic slicing for local environments *******************)
 
 lemma fqup_drops_succ: ∀b,G,K,T,i,L,U. ⬇*[↑i] L ≘ K → ⬆*[↑i] T ≘ U →
-                       â¦\83G,L,Uâ¦\84 â\8a\90+[b] ⦃G,K,T⦄.
+                       â¦\83G,L,Uâ¦\84 â¬\82+[b] ⦃G,K,T⦄.
 #b #G #K #T #i elim i -i
 [ #L #U #HLK #HTU elim (drops_inv_succ … HLK) -HLK
   #I #Y #HY #H destruct <(drops_fwd_isid … HY) -K //
@@ -33,7 +33,7 @@ lemma fqup_drops_succ: ∀b,G,K,T,i,L,U. ⬇*[↑i] L ≘ K → ⬆*[↑i] T ≘
 qed.
 
 lemma fqup_drops_strap1: ∀b,G1,G2,L1,K1,K2,T1,T2,U1,i. ⬇*[i] L1 ≘ K1 → ⬆*[i] T1 ≘ U1 →
-                         â¦\83G1,K1,T1â¦\84 â\8a\90[b] â¦\83G2,K2,T2â¦\84 â\86\92 â¦\83G1,L1,U1â¦\84 â\8a\90+[b] ⦃G2,K2,T2⦄.
+                         â¦\83G1,K1,T1â¦\84 â¬\82[b] â¦\83G2,K2,T2â¦\84 â\86\92 â¦\83G1,L1,U1â¦\84 â¬\82+[b] ⦃G2,K2,T2⦄.
 #b #G1 #G2 #L1 #K1 #K2 #T1 #T2 #U1 *
 [ #HLK1 #HTU1 #HT12
   >(drops_fwd_isid … HLK1) -L1 //
@@ -42,5 +42,5 @@ lemma fqup_drops_strap1: ∀b,G1,G2,L1,K1,K2,T1,T2,U1,i. ⬇*[i] L1 ≘ K1 → 
 ]
 qed-.
 
-lemma fqup_lref: â\88\80b,I,G,L,K,V,i. â¬\87*[i] L â\89\98 K.â\93\91{I}V â\86\92 â¦\83G,L,#iâ¦\84 â\8a\90+[b] ⦃G,K,V⦄.
+lemma fqup_lref: â\88\80b,I,G,L,K,V,i. â¬\87*[i] L â\89\98 K.â\93\91{I}V â\86\92 â¦\83G,L,#iâ¦\84 â¬\82+[b] ⦃G,K,V⦄.
 /2 width=6 by fqup_drops_strap1/ qed.