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 "terms/lift.ma".
17 (* SIZE *********************************************************************)
19 (* Note: this gives the number of abstractions and applications in M *)
20 let rec size M on M ≝ match M with
23 | Appl B A ⇒ (size B) + (size A) + 1
26 interpretation "term size"
29 lemma size_lift: ∀h,M,d. |↑[d, h] M| = |M|.
30 #h #M elim M -M normalize //