]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/grammar/lpx_sn_tc.ma
- lambdadelta: first commutation property on lazy equivalence for
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / grammar / lpx_sn_tc.ma
index fa1234d1a79cd6902ea18dc3e7b65fda51771a2e..2919d9d103b73f716e4f1c92c4110ad3f9bf7255 100644 (file)
@@ -24,7 +24,7 @@ lemma TC_lpx_sn_pair_refl: ∀R. (∀L. reflexive … (R L)) →
 #R #HR #L1 #L2 #H @(TC_star_ind … L2 H) -L2
 [ /2 width=1 by lpx_sn_refl/
 | /3 width=1 by TC_reflexive, lpx_sn_refl/
-| /3 width=5/
+| /3 width=5 by lpx_sn_pair, step/
 ]
 qed-.
 
@@ -41,17 +41,16 @@ qed-.
 lemma lpx_sn_LTC_TC_lpx_sn: ∀R. (∀L. reflexive … (R L)) →
                             ∀L1,L2. lpx_sn (LTC … R) L1 L2 →
                             TC … (lpx_sn R) L1 L2.
-#R #HR #L1 #L2 #H elim H -L1 -L2 /2 width=1/
-/2 width=1 by TC_lpx_sn_pair/
+#R #HR #L1 #L2 #H elim H -L1 -L2
+/2 width=1 by TC_lpx_sn_pair, lpx_sn_atom, inj/
 qed-.
 
 (* Inversion lemmas on transitive closure ***********************************)
 
 lemma TC_lpx_sn_inv_atom2: ∀R,L1. TC … (lpx_sn R) L1 (⋆) → L1 = ⋆.
 #R #L1 #H @(TC_ind_dx … L1 H) -L1
-[ #L1 #H lapply (lpx_sn_inv_atom2 … H) -H //
-| #L1 #L #HL1 #_ #IHL2 destruct
-  lapply (lpx_sn_inv_atom2 … HL1) -HL1 //
+[ /2 width=2 by lpx_sn_inv_atom2/
+| #L1 #L #HL1 #_ #IHL2 destruct /2 width=2 by lpx_sn_inv_atom2/
 ]
 qed-.
 
@@ -59,10 +58,10 @@ lemma TC_lpx_sn_inv_pair2: ∀R. s_rs_trans … R (lpx_sn R) →
                            ∀I,L1,K2,V2. TC  … (lpx_sn R) L1 (K2.ⓑ{I}V2) →
                            ∃∃K1,V1. TC … (lpx_sn R) K1 K2 & LTC … R K1 V1 V2 & L1 = K1. ⓑ{I} V1.
 #R #HR #I #L1 #K2 #V2 #H @(TC_ind_dx … L1 H) -L1
-[ #L1 #H elim (lpx_sn_inv_pair2 … H) -H /3 width=5/
+[ #L1 #H elim (lpx_sn_inv_pair2 … H) -H /3 width=5 by inj, ex3_2_intro/
 | #L1 #L #HL1 #_ * #K #V #HK2 #HV2 #H destruct
   elim (lpx_sn_inv_pair2 … HL1) -HL1 #K1 #V1 #HK1 #HV1 #H destruct
-  lapply (HR … HV2 … HK1) -HR -HV2 #HV2 /3 width=5/
+  lapply (HR … HV2 … HK1) -HR -HV2 /3 width=5 by TC_strap, ex3_2_intro/
 ]
 qed-.
 
@@ -78,15 +77,14 @@ lemma TC_lpx_sn_ind: ∀R. s_rs_trans … R (lpx_sn R) →
 [ #X #H >(TC_lpx_sn_inv_atom2 … H) -X //
 | #L2 #I #V2 #IHL2 #X #H
   elim (TC_lpx_sn_inv_pair2 … H) // -H -HR
-  #L1 #V1 #HL12 #HV12 #H destruct /3 width=1/
+  #L1 #V1 #HL12 #HV12 #H destruct /3 width=1 by/
 ]
 qed-.
 
 lemma TC_lpx_sn_inv_atom1: ∀R,L2. TC … (lpx_sn R) (⋆) L2 → L2 = ⋆.
 #R #L2 #H elim H -L2
-[ #L2 #H lapply (lpx_sn_inv_atom1 … H) -H //
-| #L #L2 #_ #HL2 #IHL1 destruct
-  lapply (lpx_sn_inv_atom1 … HL2) -HL2 //
+[ /2 width=2 by lpx_sn_inv_atom1/
+| #L #L2 #_ #HL2 #IHL1 destruct /2 width=2 by lpx_sn_inv_atom1/
 ]
 qed-.
 
@@ -96,7 +94,7 @@ fact TC_lpx_sn_inv_pair1_aux: ∀R. s_rs_trans … R (lpx_sn R) →
                               ∃∃K2,V2. TC … (lpx_sn R) K1 K2 & LTC … R K1 V1 V2 & L2 = K2. ⓑ{I} V2.
 #R #HR #L1 #L2 #H @(TC_lpx_sn_ind … H) // -HR -L1 -L2
 [ #J #K #W #H destruct
-| #I #L1 #L2 #V1 #V2 #HL12 #HV12 #_ #J #K #W #H destruct /2 width=5/
+| #I #L1 #L2 #V1 #V2 #HL12 #HV12 #_ #J #K #W #H destruct /2 width=5 by ex3_2_intro/
 ]
 qed-.
 
@@ -108,8 +106,7 @@ lemma TC_lpx_sn_inv_pair1: ∀R. s_rs_trans … R (lpx_sn R) →
 lemma TC_lpx_sn_inv_lpx_sn_LTC: ∀R. s_rs_trans … R (lpx_sn R) →
                                 ∀L1,L2. TC … (lpx_sn R) L1 L2 →
                                 lpx_sn (LTC … R) L1 L2.
-#R #HR #L1 #L2 #H @(TC_lpx_sn_ind … H) // -HR -L1 -L2 /2 width=1/
-qed-.
+/3 width=4 by TC_lpx_sn_ind, lpx_sn_pair/ qed-.
 
 (* Forward lemmas on transitive closure *************************************)