X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fdelayed_updating%2Fsyntax%2Fprototerm_constructors.ma;h=5bb66e16ab112f1ab94db7db70cba521bc150131;hb=9e31ac1f3f868349154b0ce2e550e2476aaf6a30;hp=f026c769e0cd3442db5b644f274f3252b37dfcf5;hpb=3bf7a0b4185dbffe5b822c907956acdbe2d1c559;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/syntax/prototerm_constructors.ma b/matita/matita/contribs/lambdadelta/delayed_updating/syntax/prototerm_constructors.ma index f026c769e..5bb66e16a 100644 --- a/matita/matita/contribs/lambdadelta/delayed_updating/syntax/prototerm_constructors.ma +++ b/matita/matita/contribs/lambdadelta/delayed_updating/syntax/prototerm_constructors.ma @@ -41,11 +41,11 @@ interpretation interpretation "outer variable reference by depth (prototerm)" - 'Hash n = (prototerm_node_0 (label_d n)). + 'Hash k = (prototerm_node_0 (label_d k)). interpretation "inner variable reference by depth (prototerm)" - 'Tau n t = (prototerm_node_1_2 (label_d n) label_m t). + 'Tau k t = (prototerm_node_1_2 (label_d k) label_m t). interpretation "name-free functional abstraction (prototerm)" @@ -57,18 +57,19 @@ interpretation (* Basic constructions *******************************************************) -lemma in_comp_iref (t) (q) (n): - q ϵ t → 𝗱n◗𝗺◗q ϵ 𝛕n.t. +lemma in_comp_iref (t) (q) (k): + q ϵ t → 𝗱k◗𝗺◗q ϵ 𝛕k.t. /2 width=3 by ex2_intro/ qed. (* Basic inversions *********************************************************) -lemma in_comp_inv_iref (t) (p) (n): - p ϵ 𝛕n.t → - ∃∃q. 𝗱n◗𝗺◗q = p & q ϵ t. -#t #p #n * #q #Hq #Hp +lemma in_comp_inv_iref (t) (p) (k): + p ϵ 𝛕k.t → + ∃∃q. 𝗱k◗𝗺◗q = p & q ϵ t. +#t #p #k * #q #Hq #Hp /2 width=3 by ex2_intro/ qed-. + (* COMMENT lemma prototerm_in_root_inv_lcons_oref: ∀p,l,n. l◗p ϵ ▵#n →