]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_drops.ma
update in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / csx_drops.ma
index ad0c072fb097cb60e0c0b52991ebd1f609474381..c9b3c1ebf2350970cd27648cde211dda6969f445 100644 (file)
@@ -12,7 +12,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
-include "static_2/relocation/lifts_tdeq.ma".
+include "static_2/relocation/lifts_teqx.ma".
 include "basic_2/rt_transition/cpx_drops.ma".
 include "basic_2/rt_computation/csx.ma".
 
@@ -22,22 +22,24 @@ include "basic_2/rt_computation/csx.ma".
 
 (* Basic_1: was just: sn3_lift *)
 (* Basic_2A1: was just: csx_lift *)
-lemma csx_lifts: ∀h,G. d_liftable1 … (csx h G).
+lemma csx_lifts (h) (G):
+      d_liftable1 … (csx h G).
 #h #G #K #T #H @(csx_ind … H) -T
 #T1 #_ #IH #b #f #L #HLK #U1 #HTU1
 @csx_intro #U2 #HU12 #HnU12
 elim (cpx_inv_lifts_sn … HU12 … HLK … HTU1) -HU12
-/4 width=7 by tdeq_lifts_bi/
+/4 width=7 by teqx_lifts_bi/
 qed-.
 
 (* Inversion lemmas with generic slicing ************************************)
 
 (* Basic_1: was just: sn3_gen_lift *)
 (* Basic_2A1: was just: csx_inv_lift *)
-lemma csx_inv_lifts: ∀h,G. d_deliftable1 … (csx h G).
+lemma csx_inv_lifts (h) (G):
+      d_deliftable1 … (csx h G).
 #h #G #L #U #H @(csx_ind … H) -U
 #U1 #_ #IH #b #f #K #HLK #T1 #HTU1
 @csx_intro #T2 #HT12 #HnT12
 elim (cpx_lifts_sn … HT12 … HLK … HTU1) -HT12
-/4 width=7 by tdeq_inv_lifts_bi/
+/4 width=7 by teqx_inv_lifts_bi/
 qed-.