]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/tests/inversion.ma
ocaml 3.09 transition
[helm.git] / helm / matita / tests / inversion.ma
index 4ea9e2deb76d59940a50eab535d75cb202228818..f717cd1df7f17969a0ca6169cc21833e711f4754 100644 (file)
@@ -1,4 +1,19 @@
-set "baseuri" "cic:/matita/tests/".
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/tests/inversion/".
+include "coq.ma".
 
 inductive nat : Set \def
    O : nat
@@ -8,26 +23,20 @@ inductive le (n:nat) : nat \to Prop \def
    leO : le n n
  | leS : \forall m. le n m \to le n (S m).
 
-alias symbol "eq" (instance 0) = "leibnitz's equality".
+alias symbol "eq" (instance 0) = "Coq's leibnitz's equality".
 
 theorem test_inversion: \forall n. le n O \to n=O.
  intros.
  (* inversion begins *)
- cut O=O.
-  (* goal 2: 0 = 0 *)
-  goal 7. reflexivity.
-  (* goal 1 *)
-  generalize in match Hcut.
-  apply (le_ind ? (\lambda x. O=x \to n=x) ? ? ? H).
-  (* first goal (left open) *)
-  intro. rewrite < H1.
-  clear Hcut.
-  (* second goal (closed) *)
-  goal 14.
-  simplify. intros.
-  discriminate H3.
-  (* inversion ends *)
+ cut (O=O);
+ [ 2: reflexivity;
+ | generalize in match Hcut.
+   apply (le_ind ? (\lambda x. O=x \to n=x) ? ? ? H);
+   [ intro. rewrite < H1. clear Hcut.
+   | simplify. intros. discriminate H3.
+   ]
   reflexivity.
+ ]
 qed.
 
 (* Piu' semplice e non lascia l'ipotesi inutile Hcut *)
@@ -36,13 +45,9 @@ theorem test_inversion2: \forall n. le n O \to n=O.
  intros.
  (* inversion begins *)
  generalize in match (refl_equal nat O).
- apply (le_ind ? (\lambda x. O=x \to n=x) ? ? ? H).
- (* first goal (left open) *)
- intro. rewrite < H1.
- (* second goal (closed) *)
- goal 13.
- simplify. intros.
- discriminate H3.
- (* inversion ends *)
+ apply (le_ind ? (\lambda x. O=x \to n=x) ? ? ? H);
+ [ intro. rewrite < H1.
+ | simplify. intros. discriminate H3.
+ ]
  reflexivity.
 qed.