set "baseuri" "cic:/matita/datatypes/bool/".
include "logic/equality.ma".
+include "higher_order_defs/functions.ma".
inductive bool : Set \def
| true : bool
intros 2.elim b.exact H. exact H.
qed.
+theorem notb_notb: \forall b:bool. notb (notb b) = b.
+intros.
+elim b;reflexivity.
+qed.
+
+theorem injective_notb: injective bool bool notb.
+unfold injective.
+intros.
+rewrite < notb_notb.
+rewrite < (notb_notb y).
+apply eq_f.
+assumption.
+qed.
+
(*CSC: the URI must disappear: there is a bug now *)
interpretation "boolean not" 'not x = (cic:/matita/datatypes/bool/notb.con x).
intros 3.elim b1.exact H. exact H.
qed.
+theorem and_true: \forall a,b:bool.
+andb a b =true \to a =true \land b= true.
+intro.elim a
+ [split
+ [reflexivity|assumption]
+ |apply False_ind.
+ apply not_eq_true_false.
+ apply sym_eq.
+ assumption
+ ]
+qed.
+
theorem andb_true_true: \forall b1,b2. (b1 \land b2) = true \to b1 = true.
intro. elim b1.
reflexivity.