X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Ftests%2Finjection.ma;h=b805df687a70ac8fa5e71eeb7bc3713269b2b372;hb=dee331ab42d5d625f32fecc3e70df013c2dd093d;hp=7f79ef3867146f4e77bd5513811178a5d09551d4;hpb=d38484e586188871d25ff9d582b1264fa58350d8;p=helm.git diff --git a/helm/software/matita/tests/injection.ma b/helm/software/matita/tests/injection.ma index 7f79ef386..b805df687 100644 --- a/helm/software/matita/tests/injection.ma +++ b/helm/software/matita/tests/injection.ma @@ -12,9 +12,9 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/test/injection/". -include "legacy/coq.ma". + +include "coq.ma". alias id "nat" = "cic:/Coq/Init/Datatypes/nat.ind#xpointer(1/1)". alias id "bool" = "cic:/Coq/Init/Datatypes/bool.ind#xpointer(1/1)". @@ -26,8 +26,8 @@ inductive t0 : Type := theorem injection_test0: ∀n,n',m,m'. k0 n m = k0 n' m' → m = m'. intros; - injection H; - assumption. + destruct H; + reflexivity. qed. inductive t : Type → Type := @@ -36,8 +36,8 @@ inductive t : Type → Type := theorem injection_test1: ∀n,n'. k n = k n' → n = n'. intros; - injection H; - assumption. + destruct H; + reflexivity. qed. inductive tt (A:Type) : Type -> Type := @@ -46,39 +46,17 @@ inductive tt (A:Type) : Type -> Type := theorem injection_test2: ∀n,n',m,m'. k1 bool n n' = k1 bool m m' → n' = m'. intros; - injection H; - assumption. + destruct H; + reflexivity. qed. inductive ttree : Type → Type := tempty: ttree nat | tnode : ∀A. ttree A → ttree A → ttree A. -(* CSC: there is an undecidable unification problem here: - consider a constructor k : \forall x. f x -> i (g x) - The head of the outtype of the injection MutCase should be (f ?1) - such that (f ?1) unifies with (g d) [ where d is the Rel that binds - the corresponding right parameter in the outtype ] - Coq dodges the problem by generating an equality between sigma-types - (that state the existence of a ?1 such that ...) -theorem injection_test3: - ∀t,t'. tnode nat t tempty = tnode nat t' tempty → t = t'. - intros; - injection H; - assumption. -qed. - -theorem injection_test3: - ∀t,t'. - tnode nat (tnode nat t t') tempty = tnode nat (tnode nat t' tempty) tempty → - t = t'. - intros; - injection H; -*) - theorem injection_test4: - ∀n,n',m,m'. k1 bool (S n) (S m) = k1 bool (S n') (S (S m')) → m = S m'. + ∀n,n',m,m'. k1 bool (S n) (S (S m)) = k1 bool (S n') (S (S (S m'))) → m = S m'. intros; - injection H; - assumption. -qed. \ No newline at end of file + destruct H; + reflexivity. +qed.