]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/static_2/static/lsubc_drops.ma
update in ground_2, static_2, basic_2, apps_2, alpha_1
[helm.git] / matita / matita / contribs / lambdadelta / static_2 / static / lsubc_drops.ma
index 9fe4888bf5b70a26425a791dfbbd217fec035783..26d627a9056f38382f3edde31bf7e3a109c973df 100644 (file)
@@ -19,12 +19,12 @@ include "static_2/static/lsubc.ma".
 
 (* Properties with generic slicing ******************************************)
 
-(* Note: the premise ð\9d\90\94â¦\83fâ¦\84 cannot be removed *)
+(* Note: the premise ð\9d\90\94â\9dªfâ\9d« cannot be removed *)
 (* Basic_1: includes: csubc_drop_conf_O *)
 (* Basic_2A1: includes: lsubc_drop_O1_trans *)
 lemma lsubc_drops_trans_isuni: ∀RP,G,L1,L2. G ⊢ L1 ⫃[RP] L2 →
-                               â\88\80b,f,K2. ð\9d\90\94â¦\83fâ¦\84 â\86\92 â¬\87*[b, f] L2 ≘ K2 →
-                               â\88\83â\88\83K1. â¬\87*[b, f] L1 ≘ K1 & G ⊢ K1 ⫃[RP] K2.
+                               â\88\80b,f,K2. ð\9d\90\94â\9dªfâ\9d« â\86\92 â\87©*[b,f] L2 ≘ K2 →
+                               â\88\83â\88\83K1. â\87©*[b,f] L1 ≘ K1 & G ⊢ K1 ⫃[RP] K2.
 #RP #G #L1 #L2 #H elim H -L1 -L2
 [ /2 width=3 by ex2_intro/
 | #I #L1 #L2 #HL12 #IH #b #f #K2 #Hf #H
@@ -48,8 +48,8 @@ qed-.
 (* Basic_1: includes: csubc_drop_conf_rev *)
 (* Basic_2A1: includes: drop_lsubc_trans *)
 lemma drops_lsubc_trans: ∀RR,RS,RP. gcp RR RS RP →
-                         â\88\80b,f,G,L1,K1. â¬\87*[b, f] L1 ≘ K1 → ∀K2. G ⊢ K1 ⫃[RP] K2 →
-                         â\88\83â\88\83L2. G â\8a¢ L1 â«\83[RP] L2 & â¬\87*[b, f] L2 ≘ K2.
+                         â\88\80b,f,G,L1,K1. â\87©*[b,f] L1 ≘ K1 → ∀K2. G ⊢ K1 ⫃[RP] K2 →
+                         â\88\83â\88\83L2. G â\8a¢ L1 â«\83[RP] L2 & â\87©*[b,f] L2 ≘ K2.
 #RR #RS #RP #HR #b #f #G #L1 #K1 #H elim H -f -L1 -K1
 [ #f #Hf #Y #H lapply (lsubc_inv_atom1 … H) -H
   #H destruct /4 width=3 by lsubc_atom, drops_atom, ex2_intro/