]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda-delta/Basic-2/substitution/tps_split.ma
- some refactoring
[helm.git] / matita / matita / contribs / lambda-delta / Basic-2 / substitution / tps_split.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_lift.ma".
16
17 (* PARTIAL SUBSTITUTION ON TERMS ********************************************)
18
19 (* Split properties *********************************************************)
20
21 lemma tps_split_up: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫ T2 → ∀i. d ≤ i → i ≤ d + e →
22                     ∃∃T. L ⊢ T1 [d, i - d] ≫ T & L ⊢ T [i, d + e - i] ≫ T2.
23 #L #T1 #T2 #d #e #H elim H -L T1 T2 d e
24 [ /2/
25 | /2/
26 | #L #K #V #V1 #V2 #i #d #e #Hdi #Hide #HLK #HV1 #HV12 #IHV12 #j #Hdj #Hjde
27   elim (lt_or_ge i j) #Hij
28   [ -HV1 Hide;
29     lapply (drop_fwd_drop2 … HLK) #HLK'
30     elim (IHV12 (j - i - 1) ? ?) -IHV12; normalize /2/ -Hjde <minus_n_O >arith_b2 // #W1 #HVW1 #HWV1
31     generalize in match HVW1 generalize in match Hij -HVW1 (**) (* rewriting in the premises, rewrites in the goal too *)
32     >(plus_minus_m_m_comm … Hdj) in ⊢ (% → % → ?) -Hdj #Hij' #HVW1
33     elim (lift_total W1 0 (i + 1)) #W2 #HW12
34     lapply (tps_lift_ge … HWV1 … HLK' HW12 HV12 ?) -HWV1 HLK' HV12 // >arith_a2 /3 width=6/
35   | -IHV12 Hdi Hdj;
36     generalize in match HV1 generalize in match Hide -HV1 Hide (**) (* rewriting in the premises, rewrites in the goal too *)
37     >(plus_minus_m_m_comm … Hjde) in ⊢ (% → % → ?) -Hjde #Hide #HV1
38     @ex2_1_intro [2: @tps_lref |1: skip | /2 width=6/ ] (**) (* /3 width=6 is too slow *)
39   ]
40 | #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide
41   elim (IHV12 i ? ?) -IHV12 // #V #HV1 #HV2
42   elim (IHT12 (i + 1) ? ?) -IHT12 [2: /2 by arith4/ |3: /2/ ] (* just /2/ is too slow *)
43   -Hdi Hide >arith_c1 >arith_c1x #T #HT1 #HT2
44   @ex2_1_intro [2,3: @tps_bind | skip ] (**) (* explicit constructors *)
45   [3: @HV1 |4: @HT1 |5: // |1,2: skip | /3 width=5/ ]
46 | #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide
47   elim (IHV12 i ? ?) -IHV12 // elim (IHT12 i ? ?) -IHT12 //
48   -Hdi Hide /3 width=5/
49 ]
50 qed.
51
52 lemma tps_inv_lift1_up: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ≫ U2 →
53                         ∀K,d,e. ↓[d, e] L ≡ K → ∀T1. ↑[d, e] T1 ≡ U1 →
54                         d ≤ dt → dt ≤ d + e → d + e ≤ dt + et →
55                         ∃∃T2. K ⊢ T1 [d, dt + et - (d + e)] ≫ T2 & ↑[d, e] T2 ≡ U2.
56 #L #U1 #U2 #dt #et #HU12 #K #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet
57 elim (tps_split_up … HU12 (d + e) ? ?) -HU12 // -Hdedet #U #HU1 #HU2
58 lapply (tps_weak … HU1 d e ? ?) -HU1 // <plus_minus_m_m_comm // -Hddt Hdtde #HU1
59 lapply (tps_inv_lift1_eq … HU1 … HTU1) -HU1 #HU1 destruct -U1;
60 elim (tps_inv_lift1_ge … HU2 … HLK … HTU1 ?) -HU2 HLK HTU1 // <minus_plus_m_m /2/
61 qed.