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/s_transition/fquq.ma".
16 include "basic_2/rt_transition/cpm_drops.ma".
17 include "basic_2/rt_transition/cpm_lsubr.ma".
18 include "basic_2/rt_transition/cpr.ma".
19 include "basic_2/rt_transition/lfpr_fqup.ma".
21 (* PARALLEL R-TRANSITION FOR LOCAL ENV.S ON REFERRED ENTRIES ****************)
23 (* Properties with supclosure ***********************************************)
25 lemma fqu_cpr_trans_dx: ∀h,b,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐[b] ⦃G2, L2, T2⦄ →
26 ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h] U2 →
27 ∃∃L,U1. ⦃G1, L1⦄ ⊢ ➡[h, T1] L & ⦃G1, L⦄ ⊢ T1 ➡[h] U1 & ⦃G1, L, U1⦄ ⊐[b] ⦃G2, L2, U2⦄.
28 #h #b #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
29 /3 width=5 by lfpr_pair, cpr_pair_sn, cpr_flat, cpm_bind, fqu_lref_O, fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, ex3_2_intro/
30 [ /5 width=5 by lsubr_cpm_trans, cpm_bind, lsubr_unit, fqu_clear, ex3_2_intro/
31 | #I #G #L #U #T #HUT #U2 #HU2 elim (cpm_lifts_sn … HU2 (Ⓣ) … HUT) -U
32 /3 width=9 by fqu_drop, drops_refl, drops_drop, ex3_2_intro/
36 (* Basic_2A1: uses: fqu_lpr_trans *)
37 lemma fqu_cpr_trans_sn: ∀h,b,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐[b] ⦃G2, L2, T2⦄ →
38 ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h] U2 →
39 ∃∃L,U1. ⦃G1, L1⦄ ⊢ ➡[h, T1] L & ⦃G1, L1⦄ ⊢ T1 ➡[h] U1 & ⦃G1, L, U1⦄ ⊐[b] ⦃G2, L2, U2⦄.
40 #h #b #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
41 /3 width=5 by lfpr_pair, cpr_pair_sn, cpr_flat, cpm_bind, fqu_lref_O, fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, ex3_2_intro/
42 [ /5 width=5 by lsubr_cpm_trans, cpm_bind, lsubr_unit, fqu_clear, ex3_2_intro/
43 | #I #G #L #U #T #HUT #U2 #HU2 elim (cpm_lifts_sn … HU2 (Ⓣ) … HUT) -U
44 /3 width=9 by fqu_drop, drops_refl, drops_drop, ex3_2_intro/
48 (* Properties with optional supclosure **************************************)
50 lemma fquq_cpr_trans_dx: ∀h,b,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮[b] ⦃G2, L2, T2⦄ →
51 ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h] U2 →
52 ∃∃L,U1. ⦃G1, L1⦄ ⊢ ➡[h, T1] L & ⦃G1, L⦄ ⊢ T1 ➡[h] U1 & ⦃G1, L, U1⦄ ⊐⸮[b] ⦃G2, L2, U2⦄.
53 #h #b #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -H
54 [ #HT12 #U2 #HTU2 elim (fqu_cpr_trans_dx … HT12 … HTU2) /3 width=5 by fqu_fquq, ex3_2_intro/
55 | * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/
59 (* Basic_2A1: uses: fquq_lpr_trans *)
60 lemma fquq_cpr_trans_sn: ∀h,b,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮[b] ⦃G2, L2, T2⦄ →
61 ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h] U2 →
62 ∃∃L,U1. ⦃G1, L1⦄ ⊢ ➡[h, T1] L & ⦃G1, L1⦄ ⊢ T1 ➡[h] U1 & ⦃G1, L, U1⦄ ⊐⸮[b] ⦃G2, L2, U2⦄.
63 #h #b #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -H
64 [ #HT12 #U2 #HTU2 elim (fqu_cpr_trans_sn … HT12 … HTU2) /3 width=5 by fqu_fquq, ex3_2_intro/
65 | * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/