]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/rt_computation/fsb_fpbg.ma
milestone in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / fsb_fpbg.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/rt_computation/fpbg_fpbs.ma".
16 include "basic_2/rt_computation/fsb_fdeq.ma".
17
18 (* STRONGLY NORMALIZING CLOSURES FOR PARALLEL RST-TRANSITION ****************)
19
20 (* Properties with parallel rst-computation for closures ********************)
21
22 lemma fsb_fpbs_trans: ∀h,G1,L1,T1. ≥[h] 𝐒⦃G1, L1, T1⦄ →
23                       ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h] ⦃G2, L2, T2⦄ → ≥[h] 𝐒⦃G2, L2, T2⦄.
24 #h #G1 #L1 #T1 #H @(fsb_ind_alt … H) -G1 -L1 -T1
25 #G1 #L1 #T1 #H1 #IH #G2 #L2 #T2 #H12
26 elim (fpbs_inv_fpbg … H12) -H12
27 [ -IH /2 width=5 by fsb_fdeq_trans/
28 | -H1 * /2 width=5 by/
29 ]
30 qed-.
31
32 (* Properties with proper parallel rst-computation for closures *************)
33
34 lemma fsb_intro_fpbg: ∀h,G1,L1,T1. (
35                          ∀G2,L2,T2. ⦃G1, L1, T1⦄ >[h] ⦃G2, L2, T2⦄ → ≥[h] 𝐒⦃G2, L2, T2⦄
36                       ) → ≥[h] 𝐒⦃G1, L1, T1⦄.
37 /4 width=1 by fsb_intro, fpb_fpbg/ qed.
38
39 (* Eliminators with proper parallel rst-computation for closures ************)
40
41 lemma fsb_ind_fpbg_fpbs: ∀h. ∀Q:relation3 genv lenv term.
42                          (∀G1,L1,T1. ≥[h] 𝐒⦃G1, L1, T1⦄ →
43                                      (∀G2,L2,T2. ⦃G1, L1, T1⦄ >[h] ⦃G2, L2, T2⦄ → Q G2 L2 T2) →
44                                      Q G1 L1 T1
45                          ) →
46                          ∀G1,L1,T1. ≥[h] 𝐒⦃G1, L1, T1⦄ → 
47                          ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h] ⦃G2, L2, T2⦄ → Q G2 L2 T2.
48 #h #Q #IH1 #G1 #L1 #T1 #H @(fsb_ind_alt … H) -G1 -L1 -T1
49 #G1 #L1 #T1 #H1 #IH #G2 #L2 #T2 #H12
50 @IH1 -IH1
51 [ -IH /2 width=5 by fsb_fpbs_trans/
52 | -H1 #G0 #L0 #T0 #H10
53   elim (fpbs_fpbg_trans … H12 … H10) -G2 -L2 -T2
54   /2 width=5 by/
55 ]
56 qed-.
57
58 lemma fsb_ind_fpbg: ∀h. ∀Q:relation3 genv lenv term.
59                     (∀G1,L1,T1. ≥[h] 𝐒⦃G1, L1, T1⦄ →
60                                 (∀G2,L2,T2. ⦃G1, L1, T1⦄ >[h] ⦃G2, L2, T2⦄ → Q G2 L2 T2) →
61                                 Q G1 L1 T1
62                     ) →
63                     ∀G1,L1,T1. ≥[h] 𝐒⦃G1, L1, T1⦄ →  Q G1 L1 T1.
64 #h #Q #IH #G1 #L1 #T1 #H @(fsb_ind_fpbg_fpbs … H) -H
65 /3 width=1 by/
66 qed-.
67
68 (* Inversion lemmas with proper parallel rst-computation for closures *******)
69
70 lemma fsb_fpbg_refl_false (h) (G) (L) (T):
71                           ≥[h] 𝐒⦃G, L, T⦄ → ⦃G, L, T⦄ >[h] ⦃G, L, T⦄ → ⊥.
72 #h #G #L #T #H
73 @(fsb_ind_fpbg … H) -G -L -T #G1 #L1 #T1 #_ #IH #H
74 /2 width=5 by/
75 qed-.