]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/tests/unifhint.ma
new apply almost there
[helm.git] / helm / software / matita / tests / unifhint.ma
index 42e9f687888ffd12ddf5cb9bdded6af59a35963d..310512b5a73003a332c028582837ed67e4cc95d9 100644 (file)
@@ -88,6 +88,18 @@ unification hint 0 (∀g:group.∀G:list (gcarr g).∀x.
       〚x,G〛 = V).
 *)
 
+include "nat/plus.ma".
+unification hint 0 (∀x,y.S x + y = S (x + y)).
+
+axiom T : nat → Prop.
+axiom F : ∀n,k.T (S (n + k)) → Prop. 
+lemma diverge: ∀k,k1.∀h:T (? + k).F ? k1 h. 
+    ?+k    = S(?+k1)
+    S?1 + k   S(?1+k1)
+    
+lemma test : ∀g:group.∀x,y:gcarr g. ∀h:P ? ((𝟙 * x) * (x^-1 * y)). 
+   start g ? ? h.
+
 lemma test : ∀g:group.∀x,y:gcarr g. ∀h:P ? ((𝟙 * x) * (x^-1 * y)). 
    start g ? ? h.