]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/grammar/lenv_length.ma
lambda_delta: partial commit ...
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / grammar / lenv_length.ma
index 780edc161d69ca33af45095714622341b1eaf3cb..faf6de02d1dfe6e55d1253c97718637c0e1fdb4f 100644 (file)
@@ -22,3 +22,31 @@ let rec length L ≝ match L with
 ].
 
 interpretation "length (local environment)" 'card L = (length L).
+
+(* Basic inversion lemmas ***************************************************)
+
+lemma length_inv_zero_dx: ∀L. |L| = 0 → L = ⋆.
+* // #L #I #V normalize <plus_n_Sm #H destruct
+qed-.
+
+lemma length_inv_zero_sn: ∀L. 0 = |L| → L = ⋆.
+* // #L #I #V normalize <plus_n_Sm #H destruct
+qed-.
+
+lemma length_inv_pos_dx: ∀d,L. |L| = d + 1 →
+                         ∃∃I,K,V. |K| = d & L = K. ⓑ{I}V.
+#d *
+[ normalize <plus_n_Sm #H destruct
+| #K #I #V #H
+  lapply (injective_plus_l … H) -H /2 width=5/
+]
+qed-.
+
+lemma length_inv_pos_sn: ∀d,L. d + 1 = |L| →
+                         ∃∃I,K,V. d = |K| & L = K. ⓑ{I}V.
+#d *
+[ normalize <plus_n_Sm #H destruct
+| #K #I #V #H
+  lapply (injective_plus_l … H) -H /2 width=5/
+]
+qed-.