From 721637c98514a77d05d080f53f226cab3a8da1e7 Mon Sep 17 00:00:00 2001 From: Frédéric Besson Date: Tue, 30 Aug 2016 17:12:27 +0200 Subject: 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. --- plugins/micromega/Lqa.v | 55 +++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 55 insertions(+) create mode 100644 plugins/micromega/Lqa.v (limited to 'plugins/micromega/Lqa.v') 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 *) +(* + (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: *) -- cgit v1.2.3