]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/tests/inversion.ma
all initialization code is now in the new matitaInit.ml module.
[helm.git] / helm / matita / tests / inversion.ma
index 888ec5abd130ae290ec3d09555a4859c0f8fbde2..4f8c4c6d42b6560bd04872e007892776bab10ab0 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,7 +23,7 @@ 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.
@@ -17,10 +32,10 @@ theorem test_inversion: \forall n. le n O \to n=O.
   (* goal 2: 0 = 0 *)
   goal 7. reflexivity.
   (* goal 1 *)
-  generalize Hcut.
+  generalize in match Hcut.
   apply (le_ind ? (\lambda x. O=x \to n=x) ? ? ? H).
   (* first goal (left open) *)
-  intro. rewrite right H1.
+  intro. rewrite < H1.
   clear Hcut.
   (* second goal (closed) *)
   goal 14.
@@ -35,10 +50,10 @@ alias id "refl_equal" = "cic:/Coq/Init/Logic/eq.ind#xpointer(1/1/1)".
 theorem test_inversion2: \forall n. le n O \to n=O.
  intros.
  (* inversion begins *)
- generalize (refl_equal nat O).
+ generalize in match (refl_equal nat O).
  apply (le_ind ? (\lambda x. O=x \to n=x) ? ? ? H).
  (* first goal (left open) *)
- intro. rewrite right H1.
+ intro. rewrite < H1.
  (* second goal (closed) *)
  goal 13.
  simplify. intros.