(* global environments *)
inductive genv: Type[0] ≝
| GAtom: genv (* empty *)
-| GBind: genv → bind2 → term → genv (* binding construction *)
+| GPair: genv → bind2 → term → genv (* binary binding construction *)
.
interpretation "sort (global environment)"
'Star = (GAtom).
interpretation "global environment binding construction (binary)"
- 'DxBind2 G I T = (GBind G I T).
+ 'DxBind2 G I T = (GPair G I T).
interpretation "abbreviation (global environment)"
- 'DxAbbr G T = (GBind G Abbr T).
+ 'DxAbbr G T = (GPair G Abbr T).
interpretation "abstraction (global environment)"
- 'DxAbst G T = (GBind G Abst T).
+ 'DxAbst G T = (GPair G Abst T).
(* Basic properties *********************************************************)
rec definition glength G on G ≝ match G with
[ GAtom ⇒ 0
-| GBind G _ _ ⇒ ↑(glength G)
+| GPair G _ _ ⇒ ↑(glength G)
].
interpretation "length (global environment)"
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basic_2/syntax/term_weight.ma".
+include "basic_2/syntax/genv.ma".
+
+(* WEIGHT OF A GLOBAL ENVIRONMENT *******************************************)
+
+rec definition gw G ≝ match G with
+[ GAtom ⇒ 0
+| GPair G I T ⇒ gw G + ♯{T}
+].
+
+interpretation "weight (global environment)" 'Weight G = (gw G).
+
+(* Basic properties *********************************************************)
+
+lemma gw_pair: ∀I,G,T. ♯{G} < ♯{G.ⓑ{I}T}.
+normalize /2 width=1 by monotonic_le_plus_r/ qed.
]
[ { "global environments" * } {
[ [ "" ] "genv_length" + "( |?| )" * ]
+ [ [ "" ] "genv_weight" + "( ♯{?} )" * ]
[ [ "" ] "genv" * ]
}
]