Library Coq.micromega.Lqa


Require Import QMicromega.
Require Import QArith.
Require Import RingMicromega.
Require Import VarMap.
Require Import DeclConstant.
Require Coq.micromega.Tauto.

Ltac rchange :=
  intros ?__wit ?__varmap ?__ff ;
  change (Tauto.eval_bf (Qeval_formula (@find Q 0%Q __varmap)) __ff) ;
  apply (QTautoChecker_sound __ff __wit).

Ltac rchecker_no_abstract := rchange ; vm_compute ; reflexivity.
Ltac rchecker_abstract := rchange ; vm_cast_no_check (eq_refl true).

Ltac rchecker := rchecker_no_abstract.

Here, lra stands for linear rational arithmetic
Ltac lra := lra_Q rchecker.

Here, nra stands for non-linear rational arithmetic
Ltac nra := xnqa rchecker.

Ltac xpsatz dom d :=
  let tac := lazymatch dom with
  | Q =>
    ((sos_Q rchecker) || (psatz_Q d rchecker))
  | _ => 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).