]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/tests/match.ma
Tactic discriminate activated in matita.
[helm.git] / helm / matita / tests / match.ma
index 21bad46a9f8ab8d2c0e5d2c236860c9805c3eaec..bc8caa22332b5c4a18bb65e7a941def31f2e47c8 100644 (file)
@@ -1,13 +1,28 @@
+(**************************************************************************)
+(*       ___                                                               *)
+(*      ||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).
+\lambda A. (A \to False).
 
 theorem absurd : \forall A,C:Prop. A \to Not A \to C.
-intros.cut False.elim Hcut.apply H1.assumption.
+intros. elim (H1 H).
 qed.
 
 inductive And (A,B:Prop) : Prop \def
@@ -97,13 +112,10 @@ 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:nat with 
+let rec plus n m \def 
+ match n with 
  [ O \Rightarrow m
- | (S p) \Rightarrow S (plus p m) ]
-in
-plus.
+ | (S p) \Rightarrow S (plus p m) ].
 
 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.
@@ -123,13 +135,10 @@ theorem assoc_plus:
 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:nat with 
+let rec times n m \def 
+ match n with 
  [ O \Rightarrow O
- | (S p) \Rightarrow (plus m (times p m)) ]
-in
-times.
+ | (S p) \Rightarrow (plus m (times p m)) ].
 
 theorem times_n_O: \forall n:nat. eq nat O (times n O).
 intros.elim n.simplify.apply refl_equal.simplify.assumption.
@@ -151,16 +160,13 @@ 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:nat with 
+let rec minus n m \def 
+ match n with 
  [ O \Rightarrow O
  | (S p) \Rightarrow 
-       [\lambda n:nat.nat] match m:nat with
+       match m with
        [O \Rightarrow (S p)
-        | (S q) \Rightarrow minus p q ]]
-in
-minus.
+        | (S q) \Rightarrow minus p q ]].
 
 theorem nat_case :
 \forall n:nat.\forall P:nat \to Prop. 
@@ -270,15 +276,13 @@ 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:nat with 
+let rec leb n m \def 
+    match n with 
     [ O \Rightarrow true
     | (S p) \Rightarrow
-       [\lambda n:nat.bool] match m:nat with 
+       match m with 
         [ O \Rightarrow false
-       | (S q) \Rightarrow leb p q]]
-in leb.
+       | (S q) \Rightarrow leb p q]].
 
 theorem le_dec: \forall n,m:nat. if_then_else (leb n m) (le n m) (Not (le n m)).
 intros.
@@ -291,6 +295,7 @@ simplify.apply le_n_S.apply H.
 simplify.intros.apply H.apply le_S_n.assumption.
 qed.
 
+(*CSC: this requires too much time
 theorem prova :
 let three \def (S (S (S O))) in
 let nine \def (times three three) in
@@ -301,5 +306,4 @@ intro.
 intro.
 intro.intro.
 normalize goal at (? ? % ?).
-
-
+*)
\ No newline at end of file