1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "basic_2/reducibility/cfpr_cpr.ma".
17 (* CONTEXT-FREE PARALLEL REDUCTION ON CLOSURES ******************************)
19 (* Advanced propertis *******************************************************)
21 lemma fpr_bind_sn: ∀L1,L2,V1,V2. ⦃L1, V1⦄ ➡ ⦃L2, V2⦄ → ∀T1,T2. T1 ➡ T2 →
22 ∀a,I. ⦃L1, ⓑ{a,I}V1.T1⦄ ➡ ⦃L2, ⓑ{a,I}V2.T2⦄.
23 #L1 #L2 #V1 #V2 #H #T1 #T2 #HT12 #a #I
24 elim (fpr_inv_all … H) /3 width=4/
27 (* Advanced forward lemmas **************************************************)
29 lemma fpr_fwd_shift_bind_minus: ∀I1,I2,L1,L2,V1,V2,T1,T2.
30 ⦃L1, -ⓑ{I1}V1.T1⦄ ➡ ⦃L2, -ⓑ{I2}V2.T2⦄ →
31 ⦃L1, V1⦄ ➡ ⦃L2, V2⦄ ∧ I1 = I2.
32 * #I2 #L1 #L2 #V1 #V2 #T1 #T2 #H
33 elim (fpr_inv_all … H) -H #L #HL1 #H #HL2
34 [ elim (cpr_inv_abbr1 … H) -H *
35 [ #V #V0 #T #HV1 #HV0 #_ #H destruct /4 width=4/
36 | #T #_ #_ #H destruct
38 | elim (cpr_inv_abst1 … H Abst V2) -H
39 #V #T #HV1 #_ #H destruct /3 width=4/
43 (* Advanced inversion lemmas ************************************************)
45 lemma fpr_inv_pair1: ∀I,K1,L2,V1,T1,T2. ⦃K1.ⓑ{I}V1, T1⦄ ➡ ⦃L2, T2⦄ →
46 ∃∃K2,V2. ⦃K1, V1⦄ ➡ ⦃K2, V2⦄ &
47 ⦃K1, -ⓑ{I}V1.T1⦄ ➡ ⦃K2, -ⓑ{I}V2.T2⦄ &
49 #I1 #K1 #X #V1 #T1 #T2 #H
50 elim (fpr_fwd_pair1 … H) -H #I2 #K2 #V2 #HT12 #H destruct
51 elim (fpr_fwd_shift_bind_minus … HT12) #HV12 #H destruct /2 width=5/
54 lemma fpr_inv_pair3: ∀I,L1,K2,V2,T1,T2. ⦃L1, T1⦄ ➡ ⦃K2.ⓑ{I}V2, T2⦄ →
55 ∃∃K1,V1. ⦃K1, V1⦄ ➡ ⦃K2, V2⦄ &
56 ⦃K1, -ⓑ{I}V1.T1⦄ ➡ ⦃K2, -ⓑ{I}V2.T2⦄ &
58 #I2 #X #K2 #V2 #T1 #T2 #H
59 elim (fpr_fwd_pair3 … H) -H #I1 #K1 #V1 #HT12 #H destruct
60 elim (fpr_fwd_shift_bind_minus … HT12) #HV12 #H destruct /2 width=5/