aboutsummaryrefslogtreecommitdiffhomepage
path: root/test-suite/success/Field.v
diff options
context:
space:
mode:
authorGravatar delahaye <delahaye@85f007b7-540e-0410-9357-904b9bb8a0f7>2001-04-20 08:26:02 +0000
committerGravatar delahaye <delahaye@85f007b7-540e-0410-9357-904b9bb8a0f7>2001-04-20 08:26:02 +0000
commit5d99ec7266c550a070798fff7116a733e5a8556c (patch)
treea7e1b113dbb0e92bc43a5e71f1c1ff5f6583158e /test-suite/success/Field.v
parentbe3d809948aac00df3c6b93e2a90d0dc0df4bcf6 (diff)
Tests pour Field avec les nombres reels
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@1644 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'test-suite/success/Field.v')
-rw-r--r--test-suite/success/Field.v71
1 files changed, 71 insertions, 0 deletions
diff --git a/test-suite/success/Field.v b/test-suite/success/Field.v
new file mode 100644
index 000000000..cf8583e28
--- /dev/null
+++ b/test-suite/success/Field.v
@@ -0,0 +1,71 @@
+(***********************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
+(* \VV/ *************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(***********************************************************************)
+
+(* $Id$ *)
+
+(**** Tests of Field with real numbers ****)
+
+Require Reals.
+
+(* Example 1 *)
+Goal (eps:R)``eps*1/(2+2)+eps*1/(2+2) == eps*1/2``.
+Proof.
+ Intros.
+ Field.
+Abort.
+
+(* Example 2 *)
+Goal (f,g:(R->R); x0,x1:R)
+ ``((f x1)-(f x0))*1/(x1-x0)+((g x1)-(g x0))*1/(x1-x0) == ((f x1)+
+ (g x1)-((f x0)+(g x0)))*1/(x1-x0)``.
+Proof.
+ Intros.
+ Field.
+Abort.
+
+(* Example 3 *)
+Goal (a,b:R)``1/(a*b)*1/1/b == 1/a``.
+Proof.
+ Intros.
+ Field.
+Abort.
+
+(* Example 4 *)
+Goal (a,b:R)``a <> 0``->``b <> 0``->``1/(a*b)/1/b == 1/a``.
+Proof.
+ Intros.
+ Field.
+Abort.
+
+(* Example 5 *)
+Goal (a:R)``1 == 1*1/a*a``.
+Proof.
+ Intros.
+ Field.
+Abort.
+
+(* Example 6 *)
+Goal (a,b:R)``b == b*/a*a``.
+Proof.
+ Intros.
+ Field.
+Abort.
+
+(* Example 7 *)
+Goal (a,b:R)``b == b*1/a*a``.
+Proof.
+ Intros.
+ Field.
+Abort.
+
+(* Example 8 *)
+Goal (x,y:R)``x*((1/x)+x/(x+y)) == -(1/y)*y*(-(x*x/(x+y))-1)``.
+Proof.
+ Intros.
+ Field.
+Abort.