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/rt_computation/cpxs_lfpx.ma".
16 include "basic_2/rt_computation/lfpxs_fqup.ma".
18 (* UNCOUNTED PARALLEL RT-COMPUTATION FOR LOCAL ENV.S ON REFERRED ENTRIES ****)
20 (* Properties with uncounted context-sensitive rt-computation for terms *****)
22 lemma lfpxs_pair: ∀h,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ⬈*[h] V2 →
23 ∀I,T. ⦃G, L.ⓑ{I}V1⦄ ⊢ ⬈*[h, T] L.ⓑ{I}V2.
24 #h #G #L #V1 #V2 #H @(cpxs_ind … H) -V2
25 /3 width=3 by lfpxs_strap1, lfpx_pair/
28 (* Basic_2A1: was just: lpxs_cpx_trans *)
29 lemma lfpxs_cpx_trans: ∀h,G. s_r_transitive … (cpx h G) (lfpxs h G).
30 /3 width=5 by s_r_trans_LTC2, lfpx_cpxs_trans/ qed-.
32 (* Note: lfpxs_cpx_conf does not hold, thus we cannot invoke s_r_trans_LTC1 *)
33 (* Basic_2A1: was just: lpxs_cpxs_trans *)
34 lemma lfpxs_cpxs_trans: ∀h,G. s_rs_transitive … (cpx h G) (lfpxs h G).
35 #h #G @s_r_to_s_rs_trans @s_r_trans_LTC2
36 @s_rs_trans_TC1 /2 width=3 by lfpx_cpxs_trans/ (**) (* full auto too slow *)
39 (* Advanced properties on uncounted rt-computation for terms ****************)
41 lemma cpxs_bind2: ∀h,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ⬈*[h] V2 →
42 ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ⬈*[h] T2 →
43 ∀p. ⦃G, L⦄ ⊢ ⓑ{p,I}V1.T1 ⬈*[h] ⓑ{p,I}V2.T2.
44 /4 width=3 by lfpxs_cpxs_trans, lfpxs_pair, cpxs_bind/ qed.
46 (* Advanced inversion lemmas on uncounted rt-computation for terms **********)
48 lemma cpxs_inv_abst1: ∀h,p,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{p}V1.T1 ⬈*[h] U2 →
49 ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ⬈*[h] V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 ⬈*[h] T2 &
51 #h #p #G #L #V1 #T1 #U2 #H @(cpxs_ind … H) -U2 /2 width=5 by ex3_2_intro/
52 #U0 #U2 #_ #HU02 * #V0 #T0 #HV10 #HT10 #H destruct
53 elim (cpx_inv_abst1 … HU02) -HU02 #V2 #T2 #HV02 #HT02 #H destruct
54 lapply (lfpxs_cpx_trans … HT02 (L.ⓛV1) ?)
55 /3 width=5 by lfpxs_pair, cpxs_trans, cpxs_strap1, ex3_2_intro/
58 lemma cpxs_inv_abbr1: ∀h,p,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{p}V1.T1 ⬈*[h] U2 → (
59 ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ⬈*[h] V2 & ⦃G, L.ⓓV1⦄ ⊢ T1 ⬈*[h] T2 &
62 ∃∃T2. ⦃G, L.ⓓV1⦄ ⊢ T1 ⬈*[h] T2 & ⬆*[1] U2 ≡ T2 & p = true.
63 #h #p #G #L #V1 #T1 #U2 #H @(cpxs_ind … H) -U2 /3 width=5 by ex3_2_intro, or_introl/
65 [ #V0 #T0 #HV10 #HT10 #H destruct
66 elim (cpx_inv_abbr1 … HU02) -HU02 *
67 [ #V2 #T2 #HV02 #HT02 #H destruct
68 lapply (lfpxs_cpx_trans … HT02 (L.ⓓV1) ?)
69 /4 width=5 by lfpxs_pair, cpxs_trans, cpxs_strap1, ex3_2_intro, or_introl/
71 lapply (lfpxs_cpx_trans … HT02 (L.ⓓV1) ?) -HT02
72 /4 width=3 by lfpxs_pair, cpxs_trans, ex3_intro, or_intror/
75 elim (cpx_lifts … HU02 (Ⓣ) … (L.ⓓV1) … HU01)
76 /4 width=3 by cpxs_strap1, drops_refl, drops_drop, ex3_intro, or_intror/