]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/grammar/lenv_weight.ma
- one axiom removed from sd
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / grammar / lenv_weight.ma
index bc8ffce87e77e97eb79f151d4c0b69d3940ffc3d..59e2e6172bc24a83e79d60cdb537e2e403087cae 100644 (file)
@@ -19,9 +19,21 @@ include "basic_2/grammar/lenv.ma".
 
 let rec lw L ≝ match L with
 [ LAtom       ⇒ 0
-| LPair L _ V ⇒ lw L + #[V]
+| LPair L _ V ⇒ lw L + #{V}
 ].
 
 interpretation "weight (local environment)" 'Weight L = (lw L).
 
-(* Basic_1: removed theorems 2: clt_cong clt_head *)
+(* Basic properties *********************************************************)
+
+lemma lw_pair: ∀I,L,V. #{L} < #{(L.ⓑ{I}V)}.
+/3 width=1/ qed.
+
+(* Basic eliminators ********************************************************)
+
+axiom lw_ind: ∀R:predicate lenv.
+              (∀L2. (∀L1. #{L1} < #{L2} → R L1) → R L2) →
+              ∀L. R L.
+
+(* Basic_1: removed theorems 2: clt_cong clt_head clt_thead *)
+(* Basic_1: note: clt_thead should be renamed clt_ctail *)