interpretation "logical not" 'not x = (Not x).
+(*
nlemma absurd : ∀A,C:Prop.A → ¬A → C.
#A; #C; #H;
nnormalize;
interpretation "logical or" 'or x y = (Or2 x y).
-ndefinition decidable : Prop → Prop ≝ λA:Prop.A ∨ ¬A.
-
nlemma or2_elim
: ∀P1,P2,Q:Prop.Or2 P1 P2 → ∀f1:P1 → Q.∀f2:P2 → Q.Q.
#P1; #P2; #Q; #H; #f1; #f2;
ninductive ex2 (A:Type) (Q,R:A → Prop) : Prop ≝
ex_intro2: ∀x:A.Q x → R x → ex2 A Q R.
-
-ndefinition iff ≝
-λA,B.(A → B) ∧ (B → A).
+*)
(* higher_order_defs/relations *)
ndefinition relation : Type → Type ≝
λA:Type.A → A → Prop.
+(*
ndefinition reflexive : ∀A:Type.∀R:relation A.Prop ≝
λA.λR.∀x:A.R x x.
+*)
ndefinition symmetric : ∀A:Type.∀R:relation A.Prop ≝
λA.λR.∀x,y:A.R x y → R y x.
+(*
ndefinition transitive : ∀A:Type.∀R:relation A.Prop ≝
λA.λR.∀x,y,z:A.R x y → R y z → R x z.
ndefinition antisymmetric : ∀A:Type.∀R:relation A.Prop ≝
λA.λR.∀x,y:A.R x y → ¬ (R y x).
+*)
(* logic/equality.ma *)
interpretation "leibnitz's non-equality" 'neq t x y = (Not (eq t x y)).
+(*
nlemma eq_f : ∀T1,T2:Type.∀x,y:T1.∀f:T1 → T2.x = y → (f x) = (f y).
#T1; #T2; #x; #y; #f; #H;
nrewrite < H;
nnormalize; #H; #H1;
napply (H (eq_f … H1)).
nqed.
+*)
nlemma symmetric_eq: ∀A:Type. symmetric A (eq A).
#A;
napply H.
nqed.
-nlemma symmetric_neq : ∀T.∀x,y:T.x ≠ y → y ≠ x.
+(*
+nlemma symmetric_neq : ∀T:Type.∀x,y:T.x ≠ y → y ≠ x.
#T; #x; #y;
nnormalize;
#H; #H1;
ndefinition associative : ∀A:Type.∀R:relationT A A.Prop ≝
λA.λR.∀x,y,z:A.R (R x y) z = R x (R y z).
+*)
+
+ninductive bool : Type ≝
+ true : bool |
+ false : bool.
+
+nlemma pippo : (true = false) → (false = true).
+ #H; ndestruct.
+nqed.