]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/lambda/terms/sequential_computation.ma
717a7602ab6f82e72e03de5ed8b6e1a55daf9423
[helm.git] / matita / matita / lib / lambda / terms / sequential_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 "lambda/terms/parallel_computation.ma".
16
17 (* SEQUENTIAL COMPUTATION (MULTISTEP) ***************************************)
18
19 definition sreds: relation term ≝ star … sred.
20
21 interpretation "sequential computation"
22    'SeqRedStar M N = (sreds M N).
23
24 lemma sreds_refl: reflexive … sreds.
25 //
26 qed.
27
28 lemma sreds_step_sn: ∀M1,M. M1 ↦ M → ∀M2. M ↦* M2 → M1 ↦* M2.
29 /2 width=3/
30 qed-.
31
32 lemma sreds_step_dx: ∀M1,M. M1 ↦* M → ∀M2. M ↦ M2 → M1 ↦* M2.
33 /2 width=3/
34 qed-.
35
36 lemma sreds_step_rc: ∀M1,M2. M1 ↦ M2 → M1 ↦* M2.
37 /2 width=1/
38 qed.
39
40 lemma lsred_compatible_abst: compatible_abst sreds.
41 /3 width=1/
42 qed.
43
44 lemma sreds_compatible_sn: compatible_sn sreds.
45 /3 width=1/
46 qed.
47
48 lemma sreds_compatible_dx: compatible_dx sreds.
49 /3 width=1/
50 qed.
51
52 lemma sreds_compatible_appl: compatible_appl sreds.
53 /3 width=3/
54 qed.
55
56 lemma sreds_lift: liftable sreds.
57 /2 width=1/
58 qed.
59
60 lemma sreds_inv_lift: deliftable_sn sreds.
61 /3 width=3 by star_deliftable_sn, sred_inv_lift/
62 qed-.
63
64 lemma sreds_dsubst: dsubstable_dx sreds.
65 /2 width=1/
66 qed.
67
68 theorem sreds_trans: transitive … sreds.
69 /2 width=3 by trans_star/
70 qed-.
71
72 (* Note: the substitution should be unparentesized *) 
73 lemma sreds_compatible_beta: ∀B1,B2. B1 ↦* B2 → ∀A1,A2. A1 ↦* A2 →
74                              @B1.𝛌.A1 ↦* ([↙B2] A2).
75 #B1 #B2 #HB12 #A1 #A2 #HA12
76 @(sreds_trans … (@B2.𝛌.A1)) /2 width=1/ -B1
77 @(sreds_step_dx … (@B2.𝛌.A2)) // /3 width=1/
78 qed.
79
80 theorem sreds_preds: ∀M1,M2. M1 ↦* M2 → M1 ⤇* M2.
81 #M1 #M2 #H @(star_ind_l … M1 H) -M1 //
82 #M1 #M #HM1 #_ #IHM2
83 @(preds_step_sn … IHM2) -M2 /2 width=2/
84 qed.
85
86 lemma pred_sreds: ∀M1,M2. M1 ⤇ M2 → M1 ↦* M2.
87 #M1 #M2 #H elim H -M1 -M2 // /2 width=1/
88 qed-.
89
90 theorem preds_sreds: ∀M1,M2. M1 ⤇* M2 → M1 ↦* M2.
91 #M1 #M2 #H elim H -M2 //
92 #M #M2 #_ #HM2 #HM1
93 lapply (pred_sreds … HM2) -HM2 #HM2
94 @(sreds_trans … HM1 … HM2)
95 qed-.
96
97 theorem sreds_conf: ∀M0,M1. M0 ↦* M1 → ∀M2. M0 ↦* M2 →
98                     ∃∃M. M1 ↦* M & M2 ↦* M.
99 #M0 #M1 #HM01 #M2 #HM02
100 lapply (sreds_preds … HM01) #HM01
101 lapply (sreds_preds … HM02) #HM02
102 elim (preds_conf … HM01 … HM02) -M0 #M #HM1 #HM2
103 lapply (preds_sreds … HM1) -HM1
104 lapply (preds_sreds … HM2) -HM2 /2 width=3/
105 qed-.