]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/computation/llprs_cprs.ma
commit of the "computation" component with lazy pointwise extensions
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / llprs_cprs.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/relocation/llpx_sn_tc.ma".
16 include "basic_2/computation/cprs_cprs.ma".
17 include "basic_2/computation/llprs.ma".
18
19 (* LAZY SN PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS ***********************)
20
21 (* Advanced properties ******************************************************)
22
23 lemma llprs_pair_dx: ∀G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡* V2 →
24                      ∀I,T. ⦃G, L.ⓑ{I}V1⦄ ⊢ ➡*[T, 0] L.ⓑ{I}V2.
25 /2 width=1 by llpx_sn_TC_pair_dx/ qed.
26
27 (* Properties on context-sensitive parallel computation for terms ***********)
28
29 lemma llprs_cpr_trans: ∀G. s_r_transitive … (cpr G) (llprs G 0).
30 /3 width=5 by cprs_llpr_trans, s_r_trans_LTC2/ qed-.
31
32 (* Basic_1: was just: pr3_pr3_pr3_t *)
33 lemma llprs_cprs_trans: ∀G. s_rs_transitive … (cpr G) (llprs G 0).
34 #G @s_r_to_s_rs_trans @s_r_trans_LTC2
35 /3 width=5 by cprs_llpr_trans, s_rs_trans_TC1/ (**) (* full auto too slow *)
36 qed-.
37
38 lemma cprs_bind2: ∀G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡* V2 →
39                   ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ➡* T2 →
40                   ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡* ⓑ{a,I}V2.T2.
41 /4 width=3 by llprs_cprs_trans, llprs_pair_dx, cprs_bind/ qed.
42
43 (* Inversion lemmas on context-sensitive parallel computation for terms *****)
44
45 (* Basic_1: was: pr3_gen_abst *)
46 lemma cprs_inv_abst1: ∀a,G,L,W1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{a}W1.T1 ➡* U2 →
47                       ∃∃W2,T2. ⦃G, L⦄ ⊢ W1 ➡* W2 & ⦃G, L.ⓛW1⦄ ⊢ T1 ➡* T2 &
48                                U2 = ⓛ{a}W2.T2.
49 #a #G #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /2 width=5 by ex3_2_intro/
50 #U0 #U2 #_ #HU02 * #V0 #T0 #HV10 #HT10 #H destruct
51 elim (cpr_inv_abst1 … HU02) -HU02 #V2 #T2 #HV02 #HT02 #H destruct
52 lapply (llprs_cpr_trans … HT02 (L.ⓛV1) ?)
53 /3 width=5 by llprs_pair_dx, cprs_trans, cprs_strap1, ex3_2_intro/
54 qed-.
55
56 lemma cprs_inv_abst: ∀a,G,L,W1,W2,T1,T2. ⦃G, L⦄ ⊢ ⓛ{a}W1.T1 ➡* ⓛ{a}W2.T2 →
57                      ⦃G, L⦄ ⊢ W1 ➡* W2 ∧ ⦃G, L.ⓛW1⦄ ⊢ T1 ➡* T2.
58 #a #G #L #W1 #W2 #T1 #T2 #H
59 elim (cprs_inv_abst1 … H) -H #W #T #HW1 #HT1 #H destruct /2 width=1 by conj/
60 qed-.
61
62 (* Basic_1: was pr3_gen_abbr *)
63 lemma cprs_inv_abbr1: ∀a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{a}V1.T1 ➡* U2 → (
64                       ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡* V2 & ⦃G, L.ⓓV1⦄ ⊢ T1 ➡* T2 &
65                                U2 = ⓓ{a}V2.T2
66                       ) ∨
67                       ∃∃T2. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡* T2 & ⇧[0, 1] U2 ≡ T2 & a = true.
68 #a #G #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5 by ex3_2_intro, or_introl/
69 #U0 #U2 #_ #HU02 * *
70 [ #V0 #T0 #HV10 #HT10 #H destruct
71   elim (cpr_inv_abbr1 … HU02) -HU02 *
72   [ #V2 #T2 #HV02 #HT02 #H destruct
73     lapply (llprs_cpr_trans … HT02 (L.ⓓV1) ?)
74     /4 width=5 by llprs_pair_dx, cprs_trans, cprs_strap1, ex3_2_intro, or_introl/
75   | #T2 #HT02 #HUT2
76     lapply (llprs_cpr_trans … HT02 (L.ⓓV1) ?) -HT02
77     /4 width=3 by llprs_pair_dx, cprs_trans, ex3_intro, or_intror/
78   ]
79 | #U1 #HTU1 #HU01
80   elim (lift_total U2 0 1) #U #HU2
81   lapply (cpr_lift … HU02 (L.ⓓV1) … HU01 … HU2) -U0
82   /4 width=3 by cprs_strap1, ldrop_drop, ex3_intro, or_intror/
83 ]
84 qed-.