]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda_delta/Basic_2/computation/csn_lift.ma
605e78db110f0c89e344cc58e84e3c12eed9d386
[helm.git] / matita / matita / contribs / lambda_delta / Basic_2 / computation / csn_lift.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/reducibility/cnf_lift.ma".
16 include "Basic_2/computation/acp.ma".
17 include "Basic_2/computation/csn.ma".
18
19 (* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERMS *****************************)
20
21 (* Advanced properties ******************************************************)
22
23 lemma csn_acp: acp cpr (eq …) (csn …).
24 @mk_acp
25 [ /2 width=1/
26 | /2 width=3/
27 | /2 width=5/
28 | @cnf_lift
29 ]
30 qed.
31
32 lemma csn_abst: ∀L,W. L ⊢ ⬇* W → ∀I,V,T. L. ⓑ{I} V ⊢ ⬇* T → L ⊢ ⬇* ⓛW. T.
33 #L #W #HW elim HW -W #W #_ #IHW #I #V #T #HT @(csn_ind … HT) -T #T #HT #IHT
34 @csn_intro #X #H1 #H2
35 elim (cpr_inv_abst1 … H1 I V) -H1
36 #W0 #T0 #HLW0 #HLT0 #H destruct
37 elim (eq_false_inv_tpair … H2) -H2
38 [ /3 width=5/
39 | -HLW0 * #H destruct /3 width=1/
40 ]
41 qed.
42
43 (* Relocation properties ****************************************************)
44
45 (* Basic_1: was: sn3_lift *)
46 lemma csn_lift: ∀L2,L1,T1,d,e. L1 ⊢ ⬇* T1 →
47                 ∀T2. ⇩[d, e] L2 ≡ L1 → ⇧[d, e] T1 ≡ T2 → L2 ⊢ ⬇* T2.
48 #L2 #L1 #T1 #d #e #H elim H -T1 #T1 #_ #IHT1 #T2 #HL21 #HT12
49 @csn_intro #T #HLT2 #HT2
50 elim (cpr_inv_lift … HL21 … HT12 … HLT2) -HLT2 #T0 #HT0 #HLT10
51 @(IHT1 … HLT10) // -L1 -L2 #H destruct
52 >(lift_mono … HT0 … HT12) in HT2; -T0 /2 width=1/
53 qed.
54
55 (* Basic_1: was: sn3_gen_lift *)
56 lemma csn_inv_lift: ∀L2,L1,T1,d,e. L1 ⊢ ⬇* T1 →
57                     ∀T2. ⇩[d, e] L1 ≡ L2 → ⇧[d, e] T2 ≡ T1 → L2 ⊢ ⬇* T2.
58 #L2 #L1 #T1 #d #e #H elim H -T1 #T1 #_ #IHT1 #T2 #HL12 #HT21
59 @csn_intro #T #HLT2 #HT2
60 elim (lift_total T d e) #T0 #HT0
61 lapply (cpr_lift … HL12 … HT21 … HT0 HLT2) -HLT2 #HLT10
62 @(IHT1 … HLT10) // -L1 -L2 #H destruct
63 >(lift_inj … HT0 … HT21) in HT2; -T0 /2 width=1/
64 qed.