]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/computation/lsx_csx.ma
670fb947324e1e92f02a33ce56a57125af4205bc
[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/lpxs_lleq.ma".
16 include "basic_2/computation/csx_alt.ma".
17 include "basic_2/computation/lsx_lpxs.ma".
18
19 (* SN EXTENDED STRONGLY NORMALIZING LOCAL ENVIRONMENTS **********************)
20
21 (* Advanced forward lemmas **************************************************)
22 (*
23 lemma lsx_fwd_bind_dx_lpxs: ∀h,g,a,I,G,L1,V. ⦃G, L1⦄ ⊢ ⬊*[h, g] V →
24                             ∀L2,T. G ⊢ ⋕⬊*[h, g, ⓑ{a,I}V.T] L2 → ⦃G, L1⦄ ⊢ ➡*[h, g] L2 →
25                             G ⊢ ⋕⬊*[h, g, T] L2.ⓑ{I}V.
26 #h #g #a #I #G #L1 #V #H @(csx_ind_alt … H) -V
27 #V1 #_ #IHV1 #L2 #T #H @(lsx_ind … H) -L2
28 #L2 #HL2 #IHL2 #HL12 @lsx_intro
29 #Y #H #HnT elim (lpxs_inv_pair1 … H) -H
30 #L3 #V3 #HL23 #HV13 #H destruct
31 lapply (lpxs_trans … HL12 … HL23) #HL13
32 elim (eq_term_dec V1 V3)
33 [ -HV13 -HL2 -IHV1 -HL12 #H destruct
34   @IHL2 -IHL2 // -HL23 -HL13 /3 width=2 by lleq_fwd_bind_O/
35 | -IHL2 -HL23 -HnT #HnV13
36   @(IHV1 … HnV13) -IHV1 -HnV13 /2 width=3 by lpxs_cpxs_trans/ -HL12 -HL13 -HV13
37   @(lsx_lpxs_trans) … HL2)
38 *)
39
40 (* Advanced inversion lemmas ************************************************)
41
42
43
44 (* Main properties **********************************************************)
45
46 axiom csx_lsx: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → G ⊢ ⋕⬊*[h, g, T] L.
47 (*
48 #h #g #G #L #T @(fqup_wf_ind_eq … G L T) -G -L -T
49 #G1 #L1 #T1 #IH #G2 #L2 * *
50 [ #k #HG #HL #HT destruct //
51 | #i #HG #HL #HT destruct
52   #H
53 | #p #HG #HL #HT destruct //
54 | #a #I #V2 #T2 #HG #HL #HT #H destruct
55   elim (csx_fwd_bind … H) -H
56   #HV2 #HT2
57 | #I #V2 #T2 #HG #HL #HT #H destruct
58   elim (csx_fwd_flat … H) -H /3 width=1 by lsx_flat/
59 *)