]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/tests/inversion.ma
A simpler implementation of inversion that does not generate the dummy
[helm.git] / helm / matita / tests / inversion.ma
index 5f173344454084a0f2fd296b9732d4ba2f16729c..9c56069b234bc1aa063c2c04b560bf45aecde005 100644 (file)
@@ -22,3 +22,14 @@ theorem test_inversion: \forall n. le n O \to n=O.
   simplify. intros.
   discriminate H3.
 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.
+ generalize (refl_equal nat O).
+ apply (le_ind ? (\lambda x. O=x \to n=x) ? ? ? H).
+ intro. reflexivity.
+ simplify. intros.
+ discriminate H3.
+qed.