]> matita.cs.unibo.it Git - helm.git/commitdiff
Added some examples for auto/paramodulation/demodulation.
authorAndrea Asperti <andrea.asperti@unibo.it>
Wed, 25 Jan 2006 08:06:07 +0000 (08:06 +0000)
committerAndrea Asperti <andrea.asperti@unibo.it>
Wed, 25 Jan 2006 08:06:07 +0000 (08:06 +0000)
helm/matita/library/SK.ma [new file with mode: 0644]
helm/matita/library/demodulation_coq.ma [new file with mode: 0644]
helm/matita/library/demodulation_matita.ma [new file with mode: 0644]

diff --git a/helm/matita/library/SK.ma b/helm/matita/library/SK.ma
new file mode 100644 (file)
index 0000000..708f92f
--- /dev/null
@@ -0,0 +1,116 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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/SK/".
+
+include "legacy/coq.ma".
+alias symbol "eq" = "Coq's leibnitz's equality".
+
+theorem SKK:
+  \forall A:Set.
+  \forall app: A \to A \to A.
+  \forall K:A. 
+  \forall S:A.
+  \forall H1: (\forall x,y:A.(app (app K x) y) = x).
+  \forall H2: (\forall x,y,z:A.
+    (app (app (app S x) y) z) = (app (app x z) (app y z))).
+  \forall x:A.
+    (app (app (app S K) K) x) = x.
+intros.auto paramodulation.
+qed.
+
+theorem bool1:
+  \forall A:Set.
+  \forall one:A.
+  \forall zero:A.
+  \forall add: A \to A \to A.
+  \forall mult: A \to A \to A.
+  \forall inv: A \to A.
+  \forall c1:(\forall x,y:A.(add x y) = (add y x)). 
+  \forall c2:(\forall x,y:A.(mult x y) = (mult y x)). 
+  \forall d1: (\forall x,y,z:A.
+              (add x (mult y z)) = (mult (add x y) (add x z))).
+  \forall d2: (\forall x,y,z:A.
+              (mult x (add y z)) = (add (mult x y) (mult x z))).  
+  \forall i1: (\forall x:A. (add x zero) = x).
+  \forall i2: (\forall x:A. (mult x one) = x).   
+  \forall inv1: (\forall x:A. (add x (inv x)) = one).  
+  \forall inv2: (\forall x:A. (mult x (inv x)) = zero). 
+  (inv zero) = one.
+intros.auto paramodulation.
+qed.
+  
+theorem bool2:
+  \forall A:Set.
+  \forall one:A.
+  \forall zero:A.
+  \forall add: A \to A \to A.
+  \forall mult: A \to A \to A.
+  \forall inv: A \to A.
+  \forall c1:(\forall x,y:A.(add x y) = (add y x)). 
+  \forall c2:(\forall x,y:A.(mult x y) = (mult y x)). 
+  \forall d1: (\forall x,y,z:A.
+              (add x (mult y z)) = (mult (add x y) (add x z))).
+  \forall d2: (\forall x,y,z:A.
+              (mult x (add y z)) = (add (mult x y) (mult x z))).  
+  \forall i1: (\forall x:A. (add x zero) = x).
+  \forall i2: (\forall x:A. (mult x one) = x).   
+  \forall inv1: (\forall x:A. (add x (inv x)) = one).  
+  \forall inv2: (\forall x:A. (mult x (inv x)) = zero).
+  \forall x:A. (mult x zero) = zero.
+intros.auto paramodulation.
+qed.
+
+theorem bool3:
+  \forall A:Set.
+  \forall one:A.
+  \forall zero:A.
+  \forall add: A \to A \to A.
+  \forall mult: A \to A \to A.
+  \forall inv: A \to A.
+  \forall c1:(\forall x,y:A.(add x y) = (add y x)). 
+  \forall c2:(\forall x,y:A.(mult x y) = (mult y x)). 
+  \forall d1: (\forall x,y,z:A.
+              (add x (mult y z)) = (mult (add x y) (add x z))).
+  \forall d2: (\forall x,y,z:A.
+              (mult x (add y z)) = (add (mult x y) (mult x z))).  
+  \forall i1: (\forall x:A. (add x zero) = x).
+  \forall i2: (\forall x:A. (mult x one) = x).   
+  \forall inv1: (\forall x:A. (add x (inv x)) = one).  
+  \forall inv2: (\forall x:A. (mult x (inv x)) = zero).
+  \forall x:A. (inv (inv x)) = x.
+intros.auto paramodulation.
+qed.
+  
+theorem bool2:
+  \forall A:Set.
+  \forall one:A.
+  \forall zero:A.
+  \forall add: A \to A \to A.
+  \forall mult: A \to A \to A.
+  \forall inv: A \to A.
+  \forall c1:(\forall x,y:A.(add x y) = (add y x)). 
+  \forall c2:(\forall x,y:A.(mult x y) = (mult y x)). 
+  \forall d1: (\forall x,y,z:A.
+              (add x (mult y z)) = (mult (add x y) (add x z))).
+  \forall d2: (\forall x,y,z:A.
+              (mult x (add y z)) = (add (mult x y) (mult x z))).  
+  \forall i1: (\forall x:A. (add x zero) = x).
+  \forall i2: (\forall x:A. (mult x one) = x).   
+  \forall inv1: (\forall x:A. (add x (inv x)) = one).  
+  \forall inv2: (\forall x:A. (mult x (inv x)) = zero). 
+  \forall x,y:A.
+    (inv (mult x y)) = (add (inv x) (inv y)).
+intros.auto paramodulation.
+qed.
diff --git a/helm/matita/library/demodulation_coq.ma b/helm/matita/library/demodulation_coq.ma
new file mode 100644 (file)
index 0000000..aa9d5f1
--- /dev/null
@@ -0,0 +1,52 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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/demodulation/".
+
+include "legacy/coq.ma".
+
+alias num = "natural number".
+alias symbol "times" = "Coq's natural times".
+alias symbol "plus" = "Coq's natural plus".
+alias symbol "eq" = "Coq's leibnitz's equality".
+alias id "nat" = "cic:/Coq/Init/Datatypes/nat.ind#xpointer(1/1)".
+alias id "S" = "cic:/Coq/Init/Datatypes/nat.ind#xpointer(1/1/2)".
+
+
+theorem p0 : \forall m:nat. m+O = m.
+intro. demodulate.
+
+theorem p: \forall m.1*m = m.
+intros.demodulate.reflexivity.
+qed.
+
+theorem p2: \forall x,y:nat.(S x)*y = (y+x*y).
+intros.demodulate.reflexivity.
+qed.
+
+theorem p1: \forall x,y:nat.(S ((S x)*y+x))=(S x)+(y*x+y).
+intros.demodulate.reflexivity.
+qed.
+
+theorem p3: \forall x,y:nat. (x+y)*(x+y) = x*x + 2*(x*y) + (y*y).
+intros.demodulate.reflexivity.
+qed.
+
+theorem p4: \forall x:nat. (x+1)*(x-1)=x*x - 1.
+intro.
+apply (nat_case x)
+[simplify.reflexivity
+|intro.demodulate.reflexivity]
+qed.
+
diff --git a/helm/matita/library/demodulation_matita.ma b/helm/matita/library/demodulation_matita.ma
new file mode 100644 (file)
index 0000000..0f4827e
--- /dev/null
@@ -0,0 +1,33 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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/demodulation_matita/".
+
+include "nat/minus.ma".
+
+theorem p2: \forall x,y:nat. x+x = (S(S O))*x.
+intros.demodulate.reflexivity.
+qed.
+
+theorem p4: \forall x:nat. (x+(S O))*(x-(S O))=x*x - (S O).
+intro.
+apply (nat_case x)
+[simplify.reflexivity
+|intro.demodulate.reflexivity]
+qed.
+
+theorem p5: \forall x,y:nat. (x+y)*(x+y) = x*x + (S(S O))*(x*y) + (y*y).
+intros.demodulate.reflexivity.
+qed.
+