]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda_delta/basic_2/substitution/ldrop_sfr.ma
e9befe700f343d46404040719aaac433826718fb
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / substitution / ldrop_sfr.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/substitution/lsubs_sfr.ma".
16 include "basic_2/substitution/ldrop.ma".
17
18 (* DROPPING *****************************************************************)
19
20 (* Properties about local env. full refinement for substitution *************)
21
22 lemma sfr_ldrop: ∀L,d,e.
23                  (∀I,K,V,i. d ≤ i → i < d + e → ⇩[0, i] L ≡ K. ⓑ{I}V → I = Abbr) →
24                  ≼ [d, e] L.
25 #L elim L -L //
26 #L #I #V #IHL #d @(nat_ind_plus … d) -d
27 [ #e @(nat_ind_plus … e) -e //
28   #e #_ #HH
29   >(HH I L V 0 ? ? ?) // /5 width=6/
30 | /5 width=6/
31 ]
32 qed.
33
34 (* Inversion lemmas about local env. full refinement for substitution *******)
35
36 lemma sfr_inv_ldrop: ∀I,L,K,V,i. ⇩[0, i] L ≡ K. ⓑ{I}V → ∀d,e. ≼ [d, e] L →
37                      d ≤ i → i < d + e → I = Abbr.
38 #I #L elim L -L
39 [ #K #V #i #H
40   lapply (ldrop_inv_atom1 … H) -H #H destruct
41 | #L #J #W #IHL #K #V #i #H
42   elim (ldrop_inv_O1 … H) -H *
43   [ -IHL #H1 #H2 #d #e #HL #Hdi #Hide destruct
44     lapply (le_n_O_to_eq … Hdi) -Hdi #H destruct
45     lapply (HL … (L.ⓓW) ?) -HL /2 width=1/ #H
46     elim (lsubs_inv_abbr1 … H ?) -H // -Hide #K #_ #H destruct //
47   | #Hi #HLK #d @(nat_ind_plus … d) -d
48     [ #e #H #_ #Hide
49       elim (sfr_inv_bind … H ?) -H [2: /2 width=2/ ] #HL #H destruct
50       @(IHL … HLK … HL) -IHL -HLK -HL // /2 width=1/
51     | #d #_ #e #H #Hdi #Hide
52       lapply (sfr_inv_skip … H ?) -H // #HL
53       @(IHL … HLK … HL) -IHL -HLK -HL /2 width=1/
54     ]
55   ]
56 ]
57 qed-.