X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fdelayed_updating%2Fsyntax%2Fprototerm_constructors.ma;h=35a79b0399624d79444c8ba3243003e8b1eb5cb6;hb=80ecd5486c6013f6c297173f41432fd1d93814ef;hp=fd86ef975a39b862aaf05ac99dc5efd298fb10ed;hpb=8d8863982ca95225551e9659ed431db046c34e81;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 fd86ef975..35a79b039 100644 --- a/matita/matita/contribs/lambdadelta/delayed_updating/syntax/prototerm_constructors.ma +++ b/matita/matita/contribs/lambdadelta/delayed_updating/syntax/prototerm_constructors.ma @@ -55,8 +55,21 @@ interpretation "application (prototerm)" 'At u t = (prototerm_node_2 label_S label_A u t). +(* Basic constructions *******************************************************) + +lemma in_comp_iref (t) (q) (n): + q ϵ t → 𝗱n◗𝗺◗q ϵ 𝛗n.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 +/2 width=3 by ex2_intro/ +qed-. +(* lemma prototerm_in_root_inv_lcons_oref: ∀p,l,n. l◗p ϵ ▵#n → ∧∧ 𝗱n = l & 𝐞 = p. @@ -98,3 +111,4 @@ lemma prototerm_in_root_inv_lcons_appl: