]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/computation/lcprs_lcprs.ma
- subject equivalence for atomic arity assignment completed!
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / computation / lcprs_lcprs.ma
index 344eae1542957e2f18e49047d560c50f304ebf2c..64851e801f9a78694e1fdc6a627d6d577856e094 100644 (file)
 (*        v         GNU General Public License Version 2                  *)
 (*                                                                        *)
 (**************************************************************************)
-(*
+
 include "basic_2/reducibility/lcpr_lcpr.ma".
-*)
 include "basic_2/computation/lcprs_cprs.ma".
 
 (* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS *************)
 
 (* Advanced properties ******************************************************)
 
-axiom lcprs_strip: ∀L,L1. L ⊢ ➡* L1 → ∀L2. L ⊢ ➡ L2 →
+lemma lcprs_strip: ∀L,L1. L ⊢ ➡* L1 → ∀L2. L ⊢ ➡ L2 →
                    ∃∃L0. L1 ⊢ ➡ L0 & L2 ⊢ ➡* L0.
-(*
 /3 width=3/ qed.
-*)
+
 (* Main properties **********************************************************)
 
+theorem lcprs_conf: ∀L,L1. L ⊢ ➡* L1 → ∀L2. L ⊢ ➡* L2 →
+                    ∃∃L0. L1 ⊢ ➡* L0 & L2 ⊢ ➡* L0.
+/3 width=3/ qed.
+
 theorem lcprs_trans: ∀L1,L. L1 ⊢ ➡* L → ∀L2. L ⊢ ➡* L2 → L1 ⊢ ➡* L2.
 /2 width=3/ qed.