+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
-(* ||A|| E.Tassi, S.Zacchiroli *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU Lesser General Public License Version 2.1 *)
-(* *)
-(**************************************************************************)
-
-
-inductive True: Prop \def
-I : True.
-
-inductive False: Prop \def .
-
-definition Not: Prop \to Prop \def
-\lambda A:Prop. (A \to False).
-
-theorem absurd : \forall A,C:Prop. A \to Not A \to C.
-intros.cut False.elim Hcut.apply H1.assumption.
-qed.
-
-inductive And (A,B:Prop) : Prop \def
- conj : A \to B \to (And A B).
-
-theorem proj1: \forall A,B:Prop. (And A B) \to A.
-intros. elim H. assumption.
-qed.
-
-theorem proj2: \forall A,B:Prop. (And A B) \to A.
-intros. elim H. assumption.
-qed.
-
-inductive Or (A,B:Prop) : Prop \def
- or_introl : A \to (Or A B)
- | or_intror : B \to (Or A B).
-
-inductive ex (A:Type) (P:A \to Prop) : Prop \def
- ex_intro: \forall x:A. P x \to ex A P.
-
-inductive ex2 (A:Type) (P,Q:A \to Prop) : Prop \def
- ex_intro2: \forall x:A. P x \to Q x \to ex2 A P Q.
-
-inductive eq (A:Type) (x:A) : A \to Prop \def
- refl_equal : eq A x x.
-
-theorem sym_eq : \forall A:Type.\forall x,y:A. eq A x y \to eq A y x.
-intros. elim H. apply refl_equal.
-qed.
-
-theorem trans_eq : \forall A:Type.
-\forall x,y,z:A. eq A x y \to eq A y z \to eq A x z.
-intros.elim H1.assumption.
-qed.
-
-theorem f_equal: \forall A,B:Type.\forall f:A\to B.
-\forall x,y:A. eq A x y \to eq B (f x) (f y).
-intros.elim H.apply refl_equal.
-qed.
-
-theorem f_equal2: \forall A,B,C:Type.\forall f:A\to B \to C.
-\forall x1,x2:A. \forall y1,y2:B.
-eq A x1 x2\to eq B y1 y2\to eq C (f x1 y1) (f x2 y2).
-intros.elim H1.elim H.apply refl_equal.
-qed.
-
-
-inductive nat : Set \def
- | O : nat
- | S : nat \to nat.
-
-definition pred: nat \to nat \def
-\lambda n:nat. match n with
-[ O \Rightarrow O
-| (S u) \Rightarrow u ].
-
-theorem pred_Sn : \forall n:nat.
-(eq nat n (pred (S n))).
-intros.apply refl_equal.
-qed.
-
-theorem injective_S : \forall n,m:nat.
-(eq nat (S n) (S m)) \to (eq nat n m).
-intros.(elim (sym_eq ? ? ? (pred_Sn n))).
-(elim (sym_eq ? ? ? (pred_Sn m))).
-apply f_equal. assumption.
-qed.
-
-theorem not_eq_S : \forall n,m:nat.
-Not (eq nat n m) \to Not (eq nat (S n) (S m)).
-intros. simplify.intros.
-apply H.apply injective_S.assumption.
-qed.
-
-definition not_zero : nat \to Prop \def
-\lambda n: nat.
- match n with
- [ O \Rightarrow False
- | (S p) \Rightarrow True ].
-
-theorem O_S : \forall n:nat. Not (eq nat O (S n)).
-intros.simplify.intros.
-cut (not_zero O).exact Hcut.elim (sym_eq ? ? ? H).
-exact I.
-qed.
-
-theorem n_Sn : \forall n:nat. Not (eq nat n (S n)).
-intros.elim n.apply O_S.apply not_eq_S.assumption.
-qed.
-
-definition plus : nat \to nat \to nat \def
-let rec plus (n,m:nat) \def
- match n with
- [ O \Rightarrow m
- | (S p) \Rightarrow S (plus p m) ]
-in
-plus.
-
-theorem plus_n_O: \forall n:nat. eq nat n (plus n O).
-intros.elim n.simplify.apply refl_equal.simplify.apply f_equal.assumption.
-qed.
-
-theorem plus_n_Sm : \forall n,m:nat. eq nat (S (plus n m)) (plus n (S m)).
-intros.elim n.simplify.apply refl_equal.simplify.apply f_equal.assumption.
-qed.
-
-theorem sym_plus: \forall n,m:nat. eq nat (plus n m) (plus m n).
-intros.elim n.simplify.apply plus_n_O.
-simplify.elim (sym_eq ? ? ? H).apply plus_n_Sm.
-qed.
-
-theorem assoc_plus:
-\forall n,m,p:nat. eq nat (plus (plus n m) p) (plus n (plus m p)).
-intros.elim n.simplify.apply refl_equal.simplify.apply f_equal.assumption.
-qed.
-
-definition times : nat \to nat \to nat \def
-let rec times (n,m:nat) \def
- match n with
- [ O \Rightarrow O
- | (S p) \Rightarrow (plus m (times p m)) ]
-in
-times.
-
-theorem times_n_O: \forall n:nat. eq nat O (times n O).
-intros.elim n.simplify.apply refl_equal.simplify.assumption.
-qed.
-
-theorem times_n_Sm :
-\forall n,m:nat. eq nat (plus n (times n m)) (times n (S m)).
-intros.elim n.simplify.apply refl_equal.
-simplify.apply f_equal.elim H.
-apply trans_eq ? ? (plus (plus e m) (times e m)).apply sym_eq.
-apply assoc_plus.apply trans_eq ? ? (plus (plus m e) (times e m)).
-apply f_equal2.
-apply sym_plus.apply refl_equal.apply assoc_plus.
-qed.
-
-theorem sym_times :
-\forall n,m:nat. eq nat (times n m) (times m n).
-intros.elim n.simplify.apply times_n_O.
-simplify.elim (sym_eq ? ? ? H).apply times_n_Sm.
-qed.
-
-definition minus : nat \to nat \to nat \def
-let rec minus (n,m:nat) \def
- [\lambda n:nat.nat] match n with
- [ O \Rightarrow O
- | (S p) \Rightarrow
- [\lambda n:nat.nat] match m with
- [O \Rightarrow (S p)
- | (S q) \Rightarrow minus p q ]]
-in
-minus.
-
-theorem nat_case :
-\forall n:nat.\forall P:nat \to Prop.
-P O \to (\forall m:nat. P (S m)) \to P n.
-intros.elim n.assumption.apply H1.
-qed.
-
-theorem nat_double_ind :
-\forall R:nat \to nat \to Prop.
-(\forall n:nat. R O n) \to
-(\forall n:nat. R (S n) O) \to
-(\forall n,m:nat. R n m \to R (S n) (S m)) \to \forall n,m:nat. R n m.
-intros.cut \forall m:nat.R n m.apply Hcut.elim n.apply H.
-apply nat_case m1.apply H1.intros.apply H2. apply H3.
-qed.
-
-inductive bool : Set \def
- | true : bool
- | false : bool.
-
-definition notn : bool \to bool \def
-\lambda b:bool.
- match b with
- [ true \Rightarrow false
- | false \Rightarrow true ].
-
-definition andb : bool \to bool \to bool\def
-\lambda b1,b2:bool.
- match b1 with
- [ true \Rightarrow
- match b2 with [true \Rightarrow true | false \Rightarrow false]
- | false \Rightarrow false ].
-
-definition orb : bool \to bool \to bool\def
-\lambda b1,b2:bool.
- match b1 with
- [ true \Rightarrow
- match b2 with [true \Rightarrow true | false \Rightarrow false]
- | false \Rightarrow false ].
-
-definition if_then_else : bool \to Prop \to Prop \to Prop \def
-\lambda b:bool.\lambda P,Q:Prop.
-match b with
-[ true \Rightarrow P
-| false \Rightarrow Q].
-
-inductive le (n:nat) : nat \to Prop \def
- | le_n : le n n
- | le_S : \forall m:nat. le n m \to le n (S m).
-
-theorem trans_le: \forall n,m,p:nat. le n m \to le m p \to le n p.
-intros.
-elim H1.assumption.
-apply le_S.assumption.
-qed.
-
-theorem le_n_S: \forall n,m:nat. le n m \to le (S n) (S m).
-intros.elim H.
-apply le_n.apply le_S.assumption.
-qed.
-
-theorem le_O_n : \forall n:nat. le O n.
-intros.elim n.apply le_n.apply le_S. assumption.
-qed.
-
-theorem le_n_Sn : \forall n:nat. le n (S n).
-intros. apply le_S.apply le_n.
-qed.
-
-theorem le_pred_n : \forall n:nat. le (pred n) n.
-intros.elim n.simplify.apply le_n.simplify.
-apply le_n_Sn.
-qed.
-
-theorem not_zero_le : \forall n,m:nat. (le (S n) m ) \to not_zero m.
-intros.elim H.exact I.exact I.
-qed.
-
-theorem le_Sn_O: \forall n:nat. Not (le (S n) O).
-intros.simplify.intros.apply not_zero_le ? O H.
-qed.
-
-theorem le_n_O_eq : \forall n:nat. (le n O) \to (eq nat O n).
-intros.cut (le n O) \to (eq nat O n).apply Hcut. assumption.
-elim n.apply refl_equal.apply False_ind.apply (le_Sn_O ? H2).
-qed.
-
-theorem le_S_n : \forall n,m:nat. le (S n) (S m) \to le n m.
-intros.cut le (pred (S n)) (pred (S m)).exact Hcut.
-elim H.apply le_n.apply trans_le ? (pred x).assumption.
-apply le_pred_n.
-qed.
-
-theorem le_Sn_n : \forall n:nat. Not (le (S n) n).
-intros.elim n.apply le_Sn_O.simplify.intro.
-cut le (S e) e.apply H.assumption.apply le_S_n.assumption.
-qed.
-
-theorem le_antisym : \forall n,m:nat. (le n m) \to (le m n) \to (eq nat n m).
-intros.cut (le n m) \to (le m n) \to (eq nat n m).exact Hcut H H1.
-apply nat_double_ind (\lambda n,m.((le n m) \to (le m n) \to eq nat n m)).
-intros.whd.intros.
-apply le_n_O_eq.assumption.
-intros.whd.intro.apply sym_eq.apply le_n_O_eq.assumption.
-intros.whd.intro.apply f_equal.apply H2.
-apply le_S_n.assumption.
-apply le_S_n.assumption.
-qed.
-
-
-definition leb : nat \to nat \to bool \def
-let rec leb (n,m:nat) \def
- [\lambda n:nat.bool] match n with
- [ O \Rightarrow true
- | (S p) \Rightarrow
- [\lambda n:nat.bool] match m with
- [ O \Rightarrow false
- | (S q) \Rightarrow leb p q]]
-in leb.
-
-theorem le_dec: \forall n,m:nat. if_then_else (leb n m) (le n m) (Not (le n m)).
-intros.
-apply (nat_double_ind
-(\lambda n,m:nat.if_then_else (leb n m) (le n m) (Not (le n m))) ? ? ? n m).
-simplify.intros.apply le_O_n.
-simplify.exact le_Sn_O.
-intros 2.simplify.elim (leb n1 m1).
-simplify.apply le_n_S.apply H.
-simplify.intro.apply H.apply le_S_n.assumption.
-
-