]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda_delta/basic_2/computation/csn_vector.ma
property S2 of strongly normalizing terms proved!
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / computation / csn_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/grammar/term_vector.ma".
16 include "basic_2/computation/csn.ma".
17
18 (* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERM VECTORS **********************)
19
20 inductive csnv (L:lenv): predicate (list term) ≝
21 | csnv_nil: csnv L ◊
22 | csn_cons: ∀Ts,T. L  ⊢ ⬇* T → csnv L Ts → csnv L (T :: Ts)
23 .
24
25 interpretation
26    "context-sensitive strong normalization (term vector)"
27    'SN L Ts = (csnv L Ts).
28
29 (* Basic properties *********************************************************)
30
31 lemma all_csnv: ∀L,Vs. all … (csn L) Vs → L ⊢ ⬇* Vs.
32 #L #Vs elim Vs -Vs //
33 #V #Vs #IHVs * /3 width=1/ 
34 qed.
35
36 (* Basic inversion lemmas ***************************************************)
37
38 fact csnv_inv_cons_aux: ∀L,Ts. L ⊢ ⬇* Ts → ∀U,Us. Ts = U :: Us →
39                         L ⊢ ⬇* U ∧ L ⊢ ⬇* Us.
40 #L #Ts * -Ts
41 [ #U #Us #H destruct
42 | #Ts #T #HT #HTs #U #Us #H destruct /2 width=1/
43 ]
44 qed.
45
46 lemma csnv_inv_cons: ∀L,T,Ts. L ⊢ ⬇* T :: Ts → L ⊢ ⬇* T ∧ L ⊢ ⬇* Ts.
47 /2 width=3/ qed-.
48
49 include "basic_2/computation/csn_cpr.ma".
50
51 lemma csn_fwd_applv: ∀L,T,Vs. L ⊢ ⬇* Ⓐ Vs. T → L ⊢ ⬇* Vs ∧ L ⊢ ⬇* T.
52 #L #T #Vs elim Vs -Vs /2 width=1/
53 #V #Vs #IHVs #HVs
54 lapply (csn_fwd_pair_sn … HVs) #HV
55 lapply (csn_fwd_flat_dx … HVs) -HVs #HVs
56 elim (IHVs HVs) -IHVs -HVs /3 width=1/
57 qed-.