]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda/labelled_hap_computation.ma
This line, and those below, will be ignored--
[helm.git] / matita / matita / contribs / lambda / labelled_hap_computation.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 "labelled_sequential_computation.ma".
16 include "labelled_hap_reduction.ma".
17
18 (* KASHIMA'S "HAP" COMPUTATION (LABELLED MULTISTEP) *************************)
19
20 (* Note: this is the "head in application" computation of:
21          R. Kashima: "A proof of the Standization Theorem in λ-Calculus". Typescript note, (2000).
22 *)
23 definition lhap: pseq → relation term ≝ lstar … lhap1.
24
25 interpretation "labelled 'hap' computation"
26    'HApStar M p N = (lhap p M N).
27
28 notation "hvbox( M break ⓗ⇀* [ term 46 p ] break term 46 N )"
29    non associative with precedence 45
30    for @{ 'HApStar $M $p $N }.
31
32 lemma lhap_step_rc: ∀p,M1,M2. M1 ⓗ⇀[p] M2 → M1 ⓗ⇀*[p::◊] M2.
33 /2 width=1/
34 qed.
35
36 lemma lhap_inv_nil: ∀s,M1,M2. M1 ⓗ⇀*[s] M2 → ◊ = s → M1 = M2.
37 /2 width=5 by lstar_inv_nil/
38 qed-.
39
40 lemma lhap_inv_cons: ∀s,M1,M2. M1 ⓗ⇀*[s] M2 → ∀q,r. q::r = s →
41                      ∃∃M. M1 ⓗ⇀[q] M & M ⓗ⇀*[r] M2.
42 /2 width=3 by lstar_inv_cons/
43 qed-.
44
45 lemma lhap_inv_step_rc: ∀p,M1,M2. M1 ⓗ⇀*[p::◊] M2 → M1 ⓗ⇀[p] M2.
46 /2 width=1 by lstar_inv_step/
47 qed-.
48
49 lemma lhap_inv_pos: ∀s,M1,M2. M1 ⓗ⇀*[s] M2 → 0 < |s| →
50                     ∃∃p,r,M. p::r = s & M1 ⓗ⇀[p] M & M ⓗ⇀*[r] M2.
51 /2 width=1 by lstar_inv_pos/
52 qed-.
53
54 lemma lhap_compatible_dx: ho_compatible_dx lhap.
55 /3 width=1/
56 qed.
57
58 lemma lhap_lift: ∀s. liftable (lhap s).
59 /2 width=1/
60 qed.
61
62 lemma lhap_inv_lift: ∀s. deliftable_sn (lhap s).
63 /3 width=3 by lstar_deliftable_sn, lhap1_inv_lift/
64 qed-.
65
66 lemma lhap_dsubst: ∀s. dsubstable_dx (lhap s).
67 /2 width=1/
68 qed.
69
70 theorem lhap_mono: ∀s. singlevalued … (lhap s).
71 /3 width=7 by lstar_singlevalued, lhap1_mono/
72 qed-.
73
74 theorem lhap_trans: ltransitive … lhap.
75 /2 width=3 by lstar_ltransitive/
76 qed-.
77
78 lemma lhap_step_dx: ∀s,M1,M. M1 ⓗ⇀*[s] M →
79                     ∀p,M2. M ⓗ⇀[p] M2 → M1 ⓗ⇀*[s@p::◊] M2.
80 #s #M1 #M #HM1 #p #M2 #HM2
81 @(lhap_trans … HM1) /2 width=1/
82 qed-.
83
84 lemma head_lsreds_lhap: ∀s,M1,M2. M1 ⇀*[s] M2 → is_head s → M1 ⓗ⇀*[s] M2.
85 #s #M1 #M2 #H @(lstar_ind_l ????????? H) -s -M1 //
86 #a #s #M1 #M #HM1 #_ #IHM2 * /3 width=3/
87 qed.
88
89 lemma lhap_inv_head: ∀s,M1,M2. M1 ⓗ⇀*[s] M2 → is_head s.
90 #s #M1 #M2 #H @(lstar_ind_l ????????? H) -s -M1 //
91 #p #s #M1 #M #HM1 #_ #IHM2
92 lapply (lhap1_inv_head … HM1) -HM1 /2 width=1/
93 qed-.
94
95 lemma lhap_inv_lsreds: ∀s,M1,M2. M1 ⓗ⇀*[s] M2 → M1 ⇀*[s] M2.
96 #s #M1 #M2 #H @(lstar_ind_l ????????? H) -s -M1 //
97 #p #s #M1 #M #HM1 #_ #IHM2
98 lapply (lhap1_inv_lsred … HM1) -HM1 /2 width=3/
99 qed-.