]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/Basic_2/grammar/lsubs.ma
- "grammar" component updated to new syntax ...
[helm.git] / matita / matita / contribs / lambda_delta / Basic_2 / grammar / lsubs.ma
index cf90a6134590c41c8ffc50b511cf5123e4fa3e49..1f750064ebd139eb0c9d472ef31dbda24c79e99e 100644 (file)
@@ -27,7 +27,9 @@ inductive lsubs: nat → nat → relation lenv ≝
               lsubs d e L1 L2 → lsubs (d + 1) e (L1. 𝕓{I1} V1) (L2. 𝕓{I2} V2)
 .
 
-interpretation "local environment refinement (substitution)" 'SubEq L1 d e L2 = (lsubs d e L1 L2).
+interpretation
+  "local environment refinement (substitution)"
+  'SubEq L1 d e L2 = (lsubs d e L1 L2).
 
 definition lsubs_conf: ∀S. (lenv → relation S) → Prop ≝ λS,R.
                        ∀L1,s1,s2. R L1 s1 s2 →
@@ -36,29 +38,29 @@ definition lsubs_conf: ∀S. (lenv → relation S) → Prop ≝ λS,R.
 (* Basic properties *********************************************************)
 
 lemma TC_lsubs_conf: ∀S,R. lsubs_conf S R → lsubs_conf S (λL. (TC … (R L))).
-#S #R #HR #L1 #s1 #s2 #H elim H -s2
+#S #R #HR #L1 #s1 #s2 #H elim H -s2
 [ /3 width=5/
 | #s #s2 #_ #Hs2 #IHs1 #L2 #d #e #HL12
-  lapply (HR … Hs2 … HL12) -HR Hs2 HL12 /3/
+  lapply (HR … Hs2 … HL12) -HR -Hs2 -HL12 /3 width=3/
 ]
 qed.
 
 lemma lsubs_bind_eq: ∀L1,L2,e. L1 [0, e] ≼ L2 → ∀I,V.
                      L1. 𝕓{I} V [0, e + 1] ≼ L2.𝕓{I} V.
-#L1 #L2 #e #HL12 #I #V elim I -I /2/
+#L1 #L2 #e #HL12 #I #V elim I -I /2 width=1/
 qed.
 
 lemma lsubs_refl: ∀d,e,L. L [d, e] ≼ L.
 #d elim d -d
-[ #e elim e -e // #e #IHe #L elim L -L /2/
-| #d #IHd #e #L elim L -L /2/
+[ #e elim e -e // #e #IHe #L elim L -L // /2 width=1/
+| #d #IHd #e #L elim L -L // /2 width=1/
 ]
 qed.
 
 lemma lsubs_skip_lt: ∀L1,L2,d,e. L1 [d - 1, e] ≼ L2 → 0 < d →
                      ∀I1,I2,V1,V2. L1. 𝕓{I1} V1 [d, e] ≼ L2. 𝕓{I2} V2.
 
-#L1 #L2 #d #e #HL12 #Hd >(plus_minus_m_m d 1) /2/
+#L1 #L2 #d #e #HL12 #Hd >(plus_minus_m_m d 1) // /2 width=1/
 qed.
 
 (* Basic inversion lemmas ***************************************************)
@@ -67,14 +69,14 @@ qed.
 
 fact lsubs_fwd_length_full1_aux: ∀L1,L2,d,e. L1 [d, e] ≼ L2 →
                                  d = 0 → e = |L1| → |L1| ≤ |L2|.
-#L1 #L2 #d #e #H elim H -H L1 L2 d e; normalize
+#L1 #L2 #d #e #H elim H -L1 -L2 -d -e normalize
 [ //
-| /2/
-| /3/
-| /3/
+| /2 width=1/
+| /3 width=1/
+| /3 width=1/
 | #L1 #L2 #_ #_ #_ #_ #d #e #_ #_ #H
   elim (plus_S_eq_O_false … H)
-]   
+]
 qed.
 
 lemma lsubs_fwd_length_full1: ∀L1,L2. L1 [0, |L1|] ≼ L2 → |L1| ≤ |L2|.
@@ -82,14 +84,14 @@ lemma lsubs_fwd_length_full1: ∀L1,L2. L1 [0, |L1|] ≼ L2 → |L1| ≤ |L2|.
 
 fact lsubs_fwd_length_full2_aux: ∀L1,L2,d,e. L1 [d, e] ≼ L2 →
                                  d = 0 → e = |L2| → |L2| ≤ |L1|.
-#L1 #L2 #d #e #H elim H -H L1 L2 d e; normalize
+#L1 #L2 #d #e #H elim H -L1 -L2 -d -e normalize
 [ //
-| /2/
-| /3/
-| /3/
+| /2 width=1/
+| /3 width=1/
+| /3 width=1/
 | #L1 #L2 #_ #_ #_ #_ #d #e #_ #_ #H
   elim (plus_S_eq_O_false … H)
-]   
+]
 qed.
 
 lemma lsubs_fwd_length_full2: ∀L1,L2. L1 [0, |L2|] ≼ L2 → |L2| ≤ |L1|.