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_2A/grammar/term_vector.ma".
16 include "basic_2A/substitution/lift.ma".
18 (* BASIC TERM VECTOR RELOCATION *********************************************)
20 inductive liftv (l,m:nat) : relation (list term) ≝
21 | liftv_nil : liftv l m (ⓔ) (ⓔ)
22 | liftv_cons: ∀T1s,T2s,T1,T2.
23 ⬆[l, m] T1 ≡ T2 → liftv l m T1s T2s →
24 liftv l m (T1 ⨮ T1s) (T2 ⨮ T2s)
27 interpretation "relocation (vector)" 'RLift l m T1s T2s = (liftv l m T1s T2s).
29 (* Basic inversion lemmas ***************************************************)
31 fact liftv_inv_nil1_aux: ∀T1s,T2s,l,m. ⬆[l, m] T1s ≡ T2s → T1s = ⓔ → T2s = ⓔ.
32 #T1s #T2s #l #m * -T1s -T2s //
33 #T1s #T2s #T1 #T2 #_ #_ #H destruct
36 lemma liftv_inv_nil1: ∀T2s,l,m. ⬆[l, m] ⓔ ≡ T2s → T2s = ⓔ.
37 /2 width=5 by liftv_inv_nil1_aux/ qed-.
39 fact liftv_inv_cons1_aux: ∀T1s,T2s,l,m. ⬆[l, m] T1s ≡ T2s →
40 ∀U1,U1s. T1s = U1 ⨮ U1s →
41 ∃∃U2,U2s. ⬆[l, m] U1 ≡ U2 & ⬆[l, m] U1s ≡ U2s &
43 #T1s #T2s #l #m * -T1s -T2s
44 [ #U1 #U1s #H destruct
45 | #T1s #T2s #T1 #T2 #HT12 #HT12s #U1 #U1s #H destruct /2 width=5 by ex3_2_intro/
49 lemma liftv_inv_cons1: ∀U1,U1s,T2s,l,m. ⬆[l, m] U1 ⨮ U1s ≡ T2s →
50 ∃∃U2,U2s. ⬆[l, m] U1 ≡ U2 & ⬆[l, m] U1s ≡ U2s &
52 /2 width=3 by liftv_inv_cons1_aux/ qed-.
54 (* Basic properties *********************************************************)
56 lemma liftv_total: ∀l,m. ∀T1s:list term. ∃T2s. ⬆[l, m] T1s ≡ T2s.
57 #l #m #T1s elim T1s -T1s
58 [ /2 width=2 by liftv_nil, ex_intro/
59 | #T1 #T1s * #T2s #HT12s
60 elim (lift_total T1 l m) /3 width=2 by liftv_cons, ex_intro/