]> matita.cs.unibo.it Git - helm.git/commitdiff
fix
authorAndrea Asperti <andrea.asperti@unibo.it>
Mon, 30 May 2005 15:57:28 +0000 (15:57 +0000)
committerAndrea Asperti <andrea.asperti@unibo.it>
Mon, 30 May 2005 15:57:28 +0000 (15:57 +0000)
helm/matita/tests/match.ma [new file with mode: 0644]

diff --git a/helm/matita/tests/match.ma b/helm/matita/tests/match.ma
new file mode 100644 (file)
index 0000000..fb0c1b0
--- /dev/null
@@ -0,0 +1,371 @@
+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.
+intro.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.
+intro. elim H. assumption.
+qed.
+
+theorem proj2: \forall A,B:Prop. (And A B) \to A.
+intro. 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.
+intro. 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.
+intro.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).
+intro.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).
+intro.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))).
+intro.apply refl_equal.
+qed.
+
+theorem injective_S : \forall n,m:nat. 
+(eq nat (S n) (S m)) \to (eq nat n m).
+intro.(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)).
+intro. simplify.intro.
+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)).
+intro.simplify.intro.
+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)).
+intro.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:nat 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).
+intro.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)).
+intro.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).
+intro.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)).
+intro.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:nat 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).
+intro.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)).
+intro.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).
+intro.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:nat with 
+ [ O \Rightarrow O
+ | (S p) \Rightarrow 
+       [\lambda n:nat.nat] match m:nat 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.
+intro.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.
+intro.cut \forall m:nat.R n m.apply Hcut.elim n.apply H.
+apply nat_case m1.apply H1.intro.apply H2. apply H3.
+qed.
+
+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.
+intro.
+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).
+intro.elim H.alias id "le_n" = "cic:/matita/andrea/le.ind#xpointer(1/1/1)".
+apply le_n.apply le_S.assumption.
+qed.
+
+theorem le_O_n : \forall n:nat. le O n.
+intro.elim n.apply le_n.apply le_S. assumption.
+qed.
+
+theorem le_n_Sn : \forall n:nat. le n (S n).
+intro. apply le_S.apply le_n.
+qed.
+
+theorem le_pred_n : \forall n:nat. le (pred n) n.
+intro.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.
+intro.elim H.exact I.exact I.
+qed.
+
+theorem le_Sn_O: \forall n:nat. Not (le (S n) O).
+intro.simplify.intro.apply not_zero_le ? O H.
+qed.
+
+theorem le_n_O_eq : \forall n:nat. (le n O) \to (eq nat O n).
+intro.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.
+intro.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).
+intro.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).
+intro.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)).
+intro.whd.intro.alias id "le_n_O_eq" = "cic:/matita/andrea/le_n_O_eq.con".
+apply le_n_O_eq.assumption.
+intro.whd.intro.apply sym_eq.apply le_n_O_eq.assumption.
+intro.whd.intro.apply f_equal.apply H2.
+apply le_S_n.assumption.
+apply le_S_n.assumption.
+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 leb : nat \to nat \to bool \def
+let rec leb (n,m:nat) \def 
+   [\lambda n:nat.bool] match n:nat with 
+    [ O \Rightarrow true
+    | (S p) \Rightarrow
+       [\lambda n:nat.bool] match m:nat with 
+        [ O \Rightarrow false
+       | (S q) \Rightarrow leb p q]]
+in leb.
+
+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].
+
+theorem le_dec: \forall n,m:nat. if_then_else (leb n m) (le n m) (Not (le n m)).
+intro.
+apply (nat_double_ind 
+(\lambda n,m:nat. if_then_else (leb n m) (le n m) (Not (le n m))) ? ? ? n m).
+intro.whd.apply le_O_n.intro.whd.apply le_Sn_O.
+
+intros 2.
+simplify.
+cut ((
+
+\lambda b:bool. [\lambda x:bool.Prop] match b:bool with 
+[ true \Rightarrow (le n1 m1) 
+| false \Rightarrow (Not(le n1 m1))]
+
+\to
+
+ [\lambda x:bool.Prop] match b:bool with 
+[ true \Rightarrow (le (S n1) (S m1)) 
+| false \Rightarrow (Not(le (S n1) (S m1)))]
+) (leb n1 m1)).
+goal 8.
+
+exact (
+[
+\lambda b:bool. [\lambda x:bool.Prop] match b:bool with 
+[ true \Rightarrow (le n1 m1) 
+| false \Rightarrow (Not(le n1 m1))]
+
+\to
+
+ [\lambda x:bool.Prop] match b:bool with 
+[ true \Rightarrow (le (S n1) (S m1)) 
+| false \Rightarrow (Not(le (S n1) (S m1)))]
+
+]
+match (leb n1 m1) : bool with 
+[ true \Rightarrow \lambda p:(le n1 m1). le_n_S n1 m1 p
+| false \Rightarrow \lambda p:(Not(le n1 m1)). \lambda q:(le (S n1) (S m1)).p(le_S_n n1 m1 q)]
+).
+apply Hcut.
+
+
+cut (eq bool (leb n1 m1) (leb (S n1) (S m1))).
+goal 8.
+simplify.apply refl_equal.
+cut 
+  (if_then_else 
+     (leb (S n1) (S m1)) 
+     (le (S n1) (S m1)) 
+     (Not (le (S n1) ( S m1)) )).
+apply Hcut1.elim Hcut.simplify.
+check ([\lambda b:bool. [\lambda b:bool.Prop]match (b:bool) with 
+     [ true \Rightarrow 
+             (if_then_else b (le n1 m1) (Not (le  n1 m1)))
+     | false \Rightarrow 
+             (if_then_else b (le n1 m1) (Not (le  n1 m1)))
+]
+]
+match (leb n1 m1) : bool with 
+[ true \Rightarrow H | false \Rightarrow H ]).
+
+
+exact (
+[\lambda b:bool. match (b:bool) with 
+     [ true \Rightarrow 
+             (if_then_else b (le (S n1 ) (S m1)) (Not ((le (S n1) (S m1)))))
+     | false \Rightarrow 
+             (if_then_else b (le (S n1 ) (S m1)) (Not ((le (S n1) (S m1)))))
+]
+]
+match (leb n1 m1) : bool with 
+[ true \Rightarrow le_n_S n1 m1 H
+| false \Rightarrow  
+      (\lambda p:(le (S n1) (S m1)).H (le_S_n n1 m1 p))]).
+
+
+
+
+elim n.simplify.apply le_O_n.elim m.
+simplify.apply le_Sn_O.simplify.
+cut (match (leb e e1):bool with 
+[true \Rightarrow (eq bool (leb e e1) true)
+| false \Rightarrow (eq bool (leb e e1) false)]).
+goal 20.
+exact [\lambda b:bool. match b with 
+[true \Rightarrow (eq bool b true)
+| false \Rightarrow (eq bool b false)]] match (leb e e1):bool with 
+[true \Rightarrow refl_equal bool true
+| false \Rightarrow refl_equal bool false].
\ No newline at end of file