]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/static_2/relocation/sex_tc.ma
update in ground static_2 basic_2 apps_2
[helm.git] / matita / matita / contribs / lambdadelta / static_2 / relocation / sex_tc.ma
index 0a95ed3b0f5720832af36230e9b3f61c57061157..86db581d2d5b42f29bf52d2b54cb2ee0eed21ebf 100644 (file)
@@ -18,7 +18,7 @@ include "static_2/relocation/sex.ma".
 (* GENERIC ENTRYWISE EXTENSION OF CONTEXT-SENSITIVE REALTIONS FOR TERMS *****)
 
 definition s_rs_transitive_isid: relation (relation3 lenv bind bind) ≝ λRN,RP.
-                                 â\88\80f. ð\9d\90\88â\9dªfâ\9d« → s_rs_transitive … RP (λ_.sex RN RP f).
+                                 â\88\80f. ð\9d\90\88â\9d¨fâ\9d© → s_rs_transitive … RP (λ_.sex RN RP f).
 
 (* Properties with transitive closure ***************************************)
 
@@ -82,7 +82,7 @@ qed.
 
 (* Basic_2A1: uses: TC_lpx_sn_ind *)
 theorem sex_tc_step_dx: ∀RN,RP. s_rs_transitive_isid RN RP →
-                        â\88\80f,L1,L. L1 âª¤[RN,RP,f] L â\86\92 ð\9d\90\88â\9dªfâ\9d« →
+                        â\88\80f,L1,L. L1 âª¤[RN,RP,f] L â\86\92 ð\9d\90\88â\9d¨fâ\9d© →
                         ∀L2. L ⪤[RN,CTC … RP,f] L2 → L1⪤ [RN,CTC … RP,f] L2.
 #RN #RP #HRP #f #L1 #L #H elim H -f -L1 -L
 [ #f #_ #Y #H -HRP >(sex_inv_atom1 … H) -Y // ]
@@ -99,7 +99,7 @@ qed-.
 (* Advanced properties ******************************************************)
 
 lemma sex_tc_dx: ∀RN,RP. s_rs_transitive_isid RN RP →
-                 â\88\80f. ð\9d\90\88â\9dªfâ\9d« → ∀L1,L2. TC … (sex RN RP f) L1 L2 → L1 ⪤[RN,CTC … RP,f] L2.
+                 â\88\80f. ð\9d\90\88â\9d¨fâ\9d© → ∀L1,L2. TC … (sex RN RP f) L1 L2 → L1 ⪤[RN,CTC … RP,f] L2.
 #RN #RP #HRP #f #Hf #L1 #L2 #H @(TC_ind_dx ??????? H) -L1
 /3 width=3 by sex_tc_step_dx, sex_tc_inj_dx/
 qed.