X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda%2Fterm.ma;h=660830d2784222dac9bc681e15b5d144aaf24edc;hb=2199f327081f49b21bdcd23d702b5e07ea4f58ce;hp=7710c05fb867084e036bee357182125af1d2e452;hpb=2b51cf74b9a5f37d0f91780ceae4b8f4d0ee38a1;p=helm.git diff --git a/matita/matita/contribs/lambda/term.ma b/matita/matita/contribs/lambda/term.ma index 7710c05fb..660830d27 100644 --- a/matita/matita/contribs/lambda/term.ma +++ b/matita/matita/contribs/lambda/term.ma @@ -18,7 +18,7 @@ include "preamble.ma". (* TERM STRUCTURE ***********************************************************) -(* Policy: term metavariables: A, B, C, D +(* Policy: term metavariables: A, B, C, D, M, N de Bruijn indexes : i, j *) inductive term: Type[0] ≝ @@ -37,17 +37,21 @@ interpretation "term construction (application)" 'Application C A = (Appl C A). notation "hvbox( # term 90 i )" - non associative with precedence 55 + non associative with precedence 90 for @{ 'VariableReferenceByIndex $i }. -notation "hvbox( 𝛌 . term 55 A )" - non associative with precedence 55 +notation "hvbox( 𝛌 . term 46 A )" + non associative with precedence 46 for @{ 'Abstraction $A }. -notation > "hvbox( 𝛌 term 55 A )" - non associative with precedence 55 - for @{ 'Abstraction $A }. - -notation "hvbox( @ term 55 C . break term 55 A )" - non associative with precedence 55 +notation "hvbox( @ term 46 C . break term 46 A )" + non associative with precedence 46 for @{ 'Application $C $A }. + +definition appl_compatible_dx: predicate (relation term) ≝ λR. + ∀B,A1,A2. R A1 A2 → R (@B.A1) (@B.A2). + +lemma star_appl_compatible_dx: ∀R. appl_compatible_dx R → + appl_compatible_dx (star … R). +#R #HR #B #A1 #A2 #H elim H -A2 // /3 width=3/ +qed.