1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| A.Asperti, C.Sacerdoti Coen, *)
8 (* ||A|| E.Tassi, S.Zacchiroli *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU Lesser General Public License Version 2.1 *)
13 (**************************************************************************)
15 inductive True: Prop \def
18 default "true" cic:/matita/logic/connectives/True.ind.
20 inductive False: Prop \def .
22 default "false" cic:/matita/logic/connectives/False.ind.
24 definition Not: Prop \to Prop \def
25 \lambda A. (A \to False).
27 interpretation "logical not" 'not x = (Not x).
29 theorem absurd : \forall A,C:Prop. A \to \lnot A \to C.
33 default "absurd" cic:/matita/logic/connectives/absurd.con.
35 theorem not_to_not : \forall A,B:Prop. (A → B) \to ¬B →¬A.
36 intros.unfold.intro.apply H1.apply (H H2).
39 default "absurd" cic:/matita/logic/connectives/absurd.con.
41 inductive And (A,B:Prop) : Prop \def
42 conj : A \to B \to (And A B).
44 interpretation "logical and" 'and x y = (And x y).
46 theorem proj1: \forall A,B:Prop. A \land B \to A.
47 intros. elim H. assumption.
50 theorem proj2: \forall A,B:Prop. A \land B \to B.
51 intros. elim H. assumption.
54 inductive Or (A,B:Prop) : Prop \def
55 or_introl : A \to (Or A B)
56 | or_intror : B \to (Or A B).
58 interpretation "logical or" 'or x y = (Or x y).
62 \forall P: A \lor B \to Prop.
63 (\forall p:A. P (or_introl ? ? p)) \to
64 (\forall q:B. P (or_intror ? ? q)) \to
65 \forall p:A \lor B. P p.
68 (match p return \lambda p.P p with
69 [(or_introl p) \Rightarrow H p
70 |(or_intror q) \Rightarrow H1 q]).
73 definition decidable : Prop \to Prop \def \lambda A:Prop. A \lor \lnot A.
75 inductive ex (A:Type) (P:A \to Prop) : Prop \def
76 ex_intro: \forall x:A. P x \to ex A P.
78 interpretation "exists" 'exists x = (ex ? x).
80 inductive ex2 (A:Type) (P,Q:A \to Prop) : Prop \def
81 ex_intro2: \forall x:A. P x \to Q x \to ex2 A P Q.
84 \lambda A,B. (A -> B) \land (B -> A).