]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/etc_new/frees/frees_lift.etc
frees_drops completed!
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / etc_new / frees / frees_lift.etc
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 "ground_2/ynat/ynat_max.ma".
16 include "basic_2/substitution/drop_drop.ma".
17 include "basic_2/multiple/frees.ma".
18
19 (* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
20
21 (* Advanced properties ******************************************************)
22
23 lemma frees_dec: ∀L,U,l,i. Decidable (frees l L U i).
24 #L #U @(f2_ind … rfw … L U) -L -U
25 #x #IH #L * *
26 [ -IH /3 width=5 by frees_inv_sort, or_intror/
27 | #j #Hx #l #i elim (ylt_split_eq i j) #Hji
28   [ -x @or_intror #H elim (ylt_yle_false … Hji)
29     lapply (frees_inv_lref_ge … H ?) -L -l /2 width=1 by ylt_fwd_le/
30   | -x /2 width=1 by or_introl/
31   | elim (ylt_split j l) #Hli
32     [ -x @or_intror #H elim (ylt_yle_false … Hji)
33       lapply (frees_inv_lref_skip … H ?) -L //
34     | elim (lt_or_ge j (|L|)) #Hj
35       [ elim (drop_O1_lt (Ⓕ) L j) // -Hj #I #K #W #HLK destruct
36         elim (IH K W … 0 (i-j-1)) -IH [1,3: /3 width=5 by frees_lref_be, drop_fwd_rfw, or_introl/ ] #HnW
37         @or_intror #H elim (frees_inv_lref_lt … H) // #Z #Y #X #_ #HLY -l
38         lapply (drop_mono … HLY … HLK) -L #H destruct /2 width=1 by/  
39       | -x @or_intror #H elim (ylt_yle_false … Hji)
40         lapply (frees_inv_lref_free … H ?) -l //
41       ]
42     ]
43   ]
44 | -IH /3 width=5 by frees_inv_gref, or_intror/
45 | #a #I #W #U #Hx #l #i destruct
46   elim (IH L W … l i) [1,3: /3 width=1 by frees_bind_sn, or_introl/ ] #HnW
47   elim (IH (L.ⓑ{I}W) U … (⫯l) (i+1)) -IH [1,3: /3 width=1 by frees_bind_dx, or_introl/ ] #HnU
48   @or_intror #H elim (frees_inv_bind … H) -H /2 width=1 by/
49 | #I #W #U #Hx #l #i destruct
50   elim (IH L W … l i) [1,3: /3 width=1 by frees_flat_sn, or_introl/ ] #HnW
51   elim (IH L U … l i) -IH [1,3: /3 width=1 by frees_flat_dx, or_introl/ ] #HnU
52   @or_intror #H elim (frees_inv_flat … H) -H /2 width=1 by/
53 ]
54 qed-.
55
56 lemma frees_S: ∀L,U,l,i. L ⊢ i ϵ 𝐅*[yinj l]⦃U⦄ → ∀I,K,W. ⬇[l] L ≡ K.ⓑ{I}W →
57                (K ⊢ ⫰(i-l) ϵ 𝐅*[0]⦃W⦄ → ⊥) → L ⊢ i ϵ 𝐅*[⫯l]⦃U⦄.
58 #L #U #l #i #H elim (frees_inv … H) -H /3 width=2 by frees_eq/
59 * #I #K #W #j #Hlj #Hji #HnU #HLK #HW #I0 #K0 #W0 #HLK0 #HnW0
60 lapply (yle_inv_inj … Hlj) -Hlj #Hlj
61 elim (le_to_or_lt_eq … Hlj) -Hlj
62 [ -I0 -K0 -W0 /3 width=9 by frees_be, yle_inj/
63 | -Hji -HnU #H destruct
64   lapply (drop_mono … HLK0 … HLK) #H destruct -I
65   elim HnW0 -L -U -HnW0 //
66 ]
67 qed.
68
69 (* Note: lemma 1250 *)
70 lemma frees_bind_dx_O: ∀a,I,L,W,U,i. L.ⓑ{I}W ⊢ ⫯i ϵ 𝐅*[0]⦃U⦄ →
71                        L ⊢ i ϵ 𝐅*[0]⦃ⓑ{a,I}W.U⦄.
72 #a #I #L #W #U #i #HU elim (frees_dec L W 0 i)
73 /4 width=5 by frees_S, frees_bind_dx, frees_bind_sn/
74 qed.