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 (**************************************************************************)
15 universe constraint Type[0] < Type[1].
16 universe constraint Type[1] < Type[2].
19 naxiom P : T → T → CProp[0].
20 naxiom Q : T → CProp[0].
23 naxiom f_com : ∀x,y. P (f x y) (f y x).
24 naxiom f_Q : ∀x,y. P x y → Q (f x y).
26 nlemma foo : ∀x,y. Q x → Q (f (f x y) (f y x)).
34 naxiom And : CProp[0] → CProp[0] → CProp[0].
35 naxiom And_intro : ∀A,B.A → B → And A B.
39 naxiom is_nat : T → CProp[0].
40 naxiom is_nat_0 : is_nat zero.
41 naxiom is_nat_S : ∀x.is_nat x → is_nat (succ x).
43 nlemma bar : ∀P:T → CProp[0].P (succ zero) → (λx.And (is_nat x) (P x)) ?.
44 ##[ #P; #H; nwhd; napply And_intro; ##] nauto.
50 nlemma baz : ∀P,Q:CProp[0].(A → P) → (And A P → Q) → Q.
56 ∀And: CProp[0] → CProp[0] → CProp[0] → CProp[0].
57 ∀And_elim : ∀a,b,c.a → b → c → And a b c.
58 ∀C: T → T → T → CProp[0].
69 (λx,y,z:T.And (A x) (B x y) (C x y z)) ???.
70 ##[ #T; #And; #And_intro; #A; #B; #C; #a; #b; #p1; #p2; #p3; #p4; #p5;