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/lveq_length.ma".
16 include "static_2/static/fsle.ma".
18 (* FREE VARIABLES INCLUSION FOR RESTRICTED CLOSURES *************************)
20 (* Properties with length for local environments ****************************)
22 lemma fsle_sort_bi: ∀L1,L2,s1,s2. |L1| = |L2| → ❨L1,⋆s1❩ ⊆ ❨L2,⋆s2❩.
23 /3 width=8 by lveq_length_eq, frees_sort, pr_sle_refl, ex4_4_intro/ qed.
25 lemma fsle_gref_bi: ∀L1,L2,l1,l2. |L1| = |L2| → ❨L1,§l1❩ ⊆ ❨L2,§l2❩.
26 /3 width=8 by lveq_length_eq, frees_gref, pr_sle_refl, ex4_4_intro/ qed.
28 lemma fsle_pair_bi: ∀K1,K2. |K1| = |K2| → ∀V1,V2. ❨K1,V1❩ ⊆ ❨K2,V2❩ →
29 ∀I1,I2. ❨K1.ⓑ[I1]V1,#O❩ ⊆ ❨K2.ⓑ[I2]V2,#O❩.
31 * #n1 #n2 #f1 #f2 #Hf1 #Hf2 #HK12 #Hf12
33 elim (lveq_inj_length … HK12) // -HK #H1 #H2 destruct
34 /3 width=12 by frees_pair, lveq_bind, pr_sle_next, ex4_4_intro/
37 lemma fsle_unit_bi: ∀K1,K2. |K1| = |K2| →
38 ∀I1,I2. ❨K1.ⓤ[I1],#O❩ ⊆ ❨K2.ⓤ[I2],#O❩.
39 /3 width=8 by frees_unit, lveq_length_eq, pr_sle_refl, ex4_4_intro/