X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Ftests%2Fcoercions_propagation.ma;h=7e3536b07071fc38d9192c0d138e8df71454b7ad;hb=b97a7976503b2d2e5cbc9199f848135a324775a8;hp=ce8f2e2a7af943c4ef1ba0520bd45b5cd732a295;hpb=5ec87fab28597f651b2792081da2264504f2dbf1;p=helm.git diff --git a/helm/software/matita/tests/coercions_propagation.ma b/helm/software/matita/tests/coercions_propagation.ma index ce8f2e2a7..7e3536b07 100644 --- a/helm/software/matita/tests/coercions_propagation.ma +++ b/helm/software/matita/tests/coercions_propagation.ma @@ -12,28 +12,26 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/test/coercions_propagation/". + include "logic/connectives.ma". include "nat/orders.ma". -alias num (instance 0) = "natural number". inductive sigma (A:Type) (P:A → Prop) : Type ≝ sigma_intro: ∀a:A. P a → sigma A P. interpretation "sigma" 'exists \eta.x = - (cic:/matita/test/coercions_propagation/sigma.ind#xpointer(1/1) _ x). + (cic:/matita/tests/coercions_propagation/sigma.ind#xpointer(1/1) _ x). definition inject ≝ λP.λa:nat.λp:P a. sigma_intro ? P ? p. -coercion cic:/matita/test/coercions_propagation/inject.con 0 1. +coercion cic:/matita/tests/coercions_propagation/inject.con 0 1. definition eject ≝ λP.λc: ∃n:nat.P n. match c with [ sigma_intro w _ ⇒ w]. -coercion cic:/matita/test/coercions_propagation/eject.con. +coercion cic:/matita/tests/coercions_propagation/eject.con. alias num (instance 0) = "natural number". - theorem test: ∃n. 0 ≤ n. apply (S O : ∃n. 0 ≤ n). autobatch. @@ -73,11 +71,11 @@ inductive NN (A:Type) : nat -> Type ≝ definition injectN ≝ λA,k.λP.λa:NN A k.λp:P a. sigma_intro ? P ? p. -coercion cic:/matita/test/coercions_propagation/injectN.con 0 1. +coercion cic:/matita/tests/coercions_propagation/injectN.con 0 1. definition ejectN ≝ λA,k.λP.λc: ∃n:NN A k.P n. match c with [ sigma_intro w _ ⇒ w]. -coercion cic:/matita/test/coercions_propagation/ejectN.con. +coercion cic:/matita/tests/coercions_propagation/ejectN.con. definition PN := λA,k.λx:NN A k. 1 <= k.