]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/lpxs.ma
reordering abstract computation properties and saturation conditions
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / lpxs.ma
index 35a8cb4829f5db8de2943380739bff31c1813bc8..ee557e4385fa4038a8df142dab4d8079c0f79c8a 100644 (file)
@@ -43,19 +43,19 @@ qed-.
 (* Basic properties *********************************************************)
 
 lemma lprs_lpxs: ∀h,g,G,L1,L2. ⦃G, L1⦄ ⊢ ➡* L2 → ⦃G, L1⦄ ⊢ ➡*[h, g] L2.
-/3 width=3/ qed.
+/3 width=3 by lpr_lpx, monotonic_TC/ qed.
 
 lemma lpx_lpxs: ∀h,g,G,L1,L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 → ⦃G, L1⦄ ⊢ ➡*[h, g] L2.
-/2 width=1/ qed.
+/2 width=1 by inj/ qed.
 
 lemma lpxs_refl: ∀h,g,G,L. ⦃G, L⦄ ⊢ ➡*[h, g] L.
-/2 width=1/ qed.
+/2 width=1 by lprs_lpxs/ qed.
 
 lemma lpxs_strap1: ∀h,g,G,L1,L,L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L → ⦃G, L⦄ ⊢ ➡[h, g] L2 → ⦃G, L1⦄ ⊢ ➡*[h, g] L2.
-/2 width=3/ qed.
+/2 width=3 by step/ qed.
 
 lemma lpxs_strap2: ∀h,g,G,L1,L,L2. ⦃G, L1⦄ ⊢ ➡[h, g] L → ⦃G, L⦄ ⊢ ➡*[h, g] L2 → ⦃G, L1⦄ ⊢ ➡*[h, g] L2.
-/2 width=3/ qed.
+/2 width=3 by TC_strap/ qed.
 
 lemma lpxs_pair_refl: ∀h,g,G,L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → ∀I,V. ⦃G, L1.ⓑ{I}V⦄ ⊢ ➡*[h, g] L2.ⓑ{I}V.
 /2 width=1 by TC_lpx_sn_pair_refl/ qed.