X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2FBasic_2%2Fgrammar%2Faarity.ma;h=15418582b1986f8699a48f153b375195ec0f1105;hb=b5db76fe31ab35bae0257cb6684c511bcc531e45;hp=42354e7008f136d5e8bbb2805e82df779388b8b0;hpb=35653f628dc3a3e665fee01acc19c660c9d555e3;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/Basic_2/grammar/aarity.ma b/matita/matita/contribs/lambda_delta/Basic_2/grammar/aarity.ma index 42354e700..15418582b 100644 --- a/matita/matita/contribs/lambda_delta/Basic_2/grammar/aarity.ma +++ b/matita/matita/contribs/lambda_delta/Basic_2/grammar/aarity.ma @@ -12,7 +12,16 @@ (* *) (**************************************************************************) -include "Ground_2/list.ma". +(* THE FORMAL SYSTEM λδ - MATITA SOURCE FILES + * Support for abstract candidates of reducibility closed: 2012 January 27 + * Confluence of context-sensitive parallel reduction closed: 2011 September 21 + * Confluence of context-free parallel reduction closed: 2011 September 6 + * Specification started: 2011 April 17 + * - Patience on me to gain peace and perfection! - + * [ suggested invocation to start formal specifications with ] + *) + +include "Ground_2/star.ma". include "Basic_2/notation.ma". (* ATOMIC ARITY *************************************************************) @@ -22,13 +31,15 @@ inductive aarity: Type[0] ≝ | APair: aarity → aarity → aarity (* binary aarity construction *) . -interpretation "aarity construction (atomic)" 'SItem = AAtom. +interpretation "aarity construction (atomic)" + 'Item0 = AAtom. -interpretation "aarity construction (binary)" 'SItem A1 A2 = (APair A1 A2). +interpretation "aarity construction (binary)" + 'SnItem2 A1 A2 = (APair A1 A2). (* Basic inversion lemmas ***************************************************) -lemma discr_apair_xy_x: ∀A,B. 𝕔 B. A = B → False. +lemma discr_apair_xy_x: ∀A,B. ②B. A = B → False. #A #B elim B -B [ #H destruct | #Y #X #IHY #_ #H destruct @@ -37,7 +48,7 @@ lemma discr_apair_xy_x: ∀A,B. 𝕔 B. A = B → False. ] qed-. -lemma discr_tpair_xy_y: ∀B,A. 𝕔 B. A = A → False. +lemma discr_tpair_xy_y: ∀B,A. ②B. A = A → False. #B #A elim A -A [ #H destruct | #Y #X #_ #IHX #H destruct