]> 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 e61e123d791d687bbc0323ea5f16fc0aa6f6c25a..f717cd1df7f17969a0ca6169cc21833e711f4754 100644 (file)
@@ -1,3 +1,20 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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
  | S : nat \to nat.
@@ -6,17 +23,31 @@ 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.
- cut O=O.
-  (* goal 2: 0 = 0 *)
-  goal 7. reflexivity.
-  (* goal 1 *)
-  generalize Hcut.      (* non attaccata. Dovrebbe dare 0=0 -> n=0 *)
-  apply (le_ind ? (\lambda x. O=x \to n=x) ? ? ? H).
-  intro. reflexivity.
-  simplify. intros.
-  discriminate H3.
-qed.
\ No newline at end of file
+ (* inversion begins *)
+ 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 *)
+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 in match (refl_equal nat O).
+ apply (le_ind ? (\lambda x. O=x \to n=x) ? ? ? H);
+ [ intro. rewrite < H1.
+ | simplify. intros. discriminate H3.
+ ]
+ reflexivity.
+qed.