]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/reducibility/ltpr.ma
- full commit for the transtive closure of ltpss!
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / reducibility / ltpr.ma
index 8009cbeb7ea1e6e94119fed123078c0dda4b309f..67986d336f87b2a39bb3bc63a5f2a9f6a8785d75 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
+include "basic_2/grammar/lenv_px.ma".
 include "basic_2/reducibility/tpr.ma".
 
 (* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************)
 
-inductive ltpr: relation lenv ≝
-| ltpr_stom: ltpr (⋆) (⋆)
-| ltpr_pair: ∀K1,K2,I,V1,V2.
-             ltpr K1 K2 → V1 ➡ V2 → ltpr (K1. ⓑ{I} V1) (K2. ⓑ{I} V2) (*ⓑ*)
-.
+definition ltpr: relation lenv ≝ lpx tpr.
 
 interpretation
   "context-free parallel reduction (environment)"
@@ -28,62 +25,30 @@ interpretation
 
 (* Basic properties *********************************************************)
 
-lemma ltpr_refl: ∀L:lenv. L ➡ L.
-#L elim L -L // /2 width=1/
-qed.
+lemma ltpr_refl: reflexive … ltpr.
+/2 width=1/ qed.
 
 (* Basic inversion lemmas ***************************************************)
 
-fact ltpr_inv_atom1_aux: ∀L1,L2. L1 ➡ L2 → L1 = ⋆ → L2 = ⋆.
-#L1 #L2 * -L1 -L2
-[ //
-| #K1 #K2 #I #V1 #V2 #_ #_ #H destruct
-]
-qed.
-
 (* Basic_1: was: wcpr0_gen_sort *)
 lemma ltpr_inv_atom1: ∀L2. ⋆ ➡ L2 → L2 = ⋆.
-/2 width=3/ qed-.
-
-fact ltpr_inv_pair1_aux: ∀L1,L2. L1 ➡ L2 → ∀K1,I,V1. L1 = K1. ⓑ{I} V1 →
-                         ∃∃K2,V2. K1 ➡ K2 & V1 ➡ V2 & L2 = K2. ⓑ{I} V2.
-#L1 #L2 * -L1 -L2
-[ #K1 #I #V1 #H destruct
-| #K1 #K2 #I #V1 #V2 #HK12 #HV12 #L #J #W #H destruct /2 width=5/
-]
-qed.
+/2 width=2 by lpx_inv_atom1/ qed-.
 
 (* Basic_1: was: wcpr0_gen_head *)
 lemma ltpr_inv_pair1: ∀K1,I,V1,L2. K1. ⓑ{I} V1 ➡ L2 →
                       ∃∃K2,V2. K1 ➡ K2 & V1 ➡ V2 & L2 = K2. ⓑ{I} V2.
-/2 width=3/ qed-.
-
-fact ltpr_inv_atom2_aux: ∀L1,L2. L1 ➡ L2 → L2 = ⋆ → L1 = ⋆.
-#L1 #L2 * -L1 -L2
-[ //
-| #K1 #K2 #I #V1 #V2 #_ #_ #H destruct
-]
-qed.
+/2 width=1 by lpx_inv_pair1/ qed-.
 
 lemma ltpr_inv_atom2: ∀L1. L1 ➡ ⋆ → L1 = ⋆.
-/2 width=3/ qed-.
-
-fact ltpr_inv_pair2_aux: ∀L1,L2. L1 ➡ L2 → ∀K2,I,V2. L2 = K2. ⓑ{I} V2 →
-                         ∃∃K1,V1. K1 ➡ K2 & V1 ➡ V2 & L1 = K1. ⓑ{I} V1.
-#L1 #L2 * -L1 -L2
-[ #K2 #I #V2 #H destruct
-| #K1 #K2 #I #V1 #V2 #HK12 #HV12 #K #J #W #H destruct /2 width=5/
-]
-qed.
+/2 width=2 by lpx_inv_atom2/ qed-.
 
 lemma ltpr_inv_pair2: ∀L1,K2,I,V2. L1 ➡ K2. ⓑ{I} V2 →
                       ∃∃K1,V1. K1 ➡ K2 & V1 ➡ V2 & L1 = K1. ⓑ{I} V1.
-/2 width=3/ qed-.
+/2 width=1 by lpx_inv_pair2/ qed-.
 
 (* Basic forward lemmas *****************************************************)
 
 lemma ltpr_fwd_length: ∀L1,L2. L1 ➡ L2 → |L1| = |L2|.
-#L1 #L2 #H elim H -L1 -L2 normalize //
-qed-.
+/2 width=2 by lpx_fwd_length/ qed-.
 
 (* Basic_1: removed theorems 2: wcpr0_getl wcpr0_getl_back *)