diff options
author | barras <barras@85f007b7-540e-0410-9357-904b9bb8a0f7> | 2006-10-02 15:55:52 +0000 |
---|---|---|
committer | barras <barras@85f007b7-540e-0410-9357-904b9bb8a0f7> | 2006-10-02 15:55:52 +0000 |
commit | 83c7941f909ebe5150868b1242669c7f4502b528 (patch) | |
tree | 1a98d6724ff5155de35c93633a669046e01e87e2 | |
parent | 08ac8a40c8bfde21889196519883718b8c5ea521 (diff) |
nouveau ring/field
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@9197 85f007b7-540e-0410-9357-904b9bb8a0f7
-rw-r--r-- | test-suite/success/Field.v | 41 | ||||
-rw-r--r-- | test-suite/success/LegacyField.v | 78 | ||||
-rw-r--r-- | test-suite/success/NatRing.v | 4 | ||||
-rw-r--r-- | test-suite/success/setoid_ring_module.v | 6 |
4 files changed, 110 insertions, 19 deletions
diff --git a/test-suite/success/Field.v b/test-suite/success/Field.v index 310dfb620..6fb922b0f 100644 --- a/test-suite/success/Field.v +++ b/test-suite/success/Field.v @@ -10,58 +10,71 @@ (**** Tests of Field with real numbers ****) -Require Import Reals. +Require Import Reals RealField. +Open Scope R_scope. (* Example 1 *) Goal forall eps : R, -(eps * (1 / (2 + 2)) + eps * (1 / (2 + 2)))%R = (eps * (1 / 2))%R. +eps * (1 / (2 + 2)) + eps * (1 / (2 + 2)) = eps * (1 / 2). Proof. intros. field. -Abort. +Qed. (* Example 2 *) Goal forall (f g : R -> R) (x0 x1 : R), -((f x1 - f x0) * (1 / (x1 - x0)) + (g x1 - g x0) * (1 / (x1 - x0)))%R = -((f x1 + g x1 - (f x0 + g x0)) * (1 / (x1 - x0)))%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 forall a b : R, (1 / (a * b) * (1 / 1 / b))%R = (1 / a)%R. +Goal forall a b : R, 1 / (a * b) * (1 / (1 / b)) = 1 / a. Proof. intros. field. Abort. + +Goal forall a b : R, 1 / (a * b) * (1 / 1 / b) = 1 / a. +Proof. + intros. + field_simplify_eq. +Abort. +Goal forall a b : R, 1 / (a * b) * (1 / 1 / b) = 1 / a. +Proof. + intros. + field_simplify (1 / (a * b) * (1 / 1 / b)). +Abort. + (* Example 4 *) Goal -forall a b : R, a <> 0%R -> b <> 0%R -> (1 / (a * b) / 1 / b)%R = (1 / a)%R. +forall a b : R, a <> 0 -> b <> 0 -> 1 / (a * b) / (1 / b) = 1 / a. Proof. intros. - field. -Abort. + field; auto. +Qed. (* Example 5 *) -Goal forall a : R, 1%R = (1 * (1 / a) * a)%R. +Goal forall a : R, 1 = 1 * (1 / a) * a. Proof. intros. field. Abort. (* Example 6 *) -Goal forall a b : R, b = (b * / a * a)%R. +Goal forall a b : R, b = b * / a * a. Proof. intros. field. Abort. (* Example 7 *) -Goal forall a b : R, b = (b * (1 / a) * a)%R. +Goal forall a b : R, b = b * (1 / a) * a. Proof. intros. field. @@ -70,8 +83,8 @@ Abort. (* Example 8 *) Goal forall x y : R, -(x * (1 / x + x / (x + y)))%R = -(- (1 / y) * y * (- (x * (x / (x + y))) - 1))%R. +x * (1 / x + x / (x + y)) = +- (1 / y) * y * (- (x * (x / (x + y))) - 1). Proof. intros. field. diff --git a/test-suite/success/LegacyField.v b/test-suite/success/LegacyField.v new file mode 100644 index 000000000..bc2cffa4d --- /dev/null +++ b/test-suite/success/LegacyField.v @@ -0,0 +1,78 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +(* $Id: Field.v 7693 2005-12-21 23:50:17Z herbelin $ *) + +(**** Tests of Field with real numbers ****) + +Require Import Reals. + +(* Example 1 *) +Goal +forall eps : R, +(eps * (1 / (2 + 2)) + eps * (1 / (2 + 2)))%R = (eps * (1 / 2))%R. +Proof. + intros. + legacy field. +Abort. + +(* Example 2 *) +Goal +forall (f g : R -> R) (x0 x1 : R), +((f x1 - f x0) * (1 / (x1 - x0)) + (g x1 - g x0) * (1 / (x1 - x0)))%R = +((f x1 + g x1 - (f x0 + g x0)) * (1 / (x1 - x0)))%R. +Proof. + intros. + legacy field. +Abort. + +(* Example 3 *) +Goal forall a b : R, (1 / (a * b) * (1 / 1 / b))%R = (1 / a)%R. +Proof. + intros. + legacy field. +Abort. + +(* Example 4 *) +Goal +forall a b : R, a <> 0%R -> b <> 0%R -> (1 / (a * b) / 1 / b)%R = (1 / a)%R. +Proof. + intros. + legacy field. +Abort. + +(* Example 5 *) +Goal forall a : R, 1%R = (1 * (1 / a) * a)%R. +Proof. + intros. + legacy field. +Abort. + +(* Example 6 *) +Goal forall a b : R, b = (b * / a * a)%R. +Proof. + intros. + legacy field. +Abort. + +(* Example 7 *) +Goal forall a b : R, b = (b * (1 / a) * a)%R. +Proof. + intros. + legacy field. +Abort. + +(* Example 8 *) +Goal +forall x y : R, +(x * (1 / x + x / (x + y)))%R = +(- (1 / y) * y * (- (x * (x / (x + y))) - 1))%R. +Proof. + intros. + legacy field. +Abort. diff --git a/test-suite/success/NatRing.v b/test-suite/success/NatRing.v index 8426c7e48..22d021d54 100644 --- a/test-suite/success/NatRing.v +++ b/test-suite/success/NatRing.v @@ -1,10 +1,10 @@ Require Import ArithRing. Lemma l1 : 2 = 1 + 1. -ring_nat. +ring. Qed. Lemma l2 : forall x : nat, S (S x) = 1 + S x. intro. -ring_nat. +ring. Qed. diff --git a/test-suite/success/setoid_ring_module.v b/test-suite/success/setoid_ring_module.v index 9dfedce35..e947c6d9c 100644 --- a/test-suite/success/setoid_ring_module.v +++ b/test-suite/success/setoid_ring_module.v @@ -1,4 +1,4 @@ -Require Import Setoid ZRing_th Ring_th. +Require Import Setoid Ring Ring_theory. Module abs_ring. @@ -28,7 +28,7 @@ Admitted. Definition cRth : ring_theory c0 c1 cadd cmul csub copp ceq. Admitted. -Add New Ring CoefRing : cRth Abstract. +Add Ring CoefRing : cRth. End abs_ring. Import abs_ring. @@ -36,5 +36,5 @@ Import abs_ring. Theorem check_setoid_ring_modules : forall a b, ceq (cadd a b) (cadd b a). intros. -setoid ring. +ring. Qed. |