1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "Basic_2/unfold/ltpss_tpss.ma".
17 (* PARTIAL UNFOLD ON LOCAL ENVIRONMENTS *************************************)
19 (* Main properties **********************************************************)
21 theorem ltpss_trans_eq: ∀L1,L,L2,d,e.
22 L1 [d, e] ≫* L → L [d, e] ≫* L2 → L1 [d, e] ≫* L2.
25 lemma ltpss_tpss2: ∀L1,L2,I,V1,V2,e.
26 L1 [0, e] ≫* L2 → L2 ⊢ V1 [0, e] ≫* V2 →
27 L1. 𝕓{I} V1 [0, e + 1] ≫* L2. 𝕓{I} V2.
28 #L1 #L2 #I #V1 #V2 #e #HL12 #H @(tpss_ind … H) -V2
30 | #V #V2 #_ #HV2 #IHV @(ltpss_trans_eq … IHV) /2/
34 lemma ltpss_tpss2_lt: ∀L1,L2,I,V1,V2,e.
35 L1 [0, e - 1] ≫* L2 → L2 ⊢ V1 [0, e - 1] ≫* V2 →
36 0 < e → L1. 𝕓{I} V1 [0, e] ≫* L2. 𝕓{I} V2.
37 #L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #He
38 >(plus_minus_m_m e 1) /2/
41 lemma ltpss_tpss1: ∀L1,L2,I,V1,V2,d,e.
42 L1 [d, e] ≫* L2 → L2 ⊢ V1 [d, e] ≫* V2 →
43 L1. 𝕓{I} V1 [d + 1, e] ≫* L2. 𝕓{I} V2.
44 #L1 #L2 #I #V1 #V2 #d #e #HL12 #H @(tpss_ind … H) -V2
46 | #V #V2 #_ #HV2 #IHV @(ltpss_trans_eq … IHV) /2/
50 lemma ltpss_tpss1_lt: ∀L1,L2,I,V1,V2,d,e.
51 L1 [d - 1, e] ≫* L2 → L2 ⊢ V1 [d - 1, e] ≫* V2 →
52 0 < d → L1. 𝕓{I} V1 [d, e] ≫* L2. 𝕓{I} V2.
53 #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #Hd
54 >(plus_minus_m_m d 1) /2/