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 "static_2/syntax/lenv.ma".
17 (* FOLD FOR RESTRICTED CLOSURES *********************************************)
19 rec definition fold L T on L ≝ match L with
21 | LBind L I ⇒ match I with
22 [ BUnit _ ⇒ fold L (-ⓛ⋆𝟎.T)
23 | BPair I V ⇒ fold L (-ⓑ[I]V.T)
27 interpretation "fold (restricted closure)" 'plus L T = (fold L T).
29 (* Basic properties *********************************************************)
31 lemma fold_atom: ∀T. ⋆ + T = T.
34 lemma fold_unit: ∀I,L,T. L.ⓤ[I]+T = L+(-ⓛ⋆𝟎.T).
37 lemma fold_pair: ∀I,L,V,T. (L.ⓑ[I]V)+T = L+(-ⓑ[I]V.T).