]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/computation/csn_tstc_vector.ma
- lambda_delta: programmed renaming to lambdadelta
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / csn_tstc_vector.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/computation/acp_cr.ma".
16 include "basic_2/computation/cprs_tstc_vector.ma".
17 include "basic_2/computation/csn_lfpr.ma".
18 include "basic_2/computation/csn_vector.ma".
19
20 (* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERM VECTORS **********************)
21
22 (* Advanced properties ******************************************************)
23
24 (* Basic_1: was only: sn3_appls_lref *)
25 lemma csn_applv_cnf: ∀L,T. 𝐒⦃T⦄ → L ⊢ 𝐍⦃T⦄ → 
26                      ∀Vs. L ⊢ ⬊* Vs → L ⊢ ⬊* ⒶVs.T.
27 #L #T #H1T #H2T #Vs elim Vs -Vs [ #_ @(csn_cnf … H2T) ] (**) (* /2 width=1/ does not work *)
28 #V #Vs #IHV #H
29 elim (csnv_inv_cons … H) -H #HV #HVs
30 @csn_appl_simple_tstc // -HV /2 width=1/ -IHV -HVs
31 #X #H #H0
32 lapply (cprs_fwd_cnf_vector … H) -H // -H1T -H2T #H
33 elim (H0 ?) -H0 //
34 qed.
35
36 (* Basic_1: was: sn3_appls_beta *)
37 lemma csn_applv_beta: ∀a,L,W. L ⊢ ⬊* W →
38                       ∀Vs,V,T. L ⊢ ⬊* ⒶVs.ⓓ{a}V.T →
39                       L ⊢ ⬊* ⒶVs. ⓐV.ⓛ{a}W. T.
40 #a #L #W #HW #Vs elim Vs -Vs /2 width=1/ -HW
41 #V0 #Vs #IHV #V #T #H1T
42 lapply (csn_fwd_pair_sn … H1T) #HV0
43 lapply (csn_fwd_flat_dx … H1T) #H2T
44 @csn_appl_simple_tstc // -HV0 /2 width=1/ -IHV -H2T
45 #X #H #H0
46 elim (cprs_fwd_beta_vector … H) -H #H
47 [ -H1T elim (H0 ?) -H0 //
48 | -H0 @(csn_cprs_trans … H1T) -H1T /2 width=1/
49 ]
50 qed.
51
52 lemma csn_applv_delta: ∀L,K,V1,i. ⇩[0, i] L ≡ K. ⓓV1 →
53                        ∀V2. ⇧[0, i + 1] V1 ≡ V2 →
54                        ∀Vs.L ⊢ ⬊* (ⒶVs. V2) → L ⊢ ⬊* (ⒶVs. #i).
55 #L #K #V1 #i #HLK #V2 #HV12 #Vs elim Vs -Vs
56 [ #H
57   lapply (ldrop_fwd_ldrop2 … HLK) #HLK0
58   lapply (csn_inv_lift … H … HLK0 HV12) -V2 -HLK0 /2 width=4/
59 | #V #Vs #IHV #H1T
60   lapply (csn_fwd_pair_sn … H1T) #HV
61   lapply (csn_fwd_flat_dx … H1T) #H2T
62   @csn_appl_simple_tstc // -HV /2 width=1/ -IHV -H2T
63   #X #H #H0
64   elim (cprs_fwd_delta_vector … HLK … HV12 … H) -HLK -HV12 -H #H
65   [ -H1T elim (H0 ?) -H0 //
66   | -H0 @(csn_cprs_trans … H1T) -H1T /2 width=1/
67   ]
68 ]
69 qed.
70
71 (* Basic_1: was: sn3_appls_abbr *) 
72 lemma csn_applv_theta: ∀a,L,V1s,V2s. ⇧[0, 1] V1s ≡ V2s →
73                        ∀V,T. L ⊢ ⬊* ⓓ{a}V. ⒶV2s. T → L ⊢ ⬊* V →
74                        L ⊢ ⬊* ⒶV1s. ⓓ{a}V. T.
75 #a #L #V1s #V2s * -V1s -V2s /2 width=1/
76 #V1s #V2s #V1 #V2 #HV12 #H 
77 generalize in match HV12; -HV12 generalize in match V2; -V2 generalize in match V1; -V1
78 elim H -V1s -V2s /2 width=3/
79 #V1s #V2s #V1 #V2 #HV12 #HV12s #IHV12s #W1 #W2 #HW12 #V #T #H #HV
80 lapply (csn_appl_theta … HW12 … H) -H -HW12 #H
81 lapply (csn_fwd_pair_sn … H) #HW1
82 lapply (csn_fwd_flat_dx … H) #H1
83 @csn_appl_simple_tstc // -HW1 /2 width=3/ -IHV12s -HV -H1 #X #H1 #H2
84 elim (cprs_fwd_theta_vector … (V2@V2s) … H1) -H1 /2 width=1/ -HV12s -HV12
85 [ -H #H elim (H2 ?) -H2 //
86 | -H2 #H1 @(csn_cprs_trans … H) -H /2 width=1/
87 ]
88 qed.
89
90 (* Basic_1: was: sn3_appls_cast *)
91 lemma csn_applv_tau: ∀L,W. L ⊢ ⬊* W →
92                      ∀Vs,T. L ⊢ ⬊* ⒶVs. T →
93                      L ⊢ ⬊* ⒶVs. ⓝW. T.
94 #L #W #HW #Vs elim Vs -Vs /2 width=1/ -HW
95 #V #Vs #IHV #T #H1T
96 lapply (csn_fwd_pair_sn … H1T) #HV
97 lapply (csn_fwd_flat_dx … H1T) #H2T
98 @csn_appl_simple_tstc // -HV /2 width=1/ -IHV -H2T
99 #X #H #H0
100 elim (cprs_fwd_tau_vector … H) -H #H
101 [ -H1T elim (H0 ?) -H0 //
102 | -H0 @(csn_cprs_trans … H1T) -H1T /2 width=1/
103 ]
104 qed.
105
106 theorem csn_acr: acr cpr (eq …) (csn …) (λL,T. L ⊢ ⬊* T).
107 @mk_acr //
108 [ /3 width=1/
109 | /2 width=1/
110 | /2 width=6/
111 | #L #V1 #V2 #HV12 #a #V #T #H #HVT
112   @(csn_applv_theta … HV12) -HV12 //
113   @(csn_abbr) //
114 | /2 width=1/
115 | @csn_lift
116 ]
117 qed.