]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/lambda/subterms/relocating_substitution.ma
decentralized notation in lambda
[helm.git] / matita / matita / lib / lambda / subterms / relocating_substitution.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 "lambda/subterms/relocation.ma".
16
17 include "lambda/notation/functions/dsubst_3.ma".
18
19 (* RELOCATING SUBSTITUTION **************************************************)
20
21 (* Policy: depth (level) metavariables: d, e (as for lift) *)
22 let rec sdsubst G d F on F ≝ match F with
23 [ SVRef b i   ⇒ tri … i d ({b}#i) (↑[i] G) ({b}#(i-1))
24 | SAbst b T   ⇒ {b}𝛌. (sdsubst G (d+1) T)
25 | SAppl b V T ⇒ {b}@ (sdsubst G d V). (sdsubst G d T)
26 ].
27
28 interpretation "relocating substitution for subterms"
29    'DSubst G d F = (sdsubst G d F).
30
31 lemma sdsubst_vref_lt: ∀b,i,d,G. i < d → [d ↙ G] {b}#i = {b}#i.
32 normalize /2 width=1/
33 qed.
34
35 lemma sdsubst_vref_eq: ∀b,i,G. [i ↙ G] {b}#i = ↑[i]G.
36 normalize //
37 qed.
38
39 lemma sdsubst_vref_gt: ∀b,i,d,G. d < i → [d ↙ G] {b}#i = {b}#(i-1).
40 normalize /2 width=1/
41 qed.
42
43 theorem sdsubst_slift_le: ∀h,G,F,d1,d2. d2 ≤ d1 →
44                           [d2 ↙ ↑[d1 - d2, h] G] ↑[d1 + 1, h] F = ↑[d1, h] [d2 ↙ G] F.
45 #h #G #F elim F -F
46 [ #b #i #d1 #d2 #Hd21 elim (lt_or_eq_or_gt i d2) #Hid2
47   [ lapply (lt_to_le_to_lt … Hid2 Hd21) -Hd21 #Hid1
48     >(sdsubst_vref_lt … Hid2) >(slift_vref_lt … Hid1) >slift_vref_lt /2 width=1/
49   | destruct >sdsubst_vref_eq >slift_vref_lt /2 width=1/
50   | >(sdsubst_vref_gt … Hid2) -Hd21 elim (lt_or_ge (i-1) d1) #Hi1d1
51     [ >(slift_vref_lt … Hi1d1) >slift_vref_lt /2 width=1/
52     | lapply (ltn_to_ltO … Hid2) #Hi
53       >(slift_vref_ge … Hi1d1) >slift_vref_ge /2 width=1/ -Hi1d1 >plus_minus // /3 width=1/
54     ]
55   ]
56 | normalize #b #T #IHT #d1 #d2 #Hd21
57   lapply (IHT (d1+1) (d2+1) ?) -IHT /2 width=1/
58 | normalize #b #V #T #IHV #IHT #d1 #d2 #Hd21
59   >IHV -IHV // >IHT -IHT //
60 ]
61 qed.
62
63 theorem sdsubst_slift_be: ∀h,G,F,d1,d2. d1 ≤ d2 → d2 ≤ d1 + h →
64                           [d2 ↙ G] ↑[d1, h + 1] F = ↑[d1, h] F.
65 #h #G #F elim F -F
66 [ #b #i #d1 #d2 #Hd12 #Hd21 elim (lt_or_ge i d1) #Hid1
67   [ lapply (lt_to_le_to_lt … Hid1 Hd12) -Hd12 -Hd21 #Hid2
68     >(slift_vref_lt … Hid1) >(slift_vref_lt … Hid1) /2 width=1/
69   | lapply (transitive_le … (i+h) Hd21 ?) -Hd12 -Hd21 /2 width=1/ #Hd2
70     >(slift_vref_ge … Hid1) >(slift_vref_ge … Hid1) -Hid1
71     >sdsubst_vref_gt // /2 width=1/
72   ]
73 | normalize #b #T #IHT #d1 #d2 #Hd12 #Hd21
74   >IHT -IHT // /2 width=1/
75 | normalize #b #V #T #IHV #IHT #d1 #d2 #Hd12 #Hd21
76   >IHV -IHV // >IHT -IHT //
77 ]
78 qed.
79
80 theorem sdsubst_slift_ge: ∀h,G,F,d1,d2. d1 + h ≤ d2 →
81                           [d2 ↙ G] ↑[d1, h] F = ↑[d1, h] [d2 - h ↙ G] F.
82 #h #G #F elim F -F
83 [ #b #i #d1 #d2 #Hd12 elim (lt_or_eq_or_gt i (d2-h)) #Hid2h
84   [ >(sdsubst_vref_lt … Hid2h) elim (lt_or_ge i d1) #Hid1
85     [ lapply (lt_to_le_to_lt … (d1+h) Hid1 ?) -Hid2h // #Hid1h
86       lapply (lt_to_le_to_lt … Hid1h Hd12) -Hid1h -Hd12 #Hid2
87       >(slift_vref_lt … Hid1) -Hid1 /2 width=1/
88     | >(slift_vref_ge … Hid1) -Hid1 -Hd12 /3 width=1/
89     ]
90   | destruct elim (le_inv_plus_l … Hd12) -Hd12 #Hd12 #Hhd2
91     >sdsubst_vref_eq >slift_vref_ge // >slift_slift_be // <plus_minus_m_m //
92   | elim (le_inv_plus_l … Hd12) -Hd12 #Hd12 #_
93     lapply (le_to_lt_to_lt … Hd12 Hid2h) -Hd12 #Hid1
94     lapply (ltn_to_ltO … Hid2h) #Hi
95     >(sdsubst_vref_gt … Hid2h)
96     >slift_vref_ge /2 width=1/ >slift_vref_ge /2 width=1/ -Hid1
97     >sdsubst_vref_gt /2 width=1/ -Hid2h >plus_minus //
98   ]
99 | normalize #b #T #IHT #d1 #d2 #Hd12
100   elim (le_inv_plus_l … Hd12) #_ #Hhd2
101   >IHT -IHT /2 width=1/ <plus_minus //
102 | normalize #b #V #T #IHV #IHT #d1 #d2 #Hd12
103   >IHV -IHV // >IHT -IHT //
104 ]
105 qed.
106
107 theorem sdsubst_sdsubst_ge: ∀G1,G2,F,d1,d2. d1 ≤ d2 →
108                             [d2 ↙ G2] [d1 ↙ G1] F = [d1 ↙ [d2 - d1 ↙ G2] G1] [d2 + 1 ↙ G2] F.
109 #G1 #G2 #F elim F -F
110 [ #b #i #d1 #d2 #Hd12 elim (lt_or_eq_or_gt i d1) #Hid1
111   [ lapply (lt_to_le_to_lt … Hid1 Hd12) -Hd12 #Hid2
112     >(sdsubst_vref_lt … Hid1) >(sdsubst_vref_lt … Hid2) >sdsubst_vref_lt /2 width=1/
113   | destruct >sdsubst_vref_eq >sdsubst_vref_lt /2 width=1/
114   | >(sdsubst_vref_gt … Hid1) elim (lt_or_eq_or_gt i (d2+1)) #Hid2
115     [ lapply (ltn_to_ltO … Hid1) #Hi
116       >(sdsubst_vref_lt … Hid2) >sdsubst_vref_lt /2 width=1/
117     | destruct /2 width=1/
118     | lapply (le_to_lt_to_lt (d1+1) … Hid2) -Hid1 /2 width=1/ -Hd12 #Hid1
119       >(sdsubst_vref_gt … Hid2) >sdsubst_vref_gt /2 width=1/
120       >sdsubst_vref_gt // /2 width=1/
121     ]
122   ]
123 | normalize #b #T #IHT #d1 #d2 #Hd12
124   lapply (IHT (d1+1) (d2+1) ?) -IHT /2 width=1/
125 | normalize #b #V #T #IHV #IHT #d1 #d2 #Hd12
126   >IHV -IHV // >IHT -IHT //
127 ]
128 qed.
129
130 theorem sdsubst_sdsubst_lt: ∀G1,G2,F,d1,d2. d2 < d1 →
131                             [d2 ↙ [d1 - d2 -1 ↙ G1] G2] [d1 ↙ G1] F = [d1 - 1 ↙ G1] [d2 ↙ G2] F.
132 #G1 #G2 #F #d1 #d2 #Hd21
133 lapply (ltn_to_ltO … Hd21) #Hd1
134 >sdsubst_sdsubst_ge in ⊢ (???%); /2 width=1/ <plus_minus_m_m //
135 qed.
136
137 definition sdsubstable_f_dx: ∀S:Type[0]. (S → ?) → predicate (relation subterms) ≝ λS,f,R.
138                              ∀G,F1,F2. R F1 F2 → ∀d. R ([d ↙ (f G)] F1) ([d ↙ (f G)] F2).
139
140 lemma lstar_sdsubstable_f_dx: ∀S1,f,S2,R. (∀a. sdsubstable_f_dx S1 f (R a)) →
141                               ∀l. sdsubstable_f_dx S1 f (lstar S2 … R l).
142 #S1 #f #S2 #R #HR #l #G #F1 #F2 #H
143 @(lstar_ind_l … l F1 H) -l -F1 // /3 width=3/
144 qed.
145 (*
146 definition sdsubstable_dx: predicate (relation subterms) ≝ λR.
147                            ∀G,F1,F2. R F1 F2 → ∀d. R ([d ↙ G] F1) ([d ↙ G] F2).
148
149 definition sdsubstable: predicate (relation subterms) ≝ λR.
150                         ∀G1,G2. R G1 G2 → ∀F1,F2. R F1 F2 → ∀d. R ([d ↙ G1] F1) ([d ↙ G2] F2).
151
152 lemma star_sdsubstable_dx: ∀R. sdsubstable_dx R → sdsubstable_dx (star … R).
153 #R #HR #G #F1 #F2 #H elim H -F2 // /3 width=3/
154 qed.
155
156 lemma lstar_sdsubstable_dx: ∀S,R. (∀a. sdsubstable_dx (R a)) →
157                             ∀l. sdsubstable_dx (lstar S … R l).
158 #S #R #HR #l #G #F1 #F2 #H
159 @(lstar_ind_l … l F1 H) -l -F1 // /3 width=3/
160 qed.
161
162 lemma star_sdsubstable: ∀R. reflexive ? R →
163                         sdsubstable R → sdsubstable (star … R).
164 #R #H1R #H2 #G1 #G2 #H elim H -G2 /3 width=1/ /3 width=5/
165 qed.
166 *)