]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda_delta/basic_2/computation/csn_tstc_vector.ma
property S2 of strongly normalizing terms proved!
[helm.git] / matita / matita / contribs / lambda_delta / 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_lcpr.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 | -L -V elim Vs -Vs //
35 ]
36 qed.
37
38 (* Basic_1: was: sn3_appls_beta *)
39 lemma csn_applv_beta: ∀L,W. L ⊢ ⬇* W →
40                       ∀Vs,V,T. L ⊢ ⬇* ⒶVs.ⓓV.T →
41                       L ⊢ ⬇* ⒶVs. ⓐV.ⓛW. T.
42 #L #W #HW #Vs elim Vs -Vs /2 width=1/ -HW
43 #V0 #Vs #IHV #V #T #H1T
44 lapply (csn_fwd_pair_sn … H1T) #HV0
45 lapply (csn_fwd_flat_dx … H1T) #H2T
46 @csn_appl_simple_tstc // -HV0 /2 width=1/ -IHV -H2T
47 [ #X #H #H0
48   elim (cprs_fwd_beta_vector … H) -H #H
49   [ -H1T elim (H0 ?) -H0 //
50   | -H0 @(csn_cprs_trans … H1T) -H1T /2 width=1/
51   ]
52 | -H1T elim Vs -Vs //
53 ]
54 qed.
55
56 lemma csn_applv_delta: ∀L,K,V1,i. ⇩[0, i] L ≡ K. ⓓV1 →
57                        ∀V2. ⇧[0, i + 1] V1 ≡ V2 →
58                        ∀Vs.L ⊢ ⬇* (ⒶVs. V2) → L ⊢ ⬇* (ⒶVs. #i).
59 #L #K #V1 #i #HLK #V2 #HV12 #Vs elim Vs -Vs
60 [ #H
61   lapply (ldrop_fwd_ldrop2 … HLK) #HLK0
62   lapply (csn_inv_lift … H … HLK0 HV12) -V2 -HLK0 /2 width=4/
63 | #V #Vs #IHV #H1T
64   lapply (csn_fwd_pair_sn … H1T) #HV
65   lapply (csn_fwd_flat_dx … H1T) #H2T
66   @csn_appl_simple_tstc // -HV /2 width=1/ -IHV -H2T
67   [ #X #H #H0
68     elim (cprs_fwd_delta_vector … HLK … HV12 … H) -HLK -HV12 -H #H
69     [ -H1T elim (H0 ?) -H0 //
70     | -H0 @(csn_cprs_trans … H1T) -H1T /2 width=1/
71     ]
72   | -L -K -V -V1 -V2 elim Vs -Vs //
73   ]
74 ]
75 qed.
76
77 (* Basic_1: was: sn3_appls_abbr *) 
78 lemma csn_applv_theta: ∀L,V1s,V2s. ⇧[0, 1] V1s ≡ V2s →
79                        ∀V,T. L ⊢ ⬇* ⓓV. ⒶV2s. T → L ⊢ ⬇* V →
80                        L ⊢ ⬇* ⒶV1s. ⓓV. T.
81 #L #V1s #V2s * -V1s -V2s /2 width=1/
82 #V1s #V2s #V1 #V2 #HV12 #H 
83 generalize in match HV12; -HV12 generalize in match V2; -V2 generalize in match V1; -V1
84 elim H -V1s -V2s /2 width=3/
85 #V1s #V2s #V1 #V2 #HV12 #HV12s #IHV12s #W1 #W2 #HW12 #V #T #H #HV
86 lapply (csn_appl_theta … HW12 … H) -H -HW12 #H
87 lapply (csn_fwd_pair_sn … H) #HW1
88 lapply (csn_fwd_flat_dx … H) #H1
89 @csn_appl_simple_tstc // -HW1 /2 width=3/ -IHV12s -HV -H1 #X #H1 #H2
90 elim (cprs_fwd_theta_vector … (V2::V2s) … H1) -H1 /2 width=1/ -HV12s -HV12
91 [ -H #H elim (H2 ?) -H2 //
92 | -H2 #H1 @(csn_cprs_trans … H) -H /2 width=1/
93 ]
94 qed.
95
96 (* Basic_1: was: sn3_appls_cast *)
97 lemma csn_applv_tau: ∀L,W. L ⊢ ⬇* W →
98                      ∀Vs,T. L ⊢ ⬇* ⒶVs. T →
99                      L ⊢ ⬇* ⒶVs. ⓣW. T.
100 #L #W #HW #Vs elim Vs -Vs /2 width=1/ -HW
101 #V #Vs #IHV #T #H1T
102 lapply (csn_fwd_pair_sn … H1T) #HV
103 lapply (csn_fwd_flat_dx … H1T) #H2T
104 @csn_appl_simple_tstc // -HV /2 width=1/ -IHV -H2T
105 [ #X #H #H0
106   elim (cprs_fwd_tau_vector … H) -H #H
107   [ -H1T elim (H0 ?) -H0 //
108   | -H0 @(csn_cprs_trans … H1T) -H1T /2 width=1/
109   ]
110 | -H1T elim Vs -Vs //
111 ]
112 qed.
113
114 theorem csn_acr: acr cpr (eq …) (csn …) (λL,T. L ⊢ ⬇* T).
115 @mk_acr //
116 [ /3 width=1/
117 | /2 width=1/
118 | /2 width=6/
119 | #L #V1 #V2 #HV12 #V #T #H #HVT
120   @(csn_applv_theta … HV12) -HV12 //
121   @(csn_abbr) //
122 | /2 width=1/
123 | @csn_lift
124 ]
125 qed.