]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda_delta/basic_2/unfold/tpss.ma
We are decapitalizing the contributions' names ...
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / unfold / tpss.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/tps.ma".
16
17 (* PARTIAL UNFOLD ON TERMS **************************************************)
18
19 definition tpss: nat → nat → lenv → relation term ≝
20                  λd,e,L. TC … (tps d e L).
21
22 interpretation "partial unfold (term)"
23    'PSubstStar L T1 d e T2 = (tpss d e L T1 T2).
24
25 (* Basic eliminators ********************************************************)
26
27 lemma tpss_ind: ∀d,e,L,T1. ∀R:predicate term. R T1 →
28                 (∀T,T2. L ⊢ T1 [d, e] ▶* T → L ⊢ T [d, e] ▶ T2 → R T → R T2) →
29                 ∀T2. L ⊢ T1 [d, e] ▶* T2 → R T2.
30 #d #e #L #T1 #R #HT1 #IHT1 #T2 #HT12 @(TC_star_ind … HT1 IHT1 … HT12) //
31 qed-.
32
33 (* Basic properties *********************************************************)
34
35 lemma tpss_strap: ∀L,T1,T,T2,d,e. 
36                   L ⊢ T1 [d, e] ▶ T → L ⊢ T [d, e] ▶* T2 → L ⊢ T1 [d, e] ▶* T2. 
37 /2 width=3/ qed.
38
39 lemma tpss_lsubs_conf: ∀L1,T1,T2,d,e. L1 ⊢ T1 [d, e] ▶* T2 →
40                        ∀L2. L1 [d, e] ≼ L2 → L2 ⊢ T1 [d, e] ▶* T2.
41 /3 width=3/ qed.
42
43 lemma tpss_refl: ∀d,e,L,T. L ⊢ T [d, e] ▶* T.
44 /2 width=1/ qed.
45
46 lemma tpss_bind: ∀L,V1,V2,d,e. L ⊢ V1 [d, e] ▶* V2 →
47                  ∀I,T1,T2. L. ⓑ{I} V2 ⊢ T1 [d + 1, e] ▶* T2 →
48                  L ⊢ ⓑ{I} V1. T1 [d, e] ▶* ⓑ{I} V2. T2.
49 #L #V1 #V2 #d #e #HV12 elim HV12 -V2
50 [ #V2 #HV12 #I #T1 #T2 #HT12 elim HT12 -T2
51   [ /3 width=5/
52   | #T #T2 #_ #HT2 #IHT @step /2 width=5/ (**) (* /3 width=5/ is too slow *)
53   ]
54 | #V #V2 #_ #HV12 #IHV #I #T1 #T2 #HT12
55   lapply (tpss_lsubs_conf … HT12 (L. ⓑ{I} V) ?) -HT12 /2 width=1/ #HT12
56   lapply (IHV … HT12) -IHV -HT12 #HT12 @step /2 width=5/ (**) (* /3 width=5/ is too slow *)
57 ]
58 qed.
59
60 lemma tpss_flat: ∀L,I,V1,V2,T1,T2,d,e.
61                  L ⊢ V1 [d, e] ▶ * V2 → L ⊢ T1 [d, e] ▶* T2 →
62                  L ⊢ ⓕ{I} V1. T1 [d, e] ▶* ⓕ{I} V2. T2.
63 #L #I #V1 #V2 #T1 #T2 #d #e #HV12 elim HV12 -V2
64 [ #V2 #HV12 #HT12 elim HT12 -T2
65   [ /3 width=1/
66   | #T #T2 #_ #HT2 #IHT @step /2 width=5/ (**) (* /3 width=5/ is too slow *)
67   ]
68 | #V #V2 #_ #HV12 #IHV #HT12
69   lapply (IHV … HT12) -IHV -HT12 #HT12 @step /2 width=5/ (**) (* /3 width=5/ is too slow *)
70 ]
71 qed.
72
73 lemma tpss_weak: ∀L,T1,T2,d1,e1. L ⊢ T1 [d1, e1] ▶* T2 →
74                  ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 →
75                  L ⊢ T1 [d2, e2] ▶* T2.
76 #L #T1 #T2 #d1 #e1 #H #d1 #d2 #Hd21 #Hde12 @(tpss_ind … H) -T2
77 [ //
78 | #T #T2 #_ #HT12 #IHT
79   lapply (tps_weak … HT12 … Hd21 Hde12) -HT12 -Hd21 -Hde12 /2 width=3/
80 ]
81 qed.
82
83 lemma tpss_weak_top: ∀L,T1,T2,d,e.
84                      L ⊢ T1 [d, e] ▶* T2 → L ⊢ T1 [d, |L| - d] ▶* T2.
85 #L #T1 #T2 #d #e #H @(tpss_ind … H) -T2
86 [ //
87 | #T #T2 #_ #HT12 #IHT
88   lapply (tps_weak_top … HT12) -HT12 /2 width=3/
89 ]
90 qed.
91
92 lemma tpss_weak_all: ∀L,T1,T2,d,e.
93                      L ⊢ T1 [d, e] ▶* T2 → L ⊢ T1 [0, |L|] ▶* T2.
94 #L #T1 #T2 #d #e #HT12
95 lapply (tpss_weak … HT12 0 (d + e) ? ?) -HT12 // #HT12
96 lapply (tpss_weak_top … HT12) //
97 qed.
98
99 (* Basic inversion lemmas ***************************************************)
100
101 (* Note: this can be derived from tpss_inv_atom1 *)
102 lemma tpss_inv_sort1: ∀L,T2,k,d,e. L ⊢ ⋆k [d, e] ▶* T2 → T2 = ⋆k.
103 #L #T2 #k #d #e #H @(tpss_ind … H) -T2
104 [ //
105 | #T #T2 #_ #HT2 #IHT destruct
106   >(tps_inv_sort1 … HT2) -HT2 //
107 ]
108 qed-.
109
110 (* Note: this can be derived from tpss_inv_atom1 *)
111 lemma tpss_inv_gref1: ∀L,T2,p,d,e. L ⊢ §p [d, e] ▶* T2 → T2 = §p.
112 #L #T2 #p #d #e #H @(tpss_ind … H) -T2
113 [ //
114 | #T #T2 #_ #HT2 #IHT destruct
115   >(tps_inv_gref1 … HT2) -HT2 //
116 ]
117 qed-.
118
119 lemma tpss_inv_bind1: ∀d,e,L,I,V1,T1,U2. L ⊢ ⓑ{I} V1. T1 [d, e] ▶* U2 →
120                       ∃∃V2,T2. L ⊢ V1 [d, e] ▶* V2 & 
121                                L. ⓑ{I} V2 ⊢ T1 [d + 1, e] ▶* T2 &
122                                U2 =  ⓑ{I} V2. T2.
123 #d #e #L #I #V1 #T1 #U2 #H @(tpss_ind … H) -U2
124 [ /2 width=5/
125 | #U #U2 #_ #HU2 * #V #T #HV1 #HT1 #H destruct
126   elim (tps_inv_bind1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H
127   lapply (tpss_lsubs_conf … HT1 (L. ⓑ{I} V2) ?) -HT1 /2 width=1/ /3 width=5/
128 ]
129 qed-.
130
131 lemma tpss_inv_flat1: ∀d,e,L,I,V1,T1,U2. L ⊢ ⓕ{I} V1. T1 [d, e] ▶* U2 →
132                       ∃∃V2,T2. L ⊢ V1 [d, e] ▶* V2 & L ⊢ T1 [d, e] ▶* T2 &
133                                U2 =  ⓕ{I} V2. T2.
134 #d #e #L #I #V1 #T1 #U2 #H @(tpss_ind … H) -U2
135 [ /2 width=5/
136 | #U #U2 #_ #HU2 * #V #T #HV1 #HT1 #H destruct
137   elim (tps_inv_flat1 … HU2) -HU2 /3 width=5/
138 ]
139 qed-.
140
141 lemma tpss_inv_refl_O2: ∀L,T1,T2,d. L ⊢ T1 [d, 0] ▶* T2 → T1 = T2.
142 #L #T1 #T2 #d #H @(tpss_ind … H) -T2
143 [ //
144 | #T #T2 #_ #HT2 #IHT <(tps_inv_refl_O2 … HT2) -HT2 //
145 ]
146 qed-.