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/lreq.ma".
16 include "basic_2/static/frees.ma".
18 (* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
20 (* Properties with ranged equivalence for local environments ****************)
22 lemma frees_lreq_conf: ∀f,L1,T. L1 ⊢ 𝐅*⦃T⦄ ≡ f → ∀L2. L1 ≡[f] L2 → L2 ⊢ 𝐅*⦃T⦄ ≡ f.
23 #f #L1 #T #H elim H -f -L1 -T
24 [ #f #I #Hf #X #H lapply (lreq_inv_atom1 … H) -H
25 #H destruct /2 width=1 by frees_atom/
26 | #f #I #L1 #V1 #s #_ #IH #X #H elim (lreq_inv_push1 … H) -H
27 /3 width=1 by frees_sort/
28 | #f #I #L1 #V1 #_ #IH #X #H elim (lreq_inv_next1 … H) -H
29 /3 width=1 by frees_zero/
30 | #f #I #L1 #V1 #i #_ #IH #X #H elim (lreq_inv_push1 … H) -H
31 /3 width=1 by frees_lref/
32 | #f #I #L1 #V1 #l #_ #IH #X #H elim (lreq_inv_push1 … H) -H
33 /3 width=1 by frees_gref/
34 | /6 width=5 by frees_bind, lreq_inv_tl, sle_lreq_trans, sor_inv_sle_dx, sor_inv_sle_sn/
35 | /5 width=5 by frees_flat, sle_lreq_trans, sor_inv_sle_dx, sor_inv_sle_sn/
39 lemma lreq_frees_trans: ∀f,L1,T. L1 ⊢ 𝐅*⦃T⦄ ≡ f → ∀L2. L2 ≡[f] L1 → L2 ⊢ 𝐅*⦃T⦄ ≡ f.
40 /3 width=3 by frees_lreq_conf, lreq_sym/ qed-.