summaryrefslogtreecommitdiff
path: root/plugins/micromega
diff options
context:
space:
mode:
authorGravatar Benjamin Barenblat <bbaren@debian.org>2018-12-29 14:31:27 -0500
committerGravatar Benjamin Barenblat <bbaren@debian.org>2018-12-29 14:31:27 -0500
commit9043add656177eeac1491a73d2f3ab92bec0013c (patch)
tree2b0092c84bfbf718eca10c81f60b2640dc8cab05 /plugins/micromega
parenta4c7f8bd98be2a200489325ff7c5061cf80ab4f3 (diff)
Imported Upstream version 8.8.2upstream/8.8.2
Diffstat (limited to 'plugins/micromega')
-rw-r--r--plugins/micromega/Env.v10
-rw-r--r--plugins/micromega/EnvRing.v26
-rw-r--r--plugins/micromega/Lia.v10
-rw-r--r--plugins/micromega/Lqa.v10
-rw-r--r--plugins/micromega/Lra.v10
-rw-r--r--plugins/micromega/MExtraction.v33
-rw-r--r--plugins/micromega/OrderedRing.v10
-rw-r--r--plugins/micromega/Psatz.v10
-rw-r--r--plugins/micromega/QMicromega.v10
-rw-r--r--plugins/micromega/RMicromega.v325
-rw-r--r--plugins/micromega/Refl.v10
-rw-r--r--plugins/micromega/RingMicromega.v13
-rw-r--r--plugins/micromega/Tauto.v10
-rw-r--r--plugins/micromega/ZCoeff.v10
-rw-r--r--plugins/micromega/ZMicromega.v10
-rw-r--r--plugins/micromega/certificate.ml10
-rw-r--r--plugins/micromega/coq_micromega.ml641
-rw-r--r--plugins/micromega/csdpcert.ml10
-rw-r--r--plugins/micromega/g_micromega.ml416
-rw-r--r--plugins/micromega/mfourier.ml12
-rw-r--r--plugins/micromega/micromega.ml342
-rw-r--r--plugins/micromega/micromega.mli139
-rw-r--r--plugins/micromega/mutils.ml10
-rw-r--r--plugins/micromega/persistent_cache.ml14
-rw-r--r--plugins/micromega/polynomial.ml10
-rw-r--r--plugins/micromega/sos.ml270
-rw-r--r--plugins/micromega/sos.mli10
-rw-r--r--plugins/micromega/sos_lib.ml4
-rw-r--r--plugins/micromega/sos_types.ml10
-rw-r--r--plugins/micromega/sos_types.mli42
-rw-r--r--plugins/micromega/vo.itarget15
31 files changed, 808 insertions, 1254 deletions
diff --git a/plugins/micromega/Env.v b/plugins/micromega/Env.v
index 7e3ef892..10326990 100644
--- a/plugins/micromega/Env.v
+++ b/plugins/micromega/Env.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* *)
(* Micromega: A reflexive tactic using the Positivstellensatz *)
diff --git a/plugins/micromega/EnvRing.v b/plugins/micromega/EnvRing.v
index 904ee4da..4042959b 100644
--- a/plugins/micromega/EnvRing.v
+++ b/plugins/micromega/EnvRing.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* F. Besson: to evaluate polynomials, the original code is using a list.
For big polynomials, this is inefficient -- linear access.
@@ -56,10 +58,18 @@ Section MakeRingPol.
Infix "?=!" := ceqb. Notation "[ x ]" := (phi x).
(* Useful tactics *)
- Add Morphism radd : radd_ext. exact (Radd_ext Reqe). Qed.
- Add Morphism rmul : rmul_ext. exact (Rmul_ext Reqe). Qed.
- Add Morphism ropp : ropp_ext. exact (Ropp_ext Reqe). Qed.
- Add Morphism rsub : rsub_ext. exact (ARsub_ext Rsth Reqe ARth). Qed.
+ Add Morphism radd with signature (req ==> req ==> req) as radd_ext.
+ Proof. exact (Radd_ext Reqe). Qed.
+
+ Add Morphism rmul with signature (req ==> req ==> req) as rmul_ext.
+ Proof. exact (Rmul_ext Reqe). Qed.
+
+ Add Morphism ropp with signature (req ==> req) as ropp_ext.
+ Proof. exact (Ropp_ext Reqe). Qed.
+
+ Add Morphism rsub with signature (req ==> req ==> req) as rsub_ext.
+ Proof. exact (ARsub_ext Rsth Reqe ARth). Qed.
+
Ltac rsimpl := gen_srewrite Rsth Reqe ARth.
Ltac add_push := gen_add_push radd Rsth Reqe ARth.
diff --git a/plugins/micromega/Lia.v b/plugins/micromega/Lia.v
index 47b6f7c7..ae05cf54 100644
--- a/plugins/micromega/Lia.v
+++ b/plugins/micromega/Lia.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* *)
(* Micromega: A reflexive tactic using the Positivstellensatz *)
diff --git a/plugins/micromega/Lqa.v b/plugins/micromega/Lqa.v
index acd2751a..caaec541 100644
--- a/plugins/micromega/Lqa.v
+++ b/plugins/micromega/Lqa.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* *)
(* Micromega: A reflexive tactic using the Positivstellensatz *)
diff --git a/plugins/micromega/Lra.v b/plugins/micromega/Lra.v
index 5b97d8ed..4ff483fb 100644
--- a/plugins/micromega/Lra.v
+++ b/plugins/micromega/Lra.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* *)
(* Micromega: A reflexive tactic using the Positivstellensatz *)
diff --git a/plugins/micromega/MExtraction.v b/plugins/micromega/MExtraction.v
index d28bb828..158ddb58 100644
--- a/plugins/micromega/MExtraction.v
+++ b/plugins/micromega/MExtraction.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* *)
(* Micromega: A reflexive tactic using the Positivstellensatz *)
@@ -14,6 +16,7 @@
(* Used to generate micromega.ml *)
+Require Extraction.
Require Import ZMicromega.
Require Import QMicromega.
Require Import RMicromega.
@@ -38,23 +41,23 @@ Extract Inductive sumor => option [ Some None ].
Let's rather use the ocaml && *)
Extract Inlined Constant andb => "(&&)".
-Require Import Reals.
+Import Reals.Rdefinitions.
-Extract Constant R => "int".
-Extract Constant R0 => "0".
-Extract Constant R1 => "1".
+Extract Constant R => "int".
+Extract Constant R0 => "0".
+Extract Constant R1 => "1".
Extract Constant Rplus => "( + )".
Extract Constant Rmult => "( * )".
Extract Constant Ropp => "fun x -> - x".
Extract Constant Rinv => "fun x -> 1 / x".
-Extraction "micromega.ml"
- List.map simpl_cone (*map_cone indexes*)
- denorm Qpower vm_add
- n_of_Z N.of_nat ZTautoChecker ZWeakChecker QTautoChecker RTautoChecker find.
-
-
-
+(** In order to avoid annoying build dependencies the actual
+ extraction is only performed as a test in the test suite. *)
+(* Extraction "plugins/micromega/micromega.ml" *)
+(* Recursive Extraction *)
+(* List.map simpl_cone (*map_cone indexes*) *)
+(* denorm Qpower vm_add *)
+(* n_of_Z N.of_nat ZTautoChecker ZWeakChecker QTautoChecker RTautoChecker find. *)
(* Local Variables: *)
(* coding: utf-8 *)
diff --git a/plugins/micromega/OrderedRing.v b/plugins/micromega/OrderedRing.v
index 72b4dcb6..62505453 100644
--- a/plugins/micromega/OrderedRing.v
+++ b/plugins/micromega/OrderedRing.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
diff --git a/plugins/micromega/Psatz.v b/plugins/micromega/Psatz.v
index 8acf0ff8..28234e7a 100644
--- a/plugins/micromega/Psatz.v
+++ b/plugins/micromega/Psatz.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* *)
(* Micromega: A reflexive tactic using the Positivstellensatz *)
diff --git a/plugins/micromega/QMicromega.v b/plugins/micromega/QMicromega.v
index b13285f5..ddf4064a 100644
--- a/plugins/micromega/QMicromega.v
+++ b/plugins/micromega/QMicromega.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* *)
(* Micromega: A reflexive tactic using the Positivstellensatz *)
diff --git a/plugins/micromega/RMicromega.v b/plugins/micromega/RMicromega.v
index 2352d78d..c2b40c73 100644
--- a/plugins/micromega/RMicromega.v
+++ b/plugins/micromega/RMicromega.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* *)
(* Micromega: A reflexive tactic using the Positivstellensatz *)
@@ -18,7 +20,7 @@ Require Import Refl.
Require Import Raxioms RIneq Rpow_def DiscrR.
Require Import QArith.
Require Import Qfield.
-
+Require Import Qreals.
Require Setoid.
(*Declare ML Module "micromega_plugin".*)
@@ -38,15 +40,8 @@ Proof.
exact Rplus_opp_r.
Qed.
-Add Ring Rring : Rsrt.
Open Scope R_scope.
-Lemma Rmult_neutral : forall x:R , 0 * x = 0.
-Proof.
- intro ; ring.
-Qed.
-
-
Lemma Rsor : SOR R0 R1 Rplus Rmult Rminus Ropp (@eq R) Rle Rlt.
Proof.
constructor; intros ; subst ; try (intuition (subst; try ring ; auto with real)).
@@ -59,142 +54,41 @@ Proof.
apply (Rlt_irrefl m) ; auto.
apply Rnot_le_lt. auto with real.
destruct (total_order_T n m) as [ [H1 | H1] | H1] ; auto.
- intros.
- rewrite <- (Rmult_neutral m).
- apply (Rmult_lt_compat_r) ; auto.
-Qed.
-
-Definition IQR := fun x : Q => (IZR (Qnum x) * / IZR (' Qden x))%R.
-
-
-Lemma Rinv_elim : forall x y z,
- y <> 0 -> (z * y = x <-> x * / y = z).
-Proof.
- intros.
- split ; intros.
- subst.
- rewrite Rmult_assoc.
- rewrite Rinv_r; auto.
- ring.
- subst.
- rewrite Rmult_assoc.
- rewrite (Rmult_comm (/ y)).
- rewrite Rinv_r ; auto.
- ring.
-Qed.
-
-Ltac INR_nat_of_P :=
- match goal with
- | H : context[INR (Pos.to_nat ?X)] |- _ =>
- revert H ;
- let HH := fresh in
- assert (HH := pos_INR_nat_of_P X) ; revert HH ; generalize (INR (Pos.to_nat X))
- | |- context[INR (Pos.to_nat ?X)] =>
- let HH := fresh in
- assert (HH := pos_INR_nat_of_P X) ; revert HH ; generalize (INR (Pos.to_nat X))
- end.
-
-Ltac add_eq expr val := set (temp := expr) ;
- generalize (eq_refl temp) ;
- unfold temp at 1 ; generalize temp ; intro val ; clear temp.
-
-Ltac Rinv_elim :=
- match goal with
- | |- context[?x * / ?y] =>
- let z := fresh "v" in
- add_eq (x * / y) z ;
- let H := fresh in intro H ; rewrite <- Rinv_elim in H
- end.
-
-Lemma Rlt_neq : forall r , 0 < r -> r <> 0.
-Proof.
- red. intros.
- subst.
- apply (Rlt_irrefl 0 H).
+ now apply Rmult_lt_0_compat.
Qed.
+Notation IQR := Q2R (only parsing).
Lemma Rinv_1 : forall x, x * / 1 = x.
Proof.
intro.
- Rinv_elim.
- subst ; ring.
- apply R1_neq_R0.
+ rewrite Rinv_1.
+ apply Rmult_1_r.
Qed.
-Lemma Qeq_true : forall x y,
- Qeq_bool x y = true ->
- IQR x = IQR y.
+Lemma Qeq_true : forall x y, Qeq_bool x y = true -> IQR x = IQR y.
Proof.
- unfold IQR.
- simpl.
- intros.
- apply Qeq_bool_eq in H.
- unfold Qeq in H.
- assert (IZR (Qnum x * ' Qden y) = IZR (Qnum y * ' Qden x))%Z.
- rewrite H. reflexivity.
- repeat rewrite mult_IZR in H0.
- simpl in H0.
- revert H0.
- repeat INR_nat_of_P.
intros.
- apply Rinv_elim in H2 ; [| apply Rlt_neq ; auto].
- rewrite <- H2.
- field.
- split ; apply Rlt_neq ; auto.
+ now apply Qeq_eqR, Qeq_bool_eq.
Qed.
Lemma Qeq_false : forall x y, Qeq_bool x y = false -> IQR x <> IQR y.
Proof.
intros.
- apply Qeq_bool_neq in H.
- intro. apply H. clear H.
- unfold Qeq,IQR in *.
- simpl in *.
- revert H0.
- repeat Rinv_elim.
- intros.
- subst.
- assert (IZR (Qnum x * ' Qden y)%Z = IZR (Qnum y * ' Qden x)%Z).
- repeat rewrite mult_IZR.
- simpl.
- rewrite <- H0. rewrite <- H.
- ring.
- apply eq_IZR ; auto.
- INR_nat_of_P; intros; apply Rlt_neq ; auto.
- INR_nat_of_P; intros ; apply Rlt_neq ; auto.
+ apply Qeq_bool_neq in H.
+ contradict H.
+ now apply eqR_Qeq.
Qed.
-
-
Lemma Qle_true : forall x y : Q, Qle_bool x y = true -> IQR x <= IQR y.
Proof.
intros.
- apply Qle_bool_imp_le in H.
- unfold Qle in H.
- unfold IQR.
- simpl in *.
- apply IZR_le in H.
- repeat rewrite mult_IZR in H.
- simpl in H.
- repeat INR_nat_of_P; intros.
- assert (Hr := Rlt_neq r H).
- assert (Hr0 := Rlt_neq r0 H0).
- replace (IZR (Qnum x) * / r) with ((IZR (Qnum x) * r0) * (/r * /r0)).
- replace (IZR (Qnum y) * / r0) with ((IZR (Qnum y) * r) * (/r * /r0)).
- apply Rmult_le_compat_r ; auto.
- apply Rmult_le_pos.
- unfold Rle. left. apply Rinv_0_lt_compat ; auto.
- unfold Rle. left. apply Rinv_0_lt_compat ; auto.
- field ; intuition.
- field ; intuition.
+ now apply Qle_Rle, Qle_bool_imp_le.
Qed.
-
-
Lemma IQR_0 : IQR 0 = 0.
Proof.
- compute. apply Rinv_1.
+ apply Rmult_0_l.
Qed.
Lemma IQR_1 : IQR 1 = 1.
@@ -202,160 +96,6 @@ Proof.
compute. apply Rinv_1.
Qed.
-Lemma IQR_plus : forall x y, IQR (x + y) = IQR x + IQR y.
-Proof.
- intros.
- unfold IQR.
- simpl in *.
- rewrite plus_IZR in *.
- rewrite mult_IZR in *.
- simpl.
- rewrite Pos2Nat.inj_mul.
- rewrite mult_INR.
- rewrite mult_IZR.
- simpl.
- repeat INR_nat_of_P.
- intros. field.
- split ; apply Rlt_neq ; auto.
-Qed.
-
-Lemma IQR_opp : forall x, IQR (- x) = - IQR x.
-Proof.
- intros.
- unfold IQR.
- simpl.
- rewrite opp_IZR.
- ring.
-Qed.
-
-Lemma IQR_minus : forall x y, IQR (x - y) = IQR x - IQR y.
-Proof.
- intros.
- unfold Qminus.
- rewrite IQR_plus.
- rewrite IQR_opp.
- ring.
-Qed.
-
-
-Lemma IQR_mult : forall x y, IQR (x * y) = IQR x * IQR y.
-Proof.
- unfold IQR ; intros.
- simpl.
- repeat rewrite mult_IZR.
- rewrite Pos2Nat.inj_mul.
- rewrite mult_INR.
- repeat INR_nat_of_P.
- intros. field ; split ; apply Rlt_neq ; auto.
-Qed.
-
-Lemma IQR_inv_lt : forall x, (0 < x)%Q ->
- IQR (/ x) = / IQR x.
-Proof.
- unfold IQR ; simpl.
- intros.
- unfold Qlt in H.
- revert H.
- simpl.
- intros.
- unfold Qinv.
- destruct x.
- destruct Qnum ; simpl in *.
- exfalso. auto with zarith.
- clear H.
- repeat INR_nat_of_P.
- intros.
- assert (HH := Rlt_neq _ H).
- assert (HH0 := Rlt_neq _ H0).
- rewrite Rinv_mult_distr ; auto.
- rewrite Rinv_involutive ; auto.
- ring.
- apply Rinv_0_lt_compat in H0.
- apply Rlt_neq ; auto.
- simpl in H.
- exfalso.
- rewrite Pos.mul_comm in H.
- compute in H.
- discriminate.
-Qed.
-
-Lemma Qinv_opp : forall x, (- (/ x) = / ( -x))%Q.
-Proof.
- destruct x ; destruct Qnum ; reflexivity.
-Qed.
-
-Lemma Qopp_involutive_strong : forall x, (- - x = x)%Q.
-Proof.
- intros.
- destruct x.
- unfold Qopp.
- simpl.
- rewrite Z.opp_involutive.
- reflexivity.
-Qed.
-
-Lemma Ropp_0 : forall r , - r = 0 -> r = 0.
-Proof.
- intros.
- rewrite <- (Ropp_involutive r).
- apply Ropp_eq_0_compat ; auto.
-Qed.
-
-Lemma IQR_x_0 : forall x, IQR x = 0 -> x == 0%Q.
-Proof.
- destruct x ; simpl.
- unfold IQR.
- simpl.
- INR_nat_of_P.
- intros.
- apply Rmult_integral in H0.
- destruct H0.
- apply eq_IZR_R0 in H0.
- subst.
- reflexivity.
- exfalso.
- apply Rinv_0_lt_compat in H.
- rewrite <- H0 in H.
- apply Rlt_irrefl in H. auto.
-Qed.
-
-
-Lemma IQR_inv_gt : forall x, (0 > x)%Q ->
- IQR (/ x) = / IQR x.
-Proof.
- intros.
- rewrite <- (Qopp_involutive_strong x).
- rewrite <- Qinv_opp.
- rewrite IQR_opp.
- rewrite IQR_inv_lt.
- repeat rewrite IQR_opp.
- rewrite Ropp_inv_permute.
- auto.
- intro.
- apply Ropp_0 in H0.
- apply IQR_x_0 in H0.
- rewrite H0 in H.
- compute in H. discriminate.
- unfold Qlt in *.
- destruct x ; simpl in *.
- auto with zarith.
-Qed.
-
-Lemma IQR_inv : forall x, ~ x == 0 ->
- IQR (/ x) = / IQR x.
-Proof.
- intros.
- assert ( 0 > x \/ 0 < x)%Q.
- destruct x ; unfold Qlt, Qeq in * ; simpl in *.
- rewrite Z.mul_1_r in *.
- destruct Qnum ; simpl in * ; intuition auto.
- right. reflexivity.
- left ; reflexivity.
- destruct H0.
- apply IQR_inv_gt ; auto.
- apply IQR_inv_lt ; auto.
-Qed.
-
Lemma IQR_inv_ext : forall x,
IQR (/ x) = (if Qeq_bool x 0 then 0 else / IQR x).
Proof.
@@ -366,18 +106,13 @@ Proof.
destruct x ; simpl.
unfold Qeq in H.
simpl in H.
- replace Qnum with 0%Z.
- compute. rewrite Rinv_1.
- reflexivity.
- rewrite <- H. ring.
+ rewrite Zmult_1_r in H.
+ rewrite H.
+ apply Rmult_0_l.
intros.
- apply IQR_inv.
- intro.
- rewrite <- Qeq_bool_iff in H0.
- congruence.
+ now apply Q2R_inv, Qeq_bool_neq.
Qed.
-
Notation to_nat := N.to_nat.
Lemma QSORaddon :
@@ -391,10 +126,10 @@ Proof.
constructor ; intros ; try reflexivity.
apply IQR_0.
apply IQR_1.
- apply IQR_plus.
- apply IQR_minus.
- apply IQR_mult.
- apply IQR_opp.
+ apply Q2R_plus.
+ apply Q2R_minus.
+ apply Q2R_mult.
+ apply Q2R_opp.
apply Qeq_true ; auto.
apply R_power_theory.
apply Qeq_false.
@@ -453,13 +188,13 @@ Proof.
apply IQR_1.
reflexivity.
unfold IQR. simpl. rewrite Rinv_1. reflexivity.
- apply IQR_plus.
- apply IQR_minus.
- apply IQR_mult.
+ apply Q2R_plus.
+ apply Q2R_minus.
+ apply Q2R_mult.
rewrite <- IHc.
apply IQR_inv_ext.
rewrite <- IHc.
- apply IQR_opp.
+ apply Q2R_opp.
Qed.
Require Import EnvRing.
diff --git a/plugins/micromega/Refl.v b/plugins/micromega/Refl.v
index 32ddb3cf..952a1b91 100644
--- a/plugins/micromega/Refl.v
+++ b/plugins/micromega/Refl.v
@@ -1,10 +1,12 @@
(* -*- coding: utf-8 -*- *)
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* *)
(* Micromega: A reflexive tactic using the Positivstellensatz *)
diff --git a/plugins/micromega/RingMicromega.v b/plugins/micromega/RingMicromega.v
index ed49c3df..f066ea46 100644
--- a/plugins/micromega/RingMicromega.v
+++ b/plugins/micromega/RingMicromega.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
@@ -19,6 +21,7 @@ Require Import List.
Require Import Bool.
Require Import OrderedRing.
Require Import Refl.
+Require Coq.micromega.Tauto.
Set Implicit Arguments.
@@ -794,7 +797,7 @@ Definition xnormalise (t:Formula C) : list (NFormula) :=
| OpLe => (psub lhs rhs ,Strict) :: nil
end.
-Require Import Coq.micromega.Tauto.
+Import Coq.micromega.Tauto.
Definition cnf_normalise (t:Formula C) : cnf (NFormula) :=
List.map (fun x => x::nil) (xnormalise t).
diff --git a/plugins/micromega/Tauto.v b/plugins/micromega/Tauto.v
index 391231af..31f55ae9 100644
--- a/plugins/micromega/Tauto.v
+++ b/plugins/micromega/Tauto.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* *)
(* Micromega: A reflexive tactic using the Positivstellensatz *)
diff --git a/plugins/micromega/ZCoeff.v b/plugins/micromega/ZCoeff.v
index bd425e6b..137453a9 100644
--- a/plugins/micromega/ZCoeff.v
+++ b/plugins/micromega/ZCoeff.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
diff --git a/plugins/micromega/ZMicromega.v b/plugins/micromega/ZMicromega.v
index 5aa8d03f..892858e6 100644
--- a/plugins/micromega/ZMicromega.v
+++ b/plugins/micromega/ZMicromega.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* *)
(* Micromega: A reflexive tactic using the Positivstellensatz *)
diff --git a/plugins/micromega/certificate.ml b/plugins/micromega/certificate.ml
index 459c72f9..9f39191f 100644
--- a/plugins/micromega/certificate.ml
+++ b/plugins/micromega/certificate.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* *)
(* Micromega: A reflexive tactic using the Positivstellensatz *)
diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml
index e4b58a56..168105e8 100644
--- a/plugins/micromega/coq_micromega.ml
+++ b/plugins/micromega/coq_micromega.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* *)
(* Micromega: A reflexive tactic using the Positivstellensatz *)
@@ -19,6 +21,8 @@
open Pp
open Mutils
open Goptions
+open Names
+open Constr
(**
* Debug flag
@@ -63,7 +67,6 @@ let _ =
let int_opt l vref =
{
- optsync = true;
optdepr = false;
optname = List.fold_right (^) l "";
optkey = l ;
@@ -73,7 +76,6 @@ let _ =
let lia_enum_opt =
{
- optsync = true;
optdepr = false;
optname = "Lia Enum";
optkey = ["Lia";"Enum"];
@@ -109,8 +111,8 @@ type 'cst atom = 'cst Micromega.formula
type 'cst formula =
| TT
| FF
- | X of Term.constr
- | A of 'cst atom * tag * Term.constr
+ | X of EConstr.constr
+ | A of 'cst atom * tag * EConstr.constr
| C of 'cst formula * 'cst formula
| D of 'cst formula * 'cst formula
| N of 'cst formula
@@ -328,9 +330,6 @@ let selecti s m =
module M =
struct
- open Coqlib
- open Term
-
(**
* Location of the Coq libraries.
*)
@@ -355,8 +354,8 @@ struct
["LRing_normalise"]]
let coq_modules =
- init_modules @
- [logic_dir] @ arith_modules @ zarith_base_modules @ mic_modules
+ Coqlib.(init_modules @
+ [logic_dir] @ arith_modules @ zarith_base_modules @ mic_modules)
let bin_module = [["Coq";"Numbers";"BinNums"]]
@@ -364,6 +363,7 @@ struct
[["Coq";"Reals" ; "Rdefinitions"];
["Coq";"Reals" ; "Rpow_def"] ;
["Coq";"Reals" ; "Raxioms"] ;
+ ["Coq";"QArith"; "Qreals"] ;
]
let z_modules = [["Coq";"ZArith";"BinInt"]]
@@ -373,7 +373,8 @@ struct
* ZMicromega.v
*)
- let init_constant = gen_constant_in_modules "ZMicromega" init_modules
+ let gen_constant_in_modules s m n = EConstr.of_constr (Universes.constr_of_global @@ Coqlib.gen_reference_in_modules s m n)
+ let init_constant = gen_constant_in_modules "ZMicromega" Coqlib.init_modules
let constant = gen_constant_in_modules "ZMicromega" coq_modules
let bin_constant = gen_constant_in_modules "ZMicromega" bin_module
let r_constant = gen_constant_in_modules "ZMicromega" r_modules
@@ -383,7 +384,6 @@ struct
let coq_and = lazy (init_constant "and")
let coq_or = lazy (init_constant "or")
let coq_not = lazy (init_constant "not")
- let coq_not_gl_ref = (Nametab.locate ( Libnames.qualid_of_string "Coq.Init.Logic.not"))
let coq_iff = lazy (init_constant "iff")
let coq_True = lazy (init_constant "True")
@@ -480,7 +480,7 @@ struct
let coq_Rinv = lazy (r_constant "Rinv")
let coq_Rpower = lazy (r_constant "pow")
let coq_IZR = lazy (r_constant "IZR")
- let coq_IQR = lazy (constant "IQR")
+ let coq_IQR = lazy (r_constant "Q2R")
let coq_PEX = lazy (constant "PEX" )
@@ -583,9 +583,9 @@ struct
| Ukn
| BadStr of string
| BadNum of int
- | BadTerm of Term.constr
+ | BadTerm of constr
| Msg of string
- | Goal of (Term.constr list ) * Term.constr * parse_error
+ | Goal of (constr list ) * constr * parse_error
let string_of_error = function
| Ukn -> "ukn"
@@ -599,11 +599,11 @@ struct
(* A simple but useful getter function *)
- let get_left_construct term =
- match Term.kind_of_term term with
+ let get_left_construct sigma term =
+ match EConstr.kind sigma term with
| Term.Construct((_,i),_) -> (i,[| |])
| Term.App(l,rst) ->
- (match Term.kind_of_term l with
+ (match EConstr.kind sigma l with
| Term.Construct((_,i),_) -> (i,rst)
| _ -> raise ParseError
)
@@ -613,11 +613,11 @@ struct
(* parse/dump/print from numbers up to expressions and formulas *)
- let rec parse_nat term =
- let (i,c) = get_left_construct term in
+ let rec parse_nat sigma term =
+ let (i,c) = get_left_construct sigma term in
match i with
| 1 -> Mc.O
- | 2 -> Mc.S (parse_nat (c.(0)))
+ | 2 -> Mc.S (parse_nat sigma (c.(0)))
| i -> raise ParseError
let pp_nat o n = Printf.fprintf o "%i" (CoqToCaml.nat n)
@@ -625,71 +625,71 @@ struct
let rec dump_nat x =
match x with
| Mc.O -> Lazy.force coq_O
- | Mc.S p -> Term.mkApp(Lazy.force coq_S,[| dump_nat p |])
+ | Mc.S p -> EConstr.mkApp(Lazy.force coq_S,[| dump_nat p |])
- let rec parse_positive term =
- let (i,c) = get_left_construct term in
+ let rec parse_positive sigma term =
+ let (i,c) = get_left_construct sigma term in
match i with
- | 1 -> Mc.XI (parse_positive c.(0))
- | 2 -> Mc.XO (parse_positive c.(0))
+ | 1 -> Mc.XI (parse_positive sigma c.(0))
+ | 2 -> Mc.XO (parse_positive sigma c.(0))
| 3 -> Mc.XH
| i -> raise ParseError
let rec dump_positive x =
match x with
| Mc.XH -> Lazy.force coq_xH
- | Mc.XO p -> Term.mkApp(Lazy.force coq_xO,[| dump_positive p |])
- | Mc.XI p -> Term.mkApp(Lazy.force coq_xI,[| dump_positive p |])
+ | Mc.XO p -> EConstr.mkApp(Lazy.force coq_xO,[| dump_positive p |])
+ | Mc.XI p -> EConstr.mkApp(Lazy.force coq_xI,[| dump_positive p |])
let pp_positive o x = Printf.fprintf o "%i" (CoqToCaml.positive x)
let dump_n x =
match x with
| Mc.N0 -> Lazy.force coq_N0
- | Mc.Npos p -> Term.mkApp(Lazy.force coq_Npos,[| dump_positive p|])
+ | Mc.Npos p -> EConstr.mkApp(Lazy.force coq_Npos,[| dump_positive p|])
let rec dump_index x =
match x with
| Mc.XH -> Lazy.force coq_xH
- | Mc.XO p -> Term.mkApp(Lazy.force coq_xO,[| dump_index p |])
- | Mc.XI p -> Term.mkApp(Lazy.force coq_xI,[| dump_index p |])
+ | Mc.XO p -> EConstr.mkApp(Lazy.force coq_xO,[| dump_index p |])
+ | Mc.XI p -> EConstr.mkApp(Lazy.force coq_xI,[| dump_index p |])
let pp_index o x = Printf.fprintf o "%i" (CoqToCaml.index x)
let pp_n o x = output_string o (string_of_int (CoqToCaml.n x))
let dump_pair t1 t2 dump_t1 dump_t2 (x,y) =
- Term.mkApp(Lazy.force coq_pair,[| t1 ; t2 ; dump_t1 x ; dump_t2 y|])
+ EConstr.mkApp(Lazy.force coq_pair,[| t1 ; t2 ; dump_t1 x ; dump_t2 y|])
- let parse_z term =
- let (i,c) = get_left_construct term in
+ let parse_z sigma term =
+ let (i,c) = get_left_construct sigma term in
match i with
| 1 -> Mc.Z0
- | 2 -> Mc.Zpos (parse_positive c.(0))
- | 3 -> Mc.Zneg (parse_positive c.(0))
+ | 2 -> Mc.Zpos (parse_positive sigma c.(0))
+ | 3 -> Mc.Zneg (parse_positive sigma c.(0))
| i -> raise ParseError
let dump_z x =
match x with
| Mc.Z0 ->Lazy.force coq_ZERO
- | Mc.Zpos p -> Term.mkApp(Lazy.force coq_POS,[| dump_positive p|])
- | Mc.Zneg p -> Term.mkApp(Lazy.force coq_NEG,[| dump_positive p|])
+ | Mc.Zpos p -> EConstr.mkApp(Lazy.force coq_POS,[| dump_positive p|])
+ | Mc.Zneg p -> EConstr.mkApp(Lazy.force coq_NEG,[| dump_positive p|])
let pp_z o x = Printf.fprintf o "%s" (Big_int.string_of_big_int (CoqToCaml.z_big_int x))
let dump_num bd1 =
- Term.mkApp(Lazy.force coq_Qmake,
- [|dump_z (CamlToCoq.bigint (numerator bd1)) ;
- dump_positive (CamlToCoq.positive_big_int (denominator bd1)) |])
+ EConstr.mkApp(Lazy.force coq_Qmake,
+ [|dump_z (CamlToCoq.bigint (numerator bd1)) ;
+ dump_positive (CamlToCoq.positive_big_int (denominator bd1)) |])
let dump_q q =
- Term.mkApp(Lazy.force coq_Qmake,
- [| dump_z q.Micromega.qnum ; dump_positive q.Micromega.qden|])
+ EConstr.mkApp(Lazy.force coq_Qmake,
+ [| dump_z q.Micromega.qnum ; dump_positive q.Micromega.qden|])
- let parse_q term =
- match Term.kind_of_term term with
- | Term.App(c, args) -> if Constr.equal c (Lazy.force coq_Qmake) then
- {Mc.qnum = parse_z args.(0) ; Mc.qden = parse_positive args.(1) }
+ let parse_q sigma term =
+ match EConstr.kind sigma term with
+ | Term.App(c, args) -> if EConstr.eq_constr sigma c (Lazy.force coq_Qmake) then
+ {Mc.qnum = parse_z sigma args.(0) ; Mc.qden = parse_positive sigma args.(1) }
else raise ParseError
| _ -> raise ParseError
@@ -711,41 +711,41 @@ struct
match cst with
| Mc.C0 -> Lazy.force coq_C0
| Mc.C1 -> Lazy.force coq_C1
- | Mc.CQ q -> Term.mkApp(Lazy.force coq_CQ, [| dump_q q |])
- | Mc.CZ z -> Term.mkApp(Lazy.force coq_CZ, [| dump_z z |])
- | Mc.CPlus(x,y) -> Term.mkApp(Lazy.force coq_CPlus, [| dump_Rcst x ; dump_Rcst y |])
- | Mc.CMinus(x,y) -> Term.mkApp(Lazy.force coq_CMinus, [| dump_Rcst x ; dump_Rcst y |])
- | Mc.CMult(x,y) -> Term.mkApp(Lazy.force coq_CMult, [| dump_Rcst x ; dump_Rcst y |])
- | Mc.CInv t -> Term.mkApp(Lazy.force coq_CInv, [| dump_Rcst t |])
- | Mc.COpp t -> Term.mkApp(Lazy.force coq_COpp, [| dump_Rcst t |])
-
- let rec parse_Rcst term =
- let (i,c) = get_left_construct term in
+ | Mc.CQ q -> EConstr.mkApp(Lazy.force coq_CQ, [| dump_q q |])
+ | Mc.CZ z -> EConstr.mkApp(Lazy.force coq_CZ, [| dump_z z |])
+ | Mc.CPlus(x,y) -> EConstr.mkApp(Lazy.force coq_CPlus, [| dump_Rcst x ; dump_Rcst y |])
+ | Mc.CMinus(x,y) -> EConstr.mkApp(Lazy.force coq_CMinus, [| dump_Rcst x ; dump_Rcst y |])
+ | Mc.CMult(x,y) -> EConstr.mkApp(Lazy.force coq_CMult, [| dump_Rcst x ; dump_Rcst y |])
+ | Mc.CInv t -> EConstr.mkApp(Lazy.force coq_CInv, [| dump_Rcst t |])
+ | Mc.COpp t -> EConstr.mkApp(Lazy.force coq_COpp, [| dump_Rcst t |])
+
+ let rec parse_Rcst sigma term =
+ let (i,c) = get_left_construct sigma term in
match i with
| 1 -> Mc.C0
| 2 -> Mc.C1
- | 3 -> Mc.CQ (parse_q c.(0))
- | 4 -> Mc.CPlus(parse_Rcst c.(0), parse_Rcst c.(1))
- | 5 -> Mc.CMinus(parse_Rcst c.(0), parse_Rcst c.(1))
- | 6 -> Mc.CMult(parse_Rcst c.(0), parse_Rcst c.(1))
- | 7 -> Mc.CInv(parse_Rcst c.(0))
- | 8 -> Mc.COpp(parse_Rcst c.(0))
+ | 3 -> Mc.CQ (parse_q sigma c.(0))
+ | 4 -> Mc.CPlus(parse_Rcst sigma c.(0), parse_Rcst sigma c.(1))
+ | 5 -> Mc.CMinus(parse_Rcst sigma c.(0), parse_Rcst sigma c.(1))
+ | 6 -> Mc.CMult(parse_Rcst sigma c.(0), parse_Rcst sigma c.(1))
+ | 7 -> Mc.CInv(parse_Rcst sigma c.(0))
+ | 8 -> Mc.COpp(parse_Rcst sigma c.(0))
| _ -> raise ParseError
- let rec parse_list parse_elt term =
- let (i,c) = get_left_construct term in
+ let rec parse_list sigma parse_elt term =
+ let (i,c) = get_left_construct sigma term in
match i with
| 1 -> []
- | 2 -> parse_elt c.(1) :: parse_list parse_elt c.(2)
+ | 2 -> parse_elt sigma c.(1) :: parse_list sigma parse_elt c.(2)
| i -> raise ParseError
let rec dump_list typ dump_elt l =
match l with
- | [] -> Term.mkApp(Lazy.force coq_nil,[| typ |])
- | e :: l -> Term.mkApp(Lazy.force coq_cons,
+ | [] -> EConstr.mkApp(Lazy.force coq_nil,[| typ |])
+ | e :: l -> EConstr.mkApp(Lazy.force coq_cons,
[| typ; dump_elt e;dump_list typ dump_elt l|])
let pp_list op cl elt o l =
@@ -775,27 +775,27 @@ struct
let dump_expr typ dump_z e =
let rec dump_expr e =
match e with
- | Mc.PEX n -> mkApp(Lazy.force coq_PEX,[| typ; dump_var n |])
- | Mc.PEc z -> mkApp(Lazy.force coq_PEc,[| typ ; dump_z z |])
- | Mc.PEadd(e1,e2) -> mkApp(Lazy.force coq_PEadd,
- [| typ; dump_expr e1;dump_expr e2|])
- | Mc.PEsub(e1,e2) -> mkApp(Lazy.force coq_PEsub,
- [| typ; dump_expr e1;dump_expr e2|])
- | Mc.PEopp e -> mkApp(Lazy.force coq_PEopp,
- [| typ; dump_expr e|])
- | Mc.PEmul(e1,e2) -> mkApp(Lazy.force coq_PEmul,
- [| typ; dump_expr e1;dump_expr e2|])
- | Mc.PEpow(e,n) -> mkApp(Lazy.force coq_PEpow,
- [| typ; dump_expr e; dump_n n|])
+ | Mc.PEX n -> EConstr.mkApp(Lazy.force coq_PEX,[| typ; dump_var n |])
+ | Mc.PEc z -> EConstr.mkApp(Lazy.force coq_PEc,[| typ ; dump_z z |])
+ | Mc.PEadd(e1,e2) -> EConstr.mkApp(Lazy.force coq_PEadd,
+ [| typ; dump_expr e1;dump_expr e2|])
+ | Mc.PEsub(e1,e2) -> EConstr.mkApp(Lazy.force coq_PEsub,
+ [| typ; dump_expr e1;dump_expr e2|])
+ | Mc.PEopp e -> EConstr.mkApp(Lazy.force coq_PEopp,
+ [| typ; dump_expr e|])
+ | Mc.PEmul(e1,e2) -> EConstr.mkApp(Lazy.force coq_PEmul,
+ [| typ; dump_expr e1;dump_expr e2|])
+ | Mc.PEpow(e,n) -> EConstr.mkApp(Lazy.force coq_PEpow,
+ [| typ; dump_expr e; dump_n n|])
in
dump_expr e
let dump_pol typ dump_c e =
let rec dump_pol e =
match e with
- | Mc.Pc n -> mkApp(Lazy.force coq_Pc, [|typ ; dump_c n|])
- | Mc.Pinj(p,pol) -> mkApp(Lazy.force coq_Pinj , [| typ ; dump_positive p ; dump_pol pol|])
- | Mc.PX(pol1,p,pol2) -> mkApp(Lazy.force coq_PX, [| typ ; dump_pol pol1 ; dump_positive p ; dump_pol pol2|]) in
+ | Mc.Pc n -> EConstr.mkApp(Lazy.force coq_Pc, [|typ ; dump_c n|])
+ | Mc.Pinj(p,pol) -> EConstr.mkApp(Lazy.force coq_Pinj , [| typ ; dump_positive p ; dump_pol pol|])
+ | Mc.PX(pol1,p,pol2) -> EConstr.mkApp(Lazy.force coq_PX, [| typ ; dump_pol pol1 ; dump_positive p ; dump_pol pol2|]) in
dump_pol e
let pp_pol pp_c o e =
@@ -814,17 +814,17 @@ struct
let z = Lazy.force typ in
let rec dump_cone e =
match e with
- | Mc.PsatzIn n -> mkApp(Lazy.force coq_PsatzIn,[| z; dump_nat n |])
- | Mc.PsatzMulC(e,c) -> mkApp(Lazy.force coq_PsatzMultC,
- [| z; dump_pol z dump_z e ; dump_cone c |])
- | Mc.PsatzSquare e -> mkApp(Lazy.force coq_PsatzSquare,
- [| z;dump_pol z dump_z e|])
- | Mc.PsatzAdd(e1,e2) -> mkApp(Lazy.force coq_PsatzAdd,
- [| z; dump_cone e1; dump_cone e2|])
- | Mc.PsatzMulE(e1,e2) -> mkApp(Lazy.force coq_PsatzMulE,
- [| z; dump_cone e1; dump_cone e2|])
- | Mc.PsatzC p -> mkApp(Lazy.force coq_PsatzC,[| z; dump_z p|])
- | Mc.PsatzZ -> mkApp( Lazy.force coq_PsatzZ,[| z|]) in
+ | Mc.PsatzIn n -> EConstr.mkApp(Lazy.force coq_PsatzIn,[| z; dump_nat n |])
+ | Mc.PsatzMulC(e,c) -> EConstr.mkApp(Lazy.force coq_PsatzMultC,
+ [| z; dump_pol z dump_z e ; dump_cone c |])
+ | Mc.PsatzSquare e -> EConstr.mkApp(Lazy.force coq_PsatzSquare,
+ [| z;dump_pol z dump_z e|])
+ | Mc.PsatzAdd(e1,e2) -> EConstr.mkApp(Lazy.force coq_PsatzAdd,
+ [| z; dump_cone e1; dump_cone e2|])
+ | Mc.PsatzMulE(e1,e2) -> EConstr.mkApp(Lazy.force coq_PsatzMulE,
+ [| z; dump_cone e1; dump_cone e2|])
+ | Mc.PsatzC p -> EConstr.mkApp(Lazy.force coq_PsatzC,[| z; dump_z p|])
+ | Mc.PsatzZ -> EConstr.mkApp(Lazy.force coq_PsatzZ,[| z|]) in
dump_cone e
let pp_psatz pp_z o e =
@@ -867,14 +867,14 @@ struct
Printf.fprintf o"(%a %a %a)" (pp_expr pp_z) l pp_op op (pp_expr pp_z) r
let dump_cstr typ dump_constant {Mc.flhs = e1 ; Mc.fop = o ; Mc.frhs = e2} =
- Term.mkApp(Lazy.force coq_Build,
- [| typ; dump_expr typ dump_constant e1 ;
- dump_op o ;
- dump_expr typ dump_constant e2|])
+ EConstr.mkApp(Lazy.force coq_Build,
+ [| typ; dump_expr typ dump_constant e1 ;
+ dump_op o ;
+ dump_expr typ dump_constant e2|])
- let assoc_const x l =
+ let assoc_const sigma x l =
try
- snd (List.find (fun (x',y) -> Constr.equal x (Lazy.force x')) l)
+ snd (List.find (fun (x',y) -> EConstr.eq_constr sigma x (Lazy.force x')) l)
with
Not_found -> raise ParseError
@@ -896,38 +896,37 @@ struct
coq_Qeq, Mc.OpEq
]
- let has_typ gl t1 typ =
- let ty = Retyping.get_type_of (Tacmach.pf_env gl) (Tacmach.project gl) t1 in
- Constr.equal ty typ
-
+ type gl = { env : Environ.env; sigma : Evd.evar_map }
let is_convertible gl t1 t2 =
- Reductionops.is_conv (Tacmach.pf_env gl) (Tacmach.project gl) t1 t2
+ Reductionops.is_conv gl.env gl.sigma t1 t2
let parse_zop gl (op,args) =
- match kind_of_term op with
- | Const (x,_) -> (assoc_const op zop_table, args.(0) , args.(1))
- | Ind((n,0),_) ->
- if Constr.equal op (Lazy.force coq_Eq) && is_convertible gl args.(0) (Lazy.force coq_Z)
+ let sigma = gl.sigma in
+ match EConstr.kind sigma op with
+ | Term.Const (x,_) -> (assoc_const sigma op zop_table, args.(0) , args.(1))
+ | Term.Ind((n,0),_) ->
+ if EConstr.eq_constr sigma op (Lazy.force coq_Eq) && is_convertible gl args.(0) (Lazy.force coq_Z)
then (Mc.OpEq, args.(1), args.(2))
else raise ParseError
| _ -> failwith "parse_zop"
let parse_rop gl (op,args) =
- match kind_of_term op with
- | Const (x,_) -> (assoc_const op rop_table, args.(0) , args.(1))
- | Ind((n,0),_) ->
- if Constr.equal op (Lazy.force coq_Eq) && is_convertible gl args.(0) (Lazy.force coq_R)
+ let sigma = gl.sigma in
+ match EConstr.kind sigma op with
+ | Term.Const (x,_) -> (assoc_const sigma op rop_table, args.(0) , args.(1))
+ | Term.Ind((n,0),_) ->
+ if EConstr.eq_constr sigma op (Lazy.force coq_Eq) && is_convertible gl args.(0) (Lazy.force coq_R)
then (Mc.OpEq, args.(1), args.(2))
else raise ParseError
| _ -> failwith "parse_zop"
let parse_qop gl (op,args) =
- (assoc_const op qop_table, args.(0) , args.(1))
+ (assoc_const gl.sigma op qop_table, args.(0) , args.(1))
- let is_constant t = (* This is an approx *)
- match kind_of_term t with
- | Construct(i,_) -> true
+ let is_constant sigma t = (* This is an approx *)
+ match EConstr.kind sigma t with
+ | Term.Construct(i,_) -> true
| _ -> false
type 'a op =
@@ -936,9 +935,9 @@ struct
| Power
| Ukn of string
- let assoc_ops x l =
+ let assoc_ops sigma x l =
try
- snd (List.find (fun (x',y) -> Constr.equal x (Lazy.force x')) l)
+ snd (List.find (fun (x',y) -> EConstr.eq_constr sigma x (Lazy.force x')) l)
with
Not_found -> Ukn "Oups"
@@ -948,14 +947,14 @@ struct
module Env =
struct
- type t = constr list
+ type t = EConstr.constr list
- let compute_rank_add env v =
+ let compute_rank_add env sigma v =
let rec _add env n v =
match env with
| [] -> ([v],n)
| e::l ->
- if eq_constr e v
+ if EConstr.eq_constr sigma e v
then (env,n)
else
let (env,n) = _add l ( n+1) v in
@@ -963,13 +962,13 @@ struct
let (env, n) = _add env 1 v in
(env, CamlToCoq.positive n)
- let get_rank env v =
+ let get_rank env sigma v =
let rec _get_rank env n =
match env with
| [] -> raise (Invalid_argument "get_rank")
| e::l ->
- if eq_constr e v
+ if EConstr.eq_constr sigma e v
then n
else _get_rank l (n+1) in
_get_rank env 1
@@ -985,9 +984,11 @@ struct
* This is the big generic function for expression parsers.
*)
- let parse_expr parse_constant parse_exp ops_spec env term =
+ let parse_expr sigma parse_constant parse_exp ops_spec env term =
if debug
- then Feedback.msg_debug (Pp.str "parse_expr: " ++ Printer.prterm term);
+ then (
+ let _, env = Pfedit.get_current_context () in
+ Feedback.msg_debug (Pp.str "parse_expr: " ++ Printer.pr_leconstr_env env sigma term));
(*
let constant_or_variable env term =
@@ -998,7 +999,7 @@ struct
(Mc.PEX n , env) in
*)
let parse_variable env term =
- let (env,n) = Env.compute_rank_add env term in
+ let (env,n) = Env.compute_rank_add env sigma term in
(Mc.PEX n , env) in
let rec parse_expr env term =
@@ -1009,12 +1010,12 @@ struct
try (Mc.PEc (parse_constant term) , env)
with ParseError ->
- match kind_of_term term with
- | App(t,args) ->
+ match EConstr.kind sigma term with
+ | Term.App(t,args) ->
(
- match kind_of_term t with
- | Const c ->
- ( match assoc_ops t ops_spec with
+ match EConstr.kind sigma t with
+ | Term.Const c ->
+ ( match assoc_ops sigma t ops_spec with
| Binop f -> combine env f (args.(0),args.(1))
| Opp -> let (expr,env) = parse_expr env args.(0) in
(Mc.PEopp expr, env)
@@ -1026,12 +1027,12 @@ struct
(power , env)
with e when CErrors.noncritical e ->
(* if the exponent is a variable *)
- let (env,n) = Env.compute_rank_add env term in (Mc.PEX n, env)
+ let (env,n) = Env.compute_rank_add env sigma term in (Mc.PEX n, env)
end
| Ukn s ->
if debug
then (Printf.printf "unknown op: %s\n" s; flush stdout;);
- let (env,n) = Env.compute_rank_add env term in (Mc.PEX n, env)
+ let (env,n) = Env.compute_rank_add env sigma term in (Mc.PEX n, env)
)
| _ -> parse_variable env term
)
@@ -1074,60 +1075,61 @@ struct
(* coq_Rdiv , (fun x y -> Mc.CMult(x,Mc.CInv y)) ;*)
]
- let rec rconstant term =
- match Term.kind_of_term term with
- | Const x ->
- if Constr.equal term (Lazy.force coq_R0)
+ let rec rconstant sigma term =
+ match EConstr.kind sigma term with
+ | Term.Const x ->
+ if EConstr.eq_constr sigma term (Lazy.force coq_R0)
then Mc.C0
- else if Constr.equal term (Lazy.force coq_R1)
+ else if EConstr.eq_constr sigma term (Lazy.force coq_R1)
then Mc.C1
else raise ParseError
- | App(op,args) ->
+ | Term.App(op,args) ->
begin
try
(* the evaluation order is important in the following *)
- let f = assoc_const op rconst_assoc in
- let a = rconstant args.(0) in
- let b = rconstant args.(1) in
+ let f = assoc_const sigma op rconst_assoc in
+ let a = rconstant sigma args.(0) in
+ let b = rconstant sigma args.(1) in
f a b
with
ParseError ->
match op with
- | op when Constr.equal op (Lazy.force coq_Rinv) ->
- let arg = rconstant args.(0) in
+ | op when EConstr.eq_constr sigma op (Lazy.force coq_Rinv) ->
+ let arg = rconstant sigma args.(0) in
if Mc.qeq_bool (Mc.q_of_Rcst arg) {Mc.qnum = Mc.Z0 ; Mc.qden = Mc.XH}
then raise ParseError (* This is a division by zero -- no semantics *)
else Mc.CInv(arg)
- | op when Constr.equal op (Lazy.force coq_IQR) -> Mc.CQ (parse_q args.(0))
- | op when Constr.equal op (Lazy.force coq_IZR) -> Mc.CZ (parse_z args.(0))
+ | op when EConstr.eq_constr sigma op (Lazy.force coq_IQR) -> Mc.CQ (parse_q sigma args.(0))
+ | op when EConstr.eq_constr sigma op (Lazy.force coq_IZR) -> Mc.CZ (parse_z sigma args.(0))
| _ -> raise ParseError
end
| _ -> raise ParseError
- let rconstant term =
+ let rconstant sigma term =
+ let _, env = Pfedit.get_current_context () in
if debug
- then Feedback.msg_debug (Pp.str "rconstant: " ++ Printer.prterm term ++ fnl ());
- let res = rconstant term in
+ then Feedback.msg_debug (Pp.str "rconstant: " ++ Printer.pr_leconstr_env env sigma term ++ fnl ());
+ let res = rconstant sigma term in
if debug then
(Printf.printf "rconstant -> %a\n" pp_Rcst res ; flush stdout) ;
res
- let parse_zexpr = parse_expr
- zconstant
+ let parse_zexpr sigma = parse_expr sigma
+ (zconstant sigma)
(fun expr x ->
- let exp = (parse_z x) in
+ let exp = (parse_z sigma x) in
match exp with
| Mc.Zneg _ -> Mc.PEc Mc.Z0
| _ -> Mc.PEpow(expr, Mc.Z.to_N exp))
zop_spec
- let parse_qexpr = parse_expr
- qconstant
+ let parse_qexpr sigma = parse_expr sigma
+ (qconstant sigma)
(fun expr x ->
- let exp = parse_z x in
+ let exp = parse_z sigma x in
match exp with
| Mc.Zneg _ ->
begin
@@ -1139,21 +1141,22 @@ struct
Mc.PEpow(expr,exp))
qop_spec
- let parse_rexpr = parse_expr
- rconstant
+ let parse_rexpr sigma = parse_expr sigma
+ (rconstant sigma)
(fun expr x ->
- let exp = Mc.N.of_nat (parse_nat x) in
+ let exp = Mc.N.of_nat (parse_nat sigma x) in
Mc.PEpow(expr,exp))
rop_spec
let parse_arith parse_op parse_expr env cstr gl =
- if debug
- then Feedback.msg_debug (Pp.str "parse_arith: " ++ Printer.prterm cstr ++ fnl ());
- match kind_of_term cstr with
- | App(op,args) ->
+ let sigma = gl.sigma in
+ if debug
+ then Feedback.msg_debug (Pp.str "parse_arith: " ++ Printer.pr_leconstr_env gl.env sigma cstr ++ fnl ());
+ match EConstr.kind sigma cstr with
+ | Term.App(op,args) ->
let (op,lhs,rhs) = parse_op gl (op,args) in
- let (e1,env) = parse_expr env lhs in
- let (e2,env) = parse_expr env rhs in
+ let (e1,env) = parse_expr sigma env lhs in
+ let (e2,env) = parse_expr sigma env rhs in
({Mc.flhs = e1; Mc.fop = op;Mc.frhs = e2},env)
| _ -> failwith "error : parse_arith(2)"
@@ -1191,6 +1194,7 @@ struct
*)
let parse_formula gl parse_atom env tg term =
+ let sigma = gl.sigma in
let parse_atom env tg t =
try
@@ -1199,34 +1203,34 @@ struct
with e when CErrors.noncritical e -> (X(t),env,tg) in
let is_prop term =
- let sort = Retyping.get_sort_of (Tacmach.pf_env gl) (Tacmach.project gl) term in
- Term.is_prop_sort sort in
+ let sort = Retyping.get_sort_of gl.env gl.sigma term in
+ Sorts.is_prop sort in
let rec xparse_formula env tg term =
- match kind_of_term term with
- | App(l,rst) ->
+ match EConstr.kind sigma term with
+ | Term.App(l,rst) ->
(match rst with
- | [|a;b|] when eq_constr l (Lazy.force coq_and) ->
+ | [|a;b|] when EConstr.eq_constr sigma l (Lazy.force coq_and) ->
let f,env,tg = xparse_formula env tg a in
let g,env, tg = xparse_formula env tg b in
mkformula_binary mkC term f g,env,tg
- | [|a;b|] when eq_constr l (Lazy.force coq_or) ->
+ | [|a;b|] when EConstr.eq_constr sigma l (Lazy.force coq_or) ->
let f,env,tg = xparse_formula env tg a in
let g,env,tg = xparse_formula env tg b in
mkformula_binary mkD term f g,env,tg
- | [|a|] when eq_constr l (Lazy.force coq_not) ->
+ | [|a|] when EConstr.eq_constr sigma l (Lazy.force coq_not) ->
let (f,env,tg) = xparse_formula env tg a in (N(f), env,tg)
- | [|a;b|] when eq_constr l (Lazy.force coq_iff) ->
+ | [|a;b|] when EConstr.eq_constr sigma l (Lazy.force coq_iff) ->
let f,env,tg = xparse_formula env tg a in
let g,env,tg = xparse_formula env tg b in
mkformula_binary mkIff term f g,env,tg
| _ -> parse_atom env tg term)
- | Prod(typ,a,b) when not (Termops.dependent (mkRel 1) b)->
+ | Term.Prod(typ,a,b) when EConstr.Vars.noccurn sigma 1 b ->
let f,env,tg = xparse_formula env tg a in
let g,env,tg = xparse_formula env tg b in
mkformula_binary mkI term f g,env,tg
- | _ when eq_constr term (Lazy.force coq_True) -> (TT,env,tg)
- | _ when eq_constr term (Lazy.force coq_False) -> (FF,env,tg)
+ | _ when EConstr.eq_constr sigma term (Lazy.force coq_True) -> (TT,env,tg)
+ | _ when EConstr.eq_constr sigma term (Lazy.force coq_False) -> (FF,env,tg)
| _ when is_prop term -> X(term),env,tg
| _ -> raise ParseError
in
@@ -1235,21 +1239,21 @@ struct
let dump_formula typ dump_atom f =
let rec xdump f =
match f with
- | TT -> mkApp(Lazy.force coq_TT,[|typ|])
- | FF -> mkApp(Lazy.force coq_FF,[|typ|])
- | C(x,y) -> mkApp(Lazy.force coq_And,[|typ ; xdump x ; xdump y|])
- | D(x,y) -> mkApp(Lazy.force coq_Or,[|typ ; xdump x ; xdump y|])
- | I(x,_,y) -> mkApp(Lazy.force coq_Impl,[|typ ; xdump x ; xdump y|])
- | N(x) -> mkApp(Lazy.force coq_Neg,[|typ ; xdump x|])
- | A(x,_,_) -> mkApp(Lazy.force coq_Atom,[|typ ; dump_atom x|])
- | X(t) -> mkApp(Lazy.force coq_X,[|typ ; t|]) in
+ | TT -> EConstr.mkApp(Lazy.force coq_TT,[|typ|])
+ | FF -> EConstr.mkApp(Lazy.force coq_FF,[|typ|])
+ | C(x,y) -> EConstr.mkApp(Lazy.force coq_And,[|typ ; xdump x ; xdump y|])
+ | D(x,y) -> EConstr.mkApp(Lazy.force coq_Or,[|typ ; xdump x ; xdump y|])
+ | I(x,_,y) -> EConstr.mkApp(Lazy.force coq_Impl,[|typ ; xdump x ; xdump y|])
+ | N(x) -> EConstr.mkApp(Lazy.force coq_Neg,[|typ ; xdump x|])
+ | A(x,_,_) -> EConstr.mkApp(Lazy.force coq_Atom,[|typ ; dump_atom x|])
+ | X(t) -> EConstr.mkApp(Lazy.force coq_X,[|typ ; t|]) in
xdump f
- let prop_env_of_formula form =
+ let prop_env_of_formula sigma form =
let rec doit env = function
| TT | FF | A(_,_,_) -> env
- | X t -> fst (Env.compute_rank_add env t)
+ | X t -> fst (Env.compute_rank_add env sigma t)
| C(f1,f2) | D(f1,f2) | I(f1,_,f2) ->
doit (doit env f1) f2
| N f -> doit env f in
@@ -1282,15 +1286,15 @@ struct
type 'cst dump_expr = (* 'cst is the type of the syntactic constants *)
{
- interp_typ : constr;
- dump_cst : 'cst -> constr;
- dump_add : constr;
- dump_sub : constr;
- dump_opp : constr;
- dump_mul : constr;
- dump_pow : constr;
- dump_pow_arg : Mc.n -> constr;
- dump_op : (Mc.op2 * Term.constr) list
+ interp_typ : EConstr.constr;
+ dump_cst : 'cst -> EConstr.constr;
+ dump_add : EConstr.constr;
+ dump_sub : EConstr.constr;
+ dump_opp : EConstr.constr;
+ dump_mul : EConstr.constr;
+ dump_pow : EConstr.constr;
+ dump_pow_arg : Mc.n -> EConstr.constr;
+ dump_op : (Mc.op2 * EConstr.constr) list
}
let dump_zexpr = lazy
@@ -1324,8 +1328,8 @@ let dump_qexpr = lazy
let add = Lazy.force coq_Rplus in
let one = Lazy.force coq_R1 in
- let mk_add x y = mkApp(add,[|x;y|]) in
- let mk_mult x y = mkApp(mult,[|x;y|]) in
+ let mk_add x y = EConstr.mkApp(add,[|x;y|]) in
+ let mk_mult x y = EConstr.mkApp(mult,[|x;y|]) in
let two = mk_add one one in
@@ -1348,13 +1352,13 @@ let rec dump_Rcst_as_R cst =
match cst with
| Mc.C0 -> Lazy.force coq_R0
| Mc.C1 -> Lazy.force coq_R1
- | Mc.CQ q -> Term.mkApp(Lazy.force coq_IQR, [| dump_q q |])
- | Mc.CZ z -> Term.mkApp(Lazy.force coq_IZR, [| dump_z z |])
- | Mc.CPlus(x,y) -> Term.mkApp(Lazy.force coq_Rplus, [| dump_Rcst_as_R x ; dump_Rcst_as_R y |])
- | Mc.CMinus(x,y) -> Term.mkApp(Lazy.force coq_Rminus, [| dump_Rcst_as_R x ; dump_Rcst_as_R y |])
- | Mc.CMult(x,y) -> Term.mkApp(Lazy.force coq_Rmult, [| dump_Rcst_as_R x ; dump_Rcst_as_R y |])
- | Mc.CInv t -> Term.mkApp(Lazy.force coq_Rinv, [| dump_Rcst_as_R t |])
- | Mc.COpp t -> Term.mkApp(Lazy.force coq_Ropp, [| dump_Rcst_as_R t |])
+ | Mc.CQ q -> EConstr.mkApp(Lazy.force coq_IQR, [| dump_q q |])
+ | Mc.CZ z -> EConstr.mkApp(Lazy.force coq_IZR, [| dump_z z |])
+ | Mc.CPlus(x,y) -> EConstr.mkApp(Lazy.force coq_Rplus, [| dump_Rcst_as_R x ; dump_Rcst_as_R y |])
+ | Mc.CMinus(x,y) -> EConstr.mkApp(Lazy.force coq_Rminus, [| dump_Rcst_as_R x ; dump_Rcst_as_R y |])
+ | Mc.CMult(x,y) -> EConstr.mkApp(Lazy.force coq_Rmult, [| dump_Rcst_as_R x ; dump_Rcst_as_R y |])
+ | Mc.CInv t -> EConstr.mkApp(Lazy.force coq_Rinv, [| dump_Rcst_as_R t |])
+ | Mc.COpp t -> EConstr.mkApp(Lazy.force coq_Ropp, [| dump_Rcst_as_R t |])
let dump_rexpr = lazy
@@ -1380,41 +1384,49 @@ let dump_rexpr = lazy
*)
-let rec make_goal_of_formula dexpr form =
+let prodn n env b =
+ let rec prodrec = function
+ | (0, env, b) -> b
+ | (n, ((v,t)::l), b) -> prodrec (n-1, l, EConstr.mkProd (v,t,b))
+ | _ -> assert false
+ in
+ prodrec (n,env,b)
+
+let make_goal_of_formula sigma dexpr form =
let vars_idx =
List.mapi (fun i v -> (v, i+1)) (ISet.elements (var_env_of_formula form)) in
(* List.iter (fun (v,i) -> Printf.fprintf stdout "var %i has index %i\n" v i) vars_idx ;*)
- let props = prop_env_of_formula form in
+ let props = prop_env_of_formula sigma form in
- let vars_n = List.map (fun (_,i) -> (Names.id_of_string (Printf.sprintf "__x%i" i)) , dexpr.interp_typ) vars_idx in
- let props_n = List.mapi (fun i _ -> (Names.id_of_string (Printf.sprintf "__p%i" (i+1))) , Term.mkProp) props in
+ let vars_n = List.map (fun (_,i) -> (Names.Id.of_string (Printf.sprintf "__x%i" i)) , dexpr.interp_typ) vars_idx in
+ let props_n = List.mapi (fun i _ -> (Names.Id.of_string (Printf.sprintf "__p%i" (i+1))) , EConstr.mkProp) props in
let var_name_pos = List.map2 (fun (idx,_) (id,_) -> id,idx) vars_idx vars_n in
let dump_expr i e =
let rec dump_expr = function
- | Mc.PEX n -> mkRel (i+(List.assoc (CoqToCaml.positive n) vars_idx))
+ | Mc.PEX n -> EConstr.mkRel (i+(List.assoc (CoqToCaml.positive n) vars_idx))
| Mc.PEc z -> dexpr.dump_cst z
- | Mc.PEadd(e1,e2) -> mkApp(dexpr.dump_add,
+ | Mc.PEadd(e1,e2) -> EConstr.mkApp(dexpr.dump_add,
[| dump_expr e1;dump_expr e2|])
- | Mc.PEsub(e1,e2) -> mkApp(dexpr.dump_sub,
+ | Mc.PEsub(e1,e2) -> EConstr.mkApp(dexpr.dump_sub,
[| dump_expr e1;dump_expr e2|])
- | Mc.PEopp e -> mkApp(dexpr.dump_opp,
- [| dump_expr e|])
- | Mc.PEmul(e1,e2) -> mkApp(dexpr.dump_mul,
- [| dump_expr e1;dump_expr e2|])
- | Mc.PEpow(e,n) -> mkApp(dexpr.dump_pow,
- [| dump_expr e; dexpr.dump_pow_arg n|])
+ | Mc.PEopp e -> EConstr.mkApp(dexpr.dump_opp,
+ [| dump_expr e|])
+ | Mc.PEmul(e1,e2) -> EConstr.mkApp(dexpr.dump_mul,
+ [| dump_expr e1;dump_expr e2|])
+ | Mc.PEpow(e,n) -> EConstr.mkApp(dexpr.dump_pow,
+ [| dump_expr e; dexpr.dump_pow_arg n|])
in dump_expr e in
let mkop op e1 e2 =
try
- Term.mkApp(List.assoc op dexpr.dump_op, [| e1; e2|])
+ EConstr.mkApp(List.assoc op dexpr.dump_op, [| e1; e2|])
with Not_found ->
- Term.mkApp(Lazy.force coq_Eq,[|dexpr.interp_typ ; e1 ;e2|]) in
+ EConstr.mkApp(Lazy.force coq_Eq,[|dexpr.interp_typ ; e1 ;e2|]) in
let dump_cstr i { Mc.flhs ; Mc.fop ; Mc.frhs } =
mkop fop (dump_expr i flhs) (dump_expr i frhs) in
@@ -1423,13 +1435,13 @@ let rec make_goal_of_formula dexpr form =
match f with
| TT -> Lazy.force coq_True
| FF -> Lazy.force coq_False
- | C(x,y) -> mkApp(Lazy.force coq_and,[|xdump pi xi x ; xdump pi xi y|])
- | D(x,y) -> mkApp(Lazy.force coq_or,[| xdump pi xi x ; xdump pi xi y|])
- | I(x,_,y) -> mkArrow (xdump pi xi x) (xdump (pi+1) (xi+1) y)
- | N(x) -> mkArrow (xdump pi xi x) (Lazy.force coq_False)
+ | C(x,y) -> EConstr.mkApp(Lazy.force coq_and,[|xdump pi xi x ; xdump pi xi y|])
+ | D(x,y) -> EConstr.mkApp(Lazy.force coq_or,[| xdump pi xi x ; xdump pi xi y|])
+ | I(x,_,y) -> EConstr.mkArrow (xdump pi xi x) (xdump (pi+1) (xi+1) y)
+ | N(x) -> EConstr.mkArrow (xdump pi xi x) (Lazy.force coq_False)
| A(x,_,_) -> dump_cstr xi x
- | X(t) -> let idx = Env.get_rank props t in
- mkRel (pi+idx) in
+ | X(t) -> let idx = Env.get_rank props sigma t in
+ EConstr.mkRel (pi+idx) in
let nb_vars = List.length vars_n in
let nb_props = List.length props_n in
@@ -1437,13 +1449,13 @@ let rec make_goal_of_formula dexpr form =
(* Printf.fprintf stdout "NBProps : %i\n" nb_props ;*)
let subst_prop p =
- let idx = Env.get_rank props p in
- mkVar (Names.id_of_string (Printf.sprintf "__p%i" idx)) in
+ let idx = Env.get_rank props sigma p in
+ EConstr.mkVar (Names.Id.of_string (Printf.sprintf "__p%i" idx)) in
let form' = map_prop subst_prop form in
- (Term.prodn nb_props (List.map (fun (x,y) -> Names.Name x,y) props_n)
- (Term.prodn nb_vars (List.map (fun (x,y) -> Names.Name x,y) vars_n)
+ (prodn nb_props (List.map (fun (x,y) -> Name.Name x,y) props_n)
+ (prodn nb_vars (List.map (fun (x,y) -> Name.Name x,y) vars_n)
(xdump (List.length vars_n) 0 form)),
List.rev props_n, List.rev var_name_pos,form')
@@ -1458,7 +1470,7 @@ let rec make_goal_of_formula dexpr form =
| [] -> acc
| (e::l) ->
let (name,expr,typ) = e in
- xset (Term.mkNamedLetIn
+ xset (EConstr.mkNamedLetIn
(Names.Id.of_string name)
expr typ acc) l in
xset concl l
@@ -1515,29 +1527,29 @@ let rec witness prover l1 l2 =
let rec apply_ids t ids =
match ids with
| [] -> t
- | i::ids -> apply_ids (Term.mkApp(t,[| Term.mkVar i |])) ids
+ | i::ids -> apply_ids (mkApp(t,[| mkVar i |])) ids
-let coq_Node = lazy
- (Coqlib.gen_constant_in_modules "VarMap"
+let coq_Node =
+ lazy (gen_constant_in_modules "VarMap"
[["Coq" ; "micromega" ; "VarMap"];["VarMap"]] "Node")
-let coq_Leaf = lazy
- (Coqlib.gen_constant_in_modules "VarMap"
+let coq_Leaf =
+ lazy (gen_constant_in_modules "VarMap"
[["Coq" ; "micromega" ; "VarMap"];["VarMap"]] "Leaf")
-let coq_Empty = lazy
- (Coqlib.gen_constant_in_modules "VarMap"
+let coq_Empty =
+ lazy (gen_constant_in_modules "VarMap"
[["Coq" ; "micromega" ;"VarMap"];["VarMap"]] "Empty")
-let coq_VarMap = lazy
- (Coqlib.gen_constant_in_modules "VarMap"
+let coq_VarMap =
+ lazy (gen_constant_in_modules "VarMap"
[["Coq" ; "micromega" ; "VarMap"] ; ["VarMap"]] "t")
let rec dump_varmap typ m =
match m with
- | Mc.Empty -> Term.mkApp(Lazy.force coq_Empty,[| typ |])
- | Mc.Leaf v -> Term.mkApp(Lazy.force coq_Leaf,[| typ; v|])
+ | Mc.Empty -> EConstr.mkApp(Lazy.force coq_Empty,[| typ |])
+ | Mc.Leaf v -> EConstr.mkApp(Lazy.force coq_Leaf,[| typ; v|])
| Mc.Node(l,o,r) ->
- Term.mkApp (Lazy.force coq_Node, [| typ; dump_varmap typ l; o ; dump_varmap typ r |])
+ EConstr.mkApp (Lazy.force coq_Node, [| typ; dump_varmap typ l; o ; dump_varmap typ r |])
let vm_of_list env =
@@ -1559,15 +1571,15 @@ let rec pp_varmap o vm =
let rec dump_proof_term = function
| Micromega.DoneProof -> Lazy.force coq_doneProof
| Micromega.RatProof(cone,rst) ->
- Term.mkApp(Lazy.force coq_ratProof, [| dump_psatz coq_Z dump_z cone; dump_proof_term rst|])
+ EConstr.mkApp(Lazy.force coq_ratProof, [| dump_psatz coq_Z dump_z cone; dump_proof_term rst|])
| Micromega.CutProof(cone,prf) ->
- Term.mkApp(Lazy.force coq_cutProof,
+ EConstr.mkApp(Lazy.force coq_cutProof,
[| dump_psatz coq_Z dump_z cone ;
dump_proof_term prf|])
| Micromega.EnumProof(c1,c2,prfs) ->
- Term.mkApp (Lazy.force coq_enumProof,
- [| dump_psatz coq_Z dump_z c1 ; dump_psatz coq_Z dump_z c2 ;
- dump_list (Lazy.force coq_proofTerm) dump_proof_term prfs |])
+ EConstr.mkApp (Lazy.force coq_enumProof,
+ [| dump_psatz coq_Z dump_z c1 ; dump_psatz coq_Z dump_z c2 ;
+ dump_list (Lazy.force coq_proofTerm) dump_proof_term prfs |])
let rec size_of_psatz = function
@@ -1627,11 +1639,11 @@ let parse_goal gl parse_arith env hyps term =
* The datastructures that aggregate theory-dependent proof values.
*)
type ('synt_c, 'prf) domain_spec = {
- typ : Term.constr; (* is the type of the interpretation domain - Z, Q, R*)
- coeff : Term.constr ; (* is the type of the syntactic coeffs - Z , Q , Rcst *)
- dump_coeff : 'synt_c -> Term.constr ;
- proof_typ : Term.constr ;
- dump_proof : 'prf -> Term.constr
+ typ : EConstr.constr; (* is the type of the interpretation domain - Z, Q, R*)
+ coeff : EConstr.constr ; (* is the type of the syntactic coeffs - Z , Q , Rcst *)
+ dump_coeff : 'synt_c -> EConstr.constr ;
+ proof_typ : EConstr.constr ;
+ dump_proof : 'prf -> EConstr.constr
}
let zz_domain_spec = lazy {
@@ -1658,8 +1670,6 @@ let rcst_domain_spec = lazy {
dump_proof = dump_psatz coq_Q dump_q
}
-open Proofview.Notations
-
(** Naive topological sort of constr according to the subterm-ordering *)
(* An element is minimal x is minimal w.r.t y if
@@ -1687,7 +1697,8 @@ let rec mk_topo_order le l =
| (Some v,l') -> v :: (mk_topo_order le l')
-let topo_sort_constr l = mk_topo_order Termops.dependent l
+let topo_sort_constr l =
+ mk_topo_order (fun c t -> Termops.dependent Evd.empty (** FIXME *) (EConstr.of_constr c) (EConstr.of_constr t)) l
(**
@@ -1697,24 +1708,23 @@ let topo_sort_constr l = mk_topo_order Termops.dependent l
let micromega_order_change spec cert cert_typ env ff (*: unit Proofview.tactic*) =
(* let ids = Util.List.map_i (fun i _ -> (Names.Id.of_string ("__v"^(string_of_int i)))) 0 env in *)
- let formula_typ = (Term.mkApp (Lazy.force coq_Cstr,[|spec.coeff|])) in
+ let formula_typ = (EConstr.mkApp (Lazy.force coq_Cstr,[|spec.coeff|])) in
let ff = dump_formula formula_typ (dump_cstr spec.coeff spec.dump_coeff) ff in
let vm = dump_varmap (spec.typ) (vm_of_list env) in
(* todo : directly generate the proof term - or generalize before conversion? *)
- Proofview.Goal.nf_enter { enter = begin fun gl ->
- let gl = Tacmach.New.of_old (fun x -> x) gl in
+ Proofview.Goal.nf_enter begin fun gl ->
Tacticals.New.tclTHENLIST
[
Tactics.change_concl
(set
[
- ("__ff", ff, Term.mkApp(Lazy.force coq_Formula, [|formula_typ |]));
- ("__varmap", vm, Term.mkApp(Lazy.force coq_VarMap, [|spec.typ|]));
+ ("__ff", ff, EConstr.mkApp(Lazy.force coq_Formula, [|formula_typ |]));
+ ("__varmap", vm, EConstr.mkApp(Lazy.force coq_VarMap, [|spec.typ|]));
("__wit", cert, cert_typ)
]
- (Tacmach.pf_concl gl))
+ (Tacmach.New.pf_concl gl))
]
- end }
+ end
(**
@@ -1833,20 +1843,20 @@ let abstract_formula hyps f =
| A(a,t,term) -> if TagSet.mem t hyps then A(a,t,term) else X(term)
| C(f1,f2) ->
(match xabs f1 , xabs f2 with
- | X a1 , X a2 -> X (Term.mkApp(Lazy.force coq_and, [|a1;a2|]))
+ | X a1 , X a2 -> X (EConstr.mkApp(Lazy.force coq_and, [|a1;a2|]))
| f1 , f2 -> C(f1,f2) )
| D(f1,f2) ->
(match xabs f1 , xabs f2 with
- | X a1 , X a2 -> X (Term.mkApp(Lazy.force coq_or, [|a1;a2|]))
+ | X a1 , X a2 -> X (EConstr.mkApp(Lazy.force coq_or, [|a1;a2|]))
| f1 , f2 -> D(f1,f2) )
| N(f) ->
(match xabs f with
- | X a -> X (Term.mkApp(Lazy.force coq_not, [|a|]))
+ | X a -> X (EConstr.mkApp(Lazy.force coq_not, [|a|]))
| f -> N f)
| I(f1,hyp,f2) ->
(match xabs f1 , hyp, xabs f2 with
| X a1 , Some _ , af2 -> af2
- | X a1 , None , X a2 -> X (Term.mkArrow a1 a2)
+ | X a1 , None , X a2 -> X (EConstr.mkArrow a1 a2)
| af1 , _ , af2 -> I(af1,hyp,af2)
)
| FF -> FF
@@ -1900,10 +1910,10 @@ let micromega_tauto negate normalise unsat deduce spec prover env polys1 polys2
if debug then
begin
Feedback.msg_notice (Pp.str "Formula....\n") ;
- let formula_typ = (Term.mkApp(Lazy.force coq_Cstr, [|spec.coeff|])) in
+ let formula_typ = (EConstr.mkApp(Lazy.force coq_Cstr, [|spec.coeff|])) in
let ff = dump_formula formula_typ
(dump_cstr spec.typ spec.dump_coeff) ff in
- Feedback.msg_notice (Printer.prterm ff);
+ Feedback.msg_notice (Printer.pr_leconstr_env gl.env gl.sigma ff);
Printf.fprintf stdout "cnf : %a\n" (pp_cnf (fun o _ -> ())) cnf_ff
end;
@@ -1925,11 +1935,11 @@ let micromega_tauto negate normalise unsat deduce spec prover env polys1 polys2
if debug then
begin
Feedback.msg_notice (Pp.str "\nAFormula\n") ;
- let formula_typ = (Term.mkApp( Lazy.force coq_Cstr,[| spec.coeff|])) in
+ let formula_typ = (EConstr.mkApp( Lazy.force coq_Cstr,[| spec.coeff|])) in
let ff' = dump_formula formula_typ
- (dump_cstr spec.typ spec.dump_coeff) ff' in
- Feedback.msg_notice (Printer.prterm ff');
- Printf.fprintf stdout "cnf : %a\n" (pp_cnf (fun o _ -> ())) cnf_ff'
+ (dump_cstr spec.typ spec.dump_coeff) ff' in
+ Feedback.msg_notice (Printer.pr_leconstr_env gl.env gl.sigma ff');
+ Printf.fprintf stdout "cnf : %a\n" (pp_cnf (fun o _ -> ())) cnf_ff'
end;
(* Even if it does not work, this does not mean it is not provable
@@ -1949,44 +1959,47 @@ let micromega_tauto negate normalise unsat deduce spec prover env polys1 polys2
Some (ids,ff',res')
-
(**
* Parse the proof environment, and call micromega_tauto
*)
+let fresh_id avoid id gl =
+ Tactics.fresh_id_in_env avoid id (Proofview.Goal.env gl)
+
let micromega_gen
parse_arith
(negate:'cst atom -> 'cst mc_cnf)
(normalise:'cst atom -> 'cst mc_cnf)
unsat deduce
spec dumpexpr prover tac =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
- let gl = Tacmach.New.of_old (fun x -> x) gl in
- let concl = Tacmach.pf_concl gl in
- let hyps = Tacmach.pf_hyps_types gl in
+ Proofview.Goal.nf_enter begin fun gl ->
+ let sigma = Tacmach.New.project gl in
+ let concl = Tacmach.New.pf_concl gl in
+ let hyps = Tacmach.New.pf_hyps_types gl in
try
- let (hyps,concl,env) = parse_goal gl parse_arith Env.empty hyps concl in
+ let gl0 = { env = Tacmach.New.pf_env gl; sigma } in
+ let (hyps,concl,env) = parse_goal gl0 parse_arith Env.empty hyps concl in
let env = Env.elements env in
let spec = Lazy.force spec in
let dumpexpr = Lazy.force dumpexpr in
- match micromega_tauto negate normalise unsat deduce spec prover env hyps concl gl with
+ match micromega_tauto negate normalise unsat deduce spec prover env hyps concl gl0 with
| None -> Tacticals.New.tclFAIL 0 (Pp.str " Cannot find witness")
| Some (ids,ff',res') ->
- let (arith_goal,props,vars,ff_arith) = make_goal_of_formula dumpexpr ff' in
+ let (arith_goal,props,vars,ff_arith) = make_goal_of_formula sigma dumpexpr ff' in
let intro (id,_) = Tactics.introduction id in
let intro_vars = Tacticals.New.tclTHENLIST (List.map intro vars) in
let intro_props = Tacticals.New.tclTHENLIST (List.map intro props) in
- let ipat_of_name id = Some (Loc.ghost, Misctypes.IntroNaming (Misctypes.IntroIdentifier id)) in
- let goal_name = Tactics.fresh_id [] (Names.Id.of_string "__arith") gl in
- let env' = List.map (fun (id,i) -> Term.mkVar id,i) vars in
+ let ipat_of_name id = Some (CAst.make @@ Misctypes.IntroNaming (Misctypes.IntroIdentifier id)) in
+ let goal_name = fresh_id Id.Set.empty (Names.Id.of_string "__arith") gl in
+ let env' = List.map (fun (id,i) -> EConstr.mkVar id,i) vars in
let tac_arith = Tacticals.New.tclTHENLIST [ intro_props ; intro_vars ;
micromega_order_change spec res'
- (Term.mkApp(Lazy.force coq_list, [|spec.proof_typ|])) env' ff_arith ] in
+ (EConstr.mkApp(Lazy.force coq_list, [|spec.proof_typ|])) env' ff_arith ] in
- let goal_props = List.rev (prop_env_of_formula ff') in
+ let goal_props = List.rev (prop_env_of_formula sigma ff') in
let goal_vars = List.map (fun (_,i) -> List.nth env (i-1)) vars in
@@ -2003,8 +2016,8 @@ let micromega_gen
[
kill_arith;
(Tacticals.New.tclTHENLIST
- [(Tactics.generalize (List.map Term.mkVar ids));
- Tactics.exact_check (Term.applist (Term.mkVar goal_name, arith_args))
+ [(Tactics.generalize (List.map EConstr.mkVar ids));
+ Tactics.exact_check (EConstr.applist (EConstr.mkVar goal_name, arith_args))
] )
]
with
@@ -2016,7 +2029,7 @@ let micromega_gen
^ "the use of a specialized external tool called csdp. \n\n"
^ "Unfortunately Coq isn't aware of the presence of any \"csdp\" executable in the path. \n\n"
^ "Csdp packages are provided by some OS distributions; binaries and source code can be downloaded from https://projects.coin-or.org/Csdp"))
- end }
+ end
let micromega_gen parse_arith
(negate:'cst atom -> 'cst mc_cnf)
@@ -2032,28 +2045,27 @@ let micromega_order_changer cert env ff =
let coeff = Lazy.force coq_Rcst in
let dump_coeff = dump_Rcst in
let typ = Lazy.force coq_R in
- let cert_typ = (Term.mkApp(Lazy.force coq_list, [|Lazy.force coq_QWitness |])) in
+ let cert_typ = (EConstr.mkApp(Lazy.force coq_list, [|Lazy.force coq_QWitness |])) in
- let formula_typ = (Term.mkApp (Lazy.force coq_Cstr,[| coeff|])) in
+ let formula_typ = (EConstr.mkApp (Lazy.force coq_Cstr,[| coeff|])) in
let ff = dump_formula formula_typ (dump_cstr coeff dump_coeff) ff in
let vm = dump_varmap (typ) (vm_of_list env) in
- Proofview.Goal.nf_enter { enter = begin fun gl ->
- let gl = Tacmach.New.of_old (fun x -> x) gl in
+ Proofview.Goal.nf_enter begin fun gl ->
Tacticals.New.tclTHENLIST
[
(Tactics.change_concl
(set
[
- ("__ff", ff, Term.mkApp(Lazy.force coq_Formula, [|formula_typ |]));
- ("__varmap", vm, Term.mkApp
- (Coqlib.gen_constant_in_modules "VarMap"
+ ("__ff", ff, EConstr.mkApp(Lazy.force coq_Formula, [|formula_typ |]));
+ ("__varmap", vm, EConstr.mkApp
+ (gen_constant_in_modules "VarMap"
[["Coq" ; "micromega" ; "VarMap"] ; ["VarMap"]] "t", [|typ|]));
("__wit", cert, cert_typ)
]
- (Tacmach.pf_concl gl)));
+ (Tacmach.New.pf_concl gl)));
(* Tacticals.New.tclTHENLIST (List.map (fun id -> (Tactics.introduction id)) ids)*)
]
- end }
+ end
let micromega_genr prover tac =
let parse_arith = parse_rarith in
@@ -2068,39 +2080,40 @@ let micromega_genr prover tac =
proof_typ = Lazy.force coq_QWitness ;
dump_proof = dump_psatz coq_Q dump_q
} in
- Proofview.Goal.nf_enter { enter = begin fun gl ->
- let gl = Tacmach.New.of_old (fun x -> x) gl in
- let concl = Tacmach.pf_concl gl in
- let hyps = Tacmach.pf_hyps_types gl in
+ Proofview.Goal.nf_enter begin fun gl ->
+ let sigma = Tacmach.New.project gl in
+ let concl = Tacmach.New.pf_concl gl in
+ let hyps = Tacmach.New.pf_hyps_types gl in
try
- let (hyps,concl,env) = parse_goal gl parse_arith Env.empty hyps concl in
+ let gl0 = { env = Tacmach.New.pf_env gl; sigma } in
+ let (hyps,concl,env) = parse_goal gl0 parse_arith Env.empty hyps concl in
let env = Env.elements env in
let spec = Lazy.force spec in
let hyps' = List.map (fun (n,f) -> (n, map_atoms (Micromega.map_Formula Micromega.q_of_Rcst) f)) hyps in
let concl' = map_atoms (Micromega.map_Formula Micromega.q_of_Rcst) concl in
- match micromega_tauto negate normalise unsat deduce spec prover env hyps' concl' gl with
+ match micromega_tauto negate normalise unsat deduce spec prover env hyps' concl' gl0 with
| None -> Tacticals.New.tclFAIL 0 (Pp.str " Cannot find witness")
| Some (ids,ff',res') ->
let (ff,ids) = formula_hyps_concl
(List.filter (fun (n,_) -> List.mem n ids) hyps) concl in
let ff' = abstract_wrt_formula ff' ff in
- let (arith_goal,props,vars,ff_arith) = make_goal_of_formula (Lazy.force dump_rexpr) ff' in
+ let (arith_goal,props,vars,ff_arith) = make_goal_of_formula sigma (Lazy.force dump_rexpr) ff' in
let intro (id,_) = Tactics.introduction id in
let intro_vars = Tacticals.New.tclTHENLIST (List.map intro vars) in
let intro_props = Tacticals.New.tclTHENLIST (List.map intro props) in
- let ipat_of_name id = Some (Loc.ghost, Misctypes.IntroNaming (Misctypes.IntroIdentifier id)) in
- let goal_name = Tactics.fresh_id [] (Names.Id.of_string "__arith") gl in
- let env' = List.map (fun (id,i) -> Term.mkVar id,i) vars in
+ let ipat_of_name id = Some (CAst.make @@ Misctypes.IntroNaming (Misctypes.IntroIdentifier id)) in
+ let goal_name = fresh_id Id.Set.empty (Names.Id.of_string "__arith") gl in
+ let env' = List.map (fun (id,i) -> EConstr.mkVar id,i) vars in
let tac_arith = Tacticals.New.tclTHENLIST [ intro_props ; intro_vars ;
micromega_order_changer res' env' ff_arith ] in
- let goal_props = List.rev (prop_env_of_formula ff') in
+ let goal_props = List.rev (prop_env_of_formula sigma ff') in
let goal_vars = List.map (fun (_,i) -> List.nth env (i-1)) vars in
@@ -2117,8 +2130,8 @@ let micromega_genr prover tac =
[
kill_arith;
(Tacticals.New.tclTHENLIST
- [(Tactics.generalize (List.map Term.mkVar ids));
- Tactics.exact_check (Term.applist (Term.mkVar goal_name, arith_args))
+ [(Tactics.generalize (List.map EConstr.mkVar ids));
+ Tactics.exact_check (EConstr.applist (EConstr.mkVar goal_name, arith_args))
] )
]
@@ -2131,7 +2144,7 @@ let micromega_genr prover tac =
^ "the use of a specialized external tool called csdp. \n\n"
^ "Unfortunately Coq isn't aware of the presence of any \"csdp\" executable in the path. \n\n"
^ "Csdp packages are provided by some OS distributions; binaries and source code can be downloaded from https://projects.coin-or.org/Csdp"))
- end }
+ end
diff --git a/plugins/micromega/csdpcert.ml b/plugins/micromega/csdpcert.ml
index 2536005e..a1245b7c 100644
--- a/plugins/micromega/csdpcert.ml
+++ b/plugins/micromega/csdpcert.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* *)
(* Micromega: A reflexive tactic using the Positivstellensatz *)
diff --git a/plugins/micromega/g_micromega.ml4 b/plugins/micromega/g_micromega.ml4
index 027f690f..81140a46 100644
--- a/plugins/micromega/g_micromega.ml4
+++ b/plugins/micromega/g_micromega.ml4
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* *)
(* Micromega: A reflexive tactic using the Positivstellensatz *)
@@ -14,9 +16,9 @@
(* *)
(************************************************************************)
-(*i camlp4deps: "grammar/grammar.cma" i*)
-
-open Constrarg
+open Ltac_plugin
+open Stdarg
+open Tacarg
DECLARE PLUGIN "micromega_plugin"
diff --git a/plugins/micromega/mfourier.ml b/plugins/micromega/mfourier.ml
index f4f9b3c2..37799441 100644
--- a/plugins/micromega/mfourier.ml
+++ b/plugins/micromega/mfourier.ml
@@ -99,7 +99,7 @@ module PSet = ISet
module System = Hashtbl.Make(Vect)
type proof =
-| Hyp of int
+| Assum of int
| Elim of var * proof * proof
| And of proof * proof
@@ -134,7 +134,7 @@ exception SystemContradiction of proof
let hyps prf =
let rec hyps prf acc =
match prf with
- | Hyp i -> ISet.add i acc
+ | Assum i -> ISet.add i acc
| Elim(_,prf1,prf2)
| And(prf1,prf2) -> hyps prf1 (hyps prf2 acc) in
hyps prf ISet.empty
@@ -143,7 +143,7 @@ let hyps prf =
(** Pretty printing *)
let rec pp_proof o prf =
match prf with
- | Hyp i -> Printf.fprintf o "H%i" i
+ | Assum i -> Printf.fprintf o "H%i" i
| Elim(v, prf1,prf2) -> Printf.fprintf o "E(%i,%a,%a)" v pp_proof prf1 pp_proof prf2
| And(prf1,prf2) -> Printf.fprintf o "A(%a,%a)" pp_proof prf1 pp_proof prf2
@@ -270,7 +270,7 @@ let norm_cstr {coeffs = v ; op = o ; cst = c} idx =
(match o with
| Eq -> Some c , Some c
| Ge -> Some c , None) ;
- prf = Hyp idx }
+ prf = Assum idx }
(** [load_system l] takes a list of constraints of type [cstr_compat]
@@ -285,7 +285,7 @@ let load_system l =
let vars = List.fold_left (fun vrs (cstr,i) ->
match norm_cstr cstr i with
- | Contradiction -> raise (SystemContradiction (Hyp i))
+ | Contradiction -> raise (SystemContradiction (Assum i))
| Redundant -> vrs
| Cstr(vect,info) ->
xadd_cstr vect info sys ;
@@ -867,7 +867,7 @@ let mk_proof hyps prf =
let rec mk_proof prf =
match prf with
- | Hyp i -> [ ([i, Int 1] , List.nth hyps i) ]
+ | Assum i -> [ ([i, Int 1] , List.nth hyps i) ]
| Elim(v,prf1,prf2) ->
let prfsl = mk_proof prf1
diff --git a/plugins/micromega/micromega.ml b/plugins/micromega/micromega.ml
index 5cf1da8e..52c6ef98 100644
--- a/plugins/micromega/micromega.ml
+++ b/plugins/micromega/micromega.ml
@@ -1,3 +1,4 @@
+
(** val negb : bool -> bool **)
let negb = function
@@ -34,8 +35,7 @@ module Coq__1 = struct
| O -> m
| S p -> S (add p m)
end
-let add = Coq__1.add
-
+include Coq__1
type positive =
| XI of positive
@@ -82,11 +82,10 @@ module Coq_Pos =
| XI q0 -> XI (add p q0)
| XO q0 -> XO (add p q0)
| XH -> XI p)
- | XH ->
- (match y with
- | XI q0 -> XO (succ q0)
- | XO q0 -> XI q0
- | XH -> XO XH)
+ | XH -> (match y with
+ | XI q0 -> XO (succ q0)
+ | XO q0 -> XI q0
+ | XH -> XO XH)
(** val add_carry : positive -> positive -> positive **)
@@ -154,10 +153,9 @@ module Coq_Pos =
| XI q0 -> succ_double_mask (sub_mask_carry p q0)
| XO q0 -> double_mask (sub_mask p q0)
| XH -> IsPos (pred_double p))
- | XH ->
- (match y with
- | XH -> IsNul
- | _ -> IsNeg)
+ | XH -> (match y with
+ | XH -> IsNul
+ | _ -> IsNeg)
(** val sub_mask_carry : positive -> positive -> mask **)
@@ -197,8 +195,7 @@ module Coq_Pos =
| XO p2 -> S (size_nat p2)
| XH -> S O
- (** val compare_cont :
- comparison -> positive -> positive -> comparison **)
+ (** val compare_cont : comparison -> positive -> positive -> comparison **)
let rec compare_cont r x y =
match x with
@@ -212,10 +209,9 @@ module Coq_Pos =
| XI q0 -> compare_cont Lt p q0
| XO q0 -> compare_cont r p q0
| XH -> Gt)
- | XH ->
- (match y with
- | XH -> r
- | _ -> Lt)
+ | XH -> (match y with
+ | XH -> r
+ | _ -> Lt)
(** val compare : positive -> positive -> comparison **)
@@ -277,14 +273,12 @@ let rec pow_pos rmul x = function
let rec nth n0 l default =
match n0 with
- | O ->
- (match l with
- | [] -> default
- | x::_ -> x)
- | S m ->
- (match l with
- | [] -> default
- | _::t0 -> nth m t0 default)
+ | O -> (match l with
+ | [] -> default
+ | x::_ -> x)
+ | S m -> (match l with
+ | [] -> default
+ | _::t0 -> nth m t0 default)
(** val map : ('a1 -> 'a2) -> 'a1 list -> 'a2 list **)
@@ -389,15 +383,13 @@ module Z =
let compare x y =
match x with
- | Z0 ->
- (match y with
- | Z0 -> Eq
- | Zpos _ -> Lt
- | Zneg _ -> Gt)
- | Zpos x' ->
- (match y with
- | Zpos y' -> Coq_Pos.compare x' y'
- | _ -> Gt)
+ | Z0 -> (match y with
+ | Z0 -> Eq
+ | Zpos _ -> Lt
+ | Zneg _ -> Gt)
+ | Zpos x' -> (match y with
+ | Zpos y' -> Coq_Pos.compare x' y'
+ | _ -> Gt)
| Zneg x' ->
(match y with
| Zneg y' -> compOpp (Coq_Pos.compare x' y')
@@ -533,10 +525,9 @@ let p1 cI =
let rec peq ceqb p p' =
match p with
- | Pc c ->
- (match p' with
- | Pc c' -> ceqb c c'
- | _ -> false)
+ | Pc c -> (match p' with
+ | Pc c' -> ceqb c c'
+ | _ -> false)
| Pinj (j, q0) ->
(match p' with
| Pinj (j', q') ->
@@ -568,8 +559,7 @@ let mkPinj_pred j p =
| XH -> p
(** val mkPX :
- 'a1 -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> positive -> 'a1 pol -> 'a1
- pol **)
+ 'a1 -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol **)
let mkPX cO ceqb p i q0 =
match p with
@@ -631,8 +621,8 @@ let rec paddI cadd pop q0 j = function
| XH -> PX (p2, i, (pop q' q0)))
(** val psubI :
- ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 pol -> 'a1 pol -> 'a1 pol)
- -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol **)
+ ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 pol -> 'a1 pol -> 'a1 pol) ->
+ 'a1 pol -> positive -> 'a1 pol -> 'a1 pol **)
let rec psubI cadd copp pop q0 j = function
| Pc c -> mkPinj j (paddC cadd (popp copp q0) c)
@@ -644,13 +634,12 @@ let rec psubI cadd copp pop q0 j = function
| PX (p2, i, q') ->
(match j with
| XI j0 -> PX (p2, i, (psubI cadd copp pop q0 (XO j0) q'))
- | XO j0 ->
- PX (p2, i, (psubI cadd copp pop q0 (Coq_Pos.pred_double j0) q'))
+ | XO j0 -> PX (p2, i, (psubI cadd copp pop q0 (Coq_Pos.pred_double j0) q'))
| XH -> PX (p2, i, (pop q' q0)))
(** val paddX :
- 'a1 -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1
- pol -> positive -> 'a1 pol -> 'a1 pol **)
+ 'a1 -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1 pol
+ -> positive -> 'a1 pol -> 'a1 pol **)
let rec paddX cO ceqb pop p' i' p = match p with
| Pc _ -> PX (p', i', p)
@@ -666,16 +655,15 @@ let rec paddX cO ceqb pop p' i' p = match p with
| Zneg k -> mkPX cO ceqb (paddX cO ceqb pop p' k p2) i q')
(** val psubX :
- 'a1 -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol ->
- 'a1 pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol **)
+ 'a1 -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol -> 'a1
+ pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol **)
let rec psubX cO copp ceqb pop p' i' p = match p with
| Pc _ -> PX ((popp copp p'), i', p)
| Pinj (j, q') ->
(match j with
| XI j0 -> PX ((popp copp p'), i', (Pinj ((XO j0), q')))
- | XO j0 ->
- PX ((popp copp p'), i', (Pinj ((Coq_Pos.pred_double j0), q')))
+ | XO j0 -> PX ((popp copp p'), i', (Pinj ((Coq_Pos.pred_double j0), q')))
| XH -> PX ((popp copp p'), i', q'))
| PX (p2, i, q') ->
(match Z.pos_sub i i' with
@@ -684,8 +672,8 @@ let rec psubX cO copp ceqb pop p' i' p = match p with
| Zneg k -> mkPX cO ceqb (psubX cO copp ceqb pop p' k p2) i q')
(** val padd :
- 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1
- pol -> 'a1 pol **)
+ 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol
+ -> 'a1 pol **)
let rec padd cO cadd ceqb p = function
| Pc c' -> paddC cadd p c'
@@ -703,8 +691,7 @@ let rec padd cO cadd ceqb p = function
| PX (p2, i, q0) ->
(match Z.pos_sub i i' with
| Z0 ->
- mkPX cO ceqb (padd cO cadd ceqb p2 p'0) i
- (padd cO cadd ceqb q0 q')
+ mkPX cO ceqb (padd cO cadd ceqb p2 p'0) i (padd cO cadd ceqb q0 q')
| Zpos k ->
mkPX cO ceqb (padd cO cadd ceqb (PX (p2, k, (p0 cO))) p'0) i'
(padd cO cadd ceqb q0 q')
@@ -713,8 +700,8 @@ let rec padd cO cadd ceqb p = function
(padd cO cadd ceqb q0 q')))
(** val psub :
- 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) ->
- ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol **)
+ 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1
+ -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol **)
let rec psub cO cadd csub copp ceqb p = function
| Pc c' -> psubC csub p c'
@@ -729,39 +716,36 @@ let rec psub cO cadd csub copp ceqb p = function
(psub cO cadd csub copp ceqb (Pinj ((XO j0), q0)) q'))
| XO j0 ->
PX ((popp copp p'0), i',
- (psub cO cadd csub copp ceqb (Pinj ((Coq_Pos.pred_double j0),
- q0)) q'))
- | XH ->
- PX ((popp copp p'0), i', (psub cO cadd csub copp ceqb q0 q')))
+ (psub cO cadd csub copp ceqb (Pinj ((Coq_Pos.pred_double j0), q0))
+ q'))
+ | XH -> PX ((popp copp p'0), i', (psub cO cadd csub copp ceqb q0 q')))
| PX (p2, i, q0) ->
(match Z.pos_sub i i' with
| Z0 ->
mkPX cO ceqb (psub cO cadd csub copp ceqb p2 p'0) i
(psub cO cadd csub copp ceqb q0 q')
| Zpos k ->
- mkPX cO ceqb
- (psub cO cadd csub copp ceqb (PX (p2, k, (p0 cO))) p'0) i'
- (psub cO cadd csub copp ceqb q0 q')
+ mkPX cO ceqb (psub cO cadd csub copp ceqb (PX (p2, k, (p0 cO))) p'0)
+ i' (psub cO cadd csub copp ceqb q0 q')
| Zneg k ->
mkPX cO ceqb
(psubX cO copp ceqb (psub cO cadd csub copp ceqb) p'0 k p2) i
(psub cO cadd csub copp ceqb q0 q')))
(** val pmulC_aux :
- 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1
- -> 'a1 pol **)
+ 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 ->
+ 'a1 pol **)
let rec pmulC_aux cO cmul ceqb p c =
match p with
| Pc c' -> Pc (cmul c' c)
| Pinj (j, q0) -> mkPinj j (pmulC_aux cO cmul ceqb q0 c)
| PX (p2, i, q0) ->
- mkPX cO ceqb (pmulC_aux cO cmul ceqb p2 c) i
- (pmulC_aux cO cmul ceqb q0 c)
+ mkPX cO ceqb (pmulC_aux cO cmul ceqb p2 c) i (pmulC_aux cO cmul ceqb q0 c)
(** val pmulC :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol
- -> 'a1 -> 'a1 pol **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol ->
+ 'a1 -> 'a1 pol **)
let pmulC cO cI cmul ceqb p c =
if ceqb c cO
@@ -769,8 +753,8 @@ let pmulC cO cI cmul ceqb p c =
else if ceqb c cI then p else pmulC_aux cO cmul ceqb p c
(** val pmulI :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol
- -> 'a1 pol -> 'a1 pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol ->
+ 'a1 pol -> 'a1 pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol **)
let rec pmulI cO cI cmul ceqb pmul0 q0 j = function
| Pc c -> mkPinj j (pmulC cO cI cmul ceqb q0 c)
@@ -791,13 +775,12 @@ let rec pmulI cO cI cmul ceqb pmul0 q0 j = function
mkPX cO ceqb (pmulI cO cI cmul ceqb pmul0 q0 XH p') i' (pmul0 q' q0))
(** val pmul :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
- 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol **)
let rec pmul cO cI cadd cmul ceqb p p'' = match p'' with
| Pc c -> pmulC cO cI cmul ceqb p c
-| Pinj (j', q') ->
- pmulI cO cI cmul ceqb (pmul cO cI cadd cmul ceqb) q' j' p
+| Pinj (j', q') -> pmulI cO cI cmul ceqb (pmul cO cI cadd cmul ceqb) q' j' p
| PX (p', i', q') ->
(match p with
| Pc c -> pmulC cO cI cmul ceqb p'' c
@@ -806,24 +789,22 @@ let rec pmul cO cI cadd cmul ceqb p p'' = match p'' with
match j with
| XI j0 -> pmul cO cI cadd cmul ceqb (Pinj ((XO j0), q0)) q'
| XO j0 ->
- pmul cO cI cadd cmul ceqb (Pinj ((Coq_Pos.pred_double j0), q0))
- q'
+ pmul cO cI cadd cmul ceqb (Pinj ((Coq_Pos.pred_double j0), q0)) q'
| XH -> pmul cO cI cadd cmul ceqb q0 q'
in
mkPX cO ceqb (pmul cO cI cadd cmul ceqb p p') i' qQ'
| PX (p2, i, q0) ->
let qQ' = pmul cO cI cadd cmul ceqb q0 q' in
- let pQ' = pmulI cO cI cmul ceqb (pmul cO cI cadd cmul ceqb) q' XH p2
- in
+ let pQ' = pmulI cO cI cmul ceqb (pmul cO cI cadd cmul ceqb) q' XH p2 in
let qP' = pmul cO cI cadd cmul ceqb (mkPinj XH q0) p' in
let pP' = pmul cO cI cadd cmul ceqb p2 p' in
padd cO cadd ceqb
- (mkPX cO ceqb (padd cO cadd ceqb (mkPX cO ceqb pP' i (p0 cO)) qP')
- i' (p0 cO)) (mkPX cO ceqb pQ' i qQ'))
+ (mkPX cO ceqb (padd cO cadd ceqb (mkPX cO ceqb pP' i (p0 cO)) qP') i'
+ (p0 cO)) (mkPX cO ceqb pQ' i qQ'))
(** val psquare :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
- 'a1 -> bool) -> 'a1 pol -> 'a1 pol **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> bool) -> 'a1 pol -> 'a1 pol **)
let rec psquare cO cI cadd cmul ceqb = function
| Pc c -> Pc (cmul c c)
@@ -852,9 +833,9 @@ let mk_X cO cI j =
mkPinj_pred j (mkX cO cI)
(** val ppow_pos :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
- 'a1 -> bool) -> ('a1 pol -> 'a1 pol) -> 'a1 pol -> 'a1 pol -> positive
- -> 'a1 pol **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> bool) -> ('a1 pol -> 'a1 pol) -> 'a1 pol -> 'a1 pol -> positive -> 'a1
+ pol **)
let rec ppow_pos cO cI cadd cmul ceqb subst_l res p = function
| XI p3 ->
@@ -868,17 +849,16 @@ let rec ppow_pos cO cI cadd cmul ceqb subst_l res p = function
| XH -> subst_l (pmul cO cI cadd cmul ceqb res p)
(** val ppow_N :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
- 'a1 -> bool) -> ('a1 pol -> 'a1 pol) -> 'a1 pol -> n -> 'a1 pol **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> bool) -> ('a1 pol -> 'a1 pol) -> 'a1 pol -> n -> 'a1 pol **)
let ppow_N cO cI cadd cmul ceqb subst_l p = function
| N0 -> p1 cI
| Npos p2 -> ppow_pos cO cI cadd cmul ceqb subst_l (p1 cI) p p2
(** val norm_aux :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
- 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExpr ->
- 'a1 pol **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExpr -> 'a1 pol **)
let rec norm_aux cO cI cadd cmul csub copp ceqb = function
| PEc c -> Pc c
@@ -899,8 +879,7 @@ let rec norm_aux cO cI cadd cmul csub copp ceqb = function
padd cO cadd ceqb (norm_aux cO cI cadd cmul csub copp ceqb pe1)
(norm_aux cO cI cadd cmul csub copp ceqb pe2)))
| PEsub (pe1, pe2) ->
- psub cO cadd csub copp ceqb
- (norm_aux cO cI cadd cmul csub copp ceqb pe1)
+ psub cO cadd csub copp ceqb (norm_aux cO cI cadd cmul csub copp ceqb pe1)
(norm_aux cO cI cadd cmul csub copp ceqb pe2)
| PEmul (pe1, pe2) ->
pmul cO cI cadd cmul ceqb (norm_aux cO cI cadd cmul csub copp ceqb pe1)
@@ -947,8 +926,8 @@ let ff =
[]::[]
(** val add_term :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 -> 'a1 clause ->
- 'a1 clause option **)
+ ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 -> 'a1 clause -> 'a1
+ clause option **)
let rec add_term unsat deduce t0 = function
| [] ->
@@ -969,8 +948,8 @@ let rec add_term unsat deduce t0 = function
| None -> None))
(** val or_clause :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 clause -> 'a1
- clause -> 'a1 clause option **)
+ ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 clause -> 'a1 clause
+ -> 'a1 clause option **)
let rec or_clause unsat deduce cl1 cl2 =
match cl1 with
@@ -981,8 +960,8 @@ let rec or_clause unsat deduce cl1 cl2 =
| None -> None)
(** val or_clause_cnf :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 clause -> 'a1 cnf
- -> 'a1 cnf **)
+ ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 clause -> 'a1 cnf ->
+ 'a1 cnf **)
let or_clause_cnf unsat deduce t0 f =
fold_right (fun e acc ->
@@ -991,8 +970,8 @@ let or_clause_cnf unsat deduce t0 f =
| None -> acc) [] f
(** val or_cnf :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 cnf -> 'a1 cnf ->
- 'a1 cnf **)
+ ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 cnf -> 'a1 cnf -> 'a1
+ cnf **)
let rec or_cnf unsat deduce f f' =
match f with
@@ -1002,12 +981,12 @@ let rec or_cnf unsat deduce f f' =
(** val and_cnf : 'a1 cnf -> 'a1 cnf -> 'a1 cnf **)
-let and_cnf f1 f2 =
- app f1 f2
+let and_cnf =
+ app
(** val xcnf :
- ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a2 cnf) ->
- ('a1 -> 'a2 cnf) -> bool -> 'a1 bFormula -> 'a2 cnf **)
+ ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a2 cnf) -> ('a1
+ -> 'a2 cnf) -> bool -> 'a1 bFormula -> 'a2 cnf **)
let rec xcnf unsat deduce normalise0 negate0 pol0 = function
| TT -> if pol0 then tt else ff
@@ -1047,9 +1026,9 @@ let rec cnf_checker checker f l =
| c::l0 -> if checker e c then cnf_checker checker f0 l0 else false)
(** val tauto_checker :
- ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a2 cnf) ->
- ('a1 -> 'a2 cnf) -> ('a2 list -> 'a3 -> bool) -> 'a1 bFormula -> 'a3
- list -> bool **)
+ ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a2 cnf) -> ('a1
+ -> 'a2 cnf) -> ('a2 list -> 'a3 -> bool) -> 'a1 bFormula -> 'a3 list ->
+ bool **)
let tauto_checker unsat deduce normalise0 negate0 checker f w =
cnf_checker checker (xcnf unsat deduce normalise0 negate0 true f) w
@@ -1085,10 +1064,9 @@ let opMult o o' =
| Equal -> Some Equal
| NonEqual -> Some NonEqual
| _ -> None)
- | Strict ->
- (match o' with
- | NonEqual -> None
- | _ -> Some o')
+ | Strict -> (match o' with
+ | NonEqual -> None
+ | _ -> Some o')
| NonStrict ->
(match o' with
| Equal -> Some Equal
@@ -1100,14 +1078,12 @@ let opMult o o' =
let opAdd o o' =
match o with
| Equal -> Some o'
- | NonEqual ->
- (match o' with
- | Equal -> Some NonEqual
- | _ -> None)
- | Strict ->
- (match o' with
- | NonEqual -> None
- | _ -> Some Strict)
+ | NonEqual -> (match o' with
+ | Equal -> Some NonEqual
+ | _ -> None)
+ | Strict -> (match o' with
+ | NonEqual -> None
+ | _ -> Some Strict)
| NonStrict ->
(match o' with
| Equal -> Some NonStrict
@@ -1134,15 +1110,14 @@ let map_option f = function
let map_option2 f o o' =
match o with
- | Some x ->
- (match o' with
- | Some x' -> f x x'
- | None -> None)
+ | Some x -> (match o' with
+ | Some x' -> f x x'
+ | None -> None)
| None -> None
(** val pexpr_times_nformula :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
- 'a1 -> bool) -> 'a1 polC -> 'a1 nFormula -> 'a1 nFormula option **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> bool) -> 'a1 polC -> 'a1 nFormula -> 'a1 nFormula option **)
let pexpr_times_nformula cO cI cplus ctimes ceqb e = function
| ef,o ->
@@ -1151,8 +1126,8 @@ let pexpr_times_nformula cO cI cplus ctimes ceqb e = function
| _ -> None)
(** val nformula_times_nformula :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
- 'a1 -> bool) -> 'a1 nFormula -> 'a1 nFormula -> 'a1 nFormula option **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> bool) -> 'a1 nFormula -> 'a1 nFormula -> 'a1 nFormula option **)
let nformula_times_nformula cO cI cplus ctimes ceqb f1 f2 =
let e1,o1 = f1 in
@@ -1161,8 +1136,8 @@ let nformula_times_nformula cO cI cplus ctimes ceqb f1 f2 =
(opMult o1 o2)
(** val nformula_plus_nformula :
- 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula ->
- 'a1 nFormula -> 'a1 nFormula option **)
+ 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula -> 'a1
+ nFormula -> 'a1 nFormula option **)
let nformula_plus_nformula cO cplus ceqb f1 f2 =
let e1,o1 = f1 in
@@ -1170,9 +1145,9 @@ let nformula_plus_nformula cO cplus ceqb f1 f2 =
map_option (fun x -> Some ((padd cO cplus ceqb e1 e2),x)) (opAdd o1 o2)
(** val eval_Psatz :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
- 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz
- -> 'a1 nFormula option **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz -> 'a1
+ nFormula option **)
let rec eval_Psatz cO cI cplus ctimes ceqb cleb l = function
| PsatzIn n0 -> Some (nth n0 l ((Pc cO),Equal))
@@ -1207,9 +1182,8 @@ let check_inconsistent cO ceqb cleb = function
| _ -> false)
(** val check_normalised_formulas :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
- 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz
- -> bool **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz -> bool **)
let check_normalised_formulas cO cI cplus ctimes ceqb cleb l cm =
match eval_Psatz cO cI cplus ctimes ceqb cleb l cm with
@@ -1227,31 +1201,30 @@ type op2 =
type 't formula = { flhs : 't pExpr; fop : op2; frhs : 't pExpr }
(** val norm :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
- 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExpr ->
- 'a1 pol **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExpr -> 'a1 pol **)
-let norm cO cI cplus ctimes cminus copp ceqb =
- norm_aux cO cI cplus ctimes cminus copp ceqb
+let norm =
+ norm_aux
(** val psub0 :
- 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) ->
- ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol **)
+ 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1
+ -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol **)
-let psub0 cO cplus cminus copp ceqb =
- psub cO cplus cminus copp ceqb
+let psub0 =
+ psub
(** val padd0 :
- 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1
- pol -> 'a1 pol **)
+ 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol
+ -> 'a1 pol **)
-let padd0 cO cplus ceqb =
- padd cO cplus ceqb
+let padd0 =
+ padd
(** val xnormalise :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
- 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula ->
- 'a1 nFormula list **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1
+ nFormula list **)
let xnormalise cO cI cplus ctimes cminus copp ceqb t0 =
let { flhs = lhs; fop = o; frhs = rhs } = t0 in
@@ -1259,11 +1232,9 @@ let xnormalise cO cI cplus ctimes cminus copp ceqb t0 =
let rhs0 = norm cO cI cplus ctimes cminus copp ceqb rhs in
(match o with
| OpEq ->
- ((psub0 cO cplus cminus copp ceqb lhs0 rhs0),Strict)::(((psub0 cO
- cplus
+ ((psub0 cO cplus cminus copp ceqb lhs0 rhs0),Strict)::(((psub0 cO cplus
cminus copp
- ceqb rhs0
- lhs0),Strict)::[])
+ ceqb rhs0 lhs0),Strict)::[])
| OpNEq -> ((psub0 cO cplus cminus copp ceqb lhs0 rhs0),Equal)::[]
| OpLe -> ((psub0 cO cplus cminus copp ceqb lhs0 rhs0),Strict)::[]
| OpGe -> ((psub0 cO cplus cminus copp ceqb rhs0 lhs0),Strict)::[]
@@ -1271,17 +1242,17 @@ let xnormalise cO cI cplus ctimes cminus copp ceqb t0 =
| OpGt -> ((psub0 cO cplus cminus copp ceqb rhs0 lhs0),NonStrict)::[])
(** val cnf_normalise :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
- 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula ->
- 'a1 nFormula cnf **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1
+ nFormula cnf **)
let cnf_normalise cO cI cplus ctimes cminus copp ceqb t0 =
map (fun x -> x::[]) (xnormalise cO cI cplus ctimes cminus copp ceqb t0)
(** val xnegate :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
- 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula ->
- 'a1 nFormula list **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1
+ nFormula list **)
let xnegate cO cI cplus ctimes cminus copp ceqb t0 =
let { flhs = lhs; fop = o; frhs = rhs } = t0 in
@@ -1290,20 +1261,18 @@ let xnegate cO cI cplus ctimes cminus copp ceqb t0 =
(match o with
| OpEq -> ((psub0 cO cplus cminus copp ceqb lhs0 rhs0),Equal)::[]
| OpNEq ->
- ((psub0 cO cplus cminus copp ceqb lhs0 rhs0),Strict)::(((psub0 cO
- cplus
+ ((psub0 cO cplus cminus copp ceqb lhs0 rhs0),Strict)::(((psub0 cO cplus
cminus copp
- ceqb rhs0
- lhs0),Strict)::[])
+ ceqb rhs0 lhs0),Strict)::[])
| OpLe -> ((psub0 cO cplus cminus copp ceqb rhs0 lhs0),NonStrict)::[]
| OpGe -> ((psub0 cO cplus cminus copp ceqb lhs0 rhs0),NonStrict)::[]
| OpLt -> ((psub0 cO cplus cminus copp ceqb rhs0 lhs0),Strict)::[]
| OpGt -> ((psub0 cO cplus cminus copp ceqb lhs0 rhs0),Strict)::[])
(** val cnf_negate :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
- 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula ->
- 'a1 nFormula cnf **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1
+ nFormula cnf **)
let cnf_negate cO cI cplus ctimes cminus copp ceqb t0 =
map (fun x -> x::[]) (xnegate cO cI cplus ctimes cminus copp ceqb t0)
@@ -1340,8 +1309,8 @@ let map_Formula c_of_S f =
{ flhs = (map_PExpr c_of_S l); fop = o; frhs = (map_PExpr c_of_S r) }
(** val simpl_cone :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 psatz
- -> 'a1 psatz **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 psatz ->
+ 'a1 psatz **)
let simpl_cone cO cI ctimes ceqb e = match e with
| PsatzSquare t0 ->
@@ -1379,8 +1348,7 @@ let simpl_cone cO cI ctimes ceqb e = match e with
| PsatzC p2 -> PsatzMulE ((PsatzC (ctimes c p2)), x)
| _ -> if ceqb cI c then t2 else PsatzMulE (t1, t2)))
| PsatzAdd (y, z0) ->
- PsatzAdd ((PsatzMulE ((PsatzC c), y)), (PsatzMulE ((PsatzC c),
- z0)))
+ PsatzAdd ((PsatzMulE ((PsatzC c), y)), (PsatzMulE ((PsatzC c), z0)))
| PsatzC c0 -> PsatzC (ctimes c c0)
| PsatzZ -> PsatzZ
| _ -> if ceqb cI c then t2 else PsatzMulE (t1, t2))
@@ -1393,10 +1361,9 @@ let simpl_cone cO cI ctimes ceqb e = match e with
| PsatzAdd (t1, t2) ->
(match t1 with
| PsatzZ -> t2
- | _ ->
- (match t2 with
- | PsatzZ -> t1
- | _ -> PsatzAdd (t1, t2)))
+ | _ -> (match t2 with
+ | PsatzZ -> t1
+ | _ -> PsatzAdd (t1, t2)))
| _ -> e
type q = { qnum : z; qden : positive }
@@ -1422,8 +1389,7 @@ let qle_bool x y =
(** val qplus : q -> q -> q **)
let qplus x y =
- { qnum =
- (Z.add (Z.mul x.qnum (Zpos y.qden)) (Z.mul y.qnum (Zpos x.qden)));
+ { qnum = (Z.add (Z.mul x.qnum (Zpos y.qden)) (Z.mul y.qnum (Zpos x.qden)));
qden = (Coq_Pos.mul x.qden y.qden) }
(** val qmult : q -> q -> q **)
@@ -1635,8 +1601,7 @@ let genCuttingPlane = function
then None
else Some ((makeCuttingPlane e),Equal)
| NonEqual -> Some ((e,Z0),op)
- | Strict ->
- Some ((makeCuttingPlane (psubC Z.sub e (Zpos XH))),NonStrict)
+ | Strict -> Some ((makeCuttingPlane (psubC Z.sub e (Zpos XH))),NonStrict)
| NonStrict -> Some ((makeCuttingPlane e),NonStrict))
(** val nformula_of_cutting_plane : ((z polC * z) * op1) -> z nFormula **)
@@ -1647,10 +1612,9 @@ let nformula_of_cutting_plane = function
(** val is_pol_Z0 : z polC -> bool **)
let is_pol_Z0 = function
-| Pc z0 ->
- (match z0 with
- | Z0 -> true
- | _ -> false)
+| Pc z0 -> (match z0 with
+ | Z0 -> true
+ | _ -> false)
| _ -> false
(** val eval_Psatz0 : z nFormula list -> zWitness -> z nFormula option **)
@@ -1730,8 +1694,8 @@ let qnormalise =
(** val qnegate : q formula -> q nFormula cnf **)
let qnegate =
- cnf_negate { qnum = Z0; qden = XH } { qnum = (Zpos XH); qden = XH }
- qplus qmult qminus qopp qeq_bool
+ cnf_negate { qnum = Z0; qden = XH } { qnum = (Zpos XH); qden = XH } qplus
+ qmult qminus qopp qeq_bool
(** val qunsat : q nFormula -> bool **)
@@ -1789,8 +1753,8 @@ let rnormalise =
(** val rnegate : q formula -> q nFormula cnf **)
let rnegate =
- cnf_negate { qnum = Z0; qden = XH } { qnum = (Zpos XH); qden = XH }
- qplus qmult qminus qopp qeq_bool
+ cnf_negate { qnum = Z0; qden = XH } { qnum = (Zpos XH); qden = XH } qplus
+ qmult qminus qopp qeq_bool
(** val runsat : q nFormula -> bool **)
diff --git a/plugins/micromega/micromega.mli b/plugins/micromega/micromega.mli
index beb042f4..96197817 100644
--- a/plugins/micromega/micromega.mli
+++ b/plugins/micromega/micromega.mli
@@ -1,3 +1,4 @@
+
val negb : bool -> bool
type nat =
@@ -168,44 +169,44 @@ val paddI :
positive -> 'a1 pol -> 'a1 pol
val psubI :
- ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 pol -> 'a1 pol -> 'a1 pol)
- -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol
+ ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 pol -> 'a1 pol -> 'a1 pol) ->
+ 'a1 pol -> positive -> 'a1 pol -> 'a1 pol
val paddX :
- 'a1 -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1
- pol -> positive -> 'a1 pol -> 'a1 pol
+ 'a1 -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1 pol
+ -> positive -> 'a1 pol -> 'a1 pol
val psubX :
- 'a1 -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol ->
- 'a1 pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol
+ 'a1 -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol -> 'a1
+ pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol
val padd :
- 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol
- -> 'a1 pol
+ 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol ->
+ 'a1 pol
val psub :
- 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) ->
- ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol
+ 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1
+ -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol
val pmulC_aux :
- 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 ->
- 'a1 pol
+ 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 -> 'a1
+ pol
val pmulC :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol ->
- 'a1 -> 'a1 pol
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1
+ -> 'a1 pol
val pmulI :
'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol ->
'a1 pol -> 'a1 pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol
val pmul :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
+ bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol
val psquare :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> bool) -> 'a1 pol -> 'a1 pol
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
+ bool) -> 'a1 pol -> 'a1 pol
type 'c pExpr =
| PEc of 'c
@@ -219,17 +220,16 @@ type 'c pExpr =
val mk_X : 'a1 -> 'a1 -> positive -> 'a1 pol
val ppow_pos :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> bool) -> ('a1 pol -> 'a1 pol) -> 'a1 pol -> 'a1 pol -> positive ->
- 'a1 pol
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
+ bool) -> ('a1 pol -> 'a1 pol) -> 'a1 pol -> 'a1 pol -> positive -> 'a1 pol
val ppow_N :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> bool) -> ('a1 pol -> 'a1 pol) -> 'a1 pol -> n -> 'a1 pol
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
+ bool) -> ('a1 pol -> 'a1 pol) -> 'a1 pol -> n -> 'a1 pol
val norm_aux :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExpr -> 'a1 pol
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
+ 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExpr -> 'a1 pol
type 'a bFormula =
| TT
@@ -256,34 +256,31 @@ val add_term :
clause option
val or_clause :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 clause -> 'a1 clause
- -> 'a1 clause option
+ ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 clause -> 'a1 clause ->
+ 'a1 clause option
val or_clause_cnf :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 clause -> 'a1 cnf ->
- 'a1 cnf
+ ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 clause -> 'a1 cnf -> 'a1
+ cnf
val or_cnf :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 cnf -> 'a1 cnf -> 'a1
- cnf
+ ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 cnf -> 'a1 cnf -> 'a1 cnf
val and_cnf : 'a1 cnf -> 'a1 cnf -> 'a1 cnf
val xcnf :
- ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a2 cnf) -> ('a1
- -> 'a2 cnf) -> bool -> 'a1 bFormula -> 'a2 cnf
+ ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a2 cnf) -> ('a1 ->
+ 'a2 cnf) -> bool -> 'a1 bFormula -> 'a2 cnf
val cnf_checker : ('a1 list -> 'a2 -> bool) -> 'a1 cnf -> 'a2 list -> bool
val tauto_checker :
- ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a2 cnf) -> ('a1
- -> 'a2 cnf) -> ('a2 list -> 'a3 -> bool) -> 'a1 bFormula -> 'a3 list ->
- bool
+ ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a2 cnf) -> ('a1 ->
+ 'a2 cnf) -> ('a2 list -> 'a3 -> bool) -> 'a1 bFormula -> 'a3 list -> bool
val cneqb : ('a1 -> 'a1 -> bool) -> 'a1 -> 'a1 -> bool
-val cltb :
- ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 -> 'a1 -> bool
+val cltb : ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 -> 'a1 -> bool
type 'c polC = 'c pol
@@ -314,30 +311,28 @@ val map_option2 :
('a1 -> 'a2 -> 'a3 option) -> 'a1 option -> 'a2 option -> 'a3 option
val pexpr_times_nformula :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> bool) -> 'a1 polC -> 'a1 nFormula -> 'a1 nFormula option
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
+ bool) -> 'a1 polC -> 'a1 nFormula -> 'a1 nFormula option
val nformula_times_nformula :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> bool) -> 'a1 nFormula -> 'a1 nFormula -> 'a1 nFormula option
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
+ bool) -> 'a1 nFormula -> 'a1 nFormula -> 'a1 nFormula option
val nformula_plus_nformula :
- 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula ->
- 'a1 nFormula -> 'a1 nFormula option
+ 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula -> 'a1
+ nFormula -> 'a1 nFormula option
val eval_Psatz :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz ->
- 'a1 nFormula option
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
+ bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz -> 'a1
+ nFormula option
val check_inconsistent :
- 'a1 -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula ->
- bool
+ 'a1 -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula -> bool
val check_normalised_formulas :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz ->
- bool
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
+ bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz -> bool
type op2 =
| OpEq
@@ -350,36 +345,36 @@ type op2 =
type 't formula = { flhs : 't pExpr; fop : op2; frhs : 't pExpr }
val norm :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExpr -> 'a1 pol
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
+ 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExpr -> 'a1 pol
val psub0 :
- 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) ->
- ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol
+ 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1
+ -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol
val padd0 :
- 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol
- -> 'a1 pol
+ 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol ->
+ 'a1 pol
val xnormalise :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1
- nFormula list
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
+ 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1 nFormula
+ list
val cnf_normalise :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1
- nFormula cnf
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
+ 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1 nFormula
+ cnf
val xnegate :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1
- nFormula list
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
+ 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1 nFormula
+ list
val cnf_negate :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1
- nFormula cnf
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
+ 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1 nFormula
+ cnf
val xdenorm : positive -> 'a1 pol -> 'a1 pExpr
@@ -390,8 +385,8 @@ val map_PExpr : ('a2 -> 'a1) -> 'a2 pExpr -> 'a1 pExpr
val map_Formula : ('a2 -> 'a1) -> 'a2 formula -> 'a1 formula
val simpl_cone :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 psatz
- -> 'a1 psatz
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 psatz ->
+ 'a1 psatz
type q = { qnum : z; qden : positive }
diff --git a/plugins/micromega/mutils.ml b/plugins/micromega/mutils.ml
index b4c6d032..82367c0b 100644
--- a/plugins/micromega/mutils.ml
+++ b/plugins/micromega/mutils.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* *)
(* Micromega: A reflexive tactic using the Positivstellensatz *)
diff --git a/plugins/micromega/persistent_cache.ml b/plugins/micromega/persistent_cache.ml
index 0e6d346a..ee5a0458 100644
--- a/plugins/micromega/persistent_cache.ml
+++ b/plugins/micromega/persistent_cache.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* *)
(* A persistent hashtable *)
@@ -149,7 +151,7 @@ let open_in f =
match read_key_elem inch with
| None -> ()
| Some (key,elem) ->
- Table.add htbl key elem ;
+ Table.replace htbl key elem ;
xload () in
try
(* Locking of the (whole) file while reading *)
@@ -195,7 +197,7 @@ let add t k e =
else
let fd = descr_of_out_channel outch in
begin
- Table.add tbl k e ;
+ Table.replace tbl k e ;
do_under_lock Write fd
(fun _ ->
Marshal.to_channel outch (k,e) [Marshal.No_sharing] ;
diff --git a/plugins/micromega/polynomial.ml b/plugins/micromega/polynomial.ml
index 90a108a3..db8b73a2 100644
--- a/plugins/micromega/polynomial.ml
+++ b/plugins/micromega/polynomial.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* *)
(* Micromega: A reflexive tactic using the Positivstellensatz *)
diff --git a/plugins/micromega/sos.ml b/plugins/micromega/sos.ml
index cc89e2b9..e1ceabe9 100644
--- a/plugins/micromega/sos.ml
+++ b/plugins/micromega/sos.ml
@@ -21,8 +21,6 @@ let debugging = ref false;;
exception Sanity;;
-exception Unsolvable;;
-
(* ------------------------------------------------------------------------- *)
(* Turn a rational into a decimal string with d sig digits. *)
(* ------------------------------------------------------------------------- *)
@@ -99,28 +97,11 @@ let vector_const c n =
if c =/ Int 0 then vector_0 n
else (n,itlist (fun k -> k |-> c) (1--n) undefined :vector);;
-let vector_1 = vector_const (Int 1);;
-
let vector_cmul c (v:vector) =
let n = dim v in
if c =/ Int 0 then vector_0 n
else n,mapf (fun x -> c */ x) (snd v)
-let vector_neg (v:vector) = (fst v,mapf minus_num (snd v) :vector);;
-
-let vector_add (v1:vector) (v2:vector) =
- let m = dim v1 and n = dim v2 in
- if m <> n then failwith "vector_add: incompatible dimensions" else
- (n,combine (+/) (fun x -> x =/ Int 0) (snd v1) (snd v2) :vector);;
-
-let vector_sub v1 v2 = vector_add v1 (vector_neg v2);;
-
-let vector_dot (v1:vector) (v2:vector) =
- let m = dim v1 and n = dim v2 in
- if m <> n then failwith "vector_add: incompatible dimensions" else
- foldl (fun a i x -> x +/ a) (Int 0)
- (combine ( */ ) (fun x -> x =/ Int 0) (snd v1) (snd v2));;
-
let vector_of_list l =
let n = List.length l in
(n,itlist2 (|->) (1--n) l undefined :vector);;
@@ -133,13 +114,6 @@ let matrix_0 (m,n) = ((m,n),undefined:matrix);;
let dimensions (m:matrix) = fst m;;
-let matrix_const c (m,n as mn) =
- if m <> n then failwith "matrix_const: needs to be square"
- else if c =/ Int 0 then matrix_0 mn
- else (mn,itlist (fun k -> (k,k) |-> c) (1--n) undefined :matrix);;
-
-let matrix_1 = matrix_const (Int 1);;
-
let matrix_cmul c (m:matrix) =
let (i,j) = dimensions m in
if c =/ Int 0 then matrix_0 (i,j)
@@ -152,8 +126,6 @@ let matrix_add (m1:matrix) (m2:matrix) =
if d1 <> d2 then failwith "matrix_add: incompatible dimensions"
else (d1,combine (+/) (fun x -> x =/ Int 0) (snd m1) (snd m2) :matrix);;
-let matrix_sub m1 m2 = matrix_add m1 (matrix_neg m2);;
-
let row k (m:matrix) =
let i,j = dimensions m in
(j,
@@ -166,20 +138,10 @@ let column k (m:matrix) =
foldl (fun a (i,j) c -> if j = k then (i |-> c) a else a) undefined (snd m)
: vector);;
-let transp (m:matrix) =
- let i,j = dimensions m in
- ((j,i),foldl (fun a (i,j) c -> ((j,i) |-> c) a) undefined (snd m) :matrix);;
-
let diagonal (v:vector) =
let n = dim v in
((n,n),foldl (fun a i c -> ((i,i) |-> c) a) undefined (snd v) : matrix);;
-let matrix_of_list l =
- let m = List.length l in
- if m = 0 then matrix_0 (0,0) else
- let n = List.length (List.hd l) in
- (m,n),itern 1 l (fun v i -> itern 1 v (fun c j -> (i,j) |-> c)) undefined;;
-
(* ------------------------------------------------------------------------- *)
(* Monomials. *)
(* ------------------------------------------------------------------------- *)
@@ -195,24 +157,8 @@ let monomial_var x = (x |=> 1 :monomial);;
let (monomial_mul:monomial->monomial->monomial) =
combine (+) (fun x -> false);;
-let monomial_pow (m:monomial) k =
- if k = 0 then monomial_1
- else mapf (fun x -> k * x) m;;
-
-let monomial_divides (m1:monomial) (m2:monomial) =
- foldl (fun a x k -> tryapplyd m2 x 0 >= k && a) true m1;;
-
-let monomial_div (m1:monomial) (m2:monomial) =
- let m = combine (+) (fun x -> x = 0) m1 (mapf (fun x -> -x) m2) in
- if foldl (fun a x k -> k >= 0 && a) true m then m
- else failwith "monomial_div: non-divisible";;
-
let monomial_degree x (m:monomial) = tryapplyd m x 0;;
-let monomial_lcm (m1:monomial) (m2:monomial) =
- (itlist (fun x -> x |-> max (monomial_degree x m1) (monomial_degree x m2))
- (union (dom m1) (dom m2)) undefined :monomial);;
-
let monomial_multidegree (m:monomial) = foldl (fun a x k -> k + a) 0 m;;
let monomial_variables m = dom m;;
@@ -252,12 +198,6 @@ let poly_cmmul (c,m) (p:poly) =
let poly_mul (p1:poly) (p2:poly) =
foldl (fun a m c -> poly_add (poly_cmmul (c,m) p2) a) poly_0 p1;;
-let poly_div (p1:poly) (p2:poly) =
- if not(poly_isconst p2) then failwith "poly_div: non-constant" else
- let c = eval undefined p2 in
- if c =/ Int 0 then failwith "poly_div: division by zero"
- else poly_cmul (Int 1 // c) p1;;
-
let poly_square p = poly_mul p p;;
let rec poly_pow p k =
@@ -266,10 +206,6 @@ let rec poly_pow p k =
else let q = poly_square(poly_pow p (k / 2)) in
if k mod 2 = 1 then poly_mul p q else q;;
-let poly_exp p1 p2 =
- if not(poly_isconst p2) then failwith "poly_exp: not a constant" else
- poly_pow p1 (Num.int_of_num (eval undefined p2));;
-
let degree x (p:poly) = foldl (fun a m c -> max (monomial_degree x m) a) 0 p;;
let multidegree (p:poly) =
@@ -282,14 +218,14 @@ let poly_variables (p:poly) =
(* Order monomials for human presentation. *)
(* ------------------------------------------------------------------------- *)
-let humanorder_varpow (x1,k1) (x2,k2) = x1 < x2 or x1 = x2 && k1 > k2;;
+let humanorder_varpow (x1,k1) (x2,k2) = x1 < x2 || x1 = x2 && k1 > k2;;
let humanorder_monomial =
let rec ord l1 l2 = match (l1,l2) with
_,[] -> true
| [],_ -> false
- | h1::t1,h2::t2 -> humanorder_varpow h1 h2 or h1 = h2 && ord t1 t2 in
- fun m1 m2 -> m1 = m2 or
+ | h1::t1,h2::t2 -> humanorder_varpow h1 h2 || h1 = h2 && ord t1 t2 in
+ fun m1 m2 -> m1 = m2 ||
ord (sort humanorder_varpow (graph m1))
(sort humanorder_varpow (graph m2));;
@@ -297,42 +233,8 @@ let humanorder_monomial =
(* Conversions to strings. *)
(* ------------------------------------------------------------------------- *)
-let string_of_vector min_size max_size (v:vector) =
- let n_raw = dim v in
- if n_raw = 0 then "[]" else
- let n = max min_size (min n_raw max_size) in
- let xs = List.map ((o) string_of_num (element v)) (1--n) in
- "[" ^ end_itlist (fun s t -> s ^ ", " ^ t) xs ^
- (if n_raw > max_size then ", ...]" else "]");;
-
-let string_of_matrix max_size (m:matrix) =
- let i_raw,j_raw = dimensions m in
- let i = min max_size i_raw and j = min max_size j_raw in
- let rstr = List.map (fun k -> string_of_vector j j (row k m)) (1--i) in
- "["^end_itlist(fun s t -> s^";\n "^t) rstr ^
- (if j > max_size then "\n ...]" else "]");;
-
let string_of_vname (v:vname): string = (v: string);;
-let rec string_of_term t =
- match t with
- Opp t1 -> "(- " ^ string_of_term t1 ^ ")"
-| Add (t1, t2) ->
- "(" ^ (string_of_term t1) ^ " + " ^ (string_of_term t2) ^ ")"
-| Sub (t1, t2) ->
- "(" ^ (string_of_term t1) ^ " - " ^ (string_of_term t2) ^ ")"
-| Mul (t1, t2) ->
- "(" ^ (string_of_term t1) ^ " * " ^ (string_of_term t2) ^ ")"
-| Inv t1 -> "(/ " ^ string_of_term t1 ^ ")"
-| Div (t1, t2) ->
- "(" ^ (string_of_term t1) ^ " / " ^ (string_of_term t2) ^ ")"
-| Pow (t1, n1) ->
- "(" ^ (string_of_term t1) ^ " ^ " ^ (string_of_int n1) ^ ")"
-| Zero -> "0"
-| Var v -> "x" ^ (string_of_vname v)
-| Const x -> string_of_num x;;
-
-
let string_of_varpow x k =
if k = 1 then string_of_vname x else string_of_vname x^"^"^string_of_int k;;
@@ -363,6 +265,7 @@ let string_of_poly (p:poly) =
(* Printers. *)
(* ------------------------------------------------------------------------- *)
+(*
let print_vector v = Format.print_string(string_of_vector 0 20 v);;
let print_matrix m = Format.print_string(string_of_matrix 20 m);;
@@ -371,7 +274,6 @@ let print_monomial m = Format.print_string(string_of_monomial m);;
let print_poly m = Format.print_string(string_of_poly m);;
-(*
#install_printer print_vector;;
#install_printer print_matrix;;
#install_printer print_monomial;;
@@ -411,19 +313,6 @@ let sdpa_of_vector (v:vector) =
end_itlist (fun x y -> x ^ " " ^ y) strs ^ "\n";;
(* ------------------------------------------------------------------------- *)
-(* String for block diagonal matrix numbered k. *)
-(* ------------------------------------------------------------------------- *)
-
-let sdpa_of_blockdiagonal k m =
- let pfx = string_of_int k ^" " in
- let ents =
- foldl (fun a (b,i,j) c -> if i > j then a else ((b,i,j),c)::a) [] m in
- let entss = sort (increasing fst) ents in
- itlist (fun ((b,i,j),c) a ->
- pfx ^ string_of_int b ^ " " ^ string_of_int i ^ " " ^ string_of_int j ^
- " " ^ decimalize 20 c ^ "\n" ^ a) entss "";;
-
-(* ------------------------------------------------------------------------- *)
(* String for a matrix numbered k, in SDPA sparse format. *)
(* ------------------------------------------------------------------------- *)
@@ -466,6 +355,7 @@ let token s =
>> (fun ((_,t),_) -> t);;
let decimal =
+ let (||) = parser_or in
let numeral = some isnum in
let decimalint = atleast 1 numeral >> ((o) Num.num_of_string implode) in
let decimalfrac = atleast 1 numeral
@@ -485,13 +375,12 @@ let mkparser p s =
let x,rst = p(explode s) in
if rst = [] then x else failwith "mkparser: unparsed input";;
-let parse_decimal = mkparser decimal;;
-
(* ------------------------------------------------------------------------- *)
(* Parse back a vector. *)
(* ------------------------------------------------------------------------- *)
-let parse_sdpaoutput,parse_csdpoutput =
+let _parse_sdpaoutput, parse_csdpoutput =
+ let (||) = parser_or in
let vector =
token "{" ++ listof decimal (token ",") "decimal" ++ token "}"
>> (fun ((_,v),_) -> vector_of_list v) in
@@ -508,23 +397,10 @@ let parse_sdpaoutput,parse_csdpoutput =
mkparser sdpaoutput,mkparser csdpoutput;;
(* ------------------------------------------------------------------------- *)
-(* Also parse the SDPA output to test success (CSDP yields a return code). *)
-(* ------------------------------------------------------------------------- *)
-
-let sdpa_run_succeeded =
- let rec skipupto dscr prs inp =
- (dscr ++ prs >> snd
- || some (fun c -> true) ++ skipupto dscr prs >> snd) inp in
- let prs = skipupto (word "phase.value" ++ token "=")
- (possibly (a "p") ++ possibly (a "d") ++
- (word "OPT" || word "FEAS")) in
- fun s -> try ignore (prs (explode s)); true with Noparse -> false;;
-
-(* ------------------------------------------------------------------------- *)
(* The default parameters. Unfortunately this goes to a fixed file. *)
(* ------------------------------------------------------------------------- *)
-let sdpa_default_parameters =
+let _sdpa_default_parameters =
"100 unsigned int maxIteration;\
\n1.0E-7 double 0.0 < epsilonStar;\
\n1.0E2 double 0.0 < lambdaStar;\
@@ -555,7 +431,7 @@ let sdpa_alt_parameters =
\n1.0E-7 double 0.0 < epsilonDash;\
\n";;
-let sdpa_params = sdpa_alt_parameters;;
+let _sdpa_params = sdpa_alt_parameters;;
(* ------------------------------------------------------------------------- *)
(* CSDP parameters; so far I'm sticking with the defaults. *)
@@ -588,10 +464,10 @@ let run_csdp dbg obj mats =
let input_file = Filename.temp_file "sos" ".dat-s" in
let output_file =
String.sub input_file 0 (String.length input_file - 6) ^ ".out"
- and params_file = Filename.concat (!temp_path) "param.csdp" in
+ and params_file = Filename.concat temp_path "param.csdp" in
file_of_string input_file (sdpa_of_problem "" obj mats);
file_of_string params_file csdp_params;
- let rv = Sys.command("cd "^(!temp_path)^"; csdp "^input_file ^
+ let rv = Sys.command("cd "^temp_path^"; csdp "^input_file ^
" " ^ output_file ^
(if dbg then "" else "> /dev/null")) in
let op = string_of_file output_file in
@@ -600,16 +476,6 @@ let run_csdp dbg obj mats =
else (Sys.remove input_file; Sys.remove output_file));
rv,res);;
-let csdp obj mats =
- let rv,res = run_csdp (!debugging) obj mats in
- (if rv = 1 or rv = 2 then failwith "csdp: Problem is infeasible"
- else if rv = 3 then ()
- (* Format.print_string "csdp warning: Reduced accuracy";
- Format.print_newline() *)
- else if rv <> 0 then failwith("csdp: error "^string_of_int rv)
- else ());
- res;;
-
(* ------------------------------------------------------------------------- *)
(* Try some apparently sensible scaling first. Note that this is purely to *)
(* get a cleaner translation to floating-point, and doesn't affect any of *)
@@ -653,21 +519,7 @@ let linear_program_basic a =
let mats = List.map (fun j -> diagonal (column j a)) (1--n)
and obj = vector_const (Int 1) m in
let rv,res = run_csdp false obj mats in
- if rv = 1 or rv = 2 then false
- else if rv = 0 then true
- else failwith "linear_program: An error occurred in the SDP solver";;
-
-(* ------------------------------------------------------------------------- *)
-(* Alternative interface testing A x >= b for matrix A, vector b. *)
-(* ------------------------------------------------------------------------- *)
-
-let linear_program a b =
- let m,n = dimensions a in
- if dim b <> m then failwith "linear_program: incompatible dimensions" else
- let mats = diagonal b :: List.map (fun j -> diagonal (column j a)) (1--n)
- and obj = vector_const (Int 1) m in
- let rv,res = run_csdp false obj mats in
- if rv = 1 or rv = 2 then false
+ if rv = 1 || rv = 2 then false
else if rv = 0 then true
else failwith "linear_program: An error occurred in the SDP solver";;
@@ -716,40 +568,6 @@ let equation_eval assig eq =
foldl (fun a v c -> a +/ value(v) */ c) (Int 0) eq;;
(* ------------------------------------------------------------------------- *)
-(* Eliminate among linear equations: return unconstrained variables and *)
-(* assignments for the others in terms of them. We give one pseudo-variable *)
-(* "one" that's used for a constant term. *)
-(* ------------------------------------------------------------------------- *)
-
-let failstore = ref [];;
-
-let eliminate_equations =
- let rec extract_first p l =
- match l with
- [] -> failwith "extract_first"
- | h::t -> if p(h) then h,t else
- let k,s = extract_first p t in
- k,h::s in
- let rec eliminate vars dun eqs =
- match vars with
- [] -> if forall is_undefined eqs then dun
- else (failstore := [vars,dun,eqs]; raise Unsolvable)
- | v::vs ->
- try let eq,oeqs = extract_first (fun e -> defined e v) eqs in
- let a = apply eq v in
- let eq' = equation_cmul (Int(-1) // a) (undefine v eq) in
- let elim e =
- let b = tryapplyd e v (Int 0) in
- if b =/ Int 0 then e else
- equation_add e (equation_cmul (minus_num b // a) eq) in
- eliminate vs ((v |-> eq') (mapf elim dun)) (List.map elim oeqs)
- with Failure _ -> eliminate vs dun eqs in
- fun one vars eqs ->
- let assig = eliminate vars undefined eqs in
- let vs = foldl (fun a x f -> subtract (dom f) [one] @ a) [] assig in
- setify vs,assig;;
-
-(* ------------------------------------------------------------------------- *)
(* Eliminate all variables, in an essentially arbitrary order. *)
(* ------------------------------------------------------------------------- *)
@@ -780,18 +598,6 @@ let eliminate_all_equations one =
setify vs,assig;;
(* ------------------------------------------------------------------------- *)
-(* Solve equations by assigning arbitrary numbers. *)
-(* ------------------------------------------------------------------------- *)
-
-let solve_equations one eqs =
- let vars,assigs = eliminate_all_equations one eqs in
- let vfn = itlist (fun v -> (v |-> Int 0)) vars (one |=> Int(-1)) in
- let ass =
- combine (+/) (fun c -> false) (mapf (equation_eval vfn) assigs) vfn in
- if forall (fun e -> equation_eval ass e =/ Int 0) eqs
- then undefine one ass else raise Sanity;;
-
-(* ------------------------------------------------------------------------- *)
(* Hence produce the "relevant" monomials: those whose squares lie in the *)
(* Newton polytope of the monomials in the input. (This is enough according *)
(* to Reznik: "Extremal PSD forms with few terms", Duke Math. Journal, *)
@@ -898,19 +704,6 @@ let epoly_pmul p q acc =
a q) acc p;;
(* ------------------------------------------------------------------------- *)
-(* Usual operations on equation-parametrized poly. *)
-(* ------------------------------------------------------------------------- *)
-
-let epoly_cmul c l =
- if c =/ Int 0 then undefined else mapf (equation_cmul c) l;;
-
-let epoly_neg = epoly_cmul (Int(-1));;
-
-let epoly_add = combine equation_add is_undefined;;
-
-let epoly_sub p q = epoly_add p (epoly_neg q);;
-
-(* ------------------------------------------------------------------------- *)
(* Convert regular polynomial. Note that we treat (0,0,0) as -1. *)
(* ------------------------------------------------------------------------- *)
@@ -953,11 +746,11 @@ let run_csdp dbg nblocks blocksizes obj mats =
let input_file = Filename.temp_file "sos" ".dat-s" in
let output_file =
String.sub input_file 0 (String.length input_file - 6) ^ ".out"
- and params_file = Filename.concat (!temp_path) "param.csdp" in
+ and params_file = Filename.concat temp_path "param.csdp" in
file_of_string input_file
(sdpa_of_blockproblem "" nblocks blocksizes obj mats);
file_of_string params_file csdp_params;
- let rv = Sys.command("cd "^(!temp_path)^"; csdp "^input_file ^
+ let rv = Sys.command("cd "^temp_path^"; csdp "^input_file ^
" " ^ output_file ^
(if dbg then "" else "> /dev/null")) in
let op = string_of_file output_file in
@@ -968,7 +761,7 @@ let run_csdp dbg nblocks blocksizes obj mats =
let csdp nblocks blocksizes obj mats =
let rv,res = run_csdp (!debugging) nblocks blocksizes obj mats in
- (if rv = 1 or rv = 2 then failwith "csdp: Problem is infeasible"
+ (if rv = 1 || rv = 2 then failwith "csdp: Problem is infeasible"
else if rv = 3 then ()
(*Format.print_string "csdp warning: Reduced accuracy";
Format.print_newline() *)
@@ -988,8 +781,6 @@ let bmatrix_cmul c bm =
let bmatrix_neg = bmatrix_cmul (Int(-1));;
-let bmatrix_sub m1 m2 = bmatrix_add m1 (bmatrix_neg m2);;
-
(* ------------------------------------------------------------------------- *)
(* Smash a block matrix into components. *)
(* ------------------------------------------------------------------------- *)
@@ -1102,15 +893,6 @@ let real_positivnullstellensatz_general linf d eqs leqs pol =
cfs,List.map (fun (a,b) -> snd a,b) msq;;
(* ------------------------------------------------------------------------- *)
-(* Iterative deepening. *)
-(* ------------------------------------------------------------------------- *)
-
-let rec deepen f n =
- try print_string "Searching with depth limit ";
- print_int n; print_newline(); f n
- with Failure _ -> deepen f (n + 1);;
-
-(* ------------------------------------------------------------------------- *)
(* The ordering so we can create canonical HOL polynomials. *)
(* ------------------------------------------------------------------------- *)
@@ -1136,10 +918,6 @@ let monomial_order =
if deg1 < deg2 then false else if deg1 > deg2 then true
else lexorder mon1 mon2;;
-let dest_poly p =
- List.map (fun (m,c) -> c,dest_monomial m)
- (sort (fun (m1,_) (m2,_) -> monomial_order m1 m2) (graph p));;
-
(* ------------------------------------------------------------------------- *)
(* Map back polynomials and their composites to HOL. *)
(* ------------------------------------------------------------------------- *)
@@ -1373,9 +1151,6 @@ let rec allpermutations l =
itlist (fun h acc -> List.map (fun t -> h::t)
(allpermutations (subtract l [h])) @ acc) l [];;
-let allvarorders l =
- List.map (fun vlis x -> index x vlis) (allpermutations l);;
-
let changevariables_monomial zoln (m:monomial) =
foldl (fun a x k -> (List.assoc x zoln |-> k) a) monomial_1 m;;
@@ -1392,15 +1167,6 @@ let sdpa_of_vector (v:vector) =
let strs = List.map (o (decimalize 20) (element v)) (1--n) in
end_itlist (fun x y -> x ^ " " ^ y) strs ^ "\n";;
-let sdpa_of_blockdiagonal k m =
- let pfx = string_of_int k ^" " in
- let ents =
- foldl (fun a (b,i,j) c -> if i > j then a else ((b,i,j),c)::a) [] m in
- let entss = sort (increasing fst) ents in
- itlist (fun ((b,i,j),c) a ->
- pfx ^ string_of_int b ^ " " ^ string_of_int i ^ " " ^ string_of_int j ^
- " " ^ decimalize 20 c ^ "\n" ^ a) entss "";;
-
let sdpa_of_matrix k (m:matrix) =
let pfx = string_of_int k ^ " 1 " in
let ms = foldr (fun (i,j) c a -> if i > j then a else ((i,j),c)::a)
@@ -1425,10 +1191,10 @@ let run_csdp dbg obj mats =
let input_file = Filename.temp_file "sos" ".dat-s" in
let output_file =
String.sub input_file 0 (String.length input_file - 6) ^ ".out"
- and params_file = Filename.concat (!temp_path) "param.csdp" in
+ and params_file = Filename.concat temp_path "param.csdp" in
file_of_string input_file (sdpa_of_problem "" obj mats);
file_of_string params_file csdp_params;
- let rv = Sys.command("cd "^(!temp_path)^"; csdp "^input_file ^
+ let rv = Sys.command("cd "^temp_path^"; csdp "^input_file ^
" " ^ output_file ^
(if dbg then "" else "> /dev/null")) in
let op = string_of_file output_file in
@@ -1439,7 +1205,7 @@ let run_csdp dbg obj mats =
let csdp obj mats =
let rv,res = run_csdp (!debugging) obj mats in
- (if rv = 1 or rv = 2 then failwith "csdp: Problem is infeasible"
+ (if rv = 1 || rv = 2 then failwith "csdp: Problem is infeasible"
else if rv = 3 then ()
(* (Format.print_string "csdp warning: Reduced accuracy";
Format.print_newline()) *)
diff --git a/plugins/micromega/sos.mli b/plugins/micromega/sos.mli
index 1ca27ea2..6e62c563 100644
--- a/plugins/micromega/sos.mli
+++ b/plugins/micromega/sos.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Sos_types
diff --git a/plugins/micromega/sos_lib.ml b/plugins/micromega/sos_lib.ml
index f54914f2..6b8b820a 100644
--- a/plugins/micromega/sos_lib.ml
+++ b/plugins/micromega/sos_lib.ml
@@ -525,7 +525,7 @@ let isspace,issep,isbra,issymb,isalpha,isnum,isalnum =
and isalnum c = Array.get ctable (charcode c) >= 16 in
isspace,issep,isbra,issymb,isalpha,isnum,isalnum;;
-let (||) parser1 parser2 input =
+let parser_or parser1 parser2 input =
try parser1 input
with Noparse -> parser2 input;;
@@ -571,7 +571,7 @@ let finished input =
(* ------------------------------------------------------------------------- *)
-let temp_path = ref Filename.temp_dir_name;;
+let temp_path = Filename.get_temp_dir_name ();;
(* ------------------------------------------------------------------------- *)
(* Convenient conversion between files and (lists of) strings. *)
diff --git a/plugins/micromega/sos_types.ml b/plugins/micromega/sos_types.ml
index 615ac5a2..dde1e6c0 100644
--- a/plugins/micromega/sos_types.ml
+++ b/plugins/micromega/sos_types.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* The type of positivstellensatz -- used to communicate with sos *)
diff --git a/plugins/micromega/sos_types.mli b/plugins/micromega/sos_types.mli
new file mode 100644
index 00000000..050ff1e4
--- /dev/null
+++ b/plugins/micromega/sos_types.mli
@@ -0,0 +1,42 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+(* The type of positivstellensatz -- used to communicate with sos *)
+
+type vname = string;;
+
+type term =
+| Zero
+| Const of Num.num
+| Var of vname
+| Inv of term
+| Opp of term
+| Add of (term * term)
+| Sub of (term * term)
+| Mul of (term * term)
+| Div of (term * term)
+| Pow of (term * int);;
+
+val output_term : out_channel -> term -> unit
+
+type positivstellensatz =
+ Axiom_eq of int
+ | Axiom_le of int
+ | Axiom_lt of int
+ | Rational_eq of Num.num
+ | Rational_le of Num.num
+ | Rational_lt of Num.num
+ | Square of term
+ | Monoid of int list
+ | Eqmul of term * positivstellensatz
+ | Sum of positivstellensatz * positivstellensatz
+ | Product of positivstellensatz * positivstellensatz;;
+
+val output_psatz : out_channel -> positivstellensatz -> unit
diff --git a/plugins/micromega/vo.itarget b/plugins/micromega/vo.itarget
deleted file mode 100644
index c9009ea4..00000000
--- a/plugins/micromega/vo.itarget
+++ /dev/null
@@ -1,15 +0,0 @@
-EnvRing.vo
-Env.vo
-OrderedRing.vo
-Psatz.vo
-QMicromega.vo
-Refl.vo
-RingMicromega.vo
-RMicromega.vo
-Tauto.vo
-VarMap.vo
-ZCoeff.vo
-ZMicromega.vo
-Lia.vo
-Lqa.vo
-Lra.vo