]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/relocation/drops_lreq.ma
partial commit in the relocation component to move some files ...
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / relocation / drops_lreq.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/drops.ma".
16 include "basic_2/relocation/lreq_lreq.ma".
17
18 (* GENERAL SLICING FOR LOCAL ENVIRONMENTS ***********************************)
19
20 definition dedropable_sn: predicate (rtmap → relation lenv) ≝
21                           λR. ∀L1,K1,c,f. ⬇*[c, f] L1 ≡ K1 → ∀K2,f1. R f1 K1 K2 →
22                           ∀f2. f ⊚ f1 ≡ f2 →
23                           ∃∃L2. R f2 L1 L2 & ⬇*[c, f] L2 ≡ K2 & L1 ≡[f] L2.
24
25 (* Properties on equivalence ************************************************)
26
27 lemma dedropable_sn_TC: ∀R. dedropable_sn R → dedropable_sn (LTC … R).
28 #R #HR #L1 #K1 #c #f #HLK1 #K2 #f1 #H elim H -K2
29 [ #K2 #HK12 #f2 #Hf elim (HR … HLK1 … HK12 … Hf) -HR -K1 -f1
30   /3 width=4 by inj, ex3_intro/
31 | #K #K2 #_ #HK2 #IH #f2 #Hf elim (IH … Hf) -IH -K1
32   #L #H1L1 #HLK #H2L1 elim (HR … HLK … HK2 … Hf) -HR -K -f1
33   /3 width=6 by lreq_trans, step, ex3_intro/
34 ]
35 qed-.
36 (*
37 lemma lreq_drop_trans_be: ∀L1,L2,l,k. L1 ⩬[l, k] L2 →
38                           ∀I,K2,W,c,i. ⬇[c, 0, i] L2 ≡ K2.ⓑ{I}W →
39                           l ≤ i → ∀k0. i + ⫯k0 = l + k →
40                           ∃∃K1. K1 ⩬[0, k0] K2 & ⬇[c, 0, i] L1 ≡ K1.ⓑ{I}W.
41 #L1 #L2 #l #k #H elim H -L1 -L2 -l -k
42 [ #l #k #J #K2 #W #c #i #H
43   elim (drop_inv_atom1 … H) -H #H destruct
44 | #I1 #I2 #L1 #L2 #V1 #V2 #_ #_ #J #K2 #W #c #i #_ #_ #k0
45   >yplus_succ2 #H elim (ysucc_inv_O_dx … H)
46 | #I #L1 #L2 #V #k #HL12 #IHL12 #J #K2 #W #c #i #H #_ >yplus_O1 #k0 #H0
47   elim (drop_inv_O1_pair1 … H) -H * #Hi #HLK1 [ -IHL12 | -HL12 ]
48   [ destruct
49     /2 width=3 by drop_pair, ex2_intro/
50   | lapply (ylt_inv_O1 … Hi) #H <H in H0; -H
51     >yplus_succ1 #H lapply (ysucc_inv_inj … H) -H <(yplus_O1 k)
52     #H0 elim (IHL12 … HLK1 … H0) -IHL12 -HLK1 -H0 //
53     /3 width=3 by drop_drop_lt, ex2_intro/
54   ]
55 | #I1 #I2 #L1 #L2 #V1 #V2 #l #k #_ #IHL12 #J #K2 #W #c #i #HLK2 #Hli #k0
56   elim (yle_inv_succ1 … Hli) -Hli
57   #Hli #Hi <Hi >yplus_succ1 >yplus_succ1 #H lapply (ysucc_inv_inj … H) -H
58   #H0 lapply (drop_inv_drop1_lt … HLK2 ?) -HLK2 /2 width=1 by ylt_O1/
59   #HLK1 elim (IHL12 … HLK1 … H0) -IHL12 -HLK1 -H0
60   /4 width=3 by ylt_O1, drop_drop_lt, ex2_intro/
61 ]
62 qed-.
63
64 lemma lreq_drop_conf_be: ∀L1,L2,l,k. L1 ⩬[l, k] L2 →
65                          ∀I,K1,W,c,i. ⬇[c, 0, i] L1 ≡ K1.ⓑ{I}W →
66                          l ≤ i → ∀k0. i + ⫯k0 = l + k →
67                          ∃∃K2. K1 ⩬[0, k0] K2 & ⬇[c, 0, i] L2 ≡ K2.ⓑ{I}W.
68 #L1 #L2 #l #k #HL12 #I #K1 #W #c #i #HLK1 #Hli #k0 #H0
69 elim (lreq_drop_trans_be … (lreq_sym … HL12) … HLK1 … H0) // -L1 -Hli -H0
70 /3 width=3 by lreq_sym, ex2_intro/
71 qed-.
72
73 lemma drop_O1_ex: ∀K2,i,L1. |L1| = |K2| + i →
74                   ∃∃L2. L1 ⩬[0, i] L2 & ⬇[i] L2 ≡ K2.
75 #K2 #i @(ynat_ind … i) -i
76 [ /3 width=3 by lreq_O2, ex2_intro/
77 | #i #IHi #Y >yplus_succ2 #Hi
78   elim (drop_O1_lt (Ⓕ) Y 0) [2: >Hi // ]
79   #I #L1 #V #H lapply (drop_inv_O2 … H) -H #H destruct
80   >length_pair in Hi; #H lapply (ysucc_inv_inj … H) -H
81   #HL1K2 elim (IHi L1) -IHi // -HL1K2
82   /3 width=5 by lreq_pair, drop_drop, ex2_intro/
83 | #L1 >yplus_Y2 #H elim (ylt_yle_false (|L1|) (∞)) //
84 ]
85 qed-.
86
87 (* Inversion lemmas on equivalence ******************************************)
88
89 lemma drop_O1_inj: ∀i,L1,L2,K. ⬇[i] L1 ≡ K → ⬇[i] L2 ≡ K → L1 ⩬[i, ∞] L2.
90 #i @(ynat_ind … i) -i
91 [ #L1 #L2 #K #H <(drop_inv_O2 … H) -K #H <(drop_inv_O2 … H) -L1 //
92 | #i #IHi * [2: #L1 #I1 #V1 ] * [2,4: #L2 #I2 #V2 ] #K #HLK1 #HLK2 //
93   lapply (drop_fwd_length … HLK1)
94   <(drop_fwd_length … HLK2) [ /4 width=5 by drop_inv_drop1, lreq_succ/ ]
95   #H [ elim (ysucc_inv_O_sn … H) | elim (ysucc_inv_O_dx … H) ]
96 | #L1 #L2 #K #H1 lapply (drop_fwd_Y2 … H1) -H1
97   #H elim (ylt_yle_false … H) //
98 ]
99 qed-.