X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Ftests%2Fng_coercions.ma;h=3e22b3410b426b980396440ed8673c0a8c244cca;hb=d86eefac7dff521eb2589b6f2dcb8a1b361be186;hp=1a16862670a8b6fa809a8517ac38df4a657881aa;hpb=dcdbb979433a61e2ef2842d96604098728824416;p=helm.git diff --git a/helm/software/matita/tests/ng_coercions.ma b/helm/software/matita/tests/ng_coercions.ma index 1a1686267..3e22b3410 100644 --- a/helm/software/matita/tests/ng_coercions.ma +++ b/helm/software/matita/tests/ng_coercions.ma @@ -12,19 +12,51 @@ (* *) (**************************************************************************) -axiom A : Type. -axiom B : Type. -axiom C : Type. +include "ng_pts.ma". -axiom f : A → B. -axiom g : B → C. +(* easy *) +naxiom T : Type. +naxiom S : Type. +naxiom R : Type. +naxiom U : Type. +naxiom c : T → S. +naxiom c1 : S → R. +naxiom c2 : R → U. -(* nlemma xxx : ∀P:C → Prop. ∀x:A. P x. *) +ncoercion foo1 : ∀_t:T.S ≝ c on _t : T to S. +ncoercion foo2 : ∀_r:R.U ≝ c2 on _r : R to U. +ncoercion foo3 : ∀_s:S.R ≝ c1 on _s : S to R. -coercion f. coercion g. +(* another *) -axiom T : Type. +naxiom nat : Type. +naxiom A : nat → Type. +naxiom B : nat → Type. +naxiom C : nat → Type. +naxiom D : nat → Type. +naxiom a : ∀n:nat. A n → B n. +naxiom a1 : ∀n:nat. B n → C n. +naxiom a2 : ∀n:nat. C n → D n. -axiom h : A → T. coercion h. +ncoercion foo1 : ∀n:nat. ∀_a:A n. B n ≝ a on _a : A ? to B ?. +ncoercion foo2 : ∀n:nat. ∀_c:C n. D n ≝ a2 on _c : C ? to D ?. +ncoercion foo3 : ∀n:nat. ∀_b:B n. C n ≝ a1 on _b : B ? to C ?. -nlemma xxx : ∀P:C → Prop. ∀x:A. P x. \ No newline at end of file +naxiom cx : ∀n,m:nat. B n → C m. + +ncoercion foo3 : ∀n,m:nat. ∀_b:B n. C m ≝ cx on _b : B ? to C ?. + +naxiom Suc : nat → nat. +naxiom cs : ∀n:nat. B n → C (Suc n). + +ncoercion foo3 : ∀n:nat. ∀_b:B n. C (Suc n) ≝ cs on _b : B ? to C ?. + +(* funclass *) +naxiom Y : Type. +naxiom W : Type. +naxiom X : Type. +naxiom f : Y → W. +naxiom g : W → X → X → X. + +ncoercion foo : ∀_y:Y. W ≝ f on _y : Y to W. +ncoercion foo : ∀_w:W. X → X → X ≝ g on _w : W to Π_.Π_.?.