diff --git a/theories/psatz_axiomfree.v b/theories/psatz_axiomfree.v new file mode 100644 index 0000000000000000000000000000000000000000..9b1d430e972ab1d850cd71c15bce7ffae7415c70 --- /dev/null +++ b/theories/psatz_axiomfree.v @@ -0,0 +1,39 @@ +(** This file is a hack that lets us use Psatz without importing all sorts of + axioms about real numbers. It has been created by copying the file + Psatz.v from the Coq distribution, removing everything defined after the lia + tactic, and removing the two lines importing RMicromega and Rdefinitions. + The original license header follows. *) + +(************************************************************************) +(* 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) 2006-2008 *) +(* *) +(************************************************************************) + +From Coq Require Import ZMicromega. +From Coq Require Import QMicromega. +From Coq Require Import QArith. +From Coq Require Import ZArith. +From Coq Require Import RingMicromega. +From Coq Require Import VarMap. +From Coq Require Tauto. +Declare ML Module "micromega_plugin". + +Ltac preprocess := + zify ; unfold Z.succ in * ; unfold Z.pred in *. + +Ltac lia := + preprocess; + xlia ; + abstract ( + intros __wit __varmap __ff ; + change (Tauto.eval_f (Zeval_formula (@find Z Z0 __varmap)) __ff) ; + apply (ZTautoChecker_sound __ff __wit); vm_cast_no_check (eq_refl true)). diff --git a/theories/tactics.v b/theories/tactics.v index b1b9224e71bc1f288c761c698e9ec67005525af5..705045be0b8d9f5493343d89b0ccfd79b36f001d 100644 --- a/theories/tactics.v +++ b/theories/tactics.v @@ -3,7 +3,7 @@ (** This file collects general purpose tactics that are used throughout the development. *) From Coq Require Import Omega. -From Coq Require Export Psatz. +From stdpp Require Export psatz_axiomfree. From stdpp Require Export decidable. Lemma f_equal_dep {A B} (f g : ∀ x : A, B x) x : f = g → f x = g x.