]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/computation/lsx_csx.ma
13e753a6cf350dd45bfc9683e3db1e26af1a06b4
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / lsx_csx.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/csx_lpxs.ma".
16 include "basic_2/computation/lcosx_cpxs.ma".
17
18 (* SN EXTENDED STRONGLY NORMALIZING LOCAL ENVIRONMENTS **********************)
19
20 (* Advanced properties ******************************************************)
21
22 lemma lsx_lref_be_lpxs: ∀h,g,I,G,K1,V,i,d. d ≤ yinj i → ⦃G, K1⦄ ⊢ ⬊*[h, g] V →
23                         ∀K2. G ⊢ ⋕⬊*[h, g, V, 0] K2 → ⦃G, K1⦄ ⊢ ➡*[h, g] K2 →
24                         ∀L2. ⇩[i] L2 ≡ K2.ⓑ{I}V → G ⊢ ⋕⬊*[h, g, #i, d] L2.
25 #h #g #I #G #K1 #V #i #d #Hdi #H @(csx_ind_alt … H) -V
26 #V0 #_ #IHV0 #K2 #H @(lsx_ind … H) -K2
27 #K0 #HK0 #IHK0 #HK10 #L0 #HLK0 @lsx_intro
28 #L2 #HL02 #HnL02 elim (lpxs_ldrop_conf … HLK0 … HL02) -HL02
29 #Y #H #HLK2 elim (lpxs_inv_pair1 … H) -H
30 #K2 #V2 #HK02 #HV02 #H destruct
31 lapply (lpxs_trans … HK10 … HK02) #HK12
32 elim (eq_term_dec V0 V2) #HnV02 destruct [ -IHV0 -HV02 -HK0 -HK10 | -IHK0 -HnL02 -HLK0 ]
33 [ /4 width=8 by lleq_lref/
34 | @(IHV0 … HnV02 … HK12 … HLK2) -IHV0 -HnV02 -HK12 -HLK2
35   /3 width=4 by lsx_cpxs_trans_O, lsx_lpxs_trans, lpxs_cpxs_trans/ (**) (* full auto too slow *)
36 ]
37 qed.
38
39 lemma lsx_lref_be: ∀h,g,I,G,K,V,i,d. d ≤ yinj i → ⦃G, K⦄ ⊢ ⬊*[h, g] V →
40                    G ⊢ ⋕⬊*[h, g, V, 0] K →
41                    ∀L. ⇩[i] L ≡ K.ⓑ{I}V → G ⊢ ⋕⬊*[h, g, #i, d] L.
42 /2 width=8 by lsx_lref_be_lpxs/ qed.
43
44 (* Main properties **********************************************************)
45
46 theorem csx_lsx: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → ∀d. G ⊢ ⋕⬊*[h, g, T, d] L.
47 #h #g #G #L #T @(fqup_wf_ind_eq … G L T) -G -L -T
48 #Z #Y #X #IH #G #L * * //
49 [ #i #HG #HL #HT #H #d destruct
50   elim (lt_or_ge i (|L|)) /2 width=1 by lsx_lref_free/
51   elim (ylt_split i d) /2 width=1 by lsx_lref_skip/
52   #Hdi #Hi elim (ldrop_O1_lt … Hi) -Hi
53   #I #K #V #HLK lapply (csx_inv_lref_bind … HLK … H) -H
54   /4 width=6 by lsx_lref_be, fqup_lref/
55 | #a #I #V #T #HG #HL #HT #H #d destruct
56   elim (csx_fwd_bind … H) -H /3 width=1 by lsx_bind/
57 | #I #V #T #HG #HL #HT #H #d destruct
58   elim (csx_fwd_flat … H) -H /3 width=1 by lsx_flat/
59 ]
60 qed.
61
62 (* Advanced eliminators *****************************************************)
63
64 fact csx_ind_lsx_aux: ∀h,g,G,T,d. ∀R:predicate lenv.
65                       (∀L1. ⦃G, L1⦄ ⊢ ⬊*[h, g] T →
66                             (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → (L1 ⋕[T, d] L2 → ⊥) → R L2) →
67                             R L1
68                       ) →
69                       ∀L. G ⊢ ⋕⬊*[h, g, T, d] L → ⦃G, L⦄ ⊢ ⬊*[h, g] T → R L.
70 #h #g #G #T #d #R #IH #L #H @(lsx_ind … H) -L
71 #L1 #_ #IHL1 #HL1 @IH -IH //
72 #L2 #HL12 #HnT @IHL1 -IHL1 /2 width=3 by csx_lpxs_conf/
73 qed-.
74
75 lemma csx_ind_lsx: ∀h,g,G,T,d. ∀R:predicate lenv.
76                    (∀L1. ⦃G, L1⦄ ⊢ ⬊*[h, g] T →
77                          (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → (L1 ⋕[T, d] L2 → ⊥) → R L2) →
78                          R L1
79                    ) →
80                    ∀L. ⦃G, L⦄ ⊢ ⬊*[h, g] T → R L.
81 #h #g #G #T #d #R #IH #L #HL @(csx_ind_lsx_aux … d … HL) /4 width=1 by csx_lsx/
82 qed-.