]> 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 6e6900f806b1a2aa57464474a1db2b7c0f9ce3b2..310512b5a73003a332c028582837ed67e4cc95d9 100644 (file)
@@ -53,9 +53,9 @@ non associative with precedence 90 for
 @{ let ${ident a} ≝ $b in $c }.
 
 unification hint 0 (∀g:group.∀x:Expr.∀G:list (gcarr g). 
-
+           V ≟ 〚x,G〛
 (* ------------------------------------ *)
-       〚Eopp x,G〛 = 〚x,G〛^-1).
+       〚Eopp x,G〛 = V^-1).
 
 
 unification hint 0 (∀g:group.∀x,y.∀G:list (gcarr g). 
@@ -79,7 +79,26 @@ unification hint 3 (∀g:group.∀G:list (gcarr g).∀n.∀x:gcarr g.
 
                V ≟ 〚Evar n, G〛 
 (* ------------------------------------ *)  
-       〚(Evar (S n)), (x :: G)〛 = V) . 
+       〚(Evar (S n)), (x :: G)〛 = V) .
+       
+(* Esempio banale di divergenza       
+unification hint 0 (∀g:group.∀G:list (gcarr g).∀x.
+     V ≟ 〚x,G〛 
+ ------------------------------------      
+      〚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.