1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
17 include "logic/equality.ma".
20 axiom A : nat -> Type.
21 axiom B : nat -> Type.
22 axiom AB : \forall x. A x = B x.
24 axiom eatA : ∀n. A n -> A O.
25 axiom eatB : ∀n. B n -> A O.
27 axiom jmcBA : ∀n,m.∀p:A n = B m.B m -> A n.
28 axiom jmcAB : ∀n,m.∀p:A n = B m.A n -> B m.
30 coercion cic:/matita/tests/coercions_open/jmcAB.con.
31 coercion cic:/matita/tests/coercions_open/jmcBA.con.
33 axiom daemon : ∀x,y:A O.x = y.
35 alias num (instance 0) = "natural number".
36 lemma xx : ∀b:B 2.∀a:A 1.eatA ? b = eatB ? a.
37 intros; [3,5,7,9: apply AB|1: apply daemon];skip.