X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Ftests%2Fmultiple_inheritance.ma;h=1aa708fb1264bc1be037d2e3f5c8241f44e90b71;hb=b6afef7e73324824025a6d7f313129d55b72cfc6;hp=614f57ca9b219d37a64c80b31aed1ff5d8b6456d;hpb=977aadd03056689f9dc71425e20bcd7ae6f49810;p=helm.git diff --git a/helm/software/matita/tests/multiple_inheritance.ma b/helm/software/matita/tests/multiple_inheritance.ma index 614f57ca9..1aa708fb1 100644 --- a/helm/software/matita/tests/multiple_inheritance.ma +++ b/helm/software/matita/tests/multiple_inheritance.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/test/". + include "logic/equality.ma". @@ -28,18 +28,19 @@ record R : Type := { r1 :> R1; r2_ : R2; with_: C r1 = K r2_ }. lemma r2 : R → R2. intro; apply mk_R2; - [ apply (C r) - | apply (eq_rect ? ? (λx.x → x → x)); - [3: symmetry; - [2: apply (with_ r) - | skip - ] - | skip - | apply (mult (r2_ r)) - ] - ]. + [ exact (C r) + | rewrite > (with_ r); exact (mult (r2_ r))] +qed. +coercion cic:/matita/tests/multiple_inheritance/r2.con. + +(* convertibility test *) +lemma conv_test : ∀r:R.C r -> K r. + intros. + change with (C (r1 r)). + change with (K (r2 r)). + change with (C (r1 r)). + assumption. qed. -coercion cic:/matita/test/r2.con. (* Let's play with it *) definition f ≝ λr:R.λx:r.plus ? (mult ? x x) x. @@ -50,5 +51,8 @@ axiom mult_idempotent: ∀r2:R2.∀x:r2. mult ? x x = x. lemma test: ∀r:R. ∀x:r. f ? x = x. intros; unfold f; - auto paramodulation. -qed. \ No newline at end of file + autobatch paramodulation. +qed. + + +