]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/substitution/frees_lift.ma
- advances on hereditarily free variables: now "frees" is primitive
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / substitution / frees_lift.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "basic_2/relocation/ldrop_ldrop.ma".
16 include "basic_2/substitution/frees.ma".
17
18 (* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
19
20 (* Advanced properties ******************************************************)
21
22 lemma frees_dec: ∀L,U,d,i. Decidable (frees d L U i).
23 #L #U @(f2_ind … rfw … L U) -L -U
24 #n #IH #L * *
25 [ -IH /3 width=5 by frees_inv_sort, or_intror/
26 | #j #Hn #d #i elim (lt_or_eq_or_gt i j) #Hji
27   [ -n @or_intror #H elim (lt_refl_false i)
28     lapply (frees_inv_lref_ge … H ?) -L -d /2 width=1 by lt_to_le/
29   | -n /2 width=1 by or_introl/
30   | elim (ylt_split j d) #Hdi
31     [ -n @or_intror #H elim (lt_refl_false i)
32       lapply (frees_inv_lref_skip … H ?) -L //
33     | elim (lt_or_ge j (|L|)) #Hj
34       [ elim (ldrop_O1_lt (Ⓕ) L j) // -Hj #I #K #W #HLK destruct
35         elim (IH K W … 0 (i-j-1)) -IH [1,3: /3 width=5 by frees_lref_be, ldrop_fwd_rfw, or_introl/ ] #HnW
36         @or_intror #H elim (frees_inv_lref_lt … H) // #Z #Y #X #_ #HLY -d
37         lapply (ldrop_mono … HLY … HLK) -L #H destruct /2 width=1 by/  
38       | -n @or_intror #H elim (lt_refl_false i)
39         lapply (frees_inv_lref_free … H ?) -d //
40       ]
41     ]
42   ]
43 | -IH /3 width=5 by frees_inv_gref, or_intror/
44 | #a #I #W #U #Hn #d #i destruct
45   elim (IH L W … d i) [1,3: /3 width=1 by frees_bind_sn, or_introl/ ] #HnW
46   elim (IH (L.ⓑ{I}W) U … (⫯d) (i+1)) -IH [1,3: /3 width=1 by frees_bind_dx, or_introl/ ] #HnU
47   @or_intror #H elim (frees_inv_bind … H) -H /2 width=1 by/
48 | #I #W #U #Hn #d #i destruct
49   elim (IH L W … d i) [1,3: /3 width=1 by frees_flat_sn, or_introl/ ] #HnW
50   elim (IH L U … d i) -IH [1,3: /3 width=1 by frees_flat_dx, or_introl/ ] #HnU
51   @or_intror #H elim (frees_inv_flat … H) -H /2 width=1 by/
52 ]
53 qed-.
54
55 lemma frees_S: ∀L,U,d,i. L ⊢ i ϵ 𝐅*[yinj d]⦃U⦄ → ∀I,K,W. ⇩[d] L ≡ K.ⓑ{I}W →
56                (K ⊢ i-d-1 ϵ 𝐅*[0]⦃W⦄ → ⊥) → L ⊢ i ϵ 𝐅*[⫯d]⦃U⦄.
57 #L #U #d #i #H elim (frees_inv … H) -H /3 width=2 by frees_eq/
58 * #I #K #W #j #Hdj #Hji #HnU #HLK #HW #I0 #K0 #W0 #HLK0 #HnW0
59 lapply (yle_inv_inj … Hdj) -Hdj #Hdj
60 elim (le_to_or_lt_eq … Hdj) -Hdj
61 [ -I0 -K0 -W0 /3 width=9 by frees_be, yle_inj/
62 | -Hji -HnU #H destruct
63   lapply (ldrop_mono … HLK0 … HLK) #H destruct -I
64   elim HnW0 -L -U -HnW0 //
65 ]
66 qed.
67
68 (* Note: lemma 1250 *)
69 lemma frees_bind_dx_O: ∀a,I,L,W,U,i. L.ⓑ{I}W ⊢ i+1 ϵ 𝐅*[0]⦃U⦄ →
70                        L ⊢ i ϵ 𝐅*[0]⦃ⓑ{a,I}W.U⦄.
71 #a #I #L #W #U #i #HU elim (frees_dec L W 0 i)
72 /4 width=5 by frees_S, frees_bind_dx, frees_bind_sn/
73 qed.