]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2A/reduction/fpbq_alt.ma
update in lambdadelta
[helm.git] / matita / matita / contribs / lambdadelta / basic_2A / reduction / fpbq_alt.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_2A/notation/relations/btpredalt_8.ma".
16 include "basic_2A/reduction/fpb_fleq.ma".
17 include "basic_2A/reduction/fpbq.ma".
18
19 (* "QRST" PARALLEL REDUCTION FOR CLOSURES ***********************************)
20
21 (* alternative definition of fpbq *)
22 definition fpbqa: ∀h. sd h → tri_relation genv lenv term ≝
23                   λh,g,G1,L1,T1,G2,L2,T2.
24                   ⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ ∨ ⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄.
25
26 interpretation
27    "'qrst' parallel reduction (closure) alternative"
28    'BTPRedAlt h g G1 L1 T1 G2 L2 T2 = (fpbqa h g G1 L1 T1 G2 L2 T2).
29
30 (* Basic properties *********************************************************)
31
32 lemma fleq_fpbq: ∀h,g,G1,G2,L1,L2,T1,T2.
33                  ⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≽[h, g] ⦃G2, L2, T2⦄.
34 #h #g #G1 #G2 #L1 #L2 #T1 #T2 * /2 width=1 by fpbq_lleq/
35 qed.
36
37 lemma fpb_fpbq: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄ →
38                 ⦃G1, L1, T1⦄ ≽[h, g] ⦃G2, L2, T2⦄.
39 #h #g #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2
40 /3 width=1 by fpbq_fquq, fpbq_cpx, fpbq_lpx, fqu_fquq/
41 qed.
42
43 (* Main properties **********************************************************)
44
45 theorem fpbq_fpbqa: ∀h,g,G1,G2,L1,L2,T1,T2.
46                     ⦃G1, L1, T1⦄ ≽[h, g] ⦃G2, L2, T2⦄ →
47                     ⦃G1, L1, T1⦄ ≽≽[h, g] ⦃G2, L2, T2⦄.
48 #h #g #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2
49 [ #G2 #L2 #T2 #H elim (fquq_inv_gen … H) -H
50   [ /3 width=1 by fpb_fqu, or_intror/
51   | * #H1 #H2 #H3 destruct /2 width=1 by or_introl/
52   ]
53 | #T2 #HT12 elim (eq_term_dec T1 T2)
54   #HnT12 destruct /4 width=1 by fpb_cpx, or_intror, or_introl/
55 | #L2 #HL12 elim (lleq_dec … T1 L1 L2 0)
56   /4 width=1 by fpb_lpx, fleq_intro, or_intror, or_introl/
57 | /3 width=1 by fleq_intro, or_introl/
58 ]
59 qed.
60
61 theorem fpbqa_inv_fpbq: ∀h,g,G1,G2,L1,L2,T1,T2.
62                         ⦃G1, L1, T1⦄ ≽≽[h, g] ⦃G2, L2, T2⦄ →
63                         ⦃G1, L1, T1⦄ ≽[h, g] ⦃G2, L2, T2⦄.
64 #h #g #G1 #G2 #L1 #L2 #T1 #T2 * /2 width=1 by fleq_fpbq, fpb_fpbq/
65 qed-.
66
67 (* Advanced eliminators *****************************************************)
68
69 lemma fpbq_ind_alt: ∀h,g,G1,G2,L1,L2,T1,T2. ∀R:Prop.
70                     (⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → R) →
71                     (⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄ → R) →
72                     ⦃G1, L1, T1⦄ ≽[h, g] ⦃G2, L2, T2⦄ → R.
73 #h #g #G1 #G2 #L1 #L2 #T1 #T2 #R #HI1 #HI2 #H elim (fpbq_fpbqa … H) /2 width=1 by/
74 qed-.
75
76 (* aternative definition of fpb *********************************************)
77
78 lemma fpb_fpbq_alt: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄ →
79                     ⦃G1, L1, T1⦄ ≽[h, g] ⦃G2, L2, T2⦄ ∧ (⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⊥).
80 /3 width=10 by fpb_fpbq, fpb_inv_fleq, conj/ qed.
81
82 lemma fpbq_inv_fpb_alt: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≽[h, g] ⦃G2, L2, T2⦄ →
83                         (⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⊥) → ⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄.
84 #h #g #G1 #G2 #L1 #L2 #T1 #T2 #H #H0 @(fpbq_ind_alt … H) -H //
85 #H elim H0 -H0 //
86 qed-.