]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda_delta/Basic_2/computation/csn_lcpr.ma
- more properties on strongly normalizing terms
[helm.git] / matita / matita / contribs / lambda_delta / Basic_2 / computation / csn_lcpr.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/lcpr_cpr.ma".
16 include "Basic_2/computation/cprs_cprs.ma".
17 include "Basic_2/computation/csn_lift.ma".
18 include "Basic_2/computation/csn_cpr.ma".
19 include "Basic_2/computation/csn_cprs.ma".
20
21 (* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERMS *****************************)
22
23 (* Advanced properties ******************************************************)
24
25 lemma csn_lcpr_conf: ∀L1,L2. L1 ⊢ ➡ L2 → ∀T. L1 ⊢ ⬇* T → L2 ⊢ ⬇* T.
26 #L1 #L2 #HL12 #T #H @(csn_ind_cprs … H) -T #T #_ #IHT
27 @csn_intro #T0 #HLT0 #HT0
28 @IHT /2 width=2/ -IHT -HT0 /2 width=3/
29 qed.
30
31 lemma csn_abbr: ∀L,V. L ⊢ ⬇* V → ∀T. L. ⓓV ⊢ ⬇* T → L ⊢ ⬇* ⓓV. T.
32 #L #V #HV elim HV -V #V #_ #IHV #T #HT @(csn_ind_cprs … HT) -T #T #HT #IHT
33 @csn_intro #X #H1 #H2
34 elim (cpr_inv_abbr1 … H1) -H1 *
35 [ #V0 #V1 #T1 #HLV0 #HLV01 #HLT1 #H destruct
36   lapply (cpr_intro … HLV0 HLV01) -HLV01 #HLV1
37   lapply (ltpr_cpr_trans (L. ⓓV) … HLT1) /2 width=1/ -V0 #HLT1
38   elim (eq_false_inv_tpair … H2) -H2
39   [ #HV1 @IHV // /2 width=1/ -HV1
40     @(csn_lcpr_conf (L. ⓓV)) /2 width=1/ -HLV1 /2 width=3/
41   | -IHV -HLV1 * #H destruct /3 width=1/
42   ]
43 | -IHV -IHT -H2 #T0 #HT0 #HLT0
44   lapply (csn_inv_lift … HT … HT0) -HT /2 width=3/
45 ]
46 qed.
47
48 fact csn_appl_beta_aux: ∀L,W. L ⊢ ⬇* W → ∀U. L ⊢ ⬇* U →
49                         ∀V,T. U = ⓓV. T → L ⊢ ⬇* ⓐV. ⓛW. T.
50 #L #W #H elim H -W #W #_ #IHW #X #H @(csn_ind_cprs … H) -X #X #HVT #IHVT #V #T #H destruct
51 lapply (csn_fwd_pair_sn … HVT) #HV
52 lapply (csn_fwd_bind_dx … HVT) #HT -HVT
53 @csn_intro #X #H #H2
54 elim (cpr_inv_appl1 … H) -H *
55 [ #V0 #Y #HLV0 #H #H0 destruct
56   elim (cpr_inv_abst1 … H Abbr V) -H #W0 #T0 #HLW0 #HLT0 #H destruct
57   elim (eq_false_inv_beta … H2) -H2
58   [ -IHVT #HW0 @IHW -IHW [1,5: // |3: skip ] -HLW0 /2 width=1/ -HW0
59     @csn_abbr /2 width=3/ -HV
60     @(csn_lcpr_conf (L. ⓓV)) /2 width=1/ -V0 /2 width=3/
61   | -IHW -HLW0 -HV -HT * #H #HVT0 destruct
62     @(IHVT … HVT0) -IHVT -HVT0 // /2 width=1/
63   ]
64 | -IHW -IHVT -H2 #V0 #W0 #T0 #T1 #HLV0 #HLT01 #H1 #H2 destruct
65   lapply (lcpr_cpr_trans (L. ⓓV) … HLT01) -HLT01 /2 width=1/ #HLT01
66   @csn_abbr /2 width=3/ -HV
67   @(csn_lcpr_conf (L. ⓓV)) /2 width=1/ -V0 /2 width=3/
68 | -IHW -IHVT -HV -HT -H2 #V0 #V1 #W0 #W1 #T0 #T1 #_ #_ #_ #_ #H destruct
69 ]
70 qed.
71
72 (* Basic_1: was: sn3_beta *)
73 lemma csn_appl_beta: ∀L,W. L ⊢ ⬇* W → ∀V,T. L ⊢ ⬇* (ⓓV. T) → (**)
74                      L ⊢ ⬇* ⓐV. ⓛW. T.
75 /2 width=3/ qed.