1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "basic_2/grammar/lenv.ma".
17 (* LENGTH OF A LOCAL ENVIRONMENT ********************************************)
19 let rec length L ≝ match L with
21 | LPair L _ _ ⇒ length L + 1
24 interpretation "length (local environment)" 'card L = (length L).
26 (* Basic inversion lemmas ***************************************************)
28 lemma length_inv_zero_dx: ∀L. |L| = 0 → L = ⋆.
29 * // #L #I #V normalize <plus_n_Sm #H destruct
32 lemma length_inv_zero_sn: ∀L. 0 = |L| → L = ⋆.
33 * // #L #I #V normalize <plus_n_Sm #H destruct
36 lemma length_inv_pos_dx: ∀d,L. |L| = d + 1 →
37 ∃∃I,K,V. |K| = d & L = K. ⓑ{I}V.
39 [ normalize <plus_n_Sm #H destruct
41 lapply (injective_plus_l … H) -H /2 width=5/
45 lemma length_inv_pos_sn: ∀d,L. d + 1 = |L| →
46 ∃∃I,K,V. d = |K| & L = K. ⓑ{I}V.
48 [ normalize <plus_n_Sm #H destruct
50 lapply (injective_plus_l … H) -H /2 width=5/