1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "basic_2/relocation/drops.ma".
17 (* GENERIC SLICING FOR LOCAL ENVIRONMENTS ***********************************)
19 (* Properties with context-sensitive equivalence for terms ******************)
21 lemma ceq_lift: d_liftable2 ceq.
22 /2 width=3 by ex2_intro/ qed-.
24 lemma ceq_inv_lift: d_deliftable2_sn ceq.
25 /2 width=3 by ex2_intro/ qed-.
27 (* Note: d_deliftable2_sn cfull does not hold *)
28 lemma cfull_lift: d_liftable2 cfull.
29 #K #T1 #T2 #_ #b #f #L #_ #U1 #_ -K -T1 -b
30 elim (lifts_total T2 f) /2 width=3 by ex2_intro/