X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2FBasic_2%2Fgrammar%2Fterm_simple.ma;fp=matita%2Fmatita%2Fcontribs%2Flambda_delta%2FBasic_2%2Fgrammar%2Fterm_simple.ma;h=c182ba36056da2b240a6ccfea49b0ef5bdccb604;hb=55dc00c1c44cc21c7ae179cb9df03e7446002c46;hp=0000000000000000000000000000000000000000;hpb=5b03651298a3943b67f49fb78dc30bb8b2780f30;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/Basic_2/grammar/term_simple.ma b/matita/matita/contribs/lambda_delta/Basic_2/grammar/term_simple.ma new file mode 100644 index 000000000..c182ba360 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/Basic_2/grammar/term_simple.ma @@ -0,0 +1,36 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "Basic-2/grammar/term.ma". + +(* SIMPLE (NEUTRAL) TERMS ***************************************************) + +inductive simple: term → Prop ≝ + | simple_atom: ∀I. simple (𝕒{I}) + | simple_flat: ∀I,V,T. simple (𝕗{I} V. T) +. + +interpretation "simple (term)" 'Simple T = (simple T). + +(* Basic inversion lemmas ***************************************************) + +fact simple_inv_bind_aux: ∀T. 𝕊[T] → ∀J,W,U. T = 𝕓{J} W. U → False. +#T * -T +[ #I #J #W #U #H destruct +| #I #V #T #J #W #U #H destruct +] +qed. + +lemma simple_inv_bind: ∀I,V,T. 𝕊[𝕓{I} V. T] → False. +/2 width=6/ qed.