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 "static_2/syntax/append.ma".
16 include "static_2/static/frees.ma".
18 (* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
20 (* Properties with append for local environments ****************************)
22 lemma frees_append_void: ∀f,K,T. K ⊢ 𝐅+❨T❩ ≘ f → ⓧ.K ⊢ 𝐅+❨T❩ ≘ f.
23 #f #K #T #H elim H -f -K -T
24 [ /2 width=1 by frees_sort/
25 | #f * /3 width=1 by frees_atom, frees_unit, frees_lref/
26 | /2 width=1 by frees_pair/
27 | /2 width=1 by frees_unit/
28 | /2 width=1 by frees_lref/
29 | /2 width=1 by frees_gref/
30 | /3 width=5 by frees_bind/
31 | /3 width=5 by frees_flat/
35 (* Inversion lemmas with append for local environments **********************)
37 fact frees_inv_append_void_aux:
38 ∀f,L,T. L ⊢ 𝐅+❨T❩ ≘ f →
39 ∀K. L = ⓧ.K → K ⊢ 𝐅+❨T❩ ≘ f.
40 #f #L #T #H elim H -f -L -T
41 [ /2 width=1 by frees_sort/
43 elim (append_inv_atom3_sn … H) -H #H1 #H2 destruct
44 | #f #I #L #V #_ #IH #K #H
45 elim (append_inv_bind3_sn … H) -H * [ | #Y ] #H1 #H2 destruct
46 /3 width=1 by frees_pair/
48 elim (append_inv_bind3_sn … H) -H * [ | #Y ] #H1 #H2 destruct
49 /2 width=1 by frees_atom, frees_unit/
50 | #f #I #L #i #Hf #IH #K #H
51 elim (append_inv_bind3_sn … H) -H * [ | #Y ] #H1 #H2 destruct
52 /3 width=1 by frees_lref, frees_lref_push/
53 | /2 width=1 by frees_gref/
54 | /3 width=5 by frees_bind/
55 | /3 width=5 by frees_flat/
59 lemma frees_inv_append_void: ∀f,K,T. ⓧ.K ⊢ 𝐅+❨T❩ ≘ f → K ⊢ 𝐅+❨T❩ ≘ f.
60 /2 width=3 by frees_inv_append_void_aux/ qed-.