X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fgrammar%2Fterm_simple.ma;h=6332ff149a2f94db65cc42751394393df1aea055;hb=ad3ca38634cfae29e8c26d0ab23cb466407eca5e;hp=328dc55a69e411fb87502bca8a4256c09e12f0a4;hpb=e8998d29ab83e7b6aa495a079193705b2f6743d3;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/grammar/term_simple.ma b/matita/matita/contribs/lambdadelta/basic_2/grammar/term_simple.ma index 328dc55a6..6332ff149 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/grammar/term_simple.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/grammar/term_simple.ma @@ -12,6 +12,7 @@ (* *) (**************************************************************************) +include "basic_2/notation/relations/simple_1.ma". include "basic_2/grammar/term.ma". (* SIMPLE (NEUTRAL) TERMS ***************************************************) @@ -24,21 +25,18 @@ inductive simple: predicate term ≝ interpretation "simple (term)" 'Simple T = (simple T). (* Basic inversion lemmas ***************************************************) -(* -lemma mt: ∀R1,R2:Prop. (R1 → R2) → (R2 → ⊥) → R1 → ⊥. -/3 width=1/ qed. -*) -fact simple_inv_bind_aux: ∀T. 𝐒⦃T⦄ → ∀a,J,W,U. T = ⓑ{a,J} W. U → ⊥. + +fact simple_inv_bind_aux: ∀T. 𝐒⦃T⦄ → ∀p,J,W,U. T = ⓑ{p,J}W.U → ⊥. #T * -T -[ #I #a #J #W #U #H destruct +[ #I #p #J #W #U #H destruct | #I #V #T #a #J #W #U #H destruct ] -qed. +qed-. -lemma simple_inv_bind: ∀a,I,V,T. 𝐒⦃ⓑ{a,I} V. T⦄ → ⊥. -/2 width=7/ qed-. (**) (* auto fails if mt is enabled *) +lemma simple_inv_bind: ∀p,I,V,T. 𝐒⦃ⓑ{p,I} V. T⦄ → ⊥. +/2 width=7 by simple_inv_bind_aux/ qed-. lemma simple_inv_pair: ∀I,V,T. 𝐒⦃②{I}V.T⦄ → ∃J. I = Flat2 J. -* /2 width=2/ #a #I #V #T #H -elim (simple_inv_bind … H) +* /2 width=2 by ex_intro/ +#p #I #V #T #H elim (simple_inv_bind … H) qed-.