]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/multiple/cpys_cpys.ma
notational change of lift, drop, and gget
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / multiple / cpys_cpys.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/cpy_cpy.ma".
16 include "basic_2/multiple/cpys_alt.ma".
17
18 (* CONTEXT-SENSITIVE EXTENDED MULTIPLE SUBSTITUTION FOR TERMS ***************)
19
20 (* Advanced inversion lemmas ************************************************)
21
22 lemma cpys_inv_SO2: ∀G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 ▶*[d, 1] T2 → ⦃G, L⦄ ⊢ T1 ▶[d, 1] T2.
23 #G #L #T1 #T2 #d #H @(cpys_ind … H) -T2 /2 width=3 by cpy_trans_ge/
24 qed-.
25
26 (* Advanced properties ******************************************************)
27
28 lemma cpys_strip_eq: ∀G,L,T0,T1,d1,e1. ⦃G, L⦄ ⊢ T0 ▶*[d1, e1] T1 →
29                      ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶[d2, e2] T2 →
30                      ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[d2, e2] T & ⦃G, L⦄ ⊢ T2 ▶*[d1, e1] T.
31 normalize /3 width=3 by cpy_conf_eq, TC_strip1/ qed-.
32
33 lemma cpys_strip_neq: ∀G,L1,T0,T1,d1,e1. ⦃G, L1⦄ ⊢ T0 ▶*[d1, e1] T1 →
34                       ∀L2,T2,d2,e2. ⦃G, L2⦄ ⊢ T0 ▶[d2, e2] T2 →
35                       (d1 + e1 ≤ d2 ∨ d2 + e2 ≤ d1) →
36                       ∃∃T. ⦃G, L2⦄ ⊢ T1 ▶[d2, e2] T & ⦃G, L1⦄ ⊢ T2 ▶*[d1, e1] T.
37 normalize /3 width=3 by cpy_conf_neq, TC_strip1/ qed-.
38
39 lemma cpys_strap1_down: ∀G,L,T1,T0,d1,e1. ⦃G, L⦄ ⊢ T1 ▶*[d1, e1] T0 →
40                         ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶[d2, e2] T2 → d2 + e2 ≤ d1 →
41                         ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[d2, e2] T & ⦃G, L⦄ ⊢ T ▶*[d1, e1] T2.
42 normalize /3 width=3 by cpy_trans_down, TC_strap1/ qed.
43
44 lemma cpys_strap2_down: ∀G,L,T1,T0,d1,e1. ⦃G, L⦄ ⊢ T1 ▶[d1, e1] T0 →
45                         ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶*[d2, e2] T2 → d2 + e2 ≤ d1 →
46                         ∃∃T. ⦃G, L⦄ ⊢ T1 ▶*[d2, e2] T & ⦃G, L⦄ ⊢ T ▶[d1, e1] T2.
47 normalize /3 width=3 by cpy_trans_down, TC_strap2/ qed-.
48
49 lemma cpys_split_up: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 →
50                      ∀i. d ≤ i → i ≤ d + e →
51                      ∃∃T. ⦃G, L⦄ ⊢ T1 ▶*[d, i - d] T & ⦃G, L⦄ ⊢ T ▶*[i, d + e - i] T2.
52 #G #L #T1 #T2 #d #e #H #i #Hdi #Hide @(cpys_ind … H) -T2
53 [ /2 width=3 by ex2_intro/
54 | #T #T2 #_ #HT12 * #T3 #HT13 #HT3
55   elim (cpy_split_up … HT12 … Hide) -HT12 -Hide #T0 #HT0 #HT02
56   elim (cpys_strap1_down … HT3 … HT0) -T /3 width=5 by cpys_strap1, ex2_intro/
57   >ymax_pre_sn_comm //
58 ]
59 qed-.
60
61 lemma cpys_inv_lift1_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 →
62                          ∀K,s,d,e. ⬇[s, d, e] L ≡ K → ∀T1. ⬆[d, e] T1 ≡ U1 →
63                          d ≤ dt → dt ≤ yinj d + e → yinj d + e ≤ dt + et →
64                          ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[d, dt + et - (yinj d + e)] T2 &
65                                ⬆[d, e] T2 ≡ U2.
66 #G #L #U1 #U2 #dt #et #HU12 #K #s #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet
67 elim (cpys_split_up … HU12 (d + e)) -HU12 // -Hdedet #U #HU1 #HU2
68 lapply (cpys_weak … HU1 d e ? ?) -HU1 // [ >ymax_pre_sn_comm // ] -Hddt -Hdtde #HU1
69 lapply (cpys_inv_lift1_eq … HU1 … HTU1) -HU1 #HU1 destruct
70 elim (cpys_inv_lift1_ge … HU2 … HLK … HTU1) -HU2 -HLK -HTU1 //
71 >yplus_minus_inj /2 width=3 by ex2_intro/
72 qed-.
73
74 (* Main properties **********************************************************)
75
76 theorem cpys_conf_eq: ∀G,L,T0,T1,d1,e1. ⦃G, L⦄ ⊢ T0 ▶*[d1, e1] T1 →
77                       ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶*[d2, e2] T2 →
78                       ∃∃T. ⦃G, L⦄ ⊢ T1 ▶*[d2, e2] T & ⦃G, L⦄ ⊢ T2 ▶*[d1, e1] T.
79 normalize /3 width=3 by cpy_conf_eq, TC_confluent2/ qed-.
80
81 theorem cpys_conf_neq: ∀G,L1,T0,T1,d1,e1. ⦃G, L1⦄ ⊢ T0 ▶*[d1, e1] T1 →
82                        ∀L2,T2,d2,e2. ⦃G, L2⦄ ⊢ T0 ▶*[d2, e2] T2 →
83                        (d1 + e1 ≤ d2 ∨ d2 + e2 ≤ d1) →
84                        ∃∃T. ⦃G, L2⦄ ⊢ T1 ▶*[d2, e2] T & ⦃G, L1⦄ ⊢ T2 ▶*[d1, e1] T.
85 normalize /3 width=3 by cpy_conf_neq, TC_confluent2/ qed-.
86
87 theorem cpys_trans_eq: ∀G,L,T1,T,T2,d,e.
88                        ⦃G, L⦄ ⊢ T1 ▶*[d, e] T → ⦃G, L⦄ ⊢ T ▶*[d, e] T2 →
89                        ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2.
90 normalize /2 width=3 by trans_TC/ qed-.
91
92 theorem cpys_trans_down: ∀G,L,T1,T0,d1,e1. ⦃G, L⦄ ⊢ T1 ▶*[d1, e1] T0 →
93                          ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶*[d2, e2] T2 → d2 + e2 ≤ d1 →
94                          ∃∃T. ⦃G, L⦄ ⊢ T1 ▶*[d2, e2] T & ⦃G, L⦄ ⊢ T ▶*[d1, e1] T2.
95 normalize /3 width=3 by cpy_trans_down, TC_transitive2/ qed-.
96
97 theorem cpys_antisym_eq: ∀G,L1,T1,T2,d,e. ⦃G, L1⦄ ⊢ T1 ▶*[d, e] T2 →
98                          ∀L2. ⦃G, L2⦄ ⊢ T2 ▶*[d, e] T1 → T1 = T2.
99 #G #L1 #T1 #T2 #d #e #H @(cpys_ind_alt … H) -G -L1 -T1 -T2 //
100 [ #I1 #G #L1 #K1 #V1 #V2 #W2 #i #d #e #Hdi #Hide #_ #_ #HVW2 #_ #L2 #HW2
101   elim (lt_or_ge (|L2|) (i+1)) #Hi [ -Hdi -Hide | ]
102   [ lapply (cpys_weak_full … HW2) -HW2 #HW2
103     lapply (cpys_weak … HW2 0 (i+1) ? ?) -HW2 //
104     [ >yplus_O1 >yplus_O1 /3 width=1 by ylt_fwd_le, ylt_inj/ ] -Hi
105     #HW2 >(cpys_inv_lift1_eq … HW2) -HW2 //
106   | elim (drop_O1_le (Ⓕ) … Hi) -Hi #K2 #HLK2
107     elim (cpys_inv_lift1_ge_up … HW2 … HLK2 … HVW2 ? ? ?) -HW2 -HLK2 -HVW2
108     /2 width=1 by ylt_fwd_le_succ, yle_succ_dx/ -Hdi -Hide
109     #X #_ #H elim (lift_inv_lref2_be … H) -H //
110   ]
111 | #a #I #G #L1 #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L2 #H elim (cpys_inv_bind1 … H) -H
112   #V #T #HV2 #HT2 #H destruct
113   lapply (IHV12 … HV2) #H destruct -IHV12 -HV2 /3 width=2 by eq_f2/
114 | #I #G #L1 #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L2 #H elim (cpys_inv_flat1 … H) -H
115   #V #T #HV2 #HT2 #H destruct /3 width=2 by eq_f2/
116 ]
117 qed-.