]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda/labelled_sequential_computation.ma
- lambda: some parts commented out, some refactoring
[helm.git] / matita / matita / contribs / lambda / labelled_sequential_computation.ma
index 1ec666483277ca5cab78dfd9c6a9b7ef6d9753ec..f305f90cb258b2aef1a1844ff38dbcd7a78afb68 100644 (file)
@@ -22,27 +22,44 @@ definition lsreds: pseq → relation term ≝ lstar … lsred.
 interpretation "labelled sequential computation"
    'SeqRedStar M s N = (lsreds s M N).
 
-notation "hvbox( M break â\87\80* [ term 46 s ] break term 46 N )"
+notation "hvbox( M break â\86¦* [ term 46 s ] break term 46 N )"
    non associative with precedence 45
    for @{ 'SeqRedStar $M $s $N }.
 
-lemma lsred_lsreds: ∀p,M1,M2. M1 ⇀[p] M2 → M1 ⇀*[p::◊] M2.
+lemma lsreds_step_rc: ∀p,M1,M2. M1 ↦[p] M2 → M1 ↦*[p::◊] M2.
 /2 width=1/
 qed.
 
-lemma lsreds_inv_nil: â\88\80s,M1,M2. M1 â\87\80*[s] M2 → ◊ = s → M1 = M2.
+lemma lsreds_inv_nil: â\88\80s,M1,M2. M1 â\86¦*[s] M2 → ◊ = s → M1 = M2.
 /2 width=5 by lstar_inv_nil/
 qed-.
 
-lemma lsreds_inv_cons: â\88\80s,M1,M2. M1 â\87\80*[s] M2 → ∀q,r. q::r = s →
-                       â\88\83â\88\83M. M1 â\87\80[q] M & M â\87\80*[r] M2.
+lemma lsreds_inv_cons: â\88\80s,M1,M2. M1 â\86¦*[s] M2 → ∀q,r. q::r = s →
+                       â\88\83â\88\83M. M1 â\86¦[q] M & M â\86¦*[r] M2.
 /2 width=3 by lstar_inv_cons/
 qed-.
 
-lemma lsreds_inv_lsred: ∀p,M1,M2. M1 ⇀*[p::◊] M2 → M1 ⇀[p] M2.
+lemma lsreds_inv_step_rc: ∀p,M1,M2. M1 ↦*[p::◊] M2 → M1 ↦[p] M2.
 /2 width=1 by lstar_inv_step/
 qed-.
 
+lemma lsreds_inv_pos: ∀s,M1,M2. M1 ↦*[s] M2 → 0 < |s| →
+                      ∃∃p,r,M. p::r = s & M1 ↦[p] M & M ↦*[r] M2.
+/2 width=1 by lstar_inv_pos/
+qed-.
+
+lemma lsred_compatible_rc: ho_compatible_rc lsreds.
+/3 width=1/
+qed.
+
+lemma lsred_compatible_sn: ho_compatible_sn lsreds.
+/3 width=1/
+qed.
+
+lemma lsred_compatible_dx: ho_compatible_dx lsreds.
+/3 width=1/
+qed.
+
 lemma lsreds_lift: ∀s. liftable (lsreds s).
 /2 width=1/
 qed.
@@ -59,13 +76,12 @@ theorem lsreds_mono: ∀s. singlevalued … (lsreds s).
 /3 width=7 by lstar_singlevalued, lsred_mono/
 qed-.
 
-theorem lsreds_trans: ∀s1,M1,M. M1 ⇀*[s1] M →
-                      ∀s2,M2. M ⇀*[s2] M2 → M1 ⇀*[s1@s2] M2.
-/2 width=3 by lstar_trans/
+theorem lsreds_trans: ltransitive … lsreds.
+/2 width=3 by lstar_ltransitive/
 qed-.
 
 (* Note: "|s|" should be unparetesized *)
-lemma lsreds_fwd_mult: â\88\80s,M1,M2. M1 â\87\80*[s] M2 → #{M2} ≤ #{M1} ^ (2 ^ (|s|)).
+lemma lsreds_fwd_mult: â\88\80s,M1,M2. M1 â\86¦*[s] M2 → #{M2} ≤ #{M1} ^ (2 ^ (|s|)).
 #s #M1 #M2 #H @(lstar_ind_l ????????? H) -s -M1 normalize //
 #p #s #M1 #M #HM1 #_ #IHM2
 lapply (lsred_fwd_mult … HM1) -p #HM1