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/notation/functions/weight_2.ma".
16 include "static_2/syntax/lenv_weight.ma".
18 (* WEIGHT OF A RESTRICTED CLOSURE *******************************************)
20 definition rfw: lenv → term → ? ≝ λL,T. ♯❨L❩ + ♯❨T❩.
22 interpretation "weight (restricted closure)" 'Weight L T = (rfw L T).
24 (* Basic properties *********************************************************)
26 lemma rfw_unfold (L) (T): ♯❨L❩ + ♯❨T❩ = ♯❨L,T❩.
29 (* Basic_1: was: flt_shift *)
30 lemma rfw_shift: ∀p,I,K,V,T. ♯❨K.ⓑ[I]V,T❩ < ♯❨K,ⓑ[p,I]V.T❩.
31 /2 width=1 by plt_plus_bi_sn/ qed.
33 lemma rfw_clear: ∀p,I1,I2,K,V,T. ♯❨K.ⓤ[I1],T❩ < ♯❨K,ⓑ[p,I2]V.T❩.
34 /2 width=1 by plt_plus_bi_sn/ qed.
36 lemma rfw_tpair_sn: ∀I,L,V,T. ♯❨L,V❩ < ♯❨L,②[I]V.T❩.
37 /2 width=1 by plt_plus_bi_sn/ qed.
39 lemma rfw_tpair_dx: ∀I,L,V,T. ♯❨L,T❩ < ♯❨L,②[I]V.T❩.
40 /2 width=1 by plt_plus_bi_sn/ qed.
42 lemma rfw_lpair_sn: ∀I,L,V,T. ♯❨L,V❩ < ♯❨L.ⓑ[I]V,T❩.
45 lemma rfw_lpair_dx: ∀I,L,V,T. ♯❨L,T❩ < ♯❨L.ⓑ[I]V,T❩.
46 /2 width=1 by plt_plus_bi_dx/ qed.
48 (* Basic_1: removed theorems 7:
49 flt_thead_sx flt_thead_dx flt_trans
50 flt_arith0 flt_arith1 flt_arith2 flt_wf_ind
52 (* Basic_1: removed local theorems 1: q_ind *)