]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda_delta/basic_2/reducibility/ltpr.ma
We are decapitalizing the contributions' names ...
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / reducibility / ltpr.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/reducibility/tpr.ma".
16
17 (* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************)
18
19 inductive ltpr: relation lenv ≝
20 | ltpr_stom: ltpr (⋆) (⋆)
21 | ltpr_pair: ∀K1,K2,I,V1,V2.
22              ltpr K1 K2 → V1 ➡ V2 → ltpr (K1. ⓑ{I} V1) (K2. ⓑ{I} V2) (*ⓑ*)
23 .
24
25 interpretation
26   "context-free parallel reduction (environment)"
27   'PRed L1 L2 = (ltpr L1 L2).
28
29 (* Basic properties *********************************************************)
30
31 lemma ltpr_refl: ∀L:lenv. L ➡ L.
32 #L elim L -L // /2 width=1/
33 qed.
34
35 (* Basic inversion lemmas ***************************************************)
36
37 fact ltpr_inv_atom1_aux: ∀L1,L2. L1 ➡ L2 → L1 = ⋆ → L2 = ⋆.
38 #L1 #L2 * -L1 -L2
39 [ //
40 | #K1 #K2 #I #V1 #V2 #_ #_ #H destruct
41 ]
42 qed.
43
44 (* Basic_1: was: wcpr0_gen_sort *)
45 lemma ltpr_inv_atom1: ∀L2. ⋆ ➡ L2 → L2 = ⋆.
46 /2 width=3/ qed-.
47
48 fact ltpr_inv_pair1_aux: ∀L1,L2. L1 ➡ L2 → ∀K1,I,V1. L1 = K1. ⓑ{I} V1 →
49                          ∃∃K2,V2. K1 ➡ K2 & V1 ➡ V2 & L2 = K2. ⓑ{I} V2.
50 #L1 #L2 * -L1 -L2
51 [ #K1 #I #V1 #H destruct
52 | #K1 #K2 #I #V1 #V2 #HK12 #HV12 #L #J #W #H destruct /2 width=5/
53 ]
54 qed.
55
56 (* Basic_1: was: wcpr0_gen_head *)
57 lemma ltpr_inv_pair1: ∀K1,I,V1,L2. K1. ⓑ{I} V1 ➡ L2 →
58                       ∃∃K2,V2. K1 ➡ K2 & V1 ➡ V2 & L2 = K2. ⓑ{I} V2.
59 /2 width=3/ qed-.
60
61 fact ltpr_inv_atom2_aux: ∀L1,L2. L1 ➡ L2 → L2 = ⋆ → L1 = ⋆.
62 #L1 #L2 * -L1 -L2
63 [ //
64 | #K1 #K2 #I #V1 #V2 #_ #_ #H destruct
65 ]
66 qed.
67
68 lemma ltpr_inv_atom2: ∀L1. L1 ➡ ⋆ → L1 = ⋆.
69 /2 width=3/ qed-.
70
71 fact ltpr_inv_pair2_aux: ∀L1,L2. L1 ➡ L2 → ∀K2,I,V2. L2 = K2. ⓑ{I} V2 →
72                          ∃∃K1,V1. K1 ➡ K2 & V1 ➡ V2 & L1 = K1. ⓑ{I} V1.
73 #L1 #L2 * -L1 -L2
74 [ #K2 #I #V2 #H destruct
75 | #K1 #K2 #I #V1 #V2 #HK12 #HV12 #K #J #W #H destruct /2 width=5/
76 ]
77 qed.
78
79 lemma ltpr_inv_pair2: ∀L1,K2,I,V2. L1 ➡ K2. ⓑ{I} V2 →
80                       ∃∃K1,V1. K1 ➡ K2 & V1 ➡ V2 & L1 = K1. ⓑ{I} V1.
81 /2 width=3/ qed-.
82
83 (* Basic forward lemmas *****************************************************)
84
85 lemma ltpr_fwd_length: ∀L1,L2. L1 ➡ L2 → |L1| = |L2|.
86 #L1 #L2 #H elim H -L1 -L2 normalize //
87 qed-.
88
89 (* Basic_1: removed theorems 2: wcpr0_getl wcpr0_getl_back *)