]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/computation/fsb_csx.ma
ceb783b0bd7409e00eee06c79da06635aa8570d3
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / fsb_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_fpbs.ma".
16 include "basic_2/computation/lsx_csx.ma".
17 include "basic_2/computation/fsb_alt.ma".
18
19 (* "QRST" STRONGLY NORMALIZING TERMS ****************************************)
20
21 (* Advanced propreties on context-sensitive extended normalizing terms ******)
22
23 lemma csx_fsb_fpbs: ∀h,g,G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, g] T1 →
24                     ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ → ⦃G2, L2⦄ ⊢ ⦥[h, g] T2.
25 #h #g #G1 #L1 #T1 #H @(csx_ind … H) -T1
26 #T1 #HT1 #IHc #G2 #L2 #T2 @(fqup_wf_ind … G2 L2 T2) -G2 -L2 -T2
27 #G0 #L0 #T0 #IHu #H10 lapply (csx_fpbs_conf … H10) // -HT1
28 #HT0 generalize in match IHu; -IHu generalize in match H10; -H10
29 @(lsx_ind … (csx_lsx … HT0 0)) -L0
30 #L0 #_ #IHl #H10 #IHu @fsb_intro
31 #G2 #L2 #T2 * -G2 -L2 -T2 [ -IHl -IHc | -IHu -IHl |  ]
32 [ /4 width=5 by fpbs_fqup_trans, fqu_fqup/
33 | #T2 #HT02 #HnT02 elim (fpbs_cpx_trans_neq … H10 … HT02 HnT02) -T0
34   /3 width=4 by/
35 | #L2 #HL02 #HnL02 @(IHl … HL02 HnL02) -IHl -HnL02 [ -IHu -IHc | ]
36   [ /3 width=3 by fpbs_lpxs_trans, lpx_lpxs/
37   | #G3 #L3 #T3 #H03 #_ elim (lpx_fqup_trans … H03 … HL02) -L2
38     #L4 #T4 elim (eq_term_dec T0 T4) [ -IHc | -IHu ]
39     [ #H destruct /5 width=5 by fsb_fpbs_trans, lpxs_fpbs, fpbs_fqup_trans, lpx_lpxs/
40     | #HnT04 #HT04 #H04 #HL43 elim (cpxs_neq_inv_step_sn … HT04 HnT04) -HT04 -HnT04
41       #T2 #HT02 #HnT02 #HT24 elim (fpbs_cpx_trans_neq … H10 … HT02 HnT02) -T0
42       lapply (fpbs_intro_alt … G3 … L4 … L3 L3 … T3 … HT24 ? ? ?) -HT24
43       /3 width=8 by fpbs_trans, lpx_lpxs, fqup_fqus/ (**) (* full auto too slow *)
44     ]
45   ]
46 ]
47 qed.
48
49 lemma csx_fsb: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → ⦃G, L⦄ ⊢ ⦥[h, g] T.
50 /2 width=5 by csx_fsb_fpbs/ qed.
51
52 (* Advanced eliminators *****************************************************)
53
54 lemma csx_ind_fpbu: ∀h,g. ∀R:relation3 genv lenv term.
55                     (∀G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, g] T1 →
56                                 (∀G2,L2,T2. ⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
57                                 R G1 L1 T1
58                     ) →
59                     ∀G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → R G L T.
60 /4 width=4 by fsb_inv_csx, csx_fsb, fsb_ind_alt/ qed-.
61
62 lemma csx_ind_fpbg: ∀h,g. ∀R:relation3 genv lenv term.
63                     (∀G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, g] T1 →
64                                 (∀G2,L2,T2. ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄ → R G2 L2 T2) →
65                                 R G1 L1 T1
66                     ) →
67                     ∀G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → R G L T.
68 /4 width=4 by fsb_inv_csx, csx_fsb, fsb_ind_fpbg/ qed-.