X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fmatita%2Ftests%2Finversion.ma;h=3e49e06685d22818abd961fec9774f7708b4ab79;hb=771ee8b9d122fa963881c876e86f90531bb7434f;hp=c76851f59fa8dd1240d87fdb04d8072a5544ba5f;hpb=ebc063e65d908c9f35619c92454dbbe76bdabd40;p=helm.git diff --git a/helm/matita/tests/inversion.ma b/helm/matita/tests/inversion.ma index c76851f59..3e49e0668 100644 --- a/helm/matita/tests/inversion.ma +++ b/helm/matita/tests/inversion.ma @@ -1,48 +1,61 @@ -set "baseuri" "cic:/matita/tests/inversion/". +(**************************************************************************) +(* ___ *) +(* ||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 *) +(* *) +(**************************************************************************) -inductive nat : Set \def - O : nat - | S : nat \to nat. - -inductive le (n:nat) : nat \to Prop \def - leO : le n n - | leS : \forall m. le n m \to le n (S m). +set "baseuri" "cic:/matita/tests/inversion_sum/". +include "legacy/coq.ma". + + +alias symbol "eq" (instance 0) = "Coq's leibnitz's equality". +alias id "O" = "cic:/Coq/Init/Datatypes/nat.ind#xpointer(1/1/1)". + +inductive sum (n:nat) : nat \to nat \to Set \def + k: \forall x,y. n = x + y \to sum n x y. -alias symbol "eq" (instance 0) = "leibnitz's equality". -theorem test_inversion: \forall n. le n O \to n=O. - intros. - (* inversion begins *) - cut O=O. - (* goal 2: 0 = 0 *) - goal 7. reflexivity. - (* goal 1 *) - generalize in match Hcut. - apply (le_ind ? (\lambda x. O=x \to n=x) ? ? ? H). - (* first goal (left open) *) - intro. rewrite < H1. - clear Hcut. - (* second goal (closed) *) - goal 14. - simplify. intros. - discriminate H3. - (* inversion ends *) - reflexivity. -qed. -(* Piu' semplice e non lascia l'ipotesi inutile Hcut *) -alias id "refl_equal" = "cic:/Coq/Init/Logic/eq.ind#xpointer(1/1/1)". -theorem test_inversion2: \forall n. le n O \to n=O. + +theorem t: \forall x,y. \forall H: sum x y O. + match H with [ (k a b p) \Rightarrow a ] = x. intros. - (* inversion begins *) - generalize in match (refl_equal nat O). - apply (le_ind ? (\lambda x. O=x \to n=x) ? ? ? H). - (* first goal (left open) *) - intro. rewrite < H1. - (* second goal (closed) *) - goal 13. + inversion H. + + (* + cut (y = y \to O = O \to match H with [ (k a b p) \Rightarrow a] = x). + apply Hcut; reflexivity. + apply + (sum_ind ? + (\lambda a,b,K. y=a \to O=b \to + match K with [ (k a b p) \Rightarrow a ] = x) + ? ? ? H). + goal 16.*) simplify. intros. - discriminate H3. - (* inversion ends *) - reflexivity. + generalize in match H1. + rewrite < H2; rewrite < H3.intro. + rewrite > H4.auto. +qed. + +theorem t1: \forall x,y. sum x y O \to x = y. +intros. + +(* +cut y=y \to O=O \to x = y. +apply Hcut.reflexivity. reflexivity. +apply (sum_ind ? (\lambda a,b,K. y=a \to O=b \to x=a) ? ? ? s).*) + +(*apply (sum_ind ? (\lambda a,b,K. y = a \to O = b \to x = a) ? ? ? s).*) +inversion s. +intros.simplify. +intros. +rewrite > H. rewrite < H2. auto. qed.