X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Fgrammar%2Fterm_simple.ma;h=328dc55a69e411fb87502bca8a4256c09e12f0a4;hb=a2144f09d1bd7022c1f2dfd4909a1fb9772c8d56;hp=83aff07c497541a142bbfaabfbb04fdbcd11cc86;hpb=770ba48ba232d7f1782629c572820a0f1bfe4fde;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 index 83aff07c4..328dc55a6 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/grammar/term_simple.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/grammar/term_simple.ma @@ -24,23 +24,21 @@ inductive simple: predicate term ≝ interpretation "simple (term)" 'Simple T = (simple T). (* Basic inversion lemmas ***************************************************) - -fact simple_inv_bind_aux: ∀T. 𝐒⦃T⦄ → ∀J,W,U. T = ⓑ{J} W. U → ⊥. +(* +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 → ⊥. #T * -T -[ #I #J #W #U #H destruct -| #I #V #T #J #W #U #H destruct +[ #I #a #J #W #U #H destruct +| #I #V #T #a #J #W #U #H destruct ] qed. -lemma simple_inv_bind: ∀I,V,T. 𝐒⦃ⓑ{I} V. T⦄ → ⊥. -/2 width=6/ 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_pair: ∀I,V,T. 𝐒⦃②{I}V.T⦄ → ∃J. I = Flat2 J. -* /2 width=2/ #I #V #T #H +* /2 width=2/ #a #I #V #T #H elim (simple_inv_bind … H) qed-. - -(* -lemma mt: ∀R1,R2:Prop. (R1 → R2) → (R2 → ⊥) → R1 → ⊥. -/3 width=1/ qed-. -*)