aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Frédéric Besson <frederic.besson@inria.fr>2016-08-30 17:12:27 +0200
committerGravatar Frédéric Besson <frederic.besson@inria.fr>2016-08-30 17:59:59 +0200
commit721637c98514a77d05d080f53f226cab3a8da1e7 (patch)
tree9a04e0482488764d39c0e24847e93f4b23f62cde
parent44ada644ef50563aa52cfcd7717d44bde29e5a20 (diff)
plugin micromega : nra also handles non-linear rational arithmetic over Q (Fixed #4985)
Lqa.v defines the tactics lra and nra working over Q. Lra.v defines the tactics lra and nra working over R.
-rw-r--r--doc/refman/Micromega.tex90
-rw-r--r--plugins/micromega/Lqa.v55
-rw-r--r--plugins/micromega/Lra.v56
-rw-r--r--plugins/micromega/Psatz.v40
-rw-r--r--plugins/micromega/coq_micromega.ml20
-rw-r--r--plugins/micromega/g_micromega.ml410
-rw-r--r--plugins/micromega/vo.itarget4
-rw-r--r--test-suite/micromega/qexample.v17
-rw-r--r--test-suite/micromega/rexample.v12
9 files changed, 218 insertions, 86 deletions
diff --git a/doc/refman/Micromega.tex b/doc/refman/Micromega.tex
index 1efed6ef7..4daf98f87 100644
--- a/doc/refman/Micromega.tex
+++ b/doc/refman/Micromega.tex
@@ -4,16 +4,19 @@
\asection{Short description of the tactics}
-\tacindex{psatz} \tacindex{lra}
+\tacindex{psatz} \tacindex{lra} \tacindex{lia} \tacindex{nia} \tacindex{nra}
\label{sec:psatz-hurry}
The {\tt Psatz} module ({\tt Require Import Psatz.}) gives access to
several tactics for solving arithmetic goals over {\tt Z}, {\tt Q}, and
{\tt R}:\footnote{Support for {\tt nat} and {\tt N} is obtained by
- pre-processing the goal with the {\tt zify} tactic.}
+ pre-processing the goal with the {\tt zify} tactic.}.
+It also possible to get the tactics for integers by a {\tt Require Import Lia}, rationals {\tt Require Import Lqa}
+and reals {\tt Require Import Lra}.
\begin{itemize}
\item {\tt lia} is a decision procedure for linear integer arithmetic (see Section~\ref{sec:lia});
\item {\tt nia} is an incomplete proof procedure for integer non-linear arithmetic (see Section~\ref{sec:nia});
-\item {\tt lra} is a decision procedure for linear (real or rational) arithmetic goals (see Section~\ref{sec:lra});
+\item {\tt lra} is a decision procedure for linear (real or rational) arithmetic (see Section~\ref{sec:lra});
+\item {\tt nra} is an incomplete proof procedure for non-linear (real or rational) arithmetic (see Section~\ref{sec:nra});
\item {\tt psatz D n} where {\tt D} is {\tt Z} or {\tt Q} or {\tt R}, and
{\tt n} is an optional integer limiting the proof search depth is is an
incomplete proof procedure for non-linear arithmetic. It is based on
@@ -114,36 +117,6 @@ The deductive power of {\tt lra} is the combined deductive power of {\tt ring\_s
%
There is also an overlap with the {\tt field} tactic {\emph e.g.}, {\tt x = 10 * x / 10} is solved by {\tt lra}.
-\asection{{\tt psatz}: a proof procedure for non-linear arithmetic}
-\label{sec:psatz}
-The {\tt psatz} tactic explores the $\mathit{Cone}$ by increasing degrees -- hence the depth parameter $n$.
-In theory, such a proof search is complete -- if the goal is provable the search eventually stops.
-Unfortunately, the external oracle is using numeric (approximate) optimization techniques that might miss a
-refutation.
-
-To illustrate the working of the tactic, consider we wish to prove the following Coq goal.
-\begin{coq_eval}
-Require Import ZArith Psatz.
-Open Scope Z_scope.
-\end{coq_eval}
-\begin{coq_example*}
-Goal forall x, -x^2 >= 0 -> x - 1 >= 0 -> False.
-\end{coq_example*}
-\begin{coq_eval}
-intro x; psatz Z 2.
-\end{coq_eval}
-Such a goal is solved by {\tt intro x; psatz Z 2}. The oracle returns the
-cone expression $2 \times (\mathbf{x-1}) + (\mathbf{x-1}) \times
-(\mathbf{x-1}) + \mathbf{-x^2}$ (polynomial hypotheses are printed in
-bold). By construction, this expression belongs to $\mathit{Cone}(\{-x^2,
-x -1\})$. Moreover, by running {\tt ring} we obtain $-1$. By
-Theorem~\ref{thm:psatz}, the goal is valid.
-%
-
-%% \paragraph{The {\tt sos} tactic} -- where {\tt sos} stands for \emph{sum of squares} -- tries to prove that a
-%% single polynomial $p$ is positive by expressing it as a sum of squares \emph{i.e.,} $\sum_{i\in S} p_i^2$.
-%% This amounts to searching for $p$ in the cone without generators \emph{i.e.}, $Cone(\{\})$.
-%
\asection{{\tt lia}: a tactic for linear integer arithmetic}
\tacindex{lia}
@@ -219,22 +192,61 @@ Our current oracle tries to find an expression $e$ with a small range $[c_1,c_2]
We generate $c_2 - c_1$ subgoals which contexts are enriched with an equation $e = i$ for $i \in [c_1,c_2]$ and
recursively search for a proof.
-\asection{{\tt nia}: a proof procedure for non-linear integer arithmetic}
-\tacindex{nia}
-\label{sec:nia}
-The {\tt nia} tactic is an {\emph experimental} proof procedure for non-linear integer arithmetic.
+
+\asection{{\tt nra}: a proof procedure for non-linear arithmetic}
+\tacindex{nra}
+\label{sec:nra}
+The {\tt nra} tactic is an {\emph experimental} proof procedure for non-linear arithmetic.
%
The tactic performs a limited amount of non-linear reasoning before running the
-linear prover of {\tt lia}.
+linear prover of {\tt lra}.
This pre-processing does the following:
\begin{itemize}
\item If the context contains an arithmetic expression of the form $e[x^2]$ where $x$ is a
monomial, the context is enriched with $x^2\ge 0$;
\item For all pairs of hypotheses $e_1\ge 0$, $e_2 \ge 0$, the context is enriched with $e_1 \times e_2 \ge 0$.
\end{itemize}
-After pre-processing, the linear prover of {\tt lia} searches for a proof
+After this pre-processing, the linear prover of {\tt lra} searches for a proof
by abstracting monomials by variables.
+\asection{{\tt nia}: a proof procedure for non-linear integer arithmetic}
+\tacindex{nia}
+\label{sec:nia}
+The {\tt nia} tactic is a proof procedure for non-linear integer arithmetic.
+%
+It performs a pre-processing similar to {\tt nra}. The obtained goal is solved using the linear integer prover {\tt lia}.
+
+\asection{{\tt psatz}: a proof procedure for non-linear arithmetic}
+\label{sec:psatz}
+The {\tt psatz} tactic explores the $\mathit{Cone}$ by increasing degrees -- hence the depth parameter $n$.
+In theory, such a proof search is complete -- if the goal is provable the search eventually stops.
+Unfortunately, the external oracle is using numeric (approximate) optimization techniques that might miss a
+refutation.
+
+To illustrate the working of the tactic, consider we wish to prove the following Coq goal.
+\begin{coq_eval}
+Require Import ZArith Psatz.
+Open Scope Z_scope.
+\end{coq_eval}
+\begin{coq_example*}
+Goal forall x, -x^2 >= 0 -> x - 1 >= 0 -> False.
+\end{coq_example*}
+\begin{coq_eval}
+intro x; psatz Z 2.
+\end{coq_eval}
+Such a goal is solved by {\tt intro x; psatz Z 2}. The oracle returns the
+cone expression $2 \times (\mathbf{x-1}) + (\mathbf{x-1}) \times
+(\mathbf{x-1}) + \mathbf{-x^2}$ (polynomial hypotheses are printed in
+bold). By construction, this expression belongs to $\mathit{Cone}(\{-x^2,
+x -1\})$. Moreover, by running {\tt ring} we obtain $-1$. By
+Theorem~\ref{thm:psatz}, the goal is valid.
+%
+
+%% \paragraph{The {\tt sos} tactic} -- where {\tt sos} stands for \emph{sum of squares} -- tries to prove that a
+%% single polynomial $p$ is positive by expressing it as a sum of squares \emph{i.e.,} $\sum_{i\in S} p_i^2$.
+%% This amounts to searching for $p$ in the cone without generators \emph{i.e.}, $Cone(\{\})$.
+%
+
%%% Local Variables:
diff --git a/plugins/micromega/Lqa.v b/plugins/micromega/Lqa.v
new file mode 100644
index 000000000..0055600a0
--- /dev/null
+++ b/plugins/micromega/Lqa.v
@@ -0,0 +1,55 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+(* *)
+(* Micromega: A reflexive tactic using the Positivstellensatz *)
+(* *)
+(* Frédéric Besson (Irisa/Inria) 2016 *)
+(* *)
+(************************************************************************)
+
+Require Import QMicromega.
+Require Import QArith.
+Require Import RingMicromega.
+Require Import VarMap.
+Require Coq.micromega.Tauto.
+Declare ML Module "micromega_plugin".
+
+(** Here, lra stands for linear rational arithmetic *)
+Ltac lra := lra_Q ;
+ (abstract(intros __wit __varmap __ff ;
+ change (Tauto.eval_f (Qeval_formula (@find Q 0%Q __varmap)) __ff) ;
+ apply (QTautoChecker_sound __ff __wit); vm_cast_no_check (eq_refl true))).
+
+(** Here, nra stands for non-linear rational arithmetic *)
+Ltac nra :=
+ xnqa ;
+ (abstract(intros __wit __varmap __ff ;
+ change (Tauto.eval_f (Qeval_formula (@find Q 0%Q __varmap)) __ff) ;
+ apply (QTautoChecker_sound __ff __wit); vm_cast_no_check (eq_refl true))).
+
+Ltac xpsatz dom d :=
+ let tac := lazymatch dom with
+ | Q =>
+ (sos_Q || psatz_Q d) ;
+ (* If csdp is not installed, the previous step might not produce any
+ progress: the rest of the tactical will then fail. Hence the 'try'. *)
+ try (abstract(intros __wit __varmap __ff ;
+ change (Tauto.eval_f (Qeval_formula (@find Q 0%Q __varmap)) __ff) ;
+ apply (QTautoChecker_sound __ff __wit); vm_cast_no_check (eq_refl true)))
+ | _ => fail "Unsupported domain"
+ end in tac.
+
+Tactic Notation "psatz" constr(dom) int_or_var(n) := xpsatz dom n.
+Tactic Notation "psatz" constr(dom) := xpsatz dom ltac:(-1).
+
+
+
+
+(* Local Variables: *)
+(* coding: utf-8 *)
+(* End: *)
diff --git a/plugins/micromega/Lra.v b/plugins/micromega/Lra.v
new file mode 100644
index 000000000..7ffe1e4ed
--- /dev/null
+++ b/plugins/micromega/Lra.v
@@ -0,0 +1,56 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+(* *)
+(* Micromega: A reflexive tactic using the Positivstellensatz *)
+(* *)
+(* Frédéric Besson (Irisa/Inria) 2016 *)
+(* *)
+(************************************************************************)
+
+Require Import RMicromega.
+Require Import QMicromega.
+Require Import Rdefinitions.
+Require Import RingMicromega.
+Require Import VarMap.
+Require Coq.micromega.Tauto.
+Declare ML Module "micromega_plugin".
+
+(** Here, lra stands for linear real arithmetic *)
+Ltac lra :=
+ unfold Rdiv in * ;
+ lra_R ;
+ (abstract((intros __wit __varmap __ff ;
+ change (Tauto.eval_f (Reval_formula (@find R 0%R __varmap)) __ff) ;
+ apply (RTautoChecker_sound __ff __wit); vm_cast_no_check (eq_refl true)))).
+
+(** Here, nra stands for non-linear real arithmetic *)
+Ltac nra :=
+ xnra ;
+ (abstract((intros __wit __varmap __ff ;
+ change (Tauto.eval_f (Reval_formula (@find R 0%R __varmap)) __ff) ;
+ apply (RTautoChecker_sound __ff __wit); vm_cast_no_check (eq_refl true)))).
+
+Ltac xpsatz dom d :=
+ let tac := lazymatch dom with
+ | R =>
+ (sos_R || psatz_R d) ;
+ (* If csdp is not installed, the previous step might not produce any
+ progress: the rest of the tactical will then fail. Hence the 'try'. *)
+ try (abstract(intros __wit __varmap __ff ;
+ change (Tauto.eval_f (Reval_formula (@find R 0%R __varmap)) __ff) ;
+ apply (RTautoChecker_sound __ff __wit); vm_cast_no_check (eq_refl true)))
+ | _ => fail "Unsupported domain"
+ end in tac.
+
+Tactic Notation "psatz" constr(dom) int_or_var(n) := xpsatz dom n.
+Tactic Notation "psatz" constr(dom) := xpsatz dom ltac:(-1).
+
+
+(* Local Variables: *)
+(* coding: utf-8 *)
+(* End: *)
diff --git a/plugins/micromega/Psatz.v b/plugins/micromega/Psatz.v
index fafd8a5f2..b1f242f58 100644
--- a/plugins/micromega/Psatz.v
+++ b/plugins/micromega/Psatz.v
@@ -8,7 +8,7 @@
(* *)
(* Micromega: A reflexive tactic using the Positivstellensatz *)
(* *)
-(* Frédéric Besson (Irisa/Inria) 2006-2008 *)
+(* Frédéric Besson (Irisa/Inria) 2006-2016 *)
(* *)
(************************************************************************)
@@ -75,35 +75,39 @@ Ltac psatzl dom :=
let tac := lazymatch dom with
| Z => lia
| Q =>
- psatzl_Q ;
- (* If csdp is not installed, the previous step might not produce any
- progress: the rest of the tactical will then fail. Hence the 'try'. *)
- try (abstract(intros __wit __varmap __ff ;
- change (Tauto.eval_f (Qeval_formula (@find Q 0%Q __varmap)) __ff) ;
- apply (QTautoChecker_sound __ff __wit); vm_cast_no_check (eq_refl true)))
+ lra_Q ;
+ (abstract(intros __wit __varmap __ff ;
+ change (Tauto.eval_f (Qeval_formula (@find Q 0%Q __varmap)) __ff) ;
+ apply (QTautoChecker_sound __ff __wit); vm_cast_no_check (eq_refl true)))
| R =>
unfold Rdiv in * ;
- psatzl_R ;
- (* If csdp is not installed, the previous step might not produce any
- progress: the rest of the tactical will then fail. Hence the 'try'. *)
- try abstract((intros __wit __varmap __ff ;
- change (Tauto.eval_f (Reval_formula (@find R 0%R __varmap)) __ff) ;
- apply (RTautoChecker_sound __ff __wit); vm_cast_no_check (eq_refl true)))
- | _ => fail "Unsupported domain"
+ lra_R ;
+ (abstract((intros __wit __varmap __ff ;
+ change (Tauto.eval_f (Reval_formula (@find R 0%R __varmap)) __ff) ;
+ apply (RTautoChecker_sound __ff __wit); vm_cast_no_check (eq_refl true))))
+| _ => fail "Unsupported domain"
end in tac.
Ltac lra :=
first [ psatzl R | psatzl Q ].
-Ltac nra :=
+Ltac nra_R :=
unfold Rdiv in * ;
xnra ;
abstract
(intros __wit __varmap __ff ;
- change (Tauto.eval_f (Reval_formula (@find R 0%R __varmap)) __ff) ;
- apply (RTautoChecker_sound __ff __wit); vm_compute ; reflexivity).
-
+ change (Tauto.eval_f (Reval_formula (@find R 0%R __varmap)) __ff) ;
+ apply (RTautoChecker_sound __ff __wit); vm_compute ; reflexivity).
+
+Ltac nra_Q :=
+ xnqa ;
+ (abstract(intros __wit __varmap __ff ;
+ change (Tauto.eval_f (Qeval_formula (@find Q 0%Q __varmap)) __ff) ;
+ apply (QTautoChecker_sound __ff __wit); vm_cast_no_check (eq_refl true))).
+
+Ltac nra :=
+ first [ nra_R | nra_Q ].
(* Local Variables: *)
diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml
index b8e5e66ca..edcb00b90 100644
--- a/plugins/micromega/coq_micromega.ml
+++ b/plugins/micromega/coq_micromega.ml
@@ -2100,11 +2100,7 @@ let tauto_lia ff =
* solvers
*)
-let psatzl_Z =
- micromega_gen parse_zarith Mc.negate Mc.normalise Mc.zunsat Mc.zdeduce zz_domain_spec
- [ linear_Z ]
-
-let psatzl_Q =
+let lra_Q =
micromega_gen parse_qarith Mc.qnegate Mc.qnormalise Mc.qunsat Mc.qdeduce qq_domain_spec
[ linear_prover_Q ]
@@ -2112,7 +2108,7 @@ let psatz_Q i =
micromega_gen parse_qarith Mc.qnegate Mc.qnormalise Mc.qunsat Mc.qdeduce qq_domain_spec
[ non_linear_prover_Q "real_nonlinear_prover" (Some i) ]
-let psatzl_R =
+let lra_R =
micromega_genr [ linear_prover_R ]
let psatz_R i =
@@ -2136,21 +2132,21 @@ let sos_R =
micromega_genr [ non_linear_prover_R "pure_sos" None ]
-let xlia =
- try
- micromega_gen parse_zarith Mc.negate Mc.normalise Mc.zunsat Mc.zdeduce zz_domain_spec
+let xlia = micromega_gen parse_zarith Mc.negate Mc.normalise Mc.zunsat Mc.zdeduce zz_domain_spec
[ linear_Z ]
- with reraise -> (*Printexc.print_backtrace stdout ;*) raise reraise
let xnlia =
- try
micromega_gen parse_zarith Mc.negate Mc.normalise Mc.zunsat Mc.zdeduce zz_domain_spec
[ nlinear_Z ]
- with reraise -> (*Printexc.print_backtrace stdout ;*) raise reraise
let nra =
micromega_genr [ nlinear_prover_R ]
+let nqa =
+ micromega_gen parse_qarith Mc.qnegate Mc.qnormalise Mc.qunsat Mc.qdeduce qq_domain_spec
+ [ nlinear_prover_R ]
+
+
(* Local Variables: *)
(* coding: utf-8 *)
diff --git a/plugins/micromega/g_micromega.ml4 b/plugins/micromega/g_micromega.ml4
index e6b5cc60d..974dcee87 100644
--- a/plugins/micromega/g_micromega.ml4
+++ b/plugins/micromega/g_micromega.ml4
@@ -37,6 +37,12 @@ TACTIC EXTEND NRA
[ "xnra" ] -> [ (Coq_micromega.nra)]
END
+TACTIC EXTEND NQA
+[ "xnqa" ] -> [ (Coq_micromega.nqa)]
+END
+
+
+
TACTIC EXTEND Sos_Z
| [ "sos_Z" ] -> [ (Coq_micromega.sos_Z) ]
END
@@ -50,11 +56,11 @@ TACTIC EXTEND Sos_R
END
TACTIC EXTEND LRA_Q
-[ "psatzl_Q" ] -> [ (Coq_micromega.psatzl_Q) ]
+[ "lra_Q" ] -> [ (Coq_micromega.lra_Q) ]
END
TACTIC EXTEND LRA_R
-[ "psatzl_R" ] -> [ (Coq_micromega.psatzl_R) ]
+[ "lra_R" ] -> [ (Coq_micromega.lra_R) ]
END
TACTIC EXTEND PsatzR
diff --git a/plugins/micromega/vo.itarget b/plugins/micromega/vo.itarget
index bf6a1a7db..cb4b2b8a5 100644
--- a/plugins/micromega/vo.itarget
+++ b/plugins/micromega/vo.itarget
@@ -10,4 +10,6 @@ Tauto.vo
VarMap.vo
ZCoeff.vo
ZMicromega.vo
-Lia.vo \ No newline at end of file
+Lia.vo
+Lqa.vo
+Lra.vo \ No newline at end of file
diff --git a/test-suite/micromega/qexample.v b/test-suite/micromega/qexample.v
index 47e6005b9..d001e8f7f 100644
--- a/test-suite/micromega/qexample.v
+++ b/test-suite/micromega/qexample.v
@@ -6,32 +6,29 @@
(* *)
(************************************************************************)
-Require Import Psatz.
+Require Import Lqa.
Require Import QArith.
Lemma plus_minus : forall x y,
0 == x + y -> 0 == x -y -> 0 == x /\ 0 == y.
Proof.
intros.
- psatzl Q.
+ lra.
Qed.
-
-
-
(* Other (simple) examples *)
Open Scope Q_scope.
Lemma binomial : forall x y:Q, ((x+y)^2 == x^2 + (2 # 1) *x*y + y^2).
Proof.
intros.
- psatzl Q.
+ lra.
Qed.
Lemma hol_light19 : forall m n, (2 # 1) * m + n == (n + m) + m.
Proof.
- intros ; psatzl Q.
+ intros ; lra.
Qed.
Open Scope Z_scope.
Open Scope Q_scope.
@@ -60,7 +57,11 @@ Lemma vcgen_25 : forall
(( 1# 1) == (-2 # 1) * i + it).
Proof.
intros.
- psatzl Q.
+ lra.
+Qed.
+
+Goal forall x : Q, x * x >= 0.
+ intro; nra.
Qed.
Goal forall x, -x^2 >= 0 -> x - 1 >= 0 -> False.
diff --git a/test-suite/micromega/rexample.v b/test-suite/micromega/rexample.v
index 2eed7e951..89c08c770 100644
--- a/test-suite/micromega/rexample.v
+++ b/test-suite/micromega/rexample.v
@@ -6,7 +6,7 @@
(* *)
(************************************************************************)
-Require Import Psatz.
+Require Import Lra.
Require Import Reals.
Open Scope R_scope.
@@ -15,7 +15,7 @@ Lemma yplus_minus : forall x y,
0 = x + y -> 0 = x -y -> 0 = x /\ 0 = y.
Proof.
intros.
- psatzl R.
+ lra.
Qed.
(* Other (simple) examples *)
@@ -23,13 +23,13 @@ Qed.
Lemma binomial : forall x y, ((x+y)^2 = x^2 + 2 *x*y + y^2).
Proof.
intros.
- psatzl R.
+ lra.
Qed.
Lemma hol_light19 : forall m n, 2 * m + n = (n + m) + m.
Proof.
- intros ; psatzl R.
+ intros ; lra.
Qed.
@@ -57,7 +57,7 @@ Lemma vcgen_25 : forall
(( 1 ) = (-2 ) * i + it).
Proof.
intros.
- psatzl R.
+ lra.
Qed.
Goal forall x, -x^2 >= 0 -> x - 1 >= 0 -> False.
@@ -72,5 +72,5 @@ Proof.
Qed.
Lemma l1 : forall x y z : R, Rabs (x - z) <= Rabs (x - y) + Rabs (y - z).
-intros; split_Rabs; psatzl R.
+intros; split_Rabs; lra.
Qed. \ No newline at end of file