]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_leq.ma
- some corrections and additions
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / relocation / ldrop_leq.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 "ground_2/ynat/ynat_plus.ma".
16 include "basic_2/grammar/leq_leq.ma".
17 include "basic_2/relocation/ldrop.ma".
18
19 (* BASIC SLICING FOR LOCAL ENVIRONMENTS *************************************)
20
21 definition dedropable_sn: predicate (relation lenv) ≝
22                           λR. ∀L1,K1,s,d,e. ⇩[s, d, e] L1 ≡ K1 → ∀K2. R K1 K2 →
23                           ∃∃L2. R L1 L2 & ⇩[s, d, e] L2 ≡ K2 & L1 ≃[d, e] L2.
24
25 (* Properties on equivalence ************************************************)
26
27 lemma leq_ldrop_trans_be: ∀L1,L2,d,e. L1 ≃[d, e] L2 →
28                           ∀I,K2,W,s,i. ⇩[s, 0, i] L2 ≡ K2.ⓑ{I}W →
29                           d ≤ i → i < d + e →
30                           ∃∃K1. K1 ≃[0, ⫰(d+e-i)] K2 & ⇩[s, 0, i] L1 ≡ K1.ⓑ{I}W.
31 #L1 #L2 #d #e #H elim H -L1 -L2 -d -e
32 [ #d #e #J #K2 #W #s #i #H
33   elim (ldrop_inv_atom1 … H) -H #H destruct
34 | #I1 #I2 #L1 #L2 #V1 #V2 #_ #_ #J #K2 #W #s #i #_ #_ #H
35   elim (ylt_yle_false … H) //
36 | #I #L1 #L2 #V #e #HL12 #IHL12 #J #K2 #W #s #i #H #_ >yplus_O1
37   elim (ldrop_inv_O1_pair1 … H) -H * #Hi #HLK1 [ -IHL12 | -HL12 ]
38   [ #_ destruct >ypred_succ
39     /2 width=3 by ldrop_pair, ex2_intro/
40   | lapply (ylt_inv_O1 i ?) /2 width=1 by ylt_inj/
41     #H <H -H #H lapply (ylt_inv_succ … H) -H
42     #Hie elim (IHL12 … HLK1) -IHL12 -HLK1 // -Hie
43     >yminus_succ <yminus_inj /3 width=3 by ldrop_drop_lt, ex2_intro/
44   ]
45 | #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #IHL12 #J #K2 #W #s #i #HLK2 #Hdi
46   elim (yle_inv_succ1 … Hdi) -Hdi
47   #Hdi #Hi <Hi >yplus_succ1 #H lapply (ylt_inv_succ … H) -H
48   #Hide lapply (ldrop_inv_drop1_lt … HLK2 ?) -HLK2 /2 width=1 by ylt_O/
49   #HLK1 elim (IHL12 … HLK1) -IHL12 -HLK1 <yminus_inj >yminus_SO2
50   /4 width=3 by ylt_O, ldrop_drop_lt, ex2_intro/
51 ]
52 qed-.
53
54 lemma leq_ldrop_conf_be: ∀L1,L2,d,e. L1 ≃[d, e] L2 →
55                          ∀I,K1,W,s,i. ⇩[s, 0, i] L1 ≡ K1.ⓑ{I}W →
56                          d ≤ i → i < d + e →
57                          ∃∃K2. K1 ≃[0, ⫰(d+e-i)] K2 & ⇩[s, 0, i] L2 ≡ K2.ⓑ{I}W.
58 #L1 #L2 #d #e #HL12 #I #K1 #W #s #i #HLK1 #Hdi #Hide
59 elim (leq_ldrop_trans_be … (leq_sym … HL12) … HLK1) // -L1 -Hdi -Hide
60 /3 width=3 by leq_sym, ex2_intro/
61 qed-.
62
63 lemma dedropable_sn_TC: ∀R. dedropable_sn R → dedropable_sn (TC … R).
64 #R #HR #L1 #K1 #s #d #e #HLK1 #K2 #H elim H -K2
65 [ #K2 #HK12 elim (HR … HLK1 … HK12) -HR -K1
66   /3 width=4 by inj, ex3_intro/
67 | #K #K2 #_ #HK2 * #L #H1L1 #HLK #H2L1 elim (HR … HLK … HK2) -HR -K
68   /3 width=6 by leq_trans, step, ex3_intro/
69 ]
70 qed-.
71
72 (* Inversion lemmas on equivalence ******************************************)
73
74 lemma ldrop_O_inj: ∀i,L1,L2,K. ⇩[i] L1 ≡ K → ⇩[i] L2 ≡ K → L1 ≃[i, ∞] L2.
75 #i @(nat_ind_plus … i) -i
76 [ #L1 #L2 #K #H <(ldrop_inv_O2 … H) -K #H <(ldrop_inv_O2 … H) -L1 //
77 | #i #IHi * [2: #L1 #I1 #V1 ] * [2,4: #L2 #I2 #V2 ] #K #HLK1 #HLK2 //
78   lapply (ldrop_fwd_length … HLK1)
79   <(ldrop_fwd_length … HLK2) [ /4 width=5 by ldrop_inv_drop1, leq_succ/ ]
80   normalize <plus_n_Sm #H destruct
81 ]
82 qed-.