]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/unfold/ltpss_sn.ma
- parallel substitution reaxiomatized
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / unfold / ltpss_sn.ma
index 0d13a5a3f5a5f697541280d7e20b6580f0b01ced..577c1506d96c67d3e3cec39ded7265c7e8189726 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
+notation "hvbox( T1 break ⊢ ▶ * [ term 46 d , break term 46 e ] break term 46 T2 )"
+   non associative with precedence 45
+   for @{ 'PSubstStarSn $T1 $d $e $T2 }.
+
 include "basic_2/unfold/tpss.ma".
 
 (* SN PARALLEL UNFOLD ON LOCAL ENVIRONMENTS *********************************)
@@ -206,7 +210,7 @@ lemma ltpss_sn_weak: ∀L1,L2,d1,e1. L1 ⊢ ▶* [d1, e1] L2 →
 ]
 qed.
 
-lemma ltpss_sn_weak_all: ∀L1,L2,d,e. L1 ⊢ ▶* [d, e] L2 → L1 ⊢ ▶* [0, |L1|] L2.
+lemma ltpss_sn_weak_full: ∀L1,L2,d,e. L1 ⊢ ▶* [d, e] L2 → L1 ⊢ ▶* [0, |L1|] L2.
 #L1 #L2 #d #e #H elim H -L1 -L2 -d -e
 // /3 width=2/ /3 width=3/
 qed.
@@ -237,7 +241,7 @@ lemma ltpss_sn_append_ge: ∀K1,K2,d,e. K1 ⊢ ▶* [d, e] K2 →
                           L1 @@ K1 ⊢ ▶* [d, e] L2 @@ K2.
 #K1 #K2 #d #e #H elim H -K1 -K2 -d -e
 [ #d #e #L1 #L2 <minus_n_O //
-| #K #I #V #L1 #L2 #_ #H 
+| #K #I #V #L1 #L2 #_ #H
   lapply (le_n_O_to_eq … H) -H normalize <plus_n_Sm #H destruct
 | #K1 #K2 #I #V1 #V2 #e #_ #_ #_ #L1 #L2 #_ #H
   lapply (le_n_O_to_eq … H) -H normalize <plus_n_Sm #H destruct