From 49cd8230fb5bed24738b03a9fd3aaa491e9e4442 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fr=C3=A9d=C3=A9ric=20Besson?= Date: Wed, 22 Jan 2025 15:21:57 +0100 Subject: [PATCH 1/6] Adapt to PR tify This is the companion of https://github.com/fajb/coq/tree/tify - Update the reflexive proof verifier for lia,lra etc. - Define instances for the [rify] tactic (Require Import Rify) --- test-suite/micromega/milp.v | 190 +++ test-suite/output/MExtraction.v | 5 +- test-suite/success/tify.v | 82 + theories/All/All.v | 5 + theories/Reals/RIneq.v | 3 + theories/Strings/PString.v | 4 +- theories/micromega/Lia.v | 1 + theories/micromega/Lra.v | 19 +- theories/micromega/MMicromega.v | 2381 +++++++++++++++++++++++++++ theories/micromega/RMicromega.v | 610 ++++++- theories/micromega/Refl.v | 14 + theories/micromega/Rify.v | 779 +++++++++ theories/micromega/RingMicromega.v | 34 +- theories/micromega/Tify.v | 20 + theories/micromega/TifyClasses.v | 286 ++++ theories/micromega/ZMicromega.v | 1465 ++-------------- theories/micromega/Zify.v | 11 +- theories/micromega/ZifyBool.v | 58 +- theories/micromega/ZifyClasses.v | 287 +--- theories/micromega/ZifyComparison.v | 16 +- theories/micromega/ZifyInst.v | 234 +-- theories/micromega/ZifyN.v | 14 +- theories/micromega/ZifyNat.v | 14 +- theories/micromega/ZifySint63.v | 50 +- theories/micromega/ZifyUint63.v | 64 +- theories/micromega/isZ.v | 74 + theories/omega/PreOmega.v | 2 +- 27 files changed, 4810 insertions(+), 1912 deletions(-) create mode 100644 test-suite/micromega/milp.v create mode 100644 test-suite/success/tify.v create mode 100644 theories/micromega/MMicromega.v create mode 100644 theories/micromega/Rify.v create mode 100644 theories/micromega/Tify.v create mode 100644 theories/micromega/TifyClasses.v create mode 100644 theories/micromega/isZ.v diff --git a/test-suite/micromega/milp.v b/test-suite/micromega/milp.v new file mode 100644 index 0000000000..a1772b7c67 --- /dev/null +++ b/test-suite/micromega/milp.v @@ -0,0 +1,190 @@ +Require Import Rbase Rify isZ Lra Zfloor. +Local Open Scope R_scope. + +Lemma up0 : up 0%R = 1%Z. +Proof. intros. rify. lra. Qed. + +Lemma up_succ r : up (r + 1)%R = (up r + 1)%Z. +Proof. intros. rify. lra. Qed. + +Lemma up_IZR z : up (IZR z) = (z + 1)%Z. +Proof. + rify. lra. +Qed. + +Lemma up_shiftz r z : up (r + IZR z)%R = (up r + z)%Z. +Proof. + rify. lra. +Qed. + + +Lemma up_succ2 r n : Zfloor (r + IZR n) = (Zfloor r + n)%Z. +Proof. + rify. lra. +Qed. + +Goal forall x1 x2 x3 x5 x6 + (H: -x2 > 0) + (H0: x3 >= 0) + (H1: -x1 + x5 + x6 = 0) + (H2: x1 + -x5 + x6 = 0) + (H4: -x1 > 0) + (H6: x2 + -x3 = 0) + (H7: 3 + 3*x1 = 0), + False. +Proof. + intros. + lra. +Qed. + + +Lemma spurious_isZ : forall + (x l X : R) + (delta : R) + (iZ1 : isZ x) + (H13 : - (X - l) <= 0) + (iZ2 : isZ X) + (H4 : l > 0) + (iZ2 : isZ X) + (H9 : delta / 2 <> 0) + (iZ2 : isZ X) + (H11 : 0 < delta / 2) + (iZ1 : isZ x) + (H10 : (X <= 0)), + False. +Proof. + intros. + lra. +Qed. + +Goal forall (k x:R) (P:Prop) + (Hyp : 0 <= k < 1) + (H2 : k < x < 1 /\ P) + (H3 : k < x < 1), + 0 < x. +Proof. + intros. + lra. +Qed. + + + Ltac cnf := match goal with + | |- Tauto.cnf_checker _ ?F _ = true => + let f:= fresh "f" in + set(f:=F); compute in f; unfold f;clear f + end. + + Ltac collect := + match goal with + | |- context [ RMicromega.xcollect_isZ ?A ?B ?F] => + let f := fresh "f" in + set (f:= RMicromega.xcollect_isZ A B F) ; compute in f ; unfold f; clear f + end. + + +Goal forall (x y:R) + (H : x <= y) + (Hlt : x < y), + y - x = 0 -> False. +Proof. + intros. + lra. +Qed. + +Lemma two_x_1 : forall (x:R) + (IZ : isZ x) + (H : 2 * x + 1 = 0), + False. +Proof. + intros. + lra. +Qed. + +Goal forall x1 x2 x3 x4, + isZ x1 -> + isZ x2 -> + isZ x3 -> + isZ x4 -> + - x1 + x2 - x3 + x4 >= 0 -> + x1 - x2 + x3 - x4 + 1 > 0 -> + - x1 + x2 - x3 + x4 > 0 -> + False. +Proof. + intros. lra. +Qed. + + + +Goal forall z1 z2 z3, Zfloor (IZR z1 - IZR z2 + IZR z3)%R = (z1 - z2 + z3)%Z. +Proof. + intros. + rify. + (** Now, some work over isZ *) + lra. +Qed. + +Goal forall (z : Z) x, IZR z <= x -> (z <= Zfloor x)%Z. +Proof. + rify. + lra. +Qed. + + + +Goal forall (z : Z) x, IZR z <= x < IZR z + 1 -> Zfloor x = z. +Proof. + intros. + rify. + lra. +Qed. + + +Goal R1 + (R1 + R1) = 3. +Proof. + intros. + rify. + reflexivity. +Qed. + +Goal forall x, + x * 3 = 1 -> + False. +Proof. + intros. + Fail lra. +Abort. + +Goal forall x, + isZ x -> + x * 3 = 1 -> + False. +Proof. + intros. + lra. +Qed. + + + +Require Import Rfunctions R_sqrt. + +Goal forall e x + (He : 0 < e) + (ef := Rmin (e / (3 * (Rabs (x) + 1))) (sqrt (e / 3)) : R) + (H : sqrt (e / 3) > 0) + (H0 : 3 * (Rabs ( x) + 1) >= 1), + 0 < ef. +Proof. + intros. + rify. nra. +Qed. + + +Goal forall x v, + v * (3 * x) = 1 -> + x >= 1 -> + v = 1 / (3 * x). +Proof. + intros. + rify. nra. +Qed. + diff --git a/test-suite/output/MExtraction.v b/test-suite/output/MExtraction.v index a31c993666..68e139540b 100644 --- a/test-suite/output/MExtraction.v +++ b/test-suite/output/MExtraction.v @@ -57,11 +57,12 @@ Extract Constant Rinv => "fun x -> 1 / x". extraction is only performed as a test in the test suite. *) Recursive Extraction Tauto.mapX Tauto.foldA Tauto.collect_annot Tauto.ids_of_formula Tauto.map_bformula + map_eFormula eFormula Tauto.abst_form - ZMicromega.cnfZ ZMicromega.Zeval_const QMicromega.cnfQ + ZMicromega.cnfZ ZMicromega.Zeval_const QMicromega.cnfQ RMicromega.cnfR List.map simpl_cone (*map_cone indexes*) denorm QArith_base.Qpower vm_add - normZ normQ normQ Z.to_N N.of_nat ZTautoChecker ZWeakChecker QTautoChecker RTautoChecker find. + normZ normQ (*normQ*) Z.to_N N.of_nat ZTautoChecker ZWeakChecker QTautoChecker RTautoChecker find. (* Local Variables: *) (* coding: utf-8 *) diff --git a/test-suite/success/tify.v b/test-suite/success/tify.v new file mode 100644 index 0000000000..8a2c1cd551 --- /dev/null +++ b/test-suite/success/tify.v @@ -0,0 +1,82 @@ +From Stdlib Require Import Reals. +From Stdlib Require Import ZArith. +From Stdlib Require Import ZifyClasses Tify. + +(* Inject nat -> Z *) + +Instance Inj_nat_Z : InjTyp nat Z. +Proof. + apply (mkinj _ _ Z.of_nat (fun x => (0 <= x)%Z)). + apply Nat2Z.is_nonneg. +Defined. +Add Tify InjTyp Inj_nat_Z. + +Instance Op_nat_add_Z : BinOp Nat.add. +Proof. + apply mkbop with (TBOp := Z.add). + simpl. apply Nat2Z.inj_add. +Defined. +Add Tify BinOp Op_nat_add_Z. + +Instance Op_eq_nat_eq_Z : BinRel (@eq nat). +Proof. +apply mkbrel with (TR := @eq Z). +simpl. split. congruence. +apply Nat2Z.inj. +Defined. +Add Tify BinRel Op_eq_nat_eq_Z. + + +(** Inject nat -> R *) + +Definition isZ (r:R) := IZR (Zfloor r) = r. + +Instance Inj_nat_R : InjTyp nat R. +Proof. + apply (mkinj _ _ INR (fun r => isZ r /\ 0 <= r)). + - intros. split. unfold isZ. + rewrite INR_IZR_INZ. + rewrite ZfloorZ. reflexivity. + apply pos_INR. +Defined. +Add Tify InjTyp Inj_nat_R. + +Instance Op_add_R : BinOp Nat.add. +Proof. + apply mkbop with (TBOp := Rplus). + simpl. intros. apply plus_INR. +Defined. +Add Tify BinOp Op_add_R. + +Instance Op_eq_nat_eq_R : BinRel (@eq nat). +Proof. +apply mkbrel with (TR := @eq R). +simpl. split; intro. congruence. +apply INR_eq; auto. +Defined. +Add Tify BinRel Op_eq_nat_eq_R. + + +(* tify selects the right injection *) + +Goal forall (x y:nat), (x + y = y + x)%nat. +Proof. + intros. + tify R. + apply Rplus_comm. +Qed. + +Goal forall (x y:nat), (x + y = y + x)%nat. +Proof. + intros. + tify Z. + apply Zplus_comm. +Qed. + +(* The R instances do not break lia. *) +From Stdlib Require Import Lia. +Goal forall (x y:nat), (x + y = y + x)%nat. +Proof. + intros. + lia. +Qed. diff --git a/theories/All/All.v b/theories/All/All.v index 42539b0b23..4ffdf72f30 100644 --- a/theories/All/All.v +++ b/theories/All/All.v @@ -47,6 +47,7 @@ From Stdlib Require Export micromega.Lqa. From Stdlib Require Export micromega.Lra. From Stdlib Require Export micromega.OrderedRing. From Stdlib Require Export micromega.Psatz. +From Stdlib Require Export micromega.MMicromega. From Stdlib Require Export micromega.QMicromega. From Stdlib Require Export micromega.RMicromega. From Stdlib Require Export micromega.Refl. @@ -57,8 +58,10 @@ From Stdlib Require Export micromega.ZArith_hints. From Stdlib Require Export micromega.ZCoeff. From Stdlib Require Export micromega.ZMicromega. From Stdlib Require Export micromega.Zify. +From Stdlib Require Export micromega.Tify. From Stdlib Require Export micromega.ZifyBool. From Stdlib Require Export micromega.ZifyClasses. +From Stdlib Require Export micromega.TifyClasses. From Stdlib Require Export micromega.ZifyComparison. From Stdlib Require Export micromega.ZifyInst. From Stdlib Require Export micromega.ZifyN. @@ -67,6 +70,8 @@ From Stdlib Require Export micromega.ZifyPow. From Stdlib Require Export micromega.ZifySint63. From Stdlib Require Export micromega.ZifyUint63. From Stdlib Require Export micromega.Ztac. +From Stdlib Require Export micromega.isZ. +From Stdlib Require Export micromega.Rify. From Stdlib Require Export funind.FunInd. From Stdlib Require Export funind.Recdef. From Stdlib Require Export extraction.ExtrHaskellBasic. diff --git a/theories/Reals/RIneq.v b/theories/Reals/RIneq.v index cbdee1920b..514e510beb 100644 --- a/theories/Reals/RIneq.v +++ b/theories/Reals/RIneq.v @@ -2522,6 +2522,9 @@ Qed. Lemma succ_IZR : forall n:Z, IZR (Z.succ n) = IZR n + 1. Proof. now intros n; unfold Z.succ; apply plus_IZR. Qed. +Lemma pred_IZR : forall n:Z, IZR (Z.pred n) = IZR n - 1. +Proof. now intros n; unfold Z.pred; apply plus_IZR. Qed. + Lemma opp_IZR : forall n:Z, IZR (- n) = - IZR n. Proof. intros [| p | p]; unfold IZR; simpl. diff --git a/theories/Strings/PString.v b/theories/Strings/PString.v index 73f8e3d5b3..eb61851b32 100644 --- a/theories/Strings/PString.v +++ b/theories/Strings/PString.v @@ -12,9 +12,9 @@ From Stdlib Require Import ZArith. #[local] Open Scope list_scope. #[local] Arguments to_Z _/ : simpl nomatch. -#[local] Instance Op_max_length : ZifyClasses.CstOp max_length := +#[local] Instance Op_max_length : TifyClasses.CstOp max_length := { TCst := 16777211%Z ; TCstInj := eq_refl }. -Add Zify CstOp Op_max_length. +Add Tify CstOp Op_max_length. #[local] Ltac case_if := lazymatch goal with diff --git a/theories/micromega/Lia.v b/theories/micromega/Lia.v index e268cef643..50c62cf8c4 100644 --- a/theories/micromega/Lia.v +++ b/theories/micromega/Lia.v @@ -1,3 +1,4 @@ + (************************************************************************) (* * The Rocq Prover / The Rocq Development Team *) (* v * Copyright INRIA, CNRS and contributors *) diff --git a/theories/micromega/Lra.v b/theories/micromega/Lra.v index 46f80c5d5a..d1450985db 100644 --- a/theories/micromega/Lra.v +++ b/theories/micromega/Lra.v @@ -21,6 +21,15 @@ From Stdlib Require Import RingMicromega. From Stdlib Require Import VarMap. From Stdlib.micromega Require Tauto. From Stdlib Require Import Rregisternames. +From Stdlib Require Import Tify. +(* With the exception of [Zfloor] and [sqrt], + importing Rify by default would be feasible but is unwise. + The goal may blow up because [Rmin], [Rmax] and [Rabs] + generate case splits. + (* From Stdlib Require Import Rify. *) + If Rify is imported, [lra] and [nra] are using [tify R]. + We provide [simple lra] and [simple nra] which do not call [tify R]. + *) Declare ML Module "rocq-runtime.plugins.micromega". @@ -29,7 +38,7 @@ Ltac rchange := let __varmap := fresh "__varmap" in let __ff := fresh "__ff" in intros __wit __varmap __ff ; - change (Tauto.eval_bf (Reval_formula (@find R 0%R __varmap)) __ff) ; + change (Tauto.eval_bf (eReval_formula (@find R 0%R __varmap)) __ff) ; apply (RTautoChecker_sound __ff __wit). Ltac rchecker_no_abstract := rchange ; vm_compute ; reflexivity. @@ -38,10 +47,14 @@ Ltac rchecker_abstract := rchange ; vm_cast_no_check (eq_refl true). Ltac rchecker := rchecker_no_abstract. (** Here, lra stands for linear real arithmetic *) -Ltac lra := unfold Rdiv in * ; xlra_R rchecker. +Ltac lra := unfold Rdiv in *; tify R ; xlra_R rchecker. (** Here, nra stands for non-linear real arithmetic *) -Ltac nra := unfold Rdiv in * ; xnra_R rchecker. +Ltac nra := unfold Rdiv in *;tify R ; xnra_R rchecker. + +Tactic Notation "simple" "nra" := unfold Rdiv in *; xnra_R rchecker. + +Tactic Notation "simple" "lra" := unfold Rdiv in *; xlra_R rchecker. Ltac xpsatz dom d := let tac := lazymatch dom with diff --git a/theories/micromega/MMicromega.v b/theories/micromega/MMicromega.v new file mode 100644 index 0000000000..14d2d94f6b --- /dev/null +++ b/theories/micromega/MMicromega.v @@ -0,0 +1,2381 @@ +(************************************************************************) +(* * The Rocq Prover / The Rocq Development Team *) +(* v * Copyright INRIA, CNRS and contributors *) +(* destruct (andb_prop _ _ id); clear id + | [ id : (_ || _)%bool = true |- _ ] => destruct (orb_prop _ _ id); clear id + end. + +Ltac inv H := inversion H ; try subst ; clear H. + +Section S. + Variables R : Type. + Variables rO rI : R. + Variables rplus rtimes rminus: R -> R -> R. + Variable ropp : R -> R. + Variables req rle rlt : R -> R -> Prop. + + Variable Rsor : SOR rO rI rplus rtimes rminus ropp req rle rlt. + + (** Injection from Z and N towards R. *) + + Fixpoint IPR (p:positive) : R := + match p with + | (p0~1)%positive => rplus rI (rplus (IPR p0) (IPR p0)) + | (p0~0)%positive => rplus (IPR p0) (IPR p0) + | 1%positive => rI + end. + + Definition IZR (z:Z) : R := + match z with + | Z0 => rO + | Zpos p => IPR p + | Zneg p => ropp (IPR p) + end. + + Definition INR (n:N) : R := + match n with + | N0 => rO + | Npos p => IPR p + end. + + Variable E : Type. + + Variable pow_phi : N -> E. + + Variable rpow : R -> E -> R. + + Variable RSORaddon : + SORaddon rO rI rplus rtimes rminus ropp req rle (* ring elements *) + 0%Z 1%Z Z.add Z.mul Z.sub Z.opp (* coefficients *) + Z.eqb Z.leb + IZR pow_phi rpow. + + + + Lemma Zsor : SOR 0 1 Z.add Z.mul Z.sub Z.opp (@eq Z) Z.le Z.lt. + Proof. + constructor ; intros ; subst; try reflexivity. + - apply InitialRing.Zsth. + - apply InitialRing.Zth. + - auto using Z.le_antisymm. + - eauto using Z.le_trans. + - apply Z.le_neq. + - apply Z.lt_trichotomy. + - apply Z.add_le_mono_l; assumption. + - apply Z.mul_pos_pos ; auto. + - discriminate. + Qed. + + Lemma ZSORaddon : + SORaddon 0 1 Z.add Z.mul Z.sub Z.opp (@eq Z) Z.le (* ring elements *) + 0%Z 1%Z Z.add Z.mul Z.sub Z.opp (* coefficients *) + Z.eqb Z.leb + (fun x => x) (fun x => x) (pow_N 1 Z.mul). + Proof. + constructor. + - constructor ; intros ; try reflexivity. + apply Z.eqb_eq ; auto. + - constructor. + reflexivity. + - intros x y. + rewrite <-Z.eqb_eq. congruence. + - apply Z.leb_le. + Qed. + + (* [isZ] defines what it means (for a variable) to range over the integers. *) + Definition isZ (r:R) := exists z, req (IZR z) r. + + Instance Eq_req : Equivalence req := SORsetoid Rsor. + + Add Morphism rplus with signature req ==> req ==> req as rplus_morph. + Proof. + exact (SORplus_wd Rsor). + Qed. + Add Morphism rtimes with signature req ==> req ==> req as rtimes_morph. + Proof. + exact (SORtimes_wd Rsor). + Qed. + Add Morphism ropp with signature req ==> req as ropp_morph. + Proof. + exact (SORopp_wd Rsor). + Qed. + Add Morphism rle with signature req ==> req ==> iff as rle_morph. + Proof. + exact (SORle_wd Rsor). + Qed. + Add Morphism rlt with signature req ==> req ==> iff as rlt_morph. + Proof. + exact (SORlt_wd Rsor). + Qed. + + Add Morphism rminus with signature req ==> req ==> req as rminus_morph. + Proof. + exact (rminus_morph Rsor). (* We already proved that minus is a morphism in OrderedRing.v *) + Qed. + + Definition Rbmorph (f1 f2: R -> R -> R) := + forall x1 x2 y1 y2, req x1 x2 -> req y1 y2 -> req (f1 x1 y1) (f2 x2 y2). + + + Add Morphism (@pow_pos R) with signature (Rbmorph) ==> req ==> @eq positive ==> req as pow_pos_morph. + Proof. + intros rt1 rt2 M x y EQ p. + induction p; simpl. + - unfold Rbmorph in M. + apply M; auto. + - apply M; auto. + - auto. + Qed. + + Add Ring SORRing : (SORrt Rsor). + + Add Morphism isZ with signature req ==> iff as isZ_morph. + Proof. + intros. unfold isZ. + split; intros; destruct H0 as (z & EQ); exists z; rewrite EQ; rewrite H; reflexivity. + Qed. + + Lemma isZ_rplus : forall r1 r2, + isZ r1 -> isZ r2 -> + isZ (rplus r1 r2). + Proof. + unfold isZ. + intros. + destruct H as (z1 & EQ1). + destruct H0 as (z2 & EQ2). + exists (z1 + z2). + rewrite (morph_add (SORrm RSORaddon)). + rewrite EQ1. rewrite EQ2. + reflexivity. + Qed. + + Lemma isZ_rtimes : forall r1 r2, + isZ r1 -> isZ r2 -> + isZ (rtimes r1 r2). + Proof. + unfold isZ. + intros. + destruct H as (z1 & EQ1). + destruct H0 as (z2 & EQ2). + exists (z1 * z2). + rewrite (morph_mul (SORrm RSORaddon)). + rewrite EQ1. rewrite EQ2. + reflexivity. + Qed. + + Lemma isZ_pow_pos : forall r p, + isZ r -> + isZ (pow_pos rtimes r p). + Proof. + intros. + induction p; simpl; now repeat apply isZ_rtimes. + Qed. + + Lemma isZ_ropp : forall r, + isZ r -> + isZ (ropp r). + Proof. + unfold isZ. + intros. + destruct H as (z1 & EQ1). + exists (- z1). + rewrite (morph_opp (SORrm RSORaddon)). + rewrite EQ1. + reflexivity. + Qed. + + Definition eval_expr := eval_pexpr rplus rtimes rminus ropp IZR pow_phi rpow. + + Definition eval_formula := + eval_formula rplus rtimes rminus ropp req rle rlt IZR pow_phi rpow. + + Definition Zeval_nformula := + eval_nformula Z0 Z.add Z.mul (@eq Z) Z.le Z.lt (fun x => x). + + Definition eval_nformula := + eval_nformula rO rplus rtimes req rle rlt IZR. + + Definition psub := psub Z0 Z.add Z.sub Z.opp Z.eqb. + Declare Equivalent Keys psub RingMicromega.psub. + + Definition popp := popp Z.opp. + Declare Equivalent Keys popp RingMicromega.popp. + + Definition padd := padd Z0 Z.add Z.eqb. + Declare Equivalent Keys padd RingMicromega.padd. + + Definition pmul := pmul 0 1 Z.add Z.mul Z.eqb. + Declare Equivalent Keys pmul RingMicromega.pmul. + + Definition norm := norm 0 1 Z.add Z.mul Z.sub Z.opp Z.eqb. + Declare Equivalent Keys norm RingMicromega.norm. + + Definition Zeval_pol := eval_pol Z.add Z.mul (fun x => x). + + Definition eval_pol := eval_pol rplus rtimes IZR. + Declare Equivalent Keys eval_pol RingMicromega.eval_pol. + + Lemma Zeval_pol_Pc : forall env z, (Zeval_pol env (Pc z)) = z. + Proof. intros. reflexivity. Qed. + + + Lemma Zeval_pol_add : forall env lhs rhs, (Zeval_pol env (padd lhs rhs)) = (Zeval_pol env lhs) + (Zeval_pol env rhs). + Proof. intros. apply (eval_pol_add Zsor ZSORaddon). Qed. + + Lemma Zeval_pol_sub : forall env lhs rhs, (Zeval_pol env (psub lhs rhs)) = (Zeval_pol env lhs) - (Zeval_pol env rhs). + Proof. intros. apply (eval_pol_sub Zsor ZSORaddon). Qed. + + Lemma eval_pol_add : forall env lhs rhs, req (eval_pol env (padd lhs rhs)) (rplus (eval_pol env lhs) (eval_pol env rhs)). + Proof. intros. apply (eval_pol_add Rsor RSORaddon). Qed. + + Lemma eval_pol_sub : forall env lhs rhs, req (eval_pol env (psub lhs rhs)) (rminus (eval_pol env lhs) (eval_pol env rhs)). + Proof. intros. unfold eval_pol. eapply (eval_pol_sub Rsor RSORaddon). Qed. + + Lemma eval_pol_mul : forall env lhs rhs, req (eval_pol env (pmul lhs rhs)) (rtimes (eval_pol env lhs) (eval_pol env rhs)). + Proof. intros. apply (eval_pol_mul Rsor RSORaddon). Qed. + + + Lemma eval_pol_norm : forall env e, req (eval_expr env e) (eval_pol env (norm e)). + Proof. intros. apply (eval_pol_norm Rsor RSORaddon). Qed. + + Definition Zunsat := check_inconsistent 0 Z.eqb Z.leb. + + Definition Zdeduce := nformula_plus_nformula 0 Z.add Z.eqb. + + Lemma Zunsat_sound : forall f, + Zunsat f = true -> forall env, eval_nformula env f -> False. + Proof. + unfold Zunsat. + intros f H env ?. + destruct f. + eapply check_inconsistent_sound with (1 := Rsor) (2 := RSORaddon) in H; eauto. + Qed. + + Lemma Zdeduce_sound : forall f1 f2 f, + Zdeduce f1 f2 = Some f -> forall env, eval_nformula env f1 -> eval_nformula env f2 -> eval_nformula env f. + Proof. + unfold Zdeduce. + intros. + eapply nformula_plus_nformula_correct in H ;eauto. + Qed. + + Definition xnnormalise (t : Formula Z) : NFormula Z := + let (lhs,o,rhs) := t in + let lhs := norm lhs in + let rhs := norm rhs in + match o with + | OpEq => (psub rhs lhs, Equal) + | OpNEq => (psub rhs lhs, NonEqual) + | OpGt => (psub lhs rhs, Strict) + | OpLt => (psub rhs lhs, Strict) + | OpGe => (psub lhs rhs, NonStrict) + | OpLe => (psub rhs lhs, NonStrict) + end. + + Lemma xnnormalise_correct : + forall env f, + eval_nformula env (xnnormalise f) <-> eval_formula env f. + Proof. + intros env f. + unfold xnnormalise. + destruct f as [lhs o rhs]. + destruct o eqn:O ; cbn ; rewrite ?eval_pol_sub; + rewrite <- !eval_pol_norm ; simpl in *; + unfold eval_expr; + generalize (eval_pexpr rplus rtimes rminus ropp IZR pow_phi rpow env lhs) as z; + generalize (eval_pexpr rplus rtimes rminus ropp IZR pow_phi rpow env rhs) as z1. + - intros. + rewrite (Rminus_eq_0 Rsor). + split ; intro H ; rewrite H; reflexivity. + - intros. + rewrite (Rminus_eq_0 Rsor). + split ; intro H ; intro H1; apply H; rewrite H1; reflexivity. + - intros. + symmetry. apply (Rle_le_minus Rsor). + - intros. symmetry. apply (Rle_le_minus Rsor). + - intros. symmetry. apply (Rlt_lt_minus Rsor). + - intros. symmetry. apply (Rlt_lt_minus Rsor). + Qed. + + Fixpoint vars (jmp : positive) (p : Pol Z) : list positive := + match p with + | Pc c => nil + | Pinj j p => vars (Pos.add j jmp) p + | PX p j q => jmp::(vars jmp p)++vars (Pos.succ jmp) q + end. + + + Lemma eval_pol_jump : forall p j env pol, + eval_pol (Env.jump p (Env.jump j env)) pol = + eval_pol (Env.jump j (Env.jump p env)) pol. + Proof. + intros. + apply env_morph. + unfold Env.jump. + intros. + f_equal. rewrite <- !Pos.add_assoc. + f_equal. apply Pos.add_comm. + Qed. + + + Lemma In_vars_add : forall pol (j x p:positive), + In x (vars j pol) -> + In (x + p)%positive (vars (p + j)%positive pol). + Proof. + induction pol. + - simpl. auto. + - simpl. intros. + apply IHpol with (p:= p0) in H. + replace (p + (p0 + j))%positive with (p0 + (p + j))%positive. + auto. + rewrite Pos.add_comm. + rewrite <- Pos.add_assoc. + f_equal. rewrite Pos.add_comm. reflexivity. + - simpl. intros. + rewrite in_app_iff in *. + destruct H as [H | [ H | H]]. + + subst. + rewrite Pos.add_comm. tauto. + + right. left. + apply IHpol1; auto. + + right. right. + rewrite <- Pos.add_succ_r. + apply IHpol2. + auto. + Qed. + + + Lemma has_var_eval : forall pol env env', + (forall x, In x (vars xH pol) -> env x = env' x) -> + eval_pol env pol = eval_pol env' pol. + Proof. + induction pol. + - simpl. reflexivity. + - simpl. + intros. + apply IHpol. + intros. + unfold Env.jump. + apply H. + apply In_vars_add; auto. + - simpl. + intros. + f_equal. + f_equal. + + eapply IHpol1. + intros. + apply H. rewrite! in_app_iff. tauto. + + unfold Env.hd,Env.jump,Env.nth. + rewrite H. reflexivity. + rewrite in_app_iff. tauto. + + + unfold Env.tail. + apply IHpol2. + intros. + unfold Env.jump. + apply H. + rewrite in_app_iff. + right. right. + change 2%positive with (1 + 1)%positive. + apply In_vars_add; auto. + Qed. + + + Fixpoint xis_integral (e : positive -> bool) (jmp:positive) (pol : Pol Z) : bool := + match pol with + | Pc _ => true + | Pinj j p0 => xis_integral e (j + jmp) p0 + | PX p0 j q => xis_integral e jmp p0 && e jmp && xis_integral e (Pos.succ jmp) q + end. + + Definition get (s: PositiveSet.t) := + fun x => PositiveSet.mem x s. + + Definition oget (s:option PositiveSet.t) (x:positive) := + match s with + | None => true + | Some s => PositiveSet.mem x s + end. + + Definition is_integral (e: option PositiveSet.t) (pol : Pol Z) := + match e with + | None => true + | Some e => xis_integral (get e) xH pol + end. + + Lemma xis_integral_jump : forall pol eb p j, + xis_integral eb (p + j) pol = true -> + xis_integral (Env.jump p eb) j pol = true. + Proof. + induction pol. + - simpl. reflexivity. + - simpl. intros. + apply IHpol. + replace (p0 + (p + j))%positive with (p + (p0 + j))%positive. + auto. + symmetry. + rewrite (Pos.add_comm p0). + rewrite <- Pos.add_assoc. + f_equal. rewrite Pos.add_comm. reflexivity. + - simpl. intros. + rewrite !andb_true_iff in *. + destruct H as ((H1 & H2) & H3). + repeat split. + + eapply IHpol1. auto. + + unfold Env.jump. rewrite Pos.add_comm. auto. + + eapply IHpol2;auto. + rewrite <- Pos.add_succ_r in H3. auto. + Qed. + + Lemma xis_integral_sound : forall pol eb j, + xis_integral eb j pol = true -> + (forall x, In x (vars j pol) -> eb x = true). + Proof. + induction pol. + - simpl. tauto. + - intros. + simpl in H,H0. + eapply IHpol;eauto. + - intros. + simpl in H,H0. + rewrite !andb_true_iff in *. + rewrite !in_app_iff in *. + destruct H as ((H1 & H2) & H3). + destruct H0 as [H0 | [H0 | H0]]. + + congruence. + + eapply IHpol1;eauto. + + eapply IHpol2;eauto. + Qed. + + Lemma is_integral_sound : forall pol eb, + is_integral eb pol = true -> + (forall x, In x (vars 1 pol) -> oget eb x = true). + Proof. + unfold is_integral. + destruct eb. + - simpl. + intros. + apply xis_integral_sound with (x:= x) in H; auto. + - reflexivity. + Qed. + + Local Open Scope Z_scope. + + Definition ceiling (a b:Z) : Z := + let (q,r) := Z.div_eucl a b in + match r with + | Z0 => q + | _ => q + 1 + end. + + Lemma Zdivide_ceiling : forall a b, (b | a) -> ceiling a b = Z.div a b. + Proof. + unfold ceiling. + intros a b H. + apply Zdivide_mod in H. + case_eq (Z.div_eucl a b). + intros z z0 H0. + change z with (fst (z,z0)). + rewrite <- H0. + change (fst (Z.div_eucl a b)) with (Z.div a b). + change z0 with (snd (z,z0)). + rewrite <- H0. + change (snd (Z.div_eucl a b)) with (Z.modulo a b). + rewrite H. + reflexivity. + Qed. + + Lemma narrow_interval_lower_bound a b x : + a > 0 -> a * x >= b -> x >= ceiling b a. + Proof. + rewrite !Z.ge_le_iff. + unfold ceiling. + intros Ha H. + generalize (Z_div_mod b a Ha). + destruct (Z.div_eucl b a) as (q,r). intros (->,(H1,H2)). + destruct r as [|r|r]. + - rewrite Z.add_0_r in H. + apply Z.mul_le_mono_pos_l in H; auto with zarith. + - assert (0 < Z.pos r) by easy. + rewrite Z.add_1_r, Z.le_succ_l. + apply Z.mul_lt_mono_pos_l with a. + + auto using Z.gt_lt. + + eapply Z.lt_le_trans. 2: eassumption. + now apply Z.lt_add_pos_r. + - now elim H1. + Qed. + (** NB: narrow_interval_upper_bound is Zdiv.Zdiv_le_lower_bound *) + + Definition ZWitness := Psatz Z. + + Inductive ZArithProof := + | DoneProof + | RatProof : ZWitness -> ZArithProof -> ZArithProof + | CutProof : ZWitness -> ZArithProof -> ZArithProof + | SplitProof : PolC Z -> ZArithProof -> ZArithProof -> ZArithProof + | EnumProof : ZWitness -> ZWitness -> list ZArithProof -> ZArithProof + | ExProof : positive -> ZArithProof -> ZArithProof + (*ExProof x : exists z t, x = z - t /\ z >= 0 /\ t >= 0 *) + . + + + Register ZArithProof as micromega.ZArithProof.type. + Register DoneProof as micromega.ZArithProof.DoneProof. + Register RatProof as micromega.ZArithProof.RatProof. + Register CutProof as micromega.ZArithProof.CutProof. + Register SplitProof as micromega.ZArithProof.SplitProof. + Register EnumProof as micromega.ZArithProof.EnumProof. + Register ExProof as micromega.ZArithProof.ExProof. + + + (* In order to compute the 'cut', we need to express a polynomial P as a * Q + b. + - b is the constant + - a is the gcd of the other coefficient. + *) + + Definition isZ0 (x:Z) := + match x with + | Z0 => true + | _ => false + end. + + Lemma isZ0_0 : forall x, isZ0 x = true <-> x = 0. + Proof. + intros x; destruct x ; simpl ; intuition congruence. + Qed. + + Lemma isZ0_n0 : forall x, isZ0 x = false <-> x <> 0. + Proof. + intros x; destruct x ; simpl ; intuition congruence. + Qed. + + Definition ZgcdM (x y : Z) := Z.max (Z.gcd x y) 1. + + + Fixpoint Zgcd_pol (p : PolC Z) : (Z * Z) := + match p with + | Pc c => (0,c) + | Pinj _ p => Zgcd_pol p + | PX p _ q => + let (g1,c1) := Zgcd_pol p in + let (g2,c2) := Zgcd_pol q in + (ZgcdM (ZgcdM g1 c1) g2 , c2) + end. + + Fixpoint Zdiv_pol (p:PolC Z) (x:Z) : PolC Z := + match p with + | Pc c => Pc (Z.div c x) + | Pinj j p => Pinj j (Zdiv_pol p x) + | PX p j q => PX (Zdiv_pol p x) j (Zdiv_pol q x) + end. + + Inductive Zdivide_pol (x:Z): PolC Z -> Prop := + | Zdiv_Pc : forall c, (x | c) -> Zdivide_pol x (Pc c) + | Zdiv_Pinj : forall p, Zdivide_pol x p -> forall j, Zdivide_pol x (Pinj j p) + | Zdiv_PX : forall p q, Zdivide_pol x p -> Zdivide_pol x q -> forall j, Zdivide_pol x (PX p j q). + + + Lemma Zdiv_pol_correct : forall a p, 0 < a -> Zdivide_pol a p -> + forall env, Zeval_pol env p = a * Zeval_pol env (Zdiv_pol p a). + Proof. + intros a p H H0. + induction H0 as [? ?|? ? IHZdivide_pol j|? ? ? IHZdivide_pol1 ? IHZdivide_pol2 j]. + - (* Pc *) + simpl. + intros. + apply Zdivide_Zdiv_eq in H0 ; auto. + - (* Pinj *) + simpl. + intros. + apply IHZdivide_pol. + - (* PX *) + simpl. + intros. + rewrite IHZdivide_pol1. + rewrite IHZdivide_pol2. + ring. + Qed. + + Lemma Zgcd_pol_ge : forall p, fst (Zgcd_pol p) >= 0. + Proof. + intros p; induction p as [c|p p1 IHp1|p1 IHp1 ? p3 IHp3]. 1-2: easy. + simpl. + case_eq (Zgcd_pol p1). + case_eq (Zgcd_pol p3). + intros. + simpl. + unfold ZgcdM. + apply Z.le_ge; transitivity 1. + - easy. + - apply Z.le_max_r. + Qed. + + Lemma Zdivide_pol_Zdivide : forall p x y, Zdivide_pol x p -> (y | x) -> Zdivide_pol y p. + Proof. + intros p x y H H0. + induction H. + - constructor. + apply Z.divide_trans with (1:= H0) ; assumption. + - constructor. auto. + - constructor ; auto. + Qed. + + Lemma Zdivide_pol_one : forall p, Zdivide_pol 1 p. + Proof. + intros p; induction p as [c| |]; constructor ; auto. + exists c. ring. + Qed. + + Lemma Zgcd_minus : forall a b c, (a | c - b ) -> (Z.gcd a b | c). + Proof. + intros a b c (q,Hq). + destruct (Zgcd_is_gcd a b) as [(a',Ha) (b',Hb) _]. + set (g:=Z.gcd a b) in *; clearbody g. + exists (q * a' + b'). + symmetry in Hq. rewrite <- Z.add_move_r in Hq. + rewrite <- Hq, Hb, Ha. ring. + Qed. + + Lemma Zdivide_pol_sub : forall p a b, + 0 < Z.gcd a b -> + Zdivide_pol a (PsubC Z.sub p b) -> + Zdivide_pol (Z.gcd a b) p. + Proof. + intros p; induction p as [c|? p IHp|p ? ? ? IHp2]. + - simpl. + intros a b H H0. inversion H0. + constructor. + apply Zgcd_minus ; auto. + - intros ? ? H H0. + constructor. + simpl in H0. inversion H0 ; subst; clear H0. + apply IHp ; auto. + - simpl. intros a b H H0. + inv H0. + constructor. + + apply Zdivide_pol_Zdivide with (1:= (ltac:(assumption) : Zdivide_pol a p)). + destruct (Zgcd_is_gcd a b) ; assumption. + + apply IHp2 ; assumption. + Qed. + + Lemma Zdivide_pol_sub_0 : forall p a, + Zdivide_pol a (PsubC Z.sub p 0) -> + Zdivide_pol a p. + Proof. + intros p; induction p as [c|? p IHp|? IHp1 ? ? IHp2]. + - simpl. + intros ? H. inversion H. + constructor. rewrite Z.sub_0_r in *. assumption. + - intros ? H. + constructor. + simpl in H. inversion H ; subst; clear H. + apply IHp ; auto. + - simpl. intros ? H. + inv H. + constructor. + + auto. + + apply IHp2 ; assumption. + Qed. + + + Lemma Zgcd_pol_div : forall p g c, + Zgcd_pol p = (g, c) -> Zdivide_pol g (PsubC Z.sub p c). + Proof. + intros p; induction p as [c|? ? IHp|p1 IHp1 ? p3 IHp2]; simpl. + - (* Pc *) + intros ? ? H. inv H. + constructor. + exists 0. now ring. + - (* Pinj *) + intros. + constructor. apply IHp ; auto. + - (* PX *) + intros g c. + case_eq (Zgcd_pol p1) ; case_eq (Zgcd_pol p3) ; intros z z0 H z1 z2 H0 H1. + inv H1. + unfold ZgcdM at 1. + destruct (Z.max_spec (Z.gcd (ZgcdM z1 z2) z) 1) as [HH1 | HH1]; cycle 1; + destruct HH1 as [HH1 HH1'] ; rewrite HH1'. + + constructor. + * apply (Zdivide_pol_Zdivide _ (ZgcdM z1 z2)). + -- unfold ZgcdM. + destruct (Z.max_spec (Z.gcd z1 z2) 1) as [HH2 | HH2]; cycle 1. + ++ destruct HH2 as [H1 H2]. + rewrite H2. + apply Zdivide_pol_sub ; auto. + apply Z.lt_le_trans with 1. + ** reflexivity. + ** trivial. + ++ destruct HH2 as [H1 H2]. rewrite H2. + apply Zdivide_pol_one. + -- unfold ZgcdM in HH1. unfold ZgcdM. + destruct (Z.max_spec (Z.gcd z1 z2) 1) as [HH2 | HH2]; cycle 1. + ++ destruct HH2 as [H1 H2]. rewrite H2 in *. + destruct (Zgcd_is_gcd (Z.gcd z1 z2) z); auto. + ++ destruct HH2 as [H1 H2]. rewrite H2. + destruct (Zgcd_is_gcd 1 z); auto. + * apply (Zdivide_pol_Zdivide _ z). + -- apply (IHp2 _ _ H); auto. + -- destruct (Zgcd_is_gcd (ZgcdM z1 z2) z); auto. + + constructor. + * apply Zdivide_pol_one. + * apply Zdivide_pol_one. + Qed. + + Lemma Zgcd_pol_correct_lt : forall p env g c, + Zgcd_pol p = (g,c) -> 0 < g -> + Zeval_pol env p = g * Zeval_pol env (Zdiv_pol (PsubC Z.sub p c) g) + c. + Proof. + intros. + rewrite <- Zdiv_pol_correct ; auto. + - rewrite (RingMicromega.PsubC_ok Zsor ZSORaddon). + unfold Zeval_pol. ring. + - apply Zgcd_pol_div ; auto. + Qed. + + Definition phi_env (S : positive -> Prop) + (env: positive -> Z) + (env': positive -> R) := + forall x, S x -> req (IZR (env x)) (env' x). + + Definition mem (l:list positive) (x:positive) := In x l. + + Lemma IZR_pow_pos : forall x p, + req (IZR (pow_pos Z.mul x p)) (pow_pos rtimes (IZR x) p). + Proof. + induction p; simpl. + - rewrite! (morph_mul (SORrm RSORaddon)). + apply rtimes_morph. + reflexivity. + apply rtimes_morph; auto. + - rewrite! (morph_mul (SORrm RSORaddon)). + apply rtimes_morph;auto. + - reflexivity. + Qed. + + + Lemma eval_Zpol_phi : forall pol env env', + phi_env (mem (vars xH pol)) env env' -> + req (IZR (Zeval_pol env pol)) (eval_pol env' pol). + Proof. + unfold phi_env,mem. + induction pol;intros env env' PHI. + - simpl. reflexivity. + - simpl. + apply IHpol. + intros x IN. + apply PHI. + now apply In_vars_add. + - simpl; simpl in PHI. + rewrite (morph_add (SORrm RSORaddon)). + rewrite (morph_mul (SORrm RSORaddon)). + apply rplus_morph. + apply rtimes_morph. + apply IHpol1. + intros x IN. apply PHI. + rewrite !in_app_iff. tauto. + unfold Env.hd,Env.nth. + rewrite IZR_pow_pos. + apply pow_pos_morph. unfold Rbmorph. intros. + rewrite H;rewrite H0;reflexivity. + apply PHI. tauto. reflexivity. + apply IHpol2. + intros. apply PHI. + rewrite !in_app_iff. + right. right. + change 2%positive with (1 + 1)%positive. + now apply In_vars_add. + Qed. + + Definition makeCuttingPlane (p : PolC Z) : PolC Z * Z := + let (g,c) := Zgcd_pol p in + if Z.gtb g Z0 + then (Zdiv_pol (PsubC Z.sub p c) g , Z.opp (ceiling (Z.opp c) g)) + else (p,Z0). + + + Definition genCuttingPlane (f : NFormula Z) : option (PolC Z * Z * Op1) := + let (e,op) := f in + match op with + | Equal => let (g,c) := Zgcd_pol e in + if andb (Z.gtb g Z0) (andb (negb (Z.eqb c Z0)) (negb (Z.eqb (Z.gcd g c) g))) + then None (* inconsistent *) + else (* Could be optimised Zgcd_pol is recomputed *) + let (p,c) := makeCuttingPlane e in + Some (p,c,Equal) + | NonEqual => Some (e,Z0,op) + | Strict => let (p,c) := makeCuttingPlane (PsubC Z.sub e 1) in + Some (p,c,NonStrict) + | NonStrict => let (p,c) := makeCuttingPlane e in + Some (p,c,NonStrict) + end. + + Definition nformula_of_cutting_plane (t : PolC Z * Z * Op1) : NFormula Z := + let (e_z, o) := t in + let (e,z) := e_z in + (padd e (Pc z) , o). + + Definition is_pol_Z0 (p : PolC Z) : bool := + match p with + | Pc Z0 => true + | _ => false + end. + + Lemma is_pol_Z0_eval_pol : forall p, is_pol_Z0 p = true -> forall env, req (eval_pol env p) rO. + Proof. + unfold is_pol_Z0. + intros p; destruct p as [z| |]; try discriminate. + destruct z ; try discriminate. + simpl. intros. apply (morph0 (SORrm RSORaddon)). + Qed. + + + Lemma is_pol_Z0_Zeval_pol : forall p, is_pol_Z0 p = true -> forall env, Zeval_pol env p = 0. + Proof. + unfold is_pol_Z0. + intros p; destruct p as [z| |]; try discriminate. + destruct z ; try discriminate. + simpl. reflexivity. + Qed. + + Definition eval_Psatz : list (NFormula Z) -> ZWitness -> option (NFormula Z) := + eval_Psatz 0 1 Z.add Z.mul Z.eqb Z.leb. + + + Definition valid_cut_sign (op:Op1) := + match op with + | Equal => true + | NonStrict => true + | Strict => true + | _ => false + end. + + Definition bound_var (v : positive) : Formula Z := + Build_Formula (PEX v) OpGe (PEc 0). + + Definition mk_eq_pos (x : positive) (y:positive) (t : positive) : Formula Z := + Build_Formula (PEX x) OpEq (PEsub (PEX y) (PEX t)). + + Fixpoint max_var (jmp : positive) (p : Pol Z) : positive := + match p with + | Pc _ => jmp + | Pinj j p => max_var (Pos.add j jmp) p + | PX p j q => Pos.max (max_var jmp p) (max_var (Pos.succ jmp) q) + end. + + Lemma pos_le_add : forall y x, + (x <= y + x)%positive. + Proof. + intros y x. + assert ((Z.pos x) <= Z.pos (x + y))%Z as H. + - rewrite <- (Z.add_0_r (Zpos x)). + rewrite <- Pos2Z.add_pos_pos. + apply Z.add_le_mono_l. + compute. congruence. + - rewrite Pos.add_comm in H. + apply H. + Qed. + + + Lemma max_var_le : forall p v, + (v <= max_var v p)%positive. + Proof. + intros p; induction p as [?|p ? IHp|? IHp1 ? ? IHp2]; simpl. + - intros. + apply Pos.le_refl. + - intros v. + specialize (IHp (p+v)%positive). + eapply Pos.le_trans ; eauto. + assert (xH + v <= p + v)%positive. + { apply Pos.add_le_mono. + - apply Pos.le_1_l. + - apply Pos.le_refl. + } + eapply Pos.le_trans ; eauto. + apply pos_le_add. + - intros v. + apply Pos.max_case_strong;intros ; auto. + specialize (IHp2 (Pos.succ v)%positive). + eapply Pos.le_trans ; eauto. + Qed. + + Lemma max_var_correct : forall p j v, + In v (vars j p) -> Pos.le v (max_var j p). + Proof. + intros p; induction p; simpl. + - tauto. + - auto. + - intros j v H. + rewrite in_app_iff in H. + destruct H as [H |[ H | H]]. + + subst. + apply Pos.max_case_strong;intros ; auto. + * apply max_var_le. + * eapply Pos.le_trans ; eauto. + apply max_var_le. + + apply Pos.max_case_strong;intros ; auto. + eapply Pos.le_trans ; eauto. + + apply Pos.max_case_strong;intros ; auto. + eapply Pos.le_trans ; eauto. + Qed. + + Definition max_var_nformulae (l : list (NFormula Z)) := + List.fold_left (fun acc f => Pos.max acc (max_var xH (fst f))) l xH. + + Section MaxVar. + + Definition F (acc : positive) (f : Pol Z * Op1) := Pos.max acc (max_var 1 (fst f)). + + Lemma max_var_nformulae_mono_aux : + forall l v acc, + (v <= acc -> + v <= fold_left F l acc)%positive. + Proof. + intros l; induction l as [|a l IHl] ; simpl ; [easy|]. + intros. + apply IHl. + unfold F. + apply Pos.max_case_strong;intros ; auto. + eapply Pos.le_trans ; eauto. + Qed. + + Lemma max_var_nformulae_mono_aux' : + forall l acc acc', + (acc <= acc' -> + fold_left F l acc <= fold_left F l acc')%positive. + Proof. + intros l; induction l as [|a l IHl]; simpl ; [easy|]. + intros. + apply IHl. + unfold F. + apply Pos.max_le_compat_r; auto. + Qed. + + Lemma max_var_nformulae_correct_aux : forall l p o v, + In (p,o) l -> In v (vars xH p) -> Pos.le v (fold_left F l 1)%positive. + Proof. + intros l p o v H H0. + generalize 1%positive as acc. + revert p o v H H0. + induction l as [|a l IHl]. + - simpl. tauto. + - simpl. + intros p o v H H0 ?. + destruct H ; subst. + + unfold F at 2. + simpl. + apply max_var_correct in H0. + apply max_var_nformulae_mono_aux. + apply Pos.max_case_strong;intros ; auto. + eapply Pos.le_trans ; eauto. + + eapply IHl ; eauto. + Qed. + + End MaxVar. + + Lemma max_var_nformalae_correct : forall l p o v, + In (p,o) l -> In v (vars xH p) -> Pos.le v (max_var_nformulae l)%positive. + Proof. + intros l p o v. + apply max_var_nformulae_correct_aux. + Qed. + + Fixpoint max_var_psatz (w : Psatz Z) : positive := + match w with + | PsatzIn _ n => xH + | PsatzSquare p => max_var xH (Psquare 0 1 Z.add Z.mul Z.eqb p) + | PsatzMulC p w => Pos.max (max_var xH p) (max_var_psatz w) + | PsatzMulE w1 w2 => Pos.max (max_var_psatz w1) (max_var_psatz w2) + | PsatzAdd w1 w2 => Pos.max (max_var_psatz w1) (max_var_psatz w2) + | _ => xH + end. + + Fixpoint max_var_prf (w : ZArithProof) : positive := + match w with + | DoneProof => xH + | RatProof w pf | CutProof w pf => Pos.max (max_var_psatz w) (max_var_prf pf) + | SplitProof p pf1 pf2 => Pos.max (max_var xH p) (Pos.max (max_var_prf pf1) (max_var_prf pf1)) + | EnumProof w1 w2 l => List.fold_left + (fun acc prf => Pos.max acc (max_var_prf prf)) l + (Pos.max (max_var_psatz w1) (max_var_psatz w2)) + | ExProof _ pf => max_var_prf pf + end. + + Definition add (x:positive) (s : option PositiveSet.t) := + match s with + | None => None + | Some s => Some (PositiveSet.add x s) + end. + + + Fixpoint ZChecker (is_integer : option (PositiveSet.t)) (l:list (NFormula Z)) (pf : ZArithProof) {struct pf} : bool := + match pf with + | DoneProof => false + | RatProof w pf => + match eval_Psatz l w with + | None => false + | Some f => + if Zunsat f then true + else ZChecker is_integer (f::l) pf + end + | CutProof w pf => + match eval_Psatz l w with + | None => false + | Some f => + if is_integral is_integer (fst f) + then + match genCuttingPlane f with + | None => true + | Some cp => + let f := nformula_of_cutting_plane cp in + if Zunsat f then true + else ZChecker is_integer (f::l) pf + end else false + end + | SplitProof p pf1 pf2 => + if is_integral is_integer p + then + match genCuttingPlane (p,NonStrict) , genCuttingPlane (popp p, NonStrict) with + | None , _ | _ , None => false + | Some cp1 , Some cp2 => + ZChecker is_integer (nformula_of_cutting_plane cp1::l) pf1 + && + ZChecker is_integer (nformula_of_cutting_plane cp2::l) pf2 + end + else false + | ExProof x prf => + let fr := max_var_nformulae l in + if Pos.leb x fr && oget is_integer x then + let z := Pos.succ fr in + let t := Pos.succ z in + let nfx := xnnormalise (mk_eq_pos x z t) in + let posz := xnnormalise (bound_var z) in + let post := xnnormalise (bound_var t) in + ZChecker (add z (add t is_integer)) (nfx::posz::post::l) prf + else false + | EnumProof w1 w2 pf => + match eval_Psatz l w1 , eval_Psatz l w2 with + | Some f1 , Some f2 => + if is_integral is_integer (fst f1) && is_integral is_integer (fst f2) + then + match genCuttingPlane f1 , genCuttingPlane f2 with + |Some (e1,z1,op1) , Some (e2,z2,op2) => + if (valid_cut_sign op1 && valid_cut_sign op2 && is_pol_Z0 (padd e1 e2)) + then + (fix label (pfs:list ZArithProof) := + fun lb ub => + match pfs with + | nil => if Z.gtb lb ub then true else false + | pf::rsr => andb (ZChecker is_integer ((psub e1 (Pc lb), Equal) :: l) pf) (label rsr (Z.add lb 1%Z) ub) + end) pf (Z.opp z1) z2 + else false + | _ , _ => true + end + else false + | _ , _ => false + end + end. + + + + Fixpoint bdepth (pf : ZArithProof) : nat := + match pf with + | DoneProof => O + | RatProof _ p => S (bdepth p) + | CutProof _ p => S (bdepth p) + | SplitProof _ p1 p2 => S (Nat.max (bdepth p1) (bdepth p2)) + | EnumProof _ _ l => S (List.fold_right (fun pf x => Nat.max (bdepth pf) x) O l) + | ExProof _ p => S (bdepth p) + end. + + Lemma in_bdepth : forall l a b y, In y l -> ltof ZArithProof bdepth y (EnumProof a b l). + Proof. + intros l; induction l as [|a l IHl]. + - (* nil *) + simpl. + tauto. + - (* cons *) + simpl. + intros a0 b y H. + destruct H as [H|H]. + + subst. + unfold ltof. + simpl. + generalize ( (fold_right + (fun (pf : ZArithProof) (x : nat) => Nat.max (bdepth pf) x) 0%nat l)). + intros. + generalize (bdepth y) ; intros. + rewrite Nat.lt_succ_r. apply Nat.le_max_l. + + generalize (IHl a0 b y H). + unfold ltof. + simpl. + generalize ( (fold_right (fun (pf : ZArithProof) (x : nat) => Nat.max (bdepth pf) x) 0%nat + l)). + intros. + eapply Nat.lt_le_trans. + * eassumption. + * rewrite <- Nat.succ_le_mono. + apply Nat.le_max_r. + Qed. + + Lemma ltof_bdepth_split_l : + forall p pf1 pf2, + ltof ZArithProof bdepth pf1 (SplitProof p pf1 pf2). + Proof. + intros. + unfold ltof. simpl. + rewrite Nat.lt_succ_r. + apply Nat.le_max_l. + Qed. + + Lemma ltof_bdepth_split_r : + forall p pf1 pf2, + ltof ZArithProof bdepth pf2 (SplitProof p pf1 pf2). + Proof. + intros. + unfold ltof. simpl. + rewrite Nat.lt_succ_r. + apply Nat.le_max_r. + Qed. + + + Lemma eval_Psatz_sound : forall env w l f', + make_conj (eval_nformula env) l -> + eval_Psatz l w = Some f' -> eval_nformula env f'. + Proof. + intros env w l f' H H0. + apply (fun H => eval_Psatz_Sound Rsor RSORaddon l _ H w) ; auto. + apply make_conj_in ; auto. + Qed. + + + Lemma Zeval_Psatz_sound : forall env w l f', + make_conj (Zeval_nformula env) l -> + eval_Psatz l w = Some f' -> Zeval_nformula env f'. + Proof. + intros env w l f' H H0. + apply (fun H => eval_Psatz_Sound Zsor ZSORaddon l _ H w) ; auto. + apply make_conj_in ; auto. + Qed. + + + + Lemma le_add_le_sub_l : + forall n m p : R, rle (rplus n p) m <-> rle p (rminus m n). + Proof. + intros. + split;intros. + - + apply (Rplus_le_mono_l Rsor) with (p := ropp n) in H. + assert (R1 : req (rplus (ropp n) (rplus n p)) p) by ring. + assert (R2 : req (rplus (ropp n) m) (rminus m n)) by ring. + now rewrite R1,R2 in H. + - apply (Rplus_le_mono_l Rsor) with (p := n) in H. + assert (R1 : req m (rplus n (rminus m n))) by ring. + now rewrite R1. + Qed. + + Lemma ropp_rminus : + forall n m : R, req (rminus m n) (rplus m (ropp n)). + Proof. + intros. + ring. + Qed. + + Lemma req_sym : forall x y, req x y -> req y x. + Proof. + intros. + rewrite H. + reflexivity. + Qed. + + Lemma rplus_phi_inv_l : forall x y r, + req (rplus y (IZR x)) (IZR r) -> + req y (IZR (r - x)). + Proof. + intros. + rewrite (morph_sub (SORrm RSORaddon)). + rewrite <- H. ring. + Qed. + + Lemma makeCuttingPlane_ns_sound : + forall env e e' c, + Zeval_nformula env (e, NonStrict) -> + makeCuttingPlane e = (e',c) -> + Zeval_nformula env (nformula_of_cutting_plane (e', c, NonStrict)). + Proof. + unfold nformula_of_cutting_plane. + unfold Zeval_nformula. unfold RingMicromega.eval_nformula. + unfold eval_op1. + intros env e e' c H H0. + rewrite (RingMicromega.eval_pol_add Zsor ZSORaddon). + simpl. + (**) + unfold makeCuttingPlane in H0. + revert H0. + case_eq (Zgcd_pol e) ; intros g c0. + case Z.gtb_spec. + - intros H0 H1 H2. + inv H2. + change (RingMicromega.eval_pol Z.add Z.mul (fun x : Z => x)) with Zeval_pol in *. + apply (Zgcd_pol_correct_lt _ env) in H1. 2: auto using Z.gt_lt. + apply Z.le_add_le_sub_l, Z.ge_le; rewrite Z.add_0_r. + apply (narrow_interval_lower_bound g (- c0) (Zeval_pol env (Zdiv_pol (PsubC Z.sub e c0) g))); auto using Z.lt_gt. + apply Z.le_ge. + rewrite <- Z.sub_0_l. + apply Z.le_sub_le_add_r. + rewrite <- H1. + assumption. + (* g <= 0 *) + - intros H0 H1 H2. inv H2. auto with zarith. + Qed. + + + Lemma cutting_plane_sound : forall env f p, + Zeval_nformula env f -> + genCuttingPlane f = Some p -> + Zeval_nformula env (nformula_of_cutting_plane p). + Proof. + unfold genCuttingPlane. + intros env f; destruct f as [e op]. + destruct op. + - (* Equal *) + intros p; destruct p as [[e' z] op]. + case_eq (Zgcd_pol e) ; intros g c. + case_eq (Z.gtb g 0 && (negb (Z.eqb c 0) && negb (Z.eqb (Z.gcd g c) g))) ; [discriminate|]. + case_eq (makeCuttingPlane e). + intros ? ? H H0 H1 H2 H3. + inv H3. + unfold makeCuttingPlane in H. + rewrite H1 in H. + revert H. + change (Zeval_pol env e = 0) in H2. + case_eq (Z.gtb g 0). + + intros H H3. + rewrite Z.gtb_lt in H. + rewrite Zgcd_pol_correct_lt with (1:= H1) in H2. 2: auto using Z.gt_lt. + unfold nformula_of_cutting_plane. + change (Zeval_pol env (padd e' (Pc z)) = 0). + inv H3. + rewrite (RingMicromega.eval_pol_add Zsor ZSORaddon). + fold Zeval_pol. + set (x:=Zeval_pol env (Zdiv_pol (PsubC Z.sub e c) g)) in *; clearbody x. + simpl. + rewrite andb_false_iff in H0. + destruct H0 as [H0|H0]. + * rewrite <-Z.gtb_lt in H ; congruence. + * rewrite andb_false_iff in H0. + destruct H0 as [H0|H0]. + -- rewrite negb_false_iff in H0. + apply Z.eqb_eq in H0. + subst. simpl. + rewrite Z.add_0_r, Z.mul_eq_0 in H2. + intuition subst; easy. + -- rewrite negb_false_iff in H0. + apply Z.eqb_eq in H0. + rewrite Zdivide_ceiling; cycle 1. + { apply Z.divide_opp_r. rewrite <-H0. apply Z.gcd_divide_r. } + apply Z.sub_move_0_r. + apply Z.div_unique_exact. + ++ now intros ->. + ++ now rewrite Z.add_move_0_r in H2. + + intros H H3. + unfold nformula_of_cutting_plane. + inv H3. + change (Zeval_pol env (padd e' (Pc 0)) = 0). + rewrite (RingMicromega.eval_pol_add Zsor ZSORaddon). + simpl. + now rewrite Z.add_0_r. + - (* NonEqual *) + intros ? H H0. + inv H0. + unfold Zeval_nformula in *. + unfold RingMicromega.eval_nformula in *. + unfold nformula_of_cutting_plane. + unfold eval_op1 in *. + rewrite (RingMicromega.eval_pol_add Zsor ZSORaddon). + simpl. now rewrite Z.add_0_r. + - (* Strict *) + intros p; destruct p as [[e' z] op]. + case_eq (makeCuttingPlane (PsubC Z.sub e 1)). + intros ? ? H H0 H1. + inv H1. + apply (makeCuttingPlane_ns_sound env) with (2:= H). + simpl in *. + rewrite (RingMicromega.PsubC_ok Zsor ZSORaddon). + now apply Z.lt_le_pred. + - (* NonStrict *) + intros p; destruct p as [[e' z] op]. + case_eq (makeCuttingPlane e). + intros ? ? H H0 H1. + inv H1. + apply (makeCuttingPlane_ns_sound env) with (2:= H). + assumption. + Qed. + + Lemma genCuttingPlaneNone : forall env f, + genCuttingPlane f = None -> + Zeval_nformula env f -> False. + Proof. + unfold genCuttingPlane. + intros env f; destruct f as [p o]. + destruct o. + - case_eq (Zgcd_pol p) ; intros g c. + case_eq (Z.gtb g 0 && (negb (Z.eqb c 0) && negb (Z.eqb (Z.gcd g c) g))). + + intros H H0 H1 H2. + flatten_bool. + match goal with [ H' : (g >? 0) = true |- ?G ] => rename H' into H3 end. + match goal with [ H' : negb (Z.eqb c 0) = true |- ?G ] => rename H' into H end. + match goal with [ H' : negb (Z.eqb (Z.gcd g c) g) = true |- ?G ] => rename H' into H5 end. + rewrite negb_true_iff in H5. + apply Z.eqb_neq in H5. + rewrite Z.gtb_lt in H3. + rewrite negb_true_iff in H. + apply Z.eqb_neq in H. + change (Zeval_pol env p = 0) in H2. + rewrite Zgcd_pol_correct_lt with (1:= H0) in H2. 2: auto using Z.gt_lt. + set (x:=Zeval_pol env (Zdiv_pol (PsubC Z.sub p c) g)) in *; clearbody x. + contradict H5. + apply Zis_gcd_gcd. + * apply Z.lt_le_incl; assumption. + * constructor; auto with zarith. + exists (-x). + rewrite Z.mul_opp_l, Z.mul_comm. + now apply Z.add_move_0_l. + (**) + + destruct (makeCuttingPlane p); discriminate. + - discriminate. + - destruct (makeCuttingPlane (PsubC Z.sub p 1)) ; discriminate. + - destruct (makeCuttingPlane p) ; discriminate. + Qed. + + + Lemma eval_nformula_mk_eq_pos : forall env x z t, + req (env x) (rminus (env z) (env t)) -> + eval_nformula env (xnnormalise (mk_eq_pos x z t)). + Proof. + intros. + rewrite xnnormalise_correct. + simpl. auto. + Qed. + + Lemma eval_nformula_bound_var : forall env x, + rle rO (env x) -> + eval_nformula env (xnnormalise (bound_var x)). + Proof. + intros. + rewrite xnnormalise_correct. + simpl. + easy. + Qed. + + + Definition agree_env (fr : positive) (env env' : positive -> R) : Prop := + forall x, Pos.le x fr -> env x = env' x. + + Lemma agree_env_subset : forall v1 v2 env env', + agree_env v1 env env' -> + Pos.le v2 v1 -> + agree_env v2 env env'. + Proof. + unfold agree_env. + intros v1 v2 env env' H ? ? ?. + apply H. + eapply Pos.le_trans ; eauto. + Qed. + + + Lemma agree_env_jump : forall fr j env env', + agree_env (fr + j) env env' -> + agree_env fr (Env.jump j env) (Env.jump j env'). + Proof. + intros fr j env env' H. + unfold agree_env ; intro. + intros. + unfold Env.jump. + apply H. + apply Pos.add_le_mono_r; auto. + Qed. + + + Lemma agree_env_tail : forall fr env env', + agree_env (Pos.succ fr) env env' -> + agree_env fr (Env.tail env) (Env.tail env'). + Proof. + intros fr env env' H. + unfold Env.tail. + apply agree_env_jump. + rewrite <- Pos.add_1_r in H. + apply H. + Qed. + + + Lemma max_var_acc : forall p i j, + (max_var (i + j) p = max_var i p + j)%positive. + Proof. + intros p; induction p as [|? ? IHp|? IHp1 ? ? IHp2]; simpl. + - reflexivity. + - intros. + rewrite ! IHp. + rewrite Pos.add_assoc. + reflexivity. + - intros. + rewrite !Pplus_one_succ_l. + rewrite ! IHp1. + rewrite ! IHp2. + rewrite ! Pos.add_assoc. + rewrite <- Pos.add_max_distr_r. + reflexivity. + Qed. + + + + Lemma agree_env_eval_nformula : + forall env env' e + (AGREE : agree_env (max_var xH (fst e)) env env'), + eval_nformula env e <-> eval_nformula env' e. + Proof. + intros env env' e; destruct e as [p o]. + simpl; intros AGREE. + fold eval_pol. + assert ((eval_pol env p) + = + (eval_pol env' p)) as H. + { + apply has_var_eval. + intros. + unfold agree_env in AGREE. + apply AGREE. apply max_var_correct; auto. + } + rewrite H. tauto. + Qed. + + Lemma agree_env_eval_nformulae : + forall env env' l + (AGREE : agree_env (max_var_nformulae l) env env'), + make_conj (eval_nformula env) l <-> + make_conj (eval_nformula env') l. + Proof. + intros env env' l; induction l as [|a l IHl]. + - simpl. tauto. + - intros. + rewrite ! make_conj_cons. + assert (eval_nformula env a <-> eval_nformula env' a) as H. + { + apply agree_env_eval_nformula. + eapply agree_env_subset ; eauto. + unfold max_var_nformulae. + simpl. + rewrite Pos.max_1_l. + apply max_var_nformulae_mono_aux. + apply Pos.le_refl. + } + rewrite H. + apply and_iff_compat_l. + apply IHl. + eapply agree_env_subset ; eauto. + unfold max_var_nformulae. + simpl. + apply max_var_nformulae_mono_aux'. + apply Pos.le_1_l. + Qed. + + + Lemma eq_true_iff_eq : + forall b1 b2 : bool, (b1 = true <-> b2 = true) <-> b1 = b2. + Proof. + intros b1 b2; destruct b1,b2 ; intuition congruence. + Qed. + + Lemma eval_nformula_split : forall env p, + eval_nformula env (p,NonStrict) \/ eval_nformula env (popp p,NonStrict). + Proof. + unfold popp. + simpl. intros. rewrite (eval_pol_opp Rsor RSORaddon). + fold eval_pol. + destruct (Rle_gt_cases Rsor rO (eval_pol env p)). + tauto. + rewrite <- (Ropp_pos_neg Rsor) in H. + rewrite (SORlt_le_neq Rsor) in H. + tauto. + Qed. + + Definition isZenv (z : option PositiveSet.t) (env: positive -> R) := + forall x, oget z x = true -> isZ (env x). + + Lemma isZenv_phi_env : forall z env pol, + isZenv z env -> + is_integral z pol = true -> + exists env', + (phi_env (mem (vars xH pol)) env' env). + Proof. + unfold isZenv. intros. + - pose proof (is_integral_sound pol z H0) as VARS. + assert (VarsZ : forall x, In x (vars 1 pol) -> isZ (env x)). + { + intros. apply H. now apply VARS. + } + assert (forall j, + (forall x : positive, In x (vars j pol) -> isZ (env x)) -> + exists env', forall x, In x (vars j pol) -> req (env x) (IZR (env' x))). + { + clear H H0 VARS VarsZ. + induction pol. + - exists (fun x => Z0). + simpl. tauto. + - simpl. + intros. + apply IHpol;auto. + - simpl ; intros. + assert (Hpol1 : forall x, In x (vars j pol1) -> isZ (env x)). + { + intros. apply H. + rewrite in_app_iff. tauto. + } + assert (Hpol2 : forall x, In x (vars (Pos.succ j) pol2) -> isZ (env x)). + { + intros. apply H. + rewrite in_app_iff. tauto. + } + apply IHpol1 in Hpol1. + apply IHpol2 in Hpol2. + destruct Hpol1 as (e1 & P1). + destruct Hpol2 as (e2 & P2). + assert (exists e3, req (env j) (IZR (e3 j))). + { + assert (isZ (env j)). + apply H. tauto. + destruct H0. exists (fun v => x). + rewrite H0;reflexivity. + } + destruct H0 as (e3 & P3). + exists (fun x => if In_dec Pos.eq_dec x (vars j pol1) + then e1 x else + if In_dec Pos.eq_dec x (vars (Pos.succ j) pol2) then e2 x + else e3 x). + intros. + rewrite in_app_iff in H0. + destruct (in_dec Pos.eq_dec x (vars j pol1)). + apply P1; auto. + destruct (in_dec Pos.eq_dec x (vars (Pos.succ j) pol2)). + apply P2;auto. + assert (j = x) by tauto. + subst. assumption. + } + destruct (H1 _ VarsZ) as (env' & EQ). + exists env'. + unfold phi_env. intros. now symmetry; apply EQ. + Qed. + + + Lemma Rlt_0_2 : rlt rO (rplus rI rI). + Proof. + assert (H : rlt rO rI) by exact (Rlt_0_1 Rsor). + apply (Rlt_trans Rsor _ rI); try easy. + rewrite <-(Rplus_0_l Rsor rI) at 1. + apply (Rplus_lt_le_mono Rsor);try easy. + apply (SORle_refl Rsor). + Qed. + + Lemma IPR_xI : + forall p : positive, + req (IPR (p~1)) + (rplus + (rtimes + (rplus rI rI) (IPR p)) rI). + Proof. + intros. simpl. ring. + Qed. + + Lemma IPR_xO : + forall p : positive, + req (IPR (p~0)) + (rtimes + (rplus rI rI) (IPR p)). + Proof. + intros. simpl. ring. + Qed. + + Lemma Rle_0_1 : rle rO rI. + Proof. + rewrite (Rle_lt_eq Rsor). + left. apply (Rlt_0_1 Rsor). + Qed. + + + Lemma IPR_ge_1 : forall p:positive, rle rI (IPR p). + Proof. + pose proof (Rlt_0_1) as H; pose proof (Rlt_0_2) as H'. + induction p as [p IH | p IH |]. + - rewrite IPR_xI. rewrite <-( Rplus_0_l Rsor rI) at 1. + apply (Rplus_le_mono_r Rsor). + apply (Rtimes_nonneg_nonneg Rsor). + + generalize Rlt_0_2. + rewrite (SORlt_le_neq Rsor). tauto. + + apply (Rle_trans Rsor _ rI); try apply IH. + apply Rle_0_1. + - simpl. rewrite <-(Rplus_0_l Rsor rI). + apply (Rplus_le_mono Rsor) ; try easy. + apply (Rle_trans Rsor _ rI); try apply IH. + apply Rle_0_1. + - simpl. apply (Rle_refl Rsor). + Qed. + + + Lemma IPR_gt_0 : forall p, rlt rO (IPR p). + Proof. + now intros p; apply (Rlt_le_trans Rsor _ rI); [apply (Rlt_0_1 Rsor) | apply IPR_ge_1]. + Qed. + + Lemma eq_IZR_R0 : forall n:Z, req (IZR n) rO -> n = 0%Z. + Proof. + intros. + destruct n as [| p | p]; unfold IZR; simpl in H; try easy. + - exfalso; apply (Rlt_neq Rsor rO (IPR p)). try easy; apply IPR_gt_0. + rewrite H ; reflexivity. + - exfalso; apply (Rlt_neq Rsor rO (IPR p)); try apply IPR_gt_0; symmetry. + apply (SORopp_wd Rsor) in H. + setoid_replace (ropp (ropp (IPR p))) with (IPR p) in H by ring. + rewrite H. ring. + Qed. + + Lemma eq_IZR : forall n m:Z, req (IZR n) (IZR m) -> n = m. + Proof. + intros n m H. + rewrite <- (Rminus_eq_0 Rsor) in H. + rewrite <- (morph_sub (SORrm RSORaddon)) in H. + apply Zminus_eq. now apply eq_IZR_R0. + Qed. + + + Lemma Rlt_irrefl : forall r, ~ rlt r r. + Proof. + intros r H. + intros. rewrite (Rlt_nge Rsor) in H. + apply H. apply (Rle_refl Rsor). + Qed. + + Lemma lt_0_IZR : forall n:Z, rlt rO (IZR n) -> (0 < n)%Z. + Proof. + intros. + revert n H. + intros [| p | p]; simpl; intros H. + - destruct (Rlt_irrefl _ H). + - now constructor. + - + rewrite (Rlt_nge Rsor _ _) in H. + destruct H. + apply (Rle_lt_eq Rsor). + left. + setoid_replace rO with (ropp rO) by ring. + rewrite <- (Ropp_lt_mono Rsor). + apply IPR_gt_0. + Qed. + + + + Lemma lt_IZR : forall n m:Z, rlt (IZR n) (IZR m) -> (n < m)%Z. + Proof. + intros z1 z2 H; apply Z.lt_0_sub. + apply lt_0_IZR. + rewrite (morph_sub (SORrm RSORaddon)). + apply (Rlt_lt_minus Rsor) in H;auto. + Qed. + + Lemma le_IZR : forall n m:Z, rle (IZR n) (IZR m) -> (n <= m)%Z. + Proof. + intros. + rewrite (Rle_lt_eq Rsor) in H. + destruct H as [H%lt_IZR | ->%eq_IZR]. + - now apply Z.lt_le_incl. + - now apply Z.le_refl. + Qed. + + Lemma Zeval_nformula_iff : forall env' env f, + phi_env (mem (vars 1 (fst f))) env' env -> + eval_nformula env f <-> Zeval_nformula env' f. + Proof. + intros. + unfold Zeval_nformula, eval_nformula,RingMicromega.eval_nformula. + fold eval_pol. fold Zeval_pol. + destruct f as (p,o). + apply eval_Zpol_phi in H. + destruct o ; simpl in *; + fold Zeval_pol; rewrite <- H. + - split ; intros. + apply eq_IZR;auto. + rewrite H0; reflexivity. + - split ; repeat intro. + apply H0. rewrite H1;reflexivity. + apply H0. now apply eq_IZR. + - split ; intros. + apply lt_IZR;auto. + rewrite (SORlt_le_neq Rsor). + split. + change rO with (IZR 0). + apply (SORcleb_morph RSORaddon). + apply Z.lt_le_incl in H0. + now apply Zbool.Zle_is_le_bool. + change rO with (IZR 0). + apply (SORcneqb_morph RSORaddon). + rewrite Z.eqb_neq. + intro. rewrite <- H1 in H0. + eapply Z.lt_irrefl;eauto. + - split ; intros. + apply le_IZR;auto. + change rO with (IZR 0). + apply (SORcleb_morph RSORaddon). + now apply Zbool.Zle_is_le_bool. + Qed. + + + + + Lemma vars_of_Zdiv_pol : forall z pol j x, + In x (vars j (Zdiv_pol pol z)) -> + In x (vars j pol). + Proof. + induction pol; simpl. + - auto. + - simpl; auto. + - intros. + rewrite in_app_iff in *. + destruct H as [H| [H |H]]. + + tauto. + + right. left;auto. + + right. right. auto. + Qed. + + + Lemma vars_of_PsubC : forall z pol j x, + In x (vars j (PsubC Z.sub pol z)) -> + In x (vars j pol). + Proof. + induction pol;simpl; auto. + - intros. + rewrite in_app_iff in *. + intuition auto. + Qed. + + Lemma vars_of_makeCuttingPlane : forall pol pol' c', + makeCuttingPlane pol = (pol', c') -> + forall x, In x (vars 1 pol') -> + In x (vars 1 pol). + Proof. + unfold makeCuttingPlane. + intros. destruct (Zgcd_pol pol). + destruct (z >? 0); try congruence. + inv H. + now apply vars_of_Zdiv_pol, vars_of_PsubC in H0. + Qed. + + + Lemma vars_of_cutting_plane : forall f p, + genCuttingPlane f = Some p -> + forall x, + In x (vars xH (fst (fst p))) -> In x (vars xH (fst f)). + Proof. + unfold genCuttingPlane. + destruct f as (pol,op). + destruct op. + - intros. + destruct (Zgcd_pol pol) as (g,c) eqn:GCD. + destruct ((g >? 0) && (negb (c =? 0) && negb (Z.gcd g c =? g))); + try discriminate. + simpl. destruct (makeCuttingPlane pol) as (pol',c') eqn:CUT. + inv H. + simpl in H0. + now apply vars_of_makeCuttingPlane with (x:=x) in CUT. + - intros. inv H. + apply H0. + - intros. + destruct (makeCuttingPlane (PsubC Z.sub pol 1)) as (pol',c') eqn:CUT. + inv H. + simpl in H0. + apply vars_of_makeCuttingPlane with (x:=x) in CUT;auto. + now apply vars_of_PsubC in CUT. + - intros. + destruct (makeCuttingPlane pol) as (pol',c') eqn:CUT. + inv H. + apply vars_of_makeCuttingPlane with (x:=x) in CUT;auto. + Qed. + + Lemma vars_of_PaddC : forall p j c x, + In x (vars j (PaddC Z.add p c)) -> + In x (vars j p). + Proof. + induction p; simpl;auto. + - intros. now apply IHp in H. + - intros. + rewrite in_app_iff in *. + destruct H as [H | [H | H]]; try tauto. + right. right. + now apply IHp2 in H. + Qed. + + Lemma phi_env_subset : forall S1 S2 env' env, + (forall x, S2 x -> S1 x) -> + phi_env S1 env' env -> + phi_env S2 env' env. + Proof. + unfold phi_env. + intros. apply H0;auto. + Qed. + + + Lemma cutting_plane_sound_R : forall env z f p, + isZenv z env -> + eval_nformula env f -> + is_integral z (fst f) = true -> + genCuttingPlane f = Some p -> + isZenv z env -> + eval_nformula env (nformula_of_cutting_plane p). + Proof. + intros. + destruct (isZenv_phi_env z env (fst f) H H1) as (env',EQ). + rewrite Zeval_nformula_iff with (env':=env'). + apply cutting_plane_sound with (f:=f); auto. + now rewrite <- Zeval_nformula_iff with (env:=env). + apply phi_env_subset with (2:=EQ). + intro. + unfold mem. + intros. + eapply vars_of_cutting_plane ;eauto. + unfold nformula_of_cutting_plane in H4. destruct p. destruct p. + simpl in *. + now apply vars_of_PaddC in H4. + Qed. + + Lemma genCuttingPlaneNone_R : forall z env f, + isZenv z env -> + is_integral z (fst f) = true -> + genCuttingPlane f = None -> + eval_nformula env f -> False. + Proof. + intros. + destruct (isZenv_phi_env z env (fst f) H H0) as (env',EQ). + rewrite Zeval_nformula_iff with (env':=env') in H2 by assumption. + revert H2. + now apply genCuttingPlaneNone. + Qed. + + Lemma is_integral_popp : forall s p, + is_integral s p = is_integral s (popp p). + Proof. + unfold is_integral. + destruct s as [s|]; try reflexivity. + intros pol. + generalize xH. + induction pol; simpl;auto. + - intros. + rewrite IHpol1. rewrite IHpol2. + reflexivity. + Qed. + + + Lemma valid_cut_sign_padd : forall env pol c op, + valid_cut_sign op = true -> + Zeval_nformula env (padd pol (Pc c), op) -> + - c <= Zeval_pol env pol. + Proof. + unfold Zeval_nformula. + unfold RingMicromega.eval_nformula. + fold eval_pol. intros. + destruct op; simpl in H ;unfold eval_op1 in H0; try discriminate. + - intros. rewrite Zeval_pol_add in H0. + simpl in H0. + setoid_replace (-c) with (-c + 0) by ring. + rewrite <- H0. + ring_simplify. + apply (SORle_refl Zsor). + - rewrite Zeval_pol_add in H0. + simpl in H0. + rewrite (SORlt_le_neq Zsor) in H0. + destruct H0. + apply (Rplus_le_mono_l Zsor) with (p:= -c) in H0. + ring_simplify in H0. auto. + - rewrite Zeval_pol_add in H0. + simpl in H0. + apply (Rplus_le_mono_l Zsor) with (p:= - c) in H0. + ring_simplify in H0. auto. + Qed. + + + + Lemma phi_env_merge : forall e1 e2 env S1 S2, + phi_env (mem S1) e1 env -> + phi_env (mem S2) e2 env -> + exists e3, phi_env (mem (S1 ++ S2)) e3 env. + Proof. + unfold phi_env. + intros. + exists (fun x => if In_dec Pos.eq_dec x S1 then e1 x else e2 x). + unfold mem in *. + intros. + destruct (in_dec Pos.eq_dec x S1). + apply H;auto. + rewrite in_app_iff in H1. destruct H1;try tauto. + apply H0;auto. + Qed. + + Lemma isZenv_add : + forall x z env + (ZX : isZ (env x)), + isZenv z env -> + isZenv (add x z) env. + Proof. + unfold isZenv in *. + intros. + unfold add,oget in H0. + destruct z. + - rewrite mem_add in H0. + destruct (Pos.eq_dec x0 x). + congruence. + apply H;auto. + - now apply H. + Qed. + + + Lemma ZChecker_sound : forall w z l, + ZChecker z l w = true -> forall env, isZenv z env -> make_impl (eval_nformula env) l False. + Proof. + intros w; induction w as [w H] using (well_founded_ind (well_founded_ltof _ bdepth)). + destruct w as [ | w pf | w pf | p pf1 pf2 | w1 w2 pf | x pf]. + - (* DoneProof *) + simpl. discriminate. + - (* RatProof *) + simpl. + intros z l. case_eq (eval_Psatz l w) ; [| discriminate]. + intros f Hf. + case_eq (Zunsat f). + + intros H0 ? ? _. + apply (checker_nf_sound Rsor RSORaddon l w). + unfold check_normalised_formulas. unfold eval_Psatz in Hf. rewrite Hf. + unfold Zunsat in H0. assumption. + + intros H0 H1 env HZ. + assert (make_impl (eval_nformula env) (f::l) False) as H2. + { apply H with (2:= H1);auto. + unfold ltof. + simpl. + auto with arith. + } + destruct f. + rewrite <- make_conj_impl in H2. + rewrite make_conj_cons in H2. + rewrite <- make_conj_impl. + intro. + apply H2. + split ; auto. + apply eval_Psatz_sound with (2:= Hf) ; assumption. + - (* CutProof *) + simpl. + intros z l. + case_eq (eval_Psatz l w) ; [ | discriminate]. + intros f' Hlc. + destruct (is_integral z (fst f')) eqn:INT. + case_eq (genCuttingPlane f'). + + intros p H0 H1 env HZ. + rewrite <- make_conj_impl. + intro. + apply (eval_Psatz_sound env) in Hlc. + apply cutting_plane_sound_R with (z:=z)(2:= Hlc) in H0;auto. + destruct (Zunsat (nformula_of_cutting_plane p)) eqn:UNSAT. + * eapply Zunsat_sound; eauto. + * + assert (make_impl (eval_nformula env) (nformula_of_cutting_plane p::l) False) as HN. + { eapply (H pf) with (env:=env) in H1 ; auto. + unfold ltof. + simpl. + auto with arith. + } + rewrite <- make_conj_impl in HN. + rewrite make_conj_cons in HN. + tauto. + * auto. + + (* genCuttingPlane = None *) + intros H0 H1 env. + rewrite <- make_conj_impl. + intros ZENV H2. + apply eval_Psatz_sound with (2:= Hlc) in H2. + apply genCuttingPlaneNone_R with (z:=z) (4:= H2) ; auto. + + discriminate. + - (* SplitProof *) + intros z l. + cbn - [genCuttingPlane]. + destruct (is_integral z p) eqn:INT ; [| discriminate]. + case_eq (genCuttingPlane (p, NonStrict)) ; [| discriminate]. + case_eq (genCuttingPlane (popp p, NonStrict)) ; [| discriminate]. + intros cp1 GCP1 cp2 GCP2 ZC1 env ZENV. + flatten_bool. + match goal with [ H' : ZChecker _ _ pf1 = true |- _ ] => rename H' into H0 end. + match goal with [ H' : ZChecker _ _ pf2 = true |- _ ] => rename H' into H1 end. + destruct (eval_nformula_split env p). + + apply H with (env:=env) in H0;auto. + * rewrite <- make_conj_impl in *. + intro ; apply H0. + rewrite make_conj_cons. split; auto. + apply (cutting_plane_sound_R _ z (p,NonStrict)) ; auto. + * apply ltof_bdepth_split_l. + + apply H with (env:=env) in H1;auto. + * rewrite <- make_conj_impl in *. + intro ; apply H1. + rewrite make_conj_cons. split; auto. + apply (cutting_plane_sound_R _ z (popp p,NonStrict)) ; auto. + rewrite <- INT. symmetry. apply is_integral_popp. + * apply ltof_bdepth_split_r. + - (* EnumProof *) + intros z l. + simpl. + case_eq (eval_Psatz l w1) ; [ | discriminate]. + case_eq (eval_Psatz l w2) ; [ | discriminate]. + intros f1 Hf1 f2 Hf2. + destruct (is_integral z (fst f2) && is_integral z (fst f1)) eqn: isINT; try discriminate. + rewrite andb_true_iff in isINT. + destruct isINT as (INT1 & INT2). + case_eq (genCuttingPlane f2). + + intros p; destruct p as [ [p1 z1] op1]. + case_eq (genCuttingPlane f1). + * intros p; destruct p as [ [p2 z2] op2]. + case_eq (valid_cut_sign op1 && valid_cut_sign op2 && is_pol_Z0 (padd p1 p2)). + -- intros Hcond. + flatten_bool. + match goal with [ H1 : is_pol_Z0 (padd p1 p2) = true |- _ ] => rename H1 into HZ0 end. + match goal with [ H2 : valid_cut_sign op1 = true |- _ ] => rename H2 into Hop1 end. + match goal with [ H3 : valid_cut_sign op2 = true |- _ ] => rename H3 into Hop2 end. + intros HCutL HCutR Hfix env. + (* get the bounds of the enum *) + rewrite <- make_conj_impl. + intros isZ H0. + assert (INT1':=INT1). + assert (INT2':=INT2). + apply isZenv_phi_env with (env:=env) in INT1. + apply isZenv_phi_env with (env:=env) in INT2. + destruct INT1 as (e1 & PHI1). + destruct INT2 as (e2 & PHI2). + destruct (phi_env_merge _ _ _ _ _ PHI1 PHI2) as (e3 & PHI3). + assert (-z1 <= Zeval_pol e3 p1 <= z2) as H1. { + split. + - apply (eval_Psatz_sound env) in Hf2 ; auto. + assert (PHI1' : + phi_env (mem (vars 1 (fst f2))) e3 env). + { eapply phi_env_subset in PHI3 ;eauto. + intros. unfold mem in *. + rewrite in_app_iff. tauto. + } + rewrite Zeval_nformula_iff with (env':=e3)in Hf2 by easy. + apply cutting_plane_sound with (1:= Hf2) in HCutR;auto. + unfold nformula_of_cutting_plane in HCutR. + eapply valid_cut_sign_padd in HCutR; eauto. + - apply (eval_Psatz_sound env) in Hf1 ; auto. + assert (HCutL' := HCutL). + apply cutting_plane_sound_R with (z:=z) (2:= Hf1) in HCutL;auto. + unfold nformula_of_cutting_plane in HCutL. + rewrite Zeval_nformula_iff with (env':=e3)in HCutL. + eapply valid_cut_sign_padd in HCutL; eauto. + apply is_pol_Z0_Zeval_pol with (env:=e3) in HZ0. + rewrite Zeval_pol_add in HZ0. + apply (Rplus_le_mono_l Zsor) with (p:= z2) in HCutL. + ring_simplify in HCutL. + rewrite <- HZ0 in HCutL. + apply (Rplus_le_mono_l Zsor) with (p:= - (Zeval_pol e3 p2)) in HCutL. + ring_simplify in HCutL. + assumption. + { + eapply phi_env_subset in PHI3 ;eauto. + intros. unfold mem in H1. unfold mem. + rewrite in_app_iff. right. + simpl in H1. + apply vars_of_PaddC in H1. + eapply vars_of_cutting_plane in HCutL'; eauto. + } + } + revert Hfix. + match goal with + | |- context[?F pf (-z1) z2 = true] => set (FF := F) + end. + intros Hfix. + assert (HH :forall x, -z1 <= x <= z2 -> exists pr, + (In pr pf /\ + ZChecker z ((PsubC Z.sub p1 x,Equal) :: l) pr = true)%Z). { + clear HZ0 Hop1 Hop2 HCutL HCutR H0 H1. + revert Hfix. + generalize (-z1). clear z1. intro z1. + revert z1 z2. + induction pf as [|a pf IHpf];simpl ;intros z1 z2 Hfix x **. + - revert Hfix. + now case (Z.gtb_spec); [ | easy ]; intros LT; elim (Zorder.Zlt_not_le _ _ LT); transitivity x. + - flatten_bool. + match goal with [ H' : _ <= x <= _ |- _ ] => rename H' into H0 end. + match goal with [ H' : FF pf (z1 + 1) z2 = true |- _ ] => rename H' into H2 end. + destruct (ZArith_dec.Z_le_lt_eq_dec _ _ (proj1 H0)) as [ LT | -> ]. + 2: exists a; auto. + rewrite <- Z.le_succ_l in LT. + assert (LE: (Z.succ z1 <= x <= z2)%Z) by intuition. + elim IHpf with (2:=H2) (3:= LE). + + intros x0 ?. + exists x0 ; split;tauto. + + intros until 1. + apply H ; auto. + cbv [ltof] in *. + cbn [bdepth] in *. + eauto using Nat.lt_le_trans, le_n_S, Nat.le_max_r. + } + (*/asser *) + destruct (HH _ H1) as [pr [Hin Hcheker]]. + assert (make_impl (eval_nformula env) ((PsubC Z.sub p1 (Zeval_pol e3 p1),Equal) :: l) False) as H2. { + eapply (H pr) ;eauto. + apply in_bdepth ; auto. + } + rewrite <- make_conj_impl in H2. + apply H2. + rewrite make_conj_cons. + split ;auto. + unfold eval_nformula. + unfold RingMicromega.eval_nformula. + simpl. + rewrite (RingMicromega.PsubC_ok Rsor RSORaddon). + rewrite eval_Zpol_phi with (env':=env). + fold eval_pol. + ring. + eapply phi_env_subset in PHI3;eauto. + intros. unfold mem in *. rewrite in_app_iff. + left. + eapply vars_of_cutting_plane in HCutR; eauto. + auto. auto. + -- discriminate. + * (* No cutting plane *) + intros H0 H1 H2 env isZ. + rewrite <- make_conj_impl. + intros H3. + apply eval_Psatz_sound with (2:= Hf1) in H3. + apply genCuttingPlaneNone_R with (z:=z)(4:= H3) ; auto. + + (* No Cutting plane (bis) *) + intros H0 H1 env isZ. + rewrite <- make_conj_impl. + intros H2. + apply eval_Psatz_sound with (2:= Hf2) in H2. + apply genCuttingPlaneNone_R with (z:=z)(4:= H2) ; auto. + - intros z l. + unfold ZChecker. + fold ZChecker. + set (fr := (max_var_nformulae l)%positive). + set (z1 := (Pos.succ fr)) in *. + set (t1 := (Pos.succ z1)) in *. + destruct ((x <=? fr) && oget z x)%positive eqn:LE ; [|congruence]. + rewrite andb_true_iff in LE. + destruct LE as (LE1 & LE2). + intros H0 env isZE. + assert (VX :=isZE x LE2). unfold isZ in VX. + destruct VX as (vx & EQ). + set (env':= fun v => if Pos.eqb v z1 + then if Z.leb vx 0 then rO else env x + else if Pos.eqb v t1 + then if Z.leb vx 0 then ropp (IZR vx) else rO + else env v). + apply H with (env:=env') in H0. + + rewrite <- make_conj_impl in *. + intro H1. + rewrite !make_conj_cons in H0. + apply H0 ; repeat split. + * + apply eval_nformula_mk_eq_pos. + unfold env'. + rewrite! Pos.eqb_refl. + replace (x=?z1)%positive with false. + 1:replace (x=?t1)%positive with false. + 1:replace (t1=?z1)%positive with false. + 1:destruct (vx <=? 0); rewrite <- ?EQ; ring. + { unfold t1. + symmetry; apply not_true_iff_false; rewrite Pos.eqb_eq; symmetry; apply Pos.succ_discr. + } + { + unfold t1, z1. + symmetry; apply not_true_iff_false; rewrite Pos.eqb_eq; intros ->. + apply Pos.leb_le, Pos.lt_succ_r in LE1; rewrite <-?Pos.succ_lt_mono in *. + pose proof Pos.lt_not_add_l fr 1; rewrite Pos.add_1_r in *; contradiction. + } + { + unfold z1. + symmetry; apply not_true_iff_false; rewrite Pos.eqb_eq; intros ->. + apply Pos.leb_le, Pos.lt_succ_r in LE1; rewrite <-?Pos.succ_lt_mono in *. + case (Pos.lt_irrefl _ LE1). + } + * + apply eval_nformula_bound_var. + unfold env'. + rewrite! Pos.eqb_refl. + destruct (vx <=? 0) eqn:EQV. + -- apply (Rle_refl Rsor). + -- rewrite <- EQ. + rewrite <- (morph0 (SORrm RSORaddon)). + apply (SORcleb_morph RSORaddon). + apply Z.leb_gt in EQV. + apply Z.leb_le. + apply Z.lt_le_incl; trivial. + * + apply eval_nformula_bound_var. + unfold env'. + rewrite! Pos.eqb_refl. + replace (t1 =? z1)%positive with false. + -- destruct (vx <=? 0) eqn:EQ1. + ++ + rewrite <- (morph0 (SORrm RSORaddon)). + rewrite <- (morph_opp (SORrm RSORaddon)). + apply (SORcleb_morph RSORaddon). + rewrite Z.leb_le in *. + rewrite Z.opp_le_mono, Z.opp_involutive; trivial. + ++ apply (Rle_refl Rsor). + -- unfold t1. + clear. + symmetry; apply not_true_iff_false; rewrite Pos.eqb_eq; symmetry; apply Pos.succ_discr. + * + rewrite (agree_env_eval_nformulae _ env') in H1;auto. + unfold agree_env; intros x0 H2. + unfold env'. + replace (x0 =? z1)%positive with false. + 1:replace (x0 =? t1)%positive with false. + 1:reflexivity. + { + unfold t1, z1. + symmetry; apply not_true_iff_false; rewrite Pos.eqb_eq; intros ->. + apply Pos.lt_succ_r in H2; rewrite <-?Pos.succ_lt_mono in *. + pose proof Pos.lt_not_add_l (max_var_nformulae l) 1; rewrite Pos.add_1_r in *; contradiction. + } + { + unfold z1, fr in *. + symmetry; apply not_true_iff_false; rewrite Pos.eqb_eq; intros ->. + apply Pos.lt_succ_r in H2; rewrite <-?Pos.succ_lt_mono in *. + case (Pos.lt_irrefl _ H2). + } + + unfold ltof. + simpl. + apply Nat.lt_succ_diag_r. + + apply isZenv_add. + unfold env'. rewrite Pos.eqb_refl. + unfold isZ. exists (if vx <=? 0 then 0 else vx). + destruct (vx <=? 0); try easy. + apply isZenv_add. + unfold env'. + replace (t1 =?z1)%positive with false. + rewrite Pos.eqb_refl. + unfold isZ. exists (if vx <=? 0 then - vx else 0). + destruct (vx <=? 0); try easy. + apply (morph_opp (SORrm RSORaddon)). + unfold t1,z1. + symmetry. rewrite Pos.eqb_neq. intro. + rewrite <- Pos.add_1_l in H1. + eapply Pos.add_no_neutral; eauto. + { unfold isZenv. + intros. + apply isZE in H1. + unfold env'. + destruct ((x0 =? z1)%positive). + destruct (vx <=? 0); auto. + exists 0. apply (morph0 (SORrm RSORaddon)). + destruct ((x0 =? t1)%positive). + destruct (vx <=? 0). + apply isZ_ropp. exists vx; reflexivity. + exists 0. apply (morph0 (SORrm RSORaddon)). + auto. + } + Qed. + +End S. + +Fixpoint xhyps_of_pt (base:nat) (acc : list nat) (pt:ZArithProof) : list nat := + match pt with + | DoneProof => acc + | RatProof c pt => xhyps_of_pt (S base ) (xhyps_of_psatz base acc c) pt + | CutProof c pt => xhyps_of_pt (S base ) (xhyps_of_psatz base acc c) pt + | SplitProof p pt1 pt2 => xhyps_of_pt (S base) (xhyps_of_pt (S base) acc pt1) pt2 + | EnumProof c1 c2 l => + let acc := xhyps_of_psatz base (xhyps_of_psatz base acc c2) c1 in + List.fold_left (xhyps_of_pt (S base)) l acc + | ExProof _ pt => xhyps_of_pt (S (S (S base ))) acc pt + end. + +Definition hyps_of_pt (pt : ZArithProof) : list nat := xhyps_of_pt 0 nil pt. + +Open Scope Z_scope. + +(** To ease bindings from ml code **) +Definition make_impl := Refl.make_impl. +Definition make_conj := Refl.make_conj. + +From Stdlib Require VarMap. + +(*Definition varmap_type := VarMap.t Z. *) +Definition env := PolEnv Z. +Definition node := @VarMap.Branch Z. +Definition empty := @VarMap.Empty Z. +Definition leaf := @VarMap.Elt Z. + +Definition coneMember := ZWitness. + +Definition eval := eval_formula. + +#[deprecated(note="Use [prod positive nat]", since="9.0")] + Definition prod_pos_nat := prod positive nat. + +#[deprecated(use=Z.to_N, since="9.0")] + Notation n_of_Z := Z.to_N (only parsing). diff --git a/theories/micromega/RMicromega.v b/theories/micromega/RMicromega.v index 2fe8662e7f..e578cee3bf 100644 --- a/theories/micromega/RMicromega.v +++ b/theories/micromega/RMicromega.v @@ -15,7 +15,7 @@ (************************************************************************) From Stdlib Require Import OrderedRing. -From Stdlib Require Import QMicromega RingMicromega. +From Stdlib Require Import QMicromega RingMicromega MMicromega. From Stdlib Require Import Refl. From Stdlib Require Import Sumbool Raxioms Rfunctions RIneq Rpow_def. From Stdlib Require Import BinNat. @@ -24,8 +24,11 @@ From Stdlib Require Import Qfield. From Stdlib Require Import Qreals. From Stdlib Require Import DeclConstant. From Stdlib Require Import Znat. - +From Stdlib Require Import MSetPositive. From Stdlib Require Setoid. +From Stdlib Require Import isZ. +Set Implicit Arguments. + Definition Rsrt : ring_theory R0 R1 Rplus Rmult Rminus Ropp (@eq R). Proof. @@ -42,7 +45,6 @@ Proof. - exact Rplus_opp_r. Qed. -Local Open Scope R_scope. Lemma Rsor : SOR R0 R1 Rplus Rmult Rminus Ropp (@eq R) Rle Rlt. Proof. @@ -59,6 +61,9 @@ Proof. - now apply Rmult_lt_0_compat. Qed. +Local Open Scope R_scope. + + Lemma Rinv_1 : forall x, x * / 1 = x. Proof. intro. @@ -134,6 +139,87 @@ Proof. - apply Qle_true. Qed. +Lemma IPR_2_eq : forall p, + IPR_2 p = MMicromega.IPR R R1 Rplus p + MMicromega.IPR R R1 Rplus p. +Proof. + induction p; simpl; auto. + - rewrite <- IHp. + ring. + - rewrite <- IHp. + ring. +Qed. + +Lemma IPR_eq : forall p, + IPR p = MMicromega.IPR R R1 Rplus p. +Proof. + unfold IPR; destruct p; simpl; + try rewrite IPR_2_eq; reflexivity. +Qed. + + +Lemma IZR_eq : forall x, + MMicromega.IZR R R0 R1 Rplus Ropp x = IZR x. +Proof. + destruct x ; simpl. + - reflexivity. + - unfold IZR. + now rewrite IPR_eq. + - unfold IZR. + f_equal. + now rewrite IPR_eq. +Qed. + +Lemma isZ_eq : forall r, + isZ.isZ r -> + MMicromega.isZ R R0 R1 Rplus Ropp eq r. +Proof. + intros. + destruct H as (z & EQ). + exists z. + now rewrite IZR_eq. +Qed. + + + + + + +Lemma ZSORaddon : + @SORaddon R + R0 R1 Rplus Rmult Rminus Ropp (@eq R) Rle (* ring elements *) + Z 0%Z 1%Z Z.add Z.mul Z.sub Z.opp (* coefficients *) + Z.eqb Z.leb + (MMicromega.IZR R R0 R1 Rplus Ropp) + nat N.to_nat pow. +Proof. + constructor. + - constructor ; intros. + + reflexivity. + + reflexivity. + + + rewrite! IZR_eq. + apply plus_IZR. + + rewrite! IZR_eq. + apply minus_IZR. + + rewrite! IZR_eq. + apply mult_IZR. + + rewrite! IZR_eq. + apply opp_IZR. + + rewrite! IZR_eq. + rewrite Z.eqb_eq in H. now f_equal. + - apply R_power_theory. + - intros. + rewrite! IZR_eq. + rewrite Z.eqb_neq in H. + intro. apply H. + apply eq_IZR; auto. + - intros. + rewrite! IZR_eq. + apply IZR_le. + now rewrite Z.leb_le in H. +Qed. + + (* Syntactic ring coefficients. *) Inductive Rcst := @@ -372,12 +458,6 @@ Qed. From Stdlib Require Import EnvRing. -Definition INZ (n:N) : R := - match n with - | N0 => IZR 0%Z - | Npos p => IZR (Zpos p) - end. - Definition Reval_expr := eval_pexpr Rplus Rmult Rminus Ropp R_of_Rcst N.to_nat pow. @@ -464,9 +544,8 @@ Definition QReval_expr := eval_pexpr Rplus Rmult Rminus Ropp Q2R N.to_nat pow. Definition QReval_formula (e: PolEnv R) (k: Tauto.kind) (ff : Formula Q) := let (lhs,o,rhs) := ff in Reval_op2 k o (QReval_expr e lhs) (QReval_expr e rhs). - Definition QReval_formula' := - eval_formula Rplus Rmult Rminus Ropp (@eq R) Rle Rlt Q2R N.to_nat pow. + RingMicromega.eval_formula Rplus Rmult Rminus Ropp (@eq R) Rle Rlt Q2R N.to_nat pow. Lemma QReval_formula_compat : forall env b f, Tauto.hold b (QReval_formula env b f) <-> QReval_formula' env f. Proof. @@ -478,17 +557,19 @@ Proof. apply Reval_pop2_eval_op2. Qed. -Definition Qeval_nformula := - eval_nformula 0 Rplus Rmult (@eq R) Rle Rlt Q2R. +Definition Qeval_nformula := + RingMicromega.eval_nformula 0 Rplus Rmult (@eq R) Rle Rlt Q2R. Lemma Reval_nformula_dec : forall env d, (Qeval_nformula env d) \/ ~ (Qeval_nformula env d). Proof. exact (fun env d =>eval_nformula_dec Rsor Q2R env d). Qed. -Definition RWitness := Psatz Q. +Definition RWitness := ZArithProof. + +(* Definition RWeakChecker := check_normalised_formulas 0%Q 1%Q Qplus Qmult Qeq_bool Qle_bool. From Stdlib Require Import List. @@ -504,6 +585,7 @@ Proof. unfold RWeakChecker in H. exact H. Qed. + *) From Stdlib.micromega Require Import Tauto. @@ -514,53 +596,499 @@ Definition runsat := check_inconsistent 0%Q Qeq_bool Qle_bool. Definition rdeduce := nformula_plus_nformula 0%Q Qplus Qeq_bool. -Definition RTautoChecker (f : BFormula (Formula Rcst) Tauto.isProp) (w: list RWitness) : bool := - @tauto_checker (Formula Q) (NFormula Q) - unit runsat rdeduce - (Rnormalise unit) (Rnegate unit) - RWitness (fun cl => RWeakChecker (List.map fst cl)) (map_bformula (map_Formula Q_of_Rcst) f) w. +Inductive eFormula (Form :Type) := +| IsZ (b:bool) (x:positive) +| IsF (f: Form). + +Definition map_eFormula {S C :Type} (F: S -> C) (f: eFormula (Formula S)) : eFormula (Formula C) := + match f with + | IsZ _ b x => IsZ _ b x + | IsF f => IsF (map_Formula F f) + end. + +Fixpoint lcm (p: Pol Q) : Z := + match p with + | Pc q => Zpos (Qden q) + | Pinj x p => lcm p + | PX p _ q => Z.lcm (lcm p) (lcm q) + end. + +Fixpoint polZ (lcm:Z) (p:Pol Q) : Pol Z := + match p with + | Pc q => Pc (Z.div (Z.mul (Qnum q) lcm) (Z.pos (Qden q))) + | Pinj x p => Pinj x (polZ lcm p) + | PX p i q => PX (polZ lcm p) i (polZ lcm q) + end. + +Lemma xpolZ_eq : forall (p:Pol Q) a env + (POS : (0 <= a)%Z) + (DIV : (lcm p | a)%Z), + Q2R (a # 1) * (RingMicromega.eval_pol Rplus Rmult Q2R env p) = + (RingMicromega.eval_pol Rplus Rmult (MMicromega.IZR R R0 R1 Rplus Ropp) env (polZ a p)). +Proof. + induction p. + - simpl. intros. + unfold Z.divide in DIV. + destruct DIV as (z & EQ). + subst. + rewrite Z.mul_assoc. + rewrite Zdiv.Z_div_mult. + rewrite <- Q2R_mult. + destruct c. simpl. + destruct z. + + simpl. + rewrite Qmult_0_l. + rewrite Z.mul_0_r. + unfold Q2R. simpl. + rewrite Rinv_1. reflexivity. + + + (*rewrite <- (Pos2Z.inj_mul p). *) + setoid_rewrite <- (Qmult_inject_Z_l (Z.pos Qden)). + setoid_rewrite <- Qmult_assoc. + setoid_rewrite Qreduce_den_l. + simpl. unfold Q2R. + rewrite IZR_eq. + rewrite Rinv_1. + rewrite Z.mul_comm. + reflexivity. + + simpl in POS. + exfalso. + revert POS. rewrite ZMicromega.le_neg. + simpl. + apply Pos2Z.pos_is_pos. + + apply Zorder.Zgt_pos_0. + - simpl. + intros. + rewrite IHp by auto. reflexivity. + - intros. + simpl. + rewrite <- IHp1; try easy. + rewrite <- IHp2; try easy. + ring. + + simpl in DIV. + generalize (Z.divide_lcm_r (lcm p1) (lcm p3)). + intros. + eapply Z.divide_trans;eauto. + + simpl in DIV. + generalize (Z.divide_lcm_l (lcm p1) (lcm p3)). + intros. + eapply Z.divide_trans;eauto. +Qed. + + +Lemma lt_0_lcm : forall p, + (0 < lcm p)%Z. +Proof. + induction p. + - simpl. + destruct c ; simpl. + apply Pos2Z.is_pos. + - easy. + - simpl. + specialize (Z.lcm_nonneg (lcm p1) (lcm p3)). + destruct (Z.eq_dec (Z.lcm (lcm p1) (lcm p3)) 0). + apply Z.lcm_eq_0 in e. + destruct e as [e|e]; rewrite e in *. + apply Z.lt_irrefl in IHp1; tauto. + apply Z.lt_irrefl in IHp2; tauto. + intro H. apply Zorder.Zle_lt_or_eq in H. + destruct H as [H | H] ;congruence. +Qed. + + + +Lemma polZ_eq : forall (p:Pol Q) env, + Q2R ((lcm p) # 1) * (RingMicromega.eval_pol Rplus Rmult Q2R env p) = + (RingMicromega.eval_pol Rplus Rmult (MMicromega.IZR R R0 R1 Rplus Ropp) env (polZ (lcm p) p)). +Proof. + intros. + apply xpolZ_eq. + apply Z.lt_le_incl. + apply lt_0_lcm. + apply Z.divide_refl. +Qed. + + +Definition nformulaZ (f : NFormula Q) : NFormula Z := + let (p,o) := f in + (polZ (lcm p) p, o). + + +Lemma nformulaZ_sound : forall env f, + Qeval_nformula env f -> + eval_nformula R R0 R1 Rplus Rmult Ropp eq Rle Rlt env (nformulaZ f). +Proof. + unfold Qeval_nformula. + destruct f as (f & o). + unfold nformulaZ. + simpl. + rewrite <- polZ_eq. + assert (LCMPOS := lt_0_lcm f). + destruct o ; simpl. + - intros H. rewrite H. + rewrite Rmult_0_r. + reflexivity. + - intro H. intro H'. + apply H. + apply Rmult_integral in H'. + destruct H' as [H' | H'];[|easy]. + unfold Q2R in H'. + unfold Qnum,Qden in H'. + rewrite Rinv_1 in H'. + apply eq_IZR in H'. rewrite H' in LCMPOS. + apply Z.lt_irrefl in LCMPOS ; tauto. + - intro H. + replace R0 with (R0 * R0) by ring. + apply Rmult_ge_0_gt_0_lt_compat; try easy. + apply Rlt_gt. + unfold Q2R. simpl. rewrite Rinv_1. + now apply IZR_lt. + unfold Q2R. simpl. rewrite Rinv_1. + change R0 with (IZR 0). + now apply IZR_lt. + - intro H. + replace R0 with (R0 * R0) by ring. + apply Rmult_le_compat; try easy. + unfold Q2R. simpl. + rewrite Rinv_1. + change R0 with (IZR 0). + apply IZR_le. + now apply Z.lt_le_incl. +Qed. + + +Fixpoint xcollect_isZ (s:PositiveSet.t) (acc : list (NFormula Z)) (l:list (eFormula (NFormula Q) * unit)) : PositiveSet.t * list (NFormula Z) := + match l with + | nil => (s,acc) + | cons (e,_) l => match e with + | IsZ _ b x => xcollect_isZ (if b then (PositiveSet.add x s) else s) acc l + | IsF f => xcollect_isZ s (nformulaZ f :: acc) l + end + end. + +Definition QCheck (l : list (eFormula (NFormula Q) * unit)) := + let (s,cl) := xcollect_isZ PositiveSet.empty nil l in + ZChecker (Some s) cl. + +Definition erunsat (f : eFormula (NFormula Q)) := + match f with + | IsZ _ _ _ => false + | IsF f => runsat f + end. + +Definition erdeduce (f1 f2 : eFormula (NFormula Q)) := + match f1,f2 with + | IsF f1, IsF f2 => match rdeduce f1 f2 with + | None => None + | Some f => Some (IsF f) + end + | _ , _ => None + end. + +Definition map_cnf (T:Type) {A B:Type} (F : A -> B) (l : cnf A T) : cnf B T := + List.map (List.map (fun '(a,t) => (F a , t))) l. + +Definition eRnormalise (T:Type) (f: eFormula (Formula Q)) (t:T) : cnf (eFormula (NFormula Q)) T := + match f with + | IsZ _ b z => cons (cons (IsZ _ (negb b) z, t) nil) nil + | IsF f => map_cnf (@IsF (NFormula Q)) (Rnormalise f t) + end. + +Definition eRnegate (T:Type) (f: eFormula (Formula Q)) (t:T) : cnf (eFormula (NFormula Q)) T := + match f with + | IsZ _ b z => cons (cons (IsZ _ b z, t) nil) nil + | IsF f => map_cnf (@IsF (NFormula Q)) (Rnegate f t) + end. + + +Definition cnfR (Annot:Type) (TX : Tauto.kind -> Type) (AF:Type) (k:Tauto.kind) + (f: TFormula (eFormula (Formula Q)) Annot TX AF k) := + rxcnf erunsat erdeduce (@eRnormalise Annot) (@eRnegate Annot) true f. + +Definition RTautoChecker (f : BFormula (eFormula (Formula Rcst)) Tauto.isProp) (w: list RWitness) : bool := + @tauto_checker (eFormula (Formula Q)) (eFormula (NFormula Q)) + unit erunsat erdeduce + (@eRnormalise unit) (@eRnegate unit) + RWitness (QCheck) (map_bformula (map_eFormula Q_of_Rcst) f) w. + +Definition Qeval_formula (e: PolEnv R) (k:kind) (ff : Formula Q) : rtyp k := + let (lhs, o, rhs) := ff in Reval_op2 k o (QReval_expr e lhs) (QReval_expr e rhs). + + +Definition eval_eformula_k {F: Type} (eval : PolEnv R -> forall (k:kind),F -> rtyp k) (e:PolEnv R) (k:kind) (f : eFormula F) : + rtyp k := + match f with + | IsZ _ b x => match k with + | isProp => if b then isZ (e x) else ~ isZ (e x) + | isBool => if b then isZb (e x) else negb (isZb (e x)) + end + | IsF f => eval e k f + end. + +Definition eval_eformula {F: Type} (eval : PolEnv R -> F -> Prop) (e:PolEnv R) (f : eFormula F) : Prop := + match f with + | IsZ _ b x => if b then isZ (e x) else ~ isZ (e x) + | IsF f => eval e f + end. -Lemma RTautoChecker_sound : forall f w, RTautoChecker f w = true -> forall env, eval_bf (Reval_formula env) f. + +Lemma eval_eformula_morph : forall {F: Type} (eval1: PolEnv R -> forall (k:kind),F -> rtyp k) + (eval2 : PolEnv R -> F -> Prop), + (forall e k f, hold k (eval1 e k f ) <-> (eval2 e f)) -> + forall e k f, + hold k (eval_eformula_k eval1 e k f) <-> (eval_eformula eval2 e f). +Proof. + destruct f; simpl. + - destruct k; simpl. tauto. + destruct b ; + rewrite isZ_isZb. unfold is_true. + tauto. + unfold is_true. rewrite eq_true_not_negb_iff. + tauto. + - apply H. +Qed. + + + +Definition eQeval_formula (e : PolEnv R) (k:kind) (ff:eFormula (Formula Q)) : rtyp k := + eval_eformula_k Qeval_formula e k ff. + +Definition eQeval_nformula (e : PolEnv R) (ff:eFormula (NFormula Q)) : Prop := + eval_eformula Qeval_nformula e ff. + + +Lemma eval_eformula_dec : forall (T: Type) eval env (ff: eFormula T), + (forall f, eval env f \/ ~ eval env f) -> + eval_eformula eval env ff \/ ~ eval_eformula eval env ff. +Proof. + intros. destruct ff; simpl; auto. + - destruct b; + rewrite isZ_isZb;destruct (isZb (env x)); intuition congruence. +Qed. + + + +Definition eReval_formula (e : PolEnv R) (k:kind) (ff:eFormula (Formula Rcst)) : rtyp k := + eval_eformula_k Reval_formula e k ff. + + +(*Definition eZeval_nformula (e : PolEnv R) (k:kind) (ff:eFormula (NFormula Z)) : rtyp k := + match ff with + | IsZ _ x => match k with + | isProp => isZ (e x) + | isBool => isZb (e x) + end + | IsF f => Qeval_nformula e k f + end. +*) + +Lemma eval_eformula_map_eFormula : forall env a, + eval_eformula QReval_formula' env (map_eFormula Q_of_Rcst a) <-> eval_eformula Reval_formula' env a. +Proof. + destruct a; simpl. + - tauto. + - unfold QReval_formula'. + rewrite <- eval_formulaSC with (phiS := R_of_Rcst). + + tauto. + + intro. now rewrite Q_of_RcstR. +Qed. + +Lemma eval_cnf_map_cnf : forall (A:Type) eval env f, + eval_cnf (Annot:=A) (eval_eformula eval) env (map_cnf (IsF (Form:=NFormula Q)) f) <-> + eval_cnf (Annot:=A) eval env f. +Proof. + intros. + unfold map_cnf. + unfold eval_cnf. + induction f. + - simpl. tauto. + - + rewrite List.map_cons. + rewrite ! make_conj_cons. + rewrite IHf. + assert (eval_clause (eval_eformula eval) env (ListDef.map (fun '(a0, t) => (IsF a0, t)) a) + <-> eval_clause eval env a). + { clear IHf. + unfold eval_clause. + induction a. + - tauto. + - rewrite List.map_cons. + rewrite ! make_conj_cons. + tauto. + } + tauto. +Qed. + +Lemma xcollect_isZ_nformulaZ : forall cl s cz s' cz', + xcollect_isZ s cz cl = (s', cz') -> + forall f, List.In f cz' -> List.In f cz \/ exists f', List.In (IsF f',tt) cl /\ f = nformulaZ f'. +Proof. + induction cl. + - simpl. intros. inv H. + tauto. + - simpl. destruct a. destruct e. + + intros. + apply IHcl with (f:=f) in H;try easy. + destruct H; [tauto|]. + destruct H as (f' & IN & EQ). + right. exists f'. + tauto. + + intros. + apply IHcl with (f:=f0) in H;try easy. + destruct H. + destruct H. + * right. + exists f. + destruct u. + split. tauto. congruence. + * tauto. + * destruct H as (f' & IN & EQ). + right. exists f'. split;auto. +Qed. + + +Lemma collect_isZ_nformulaZ : forall cl s' cz', + xcollect_isZ PositiveSet.empty nil cl = (s', cz') -> + forall f, List.In f cz' -> exists f', List.In (IsF f',tt) cl /\ f = nformulaZ f'. +Proof. + intros. + apply xcollect_isZ_nformulaZ with (f:=f) in H; try easy. + simpl in H. tauto. +Qed. + + + +Lemma xcollect_isZ_sound : forall cl s cz s' cz', + xcollect_isZ s cz cl = (s', cz') -> + forall x, PositiveSet.mem x s' = true -> + PositiveSet.mem x s = true \/ List.In (IsZ _ true x,tt) cl. +Proof. + induction cl. + - simpl. intros. inv H. tauto. + - simpl. destruct a. destruct e. + + intros. + apply IHcl with (x:= x0) in H;try easy. + destruct H. + destruct b. + rewrite PositiveSetAddon.mem_add in H. + destruct (Pos.eq_dec x0 x). + subst. + right. destruct u. tauto. + tauto. + tauto. + tauto. + + intros. + apply IHcl with (x:=x) in H; tauto. +Qed. + +Lemma collect_isZ_sound : forall cl s' cz', + xcollect_isZ PositiveSet.empty nil cl = (s', cz') -> + forall x, PositiveSet.mem x s' = true -> List.In (IsZ _ true x,tt) cl. +Proof. + intros. + eapply xcollect_isZ_sound with (x:=x) in H; try assumption. + destruct H; [discriminate| assumption]. +Qed. + + +Lemma RTautoChecker_sound : forall f w, RTautoChecker f w = true -> forall env, eval_bf (eReval_formula env) f. Proof. intros f w. unfold RTautoChecker. intros TC env. - apply tauto_checker_sound with (eval:=QReval_formula) (eval':= Qeval_nformula) (env := env) in TC. - - change (eval_f e_rtyp (QReval_formula env)) + eapply tauto_checker_sound with (eval:=eQeval_formula) (eval':= eQeval_nformula) (env := env) in TC. + - change (eval_f e_rtyp (eQeval_formula env)) with - (eval_bf (QReval_formula env)) in TC. + (eval_bf (eQeval_formula env)) in TC. rewrite eval_bf_map in TC. unfold eval_bf in TC. - rewrite eval_f_morph with (ev':= Reval_formula env) in TC ; auto. + rewrite eval_f_morph with (ev':= eReval_formula env) in TC ; auto. intros. apply Tauto.hold_eiff. - rewrite QReval_formula_compat. - unfold QReval_formula'. - rewrite <- eval_formulaSC with (phiS := R_of_Rcst). - + rewrite Reval_formula_compat. - tauto. - + intro. rewrite Q_of_RcstR. reflexivity. - - - apply Reval_nformula_dec. - - destruct t. + unfold eQeval_formula. + rewrite eval_eformula_morph + with (eval2 := QReval_formula') by apply QReval_formula_compat. + unfold eReval_formula. + rewrite eval_eformula_morph + with (eval2 := Reval_formula') by apply Reval_formula_compat. + apply eval_eformula_map_eFormula. + - + intros. + apply eval_eformula_dec. + apply Reval_nformula_dec. + - destruct t as [|t]; try discriminate. + simpl. destruct t. apply (check_inconsistent_sound Rsor QSORaddon) ; auto. - - unfold rdeduce. + - destruct t as [|t]; try discriminate. + unfold erdeduce. intros. revert H. + destruct t' ; try discriminate. + simpl in H0,H1. + destruct (rdeduce t f0) eqn:EQ;try discriminate. + intro. inv H. revert EQ. eapply (nformula_plus_nformula_correct Rsor QSORaddon); eauto. - - intros. - rewrite QReval_formula_compat. - eapply (cnf_normalise_correct Rsor QSORaddon) ; eauto. - - intros. rewrite Tauto.hold_eNOT. rewrite QReval_formula_compat. - now eapply (cnf_negate_correct Rsor QSORaddon); eauto. + unfold eQeval_formula. + rewrite eval_eformula_morph by apply QReval_formula_compat. + destruct t as [| t]; simpl in H; simpl. + + rewrite <- eval_cnf_cons_iff in H. + simpl in H. + unfold eval_tt in H. simpl in H. + destruct H. + destruct b0; simpl in H; + generalize (isZ_dec (env0 x)); tauto. + + unfold eQeval_nformula in H. + rewrite eval_cnf_map_cnf in H. + eapply (cnf_normalise_correct Rsor QSORaddon) ; eauto. + - intros. rewrite Tauto.hold_eNOT. + destruct t as [| t]; simpl in H; simpl. + + rewrite <- eval_cnf_cons_iff in H. + simpl in H. + unfold eval_tt in H. simpl in H. + destruct b,b0; simpl in H; simpl; try tauto; + rewrite isZ_isZb in H;unfold is_true. + tauto. + destruct (isZb (env0 x)); simpl; intuition congruence. + + unfold eQeval_nformula in H. + rewrite eval_cnf_map_cnf in H. + rewrite QReval_formula_compat. + now eapply (cnf_negate_correct Rsor QSORaddon); eauto. - intros t w0. unfold eval_tt. intros. - rewrite make_impl_map with (eval := Qeval_nformula env0). - + eapply RWeakChecker_sound; eauto. + rewrite make_impl_map with (eval := eQeval_nformula env0). + + unfold QCheck in H. + destruct (xcollect_isZ PositiveSet.empty nil t) as (s,cl) eqn:COLLECT. + rewrite <- make_conj_impl. + intro CONJ. + apply ZChecker_sound with (1:=Rsor) (2:= ZSORaddon) (env:=env0) in H. + rewrite <- make_conj_impl in H. + apply H;clear H. + * apply make_conj_in_rev. + intros p IN. + apply collect_isZ_nformulaZ with (f:=p) in COLLECT. + destruct COLLECT as (p' & IN' & EQ). + subst. apply make_conj_in with (p:= IsF p') in CONJ. + simpl in CONJ. + apply nformulaZ_sound;auto. + rewrite List.in_map_iff. + exists (IsF p',tt). + simpl. split;auto. auto. + * unfold isZenv. + intros. simpl in H0. + apply collect_isZ_sound with (x:=x) in COLLECT; auto. + apply make_conj_in with (p:= (IsZ _ true x)) in CONJ. + simpl in CONJ. + now apply isZ_eq. + rewrite List.in_map_iff. + exists (IsZ _ true x, tt). + split ;auto. + tauto. Qed. +Register eFormula as micromega.eFormula.type. +Register IsZ as micromega.eFormula.IsZ. +Register IsF as micromega.eFormula.IsF. (* Local Variables: *) diff --git a/theories/micromega/Refl.v b/theories/micromega/Refl.v index 4041340f9d..4dd0dee2d0 100644 --- a/theories/micromega/Refl.v +++ b/theories/micromega/Refl.v @@ -99,6 +99,20 @@ Proof. * apply IHl; auto. Qed. +Lemma make_conj_in_rev : forall (A : Type) (eval : A -> Prop) (l : list A), + (forall p, In p l -> eval p) -> make_conj eval l. +Proof. + intros A eval l; induction l as [|? l IHl]. + - simpl. + tauto. + - rewrite make_conj_cons. + intros H. + split. apply H. simpl;tauto. + apply IHl. intros; apply H ; simpl;tauto. +Qed. + + + Lemma make_conj_app : forall A eval l1 l2, @make_conj A eval (l1 ++ l2) <-> @make_conj A eval l1 /\ @make_conj A eval l2. Proof. intros A eval l1; induction l1 as [|a l1 IHl1]. diff --git a/theories/micromega/Rify.v b/theories/micromega/Rify.v new file mode 100644 index 0000000000..4dbca0ac96 --- /dev/null +++ b/theories/micromega/Rify.v @@ -0,0 +1,779 @@ +(************************************************************************) +(* * The Rocq Prover / The Rocq Development Team *) +(* v * Copyright INRIA, CNRS and contributors *) +(* isZ x /\ 1 <= x)). + - intros. unfold isZ. split. eexists (Zpos x). reflexivity. + apply IPR_ge_1. +Defined. +Add Tify InjTyp Inj_P_R. + + +#[global] Instance Inj_nat_R : InjTyp nat R. +Proof. + apply (mkinj _ _ INR (fun x => isZ x /\ 0 <= x)). + - intros. unfold isZ. split. + exists (Z.of_nat x). symmetry. apply INR_IZR_INZ. + apply pos_INR. +Defined. +Add Tify InjTyp Inj_nat_R. + +Definition InjNR (n:N) : R := + match n with + | N0 => IZR 0%Z + | Npos p => IZR (Zpos p) + end. + +Lemma IZR_of_N : forall x, IZR (Z.of_N x) = InjNR x. +Proof. + destruct x; simpl; reflexivity. +Qed. + +Lemma InjNR_pos : forall x, + 0 <= InjNR x. +Proof. + destruct x; simpl. + apply Rle_refl. + unfold IZR. + apply Rle_trans with (r2:= R1). + apply Rle_0_1. + apply (IPR_ge_1 p). +Qed. + +#[global] +Instance Inj_N_R : InjTyp N R. +Proof. + apply (mkinj _ _ InjNR (fun x => isZ x /\ 0 <= x)). + - intros. unfold isZ. + split. + exists (Z.of_N x). apply IZR_of_N. apply InjNR_pos. +Defined. +Add Tify InjTyp Inj_N_R. + +#[global] +Instance Inj_R_R : InjTyp R R. +Proof. + apply (mkinj _ _ (fun x => x) (fun x => True) (fun _ => I)). +Defined. + +(** [R -> R -> Prop] : [@eq R], [Rle], [Rge], [Rlt], [Rgt] *) + +#[global] +Instance Op_eqR : BinRel (@eq R). +Proof. +apply mkbrel with (TR := @eq R). +simpl. tauto. +Defined. +Add Tify BinRel Op_eqR. + +#[global] +Instance Op_Rle : BinRel (Rle). +Proof. +apply mkbrel with (TR := Rle). +simpl. tauto. +Defined. +Add Tify BinRel Op_Rle. + +#[global] +Instance Op_Rlt : BinRel (Rlt). +Proof. +apply mkbrel with (TR := Rlt). +simpl. tauto. +Defined. +Add Tify BinRel Op_Rlt. + +#[global] +Instance Op_Rgt : BinRel (Rgt). +Proof. +apply mkbrel with (TR := Rgt). +simpl. tauto. +Defined. +Add Tify BinRel Op_Rgt. + +#[global] +Instance Op_Rge : BinRel (Rge). +Proof. +apply mkbrel with (TR := Rge). +simpl. tauto. +Defined. +Add Tify BinRel Op_Rge. + +(* [R -> R] : [Ropp], [Rabs], [sqrt], [Rinv] *) + +#[global] +Instance Op_Ropp : UnOp Ropp. +Proof. + apply mkuop with (TUOp := Ropp). + simpl. reflexivity. +Defined. +Add Tify UnOp Op_Ropp. + +#[global] +Instance Op_Rabs : UnOp Rabs. +Proof. + apply mkuop with (TUOp := Rabs). + reflexivity. +Defined. +Add Tify UnOp Op_Rabs. + +#[global] +Instance Op_Rinv : UnOp Rinv. +Proof. + apply mkuop with (TUOp := Rinv). + simpl. reflexivity. +Defined. +Add Tify UnOp Op_Rinv. + + +(* [R -> R -> R] : [Rplus] [Rminus] [Rmult] [Rmin] [Rmax] [Rdiv] *) + +#[global] +Instance Op_Rplus : BinOp Rplus. +Proof. + apply mkbop with (TBOp := Rplus). + reflexivity. +Defined. +Add Tify BinOp Op_Rplus. + +#[global] +Instance Op_Rminus : BinOp Rminus. +Proof. + apply mkbop with (TBOp := Rminus). + reflexivity. +Defined. +Add Tify BinOp Op_Rminus. + + +#[global] +Instance Op_Rmult : BinOp Rmult. +Proof. + apply mkbop with (TBOp := Rmult). + reflexivity. +Defined. +Add Tify BinOp Op_Rmult. + +#[global] +Instance Op_Rmin : BinOp Rmin. +Proof. + apply mkbop with (TBOp := Rmin). + reflexivity. +Defined. +Add Tify BinOp Op_Rmin. + +#[global] +Instance Op_Rmax : BinOp Rmax. +Proof. + apply mkbop with (TBOp := Rmax). + reflexivity. +Defined. +Add Tify BinOp Op_Rmax. + +#[global] +Instance Op_div : BinOp Rdiv. +Proof. + apply mkbop with (TBOp := (fun x y => x * / y)). + reflexivity. +Defined. +Add Tify BinOp Op_div. + +(** [Z -> Z -> Prop] : [@eq Z], [Z.le], [Z.gt], *) + +#[global] +Instance Op_eq : BinRel (@eq Z). +Proof. +apply mkbrel with (TR := @eq R). +simpl. split; intro. congruence. +apply eq_IZR; auto. +Defined. +Add Tify BinRel Op_eq. + +#[global] +Instance Op_le : BinRel Z.le. +Proof. + apply mkbrel with (TR := Rle). + split;intros. + apply IZR_le; auto. + apply le_IZR;auto. +Defined. +Add Tify BinRel Op_le. + +#[global] +Instance Op_ge : BinRel Z.ge. +Proof. + apply mkbrel with (TR := Rge). + split;intros. + apply IZR_ge; auto. + apply Z.le_ge. + apply le_IZR;auto. + now apply Rge_le. +Defined. +Add Tify BinRel Op_ge. + + +#[global] +Instance Op_gt : BinRel Z.gt. +Proof. +apply mkbrel with (TR := Rgt). +unfold Rgt. +intros. simpl. split. intro. apply IZR_lt. +rewrite <- Z.gt_lt_iff. auto. +intros. apply lt_IZR in H. rewrite <- Z.gt_lt_iff in H. auto. +Defined. +Add Tify BinRel Op_gt. + +#[global] +Instance Op_lt : BinRel Z.lt. +Proof. +apply mkbrel with (TR := Rlt). +intros. simpl. split. intro. now apply IZR_lt. +intros. now apply lt_IZR in H. +Defined. +Add Tify BinRel Op_lt. + +(* [Z -> R] : [IZR] *) + +#[global] Instance UnOp_IZR : UnOp IZR. +Proof. + apply mkuop with (TUOp:= (fun x => x)). + reflexivity. +Defined. +Add Tify UnOp UnOp_IZR. + +(* [R -> Z] : [up] *) + +Definition Rup (x:R) := IZR (up x). + +#[global] +Instance Op_up : UnOp up. +Proof. +apply mkuop with (TUOp := Rup). +unfold inj,Rup. simpl. +reflexivity. +Defined. +Add Tify UnOp Op_up. + +(* [Z] : Z0 *) + +#[global] +Instance Op_Z0 : CstOp Z0. +Proof. + apply mkcst with (TCst := R0). + simpl. reflexivity. +Defined. +Add Tify CstOp Op_Z0. + +(* [Z -> Z] : [Z.opp] [Z.abs] [Z.succ] [Z.pred] *) + +#[global] + Instance Op_Z_opp : UnOp Z.opp. +Proof. +apply mkuop with (TUOp := Ropp). +apply opp_IZR. +Defined. +Add Tify UnOp Op_Z_opp. + +#[global] + Instance Op_Z_abs : UnOp Z.abs. +Proof. +apply mkuop with (TUOp := Rabs). +apply abs_IZR. +Defined. +Add Tify UnOp Op_Z_abs. + +#[global] + Instance Op_Z_succ : UnOp Z.succ. +Proof. + apply mkuop with (TUOp := Rplus R1). + simpl. intros. + rewrite succ_IZR. + rewrite Rplus_comm. + reflexivity. +Defined. +Add Tify UnOp Op_Z_succ. + +#[global] + Instance Op_Z_pred : UnOp Z.pred. +Proof. + apply mkuop with (TUOp := fun x => x - R1). + simpl. intros. + apply pred_IZR. +Defined. +Add Tify UnOp Op_Z_pred. + +(* [positive -> Z] : [Z.pos] [Z.neg] *) + +#[global] +Instance Op_Zpos : UnOp Zpos. +Proof. + apply mkuop with (TUOp := fun x => x). + simpl. reflexivity. +Defined. +Add Tify UnOp Op_Zpos. + +#[global] +Instance Op_Zneg : UnOp Zneg. +Proof. + apply mkuop with (TUOp := Ropp). + simpl. reflexivity. +Defined. +Add Tify UnOp Op_Zneg. + +(* [Z -> Z -> Z] : [Z.add] [Z.sub] [Z.mul] [Z.max] [Z.min] *) + +#[global] +Instance Op_Zadd : BinOp Z.add. +Proof. + apply mkbop with (TBOp := Rplus). + simpl. intros. rewrite plus_IZR. + reflexivity. +Defined. +Add Tify BinOp Op_Zadd. + +#[global] +Instance Op_Zsub : BinOp Z.sub. +Proof. + apply mkbop with (TBOp := Rminus). + simpl. intros. apply minus_IZR. +Defined. +Add Tify BinOp Op_Zsub. + +#[global] +Instance Op_Zmul : BinOp Z.mul. +Proof. + apply mkbop with (TBOp := Rmult). + simpl. intros. apply mult_IZR. +Defined. +Add Tify BinOp Op_Zmul. + +Lemma max_IZR : forall x y, + IZR (Z.max x y) = Rmax (IZR x) (IZR y). +Proof. + intros. + apply Rmax_case_strong; + intros H; apply le_IZR in H; + f_equal. + now apply Z.max_l. + now apply Z.max_r. +Qed. + +#[global] +Instance Op_Zmax : BinOp Z.max. +Proof. + apply mkbop with (TBOp := Rmax). + simpl. intros. apply max_IZR. +Defined. +Add Tify BinOp Op_Zmax. + +Lemma min_IZR : forall x y, + IZR (Z.min x y) = Rmin (IZR x) (IZR y). +Proof. + intros. + apply Rmin_case_strong; + intros H; apply le_IZR in H; + f_equal. + now apply Z.min_l. + now apply Z.min_r. +Qed. + +#[global] +Instance Op_Zmin : BinOp Z.min. +Proof. + apply mkbop with (TBOp := Rmin). + simpl. intros. apply min_IZR. +Defined. +Add Tify BinOp Op_Zmin. + + +(** minimal support for positive [xH], [xO], [xI], [IPR] *) + +#[global] +Instance Op_xH : CstOp xH. +Proof. + apply mkcst with (TCst := R1). + reflexivity. +Defined. +Add Tify CstOp Op_xH. + +#[global] +Instance Op_xI : UnOp xI. +Proof. + apply mkuop with (TUOp := fun x => 2 * x + 1). + simpl. apply IPR_xI. +Defined. +Add Tify UnOp Op_xI. + +#[global] +Instance Op_xO : UnOp xO. +Proof. + apply mkuop with (TUOp := fun x => 2 * x). + simpl. apply IPR_xO. +Defined. +Add Tify UnOp Op_xO. + +#[global] Instance UnOp_IPR : UnOp IPR. +Proof. + apply mkuop with (TUOp:= (fun x => x)). + reflexivity. +Defined. +Add Tify UnOp UnOp_IPR. + +(* [nat] : O *) + +#[global] +Instance Op_O : CstOp O. +Proof. + apply mkcst with (TCst := R0). + reflexivity. +Defined. +Add Tify CstOp Op_O. + +(* [nat -> nat] : S *) + +#[global] +Instance Op_S : UnOp S. +Proof. + apply mkuop with (TUOp := Rplus R1). + intro. + rewrite S_INR. rewrite Rplus_comm. reflexivity. +Defined. +Add Tify UnOp Op_S. + +#[global] +Instance Op_pred : UnOp Init.Nat.pred. +Proof. + apply mkuop with (TUOp := (fun x => Rmax 0 (x - R1))). + intro. simpl. + destruct x. + - simpl. rewrite Rmax_left. reflexivity. + rewrite Rminus_0_l. + change 0 with (IZR (- 0)). + rewrite Ropp_Ropp_IZR. + apply Ropp_le_contravar. + apply Rle_0_1. + - rewrite S_INR. + simpl. + rewrite Rmax_right. ring. + replace (INR x + 1 - R1) with (INR x) by ring. + apply (pos_INR x). +Defined. +Add Tify UnOp Op_pred. + +(* [Z -> nat] : [Z.to_nat] *) + +Lemma INR_to_nat_Rmax : forall x, + INR (Z.to_nat x) = Rmax (IZR x) 0. +Proof. + intros. + rewrite INR_IZR_INZ. + rewrite ZifyInst.of_nat_to_nat_eq. + now rewrite Z.max_comm, max_IZR. +Qed. + +#[global] +Instance Op_Z_to_nat : UnOp Z.to_nat. +Proof. + apply mkuop with (TUOp := fun x => Rmax x 0). + apply INR_to_nat_Rmax. +Defined. +Add Tify UnOp Op_Z_to_nat. + +(* [nat -> nat -> Prop] : [@eq nat], [le], [ge], [lt], [gt] *) + +#[global] +Instance Op_nat_eq : BinRel (@eq nat). +Proof. + apply mkbrel with (TR := (@eq R)). + split;intros. + congruence. + now apply INR_eq. +Defined. +Add Tify BinRel Op_nat_eq. + + +#[global] +Instance Op_nat_le : BinRel Peano.le. +Proof. + apply mkbrel with (TR := Rle). + split;intros. + simpl. + apply le_INR; auto. + apply INR_le;auto. +Defined. +Add Tify BinRel Op_nat_le. + +#[global] +Instance Op_nat_ge : BinRel Peano.ge. +Proof. + apply mkbrel with (TR := Rge). + split;intros. + simpl. + apply Rle_ge. + apply le_INR; auto. + apply INR_le;auto. + now apply Rge_le. +Defined. +Add Tify BinRel Op_nat_ge. + + +#[global] +Instance Op_nat_lt : BinRel Peano.lt. +Proof. + apply mkbrel with (TR := Rlt). + split;intros. + simpl. + apply lt_INR; auto. + apply INR_lt;auto. +Defined. +Add Tify BinRel Op_nat_lt. + +#[global] +Instance Op_nat_gt : BinRel Peano.gt. +Proof. + apply mkbrel with (TR := Rgt). + split;intros. + simpl. + apply Rlt_gt. + apply lt_INR; auto. + apply INR_lt;auto. +Defined. +Add Tify BinRel Op_nat_gt. + +(** [nat -> nat -> nat] : [add], [mul] *) + +#[global] +Instance Op_nat_add : BinOp Init.Nat.add. +Proof. + apply mkbop with (TBOp := Rplus). + simpl. intros. rewrite plus_INR. + reflexivity. +Defined. +Add Tify BinOp Op_nat_add. + +Lemma sub_INR : forall x y, + INR (x - y) = Rmax 0 ((INR x) - (INR y)). +Proof. + intros. + apply Rmax_case_strong. + - revert y. + induction x. + + intros. reflexivity. + + intros. destruct y. + * rewrite INR_0 in H. + rewrite S_INR in H. + exfalso. + pose proof (pos_INR x) as Px. + pose proof (Rplus_le_compat _ _ _ _ H Px) as A. + apply Rplus_le_compat with (1 := Rle_refl (-(INR x))) in A. + ring_simplify in A. + revert A. apply Rlt_not_le. + apply Rlt_0_1. + * rewrite Nat.sub_succ. + apply IHx. + rewrite !S_INR in H. + now replace (INR x + 1 - (INR y + 1)) with (INR x - INR y) in H by ring. + - intros. rewrite minus_INR. reflexivity. + apply INR_le. + apply Rplus_le_reg_r with (r := - INR y). + rewrite Rplus_opp_r. + apply H. +Qed. + +Instance Op_nat_sub : BinOp Init.Nat.sub. +Proof. + apply mkbop with (TBOp := fun x y => Rmax 0 (Rminus x y)). + simpl. intros. apply sub_INR. +Defined. +Add Tify BinOp Op_nat_sub. + +#[global] +Instance Op_nat_mul : BinOp Init.Nat.mul. +Proof. + apply mkbop with (TBOp := Rmult). + simpl. intros. rewrite mult_INR. + reflexivity. +Defined. +Add Tify BinOp Op_nat_mul. + +(* [nat -> R] : [INR] *) + +#[global] +Instance Op_INR : UnOp INR. +Proof. + apply mkuop with (TUOp := fun x => x). + simpl. reflexivity. +Defined. +Add Tify UnOp Op_INR. + +(** Specification for operators over R *) + + +Lemma Rup_Spec : forall x : R, isZ (Rup x) /\ Rup x > x /\ Rup x - x <= 1%R. +Proof. +intros. +split. unfold isZ, Rup. eexists. reflexivity. +unfold Rup. apply (archimed x). +Qed. + +Instance RupSpec : UnOpSpec Rup. +Proof. apply mkuspec with (UPred := (fun n r => isZ r /\ r > n /\ r - n <= 1)%R). + intros. + split. unfold isZ, Rup. eexists. reflexivity. + unfold Rup. apply (archimed x). +Defined. +Add Tify UnOpSpec RupSpec. + +#[global] +Instance RabsSpec : UnOpSpec Rabs. +Proof. apply mkuspec with (UPred := (fun n r => (n >= 0 -> r = n) /\ (n < 0 -> r = - n)%R)). + split. + apply Rabs_right. + apply Rabs_left. +Defined. +Add Tify UnOpSpec RabsSpec. + +#[global] +Instance RminSpec : BinOpSpec Rmin. +Proof. apply mkbspec with (BPred := (fun x y r => (x <= y -> r = x) /\ (y <= x -> r = y)%R)). + split. + apply Rmin_left. + apply Rmin_right. +Defined. +Add Tify BinOpSpec RminSpec. + +#[global] +Instance RmaxSpec : BinOpSpec Rmax. +Proof. apply mkbspec with (BPred := (fun x y r => (y <= x -> r = x) /\ (x <= y -> r = y)%R)). + split. + apply Rmax_left. + apply Rmax_right. +Defined. +Add Tify BinOpSpec RmaxSpec. + + +#[global] +Instance RinvSpec : UnOpSpec Rinv. +Proof. apply mkuspec with (UPred := (fun x r => (x <> 0 -> r * x = 1)%R)). + intros. now apply Rinv_l. +Defined. +Add Tify UnOpSpec RinvSpec. + +#[global] +Instance RdivSpec : BinOpSpec Rdiv. +Proof. apply mkbspec with (BPred := (fun x y r => (y <> 0 -> r * y = x)%R)). + intros. unfold Rdiv. rewrite Rmult_assoc. + rewrite Rinv_l by assumption. ring. +Defined. +Add Tify BinOpSpec RdivSpec. + + +From Stdlib Require Import Lra Zfloor R_sqrt. + +(** Support for [Zfloor] and [sqrt] *) + +Definition Rfloor (x:R) := IZR (Zfloor x). + +#[global] +Instance Op_Zfloor : UnOp Zfloor. +Proof. +apply mkuop with (TUOp := Rfloor). +unfold inj,Rfloor. simpl. +reflexivity. +Defined. +Add Tify UnOp Op_Zfloor. + + +#[global] +Instance Op_sqrt : UnOp sqrt. +Proof. + apply mkuop with (TUOp := sqrt). + simpl. reflexivity. +Defined. +Add Tify UnOp Op_sqrt. + + + +Lemma sqrtSpec_complete : forall x r, +(( 0 <= r /\ (0 <= x -> r * r = x) /\ (x <= 0 -> r = 0)) <-> r = sqrt x). +Proof. + split;intros. + - destruct H as (H1 & H2 & H3). + assert (C: 0 <= x \/ x <= 0). + { destruct (Rle_gt_dec 0 x). + tauto. right. + apply Rlt_le. now apply Rgt_lt. + } + destruct C as [C | C]. + rewrite <- H2 by assumption. + rewrite sqrt_square;auto. + rewrite sqrt_neg_0;auto. + - subst. + split. + apply sqrt_pos. + split. + intro. apply sqrt_def;auto. + apply sqrt_neg_0. +Qed. + +#[global] +Instance sqrtSpec : UnOpSpec sqrt. +Proof. apply mkuspec with (UPred := (fun x r => ( 0 <= r /\ (0 <= x -> r * r = x) /\ (x <= 0 -> r = 0)%R))). + intros. + rewrite sqrtSpec_complete. reflexivity. +Defined. +Add Tify UnOpSpec sqrtSpec. + +Lemma Rfloor_Spec : forall x : R, isZ (Rfloor x) /\ (Rfloor x <= x < Rfloor x + 1)%R. +Proof. +intros. +split. unfold isZ, Rfloor. eexists. reflexivity. +unfold Rfloor. +apply Zfloor_bound. +Qed. + +Lemma Rfloor_complete: forall x y : R, (isZ y /\ y <= x < y + 1) <-> y = Rfloor x. +Proof. + intros;unfold Rfloor; split;intros. + destruct H as (isZ & C). + destruct isZ. + subst. + f_equal. + symmetry. + apply Zfloor_eq. + auto. + subst. + apply Rfloor_Spec. +Qed. + +#[global] +Instance RfloorSpec : UnOpSpec Rfloor. +Proof. apply mkuspec with (UPred := (fun n r => isZ r /\ r <= n < r+1)%R). + apply Rfloor_Spec. +Defined. +Add Tify UnOpSpec RfloorSpec. diff --git a/theories/micromega/RingMicromega.v b/theories/micromega/RingMicromega.v index d45878799b..bab14ec9d0 100644 --- a/theories/micromega/RingMicromega.v +++ b/theories/micromega/RingMicromega.v @@ -683,30 +683,30 @@ Definition pmul := Pmul cO cI cplus ctimes ceqb. Definition popp := Popp copp. Definition normalise (f : Formula C) : NFormula := -let (lhs, op, rhs) := f in + let (lhs, op, rhs) := f in let lhs := norm lhs in - let rhs := norm rhs in + let rhs := norm rhs in match op with - | OpEq => (psub lhs rhs, Equal) + | OpEq => (psub lhs rhs, Equal) | OpNEq => (psub lhs rhs, NonEqual) - | OpLe => (psub rhs lhs, NonStrict) - | OpGe => (psub lhs rhs, NonStrict) - | OpGt => (psub lhs rhs, Strict) - | OpLt => (psub rhs lhs, Strict) + | OpLe => (psub rhs lhs, NonStrict) + | OpGe => (psub lhs rhs, NonStrict) + | OpGt => (psub lhs rhs, Strict) + | OpLt => (psub rhs lhs, Strict) end. Definition negate (f : Formula C) : NFormula := -let (lhs, op, rhs) := f in + let (lhs, op, rhs) := f in let lhs := norm lhs in - let rhs := norm rhs in - match op with - | OpEq => (psub rhs lhs, NonEqual) - | OpNEq => (psub rhs lhs, Equal) - | OpLe => (psub lhs rhs, Strict) (* e <= e' == ~ e > e' *) - | OpGe => (psub rhs lhs, Strict) - | OpGt => (psub rhs lhs, NonStrict) - | OpLt => (psub lhs rhs, NonStrict) - end. + let rhs := norm rhs in + match op with + | OpEq => (psub rhs lhs, NonEqual) + | OpNEq => (psub rhs lhs, Equal) + | OpLe => (psub lhs rhs, Strict) (* e <= e' == ~ e > e' *) + | OpGe => (psub rhs lhs, Strict) + | OpGt => (psub rhs lhs, NonStrict) + | OpLt => (psub lhs rhs, NonStrict) + end. Lemma eval_pol_sub : forall env lhs rhs, eval_pol env (psub lhs rhs) == eval_pol env lhs - eval_pol env rhs. Proof. diff --git a/theories/micromega/Tify.v b/theories/micromega/Tify.v new file mode 100644 index 0000000000..fb41c90004 --- /dev/null +++ b/theories/micromega/Tify.v @@ -0,0 +1,20 @@ +(************************************************************************) +(* * The Rocq Prover / The Rocq Development Team *) +(* v * Copyright INRIA, CNRS and contributors *) +(* S -> S]. + Another limitation is that our injection theorems e.g. [TBOpInj], + are using Leibniz equality; the payoff is that there is no need for morphisms... + *) + +(** An injection [InjTyp S T] declares an injection + from source type S to target type T. +*) +Class InjTyp (S : Type) (T : Type) := + mkinj { + (* [inj] is the injection function *) + inj : S -> T; + pred : T -> Prop; + (* [cstr] states that [pred] holds for any injected element. + [cstr (inj x)] is introduced in the goal for any leaf + term of the form [inj x] + *) + cstr : forall x, pred (inj x) + }. + +(** [BinOp Op] declares a source operator [Op: S1 -> S2 -> S3]. + *) +Class BinOp {S1 S2 S3 T1 T2 T3:Type} (Op : S1 -> S2 -> S3) {I1 : InjTyp S1 T1} {I2 : InjTyp S2 T2} {I3 : InjTyp S3 T3} := + mkbop { + (* [TBOp] is the target operator after injection of operands. *) + TBOp : T1 -> T2 -> T3; + (* [TBOpInj] states the correctness of the injection. *) + TBOpInj : forall (n:S1) (m:S2), inj (Op n m) = TBOp (inj n) (inj m) + }. + +(** [Unop Op] declares a source operator [Op : S1 -> S2]. *) +Class UnOp {S1 S2 T1 T2:Type} (Op : S1 -> S2) {I1 : InjTyp S1 T1} {I2 : InjTyp S2 T2} := + mkuop { + (* [TUOp] is the target operator after injection of operands. *) + TUOp : T1 -> T2; + (* [TUOpInj] states the correctness of the injection. *) + TUOpInj : forall (x:S1), inj (Op x) = TUOp (inj x) + }. + +(** [CstOp Op] declares a source constant [Op : S]. *) +Class CstOp {S T:Type} (Op : S) {I : InjTyp S T} := + mkcst { + (* [TCst] is the target constant. *) + TCst : T; + (* [TCstInj] states the correctness of the injection. *) + TCstInj : inj Op = TCst + }. + +(** In the framework, [Prop] is mapped to [Prop] and the injection is phrased in + terms of [=] instead of [<->]. +*) + +(** [BinRel R] declares the injection of a binary relation. *) +Class BinRel {S:Type} {T:Type} (R : S -> S -> Prop) {I : InjTyp S T} := + mkbrel { + TR : T -> T -> Prop; + TRInj : forall n m : S, R n m <-> TR (@inj _ _ I n) (inj m) + }. + +(** [PropOp Op] declares morphisms for [<->]. + This will be used to deal with e.g. [and], [or],... *) + +Class PropOp (Op : Prop -> Prop -> Prop) := + mkprop { + op_iff : forall (p1 p2 q1 q2:Prop), (p1 <-> q1) -> (p2 <-> q2) -> (Op p1 p2 <-> Op q1 q2) + }. + +Class PropUOp (Op : Prop -> Prop) := + mkuprop { + uop_iff : forall (p1 q1 :Prop), (p1 <-> q1) -> (Op p1 <-> Op q1) + }. + + + +(** Once the term is injected, terms can be replaced by their specification. + NB1: The Ltac code is currently limited to (Op: Z -> Z -> Z) + NB2: This is not sufficient to cope with [Z.div] or [Z.mod] + *) +Class BinOpSpec {T1 T2 T3: Type} (Op : T1 -> T2 -> T3) := + mkbspec { + BPred : T1 -> T2 -> T3 -> Prop; + BSpec : forall x y, BPred x y (Op x y) + }. + +Class UnOpSpec {T1 T2: Type} (Op : T1 -> T2) := + mkuspec { + UPred : T1 -> T2 -> Prop; + USpec : forall x, UPred x (Op x) + }. + +(** After injections, e.g. nat -> Z, + the fact that Z.of_nat x * Z.of_nat y is positive is lost. + This information can be recovered using instance of the [Saturate] class. +*) +Class Saturate {T: Type} (Op : T -> T -> T) := + mksat { + (** Given [Op x y], + - [PArg1] is the pre-condition of x + - [PArg2] is the pre-condition of y + - [PRes] is the pos-condition of (Op x y) *) + PArg1 : T -> Prop; + PArg2 : T -> Prop; + PRes : T -> T -> T -> Prop; + (** [SatOk] states the correctness of the reasoning *) + SatOk : forall x y, PArg1 x -> PArg2 y -> PRes x y (Op x y) + }. +(* )Arguments PRes {_ _} _. *) + +(* The [ZifyInst.saturate] iterates over all the instances + and for every pattern of the form + [H1 : PArg1 ?x , H2 : PArg2 ?y , T : context[Op ?x ?y] |- _ ] + [H1 : PArg1 ?x , H2 : PArg2 ?y |- context[Op ?x ?y] ] + asserts (SatOK x y H1 H2) *) + +(** The rest of the file is for internal use by the ML tactic. + There are data-structures and lemmas used to inductively construct + the injected terms. *) + +(** The data-structures [injterm] and [injected_prop] + are used to store source and target expressions together + with a correctness proof. *) + +Record injterm {S T: Type} (I : S -> T) := + mkinjterm { source : S ; target : T ; inj_ok : I source = target}. + +Record injprop := + mkinjprop { + source_prop : Prop ; target_prop : Prop ; + injprop_ok : source_prop <-> target_prop}. + + +(** Lemmas for building rewrite rules. *) + +Definition PropOp_iff (Op : Prop -> Prop -> Prop) := + forall (p1 p2 q1 q2:Prop), (p1 <-> q1) -> (p2 <-> q2) -> (Op p1 p2 <-> Op q1 q2). + +Definition PropUOp_iff (Op : Prop -> Prop) := + forall (p1 q1 :Prop), (p1 <-> q1) -> (Op p1 <-> Op q1). + +Lemma mkapp2 (S1 S2 S3 T1 T2 T3 : Type) (Op : S1 -> S2 -> S3) + (I1 : S1 -> T1) (I2 : S2 -> T2) (I3 : S3 -> T3) + (TBOP : T1 -> T2 -> T3) + (TBOPINJ : forall n m, I3 (Op n m) = TBOP (I1 n) (I2 m)) + (s1 : S1) (t1 : T1) (P1: I1 s1 = t1) + (s2 : S2) (t2 : T2) (P2: I2 s2 = t2): I3 (Op s1 s2) = TBOP t1 t2. +Proof. + subst. apply TBOPINJ. +Qed. + +Lemma mkapp (S1 S2 T1 T2 : Type) (OP : S1 -> S2) + (I1 : S1 -> T1) + (I2 : S2 -> T2) + (TUOP : T1 -> T2) + (TUOPINJ : forall n, I2 (OP n) = TUOP (I1 n)) + (s1: S1) (t1: T1) (P1: I1 s1 = t1): I2 (OP s1) = TUOP t1. +Proof. + subst. apply TUOPINJ. +Qed. + +Lemma mkrel (S T : Type) (R : S -> S -> Prop) + (I : S -> T) + (TR : T -> T -> Prop) + (TRINJ : forall n m : S, R n m <-> TR (I n) (I m)) + (s1 : S) (t1 : T) (P1 : I s1 = t1) + (s2 : S) (t2 : T) (P2 : I s2 = t2): + R s1 s2 <-> TR t1 t2. +Proof. + subst. + apply TRINJ. +Qed. + +(** Hardcoded support and lemma for propositional logic *) + +Lemma and_morph : forall (s1 s2 t1 t2:Prop), s1 <-> t1 -> s2 <-> t2 -> ((s1 /\ s2) <-> (t1 /\ t2)). +Proof. + intros. tauto. +Qed. + +Lemma or_morph : forall (s1 s2 t1 t2:Prop), s1 <-> t1 -> s2 <-> t2 -> ((s1 \/ s2) <-> (t1 \/ t2)). +Proof. + intros. tauto. +Qed. + +Lemma impl_morph : forall (s1 s2 t1 t2:Prop), s1 <-> t1 -> s2 <-> t2 -> ((s1 -> s2) <-> (t1 -> t2)). +Proof. + intros. tauto. +Qed. + +Lemma iff_morph : forall (s1 s2 t1 t2:Prop), s1 <-> t1 -> s2 <-> t2 -> ((s1 <-> s2) <-> (t1 <-> t2)). +Proof. + intros. tauto. +Qed. + +Lemma not_morph : forall (s1 t1:Prop), s1 <-> t1 -> (not s1) <-> (not t1). +Proof. + intros. tauto. +Qed. + +Lemma eq_iff : forall (P Q : Prop), P = Q -> (P <-> Q). +Proof. + intros P Q H. + rewrite H. + apply iff_refl. +Defined. + +Lemma rew_iff (P Q : Prop) (IFF : P <-> Q) : P -> Q. +Proof. + exact (fun H => proj1 IFF H). +Qed. + +Lemma rew_iff_rev (P Q : Prop) (IFF : P <-> Q) : Q -> P. +Proof. + exact (fun H => proj2 IFF H). +Qed. + + + +(** Registering constants for use by the plugin *) +Register eq_iff as TifyClasses.eq_iff. +Register target_prop as TifyClasses.target_prop. +Register mkrel as TifyClasses.mkrel. +Register target as TifyClasses.target. +Register mkapp2 as TifyClasses.mkapp2. +Register mkapp as TifyClasses.mkapp. +Register op_iff as TifyClasses.op_iff. +Register uop_iff as TifyClasses.uop_iff. +Register TR as TifyClasses.TR. +Register TBOp as TifyClasses.TBOp. +Register TUOp as TifyClasses.TUOp. +Register TCst as TifyClasses.TCst. +Register injprop_ok as TifyClasses.injprop_ok. +Register inj_ok as TifyClasses.inj_ok. +Register source as TifyClasses.source. +Register source_prop as TifyClasses.source_prop. +Register inj as TifyClasses.inj. +Register TRInj as TifyClasses.TRInj. +Register TUOpInj as TifyClasses.TUOpInj. +Register not as TifyClasses.not. +Register mkinjterm as TifyClasses.mkinjterm. +Register eq_refl as TifyClasses.eq_refl. +Register eq as TifyClasses.eq. +Register mkinjprop as TifyClasses.mkinjprop. +Register iff_refl as TifyClasses.iff_refl. +Register rew_iff as TifyClasses.rew_iff. +Register rew_iff_rev as TifyClasses.rew_iff_rev. +Register source_prop as TifyClasses.source_prop. +Register injprop_ok as TifyClasses.injprop_ok. +Register iff as TifyClasses.iff. + +Register InjTyp as TifyClasses.InjTyp. +Register BinOp as TifyClasses.BinOp. +Register UnOp as TifyClasses.UnOp. +Register CstOp as TifyClasses.CstOp. +Register BinRel as TifyClasses.BinRel. +Register PropOp as TifyClasses.PropOp. +Register PropUOp as TifyClasses.PropUOp. +Register BinOpSpec as TifyClasses.BinOpSpec. +Register UnOpSpec as TifyClasses.UnOpSpec. +Register Saturate as TifyClasses.Saturate. + + +(** Propositional logic *) +Register and as TifyClasses.and. +Register and_morph as TifyClasses.and_morph. +Register or as TifyClasses.or. +Register or_morph as TifyClasses.or_morph. +Register iff as TifyClasses.iff. +Register iff_morph as TifyClasses.iff_morph. +Register impl_morph as TifyClasses.impl_morph. +Register not as TifyClasses.not. +Register not_morph as TifyClasses.not_morph. +Register True as TifyClasses.True. +Register I as TifyClasses.I. diff --git a/theories/micromega/ZMicromega.v b/theories/micromega/ZMicromega.v index aaf25a5536..f68d3f82d7 100644 --- a/theories/micromega/ZMicromega.v +++ b/theories/micromega/ZMicromega.v @@ -14,18 +14,19 @@ (* *) (************************************************************************) - From Stdlib Require Import List. From Stdlib Require Import Bool. From Stdlib Require Import OrderedRing. From Stdlib Require Import RingMicromega. -From Stdlib Require Import ZArithRing. +From Stdlib Require Import MMicromega. From Stdlib Require Import ZCoeff. From Stdlib Require Import Refl. From Stdlib Require Import BinInt. From Stdlib Require InitialRing. From Stdlib.micromega Require Import Tauto. Local Open Scope Z_scope. +Set Implicit Arguments. + Ltac flatten_bool := repeat match goal with @@ -65,41 +66,42 @@ Proof. { subst x. split; reflexivity. } Qed. +Lemma IPR_eq : forall p, IPR Z 1 Z.add p = Z.pos p. +Proof. + induction p; simpl;auto. + - rewrite IHp. rewrite Pos2Z.add_pos_pos. + now rewrite Pos.add_diag. + - rewrite IHp. rewrite Pos2Z.add_pos_pos. + now rewrite Pos.add_diag. +Qed. - - -From Stdlib Require Import EnvRing. - -Lemma Zsor : SOR 0 1 Z.add Z.mul Z.sub Z.opp (@eq Z) Z.le Z.lt. +Lemma IZR_eq : forall c, IZR Z 0 1 Z.add Z.opp c = c. Proof. - constructor ; intros ; subst; try reflexivity. - - apply InitialRing.Zsth. - - apply InitialRing.Zth. - - auto using Z.le_antisymm. - - eauto using Z.le_trans. - - apply Z.le_neq. - - apply Z.lt_trichotomy. - - apply Z.add_le_mono_l; assumption. - - apply Z.mul_pos_pos ; auto. - - discriminate. + destruct c; simpl;auto. + - now rewrite IPR_eq. + - now rewrite IPR_eq. Qed. + + Lemma ZSORaddon : - SORaddon 0 1 Z.add Z.mul Z.sub Z.opp (@eq Z) Z.le (* ring elements *) - 0%Z 1%Z Z.add Z.mul Z.sub Z.opp (* coefficients *) - Z.eqb Z.leb - (fun x => x) (fun x => x) (pow_N 1 Z.mul). + SORaddon 0 1 Z.add Z.mul Z.sub Z.opp eq Z.le 0 1 Z.add Z.mul Z.sub Z.opp Z.eqb Z.leb (IZR Z 0 1 Z.add Z.opp) (fun x => x) (Ring_theory.pow_N 1 Z.mul). Proof. - constructor. - - constructor ; intros ; try reflexivity. - apply Z.eqb_eq ; auto. - - constructor. - reflexivity. - - intros x y. - rewrite <-Z.eqb_eq. congruence. - - apply Z.leb_le. +(* destruct ZSORaddon;*) constructor;try assumption. + - constructor; + intros; rewrite !IZR_eq; auto. + now apply Z.eqb_eq. + - constructor;intros; auto. + - intros. rewrite !IZR_eq; auto. + now apply Z.eqb_neq. + - intros. rewrite !IZR_eq; auto. + now apply Zbool.Zle_bool_imp_le. Qed. + +From Stdlib Require Import EnvRing. + + Fixpoint Zeval_expr (env : PolEnv Z) (e: PExpr Z) : Z := match e with | PEc c => c @@ -113,7 +115,7 @@ Fixpoint Zeval_expr (env : PolEnv Z) (e: PExpr Z) : Z := Strategy expand [ Zeval_expr ]. -Definition eval_expr := eval_pexpr Z.add Z.mul Z.sub Z.opp (fun x => x) (fun x => x) (pow_N 1 Z.mul). +Definition eval_expr := eval_pexpr Z.add Z.mul Z.sub Z.opp (IZR Z 0 1 Z.add Z.opp) (fun x => x) (pow_N 1 Z.mul). Fixpoint Zeval_const (e: PExpr Z) : option Z := match e with @@ -142,9 +144,11 @@ Proof. rewrite ?IHp, ?Z.mul_assoc; auto using Z.mul_comm, f_equal2. Qed. + Lemma Zeval_expr_compat : forall env e, Zeval_expr env e = eval_expr env e. Proof. intros env e; induction e ; simpl ; try congruence. + - now rewrite IZR_eq. - reflexivity. - rewrite ZNpower. congruence. Qed. @@ -204,7 +208,7 @@ Definition Zeval_formula (env : PolEnv Z) (k: Tauto.kind) (f : Formula Z):= (Zeval_op2 k op) (Zeval_expr env lhs) (Zeval_expr env rhs). Definition Zeval_formula' := - eval_formula Z.add Z.mul Z.sub Z.opp (@eq Z) Z.le Z.lt (fun x => x) (fun x => x) (pow_N 1 Z.mul). + @MMicromega.eval_formula Z 0 1 Z.add Z.mul Z.sub Z.opp (@eq Z) Z.le Z.lt N (fun x => x) (pow_N 1 Z.mul). Lemma Zeval_formula_compat : forall env k f, Tauto.hold k (Zeval_formula env k f) <-> Zeval_formula env Tauto.isProp f. Proof. @@ -231,9 +235,8 @@ Proof. intuition auto using Z.le_ge, Z.ge_le, Z.lt_gt, Z.gt_lt. Qed. - Definition eval_nformula := - eval_nformula 0 Z.add Z.mul (@eq Z) Z.le Z.lt (fun x => x) . + eval_nformula Z 0 1 Z.add Z.mul Z.opp (@eq Z) Z.le Z.lt. Definition Zeval_op1 (o : Op1) : Z -> Prop := match o with @@ -265,7 +268,7 @@ Proof. unfold ZWeakChecker in H. exact H. Qed. - +(* Definition psub := psub Z0 Z.add Z.sub Z.opp Z.eqb. Declare Equivalent Keys psub RingMicromega.psub. @@ -276,10 +279,10 @@ Definition padd := padd Z0 Z.add Z.eqb. Declare Equivalent Keys padd RingMicromega.padd. Definition pmul := pmul 0 1 Z.add Z.mul Z.eqb. - -Definition normZ := norm 0 1 Z.add Z.mul Z.sub Z.opp Z.eqb. + *) +Definition normZ := RingMicromega.norm 0 1 Z.add Z.mul Z.sub Z.opp Z.eqb. Declare Equivalent Keys normZ RingMicromega.norm. - +(* Definition eval_pol := eval_pol Z.add Z.mul (fun x => x). Declare Equivalent Keys eval_pol RingMicromega.eval_pol. @@ -310,9 +313,9 @@ Qed. Definition Zunsat := check_inconsistent 0 Z.eqb Z.leb. -Definition Zdeduce := nformula_plus_nformula 0 Z.add Z.eqb. + *) -Lemma Zunsat_sound : forall f, +(*Lemma Zunsat_sound : forall f, Zunsat f = true -> forall env, eval_nformula env f -> False. Proof. unfold Zunsat. @@ -320,8 +323,10 @@ Proof. destruct f. eapply check_inconsistent_sound with (1 := Zsor) (2 := ZSORaddon) in H; eauto. Qed. +*) + -Definition xnnormalise (t : Formula Z) : NFormula Z := +(*Definition xnnormalise (t : Formula Z) : NFormula Z := let (lhs,o,rhs) := t in let lhs := normZ lhs in let rhs := normZ rhs in @@ -368,7 +373,8 @@ Proof. - apply Z.lt_0_sub. - apply Z.lt_0_sub. Qed. - + *) +(* Definition xnormalise (f: NFormula Z) : list (NFormula Z) := let (e,o) := f in match o with @@ -378,11 +384,12 @@ Definition xnormalise (f: NFormula Z) : list (NFormula Z) := | NonEqual => (e, Equal)::nil end. -Lemma eval_pol_Pc : forall env z, +(*Lemma eval_pol_Pc : forall env z, eval_pol env (Pc z) = z. Proof. reflexivity. Qed. + *) Lemma xnormalise_correct : forall env f, (make_conj (fun x => eval_nformula env x -> False) (xnormalise f)) <-> eval_nformula env f. @@ -397,7 +404,7 @@ Proof. - rewrite le_neg, lt_le_iff. rewrite Z.sub_opp_l, Z.sub_sub_distr. reflexivity. Qed. - +*) From Stdlib.micromega Require Import Tauto. @@ -426,8 +433,9 @@ Proof. unfold F at 1. destruct (Zunsat a) eqn:EQ. + rewrite IHf. - unfold E at 1. - specialize (Zunsat_sound _ EQ env). + assert (eval_nformula env a -> False). + { now eapply Zunsat_sound with (1:= Zsor) (2:=ZSORaddon). } + unfold E at 2. tauto. + rewrite <- eval_cnf_cons_iff. @@ -438,23 +446,83 @@ Proof. tauto. Qed. + +(** If all the variables range over the integers, + we can turn a [e > 0] constraint into [e - 1 >= 0] *) + +Definition xnormalise (f: NFormula Z) : list (NFormula Z) := + let (e,o) := f in + match o with + | Equal => (psub e (Pc 1),NonStrict) :: (psub (Pc (-1)) e, NonStrict) :: nil + | NonStrict => ((psub (Pc (-1)) e,NonStrict)::nil) + | Strict => ((psub (Pc 0)) e, NonStrict)::nil + | NonEqual => (e, Equal)::nil + end. + +Lemma xnormalise_correct : forall env f, + (make_conj (fun x => Zeval_nformula env x -> False) (xnormalise f)) <-> Zeval_nformula env f. +Proof. + intros env f. + destruct f as [e o]; destruct o eqn:Op; cbn - [psub]; + repeat rewrite Zeval_pol_sub; fold Zeval_pol; repeat rewrite Zeval_pol_Pc; + generalize (Zeval_pol env e) as x; intro. + - + apply eq_cnf. + - unfold not. tauto. + - rewrite le_neg. rewrite (Z.sub_0_l x), Z.opp_involutive; reflexivity. + - rewrite le_neg, lt_le_iff. + rewrite Z.sub_opp_l, Z.sub_sub_distr. reflexivity. +Qed. + Definition normalise {T : Type} (t:Formula Z) (tg:T) : cnf (NFormula Z) T := let f := xnnormalise t in if Zunsat f then cnf_ff _ _ else cnf_of_list tg (xnormalise f). +Lemma eval_pol_eq : forall p env, + (RingMicromega.eval_pol Z.add Z.mul (fun x : Z => x) env p) = + (RingMicromega.eval_pol Z.add Z.mul (IZR Z 0 1 Z.add Z.opp) env p). +Proof. + induction p; simpl; intros. + - now rewrite IZR_eq. + - now rewrite IHp. + - now rewrite IHp1,IHp2. +Qed. + +Lemma eval_nformula_compat : forall env f, + eval_nformula env f = Zeval_nformula env f. +Proof. + unfold eval_nformula,Zeval_nformula. + unfold MMicromega.eval_nformula. + destruct f; simpl. + rewrite eval_pol_eq. + reflexivity. +Qed. + Lemma normalise_correct : forall (T: Type) env t (tg:T), eval_cnf eval_nformula env (normalise t tg) <-> Zeval_formula env Tauto.isProp t. Proof. intros T env t tg. - rewrite <- xnnormalise_correct. + rewrite Zeval_formula_compat'. + unfold Zeval_formula'. + rewrite <- xnnormalise_correct with (1:= Zsor) (2:=ZSORaddon). unfold normalise. generalize (xnnormalise t) as f;intro f. destruct (Zunsat f) eqn:U. - - assert (US := Zunsat_sound _ U env). + - + assert (eval_nformula env f -> False). + { now eapply Zunsat_sound with (1:= Zsor) (2:=ZSORaddon). } rewrite eval_cnf_ff. tauto. - rewrite cnf_of_list_correct. - apply xnormalise_correct. + fold eval_nformula. + rewrite eval_nformula_compat. + rewrite <- xnormalise_correct. + induction (xnormalise f). + + simpl. tauto. + + rewrite ! make_conj_cons. + rewrite IHl. + rewrite eval_nformula_compat. + tauto. Qed. @@ -476,9 +544,10 @@ Lemma xnegate_correct : forall env f, (make_conj (fun x => eval_nformula env x -> False) (xnegate f)) <-> ~ eval_nformula env f. Proof. intros env f. - destruct f as [e o]; destruct o eqn:Op; cbn - [psub]; - repeat rewrite eval_pol_sub; fold eval_pol; repeat rewrite eval_pol_Pc; - generalize (eval_pol env e) as x; intro x. + destruct f as [e o]; destruct o eqn:Op;cbn - [psub]; + rewrite <- !eval_pol_eq; + repeat rewrite Zeval_pol_sub; fold Zeval_pol; repeat rewrite Zeval_pol_Pc; + generalize (Zeval_pol env e) as x; intro. - tauto. - rewrite eq_cnf. destruct (Z.eq_decidable x 0);tauto. @@ -490,13 +559,16 @@ Qed. Lemma negate_correct : forall T env t (tg:T), eval_cnf eval_nformula env (negate t tg) <-> ~ Zeval_formula env Tauto.isProp t. Proof. intros T env t tg. - rewrite <- xnnormalise_correct. + rewrite Zeval_formula_compat'. + unfold Zeval_formula'. + rewrite <- xnnormalise_correct with (1:=Zsor) (2:=ZSORaddon). unfold negate. generalize (xnnormalise t) as f;intro f. destruct (Zunsat f) eqn:U. - - assert (US := Zunsat_sound _ U env). - rewrite eval_cnf_tt. - tauto. + - rewrite eval_cnf_tt. + split;auto. + intro. unfold not. + now apply Zunsat_sound with (1:=Zsor) (2:=ZSORaddon). - rewrite cnf_of_list_correct. apply xnegate_correct. Qed. @@ -504,1246 +576,19 @@ Qed. Definition cnfZ (Annot: Type) (TX : Tauto.kind -> Type) (AF : Type) (k: Tauto.kind) (f : TFormula (Formula Z) Annot TX AF k) := rxcnf Zunsat Zdeduce normalise negate true f. -Definition ZweakTautoChecker (w: list ZWitness) (f : BFormula (Formula Z) Tauto.isProp) : bool := - @tauto_checker (Formula Z) (NFormula Z) unit Zunsat Zdeduce normalise negate ZWitness (fun cl => ZWeakChecker (List.map fst cl)) f w. - -(* To get a complete checker, the proof format has to be enriched *) - -From Stdlib Require Import Zdiv. -Local Open Scope Z_scope. - -Definition ceiling (a b:Z) : Z := - let (q,r) := Z.div_eucl a b in - match r with - | Z0 => q - | _ => q + 1 - end. - - -From Stdlib Require Import Znumtheory. - -Lemma Zdivide_ceiling : forall a b, (b | a) -> ceiling a b = Z.div a b. -Proof. - unfold ceiling. - intros a b H. - apply Zdivide_mod in H. - case_eq (Z.div_eucl a b). - intros z z0 H0. - change z with (fst (z,z0)). - rewrite <- H0. - change (fst (Z.div_eucl a b)) with (Z.div a b). - change z0 with (snd (z,z0)). - rewrite <- H0. - change (snd (Z.div_eucl a b)) with (Z.modulo a b). - rewrite H. - reflexivity. -Qed. - -Lemma narrow_interval_lower_bound a b x : - a > 0 -> a * x >= b -> x >= ceiling b a. -Proof. - rewrite !Z.ge_le_iff. - unfold ceiling. - intros Ha H. - generalize (Z_div_mod b a Ha). - destruct (Z.div_eucl b a) as (q,r). intros (->,(H1,H2)). - destruct r as [|r|r]. - - rewrite Z.add_0_r in H. - apply Z.mul_le_mono_pos_l in H; auto with zarith. - - assert (0 < Z.pos r) by easy. - rewrite Z.add_1_r, Z.le_succ_l. - apply Z.mul_lt_mono_pos_l with a. - + auto using Z.gt_lt. - + eapply Z.lt_le_trans. 2: eassumption. - now apply Z.lt_add_pos_r. - - now elim H1. -Qed. - -(** NB: narrow_interval_upper_bound is Zdiv.Zdiv_le_lower_bound *) - -Inductive ZArithProof := -| DoneProof -| RatProof : ZWitness -> ZArithProof -> ZArithProof -| CutProof : ZWitness -> ZArithProof -> ZArithProof -| SplitProof : PolC Z -> ZArithProof -> ZArithProof -> ZArithProof -| EnumProof : ZWitness -> ZWitness -> list ZArithProof -> ZArithProof -| ExProof : positive -> ZArithProof -> ZArithProof -(*ExProof x : exists z t, x = z - t /\ z >= 0 /\ t >= 0 *) -. - - -Register ZArithProof as micromega.ZArithProof.type. -Register DoneProof as micromega.ZArithProof.DoneProof. -Register RatProof as micromega.ZArithProof.RatProof. -Register CutProof as micromega.ZArithProof.CutProof. -Register SplitProof as micromega.ZArithProof.SplitProof. -Register EnumProof as micromega.ZArithProof.EnumProof. -Register ExProof as micromega.ZArithProof.ExProof. - - - -(* In order to compute the 'cut', we need to express a polynomial P as a * Q + b. - - b is the constant - - a is the gcd of the other coefficient. -*) -From Stdlib Require Import Znumtheory. - -Definition isZ0 (x:Z) := - match x with - | Z0 => true - | _ => false - end. - -Lemma isZ0_0 : forall x, isZ0 x = true <-> x = 0. -Proof. - intros x; destruct x ; simpl ; intuition congruence. -Qed. - -Lemma isZ0_n0 : forall x, isZ0 x = false <-> x <> 0. -Proof. - intros x; destruct x ; simpl ; intuition congruence. -Qed. - -Definition ZgcdM (x y : Z) := Z.max (Z.gcd x y) 1. - - -Fixpoint Zgcd_pol (p : PolC Z) : (Z * Z) := - match p with - | Pc c => (0,c) - | Pinj _ p => Zgcd_pol p - | PX p _ q => - let (g1,c1) := Zgcd_pol p in - let (g2,c2) := Zgcd_pol q in - (ZgcdM (ZgcdM g1 c1) g2 , c2) - end. - -(*Eval compute in (Zgcd_pol ((PX (Pc (-2)) 1 (Pc 4)))).*) - - -Fixpoint Zdiv_pol (p:PolC Z) (x:Z) : PolC Z := - match p with - | Pc c => Pc (Z.div c x) - | Pinj j p => Pinj j (Zdiv_pol p x) - | PX p j q => PX (Zdiv_pol p x) j (Zdiv_pol q x) - end. - -Inductive Zdivide_pol (x:Z): PolC Z -> Prop := -| Zdiv_Pc : forall c, (x | c) -> Zdivide_pol x (Pc c) -| Zdiv_Pinj : forall p, Zdivide_pol x p -> forall j, Zdivide_pol x (Pinj j p) -| Zdiv_PX : forall p q, Zdivide_pol x p -> Zdivide_pol x q -> forall j, Zdivide_pol x (PX p j q). - - -Lemma Zdiv_pol_correct : forall a p, 0 < a -> Zdivide_pol a p -> - forall env, eval_pol env p = a * eval_pol env (Zdiv_pol p a). -Proof. - intros a p H H0. - induction H0 as [? ?|? ? IHZdivide_pol j|? ? ? IHZdivide_pol1 ? IHZdivide_pol2 j]. - - (* Pc *) - simpl. - intros. - apply Zdivide_Zdiv_eq ; auto. - - (* Pinj *) - simpl. - intros. - apply IHZdivide_pol. - - (* PX *) - simpl. - intros. - rewrite IHZdivide_pol1. - rewrite IHZdivide_pol2. - ring. -Qed. - -Lemma Zgcd_pol_ge : forall p, fst (Zgcd_pol p) >= 0. -Proof. - intros p; induction p as [c|p p1 IHp1|p1 IHp1 ? p3 IHp3]. 1-2: easy. - simpl. - case_eq (Zgcd_pol p1). - case_eq (Zgcd_pol p3). - intros. - simpl. - unfold ZgcdM. - apply Z.le_ge; transitivity 1. - - easy. - - apply Z.le_max_r. -Qed. - -Lemma Zdivide_pol_Zdivide : forall p x y, Zdivide_pol x p -> (y | x) -> Zdivide_pol y p. -Proof. - intros p x y H H0. - induction H. - - constructor. - apply Z.divide_trans with (1:= H0) ; assumption. - - constructor. auto. - - constructor ; auto. -Qed. - -Lemma Zdivide_pol_one : forall p, Zdivide_pol 1 p. -Proof. - intros p; induction p as [c| |]; constructor ; auto. - exists c. ring. -Qed. - -Lemma Zgcd_minus : forall a b c, (a | c - b ) -> (Z.gcd a b | c). -Proof. - intros a b c (q,Hq). - destruct (Zgcd_is_gcd a b) as [(a',Ha) (b',Hb) _]. - set (g:=Z.gcd a b) in *; clearbody g. - exists (q * a' + b'). - symmetry in Hq. rewrite <- Z.add_move_r in Hq. - rewrite <- Hq, Hb, Ha. ring. -Qed. - -Lemma Zdivide_pol_sub : forall p a b, - 0 < Z.gcd a b -> - Zdivide_pol a (PsubC Z.sub p b) -> - Zdivide_pol (Z.gcd a b) p. -Proof. - intros p; induction p as [c|? p IHp|p ? ? ? IHp2]. - - simpl. - intros a b H H0. inversion H0. - constructor. - apply Zgcd_minus ; auto. - - intros ? ? H H0. - constructor. - simpl in H0. inversion H0 ; subst; clear H0. - apply IHp ; auto. - - simpl. intros a b H H0. - inv H0. - constructor. - + apply Zdivide_pol_Zdivide with (1:= (ltac:(assumption) : Zdivide_pol a p)). - destruct (Zgcd_is_gcd a b) ; assumption. - + apply IHp2 ; assumption. -Qed. - -Lemma Zdivide_pol_sub_0 : forall p a, - Zdivide_pol a (PsubC Z.sub p 0) -> - Zdivide_pol a p. -Proof. - intros p; induction p as [c|? p IHp|? IHp1 ? ? IHp2]. - - simpl. - intros ? H. inversion H. - constructor. rewrite Z.sub_0_r in *. assumption. - - intros ? H. - constructor. - simpl in H. inversion H ; subst; clear H. - apply IHp ; auto. - - simpl. intros ? H. - inv H. - constructor. - + auto. - + apply IHp2 ; assumption. -Qed. - - -Lemma Zgcd_pol_div : forall p g c, - Zgcd_pol p = (g, c) -> Zdivide_pol g (PsubC Z.sub p c). -Proof. - intros p; induction p as [c|? ? IHp|p1 IHp1 ? p3 IHp2]; simpl. - - (* Pc *) - intros ? ? H. inv H. - constructor. - exists 0. now ring. - - (* Pinj *) - intros. - constructor. apply IHp ; auto. - - (* PX *) - intros g c. - case_eq (Zgcd_pol p1) ; case_eq (Zgcd_pol p3) ; intros z z0 H z1 z2 H0 H1. - inv H1. - unfold ZgcdM at 1. - destruct (Z.max_spec (Z.gcd (ZgcdM z1 z2) z) 1) as [HH1 | HH1]; cycle 1; - destruct HH1 as [HH1 HH1'] ; rewrite HH1'. - + constructor. - * apply (Zdivide_pol_Zdivide _ (ZgcdM z1 z2)). - -- unfold ZgcdM. - destruct (Z.max_spec (Z.gcd z1 z2) 1) as [HH2 | HH2]; cycle 1. - ++ destruct HH2 as [H1 H2]. - rewrite H2. - apply Zdivide_pol_sub ; auto. - apply Z.lt_le_trans with 1. - ** reflexivity. - ** trivial. - ++ destruct HH2 as [H1 H2]. rewrite H2. - apply Zdivide_pol_one. - -- unfold ZgcdM in HH1. unfold ZgcdM. - destruct (Z.max_spec (Z.gcd z1 z2) 1) as [HH2 | HH2]; cycle 1. - ++ destruct HH2 as [H1 H2]. rewrite H2 in *. - destruct (Zgcd_is_gcd (Z.gcd z1 z2) z); auto. - ++ destruct HH2 as [H1 H2]. rewrite H2. - destruct (Zgcd_is_gcd 1 z); auto. - * apply (Zdivide_pol_Zdivide _ z). - -- apply (IHp2 _ _ H); auto. - -- destruct (Zgcd_is_gcd (ZgcdM z1 z2) z); auto. - + constructor. - * apply Zdivide_pol_one. - * apply Zdivide_pol_one. -Qed. - - - - -Lemma Zgcd_pol_correct_lt : forall p env g c, Zgcd_pol p = (g,c) -> 0 < g -> eval_pol env p = g * (eval_pol env (Zdiv_pol (PsubC Z.sub p c) g)) + c. -Proof. - intros. - rewrite <- Zdiv_pol_correct ; auto. - - rewrite (RingMicromega.PsubC_ok Zsor ZSORaddon). - unfold eval_pol. ring. - (**) - - apply Zgcd_pol_div ; auto. -Qed. - - - -Definition makeCuttingPlane (p : PolC Z) : PolC Z * Z := - let (g,c) := Zgcd_pol p in - if Z.gtb g Z0 - then (Zdiv_pol (PsubC Z.sub p c) g , Z.opp (ceiling (Z.opp c) g)) - else (p,Z0). - - -Definition genCuttingPlane (f : NFormula Z) : option (PolC Z * Z * Op1) := - let (e,op) := f in - match op with - | Equal => let (g,c) := Zgcd_pol e in - if andb (Z.gtb g Z0) (andb (negb (Z.eqb c Z0)) (negb (Z.eqb (Z.gcd g c) g))) - then None (* inconsistent *) - else (* Could be optimised Zgcd_pol is recomputed *) - let (p,c) := makeCuttingPlane e in - Some (p,c,Equal) - | NonEqual => Some (e,Z0,op) - | Strict => let (p,c) := makeCuttingPlane (PsubC Z.sub e 1) in - Some (p,c,NonStrict) - | NonStrict => let (p,c) := makeCuttingPlane e in - Some (p,c,NonStrict) - end. - -Definition nformula_of_cutting_plane (t : PolC Z * Z * Op1) : NFormula Z := - let (e_z, o) := t in - let (e,z) := e_z in - (padd e (Pc z) , o). - -Definition is_pol_Z0 (p : PolC Z) : bool := - match p with - | Pc Z0 => true - | _ => false - end. - -Lemma is_pol_Z0_eval_pol : forall p, is_pol_Z0 p = true -> forall env, eval_pol env p = 0. -Proof. - unfold is_pol_Z0. - intros p; destruct p as [z| |]; try discriminate. - destruct z ; try discriminate. - reflexivity. -Qed. - - -Definition eval_Psatz : list (NFormula Z) -> ZWitness -> option (NFormula Z) := - eval_Psatz 0 1 Z.add Z.mul Z.eqb Z.leb. - - -Definition valid_cut_sign (op:Op1) := - match op with - | Equal => true - | NonStrict => true - | _ => false - end. - - -Definition bound_var (v : positive) : Formula Z := - Build_Formula (PEX v) OpGe (PEc 0). - -Definition mk_eq_pos (x : positive) (y:positive) (t : positive) : Formula Z := - Build_Formula (PEX x) OpEq (PEsub (PEX y) (PEX t)). - - -Fixpoint vars (jmp : positive) (p : Pol Z) : list positive := - match p with - | Pc c => nil - | Pinj j p => vars (Pos.add j jmp) p - | PX p j q => jmp::(vars jmp p)++vars (Pos.succ jmp) q - end. - -Fixpoint max_var (jmp : positive) (p : Pol Z) : positive := - match p with - | Pc _ => jmp - | Pinj j p => max_var (Pos.add j jmp) p - | PX p j q => Pos.max (max_var jmp p) (max_var (Pos.succ jmp) q) - end. - -Lemma pos_le_add : forall y x, - (x <= y + x)%positive. -Proof. - intros y x. - assert ((Z.pos x) <= Z.pos (x + y))%Z as H. - - rewrite <- (Z.add_0_r (Zpos x)). - rewrite <- Pos2Z.add_pos_pos. - apply Z.add_le_mono_l. - compute. congruence. - - rewrite Pos.add_comm in H. - apply H. -Qed. - - -Lemma max_var_le : forall p v, - (v <= max_var v p)%positive. -Proof. - intros p; induction p as [?|p ? IHp|? IHp1 ? ? IHp2]; simpl. - - intros. - apply Pos.le_refl. - - intros v. - specialize (IHp (p+v)%positive). - eapply Pos.le_trans ; eauto. - assert (xH + v <= p + v)%positive. - { apply Pos.add_le_mono. - - apply Pos.le_1_l. - - apply Pos.le_refl. - } - eapply Pos.le_trans ; eauto. - apply pos_le_add. - - intros v. - apply Pos.max_case_strong;intros ; auto. - specialize (IHp2 (Pos.succ v)%positive). - eapply Pos.le_trans ; eauto. -Qed. - -Lemma max_var_correct : forall p j v, - In v (vars j p) -> Pos.le v (max_var j p). -Proof. - intros p; induction p; simpl. - - tauto. - - auto. - - intros j v H. - rewrite in_app_iff in H. - destruct H as [H |[ H | H]]. - + subst. - apply Pos.max_case_strong;intros ; auto. - * apply max_var_le. - * eapply Pos.le_trans ; eauto. - apply max_var_le. - + apply Pos.max_case_strong;intros ; auto. - eapply Pos.le_trans ; eauto. - + apply Pos.max_case_strong;intros ; auto. - eapply Pos.le_trans ; eauto. -Qed. - -Definition max_var_nformulae (l : list (NFormula Z)) := - List.fold_left (fun acc f => Pos.max acc (max_var xH (fst f))) l xH. - -Section MaxVar. - - Definition F (acc : positive) (f : Pol Z * Op1) := Pos.max acc (max_var 1 (fst f)). - - Lemma max_var_nformulae_mono_aux : - forall l v acc, - (v <= acc -> - v <= fold_left F l acc)%positive. - Proof. - intros l; induction l as [|a l IHl] ; simpl ; [easy|]. - intros. - apply IHl. - unfold F. - apply Pos.max_case_strong;intros ; auto. - eapply Pos.le_trans ; eauto. - Qed. - - Lemma max_var_nformulae_mono_aux' : - forall l acc acc', - (acc <= acc' -> - fold_left F l acc <= fold_left F l acc')%positive. - Proof. - intros l; induction l as [|a l IHl]; simpl ; [easy|]. - intros. - apply IHl. - unfold F. - apply Pos.max_le_compat_r; auto. - Qed. - - - - - Lemma max_var_nformulae_correct_aux : forall l p o v, - In (p,o) l -> In v (vars xH p) -> Pos.le v (fold_left F l 1)%positive. - Proof. - intros l p o v H H0. - generalize 1%positive as acc. - revert p o v H H0. - induction l as [|a l IHl]. - - simpl. tauto. - - simpl. - intros p o v H H0 ?. - destruct H ; subst. - + unfold F at 2. - simpl. - apply max_var_correct in H0. - apply max_var_nformulae_mono_aux. - apply Pos.max_case_strong;intros ; auto. - eapply Pos.le_trans ; eauto. - + eapply IHl ; eauto. - Qed. - -End MaxVar. - -Lemma max_var_nformalae_correct : forall l p o v, - In (p,o) l -> In v (vars xH p) -> Pos.le v (max_var_nformulae l)%positive. -Proof. - intros l p o v. - apply max_var_nformulae_correct_aux. -Qed. - - -Fixpoint max_var_psatz (w : Psatz Z) : positive := - match w with - | PsatzIn _ n => xH - | PsatzSquare p => max_var xH (Psquare 0 1 Z.add Z.mul Z.eqb p) - | PsatzMulC p w => Pos.max (max_var xH p) (max_var_psatz w) - | PsatzMulE w1 w2 => Pos.max (max_var_psatz w1) (max_var_psatz w2) - | PsatzAdd w1 w2 => Pos.max (max_var_psatz w1) (max_var_psatz w2) - | _ => xH - end. - -Fixpoint max_var_prf (w : ZArithProof) : positive := - match w with - | DoneProof => xH - | RatProof w pf | CutProof w pf => Pos.max (max_var_psatz w) (max_var_prf pf) - | SplitProof p pf1 pf2 => Pos.max (max_var xH p) (Pos.max (max_var_prf pf1) (max_var_prf pf1)) - | EnumProof w1 w2 l => List.fold_left - (fun acc prf => Pos.max acc (max_var_prf prf)) l - (Pos.max (max_var_psatz w1) (max_var_psatz w2)) - | ExProof _ pf => max_var_prf pf - end. - - -Fixpoint ZChecker (l:list (NFormula Z)) (pf : ZArithProof) {struct pf} : bool := - match pf with - | DoneProof => false - | RatProof w pf => - match eval_Psatz l w with - | None => false - | Some f => - if Zunsat f then true - else ZChecker (f::l) pf - end - | CutProof w pf => - match eval_Psatz l w with - | None => false - | Some f => - match genCuttingPlane f with - | None => true - | Some cp => ZChecker (nformula_of_cutting_plane cp::l) pf - end - end - | SplitProof p pf1 pf2 => - match genCuttingPlane (p,NonStrict) , genCuttingPlane (popp p, NonStrict) with - | None , _ | _ , None => false - | Some cp1 , Some cp2 => - ZChecker (nformula_of_cutting_plane cp1::l) pf1 - && - ZChecker (nformula_of_cutting_plane cp2::l) pf2 - end - | ExProof x prf => - let fr := max_var_nformulae l in - if Pos.leb x fr then - let z := Pos.succ fr in - let t := Pos.succ z in - let nfx := xnnormalise (mk_eq_pos x z t) in - let posz := xnnormalise (bound_var z) in - let post := xnnormalise (bound_var t) in - ZChecker (nfx::posz::post::l) prf - else false - | EnumProof w1 w2 pf => - match eval_Psatz l w1 , eval_Psatz l w2 with - | Some f1 , Some f2 => - match genCuttingPlane f1 , genCuttingPlane f2 with - |Some (e1,z1,op1) , Some (e2,z2,op2) => - if (valid_cut_sign op1 && valid_cut_sign op2 && is_pol_Z0 (padd e1 e2)) - then - (fix label (pfs:list ZArithProof) := - fun lb ub => - match pfs with - | nil => if Z.gtb lb ub then true else false - | pf::rsr => andb (ZChecker ((psub e1 (Pc lb), Equal) :: l) pf) (label rsr (Z.add lb 1%Z) ub) - end) pf (Z.opp z1) z2 - else false - | _ , _ => true - end - | _ , _ => false - end -end. - - - -Fixpoint bdepth (pf : ZArithProof) : nat := - match pf with - | DoneProof => O - | RatProof _ p => S (bdepth p) - | CutProof _ p => S (bdepth p) - | SplitProof _ p1 p2 => S (Nat.max (bdepth p1) (bdepth p2)) - | EnumProof _ _ l => S (List.fold_right (fun pf x => Nat.max (bdepth pf) x) O l) - | ExProof _ p => S (bdepth p) - end. - -From Stdlib Require Import PeanoNat Wf_nat. - -Lemma in_bdepth : forall l a b y, In y l -> ltof ZArithProof bdepth y (EnumProof a b l). -Proof. - intros l; induction l as [|a l IHl]. - - (* nil *) - simpl. - tauto. - - (* cons *) - simpl. - intros a0 b y H. - destruct H as [H|H]. - + subst. - unfold ltof. - simpl. - generalize ( (fold_right - (fun (pf : ZArithProof) (x : nat) => Nat.max (bdepth pf) x) 0%nat l)). - intros. - generalize (bdepth y) ; intros. - rewrite Nat.lt_succ_r. apply Nat.le_max_l. - + generalize (IHl a0 b y H). - unfold ltof. - simpl. - generalize ( (fold_right (fun (pf : ZArithProof) (x : nat) => Nat.max (bdepth pf) x) 0%nat - l)). - intros. - eapply Nat.lt_le_trans. - * eassumption. - * rewrite <- Nat.succ_le_mono. - apply Nat.le_max_r. -Qed. - -Lemma ltof_bdepth_split_l : - forall p pf1 pf2, - ltof ZArithProof bdepth pf1 (SplitProof p pf1 pf2). -Proof. - intros. - unfold ltof. simpl. - rewrite Nat.lt_succ_r. - apply Nat.le_max_l. -Qed. - -Lemma ltof_bdepth_split_r : - forall p pf1 pf2, - ltof ZArithProof bdepth pf2 (SplitProof p pf1 pf2). -Proof. - intros. - unfold ltof. simpl. - rewrite Nat.lt_succ_r. - apply Nat.le_max_r. -Qed. - - -Lemma eval_Psatz_sound : forall env w l f', - make_conj (eval_nformula env) l -> - eval_Psatz l w = Some f' -> eval_nformula env f'. -Proof. - intros env w l f' H H0. - apply (fun H => eval_Psatz_Sound Zsor ZSORaddon l _ H w) ; auto. - apply make_conj_in ; auto. -Qed. - -Lemma makeCuttingPlane_ns_sound : forall env e e' c, - eval_nformula env (e, NonStrict) -> - makeCuttingPlane e = (e',c) -> - eval_nformula env (nformula_of_cutting_plane (e', c, NonStrict)). -Proof. - unfold nformula_of_cutting_plane. - unfold eval_nformula. unfold RingMicromega.eval_nformula. - unfold eval_op1. - intros env e e' c H H0. - rewrite (RingMicromega.eval_pol_add Zsor ZSORaddon). - simpl. - (**) - unfold makeCuttingPlane in H0. - revert H0. - case_eq (Zgcd_pol e) ; intros g c0. - case Z.gtb_spec. - - intros H0 H1 H2. - inv H2. - change (RingMicromega.eval_pol Z.add Z.mul (fun x : Z => x)) with eval_pol in *. - apply (Zgcd_pol_correct_lt _ env) in H1. 2: auto using Z.gt_lt. - apply Z.le_add_le_sub_l, Z.ge_le; rewrite Z.add_0_r. - apply (narrow_interval_lower_bound g (- c0) (eval_pol env (Zdiv_pol (PsubC Z.sub e c0) g))); auto using Z.lt_gt. - apply Z.le_ge. - rewrite <- Z.sub_0_l. - apply Z.le_sub_le_add_r. - rewrite <- H1. - assumption. - (* g <= 0 *) - - intros H0 H1 H2. inv H2. auto with zarith. -Qed. - -Lemma cutting_plane_sound : forall env f p, - eval_nformula env f -> - genCuttingPlane f = Some p -> - eval_nformula env (nformula_of_cutting_plane p). -Proof. - unfold genCuttingPlane. - intros env f; destruct f as [e op]. - destruct op. - - (* Equal *) - intros p; destruct p as [[e' z] op]. - case_eq (Zgcd_pol e) ; intros g c. - case_eq (Z.gtb g 0 && (negb (Z.eqb c 0) && negb (Z.eqb (Z.gcd g c) g))) ; [discriminate|]. - case_eq (makeCuttingPlane e). - intros ? ? H H0 H1 H2 H3. - inv H3. - unfold makeCuttingPlane in H. - rewrite H1 in H. - revert H. - change (eval_pol env e = 0) in H2. - case_eq (Z.gtb g 0). - + intros H H3. - rewrite Z.gtb_lt in H. - rewrite Zgcd_pol_correct_lt with (1:= H1) in H2. 2: auto using Z.gt_lt. - unfold nformula_of_cutting_plane. - change (eval_pol env (padd e' (Pc z)) = 0). - inv H3. - rewrite eval_pol_add. - set (x:=eval_pol env (Zdiv_pol (PsubC Z.sub e c) g)) in *; clearbody x. - simpl. - rewrite andb_false_iff in H0. - destruct H0 as [H0|H0]. - * rewrite <-Z.gtb_lt in H ; congruence. - * rewrite andb_false_iff in H0. - destruct H0 as [H0|H0]. - -- rewrite negb_false_iff in H0. - apply Z.eqb_eq in H0. - subst. simpl. - rewrite Z.add_0_r, Z.mul_eq_0 in H2. - intuition subst; easy. - -- rewrite negb_false_iff in H0. - apply Z.eqb_eq in H0. - rewrite Zdivide_ceiling; cycle 1. - { apply Z.divide_opp_r. rewrite <-H0. apply Z.gcd_divide_r. } - apply Z.sub_move_0_r. - apply Z.div_unique_exact. - ++ now intros ->. - ++ now rewrite Z.add_move_0_r in H2. - + intros H H3. - unfold nformula_of_cutting_plane. - inv H3. - change (eval_pol env (padd e' (Pc 0)) = 0). - rewrite eval_pol_add. - simpl. - now rewrite Z.add_0_r. - - (* NonEqual *) - intros ? H H0. - inv H0. - unfold eval_nformula in *. - unfold RingMicromega.eval_nformula in *. - unfold nformula_of_cutting_plane. - unfold eval_op1 in *. - rewrite (RingMicromega.eval_pol_add Zsor ZSORaddon). - simpl. now rewrite Z.add_0_r. - - (* Strict *) - intros p; destruct p as [[e' z] op]. - case_eq (makeCuttingPlane (PsubC Z.sub e 1)). - intros ? ? H H0 H1. - inv H1. - apply (makeCuttingPlane_ns_sound env) with (2:= H). - simpl in *. - rewrite (RingMicromega.PsubC_ok Zsor ZSORaddon). - now apply Z.lt_le_pred. - - (* NonStrict *) - intros p; destruct p as [[e' z] op]. - case_eq (makeCuttingPlane e). - intros ? ? H H0 H1. - inv H1. - apply (makeCuttingPlane_ns_sound env) with (2:= H). - assumption. -Qed. - -Lemma genCuttingPlaneNone : forall env f, - genCuttingPlane f = None -> - eval_nformula env f -> False. -Proof. - unfold genCuttingPlane. - intros env f; destruct f as [p o]. - destruct o. - - case_eq (Zgcd_pol p) ; intros g c. - case_eq (Z.gtb g 0 && (negb (Z.eqb c 0) && negb (Z.eqb (Z.gcd g c) g))). - + intros H H0 H1 H2. - flatten_bool. - match goal with [ H' : (g >? 0) = true |- ?G ] => rename H' into H3 end. - match goal with [ H' : negb (Z.eqb c 0) = true |- ?G ] => rename H' into H end. - match goal with [ H' : negb (Z.eqb (Z.gcd g c) g) = true |- ?G ] => rename H' into H5 end. - rewrite negb_true_iff in H5. - apply Z.eqb_neq in H5. - rewrite Z.gtb_lt in H3. - rewrite negb_true_iff in H. - apply Z.eqb_neq in H. - change (eval_pol env p = 0) in H2. - rewrite Zgcd_pol_correct_lt with (1:= H0) in H2. 2: auto using Z.gt_lt. - set (x:=eval_pol env (Zdiv_pol (PsubC Z.sub p c) g)) in *; clearbody x. - contradict H5. - apply Zis_gcd_gcd. - * apply Z.lt_le_incl; assumption. - * constructor; auto with zarith. - exists (-x). - rewrite Z.mul_opp_l, Z.mul_comm. - now apply Z.add_move_0_l. - (**) - + destruct (makeCuttingPlane p); discriminate. - - discriminate. - - destruct (makeCuttingPlane (PsubC Z.sub p 1)) ; discriminate. - - destruct (makeCuttingPlane p) ; discriminate. -Qed. - -Lemma eval_nformula_mk_eq_pos : forall env x z t, - env x = env z - env t -> - eval_nformula env (xnnormalise (mk_eq_pos x z t)). -Proof. - intros. - rewrite xnnormalise_correct. - simpl. auto. -Qed. - -Lemma eval_nformula_bound_var : forall env x, - env x >= 0 -> - eval_nformula env (xnnormalise (bound_var x)). -Proof. - intros. - rewrite xnnormalise_correct. - simpl. auto. -Qed. - - -Definition agree_env (fr : positive) (env env' : positive -> Z) : Prop := - forall x, Pos.le x fr -> env x = env' x. - -Lemma agree_env_subset : forall v1 v2 env env', - agree_env v1 env env' -> - Pos.le v2 v1 -> - agree_env v2 env env'. -Proof. - unfold agree_env. - intros v1 v2 env env' H ? ? ?. - apply H. - eapply Pos.le_trans ; eauto. -Qed. - +Definition ZTautoChecker (f : BFormula (Formula Z) Tauto.isProp) (w: list ZArithProof): bool := + @tauto_checker (Formula Z) (NFormula Z) unit Zunsat Zdeduce normalise negate ZArithProof (fun cl => ZChecker None (List.map fst cl)) f w. -Lemma agree_env_jump : forall fr j env env', - agree_env (fr + j) env env' -> - agree_env fr (Env.jump j env) (Env.jump j env'). +Lemma isZenv_None : forall env, + isZenv Z 0 1 Z.add Z.opp eq None env. Proof. - intros fr j env env' H. - unfold agree_env ; intro. + intro. + unfold isZenv. intros. - unfold Env.jump. - apply H. - apply Pos.add_le_mono_r; auto. -Qed. - - -Lemma agree_env_tail : forall fr env env', - agree_env (Pos.succ fr) env env' -> - agree_env fr (Env.tail env) (Env.tail env'). -Proof. - intros fr env env' H. - unfold Env.tail. - apply agree_env_jump. - rewrite <- Pos.add_1_r in H. - apply H. -Qed. - - -Lemma max_var_acc : forall p i j, - (max_var (i + j) p = max_var i p + j)%positive. -Proof. - intros p; induction p as [|? ? IHp|? IHp1 ? ? IHp2]; simpl. - - reflexivity. - - intros. - rewrite ! IHp. - rewrite Pos.add_assoc. - reflexivity. - - intros. - rewrite !Pplus_one_succ_l. - rewrite ! IHp1. - rewrite ! IHp2. - rewrite ! Pos.add_assoc. - rewrite <- Pos.add_max_distr_r. - reflexivity. -Qed. - - - -Lemma agree_env_eval_nformula : - forall env env' e - (AGREE : agree_env (max_var xH (fst e)) env env'), - eval_nformula env e <-> eval_nformula env' e. -Proof. - intros env env' e; destruct e as [p o]. - simpl; intros AGREE. - assert ((RingMicromega.eval_pol Z.add Z.mul (fun x : Z => x) env p) - = - (RingMicromega.eval_pol Z.add Z.mul (fun x : Z => x) env' p)) as H. - { - revert env env' AGREE. - generalize xH. - induction p as [?|p ? IHp|? IHp1 ? ? IHp2]; simpl. - - reflexivity. - - intros p1 **. - apply (IHp p1). - apply agree_env_jump. - eapply agree_env_subset; eauto. - rewrite (Pos.add_comm p). - rewrite max_var_acc. - apply Pos.le_refl. - - intros p ? ? AGREE. - f_equal;[f_equal|]. - + apply (IHp1 p). - eapply agree_env_subset; eauto. - apply Pos.le_max_l. - - + f_equal. - unfold Env.hd. - unfold Env.nth. - apply AGREE. - apply Pos.le_1_l. - - - + apply (IHp2 p). - apply agree_env_tail. - eapply agree_env_subset; eauto. - rewrite !Pplus_one_succ_r. - rewrite max_var_acc. - apply Pos.le_max_r. - } - - rewrite H. tauto. -Qed. - -Lemma agree_env_eval_nformulae : - forall env env' l - (AGREE : agree_env (max_var_nformulae l) env env'), - make_conj (eval_nformula env) l <-> - make_conj (eval_nformula env') l. -Proof. - intros env env' l; induction l as [|a l IHl]. - - simpl. tauto. - - intros. - rewrite ! make_conj_cons. - assert (eval_nformula env a <-> eval_nformula env' a) as H. - { - apply agree_env_eval_nformula. - eapply agree_env_subset ; eauto. - unfold max_var_nformulae. - simpl. - rewrite Pos.max_1_l. - apply max_var_nformulae_mono_aux. - apply Pos.le_refl. - } - rewrite H. - apply and_iff_compat_l. - apply IHl. - eapply agree_env_subset ; eauto. - unfold max_var_nformulae. - simpl. - apply max_var_nformulae_mono_aux'. - apply Pos.le_1_l. -Qed. - - -Lemma eq_true_iff_eq : - forall b1 b2 : bool, (b1 = true <-> b2 = true) <-> b1 = b2. -Proof. - intros b1 b2; destruct b1,b2 ; intuition congruence. -Qed. - -Lemma eval_nformula_split : forall env p, - eval_nformula env (p,NonStrict) \/ eval_nformula env (popp p,NonStrict). -Proof. - unfold popp. - simpl. intros. rewrite (eval_pol_opp Zsor ZSORaddon). - rewrite Z.opp_nonneg_nonpos. - apply Z.le_ge_cases. -Qed. - - - - -Lemma ZChecker_sound : forall w l, - ZChecker l w = true -> forall env, make_impl (eval_nformula env) l False. -Proof. - intros w; induction w as [w H] using (well_founded_ind (well_founded_ltof _ bdepth)). - destruct w as [ | w pf | w pf | p pf1 pf2 | w1 w2 pf | x pf]. - - (* DoneProof *) - simpl. discriminate. - - (* RatProof *) - simpl. - intros l. case_eq (eval_Psatz l w) ; [| discriminate]. - intros f Hf. - case_eq (Zunsat f). - + intros H0 ? ?. - apply (checker_nf_sound Zsor ZSORaddon l w). - unfold check_normalised_formulas. unfold eval_Psatz in Hf. rewrite Hf. - unfold Zunsat in H0. assumption. - + intros H0 H1 env. - assert (make_impl (eval_nformula env) (f::l) False) as H2. - { apply H with (2:= H1). - unfold ltof. - simpl. - auto with arith. - } - destruct f. - rewrite <- make_conj_impl in H2. - rewrite make_conj_cons in H2. - rewrite <- make_conj_impl. - intro. - apply H2. - split ; auto. - apply eval_Psatz_sound with (2:= Hf) ; assumption. - - (* CutProof *) - simpl. - intros l. - case_eq (eval_Psatz l w) ; [ | discriminate]. - intros f' Hlc. - case_eq (genCuttingPlane f'). - + intros p H0 H1 env. - assert (make_impl (eval_nformula env) (nformula_of_cutting_plane p::l) False) as H2. - { eapply (H pf) ; auto. - unfold ltof. - simpl. - auto with arith. - } - rewrite <- make_conj_impl in H2. - rewrite make_conj_cons in H2. - rewrite <- make_conj_impl. - intro. - apply H2. - split ; auto. - apply (eval_Psatz_sound env) in Hlc. - * apply cutting_plane_sound with (1:= Hlc) (2:= H0). - * auto. - + (* genCuttingPlane = None *) - intros H0 H1 env. - rewrite <- make_conj_impl. - intros H2. - apply eval_Psatz_sound with (2:= Hlc) in H2. - apply genCuttingPlaneNone with (2:= H2) ; auto. - - (* SplitProof *) - intros l. - cbn - [genCuttingPlane]. - case_eq (genCuttingPlane (p, NonStrict)) ; [| discriminate]. - case_eq (genCuttingPlane (popp p, NonStrict)) ; [| discriminate]. - intros cp1 GCP1 cp2 GCP2 ZC1 env. - flatten_bool. - match goal with [ H' : ZChecker _ pf1 = true |- _ ] => rename H' into H0 end. - match goal with [ H' : ZChecker _ pf2 = true |- _ ] => rename H' into H1 end. - destruct (eval_nformula_split env p). - + apply (fun H' ck => H _ H' _ ck env) in H0. - * rewrite <- make_conj_impl in *. - intro ; apply H0. - rewrite make_conj_cons. split; auto. - apply (cutting_plane_sound _ (p,NonStrict)) ; auto. - * apply ltof_bdepth_split_l. - + apply (fun H' ck => H _ H' _ ck env) in H1. - * rewrite <- make_conj_impl in *. - intro ; apply H1. - rewrite make_conj_cons. split; auto. - apply (cutting_plane_sound _ (popp p,NonStrict)) ; auto. - * apply ltof_bdepth_split_r. - - (* EnumProof *) - intros l. - simpl. - case_eq (eval_Psatz l w1) ; [ | discriminate]. - case_eq (eval_Psatz l w2) ; [ | discriminate]. - intros f1 Hf1 f2 Hf2. - case_eq (genCuttingPlane f2). - + intros p; destruct p as [ [p1 z1] op1]. - case_eq (genCuttingPlane f1). - * intros p; destruct p as [ [p2 z2] op2]. - case_eq (valid_cut_sign op1 && valid_cut_sign op2 && is_pol_Z0 (padd p1 p2)). - -- intros Hcond. - flatten_bool. - match goal with [ H1 : is_pol_Z0 (padd p1 p2) = true |- _ ] => rename H1 into HZ0 end. - match goal with [ H2 : valid_cut_sign op1 = true |- _ ] => rename H2 into Hop1 end. - match goal with [ H3 : valid_cut_sign op2 = true |- _ ] => rename H3 into Hop2 end. - intros HCutL HCutR Hfix env. - (* get the bounds of the enum *) - rewrite <- make_conj_impl. - intro H0. - assert (-z1 <= eval_pol env p1 <= z2) as H1. { - split. - - apply (eval_Psatz_sound env) in Hf2 ; auto. - apply cutting_plane_sound with (1:= Hf2) in HCutR. - unfold nformula_of_cutting_plane in HCutR. - unfold eval_nformula in HCutR. - unfold RingMicromega.eval_nformula in HCutR. - change (RingMicromega.eval_pol Z.add Z.mul (fun x : Z => x)) with eval_pol in HCutR. - unfold eval_op1 in HCutR. - destruct op1 ; simpl in Hop1 ; try discriminate; - rewrite eval_pol_add in HCutR; simpl in HCutR. - + rewrite Z.add_move_0_l in HCutR; rewrite HCutR, Z.opp_involutive; reflexivity. - + now apply Z.le_sub_le_add_r in HCutR. - (**) - - apply (fun H => is_pol_Z0_eval_pol _ H env) in HZ0. - rewrite eval_pol_add, Z.add_move_r, Z.sub_0_l in HZ0. - rewrite HZ0. - apply (eval_Psatz_sound env) in Hf1 ; auto. - apply cutting_plane_sound with (1:= Hf1) in HCutL. - unfold nformula_of_cutting_plane in HCutL. - unfold eval_nformula in HCutL. - unfold RingMicromega.eval_nformula in HCutL. - change (RingMicromega.eval_pol Z.add Z.mul (fun x : Z => x)) with eval_pol in HCutL. - unfold eval_op1 in HCutL. - rewrite eval_pol_add in HCutL. simpl in HCutL. - destruct op2 ; simpl in Hop2 ; try discriminate. - + rewrite Z.add_move_r, Z.sub_0_l in HCutL. - now rewrite HCutL, Z.opp_involutive. - + now rewrite <- Z.le_sub_le_add_l in HCutL. - } - revert Hfix. - match goal with - | |- context[?F pf (-z1) z2 = true] => set (FF := F) - end. - intros Hfix. - assert (HH :forall x, -z1 <= x <= z2 -> exists pr, - (In pr pf /\ - ZChecker ((PsubC Z.sub p1 x,Equal) :: l) pr = true)%Z). { - clear HZ0 Hop1 Hop2 HCutL HCutR H0 H1. - revert Hfix. - generalize (-z1). clear z1. intro z1. - revert z1 z2. - induction pf as [|a pf IHpf];simpl ;intros z1 z2 Hfix x **. - - revert Hfix. - now case (Z.gtb_spec); [ | easy ]; intros LT; elim (Zorder.Zlt_not_le _ _ LT); transitivity x. - - flatten_bool. - match goal with [ H' : _ <= x <= _ |- _ ] => rename H' into H0 end. - match goal with [ H' : FF pf (z1 + 1) z2 = true |- _ ] => rename H' into H2 end. - destruct (ZArith_dec.Z_le_lt_eq_dec _ _ (proj1 H0)) as [ LT | -> ]. - 2: exists a; auto. - rewrite <- Z.le_succ_l in LT. - assert (LE: (Z.succ z1 <= x <= z2)%Z) by intuition. - elim IHpf with (2:=H2) (3:= LE). - + intros x0 ?. - exists x0 ; split;tauto. - + intros until 1. - apply H ; auto. - cbv [ltof] in *. - cbn [bdepth] in *. - eauto using Nat.lt_le_trans, le_n_S, Nat.le_max_r. - } - (*/asser *) - destruct (HH _ H1) as [pr [Hin Hcheker]]. - assert (make_impl (eval_nformula env) ((PsubC Z.sub p1 (eval_pol env p1),Equal) :: l) False) as H2. { - eapply (H pr) ;auto. - apply in_bdepth ; auto. - } - rewrite <- make_conj_impl in H2. - apply H2. - rewrite make_conj_cons. - split ;auto. - unfold eval_nformula. - unfold RingMicromega.eval_nformula. - simpl. - rewrite (RingMicromega.PsubC_ok Zsor ZSORaddon). - unfold eval_pol. ring. - -- discriminate. - * (* No cutting plane *) - intros H0 H1 H2 env. - rewrite <- make_conj_impl. - intros H3. - apply eval_Psatz_sound with (2:= Hf1) in H3. - apply genCuttingPlaneNone with (2:= H3) ; auto. - + (* No Cutting plane (bis) *) - intros H0 H1 env. - rewrite <- make_conj_impl. - intros H2. - apply eval_Psatz_sound with (2:= Hf2) in H2. - apply genCuttingPlaneNone with (2:= H2) ; auto. - - intros l. - unfold ZChecker. - fold ZChecker. - set (fr := (max_var_nformulae l)%positive). - set (z1 := (Pos.succ fr)) in *. - set (t1 := (Pos.succ z1)) in *. - destruct (x <=? fr)%positive eqn:LE ; [|congruence]. - intros H0 env. - set (env':= fun v => if Pos.eqb v z1 - then if Z.leb (env x) 0 then 0 else env x - else if Pos.eqb v t1 - then if Z.leb (env x) 0 then -(env x) else 0 - else env v). - apply (fun H' ck => H _ H' _ ck env') in H0. - + rewrite <- make_conj_impl in *. - intro H1. - rewrite !make_conj_cons in H0. - apply H0 ; repeat split. - * - apply eval_nformula_mk_eq_pos. - unfold env'. - rewrite! Pos.eqb_refl. - replace (x=?z1)%positive with false. - 1:replace (x=?t1)%positive with false. - 1:replace (t1=?z1)%positive with false. - 1:destruct (env x <=? 0); ring. - { unfold t1. - symmetry; apply not_true_iff_false; rewrite Pos.eqb_eq; symmetry; apply Pos.succ_discr. - } - { - unfold t1, z1. - symmetry; apply not_true_iff_false; rewrite Pos.eqb_eq; intros ->. - apply Pos.leb_le, Pos.lt_succ_r in LE; rewrite <-?Pos.succ_lt_mono in *. - pose proof Pos.lt_not_add_l fr 1; rewrite Pos.add_1_r in *; contradiction. - } - { - unfold z1. - symmetry; apply not_true_iff_false; rewrite Pos.eqb_eq; intros ->. - apply Pos.leb_le, Pos.lt_succ_r in LE; rewrite <-?Pos.succ_lt_mono in *. - case (Pos.lt_irrefl _ LE). - } - * - apply eval_nformula_bound_var. - unfold env'. - rewrite! Pos.eqb_refl. - destruct (env x <=? 0) eqn:EQ. - -- compute. congruence. - -- rewrite Z.leb_gt in EQ. - apply Z.ge_le_iff, Z.lt_le_incl; trivial. - * - apply eval_nformula_bound_var. - unfold env'. - rewrite! Pos.eqb_refl. - replace (t1 =? z1)%positive with false. - -- destruct (env x <=? 0) eqn:EQ. - ++ rewrite Z.leb_le in EQ. - apply Z.ge_le_iff. rewrite Z.opp_le_mono, Z.opp_involutive; trivial. - ++ compute; congruence. - -- unfold t1. - clear. - symmetry; apply not_true_iff_false; rewrite Pos.eqb_eq; symmetry; apply Pos.succ_discr. - * - rewrite (agree_env_eval_nformulae _ env') in H1;auto. - unfold agree_env; intros x0 H2. - unfold env'. - replace (x0 =? z1)%positive with false. - 1:replace (x0 =? t1)%positive with false. - 1:reflexivity. - { - unfold t1, z1. - symmetry; apply not_true_iff_false; rewrite Pos.eqb_eq; intros ->. - apply Pos.lt_succ_r in H2; rewrite <-?Pos.succ_lt_mono in *. - pose proof Pos.lt_not_add_l (max_var_nformulae l) 1; rewrite Pos.add_1_r in *; contradiction. - } - { - unfold z1, fr in *. - symmetry; apply not_true_iff_false; rewrite Pos.eqb_eq; intros ->. - apply Pos.lt_succ_r in H2; rewrite <-?Pos.succ_lt_mono in *. - case (Pos.lt_irrefl _ H2). - } - + unfold ltof. - simpl. - apply Nat.lt_succ_diag_r. + exists (env x). + rewrite IZR_eq. reflexivity. Qed. -Definition ZTautoChecker (f : BFormula (Formula Z) Tauto.isProp) (w: list ZArithProof): bool := - @tauto_checker (Formula Z) (NFormula Z) unit Zunsat Zdeduce normalise negate ZArithProof (fun cl => ZChecker (List.map fst cl)) f w. Lemma ZTautoChecker_sound : forall f w, ZTautoChecker f w = true -> forall env, eval_bf (Zeval_formula env) f. Proof. @@ -1770,43 +615,7 @@ Proof. unfold eval_tt. intros H env. rewrite (make_impl_map (eval_nformula env)). - + eapply ZChecker_sound; eauto. + + eapply ZChecker_sound with (1:=Zsor) (2:=ZSORaddon); eauto. + apply isZenv_None. + tauto. Qed. -Fixpoint xhyps_of_pt (base:nat) (acc : list nat) (pt:ZArithProof) : list nat := - match pt with - | DoneProof => acc - | RatProof c pt => xhyps_of_pt (S base ) (xhyps_of_psatz base acc c) pt - | CutProof c pt => xhyps_of_pt (S base ) (xhyps_of_psatz base acc c) pt - | SplitProof p pt1 pt2 => xhyps_of_pt (S base) (xhyps_of_pt (S base) acc pt1) pt2 - | EnumProof c1 c2 l => - let acc := xhyps_of_psatz base (xhyps_of_psatz base acc c2) c1 in - List.fold_left (xhyps_of_pt (S base)) l acc - | ExProof _ pt => xhyps_of_pt (S (S (S base ))) acc pt - end. - -Definition hyps_of_pt (pt : ZArithProof) : list nat := xhyps_of_pt 0 nil pt. - -Open Scope Z_scope. - -(** To ease bindings from ml code **) -Definition make_impl := Refl.make_impl. -Definition make_conj := Refl.make_conj. - -From Stdlib Require VarMap. - -(*Definition varmap_type := VarMap.t Z. *) -Definition env := PolEnv Z. -Definition node := @VarMap.Branch Z. -Definition empty := @VarMap.Empty Z. -Definition leaf := @VarMap.Elt Z. - -Definition coneMember := ZWitness. - -Definition eval := eval_formula. - -#[deprecated(note="Use [prod positive nat]", since="9.0")] -Definition prod_pos_nat := prod positive nat. - -#[deprecated(use=Z.to_N, since="9.0")] -Notation n_of_Z := Z.to_N (only parsing). diff --git a/theories/micromega/Zify.v b/theories/micromega/Zify.v index 378122e8a6..1389596bc0 100644 --- a/theories/micromega/Zify.v +++ b/theories/micromega/Zify.v @@ -8,8 +8,8 @@ (* * (see LICENSE file for the text of the license) *) (************************************************************************) -From Stdlib Require Import ZifyClasses ZifyInst. -Declare ML Module "rocq-runtime.plugins.zify". +From Stdlib Require Import Tify ZifyInst. +From Stdlib Require Import BinInt. (** [zify_pre_hook] and [zify_post_hook] are there to be redefined. *) Ltac zify_pre_hook := idtac. @@ -28,11 +28,6 @@ Ltac zify_to_euclidean_division_equations := end. -Ltac zify := intros; - zify_pre_hook ; - zify_elim_let ; - zify_op ; - (zify_iter_specs) ; - zify_saturate; +Ltac zify := zify_pre_hook ;tify Z; zify_to_euclidean_division_equations ; zify_post_hook. diff --git a/theories/micromega/ZifyBool.v b/theories/micromega/ZifyBool.v index fd1939fdfe..2a1f644de7 100644 --- a/theories/micromega/ZifyBool.v +++ b/theories/micromega/ZifyBool.v @@ -8,33 +8,33 @@ (* * (see LICENSE file for the text of the license) *) (************************************************************************) From Stdlib Require Import Bool BinNat BinInt Znat. -From Stdlib Require Import Zify ZifyClasses. +From Stdlib Require Import Zify TifyClasses. From Stdlib Require Import ZifyInst. Local Open Scope Z_scope. -(* Instances of [ZifyClasses] for dealing with boolean operators. *) +(* Instances of [TifyClasses] for dealing with boolean operators. *) #[global] Instance Inj_bool_bool : InjTyp bool bool := { inj b := b ; pred b := b = true \/ b = false ; cstr b := ltac:(destruct b; tauto) }. -Add Zify InjTyp Inj_bool_bool. +Add Tify InjTyp Inj_bool_bool. (** Boolean operators *) #[global] Instance Op_andb : BinOp andb := { TBOp := andb ; TBOpInj _ _ := eq_refl}. -Add Zify BinOp Op_andb. +Add Tify BinOp Op_andb. #[global] Instance Op_orb : BinOp orb := { TBOp := orb ; TBOpInj _ _ := eq_refl}. -Add Zify BinOp Op_orb. +Add Tify BinOp Op_orb. #[global] Instance Op_implb : BinOp implb := { TBOp := implb; TBOpInj _ _ := eq_refl }. -Add Zify BinOp Op_implb. +Add Tify BinOp Op_implb. Lemma xorb_eq b1 b2 : xorb b1 b2 = andb (orb b1 b2) (negb (eqb b1 b2)). Proof. @@ -44,93 +44,93 @@ Qed. #[global] Instance Op_xorb : BinOp xorb := { TBOp x y := andb (orb x y) (negb (eqb x y)); TBOpInj := xorb_eq }. -Add Zify BinOp Op_xorb. +Add Tify BinOp Op_xorb. #[global] Instance Op_eqb : BinOp eqb := { TBOp := eqb; TBOpInj _ _ := eq_refl }. -Add Zify BinOp Op_eqb. +Add Tify BinOp Op_eqb. #[global] Instance Op_negb : UnOp negb := { TUOp := negb ; TUOpInj _ := eq_refl}. -Add Zify UnOp Op_negb. +Add Tify UnOp Op_negb. #[global] Instance Op_eq_bool : BinRel (@eq bool) := {TR := @eq bool ; TRInj b1 b2 := iff_refl (b1 = b2) }. -Add Zify BinRel Op_eq_bool. +Add Tify BinRel Op_eq_bool. #[global] Instance Op_true : CstOp true := { TCst := true ; TCstInj := eq_refl }. -Add Zify CstOp Op_true. +Add Tify CstOp Op_true. #[global] Instance Op_false : CstOp false := { TCst := false ; TCstInj := eq_refl }. -Add Zify CstOp Op_false. +Add Tify CstOp Op_false. (** Comparison over Z *) #[global] Instance Op_Zeqb : BinOp Z.eqb := { TBOp := Z.eqb ; TBOpInj _ _ := eq_refl }. -Add Zify BinOp Op_Zeqb. +Add Tify BinOp Op_Zeqb. #[global] Instance Op_Zleb : BinOp Z.leb := { TBOp := Z.leb; TBOpInj _ _ := eq_refl }. -Add Zify BinOp Op_Zleb. +Add Tify BinOp Op_Zleb. #[global] Instance Op_Zgeb : BinOp Z.geb := { TBOp := Z.geb; TBOpInj _ _ := eq_refl }. -Add Zify BinOp Op_Zgeb. +Add Tify BinOp Op_Zgeb. #[global] Instance Op_Zltb : BinOp Z.ltb := { TBOp := Z.ltb ; TBOpInj _ _ := eq_refl }. -Add Zify BinOp Op_Zltb. +Add Tify BinOp Op_Zltb. #[global] Instance Op_Zgtb : BinOp Z.gtb := { TBOp := Z.gtb; TBOpInj _ _ := eq_refl }. -Add Zify BinOp Op_Zgtb. +Add Tify BinOp Op_Zgtb. (** Comparison over N *) #[global] Instance Op_Neqb : BinOp N.eqb := { TBOp := Z.eqb; TBOpInj n m := ltac:(now destruct n, m) }. -Add Zify BinOp Op_Neqb. +Add Tify BinOp Op_Neqb. #[global] Instance Op_Nleb : BinOp N.leb := { TBOp := Z.leb; TBOpInj n m := ltac:(now destruct n, m) }. -Add Zify BinOp Op_Nleb. +Add Tify BinOp Op_Nleb. #[global] Instance Op_Nltb : BinOp N.ltb := { TBOp := Z.ltb; TBOpInj n m := ltac:(now destruct n, m) }. -Add Zify BinOp Op_Nltb. +Add Tify BinOp Op_Nltb. (** Comparison over positive *) #[global] Instance Op_Pos_eqb : BinOp Pos.eqb := { TBOp := Z.eqb; TBOpInj _ _ := eq_refl }. -Add Zify BinOp Op_Pos_eqb. +Add Tify BinOp Op_Pos_eqb. #[global] Instance Op_Pos_leb : BinOp Pos.leb := { TBOp := Z.leb; TBOpInj _ _ := eq_refl }. -Add Zify BinOp Op_Pos_leb. +Add Tify BinOp Op_Pos_leb. #[global] Instance Op_Pos_ltb : BinOp Pos.ltb := { TBOp := Z.ltb; TBOpInj _ _ := eq_refl }. -Add Zify BinOp Op_Pos_ltb. +Add Tify BinOp Op_Pos_ltb. (** Comparison over nat *) @@ -161,17 +161,17 @@ Qed. #[global] Instance Op_nat_eqb : BinOp Nat.eqb := { TBOp := Z.eqb; TBOpInj := Z_of_nat_eqb_iff }. -Add Zify BinOp Op_nat_eqb. +Add Tify BinOp Op_nat_eqb. #[global] Instance Op_nat_leb : BinOp Nat.leb := { TBOp := Z.leb; TBOpInj := Z_of_nat_leb_iff }. -Add Zify BinOp Op_nat_leb. +Add Tify BinOp Op_nat_leb. #[global] Instance Op_nat_ltb : BinOp Nat.ltb := { TBOp := Z.ltb; TBOpInj := Z_of_nat_ltb_iff }. -Add Zify BinOp Op_nat_ltb. +Add Tify BinOp Op_nat_ltb. Lemma b2n_b2z x : Z.of_nat (Nat.b2n x) = Z.b2z x. Proof. @@ -181,12 +181,12 @@ Qed. #[global] Instance Op_b2n : UnOp Nat.b2n := { TUOp := Z.b2z; TUOpInj := b2n_b2z }. -Add Zify UnOp Op_b2n. +Add Tify UnOp Op_b2n. #[global] Instance Op_b2z : UnOp Z.b2z := { TUOp := Z.b2z; TUOpInj _ := eq_refl }. -Add Zify UnOp Op_b2z. +Add Tify UnOp Op_b2z. Lemma b2z_spec b : (b = true /\ Z.b2z b = 1) \/ (b = false /\ Z.b2z b = 0). Proof. @@ -197,7 +197,7 @@ Qed. Instance b2zSpec : UnOpSpec Z.b2z := { UPred b r := (b = true /\ r = 1) \/ (b = false /\ r = 0); USpec := b2z_spec }. -Add Zify UnOpSpec b2zSpec. +Add Tify UnOpSpec b2zSpec. Ltac elim_bool_cstr := repeat match goal with diff --git a/theories/micromega/ZifyClasses.v b/theories/micromega/ZifyClasses.v index eace54be57..5cdfcc8c39 100644 --- a/theories/micromega/ZifyClasses.v +++ b/theories/micromega/ZifyClasses.v @@ -1,286 +1,3 @@ -(************************************************************************) -(* * The Rocq Prover / The Rocq Development Team *) -(* v * Copyright INRIA, CNRS and contributors *) -(* S -> S]. - Another limitation is that our injection theorems e.g. [TBOpInj], - are using Leibniz equality; the payoff is that there is no need for morphisms... - *) - -(** An injection [InjTyp S T] declares an injection - from source type S to target type T. -*) -Class InjTyp (S : Type) (T : Type) := - mkinj { - (* [inj] is the injection function *) - inj : S -> T; - pred : T -> Prop; - (* [cstr] states that [pred] holds for any injected element. - [cstr (inj x)] is introduced in the goal for any leaf - term of the form [inj x] - *) - cstr : forall x, pred (inj x) - }. - -(** [BinOp Op] declares a source operator [Op: S1 -> S2 -> S3]. - *) -Class BinOp {S1 S2 S3 T1 T2 T3:Type} (Op : S1 -> S2 -> S3) {I1 : InjTyp S1 T1} {I2 : InjTyp S2 T2} {I3 : InjTyp S3 T3} := - mkbop { - (* [TBOp] is the target operator after injection of operands. *) - TBOp : T1 -> T2 -> T3; - (* [TBOpInj] states the correctness of the injection. *) - TBOpInj : forall (n:S1) (m:S2), inj (Op n m) = TBOp (inj n) (inj m) - }. - -(** [Unop Op] declares a source operator [Op : S1 -> S2]. *) -Class UnOp {S1 S2 T1 T2:Type} (Op : S1 -> S2) {I1 : InjTyp S1 T1} {I2 : InjTyp S2 T2} := - mkuop { - (* [TUOp] is the target operator after injection of operands. *) - TUOp : T1 -> T2; - (* [TUOpInj] states the correctness of the injection. *) - TUOpInj : forall (x:S1), inj (Op x) = TUOp (inj x) - }. - -(** [CstOp Op] declares a source constant [Op : S]. *) -Class CstOp {S T:Type} (Op : S) {I : InjTyp S T} := - mkcst { - (* [TCst] is the target constant. *) - TCst : T; - (* [TCstInj] states the correctness of the injection. *) - TCstInj : inj Op = TCst - }. - -(** In the framework, [Prop] is mapped to [Prop] and the injection is phrased in - terms of [=] instead of [<->]. -*) - -(** [BinRel R] declares the injection of a binary relation. *) -Class BinRel {S:Type} {T:Type} (R : S -> S -> Prop) {I : InjTyp S T} := - mkbrel { - TR : T -> T -> Prop; - TRInj : forall n m : S, R n m <-> TR (@inj _ _ I n) (inj m) - }. - -(** [PropOp Op] declares morphisms for [<->]. - This will be used to deal with e.g. [and], [or],... *) - -Class PropOp (Op : Prop -> Prop -> Prop) := - mkprop { - op_iff : forall (p1 p2 q1 q2:Prop), (p1 <-> q1) -> (p2 <-> q2) -> (Op p1 p2 <-> Op q1 q2) - }. - -Class PropUOp (Op : Prop -> Prop) := - mkuprop { - uop_iff : forall (p1 q1 :Prop), (p1 <-> q1) -> (Op p1 <-> Op q1) - }. - - - -(** Once the term is injected, terms can be replaced by their specification. - NB1: The Ltac code is currently limited to (Op: Z -> Z -> Z) - NB2: This is not sufficient to cope with [Z.div] or [Z.mod] - *) -Class BinOpSpec {T1 T2 T3: Type} (Op : T1 -> T2 -> T3) := - mkbspec { - BPred : T1 -> T2 -> T3 -> Prop; - BSpec : forall x y, BPred x y (Op x y) - }. - -Class UnOpSpec {T1 T2: Type} (Op : T1 -> T2) := - mkuspec { - UPred : T1 -> T2 -> Prop; - USpec : forall x, UPred x (Op x) - }. - -(** After injections, e.g. nat -> Z, - the fact that Z.of_nat x * Z.of_nat y is positive is lost. - This information can be recovered using instance of the [Saturate] class. -*) -Class Saturate {T: Type} (Op : T -> T -> T) := - mksat { - (** Given [Op x y], - - [PArg1] is the pre-condition of x - - [PArg2] is the pre-condition of y - - [PRes] is the pos-condition of (Op x y) *) - PArg1 : T -> Prop; - PArg2 : T -> Prop; - PRes : T -> T -> T -> Prop; - (** [SatOk] states the correctness of the reasoning *) - SatOk : forall x y, PArg1 x -> PArg2 y -> PRes x y (Op x y) - }. -(* )Arguments PRes {_ _} _. *) - -(* The [ZifyInst.saturate] iterates over all the instances - and for every pattern of the form - [H1 : PArg1 ?x , H2 : PArg2 ?y , T : context[Op ?x ?y] |- _ ] - [H1 : PArg1 ?x , H2 : PArg2 ?y |- context[Op ?x ?y] ] - asserts (SatOK x y H1 H2) *) - -(** The rest of the file is for internal use by the ML tactic. - There are data-structures and lemmas used to inductively construct - the injected terms. *) - -(** The data-structures [injterm] and [injected_prop] - are used to store source and target expressions together - with a correctness proof. *) - -Record injterm {S T: Type} (I : S -> T) := - mkinjterm { source : S ; target : T ; inj_ok : I source = target}. - -Record injprop := - mkinjprop { - source_prop : Prop ; target_prop : Prop ; - injprop_ok : source_prop <-> target_prop}. - - -(** Lemmas for building rewrite rules. *) - -Definition PropOp_iff (Op : Prop -> Prop -> Prop) := - forall (p1 p2 q1 q2:Prop), (p1 <-> q1) -> (p2 <-> q2) -> (Op p1 p2 <-> Op q1 q2). - -Definition PropUOp_iff (Op : Prop -> Prop) := - forall (p1 q1 :Prop), (p1 <-> q1) -> (Op p1 <-> Op q1). - -Lemma mkapp2 (S1 S2 S3 T1 T2 T3 : Type) (Op : S1 -> S2 -> S3) - (I1 : S1 -> T1) (I2 : S2 -> T2) (I3 : S3 -> T3) - (TBOP : T1 -> T2 -> T3) - (TBOPINJ : forall n m, I3 (Op n m) = TBOP (I1 n) (I2 m)) - (s1 : S1) (t1 : T1) (P1: I1 s1 = t1) - (s2 : S2) (t2 : T2) (P2: I2 s2 = t2): I3 (Op s1 s2) = TBOP t1 t2. -Proof. - subst. apply TBOPINJ. -Qed. - -Lemma mkapp (S1 S2 T1 T2 : Type) (OP : S1 -> S2) - (I1 : S1 -> T1) - (I2 : S2 -> T2) - (TUOP : T1 -> T2) - (TUOPINJ : forall n, I2 (OP n) = TUOP (I1 n)) - (s1: S1) (t1: T1) (P1: I1 s1 = t1): I2 (OP s1) = TUOP t1. -Proof. - subst. apply TUOPINJ. -Qed. - -Lemma mkrel (S T : Type) (R : S -> S -> Prop) - (I : S -> T) - (TR : T -> T -> Prop) - (TRINJ : forall n m : S, R n m <-> TR (I n) (I m)) - (s1 : S) (t1 : T) (P1 : I s1 = t1) - (s2 : S) (t2 : T) (P2 : I s2 = t2): - R s1 s2 <-> TR t1 t2. -Proof. - subst. - apply TRINJ. -Qed. - -(** Hardcoded support and lemma for propositional logic *) - -Lemma and_morph : forall (s1 s2 t1 t2:Prop), s1 <-> t1 -> s2 <-> t2 -> ((s1 /\ s2) <-> (t1 /\ t2)). -Proof. - intros. tauto. -Qed. - -Lemma or_morph : forall (s1 s2 t1 t2:Prop), s1 <-> t1 -> s2 <-> t2 -> ((s1 \/ s2) <-> (t1 \/ t2)). -Proof. - intros. tauto. -Qed. - -Lemma impl_morph : forall (s1 s2 t1 t2:Prop), s1 <-> t1 -> s2 <-> t2 -> ((s1 -> s2) <-> (t1 -> t2)). -Proof. - intros. tauto. -Qed. - -Lemma iff_morph : forall (s1 s2 t1 t2:Prop), s1 <-> t1 -> s2 <-> t2 -> ((s1 <-> s2) <-> (t1 <-> t2)). -Proof. - intros. tauto. -Qed. - -Lemma not_morph : forall (s1 t1:Prop), s1 <-> t1 -> (not s1) <-> (not t1). -Proof. - intros. tauto. -Qed. - -Lemma eq_iff : forall (P Q : Prop), P = Q -> (P <-> Q). -Proof. - intros P Q H. - rewrite H. - apply iff_refl. -Defined. - -Lemma rew_iff (P Q : Prop) (IFF : P <-> Q) : P -> Q. -Proof. - exact (fun H => proj1 IFF H). -Qed. - -Lemma rew_iff_rev (P Q : Prop) (IFF : P <-> Q) : Q -> P. -Proof. - exact (fun H => proj2 IFF H). -Qed. - - - -(** Registering constants for use by the plugin *) -Register eq_iff as ZifyClasses.eq_iff. -Register target_prop as ZifyClasses.target_prop. -Register mkrel as ZifyClasses.mkrel. -Register target as ZifyClasses.target. -Register mkapp2 as ZifyClasses.mkapp2. -Register mkapp as ZifyClasses.mkapp. -Register op_iff as ZifyClasses.op_iff. -Register uop_iff as ZifyClasses.uop_iff. -Register TR as ZifyClasses.TR. -Register TBOp as ZifyClasses.TBOp. -Register TUOp as ZifyClasses.TUOp. -Register TCst as ZifyClasses.TCst. -Register injprop_ok as ZifyClasses.injprop_ok. -Register inj_ok as ZifyClasses.inj_ok. -Register source as ZifyClasses.source. -Register source_prop as ZifyClasses.source_prop. -Register inj as ZifyClasses.inj. -Register TRInj as ZifyClasses.TRInj. -Register TUOpInj as ZifyClasses.TUOpInj. -Register not as ZifyClasses.not. -Register mkinjterm as ZifyClasses.mkinjterm. -Register eq_refl as ZifyClasses.eq_refl. -Register eq as ZifyClasses.eq. -Register mkinjprop as ZifyClasses.mkinjprop. -Register iff_refl as ZifyClasses.iff_refl. -Register rew_iff as ZifyClasses.rew_iff. -Register rew_iff_rev as ZifyClasses.rew_iff_rev. -Register source_prop as ZifyClasses.source_prop. -Register injprop_ok as ZifyClasses.injprop_ok. -Register iff as ZifyClasses.iff. - -Register InjTyp as ZifyClasses.InjTyp. -Register BinOp as ZifyClasses.BinOp. -Register UnOp as ZifyClasses.UnOp. -Register CstOp as ZifyClasses.CstOp. -Register BinRel as ZifyClasses.BinRel. -Register PropOp as ZifyClasses.PropOp. -Register PropUOp as ZifyClasses.PropUOp. -Register BinOpSpec as ZifyClasses.BinOpSpec. -Register UnOpSpec as ZifyClasses.UnOpSpec. -Register Saturate as ZifyClasses.Saturate. - - -(** Propositional logic *) -Register and as ZifyClasses.and. -Register and_morph as ZifyClasses.and_morph. -Register or as ZifyClasses.or. -Register or_morph as ZifyClasses.or_morph. -Register iff as ZifyClasses.iff. -Register iff_morph as ZifyClasses.iff_morph. -Register impl_morph as ZifyClasses.impl_morph. -Register not as ZifyClasses.not. -Register not_morph as ZifyClasses.not_morph. -Register True as ZifyClasses.True. -Register I as ZifyClasses.I. +From Stdlib Require Export TifyClasses. diff --git a/theories/micromega/ZifyComparison.v b/theories/micromega/ZifyComparison.v index a34d33292a..994d42ad98 100644 --- a/theories/micromega/ZifyComparison.v +++ b/theories/micromega/ZifyComparison.v @@ -9,7 +9,7 @@ (************************************************************************) From Stdlib Require Import Bool BinInt. -From Stdlib Require Import Zify ZifyClasses. +From Stdlib Require Import Zify TifyClasses. From Stdlib Require Import Lia. Local Open Scope Z_scope. @@ -29,7 +29,7 @@ Qed. #[global] Instance Inj_comparison_Z : InjTyp comparison Z := { inj := Z_of_comparison ; pred :=(fun x => -1 <= x <= 1) ; cstr := Z_of_comparison_bound}. -Add Zify InjTyp Inj_comparison_Z. +Add Tify InjTyp Inj_comparison_Z. Definition ZcompareZ (x y : Z) := Z_of_comparison (Z.compare x y). @@ -37,27 +37,27 @@ Definition ZcompareZ (x y : Z) := #[global] Program Instance BinOp_Zcompare : BinOp Z.compare := { TBOp := ZcompareZ }. -Add Zify BinOp BinOp_Zcompare. +Add Tify BinOp BinOp_Zcompare. #[global] Instance Op_eq_comparison : BinRel (@eq comparison) := {TR := @eq Z ; TRInj := ltac:(intros [] []; simpl ; intuition congruence) }. -Add Zify BinRel Op_eq_comparison. +Add Tify BinRel Op_eq_comparison. #[global] Instance Op_Eq : CstOp Eq := { TCst := 0 ; TCstInj := eq_refl }. -Add Zify CstOp Op_Eq. +Add Tify CstOp Op_Eq. #[global] Instance Op_Lt : CstOp Lt := { TCst := -1 ; TCstInj := eq_refl }. -Add Zify CstOp Op_Lt. +Add Tify CstOp Op_Lt. #[global] Instance Op_Gt : CstOp Gt := { TCst := 1 ; TCstInj := eq_refl }. -Add Zify CstOp Op_Gt. +Add Tify CstOp Op_Gt. Lemma Zcompare_spec : forall x y, @@ -86,4 +86,4 @@ Instance ZcompareSpec : BinOpSpec ZcompareZ := /\ (x < y -> r = -1) ; BSpec := Zcompare_spec|}. -Add Zify BinOpSpec ZcompareSpec. +Add Tify BinOpSpec ZcompareSpec. diff --git a/theories/micromega/ZifyInst.v b/theories/micromega/ZifyInst.v index a6f066fe8c..43c514b893 100644 --- a/theories/micromega/ZifyInst.v +++ b/theories/micromega/ZifyInst.v @@ -8,12 +8,12 @@ (* * (see LICENSE file for the text of the license) *) (************************************************************************) -(* Instances of [ZifyClasses] for emulating the existing zify. +(* Instances of [TifyClasses] for emulating the existing zify. Each instance is registered using a Add 'class' 'name_of_instance'. *) From Stdlib Require Import Arith BinInt BinNat Zeven Znat Nnat. -From Stdlib Require Import ZifyClasses. +From Stdlib Require Import TifyClasses. Declare ML Module "rocq-runtime.plugins.zify". Local Open Scope Z_scope. @@ -26,154 +26,154 @@ Ltac refl := #[global] Instance Inj_Z_Z : InjTyp Z Z := mkinj _ _ (fun x => x) (fun x => True ) (fun _ => I). -Add Zify InjTyp Inj_Z_Z. +Add Tify InjTyp Inj_Z_Z. (** Support for nat *) #[global] Instance Inj_nat_Z : InjTyp nat Z := mkinj _ _ Z.of_nat (fun x => 0 <= x ) Nat2Z.is_nonneg. -Add Zify InjTyp Inj_nat_Z. +Add Tify InjTyp Inj_nat_Z. (* zify_nat_rel *) #[global] Instance Op_ge : BinRel ge := { TR := Z.ge; TRInj := Nat2Z.inj_ge }. -Add Zify BinRel Op_ge. +Add Tify BinRel Op_ge. #[global] Instance Op_lt : BinRel lt := { TR := Z.lt; TRInj := Nat2Z.inj_lt }. -Add Zify BinRel Op_lt. +Add Tify BinRel Op_lt. #[global] Instance Op_Nat_lt : BinRel Nat.lt := Op_lt. -Add Zify BinRel Op_Nat_lt. +Add Tify BinRel Op_Nat_lt. #[global] Instance Op_gt : BinRel gt := { TR := Z.gt; TRInj := Nat2Z.inj_gt }. -Add Zify BinRel Op_gt. +Add Tify BinRel Op_gt. #[global] Instance Op_le : BinRel le := { TR := Z.le; TRInj := Nat2Z.inj_le }. -Add Zify BinRel Op_le. +Add Tify BinRel Op_le. #[global] Instance Op_Nat_le : BinRel Nat.le := Op_le. -Add Zify BinRel Op_Nat_le. +Add Tify BinRel Op_Nat_le. #[global] Instance Op_eq_nat : BinRel (@eq nat) := { TR := @eq Z ; TRInj x y := iff_sym (Nat2Z.inj_iff x y) }. -Add Zify BinRel Op_eq_nat. +Add Tify BinRel Op_eq_nat. #[global] Instance Op_Nat_eq : BinRel (Nat.eq) := Op_eq_nat. -Add Zify BinRel Op_Nat_eq. +Add Tify BinRel Op_Nat_eq. (* zify_nat_op *) #[global] Instance Op_plus : BinOp Nat.add := { TBOp := Z.add; TBOpInj := Nat2Z.inj_add }. -Add Zify BinOp Op_plus. +Add Tify BinOp Op_plus. #[global] Instance Op_sub : BinOp Nat.sub := { TBOp n m := Z.max 0 (n - m) ; TBOpInj := Nat2Z.inj_sub_max }. -Add Zify BinOp Op_sub. +Add Tify BinOp Op_sub. #[global] Instance Op_mul : BinOp Nat.mul := { TBOp := Z.mul ; TBOpInj := Nat2Z.inj_mul }. -Add Zify BinOp Op_mul. +Add Tify BinOp Op_mul. #[global] Instance Op_min : BinOp Nat.min := { TBOp := Z.min ; TBOpInj := Nat2Z.inj_min }. -Add Zify BinOp Op_min. +Add Tify BinOp Op_min. #[global] Instance Op_max : BinOp Nat.max := { TBOp := Z.max ; TBOpInj := Nat2Z.inj_max }. -Add Zify BinOp Op_max. +Add Tify BinOp Op_max. #[global] Instance Op_pred : UnOp Nat.pred := { TUOp n := Z.max 0 (n - 1) ; TUOpInj := Nat2Z.inj_pred_max }. -Add Zify UnOp Op_pred. +Add Tify UnOp Op_pred. #[global] Instance Op_S : UnOp S := { TUOp x := Z.add x 1 ; TUOpInj := Nat2Z.inj_succ }. -Add Zify UnOp Op_S. +Add Tify UnOp Op_S. #[global] Instance Op_O : CstOp O := { TCst := Z0 ; TCstInj := eq_refl (Z.of_nat 0) }. -Add Zify CstOp Op_O. +Add Tify CstOp Op_O. #[global] Instance Op_Z_abs_nat : UnOp Z.abs_nat := { TUOp := Z.abs ; TUOpInj := Zabs2Nat.id_abs }. -Add Zify UnOp Op_Z_abs_nat. +Add Tify UnOp Op_Z_abs_nat. #[global] Instance Op_nat_div2 : UnOp Nat.div2 := { TUOp x := x / 2 ; TUOpInj x := ltac:(now rewrite Nat2Z.inj_div2, Z.div2_div) }. -Add Zify UnOp Op_nat_div2. +Add Tify UnOp Op_nat_div2. #[global] Instance Op_nat_double : UnOp Nat.double := {| TUOp := Z.mul 2 ; TUOpInj := Nat2Z.inj_double |}. -Add Zify UnOp Op_nat_double. +Add Tify UnOp Op_nat_double. (** Support for positive *) #[global] Instance Inj_pos_Z : InjTyp positive Z := { inj := Zpos ; pred x := 0 < x ; cstr := Pos2Z.pos_is_pos }. -Add Zify InjTyp Inj_pos_Z. +Add Tify InjTyp Inj_pos_Z. #[global] Instance Op_pos_to_nat : UnOp Pos.to_nat := {TUOp x := x ; TUOpInj := positive_nat_Z}. -Add Zify UnOp Op_pos_to_nat. +Add Tify UnOp Op_pos_to_nat. #[global] Instance Inj_N_Z : InjTyp N Z := mkinj _ _ Z.of_N (fun x => 0 <= x ) N2Z.is_nonneg. -Add Zify InjTyp Inj_N_Z. +Add Tify InjTyp Inj_N_Z. #[global] Instance Op_N_to_nat : UnOp N.to_nat := { TUOp x := x ; TUOpInj := N_nat_Z }. -Add Zify UnOp Op_N_to_nat. +Add Tify UnOp Op_N_to_nat. (* zify_positive_rel *) #[global] Instance Op_pos_ge : BinRel Pos.ge := { TR := Z.ge; TRInj x y := iff_refl (Z.pos x >= Z.pos y) }. -Add Zify BinRel Op_pos_ge. +Add Tify BinRel Op_pos_ge. #[global] Instance Op_pos_lt : BinRel Pos.lt := { TR := Z.lt; TRInj x y := iff_refl (Z.pos x < Z.pos y) }. -Add Zify BinRel Op_pos_lt. +Add Tify BinRel Op_pos_lt. #[global] Instance Op_pos_gt : BinRel Pos.gt := { TR := Z.gt; TRInj x y := iff_refl (Z.pos x > Z.pos y) }. -Add Zify BinRel Op_pos_gt. +Add Tify BinRel Op_pos_gt. #[global] Instance Op_pos_le : BinRel Pos.le := { TR := Z.le; TRInj x y := iff_refl (Z.pos x <= Z.pos y) }. -Add Zify BinRel Op_pos_le. +Add Tify BinRel Op_pos_le. Lemma eq_pos_inj x y : x = y <-> Z.pos x = Z.pos y. Proof. @@ -183,265 +183,265 @@ Qed. #[global] Instance Op_eq_pos : BinRel (@eq positive) := { TR := @eq Z ; TRInj := eq_pos_inj }. -Add Zify BinRel Op_eq_pos. +Add Tify BinRel Op_eq_pos. (* zify_positive_op *) #[global] Instance Op_Z_of_N : UnOp Z.of_N := { TUOp x := x ; TUOpInj x := eq_refl (Z.of_N x) }. -Add Zify UnOp Op_Z_of_N. +Add Tify UnOp Op_Z_of_N. #[global] Instance Op_Z_to_N : UnOp Z.to_N := { TUOp x := Z.max 0 x ; TUOpInj x := ltac:(now destruct x) }. -Add Zify UnOp Op_Z_to_N. +Add Tify UnOp Op_Z_to_N. #[global] Instance Op_Z_neg : UnOp Z.neg := { TUOp := Z.opp ; TUOpInj x := eq_refl (Zneg x) }. -Add Zify UnOp Op_Z_neg. +Add Tify UnOp Op_Z_neg. #[global] Instance Op_Z_pos : UnOp Z.pos := { TUOp x := x ; TUOpInj x := eq_refl (Z.pos x) }. -Add Zify UnOp Op_Z_pos. +Add Tify UnOp Op_Z_pos. #[global] Instance Op_pos_succ : UnOp Pos.succ := { TUOp x := x + 1 ; TUOpInj := Pos2Z.inj_succ }. -Add Zify UnOp Op_pos_succ. +Add Tify UnOp Op_pos_succ. #[global] Instance Op_pos_pred_double : UnOp Pos.pred_double := { TUOp x := 2 * x - 1 ; TUOpInj _ := eq_refl }. -Add Zify UnOp Op_pos_pred_double. +Add Tify UnOp Op_pos_pred_double. #[global] Instance Op_pos_pred : UnOp Pos.pred := { TUOp x := Z.max 1 (x - 1) ; TUOpInj x := ltac:(rewrite <- Pos.sub_1_r; apply Pos2Z.inj_sub_max) }. -Add Zify UnOp Op_pos_pred. +Add Tify UnOp Op_pos_pred. #[global] Instance Op_pos_predN : UnOp Pos.pred_N := { TUOp x := x - 1 ; TUOpInj x := ltac: (now destruct x; rewrite N.pos_pred_spec) }. -Add Zify UnOp Op_pos_predN. +Add Tify UnOp Op_pos_predN. #[global] Instance Op_pos_of_succ_nat : UnOp Pos.of_succ_nat := { TUOp x := x + 1 ; TUOpInj := Zpos_P_of_succ_nat }. -Add Zify UnOp Op_pos_of_succ_nat. +Add Tify UnOp Op_pos_of_succ_nat. #[global] Instance Op_pos_of_nat : UnOp Pos.of_nat := { TUOp x := Z.max 1 x ; TUOpInj x := ltac: (now destruct x; [|rewrite <- Pos.of_nat_succ, <- (Nat2Z.inj_max 1)]) }. -Add Zify UnOp Op_pos_of_nat. +Add Tify UnOp Op_pos_of_nat. #[global] Instance Op_pos_add : BinOp Pos.add := { TBOp := Z.add ; TBOpInj _ _ := eq_refl }. -Add Zify BinOp Op_pos_add. +Add Tify BinOp Op_pos_add. #[global] Instance Op_pos_add_carry : BinOp Pos.add_carry := { TBOp x y := x + y + 1 ; TBOpInj := ltac:(now intros; rewrite Pos.add_carry_spec, Pos2Z.inj_succ) }. -Add Zify BinOp Op_pos_add_carry. +Add Tify BinOp Op_pos_add_carry. #[global] Instance Op_pos_sub : BinOp Pos.sub := { TBOp n m := Z.max 1 (n - m) ; TBOpInj := Pos2Z.inj_sub_max }. -Add Zify BinOp Op_pos_sub. +Add Tify BinOp Op_pos_sub. #[global] Instance Op_pos_mul : BinOp Pos.mul := { TBOp := Z.mul ; TBOpInj _ _ := eq_refl }. -Add Zify BinOp Op_pos_mul. +Add Tify BinOp Op_pos_mul. #[global] Instance Op_pos_min : BinOp Pos.min := { TBOp := Z.min ; TBOpInj := Pos2Z.inj_min }. -Add Zify BinOp Op_pos_min. +Add Tify BinOp Op_pos_min. #[global] Instance Op_pos_max : BinOp Pos.max := { TBOp := Z.max ; TBOpInj := Pos2Z.inj_max }. -Add Zify BinOp Op_pos_max. +Add Tify BinOp Op_pos_max. #[global] Instance Op_pos_pow : BinOp Pos.pow := { TBOp := Z.pow ; TBOpInj := Pos2Z.inj_pow }. -Add Zify BinOp Op_pos_pow. +Add Tify BinOp Op_pos_pow. #[global] Instance Op_pos_square : UnOp Pos.square := { TUOp := Z.square ; TUOpInj := Pos2Z.inj_square }. -Add Zify UnOp Op_pos_square. +Add Tify UnOp Op_pos_square. #[global] Instance Op_Pos_Nsucc_double : UnOp Pos.Nsucc_double := { TUOp x := 2 * x + 1 ; TUOpInj x := ltac:(now destruct x) }. -Add Zify UnOp Op_Pos_Nsucc_double. +Add Tify UnOp Op_Pos_Nsucc_double. #[global] Instance Op_Pos_Ndouble : UnOp Pos.Ndouble := { TUOp x := 2 * x ; TUOpInj x := ltac:(now destruct x) }. -Add Zify UnOp Op_Pos_Ndouble. +Add Tify UnOp Op_Pos_Ndouble. #[global] Instance Op_xO : UnOp xO := { TUOp x := 2 * x ; TUOpInj _ := eq_refl }. -Add Zify UnOp Op_xO. +Add Tify UnOp Op_xO. #[global] Instance Op_xI : UnOp xI := { TUOp x := 2 * x + 1 ; TUOpInj _ := eq_refl }. -Add Zify UnOp Op_xI. +Add Tify UnOp Op_xI. #[global] Instance Op_xH : CstOp xH := { TCst := 1%Z ; TCstInj := eq_refl }. -Add Zify CstOp Op_xH. +Add Tify CstOp Op_xH. #[global] Instance Op_Z_of_nat : UnOp Z.of_nat:= { TUOp x := x ; TUOpInj x := eq_refl (Z.of_nat x) }. -Add Zify UnOp Op_Z_of_nat. +Add Tify UnOp Op_Z_of_nat. (* zify_N_rel *) #[global] Instance Op_N_ge : BinRel N.ge := { TR := Z.ge ; TRInj := N2Z.inj_ge }. -Add Zify BinRel Op_N_ge. +Add Tify BinRel Op_N_ge. #[global] Instance Op_N_lt : BinRel N.lt := { TR := Z.lt ; TRInj := N2Z.inj_lt }. -Add Zify BinRel Op_N_lt. +Add Tify BinRel Op_N_lt. #[global] Instance Op_N_gt : BinRel N.gt := { TR := Z.gt ; TRInj := N2Z.inj_gt }. -Add Zify BinRel Op_N_gt. +Add Tify BinRel Op_N_gt. #[global] Instance Op_N_le : BinRel N.le := { TR := Z.le ; TRInj := N2Z.inj_le }. -Add Zify BinRel Op_N_le. +Add Tify BinRel Op_N_le. #[global] Instance Op_eq_N : BinRel (@eq N) := { TR := @eq Z ; TRInj x y := iff_sym (N2Z.inj_iff x y) }. -Add Zify BinRel Op_eq_N. +Add Tify BinRel Op_eq_N. (* zify_N_op *) #[global] Instance Op_N_N0 : CstOp N0 := { TCst := Z0 ; TCstInj := eq_refl }. -Add Zify CstOp Op_N_N0. +Add Tify CstOp Op_N_N0. #[global] Instance Op_N_Npos : UnOp Npos := { TUOp x := x ; TUOpInj _ := eq_refl }. -Add Zify UnOp Op_N_Npos. +Add Tify UnOp Op_N_Npos. #[global] Instance Op_N_of_nat : UnOp N.of_nat := { TUOp x := x ; TUOpInj := nat_N_Z }. -Add Zify UnOp Op_N_of_nat. +Add Tify UnOp Op_N_of_nat. #[global] Instance Op_Z_abs_N : UnOp Z.abs_N := { TUOp := Z.abs ; TUOpInj := N2Z.inj_abs_N }. -Add Zify UnOp Op_Z_abs_N. +Add Tify UnOp Op_Z_abs_N. #[global] Instance Op_N_pos : UnOp N.pos := { TUOp x := x ; TUOpInj _ := eq_refl }. -Add Zify UnOp Op_N_pos. +Add Tify UnOp Op_N_pos. #[global] Instance Op_N_add : BinOp N.add := { TBOp := Z.add ; TBOpInj := N2Z.inj_add }. -Add Zify BinOp Op_N_add. +Add Tify BinOp Op_N_add. #[global] Instance Op_N_min : BinOp N.min := { TBOp := Z.min ; TBOpInj := N2Z.inj_min }. -Add Zify BinOp Op_N_min. +Add Tify BinOp Op_N_min. #[global] Instance Op_N_max : BinOp N.max := { TBOp := Z.max ; TBOpInj := N2Z.inj_max }. -Add Zify BinOp Op_N_max. +Add Tify BinOp Op_N_max. #[global] Instance Op_N_mul : BinOp N.mul := { TBOp := Z.mul ; TBOpInj := N2Z.inj_mul }. -Add Zify BinOp Op_N_mul. +Add Tify BinOp Op_N_mul. #[global] Instance Op_N_sub : BinOp N.sub := { TBOp x y := Z.max 0 (x - y) ; TBOpInj := N2Z.inj_sub_max }. -Add Zify BinOp Op_N_sub. +Add Tify BinOp Op_N_sub. #[global] Instance Op_N_div : BinOp N.div := { TBOp := Z.div ; TBOpInj := N2Z.inj_div }. -Add Zify BinOp Op_N_div. +Add Tify BinOp Op_N_div. #[global] Instance Op_N_mod : BinOp N.modulo := { TBOp := Z.rem ; TBOpInj := N2Z.inj_rem }. -Add Zify BinOp Op_N_mod. +Add Tify BinOp Op_N_mod. #[global] Instance Op_N_pred : UnOp N.pred := { TUOp x := Z.max 0 (x - 1) ; TUOpInj x := ltac:(rewrite N.pred_sub; apply N2Z.inj_sub_max) }. -Add Zify UnOp Op_N_pred. +Add Tify UnOp Op_N_pred. #[global] Instance Op_N_succ : UnOp N.succ := { TUOp x := x + 1 ; TUOpInj := N2Z.inj_succ }. -Add Zify UnOp Op_N_succ. +Add Tify UnOp Op_N_succ. #[global] Instance Op_N_succ_double : UnOp N.succ_double := { TUOp x := 2 * x + 1 ; TUOpInj x := ltac:(now destruct x) }. -Add Zify UnOp Op_N_succ_double. +Add Tify UnOp Op_N_succ_double. #[global] Instance Op_N_double : UnOp N.double := { TUOp x := 2 * x ; TUOpInj x := ltac:(now destruct x) }. -Add Zify UnOp Op_N_double. +Add Tify UnOp Op_N_double. #[global] Instance Op_N_succ_pos : UnOp N.succ_pos := { TUOp x := x + 1 ; TUOpInj x := ltac:(now destruct x; simpl; [| rewrite Pplus_one_succ_r]) }. -Add Zify UnOp Op_N_succ_pos. +Add Tify UnOp Op_N_succ_pos. #[global] Instance Op_N_div2 : UnOp N.div2 := { TUOp x := x / 2 ; TUOpInj x := ltac:(now rewrite N2Z.inj_div2, Z.div2_div) }. -Add Zify UnOp Op_N_div2. +Add Tify UnOp Op_N_div2. #[global] Instance Op_N_pow : BinOp N.pow := { TBOp := Z.pow ; TBOpInj := N2Z.inj_pow }. -Add Zify BinOp Op_N_pow. +Add Tify BinOp Op_N_pow. #[global] Instance Op_N_square : UnOp N.square := { TUOp x := x * x ; TUOpInj x := ltac:(now rewrite N.square_spec, N2Z.inj_mul) }. -Add Zify UnOp Op_N_square. +Add Tify UnOp Op_N_square. (** Support for Z - injected to itself *) @@ -449,142 +449,142 @@ Add Zify UnOp Op_N_square. #[global] Instance Op_Z_ge : BinRel Z.ge := { TR := Z.ge ; TRInj x y := iff_refl (x>= y) }. -Add Zify BinRel Op_Z_ge. +Add Tify BinRel Op_Z_ge. #[global] Instance Op_Z_lt : BinRel Z.lt := { TR := Z.lt ; TRInj x y := iff_refl (x < y) }. -Add Zify BinRel Op_Z_lt. +Add Tify BinRel Op_Z_lt. #[global] Instance Op_Z_gt : BinRel Z.gt := { TR := Z.gt ;TRInj x y := iff_refl (x > y) }. -Add Zify BinRel Op_Z_gt. +Add Tify BinRel Op_Z_gt. #[global] Instance Op_Z_le : BinRel Z.le := { TR := Z.le ;TRInj x y := iff_refl (x <= y) }. -Add Zify BinRel Op_Z_le. +Add Tify BinRel Op_Z_le. #[global] Instance Op_eqZ : BinRel (@eq Z) := { TR := @eq Z ; TRInj x y := iff_refl (x = y) }. -Add Zify BinRel Op_eqZ. +Add Tify BinRel Op_eqZ. #[global] Instance Op_Z_Z0 : CstOp Z0 := { TCst := Z0 ; TCstInj := eq_refl }. -Add Zify CstOp Op_Z_Z0. +Add Tify CstOp Op_Z_Z0. #[global] Instance Op_Z_add : BinOp Z.add := { TBOp := Z.add ; TBOpInj _ _ := eq_refl }. -Add Zify BinOp Op_Z_add. +Add Tify BinOp Op_Z_add. #[global] Instance Op_Z_min : BinOp Z.min := { TBOp := Z.min ; TBOpInj _ _ := eq_refl }. -Add Zify BinOp Op_Z_min. +Add Tify BinOp Op_Z_min. #[global] Instance Op_Z_max : BinOp Z.max := { TBOp := Z.max ; TBOpInj _ _ := eq_refl }. -Add Zify BinOp Op_Z_max. +Add Tify BinOp Op_Z_max. #[global] Instance Op_Z_mul : BinOp Z.mul := { TBOp := Z.mul ; TBOpInj _ _ := eq_refl }. -Add Zify BinOp Op_Z_mul. +Add Tify BinOp Op_Z_mul. #[global] Instance Op_Z_sub : BinOp Z.sub := { TBOp := Z.sub ; TBOpInj _ _ := eq_refl }. -Add Zify BinOp Op_Z_sub. +Add Tify BinOp Op_Z_sub. #[global] Instance Op_Z_div : BinOp Z.div := { TBOp := Z.div ; TBOpInj _ _ := eq_refl }. -Add Zify BinOp Op_Z_div. +Add Tify BinOp Op_Z_div. #[global] Instance Op_Z_mod : BinOp Z.modulo := { TBOp := Z.modulo ; TBOpInj _ _ := eq_refl }. -Add Zify BinOp Op_Z_mod. +Add Tify BinOp Op_Z_mod. #[global] Instance Op_Z_rem : BinOp Z.rem := { TBOp := Z.rem ; TBOpInj _ _ := eq_refl }. -Add Zify BinOp Op_Z_rem. +Add Tify BinOp Op_Z_rem. #[global] Instance Op_Z_quot : BinOp Z.quot := { TBOp := Z.quot ; TBOpInj _ _ := eq_refl }. -Add Zify BinOp Op_Z_quot. +Add Tify BinOp Op_Z_quot. #[global] Instance Op_Z_succ : UnOp Z.succ := { TUOp x := x + 1 ; TUOpInj _ := eq_refl }. -Add Zify UnOp Op_Z_succ. +Add Tify UnOp Op_Z_succ. #[global] Instance Op_Z_pred : UnOp Z.pred := { TUOp x := x - 1 ; TUOpInj _ := eq_refl }. -Add Zify UnOp Op_Z_pred. +Add Tify UnOp Op_Z_pred. #[global] Instance Op_Z_opp : UnOp Z.opp := { TUOp := Z.opp ; TUOpInj _ := eq_refl }. -Add Zify UnOp Op_Z_opp. +Add Tify UnOp Op_Z_opp. #[global] Instance Op_Z_abs : UnOp Z.abs := { TUOp := Z.abs ; TUOpInj _ := eq_refl }. -Add Zify UnOp Op_Z_abs. +Add Tify UnOp Op_Z_abs. #[global] Instance Op_Z_sgn : UnOp Z.sgn := { TUOp := Z.sgn ; TUOpInj _ := eq_refl }. -Add Zify UnOp Op_Z_sgn. +Add Tify UnOp Op_Z_sgn. #[global] Instance Op_Z_pow : BinOp Z.pow := { TBOp := Z.pow ; TBOpInj _ _ := eq_refl }. -Add Zify BinOp Op_Z_pow. +Add Tify BinOp Op_Z_pow. #[global] Instance Op_Z_pow_pos : BinOp Z.pow_pos := { TBOp := Z.pow ; TBOpInj _ _ := eq_refl }. -Add Zify BinOp Op_Z_pow_pos. +Add Tify BinOp Op_Z_pow_pos. #[global] Instance Op_Z_double : UnOp Z.double := { TUOp := Z.mul 2 ; TUOpInj := Z.double_spec }. -Add Zify UnOp Op_Z_double. +Add Tify UnOp Op_Z_double. #[global] Instance Op_Z_pred_double : UnOp Z.pred_double := { TUOp x := 2 * x - 1 ; TUOpInj := Z.pred_double_spec }. -Add Zify UnOp Op_Z_pred_double. +Add Tify UnOp Op_Z_pred_double. #[global] Instance Op_Z_succ_double : UnOp Z.succ_double := { TUOp x := 2 * x + 1 ; TUOpInj := Z.succ_double_spec }. -Add Zify UnOp Op_Z_succ_double. +Add Tify UnOp Op_Z_succ_double. #[global] Instance Op_Z_square : UnOp Z.square := { TUOp x := x * x ; TUOpInj := Z.square_spec }. -Add Zify UnOp Op_Z_square. +Add Tify UnOp Op_Z_square. #[global] Instance Op_Z_div2 : UnOp Z.div2 := { TUOp x := x / 2 ; TUOpInj := Z.div2_div }. -Add Zify UnOp Op_Z_div2. +Add Tify UnOp Op_Z_div2. #[global] Instance Op_Z_quot2 : UnOp Z.quot2 := { TUOp x := Z.quot x 2 ; TUOpInj := Zeven.Zquot2_quot }. -Add Zify UnOp Op_Z_quot2. +Add Tify UnOp Op_Z_quot2. Lemma of_nat_to_nat_eq x : Z.of_nat (Z.to_nat x) = Z.max 0 x. Proof. @@ -597,37 +597,37 @@ Qed. #[global] Instance Op_Z_to_nat : UnOp Z.to_nat := { TUOp x := Z.max 0 x ; TUOpInj := of_nat_to_nat_eq }. -Add Zify UnOp Op_Z_to_nat. +Add Tify UnOp Op_Z_to_nat. #[global] Instance Op_Z_to_pos : UnOp Z.to_pos := { TUOp x := Z.max 1 x ; TUOpInj x := ltac:(now simpl; destruct x; [| rewrite <- Pos2Z.inj_max; rewrite Pos.max_1_l |]) }. -Add Zify UnOp Op_Z_to_pos. +Add Tify UnOp Op_Z_to_pos. (** Specification of derived operators over Z *) #[global] Instance ZmaxSpec : BinOpSpec Z.max := { BPred n m r := n < m /\ r = m \/ m <= n /\ r = n ; BSpec := Z.max_spec }. -Add Zify BinOpSpec ZmaxSpec. +Add Tify BinOpSpec ZmaxSpec. #[global] Instance ZminSpec : BinOpSpec Z.min := { BPred n m r := n < m /\ r = n \/ m <= n /\ r = m ; BSpec := Z.min_spec }. -Add Zify BinOpSpec ZminSpec. +Add Tify BinOpSpec ZminSpec. #[global] Instance ZsgnSpec : UnOpSpec Z.sgn := { UPred n r := 0 < n /\ r = 1 \/ 0 = n /\ r = 0 \/ n < 0 /\ r = - 1 ; USpec := Z.sgn_spec }. -Add Zify UnOpSpec ZsgnSpec. +Add Tify UnOpSpec ZsgnSpec. #[global] Instance ZabsSpec : UnOpSpec Z.abs := { UPred n r := 0 <= n /\ r = n \/ n < 0 /\ r = - n ; USpec := Z.abs_spec }. -Add Zify UnOpSpec ZabsSpec. +Add Tify UnOpSpec ZabsSpec. (** Saturate positivity constraints *) @@ -637,7 +637,7 @@ Instance SatPowPos : Saturate Z.pow := PArg2 y := 0 <= y; PRes _ _ r := 0 < r; SatOk := fun x y => Z.pow_pos_nonneg x y}. -Add Zify Saturate SatPowPos. +Add Tify Saturate SatPowPos. #[global] Instance SatPowNonneg : Saturate Z.pow := @@ -645,6 +645,6 @@ Instance SatPowNonneg : Saturate Z.pow := PArg2 y := True; PRes _ _ r := 0 <= r; SatOk a b Ha _ := @Z.pow_nonneg a b Ha }. -Add Zify Saturate SatPowNonneg. +Add Tify Saturate SatPowNonneg. (* TODO #14736 for compatibility only, should be removed after deprecation *) diff --git a/theories/micromega/ZifyN.v b/theories/micromega/ZifyN.v index 69a06d23c3..6848b6e049 100644 --- a/theories/micromega/ZifyN.v +++ b/theories/micromega/ZifyN.v @@ -8,10 +8,10 @@ (* * (see LICENSE file for the text of the license) *) (************************************************************************) -(* Instances of [ZifyClasses] for dealing with advanced [N] operators. *) +(* Instances of [TifyClasses] for dealing with advanced [N] operators. *) From Stdlib Require Import BinNat BinInt Znat Zdiv. -From Stdlib Require Import ZifyClasses ZifyInst Zify. +From Stdlib Require Import TifyClasses ZifyInst Zify. Ltac zify_convert_to_euclidean_division_equations_flag ::= constr:(true). @@ -23,17 +23,17 @@ Existing Instance Inj_N_Z. #[global] Instance Op_N_div : BinOp N.div := {| TBOp := Z.div ; TBOpInj := N2Z.inj_div |}. -Add Zify BinOp Op_N_div. +Add Tify BinOp Op_N_div. #[global] Instance Op_N_mod : BinOp N.modulo := {| TBOp := Z.rem ; TBOpInj := N2Z.inj_rem |}. -Add Zify BinOp Op_N_mod. +Add Tify BinOp Op_N_mod. #[global] Instance Op_N_pow : BinOp N.pow := {| TBOp := Z.pow ; TBOpInj := N2Z.inj_pow|}. -Add Zify BinOp Op_N_pow. +Add Tify BinOp Op_N_pow. #[local] Open Scope Z_scope. @@ -45,7 +45,7 @@ Instance SatDiv : Saturate Z.div := PRes := fun _ _ r => 0 <= r; SatOk := Z_div_nonneg_nonneg |}. -Add Zify Saturate SatDiv. +Add Tify Saturate SatDiv. #[global] Instance SatMod : Saturate Z.modulo := @@ -55,4 +55,4 @@ Instance SatMod : Saturate Z.modulo := PRes := fun _ _ r => 0 <= r; SatOk := Z_mod_nonneg_nonneg |}. -Add Zify Saturate SatMod. +Add Tify Saturate SatMod. diff --git a/theories/micromega/ZifyNat.v b/theories/micromega/ZifyNat.v index 6c32015ed2..f9ceebd858 100644 --- a/theories/micromega/ZifyNat.v +++ b/theories/micromega/ZifyNat.v @@ -8,10 +8,10 @@ (* * (see LICENSE file for the text of the license) *) (************************************************************************) -(* Instances of [ZifyClasses] for dealing with advanced [nat] operators. *) +(* Instances of [TifyClasses] for dealing with advanced [nat] operators. *) From Stdlib Require Import BinInt Znat Zdiv. -From Stdlib Require Import ZifyClasses ZifyInst Zify. +From Stdlib Require Import TifyClasses ZifyInst Zify. Ltac zify_convert_to_euclidean_division_equations_flag ::= constr:(true). @@ -23,17 +23,17 @@ Existing Instance Inj_nat_Z. #[global] Instance Op_mod : BinOp Nat.modulo := {| TBOp := Z.modulo ; TBOpInj := Nat2Z.inj_mod |}. -Add Zify BinOp Op_mod. +Add Tify BinOp Op_mod. #[global] Instance Op_div : BinOp Nat.div := {| TBOp := Z.div ; TBOpInj := Nat2Z.inj_div |}. -Add Zify BinOp Op_div. +Add Tify BinOp Op_div. #[global] Instance Op_pow : BinOp Nat.pow := {| TBOp := Z.pow ; TBOpInj := Nat2Z.inj_pow |}. -Add Zify BinOp Op_pow. +Add Tify BinOp Op_pow. #[local] Open Scope Z_scope. @@ -45,7 +45,7 @@ Instance SatDiv : Saturate Z.div := PRes := fun _ _ r => 0 <= r; SatOk := Z_div_nonneg_nonneg |}. -Add Zify Saturate SatDiv. +Add Tify Saturate SatDiv. #[global] Instance SatMod : Saturate Z.modulo := @@ -55,4 +55,4 @@ Instance SatMod : Saturate Z.modulo := PRes := fun _ _ r => 0 <= r; SatOk := Z_mod_nonneg_nonneg |}. -Add Zify Saturate SatMod. +Add Tify Saturate SatMod. diff --git a/theories/micromega/ZifySint63.v b/theories/micromega/ZifySint63.v index 15a860a3fb..6ce17f9a17 100644 --- a/theories/micromega/ZifySint63.v +++ b/theories/micromega/ZifySint63.v @@ -1,7 +1,7 @@ From Stdlib Require Import ZArith. From Stdlib Require Import Sint63. From Stdlib Require Import ZifyBool. -Import ZifyClasses. +Import TifyClasses. Lemma to_Z_bounded (x : int) : (-4611686018427387904 <= to_Z x <= 4611686018427387903)%Z. @@ -11,32 +11,32 @@ Proof. now apply to_Z_bounded. Qed. Instance Inj_int_Z : InjTyp int Z := mkinj _ _ to_Z (fun x => -4611686018427387904 <= x <= 4611686018427387903)%Z to_Z_bounded. -Add Zify InjTyp Inj_int_Z. +Add Tify InjTyp Inj_int_Z. #[global] Instance Op_max_int : CstOp max_int := { TCst := 4611686018427387903 ; TCstInj := eq_refl }. -Add Zify CstOp Op_max_int. +Add Tify CstOp Op_max_int. #[global] Instance Op_min_int : CstOp min_int := { TCst := -4611686018427387904 ; TCstInj := eq_refl }. -Add Zify CstOp Op_min_int. +Add Tify CstOp Op_min_int. #[global] Instance Op_digits : CstOp digits := { TCst := 63 ; TCstInj := eq_refl }. -Add Zify CstOp Op_digits. +Add Tify CstOp Op_digits. #[global] Instance Op_size : CstOp size := { TCst := 63 ; TCstInj := eq_refl }. -Add Zify CstOp Op_size. +Add Tify CstOp Op_size. #[global] Instance Op_wB : CstOp wB := { TCst := 2^63 ; TCstInj := eq_refl }. -Add Zify CstOp Op_wB. +Add Tify CstOp Op_wB. Lemma ltb_lt : forall n m, (n (to_Z n = to_Z m)%sint63. Proof. @@ -89,7 +89,7 @@ Qed. #[global] Instance Op_eq : BinRel (@eq int) := {| TR := @eq Z; TRInj := eq_int_inj |}. -Add Zify BinRel Op_eq. +Add Tify BinRel Op_eq. Notation cmodwB x := ((x + 4611686018427387904) mod 9223372036854775808 - 4611686018427387904)%Z. @@ -97,42 +97,42 @@ Notation cmodwB x := #[global] Instance Op_add : BinOp add := {| TBOp := fun x y => cmodwB (x + y); TBOpInj := add_spec |}%Z. -Add Zify BinOp Op_add. +Add Tify BinOp Op_add. #[global] Instance Op_sub : BinOp sub := {| TBOp := fun x y => cmodwB (x - y); TBOpInj := sub_spec |}%Z. -Add Zify BinOp Op_sub. +Add Tify BinOp Op_sub. #[global] Instance Op_opp : UnOp Uint63.opp := {| TUOp := fun x => cmodwB (- x); TUOpInj := (sub_spec 0) |}%Z. -Add Zify UnOp Op_opp. +Add Tify UnOp Op_opp. #[global] Instance Op_succ : UnOp succ := {| TUOp := fun x => cmodwB (x + 1); TUOpInj := succ_spec |}%Z. -Add Zify UnOp Op_succ. +Add Tify UnOp Op_succ. #[global] Instance Op_pred : UnOp Uint63.pred := {| TUOp := fun x => cmodwB (x - 1); TUOpInj := pred_spec |}%Z. -Add Zify UnOp Op_pred. +Add Tify UnOp Op_pred. #[global] Instance Op_mul : BinOp mul := {| TBOp := fun x y => cmodwB (x * y); TBOpInj := mul_spec |}%Z. -Add Zify BinOp Op_mul. +Add Tify BinOp Op_mul. #[global] Instance Op_mod : BinOp PrimInt63.mods := {| TBOp := Z.rem ; TBOpInj := mod_spec |}. -Add Zify BinOp Op_mod. +Add Tify BinOp Op_mod. #[global] Instance Op_asr : BinOp asr := {| TBOp := fun x y => x / 2^ y ; TBOpInj := asr_spec |}%Z. -Add Zify BinOp Op_asr. +Add Tify BinOp Op_asr. Definition quots (x d : Z) : Z := if ((x =? -4611686018427387904)%Z && (d =? -1)%Z)%bool then @@ -155,7 +155,7 @@ Qed. #[global] Instance Op_div : BinOp div := {| TBOp := quots ; TBOpInj := div_quots |}. -Add Zify BinOp Op_div. +Add Tify BinOp Op_div. Lemma quots_spec (x y : Z) : ((x = -4611686018427387904 /\ y = -1 /\ quots x y = -4611686018427387904) @@ -172,17 +172,17 @@ Instance quotsSpec : BinOpSpec quots := ((x = -4611686018427387904 /\ d = -1 /\ r = -4611686018427387904) \/ ((x <> -4611686018427387904 \/ d <> -1) /\ r = Z.quot x d))%Z; BSpec := quots_spec |}. -Add Zify BinOpSpec quotsSpec. +Add Tify BinOpSpec quotsSpec. #[global] Instance Op_of_Z : UnOp of_Z := { TUOp := fun x => cmodwB x; TUOpInj := of_Z_spec }. -Add Zify UnOp Op_of_Z. +Add Tify UnOp Op_of_Z. #[global] Instance Op_to_Z : UnOp to_Z := { TUOp := fun x => x ; TUOpInj := fun x : int => eq_refl }. -Add Zify UnOp Op_to_Z. +Add Tify UnOp Op_to_Z. Lemma is_zeroE : forall n : int, is_zero n = (to_Z n =? 0)%Z. Proof. @@ -195,11 +195,11 @@ Qed. #[global] Instance Op_is_zero : UnOp is_zero := { TUOp := (Z.eqb 0) ; TUOpInj := is_zeroE }. -Add Zify UnOp Op_is_zero. +Add Tify UnOp Op_is_zero. #[global] Instance Op_abs : UnOp abs := { TUOp := fun x => cmodwB (Z.abs x) ; TUOpInj := abs_spec }. -Add Zify UnOp Op_abs. +Add Tify UnOp Op_abs. Ltac Zify.zify_convert_to_euclidean_division_equations_flag ::= constr:(true). diff --git a/theories/micromega/ZifyUint63.v b/theories/micromega/ZifyUint63.v index aaac88c953..e202c6a79b 100644 --- a/theories/micromega/ZifyUint63.v +++ b/theories/micromega/ZifyUint63.v @@ -1,7 +1,7 @@ From Stdlib Require Import ZArith. From Stdlib Require Import Uint63. From Stdlib Require Import ZifyBool. -Import ZifyClasses. +Import TifyClasses. Lemma to_Z_bounded : forall x, (0 <= to_Z x < 9223372036854775808)%Z. Proof. apply to_Z_bounded. Qed. @@ -9,27 +9,27 @@ Proof. apply to_Z_bounded. Qed. #[global] Instance Inj_int_Z : InjTyp int Z := mkinj _ _ to_Z (fun x => 0 <= x < 9223372036854775808)%Z to_Z_bounded. -Add Zify InjTyp Inj_int_Z. +Add Tify InjTyp Inj_int_Z. #[global] Instance Op_max_int : CstOp max_int := { TCst := 9223372036854775807 ; TCstInj := eq_refl }. -Add Zify CstOp Op_max_int. +Add Tify CstOp Op_max_int. #[global] Instance Op_digits : CstOp digits := { TCst := 63 ; TCstInj := eq_refl }. -Add Zify CstOp Op_digits. +Add Tify CstOp Op_digits. #[global] Instance Op_size : CstOp size := { TCst := 63 ; TCstInj := eq_refl }. -Add Zify CstOp Op_size. +Add Tify CstOp Op_size. #[global] Instance Op_wB : CstOp wB := { TCst := 2^63 ; TCstInj := eq_refl }. -Add Zify CstOp Op_wB. +Add Tify CstOp Op_wB. Lemma ltb_lt : forall n m, (n (φ n = φ m)%uint63. Proof. @@ -82,112 +82,112 @@ Qed. #[global] Instance Op_eq : BinRel (@eq int) := {| TR := @eq Z; TRInj := eq_int_inj |}. -Add Zify BinRel Op_eq. +Add Tify BinRel Op_eq. #[global] Instance Op_add : BinOp add := {| TBOp := fun x y => (x + y) mod 9223372036854775808%Z; TBOpInj := add_spec |}%Z. -Add Zify BinOp Op_add. +Add Tify BinOp Op_add. #[global] Instance Op_sub : BinOp sub := {| TBOp := fun x y => (x - y) mod 9223372036854775808%Z; TBOpInj := sub_spec |}%Z. -Add Zify BinOp Op_sub. +Add Tify BinOp Op_sub. #[global] Instance Op_opp : UnOp Uint63.opp := {| TUOp := (fun x => (- x) mod 9223372036854775808)%Z; TUOpInj := (sub_spec 0) |}%Z. -Add Zify UnOp Op_opp. +Add Tify UnOp Op_opp. #[global] Instance Op_oppcarry : UnOp oppcarry := {| TUOp := (fun x => 2^63 - x - 1)%Z; TUOpInj := oppcarry_spec |}%Z. -Add Zify UnOp Op_oppcarry. +Add Tify UnOp Op_oppcarry. #[global] Instance Op_succ : UnOp succ := {| TUOp := (fun x => (x + 1) mod 2^63)%Z; TUOpInj := succ_spec |}%Z. -Add Zify UnOp Op_succ. +Add Tify UnOp Op_succ. #[global] Instance Op_pred : UnOp Uint63.pred := {| TUOp := (fun x => (x - 1) mod 2^63)%Z; TUOpInj := pred_spec |}%Z. -Add Zify UnOp Op_pred. +Add Tify UnOp Op_pred. #[global] Instance Op_mul : BinOp mul := {| TBOp := fun x y => (x * y) mod 9223372036854775808%Z; TBOpInj := mul_spec |}%Z. -Add Zify BinOp Op_mul. +Add Tify BinOp Op_mul. #[global] Instance Op_gcd : BinOp gcd:= {| TBOp := (fun x y => Zgcd_alt.Zgcdn (2 * 63)%nat y x) ; TBOpInj := to_Z_gcd |}. -Add Zify BinOp Op_gcd. +Add Tify BinOp Op_gcd. #[global] Instance Op_mod : BinOp Uint63.mod := {| TBOp := Z.modulo ; TBOpInj := mod_spec |}. -Add Zify BinOp Op_mod. +Add Tify BinOp Op_mod. #[global] Instance Op_subcarry : BinOp subcarry := {| TBOp := (fun x y => (x - y - 1) mod 2^63)%Z ; TBOpInj := subcarry_spec |}. -Add Zify BinOp Op_subcarry. +Add Tify BinOp Op_subcarry. #[global] Instance Op_addcarry : BinOp addcarry := {| TBOp := (fun x y => (x + y + 1) mod 2^63)%Z ; TBOpInj := addcarry_spec |}. -Add Zify BinOp Op_addcarry. +Add Tify BinOp Op_addcarry. #[global] Instance Op_lsr : BinOp lsr := {| TBOp := (fun x y => x / 2^ y)%Z ; TBOpInj := lsr_spec |}. -Add Zify BinOp Op_lsr. +Add Tify BinOp Op_lsr. #[global] Instance Op_lsl : BinOp lsl := {| TBOp := (fun x y => (x * 2^ y) mod 2^ 63)%Z ; TBOpInj := lsl_spec |}. -Add Zify BinOp Op_lsl. +Add Tify BinOp Op_lsl. #[global] Instance Op_lor : BinOp Uint63.lor := {| TBOp := Z.lor ; TBOpInj := lor_spec' |}. -Add Zify BinOp Op_lor. +Add Tify BinOp Op_lor. #[global] Instance Op_land : BinOp Uint63.land := {| TBOp := Z.land ; TBOpInj := land_spec' |}. -Add Zify BinOp Op_land. +Add Tify BinOp Op_land. #[global] Instance Op_lxor : BinOp Uint63.lxor := {| TBOp := Z.lxor ; TBOpInj := lxor_spec' |}. -Add Zify BinOp Op_lxor. +Add Tify BinOp Op_lxor. #[global] Instance Op_div : BinOp div := {| TBOp := Z.div ; TBOpInj := div_spec |}. -Add Zify BinOp Op_div. +Add Tify BinOp Op_div. #[global] Instance Op_bit : BinOp bit := {| TBOp := Z.testbit ; TBOpInj := bitE |}. -Add Zify BinOp Op_bit. +Add Tify BinOp Op_bit. #[global] Instance Op_of_Z : UnOp of_Z := { TUOp := (fun x => x mod 9223372036854775808)%Z; TUOpInj := of_Z_spec }. -Add Zify UnOp Op_of_Z. +Add Tify UnOp Op_of_Z. #[global] Instance Op_to_Z : UnOp to_Z := { TUOp := fun x => x ; TUOpInj := fun x : int => eq_refl }. -Add Zify UnOp Op_to_Z. +Add Tify UnOp Op_to_Z. #[global] Instance Op_is_zero : UnOp is_zero := { TUOp := (Z.eqb 0) ; TUOpInj := is_zeroE }. -Add Zify UnOp Op_is_zero. +Add Tify UnOp Op_is_zero. Lemma is_evenE : forall x, is_even x = Z.even φ (x)%uint63. @@ -203,7 +203,7 @@ Qed. #[global] Instance Op_is_even : UnOp is_even := { TUOp := Z.even ; TUOpInj := is_evenE }. -Add Zify UnOp Op_is_even. +Add Tify UnOp Op_is_even. Ltac Zify.zify_convert_to_euclidean_division_equations_flag ::= constr:(true). diff --git a/theories/micromega/isZ.v b/theories/micromega/isZ.v new file mode 100644 index 0000000000..79674d4d7c --- /dev/null +++ b/theories/micromega/isZ.v @@ -0,0 +1,74 @@ +(************************************************************************) +(* * The Rocq Prover / The Rocq Development Team *) +(* v * Copyright INRIA, CNRS and contributors *) +(* isZ' r. +Proof. + unfold isZ,isZ'. + split ; intros. + - destruct H. subst. + destruct (archimed (IZR x)). + apply Rgt_lt in H. + apply lt_IZR in H. + rewrite <- minus_IZR in H0. + apply le_IZR in H0. + apply f_equal. + (* proof over Z *) + symmetry. + apply Z.le_antisymm. + apply Zorder.Zlt_succ_le. + now ring_simplify. + apply Zorder.Zle_minus_le_0 in H0. + apply Zorder.Zle_0_minus_le. + now replace (1 - (up (IZR x) - x))%Z with (x - (up (IZR x) - 1))%Z in H0 by ring. + - exists (up r - 1)%Z. + congruence. +Qed. + +Lemma isZ_dec : forall r, isZ r \/ ~ isZ r. +Proof. + intros. + rewrite isZ_eq. + unfold isZ'. + apply Req_dec. +Qed. + + +Definition isZb (r:R) := if Req_dec_T r (IZR (up r - 1)) then true else false. + +Lemma isZ_isZb : forall r, isZ r <-> isZb r = true. +Proof. + intros. + rewrite isZ_eq. + unfold isZ',isZb. + split ; intros. + - destruct (Req_dec_T r (IZR (up r - 1))). + reflexivity. + congruence. + - destruct (Req_dec_T r (IZR (up r - 1))) ; try discriminate. + congruence. +Qed. + +Lemma isZb_IZR : forall x, isZb (IZR x) = true. +Proof. + intros. + rewrite <- isZ_isZb. + exists x. reflexivity. +Qed. diff --git a/theories/omega/PreOmega.v b/theories/omega/PreOmega.v index a887994cff..6ed93255d8 100644 --- a/theories/omega/PreOmega.v +++ b/theories/omega/PreOmega.v @@ -228,7 +228,7 @@ Module Z. to_euclidean_division_equations_with euclidean_division_equations_flags.default. End Z. -From Stdlib Require Import ZifyClasses ZifyInst. +From Stdlib Require Import TifyClasses ZifyInst. From Stdlib Require Zify. Ltac Zify.zify_internal_to_euclidean_division_equations ::= Z.to_euclidean_division_equations. From d0ab6c3d8dd35d4ceb0b42d5c9b9ff765a926899 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fr=C3=A9d=C3=A9ric=20Besson?= Date: Thu, 27 Feb 2025 10:57:05 +0100 Subject: [PATCH 2/6] [Cos_rel] drop dependency on Lra lra, nra -> field,ring --- theories/Reals/Cos_rel.v | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/theories/Reals/Cos_rel.v b/theories/Reals/Cos_rel.v index b2e4fee31e..e318b8e296 100644 --- a/theories/Reals/Cos_rel.v +++ b/theories/Reals/Cos_rel.v @@ -12,7 +12,7 @@ From Stdlib Require Import Rbase. From Stdlib Require Import Rfunctions. From Stdlib Require Import SeqSeries. From Stdlib Require Import Rtrigo_def. -From Stdlib Require Import Lia Lra. +From Stdlib Require Import Lia. From Stdlib Require Import Arith.Factorial. Local Open Scope R_scope. @@ -77,7 +77,7 @@ replace y ^ (2 * (S n - l)))) (pred (S n - k))) ( pred (S n))) with (Reste1 x y (S n)). 2:{ unfold Reste1; apply sum_eq; intros. - apply sum_eq; intros. nra. } + apply sum_eq; intros. ring. } replace (sum_f_R0 (fun k:nat => @@ -89,7 +89,7 @@ replace y ^ (2 * (n - l) + 1))) (pred (n - k))) ( pred n)) with (Reste2 x y n). 2:{ unfold Reste2; apply sum_eq; intros. - apply sum_eq; intros. nra. } + apply sum_eq; intros. ring. } replace (sum_f_R0 (fun k:nat => @@ -153,7 +153,7 @@ replace unfold C1. apply sum_eq; intros. induction i as [| i Hreci]. - { unfold C; simpl. nra. } + { unfold C; simpl. field. } unfold sin_nnn. rewrite <- Rmult_plus_distr_l. apply Rmult_eq_compat_l. From 05037406a89cf74f7113b81aae1005ae54f5391f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fr=C3=A9d=C3=A9ric=20Besson?= Date: Thu, 27 Mar 2025 10:01:11 +0100 Subject: [PATCH 3/6] [zify] add Pcompare --- theories/micromega/ZifyComparison.v | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/theories/micromega/ZifyComparison.v b/theories/micromega/ZifyComparison.v index 994d42ad98..7e7dab332e 100644 --- a/theories/micromega/ZifyComparison.v +++ b/theories/micromega/ZifyComparison.v @@ -31,6 +31,8 @@ Instance Inj_comparison_Z : InjTyp comparison Z := { inj := Z_of_comparison ; pred :=(fun x => -1 <= x <= 1) ; cstr := Z_of_comparison_bound}. Add Tify InjTyp Inj_comparison_Z. + + Definition ZcompareZ (x y : Z) := Z_of_comparison (Z.compare x y). @@ -39,6 +41,12 @@ Program Instance BinOp_Zcompare : BinOp Z.compare := { TBOp := ZcompareZ }. Add Tify BinOp BinOp_Zcompare. +#[global] +Program Instance BinOp_Pcompare : BinOp Pos.compare := + { TBOp := ZcompareZ }. +Add Tify BinOp BinOp_Pcompare. + + #[global] Instance Op_eq_comparison : BinRel (@eq comparison) := {TR := @eq Z ; TRInj := ltac:(intros [] []; simpl ; intuition congruence) }. From dc34c3a5001ddca2fe2c20d714acc2c040d25c18 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fr=C3=A9d=C3=A9ric=20Besson?= Date: Thu, 27 Mar 2025 16:16:13 +0100 Subject: [PATCH 4/6] [zify] Notation from Zify -> Tify --- theories/micromega/ZifyClasses.v | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/theories/micromega/ZifyClasses.v b/theories/micromega/ZifyClasses.v index 5cdfcc8c39..7b93aa4598 100644 --- a/theories/micromega/ZifyClasses.v +++ b/theories/micromega/ZifyClasses.v @@ -1,3 +1,14 @@ Attributes deprecated(since="9.1", note="Use TifyClasses instead."). From Stdlib Require Export TifyClasses. + +Notation InjTyp := TifyClasses.InjTyp. +Notation BinOp := TifyClasses.BinOp. +Notation UnOp := TifyClasses.UnOp. +Notation CstOp := TifyClasses.CstOp. +Notation BinRel := TifyClasses.BinRel. +Notation PropOp := TifyClasses.PropOp. +Notation PropUOp := TifyClasses.PropUOp. +Notation BinOpSpec := TifyClasses.BinOpSpec. +Notation UnOpSpec := TifyClasses.UnOpSpec. +Notation Saturate := TifyClasses.Saturate. From f796e057d8be7dd39ef1e0ef768114105a7bcaa6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fr=C3=A9d=C3=A9ric=20Besson?= Date: Thu, 27 Mar 2025 16:16:54 +0100 Subject: [PATCH 5/6] fix test-suite --- test-suite/bugs/bug_16803.v | 2 +- test-suite/bugs/bug_18151.v | 20 +- test-suite/micromega/bug_18158.v | 50 +-- test-suite/micromega/milp.v | 77 +++- test-suite/micromega/rexample.v | 3 +- test-suite/micromega/witness_tactics.v | 2 +- test-suite/output/MExtraction.out | 549 ++++++++++++++++--------- test-suite/success/tify.v | 2 +- 8 files changed, 476 insertions(+), 229 deletions(-) diff --git a/test-suite/bugs/bug_16803.v b/test-suite/bugs/bug_16803.v index b5f63a8ed5..02a403b3f7 100644 --- a/test-suite/bugs/bug_16803.v +++ b/test-suite/bugs/bug_16803.v @@ -1,6 +1,6 @@ From Stdlib Require Import Lia. From Stdlib Require Import ZArith. -Import ZifyClasses. +Import TifyClasses. Module Test1. diff --git a/test-suite/bugs/bug_18151.v b/test-suite/bugs/bug_18151.v index 1a92ddc2e0..b27cc01d2d 100644 --- a/test-suite/bugs/bug_18151.v +++ b/test-suite/bugs/bug_18151.v @@ -3,12 +3,12 @@ From Stdlib Require Import Zify ZifyClasses. Import Z. Open Scope Z_scope. -#[global] Instance sat_mod_le : ZifyClasses.Saturate BinIntDef.Z.modulo := +#[global] Instance sat_mod_le : Saturate BinIntDef.Z.modulo := {| - ZifyClasses.PArg1 := fun a => 0 <= a; - ZifyClasses.PArg2 := fun b => 0 < b; - ZifyClasses.PRes := fun a b ab => ab <= a; - ZifyClasses.SatOk := mod_le + PArg1 := fun a => 0 <= a; + PArg2 := fun b => 0 < b; + PRes := fun a b ab => ab <= a; + SatOk := mod_le |}. Add Zify Saturate sat_mod_le. @@ -16,12 +16,12 @@ Lemma shiftr_lbound a n: 0 <= a -> True -> 0 <= (Z.shiftr a n). Proof. now intros; apply Z.shiftr_nonneg. Qed. -#[global] Instance sat_shiftr_lbound : ZifyClasses.Saturate BinIntDef.Z.shiftr := +#[global] Instance sat_shiftr_lbound : Saturate BinIntDef.Z.shiftr := {| - ZifyClasses.PArg1 := fun a => 0 <= a; - ZifyClasses.PArg2 := fun b => True; - ZifyClasses.PRes := fun a b ab => 0 <= ab; - ZifyClasses.SatOk := shiftr_lbound + PArg1 := fun a => 0 <= a; + PArg2 := fun b => True; + PRes := fun a b ab => 0 <= ab; + SatOk := shiftr_lbound |}. Add Zify Saturate sat_shiftr_lbound. diff --git a/test-suite/micromega/bug_18158.v b/test-suite/micromega/bug_18158.v index 002e0f373f..9e8eecf480 100644 --- a/test-suite/micromega/bug_18158.v +++ b/test-suite/micromega/bug_18158.v @@ -40,42 +40,42 @@ Proof. intros; apply shiftr_ubound; assumption. Qed. -#[global] Instance sat_shiftr_lbound : ZifyClasses.Saturate BinIntDef.Z.shiftr := +#[global] Instance sat_shiftr_lbound : TifyClasses.Saturate BinIntDef.Z.shiftr := {| - ZifyClasses.PArg1 := fun a => 0 <= a; - ZifyClasses.PArg2 := fun b => True; - ZifyClasses.PRes := fun a b ab => 0 <= ab; - ZifyClasses.SatOk := shiftr_lbound + TifyClasses.PArg1 := fun a => 0 <= a; + TifyClasses.PArg2 := fun b => True; + TifyClasses.PRes := fun a b ab => 0 <= ab; + TifyClasses.SatOk := shiftr_lbound |}. -Add Zify Saturate sat_shiftr_lbound. +Add Tify Saturate sat_shiftr_lbound. -#[global] Instance sat_shiftr_contr_8 : ZifyClasses.Saturate BinIntDef.Z.shiftr := +#[global] Instance sat_shiftr_contr_8 : TifyClasses.Saturate BinIntDef.Z.shiftr := {| - ZifyClasses.PArg1 := fun a => 0 <= a < 2 ^ 8; - ZifyClasses.PArg2 := fun b => 0 <= b; - ZifyClasses.PRes := fun a b ab => ab < 2 ^ 8; - ZifyClasses.SatOk := shiftrContractive8 + TifyClasses.PArg1 := fun a => 0 <= a < 2 ^ 8; + TifyClasses.PArg2 := fun b => 0 <= b; + TifyClasses.PRes := fun a b ab => ab < 2 ^ 8; + TifyClasses.SatOk := shiftrContractive8 |}. -Add Zify Saturate sat_shiftr_contr_8. +Add Tify Saturate sat_shiftr_contr_8. -#[global] Instance sat_shiftr_contr_16 : ZifyClasses.Saturate BinIntDef.Z.shiftr := +#[global] Instance sat_shiftr_contr_16 : TifyClasses.Saturate BinIntDef.Z.shiftr := {| - ZifyClasses.PArg1 := fun a => 0 <= a < 2 ^ 16; - ZifyClasses.PArg2 := fun b => 0 <= b; - ZifyClasses.PRes := fun a b ab => ab < 2 ^ 16; - ZifyClasses.SatOk := shiftrContractive16 + TifyClasses.PArg1 := fun a => 0 <= a < 2 ^ 16; + TifyClasses.PArg2 := fun b => 0 <= b; + TifyClasses.PRes := fun a b ab => ab < 2 ^ 16; + TifyClasses.SatOk := shiftrContractive16 |}. -Add Zify Saturate sat_shiftr_contr_16. +Add Tify Saturate sat_shiftr_contr_16. -#[global] Instance sat_shiftr_contr_32 : ZifyClasses.Saturate BinIntDef.Z.shiftr := +#[global] Instance sat_shiftr_contr_32 : TifyClasses.Saturate BinIntDef.Z.shiftr := {| - ZifyClasses.PArg1 := fun a => 0 <= a < 2 ^ 32; - ZifyClasses.PArg2 := fun b => 0 <= b; - ZifyClasses.PRes := fun a b ab => ab < 2 ^ 32; - ZifyClasses.SatOk := shiftrContractive32 + TifyClasses.PArg1 := fun a => 0 <= a < 2 ^ 32; + TifyClasses.PArg2 := fun b => 0 <= b; + TifyClasses.PRes := fun a b ab => ab < 2 ^ 32; + TifyClasses.SatOk := shiftrContractive32 |}. -Add Zify Saturate sat_shiftr_contr_32. +Add Tify Saturate sat_shiftr_contr_32. Goal forall x y , @@ -85,7 +85,7 @@ Goal forall x y , -> Z.le (Z.shiftr y 8) 255 -> Z.le (Z.shiftr x 24) 255. intros. - Zify.zify_saturate. + Tify.tify_saturate. (* [xlia zchecker] used to raise a [Stack overflow] error. It is supposed to fail normally. *) assert_fails (xlia zchecker). Abort. diff --git a/test-suite/micromega/milp.v b/test-suite/micromega/milp.v index a1772b7c67..e62681949c 100644 --- a/test-suite/micromega/milp.v +++ b/test-suite/micromega/milp.v @@ -1,26 +1,89 @@ Require Import Rbase Rify isZ Lra Zfloor. Local Open Scope R_scope. +Unset Lra Cache. + Lemma up0 : up 0%R = 1%Z. -Proof. intros. rify. lra. Qed. +Proof. intros. lra. Qed. Lemma up_succ r : up (r + 1)%R = (up r + 1)%Z. -Proof. intros. rify. lra. Qed. +Proof. intros. lra. Qed. Lemma up_IZR z : up (IZR z) = (z + 1)%Z. Proof. - rify. lra. + lra. Qed. Lemma up_shiftz r z : up (r + IZR z)%R = (up r + z)%Z. Proof. - rify. lra. + lra. +Qed. + +Lemma div_inv : forall y x, + 1 = 1 + (y - x) / x -> + (y - x) * / x = 0. +Proof. intros. lra. Qed. + +Goal forall (x1 : R) (x2 : R) (x3 : R) (x4 : R) (x5 : R) (x6 : R) +(H0 : -x2 + x6 = R0) +(H1 : x6 + -x4 = R0) +(H2 : x2*x3 + -x1*x3 = R0) +(H3 : x4 + -x2*x5 = R0) +(H5 : x2*x3 + -x1*x3 > R0) +, False +. +Proof. + intros. + lra. +Qed. + +Goal forall (x1 : R) (x2 : R) (x3 : R) (x4 : R) +(H0 : x4 >= 0) +(H1 : 1 + -x4 > 0) +(H2 : -x4 + x1 > 0) +(H3 : 1 + -x1 > 0) +(H4 : x2 + -x3 = 0) +(H5 : -x1 > 0) +, False +. +Proof. + intros. + lra. Qed. +Goal forall (x1 : R) (x2 : R) (x3 : R) (x4 : R) +(H0 : x4 >= 0) +(H1 : 1 + -x4 > 0) +(H2 : -x4 + x1 > 0) +(H3 : 1 + -x1 > 0) +(H4 : x2 + -x3 = 0) +(H5 : -x1 > 0) +, False +. +Proof. + intros. + lra. +Qed. + + + + +Goal forall (x1 : R) (x2 : R) (x3 : R) (x4 : R) (x5 : R) (x6 : R) +(H0 : -x2 + x6 = R0) +(H1 : x6 + -x4 = R0) +(H2 : x2*x3 + -x1*x3 = R0) +(H3 : x4 + -x2*x5 = R0) +(H5 : x2*x3 + -x1*x3 > R0) +, False +. +Proof. + lra. +Qed. + Lemma up_succ2 r n : Zfloor (r + IZR n) = (Zfloor r + n)%Z. Proof. - rify. lra. + lra. Qed. Goal forall x1 x2 x3 x5 x6 @@ -175,7 +238,7 @@ Goal forall e x 0 < ef. Proof. intros. - rify. nra. + nra. Qed. @@ -185,6 +248,6 @@ Goal forall x v, v = 1 / (3 * x). Proof. intros. - rify. nra. + nra. Qed. diff --git a/test-suite/micromega/rexample.v b/test-suite/micromega/rexample.v index 22eeaa32ac..c536de09c7 100644 --- a/test-suite/micromega/rexample.v +++ b/test-suite/micromega/rexample.v @@ -75,11 +75,12 @@ Proof. lra. Qed. -Goal forall x, -x^2 >= 0 -> x - 1 >= 0 -> False. +(*Goal forall x, -x^2 >= 0 -> x - 1 >= 0 -> False. Proof. intros. psatz R 3. Qed. + *) Goal forall x, -x^2 >= 0 -> x - 1 >= 0 -> False. Proof. diff --git a/test-suite/micromega/witness_tactics.v b/test-suite/micromega/witness_tactics.v index c4c12066e5..033f4caa95 100644 --- a/test-suite/micromega/witness_tactics.v +++ b/test-suite/micromega/witness_tactics.v @@ -1,6 +1,6 @@ From Stdlib Require Import ZArith QArith. From Stdlib.micromega Require Import RingMicromega EnvRing Tauto. -From Stdlib.micromega Require Import ZMicromega QMicromega. +From Stdlib.micromega Require Import ZMicromega QMicromega MMicromega. Declare ML Module "rocq-runtime.plugins.micromega". diff --git a/test-suite/output/MExtraction.out b/test-suite/output/MExtraction.out index 7a08d63147..ca26db7c10 100644 --- a/test-suite/output/MExtraction.out +++ b/test-suite/output/MExtraction.out @@ -688,6 +688,11 @@ module Z = | Z0 -> abs a | Zpos b0 -> Zpos (Coq_Pos.gcd a0 b0) | Zneg b0 -> Zpos (Coq_Pos.gcd a0 b0)) + + (** val lcm : z -> z -> z **) + + let lcm a b = + abs (mul a (div b (gcd a b))) end type 'c pExpr = @@ -1164,7 +1169,7 @@ let rec map_bformula _ fct = function | TT k -> TT k | FF k -> FF k | X (k, p) -> X (k, p) -| A (k, a, t0) -> A (k, (fct a), t0) +| A (k, a, t1) -> A (k, (fct a), t1) | AND (k0, f1, f2) -> AND (k0, (map_bformula k0 fct f1), (map_bformula k0 fct f2)) | OR (k0, f1, f2) -> @@ -1195,21 +1200,21 @@ let cnf_ff = ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1 * 'a2) -> ('a1, 'a2) clause -> ('a1, 'a2) clause option **) -let rec add_term unsat deduce t0 = function +let rec add_term unsat deduce t1 = function | [] -> - (match deduce (fst t0) (fst t0) with - | Some u -> if unsat u then None else Some (t0::[]) - | None -> Some (t0::[])) + (match deduce (fst t1) (fst t1) with + | Some u -> if unsat u then None else Some (t1::[]) + | None -> Some (t1::[])) | t'::cl0 -> - (match deduce (fst t0) (fst t') with + (match deduce (fst t1) (fst t') with | Some u -> if unsat u then None - else (match add_term unsat deduce t0 cl0 with + else (match add_term unsat deduce t1 cl0 with | Some cl' -> Some (t'::cl') | None -> None) | None -> - (match add_term unsat deduce t0 cl0 with + (match add_term unsat deduce t1 cl0 with | Some cl' -> Some (t'::cl') | None -> None)) @@ -1220,8 +1225,8 @@ let rec add_term unsat deduce t0 = function let rec or_clause unsat deduce cl1 cl2 = match cl1 with | [] -> Some cl2 - | t0::cl -> - (match add_term unsat deduce t0 cl2 with + | t1::cl -> + (match add_term unsat deduce t1 cl2 with | Some cl' -> or_clause unsat deduce cl cl' | None -> None) @@ -1229,9 +1234,9 @@ let rec or_clause unsat deduce cl1 cl2 = ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) clause -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf **) -let xor_clause_cnf unsat deduce t0 f = +let xor_clause_cnf unsat deduce t1 f = fold_left (fun acc e -> - match or_clause unsat deduce t0 e with + match or_clause unsat deduce t1 e with | Some cl -> cl::acc | None -> acc) f [] @@ -1239,10 +1244,10 @@ let xor_clause_cnf unsat deduce t0 f = ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) clause -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf **) -let or_clause_cnf unsat deduce t0 f = - match t0 with +let or_clause_cnf unsat deduce t1 f = + match t1 with | [] -> f - | _::_ -> xor_clause_cnf unsat deduce t0 f + | _::_ -> xor_clause_cnf unsat deduce t1 f (** val or_cnf : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) cnf -> ('a1, @@ -1350,7 +1355,7 @@ let rec xcnf unsat deduce normalise1 negate0 pol0 _ = function | TT _ -> if pol0 then cnf_tt else cnf_ff | FF _ -> if pol0 then cnf_ff else cnf_tt | X (_, _) -> cnf_ff -| A (_, x, t0) -> if pol0 then normalise1 x t0 else negate0 x t0 +| A (_, x, t1) -> if pol0 then normalise1 x t1 else negate0 x t1 | AND (k0, e1, e2) -> mk_and unsat deduce (fun x x0 x1 -> xcnf unsat deduce normalise1 negate0 x x0 x1) k0 pol0 e1 e2 @@ -1382,21 +1387,21 @@ let rec xcnf unsat deduce normalise1 negate0 pol0 _ = function ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1 * 'a2) -> ('a1, 'a2) clause -> (('a1, 'a2) clause, 'a2 trace) sum **) -let rec radd_term unsat deduce t0 = function +let rec radd_term unsat deduce t1 = function | [] -> - (match deduce (fst t0) (fst t0) with - | Some u -> if unsat u then Inr (Push ((snd t0), Null)) else Inl (t0::[]) - | None -> Inl (t0::[])) + (match deduce (fst t1) (fst t1) with + | Some u -> if unsat u then Inr (Push ((snd t1), Null)) else Inl (t1::[]) + | None -> Inl (t1::[])) | t'::cl0 -> - (match deduce (fst t0) (fst t') with + (match deduce (fst t1) (fst t') with | Some u -> if unsat u - then Inr (Push ((snd t0), (Push ((snd t'), Null)))) - else (match radd_term unsat deduce t0 cl0 with + then Inr (Push ((snd t1), (Push ((snd t'), Null)))) + else (match radd_term unsat deduce t1 cl0 with | Inl cl' -> Inl (t'::cl') | Inr l -> Inr l) | None -> - (match radd_term unsat deduce t0 cl0 with + (match radd_term unsat deduce t1 cl0 with | Inl cl' -> Inl (t'::cl') | Inr l -> Inr l)) @@ -1407,8 +1412,8 @@ let rec radd_term unsat deduce t0 = function let rec ror_clause unsat deduce cl1 cl2 = match cl1 with | [] -> Inl cl2 - | t0::cl -> - (match radd_term unsat deduce t0 cl2 with + | t1::cl -> + (match radd_term unsat deduce t1 cl2 with | Inl cl' -> ror_clause unsat deduce cl cl' | Inr l -> Inr l) @@ -1416,10 +1421,10 @@ let rec ror_clause unsat deduce cl1 cl2 = ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1 * 'a2) list -> ('a1, 'a2) clause list -> ('a1, 'a2) clause list * 'a2 trace **) -let xror_clause_cnf unsat deduce t0 f = +let xror_clause_cnf unsat deduce t1 f = fold_left (fun pat e -> let acc,tg = pat in - (match ror_clause unsat deduce t0 e with + (match ror_clause unsat deduce t1 e with | Inl cl -> (cl::acc),tg | Inr l -> acc,(Merge (tg, l)))) f ([],Null) @@ -1428,10 +1433,10 @@ let xror_clause_cnf unsat deduce t0 f = ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1 * 'a2) list -> ('a1, 'a2) clause list -> ('a1, 'a2) clause list * 'a2 trace **) -let ror_clause_cnf unsat deduce t0 f = - match t0 with +let ror_clause_cnf unsat deduce t1 f = + match t1 with | [] -> f,Null - | _::_ -> xror_clause_cnf unsat deduce t0 f + | _::_ -> xror_clause_cnf unsat deduce t1 f (** val ror_cnf : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) clause list -> @@ -1441,9 +1446,9 @@ let rec ror_cnf unsat deduce f f' = match f with | [] -> cnf_tt,Null | e::rst -> - let rst_f',t0 = ror_cnf unsat deduce rst f' in + let rst_f',t1 = ror_cnf unsat deduce rst f' in let e_f',t' = ror_clause_cnf unsat deduce e f' in - (rev_append rst_f' e_f'),(Merge (t0, t')) + (rev_append rst_f' e_f'),(Merge (t1, t')) (** val ror_cnf_opt : ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) cnf -> ('a1, @@ -1533,8 +1538,8 @@ let rec rxcnf unsat deduce normalise1 negate0 polarity _ = function | TT _ -> if polarity then cnf_tt,Null else cnf_ff,Null | FF _ -> if polarity then cnf_ff,Null else cnf_tt,Null | X (_, _) -> cnf_ff,Null -| A (_, x, t0) -> - ratom (if polarity then normalise1 x t0 else negate0 x t0) t0 +| A (_, x, t1) -> + ratom (if polarity then normalise1 x t1 else negate0 x t1) t1 | AND (k0, e1, e2) -> rxcnf_and unsat deduce (fun x x0 x1 -> rxcnf unsat deduce normalise1 negate0 x x0 x1) polarity k0 e1 e2 @@ -1570,7 +1575,7 @@ let rec aformula to_constr _ = function | TT b -> to_constr.mkTT b | FF b -> to_constr.mkFF b | X (_, p) -> p -| A (b, x, t0) -> to_constr.mkA b x t0 +| A (b, x, t1) -> to_constr.mkA b x t1 | AND (k0, f1, f2) -> to_constr.mkAND k0 (aformula to_constr k0 f1) (aformula to_constr k0 f2) | OR (k0, f1, f2) -> @@ -1644,8 +1649,8 @@ let mk_arrow o k f1 f2 = 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4) tFormula **) let rec abst_simpl to_constr needA _ = function -| A (k, x, t0) -> - if needA t0 then A (k, x, t0) else X (k, (to_constr.mkA k x t0)) +| A (k, x, t1) -> + if needA t1 then A (k, x, t1) else X (k, (to_constr.mkA k x t1)) | AND (k0, f1, f2) -> AND (k0, (abst_simpl to_constr needA k0 f1), (abst_simpl to_constr needA k0 f2)) @@ -1754,8 +1759,8 @@ let rec abst_form to_constr needA pol0 _ = function | TT k -> if pol0 then TT k else X (k, (to_constr.mkTT k)) | FF k -> if pol0 then X (k, (to_constr.mkFF k)) else FF k | X (k, p) -> X (k, p) -| A (k, x, t0) -> - if needA t0 then A (k, x, t0) else X (k, (to_constr.mkA k x t0)) +| A (k, x, t1) -> + if needA t1 then A (k, x, t1) else X (k, (to_constr.mkA k x t1)) | AND (k0, f1, f2) -> abst_and to_constr (abst_form to_constr needA) pol0 k0 f1 f2 | OR (k0, f1, f2) -> @@ -2036,8 +2041,8 @@ let cnf_of_list cO ceqb cleb l tg = -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a2 -> ('a1 nFormula, 'a2) cnf **) -let cnf_normalise cO cI cplus ctimes cminus copp ceqb cleb t0 tg = - let f = normalise cO cI cplus ctimes cminus copp ceqb t0 in +let cnf_normalise cO cI cplus ctimes cminus copp ceqb cleb t1 tg = + let f = normalise cO cI cplus ctimes cminus copp ceqb t1 in if check_inconsistent cO ceqb cleb f then cnf_ff else cnf_of_list cO ceqb cleb (xnormalise copp f) tg @@ -2047,8 +2052,8 @@ let cnf_normalise cO cI cplus ctimes cminus copp ceqb cleb t0 tg = -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a2 -> ('a1 nFormula, 'a2) cnf **) -let cnf_negate cO cI cplus ctimes cminus copp ceqb cleb t0 tg = - let f = normalise cO cI cplus ctimes cminus copp ceqb t0 in +let cnf_negate cO cI cplus ctimes cminus copp ceqb cleb t1 tg = + let f = normalise cO cI cplus ctimes cminus copp ceqb t1 in if check_inconsistent cO ceqb cleb f then cnf_tt else cnf_of_list cO ceqb cleb (xnegate copp f) tg @@ -2089,10 +2094,10 @@ let map_Formula c_of_S f = 'a1 psatz **) let simpl_cone cO cI ctimes ceqb e = match e with -| PsatzSquare t0 -> - (match t0 with +| PsatzSquare t1 -> + (match t1 with | Pc c -> if ceqb cO c then PsatzZ else PsatzC (ctimes c c) - | _ -> PsatzSquare t0) + | _ -> PsatzSquare t1) | PsatzMulE (t1, t2) -> (match t1 with | PsatzMulE (x, x0) -> @@ -2142,12 +2147,50 @@ let simpl_cone cO cI ctimes ceqb e = match e with | _ -> PsatzAdd (t1, t2))) | _ -> e -type 'a t = +module PositiveSet = + struct + type tree = + | Leaf + | Node of tree * bool * tree + + type t = tree + + (** val empty : t **) + + let empty = + Leaf + + (** val mem : positive -> t -> bool **) + + let rec mem i = function + | Leaf -> false + | Node (l, o, r) -> + (match i with + | XI i0 -> mem i0 r + | XO i0 -> mem i0 l + | XH -> o) + + (** val add : positive -> t -> t **) + + let rec add i = function + | Leaf -> + (match i with + | XI i0 -> Node (Leaf, false, (add i0 Leaf)) + | XO i0 -> Node ((add i0 Leaf), false, Leaf) + | XH -> Node (Leaf, true, Leaf)) + | Node (l, o, r) -> + (match i with + | XI i0 -> Node (l, o, (add i0 r)) + | XO i0 -> Node ((add i0 l), o, r) + | XH -> Node (l, true, r)) + end + +type 'a t0 = | Empty | Elt of 'a -| Branch of 'a t * 'a * 'a t +| Branch of 'a t0 * 'a * 'a t0 -(** val find : 'a1 -> 'a1 t -> positive -> 'a1 **) +(** val find : 'a1 -> 'a1 t0 -> positive -> 'a1 **) let rec find default vm p = match vm with @@ -2159,7 +2202,7 @@ let rec find default vm p = | XO p2 -> find default l p2 | XH -> e) -(** val singleton : 'a1 -> positive -> 'a1 -> 'a1 t **) +(** val singleton : 'a1 -> positive -> 'a1 -> 'a1 t0 **) let rec singleton default x v = match x with @@ -2167,7 +2210,7 @@ let rec singleton default x v = | XO p -> Branch ((singleton default p v), default, Empty) | XH -> Elt v -(** val vm_add : 'a1 -> positive -> 'a1 -> 'a1 t -> 'a1 t **) +(** val vm_add : 'a1 -> positive -> 'a1 -> 'a1 t0 -> 'a1 t0 **) let rec vm_add default x v = function | Empty -> singleton default x v @@ -2182,28 +2225,6 @@ let rec vm_add default x v = function | XO p -> Branch ((vm_add default p v l), o, r) | XH -> Branch (l, v, r)) -(** val zeval_const : z pExpr -> z option **) - -let rec zeval_const = function -| PEc c -> Some c -| PEX _ -> None -| PEadd (e1, e2) -> - map_option2 (fun x y -> Some (Z.add x y)) (zeval_const e1) (zeval_const e2) -| PEsub (e1, e2) -> - map_option2 (fun x y -> Some (Z.sub x y)) (zeval_const e1) (zeval_const e2) -| PEmul (e1, e2) -> - map_option2 (fun x y -> Some (Z.mul x y)) (zeval_const e1) (zeval_const e2) -| PEopp e0 -> map_option (fun x -> Some (Z.opp x)) (zeval_const e0) -| PEpow (e1, n0) -> - map_option (fun x -> Some (Z.pow x (Z.of_N n0))) (zeval_const e1) - -type zWitness = z psatz - -(** val zWeakChecker : z nFormula list -> z psatz -> bool **) - -let zWeakChecker = - check_normalised_formulas Z0 (Zpos XH) Z.add Z.mul Z.eqb Z.leb - (** val psub1 : z pol -> z pol -> z pol **) let psub1 = @@ -2219,9 +2240,9 @@ let popp1 = let padd1 = padd0 Z0 Z.add Z.eqb -(** val normZ : z pExpr -> z pol **) +(** val norm0 : z pExpr -> z pol **) -let normZ = +let norm0 = norm Z0 (Zpos XH) Z.add Z.mul Z.sub Z.opp Z.eqb (** val zunsat : z nFormula -> bool **) @@ -2236,10 +2257,10 @@ let zdeduce = (** val xnnormalise : z formula -> z nFormula **) -let xnnormalise t0 = - let { flhs = lhs; fop = o; frhs = rhs } = t0 in - let lhs0 = normZ lhs in - let rhs0 = normZ rhs in +let xnnormalise t1 = + let { flhs = lhs; fop = o; frhs = rhs } = t1 in + let lhs0 = norm0 lhs in + let rhs0 = norm0 rhs in (match o with | OpEq -> (psub1 rhs0 lhs0),Equal | OpNEq -> (psub1 rhs0 lhs0),NonEqual @@ -2248,52 +2269,33 @@ let xnnormalise t0 = | OpLt -> (psub1 rhs0 lhs0),Strict | OpGt -> (psub1 lhs0 rhs0),Strict) -(** val xnormalise0 : z nFormula -> z nFormula list **) - -let xnormalise0 = function -| e,o -> - (match o with - | Equal -> - ((psub1 e (Pc (Zpos XH))),NonStrict)::(((psub1 (Pc (Zneg XH)) e),NonStrict)::[]) - | NonEqual -> (e,Equal)::[] - | Strict -> ((psub1 (Pc Z0) e),NonStrict)::[] - | NonStrict -> ((psub1 (Pc (Zneg XH)) e),NonStrict)::[]) - -(** val cnf_of_list0 : - 'a1 -> z nFormula list -> (z nFormula * 'a1) list list **) - -let cnf_of_list0 tg l = - fold_right (fun x acc -> if zunsat x then acc else ((x,tg)::[])::acc) - cnf_tt l +(** val xis_integral : (positive -> bool) -> positive -> z pol -> bool **) -(** val normalise0 : z formula -> 'a1 -> (z nFormula, 'a1) cnf **) - -let normalise0 t0 tg = - let f = xnnormalise t0 in - if zunsat f then cnf_ff else cnf_of_list0 tg (xnormalise0 f) +let rec xis_integral e jmp = function +| Pc _ -> true +| Pinj (j, p2) -> xis_integral e (Coq_Pos.add j jmp) p2 +| PX (p2, _, q0) -> + (&&) ((&&) (xis_integral e jmp p2) (e jmp)) + (xis_integral e (Coq_Pos.succ jmp) q0) -(** val xnegate0 : z nFormula -> z nFormula list **) +(** val get : PositiveSet.t -> positive -> bool **) -let xnegate0 = function -| e,o -> - (match o with - | NonEqual -> - ((psub1 e (Pc (Zpos XH))),NonStrict)::(((psub1 (Pc (Zneg XH)) e),NonStrict)::[]) - | Strict -> ((psub1 e (Pc (Zpos XH))),NonStrict)::[] - | x -> (e,x)::[]) +let get s x = + PositiveSet.mem x s -(** val negate : z formula -> 'a1 -> (z nFormula, 'a1) cnf **) +(** val oget : PositiveSet.t option -> positive -> bool **) -let negate t0 tg = - let f = xnnormalise t0 in - if zunsat f then cnf_tt else cnf_of_list0 tg (xnegate0 f) +let oget s x = + match s with + | Some s0 -> PositiveSet.mem x s0 + | None -> true -(** val cnfZ : - kind -> (z formula, 'a1, 'a2, 'a3) tFormula -> (z nFormula, 'a1) - cnf * 'a1 trace **) +(** val is_integral : PositiveSet.t option -> z pol -> bool **) -let cnfZ k f = - rxcnf zunsat zdeduce normalise0 negate true k f +let is_integral e pol0 = + match e with + | Some e0 -> xis_integral (get e0) XH pol0 + | None -> true (** val ceiling : z -> z -> z **) @@ -2303,6 +2305,8 @@ let ceiling a b = | Z0 -> q0 | _ -> Z.add q0 (Zpos XH)) +type zWitness = z psatz + type zArithProof = | DoneProof | RatProof of zWitness * zArithProof @@ -2377,9 +2381,8 @@ let eval_Psatz0 = (** val valid_cut_sign : op1 -> bool **) let valid_cut_sign = function -| Equal -> true -| NonStrict -> true -| _ -> false +| NonEqual -> false +| _ -> true (** val bound_var : positive -> z formula **) @@ -2388,8 +2391,8 @@ let bound_var v = (** val mk_eq_pos : positive -> positive -> positive -> z formula **) -let mk_eq_pos x y t0 = - { flhs = (PEX x); fop = OpEq; frhs = (PEsub ((PEX y), (PEX t0))) } +let mk_eq_pos x y t1 = + { flhs = (PEX x); fop = OpEq; frhs = (PEsub ((PEX y), (PEX t1))) } (** val max_var : positive -> z pol -> positive **) @@ -2404,73 +2407,164 @@ let rec max_var jmp = function let max_var_nformulae l = fold_left (fun acc f -> Coq_Pos.max acc (max_var XH (fst f))) l XH -(** val zChecker : z nFormula list -> zArithProof -> bool **) +(** val add0 : positive -> PositiveSet.t option -> PositiveSet.t option **) + +let add0 x = function +| Some s0 -> Some (PositiveSet.add x s0) +| None -> None + +(** val zChecker : + PositiveSet.t option -> z nFormula list -> zArithProof -> bool **) -let rec zChecker l = function +let rec zChecker is_integer l = function | DoneProof -> false | RatProof (w, pf0) -> (match eval_Psatz0 l w with - | Some f -> if zunsat f then true else zChecker (f::l) pf0 + | Some f -> if zunsat f then true else zChecker is_integer (f::l) pf0 | None -> false) | CutProof (w, pf0) -> (match eval_Psatz0 l w with | Some f -> - (match genCuttingPlane f with - | Some cp -> zChecker ((nformula_of_cutting_plane cp)::l) pf0 - | None -> true) + if is_integral is_integer (fst f) + then (match genCuttingPlane f with + | Some cp -> + let f0 = nformula_of_cutting_plane cp in + if zunsat f0 then true else zChecker is_integer (f0::l) pf0 + | None -> true) + else false | None -> false) | SplitProof (p, pf1, pf2) -> - (match genCuttingPlane (p,NonStrict) with - | Some cp1 -> - (match genCuttingPlane ((popp1 p),NonStrict) with - | Some cp2 -> - (&&) (zChecker ((nformula_of_cutting_plane cp1)::l) pf1) - (zChecker ((nformula_of_cutting_plane cp2)::l) pf2) - | None -> false) - | None -> false) + if is_integral is_integer p + then (match genCuttingPlane (p,NonStrict) with + | Some cp1 -> + (match genCuttingPlane ((popp1 p),NonStrict) with + | Some cp2 -> + (&&) + (zChecker is_integer ((nformula_of_cutting_plane cp1)::l) pf1) + (zChecker is_integer ((nformula_of_cutting_plane cp2)::l) pf2) + | None -> false) + | None -> false) + else false | EnumProof (w1, w2, pf0) -> (match eval_Psatz0 l w1 with | Some f1 -> (match eval_Psatz0 l w2 with | Some f2 -> - (match genCuttingPlane f1 with - | Some p -> - let p2,op3 = p in - let e1,z1 = p2 in - (match genCuttingPlane f2 with - | Some p3 -> - let p4,op4 = p3 in - let e2,z2 = p4 in - if (&&) ((&&) (valid_cut_sign op3) (valid_cut_sign op4)) - (is_pol_Z0 (padd1 e1 e2)) - then let rec label pfs lb ub = - match pfs with - | [] -> Z.gtb lb ub - | pf1::rsr -> - (&&) (zChecker (((psub1 e1 (Pc lb)),Equal)::l) pf1) - (label rsr (Z.add lb (Zpos XH)) ub) - in label pf0 (Z.opp z1) z2 - else false - | None -> true) - | None -> true) + if (&&) (is_integral is_integer (fst f1)) + (is_integral is_integer (fst f2)) + then (match genCuttingPlane f1 with + | Some p -> + let p2,op3 = p in + let e1,z1 = p2 in + (match genCuttingPlane f2 with + | Some p3 -> + let p4,op4 = p3 in + let e2,z2 = p4 in + if (&&) ((&&) (valid_cut_sign op3) (valid_cut_sign op4)) + (is_pol_Z0 (padd1 e1 e2)) + then let rec label pfs lb ub = + match pfs with + | [] -> Z.gtb lb ub + | pf1::rsr -> + (&&) + (zChecker is_integer + (((psub1 e1 (Pc lb)),Equal)::l) pf1) + (label rsr (Z.add lb (Zpos XH)) ub) + in label pf0 (Z.opp z1) z2 + else false + | None -> true) + | None -> true) + else false | None -> false) | None -> false) | ExProof (x, prf) -> let fr = max_var_nformulae l in - if Coq_Pos.leb x fr + if (&&) (Coq_Pos.leb x fr) (oget is_integer x) then let z0 = Coq_Pos.succ fr in - let t0 = Coq_Pos.succ z0 in - let nfx = xnnormalise (mk_eq_pos x z0 t0) in + let t1 = Coq_Pos.succ z0 in + let nfx = xnnormalise (mk_eq_pos x z0 t1) in let posz = xnnormalise (bound_var z0) in - let post = xnnormalise (bound_var t0) in - zChecker (nfx::(posz::(post::l))) prf + let post = xnnormalise (bound_var t1) in + zChecker (add0 z0 (add0 t1 is_integer)) (nfx::(posz::(post::l))) prf else false +(** val zeval_const : z pExpr -> z option **) + +let rec zeval_const = function +| PEc c -> Some c +| PEX _ -> None +| PEadd (e1, e2) -> + map_option2 (fun x y -> Some (Z.add x y)) (zeval_const e1) (zeval_const e2) +| PEsub (e1, e2) -> + map_option2 (fun x y -> Some (Z.sub x y)) (zeval_const e1) (zeval_const e2) +| PEmul (e1, e2) -> + map_option2 (fun x y -> Some (Z.mul x y)) (zeval_const e1) (zeval_const e2) +| PEopp e0 -> map_option (fun x -> Some (Z.opp x)) (zeval_const e0) +| PEpow (e1, n0) -> + map_option (fun x -> Some (Z.pow x (Z.of_N n0))) (zeval_const e1) + +(** val zWeakChecker : z nFormula list -> z psatz -> bool **) + +let zWeakChecker = + check_normalised_formulas Z0 (Zpos XH) Z.add Z.mul Z.eqb Z.leb + +(** val normZ : z pExpr -> z pol **) + +let normZ = + norm Z0 (Zpos XH) Z.add Z.mul Z.sub Z.opp Z.eqb + +(** val cnf_of_list0 : + 'a1 -> z nFormula list -> (z nFormula * 'a1) list list **) + +let cnf_of_list0 tg l = + fold_right (fun x acc -> if zunsat x then acc else ((x,tg)::[])::acc) + cnf_tt l + +(** val xnormalise0 : z nFormula -> z nFormula list **) + +let xnormalise0 = function +| e,o -> + (match o with + | Equal -> + ((psub1 e (Pc (Zpos XH))),NonStrict)::(((psub1 (Pc (Zneg XH)) e),NonStrict)::[]) + | NonEqual -> (e,Equal)::[] + | Strict -> ((psub1 (Pc Z0) e),NonStrict)::[] + | NonStrict -> ((psub1 (Pc (Zneg XH)) e),NonStrict)::[]) + +(** val normalise0 : z formula -> 'a1 -> (z nFormula, 'a1) cnf **) + +let normalise0 t1 tg = + let f = xnnormalise t1 in + if zunsat f then cnf_ff else cnf_of_list0 tg (xnormalise0 f) + +(** val xnegate0 : z nFormula -> z nFormula list **) + +let xnegate0 = function +| e,o -> + (match o with + | NonEqual -> + ((psub1 e (Pc (Zpos XH))),NonStrict)::(((psub1 (Pc (Zneg XH)) e),NonStrict)::[]) + | Strict -> ((psub1 e (Pc (Zpos XH))),NonStrict)::[] + | x -> (e,x)::[]) + +(** val negate : z formula -> 'a1 -> (z nFormula, 'a1) cnf **) + +let negate t1 tg = + let f = xnnormalise t1 in + if zunsat f then cnf_tt else cnf_of_list0 tg (xnegate0 f) + +(** val cnfZ : + kind -> (z formula, 'a1, 'a2, 'a3) tFormula -> (z nFormula, 'a1) + cnf * 'a1 trace **) + +let cnfZ k f = + rxcnf zunsat zdeduce normalise0 negate true k f + (** val zTautoChecker : z formula bFormula -> zArithProof list -> bool **) let zTautoChecker f w = tauto_checker zunsat zdeduce normalise0 negate (fun cl -> - zChecker (map fst cl)) f w + zChecker None (map fst cl)) f w type q = { qnum : z; qden : positive } @@ -2535,15 +2629,15 @@ let qWeakChecker = (** val qnormalise : q formula -> 'a1 -> (q nFormula, 'a1) cnf **) -let qnormalise t0 tg = +let qnormalise t1 tg = cnf_normalise { qnum = Z0; qden = XH } { qnum = (Zpos XH); qden = XH } - qplus qmult qminus qopp qeq_bool qle_bool t0 tg + qplus qmult qminus qopp qeq_bool qle_bool t1 tg (** val qnegate : q formula -> 'a1 -> (q nFormula, 'a1) cnf **) -let qnegate t0 tg = +let qnegate t1 tg = cnf_negate { qnum = Z0; qden = XH } { qnum = (Zpos XH); qden = XH } qplus - qmult qminus qopp qeq_bool qle_bool t0 tg + qmult qminus qopp qeq_bool qle_bool t1 tg (** val qunsat : q nFormula -> bool **) @@ -2606,25 +2700,19 @@ let rec q_of_Rcst = function | CInv r0 -> qinv (q_of_Rcst r0) | COpp r0 -> qopp (q_of_Rcst r0) -type rWitness = q psatz - -(** val rWeakChecker : q nFormula list -> q psatz -> bool **) - -let rWeakChecker = - check_normalised_formulas { qnum = Z0; qden = XH } { qnum = (Zpos XH); - qden = XH } qplus qmult qeq_bool qle_bool +type rWitness = zArithProof (** val rnormalise : q formula -> 'a1 -> (q nFormula, 'a1) cnf **) -let rnormalise t0 tg = +let rnormalise t1 tg = cnf_normalise { qnum = Z0; qden = XH } { qnum = (Zpos XH); qden = XH } - qplus qmult qminus qopp qeq_bool qle_bool t0 tg + qplus qmult qminus qopp qeq_bool qle_bool t1 tg (** val rnegate : q formula -> 'a1 -> (q nFormula, 'a1) cnf **) -let rnegate t0 tg = +let rnegate t1 tg = cnf_negate { qnum = Z0; qden = XH } { qnum = (Zpos XH); qden = XH } qplus - qmult qminus qopp qeq_bool qle_bool t0 tg + qmult qminus qopp qeq_bool qle_bool t1 tg (** val runsat : q nFormula -> bool **) @@ -2636,10 +2724,105 @@ let runsat = let rdeduce = nformula_plus_nformula { qnum = Z0; qden = XH } qplus qeq_bool -(** val rTautoChecker : rcst formula bFormula -> rWitness list -> bool **) +type 'form eFormula = +| IsZ of bool * positive +| IsF of 'form + +(** val map_eFormula : + ('a1 -> 'a2) -> 'a1 formula eFormula -> 'a2 formula eFormula **) + +let map_eFormula f = function +| IsZ (b, x) -> IsZ (b, x) +| IsF f1 -> IsF (map_Formula f f1) + +(** val lcm0 : q pol -> z **) + +let rec lcm0 = function +| Pc q0 -> Zpos q0.qden +| Pinj (_, p2) -> lcm0 p2 +| PX (p2, _, q0) -> Z.lcm (lcm0 p2) (lcm0 q0) + +(** val polZ : z -> q pol -> z pol **) + +let rec polZ lcm1 = function +| Pc q0 -> Pc (Z.div (Z.mul q0.qnum lcm1) (Zpos q0.qden)) +| Pinj (x, p2) -> Pinj (x, (polZ lcm1 p2)) +| PX (p2, i, q0) -> PX ((polZ lcm1 p2), i, (polZ lcm1 q0)) + +(** val nformulaZ : q nFormula -> z nFormula **) + +let nformulaZ = function +| p,o -> (polZ (lcm0 p) p),o + +(** val xcollect_isZ : + PositiveSet.t -> z nFormula list -> (q nFormula eFormula * unit0) list -> + PositiveSet.t * z nFormula list **) + +let rec xcollect_isZ s acc = function +| [] -> s,acc +| p::l0 -> + let e,_ = p in + (match e with + | IsZ (b, x) -> xcollect_isZ (if b then PositiveSet.add x s else s) acc l0 + | IsF f -> xcollect_isZ s ((nformulaZ f)::acc) l0) + +(** val qCheck : (q nFormula eFormula * unit0) list -> zArithProof -> bool **) + +let qCheck l = + let s,cl = xcollect_isZ PositiveSet.empty [] l in zChecker (Some s) cl + +(** val erunsat : q nFormula eFormula -> bool **) + +let erunsat = function +| IsZ (_, _) -> false +| IsF f0 -> runsat f0 + +(** val erdeduce : + q nFormula eFormula -> q nFormula eFormula -> q nFormula eFormula option **) + +let erdeduce f1 f2 = + match f1 with + | IsZ (_, _) -> None + | IsF f3 -> + (match f2 with + | IsZ (_, _) -> None + | IsF f4 -> + (match rdeduce f3 f4 with + | Some f -> Some (IsF f) + | None -> None)) + +(** val map_cnf : ('a2 -> 'a3) -> ('a2, 'a1) cnf -> ('a3, 'a1) cnf **) + +let map_cnf f l = + map (map (fun pat -> let a,t1 = pat in (f a),t1)) l + +(** val eRnormalise : + q formula eFormula -> 'a1 -> (q nFormula eFormula, 'a1) cnf **) + +let eRnormalise f t1 = + match f with + | IsZ (b, z0) -> (((IsZ ((negb b), z0)),t1)::[])::[] + | IsF f0 -> map_cnf (fun x -> IsF x) (rnormalise f0 t1) + +(** val eRnegate : + q formula eFormula -> 'a1 -> (q nFormula eFormula, 'a1) cnf **) + +let eRnegate f t1 = + match f with + | IsZ (b, z0) -> (((IsZ (b, z0)),t1)::[])::[] + | IsF f0 -> map_cnf (fun x -> IsF x) (rnegate f0 t1) + +(** val cnfR : + kind -> (q formula eFormula, 'a1, 'a2, 'a3) tFormula -> (q nFormula + eFormula, 'a1) cnf * 'a1 trace **) + +let cnfR k f = + rxcnf erunsat erdeduce eRnormalise eRnegate true k f + +(** val rTautoChecker : + rcst formula eFormula bFormula -> rWitness list -> bool **) let rTautoChecker f w = - tauto_checker runsat rdeduce rnormalise rnegate (fun cl -> - rWeakChecker (map fst cl)) - (map_bformula IsProp (map_Formula q_of_Rcst) f) w + tauto_checker erunsat erdeduce eRnormalise eRnegate qCheck + (map_bformula IsProp (map_eFormula q_of_Rcst) f) w diff --git a/test-suite/success/tify.v b/test-suite/success/tify.v index 8a2c1cd551..e354929dd4 100644 --- a/test-suite/success/tify.v +++ b/test-suite/success/tify.v @@ -33,7 +33,7 @@ Definition isZ (r:R) := IZR (Zfloor r) = r. Instance Inj_nat_R : InjTyp nat R. Proof. - apply (mkinj _ _ INR (fun r => isZ r /\ 0 <= r)). + apply (mkinj _ _ INR (fun r => isZ r /\ 0 <= r)%R). - intros. split. unfold isZ. rewrite INR_IZR_INZ. rewrite ZfloorZ. reflexivity. From f0b5f5b517853fad0d5507bb8168b4416d561769 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fr=C3=A9d=C3=A9ric=20Besson?= Date: Fri, 28 Mar 2025 17:39:19 +0100 Subject: [PATCH 6/6] [fix] for backward compat --- theories/micromega/ZMicromega.v | 3 +++ theories/micromega/ZifyInst.v | 1 + 2 files changed, 4 insertions(+) diff --git a/theories/micromega/ZMicromega.v b/theories/micromega/ZMicromega.v index f68d3f82d7..96f04658fa 100644 --- a/theories/micromega/ZMicromega.v +++ b/theories/micromega/ZMicromega.v @@ -576,6 +576,9 @@ Qed. Definition cnfZ (Annot: Type) (TX : Tauto.kind -> Type) (AF : Type) (k: Tauto.kind) (f : TFormula (Formula Z) Annot TX AF k) := rxcnf Zunsat Zdeduce normalise negate true f. +#[deprecated(since="9.1", note="Use MMicromega.ZArithProof instead.")] +Definition ZArithProof := ZArithProof. + Definition ZTautoChecker (f : BFormula (Formula Z) Tauto.isProp) (w: list ZArithProof): bool := @tauto_checker (Formula Z) (NFormula Z) unit Zunsat Zdeduce normalise negate ZArithProof (fun cl => ZChecker None (List.map fst cl)) f w. diff --git a/theories/micromega/ZifyInst.v b/theories/micromega/ZifyInst.v index 43c514b893..505a270a85 100644 --- a/theories/micromega/ZifyInst.v +++ b/theories/micromega/ZifyInst.v @@ -14,6 +14,7 @@ From Stdlib Require Import Arith BinInt BinNat Zeven Znat Nnat. From Stdlib Require Import TifyClasses. +From Stdlib Require Import ZifyClasses. (* for backward compatibility - to be removed *) Declare ML Module "rocq-runtime.plugins.zify". Local Open Scope Z_scope.