diff options
author | Samuel Mimram <smimram@debian.org> | 2008-07-25 15:12:53 +0200 |
---|---|---|
committer | Samuel Mimram <smimram@debian.org> | 2008-07-25 15:12:53 +0200 |
commit | a0cfa4f118023d35b767a999d5a2ac4b082857b4 (patch) | |
tree | dabcac548e299fee1da464c93b3dba98484f45b1 /contrib/micromega | |
parent | 2281410e38ef99d025ea77194585a9bc019fdaa9 (diff) |
Imported Upstream version 8.2~beta3+dfsgupstream/8.2.beta3+dfsg
Diffstat (limited to 'contrib/micromega')
26 files changed, 12928 insertions, 0 deletions
diff --git a/contrib/micromega/CheckerMaker.v b/contrib/micromega/CheckerMaker.v new file mode 100644 index 00000000..93b4d213 --- /dev/null +++ b/contrib/micromega/CheckerMaker.v @@ -0,0 +1,129 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* *) +(* Micromega: A reflexive tactic using the Positivstellensatz *) +(* *) +(* Frédéric Besson (Irisa/Inria) 2006-2008 *) +(* *) +(************************************************************************) + +Require Import Setoid. +Require Import Decidable. +Require Import List. +Require Import Refl. + +Set Implicit Arguments. + +Section CheckerMaker. + +(* 'Formula' is a syntactic representation of a certain kind of propositions. *) +Variable Formula : Type. + +Variable Env : Type. + +Variable eval : Env -> Formula -> Prop. + +Variable Formula' : Type. + +Variable eval' : Env -> Formula' -> Prop. + +Variable normalise : Formula -> Formula'. + +Variable negate : Formula -> Formula'. + +Hypothesis normalise_sound : + forall (env : Env) (t : Formula), eval env t -> eval' env (normalise t). + +Hypothesis negate_correct : + forall (env : Env) (t : Formula), eval env t <-> ~ (eval' env (negate t)). + +Variable Witness : Type. + +Variable check_formulas' : list Formula' -> Witness -> bool. + +Hypothesis check_formulas'_sound : + forall (l : list Formula') (w : Witness), + check_formulas' l w = true -> + forall env : Env, make_impl (eval' env) l False. + +Definition normalise_list : list Formula -> list Formula' := map normalise. +Definition negate_list : list Formula -> list Formula' := map negate. + +Definition check_formulas (l : list Formula) (w : Witness) : bool := + check_formulas' (map normalise l) w. + +(* Contraposition of normalise_sound for lists *) +Lemma normalise_sound_contr : forall (env : Env) (l : list Formula), + make_impl (eval' env) (map normalise l) False -> make_impl (eval env) l False. +Proof. +intros env l; induction l as [| t l IH]; simpl in *. +trivial. +intros H1 H2. apply IH. apply H1. now apply normalise_sound. +Qed. + +Theorem check_formulas_sound : + forall (l : list Formula) (w : Witness), + check_formulas l w = true -> forall env : Env, make_impl (eval env) l False. +Proof. +unfold check_formulas; intros l w H env. destruct l as [| t l]; simpl in *. +pose proof (check_formulas'_sound H env) as H1; now simpl in H1. +intro H1. apply normalise_sound in H1. +pose proof (check_formulas'_sound H env) as H2; simpl in H2. +apply H2 in H1. now apply normalise_sound_contr. +Qed. + +(* In check_conj_formulas', t2 is supposed to be a list of negations of +formulas. If, for example, t1 = [A1, A2] and t2 = [~ B1, ~ B2], then +check_conj_formulas' checks that each of [~ B1, A1, A2] and [~ B2, A1, A2] is +inconsistent. This means that A1 /\ A2 -> B1 and A1 /\ A2 -> B1, i.e., that +A1 /\ A2 -> B1 /\ B2. *) + +Fixpoint check_conj_formulas' + (t1 : list Formula') (wits : list Witness) (t2 : list Formula') {struct wits} : bool := +match t2 with +| nil => true +| t':: rt2 => + match wits with + | nil => false + | w :: rwits => + match check_formulas' (t':: t1) w with + | true => check_conj_formulas' t1 rwits rt2 + | false => false + end + end +end. + +(* checks whether the conjunction of t1 implies the conjunction of t2 *) + +Definition check_conj_formulas + (t1 : list Formula) (wits : list Witness) (t2 : list Formula) : bool := + check_conj_formulas' (normalise_list t1) wits (negate_list t2). + +Theorem check_conj_formulas_sound : + forall (t1 : list Formula) (t2 : list Formula) (wits : list Witness), + check_conj_formulas t1 wits t2 = true -> + forall env : Env, make_impl (eval env) t1 (make_conj (eval env) t2). +Proof. +intro t1; induction t2 as [| a2 t2' IH]. +intros; apply make_impl_true. +intros wits H env. +unfold check_conj_formulas in H; simpl in H. +destruct wits as [| w ws]; simpl in H. discriminate. +case_eq (check_formulas' (negate a2 :: normalise_list t1) w); +intro H1; rewrite H1 in H; [| discriminate]. +assert (H2 : make_impl (eval' env) (negate a2 :: normalise_list t1) False) by +now apply check_formulas'_sound with (w := w). clear H1. +pose proof (IH ws H env) as H1. simpl in H2. +assert (H3 : eval' env (negate a2) -> make_impl (eval env) t1 False) +by auto using normalise_sound_contr. clear H2. +rewrite <- make_conj_impl in *. +rewrite make_conj_cons. intro H2. split. +apply <- negate_correct. intro; now elim H3. exact (H1 H2). +Qed. + +End CheckerMaker. diff --git a/contrib/micromega/Env.v b/contrib/micromega/Env.v new file mode 100644 index 00000000..40db9e46 --- /dev/null +++ b/contrib/micromega/Env.v @@ -0,0 +1,182 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* *) +(* Micromega: A reflexive tactic using the Positivstellensatz *) +(* *) +(* Frédéric Besson (Irisa/Inria) 2006-2008 *) +(* *) +(************************************************************************) + +Require Import ZArith. +Require Import Coq.Arith.Max. +Require Import List. +Set Implicit Arguments. + +(* I have addded a Leaf constructor to the varmap data structure (/contrib/ring/Quote.v) + -- this is harmless and spares a lot of Empty. + This means smaller proof-terms. + BTW, by dropping the polymorphism, I get small (yet noticeable) speed-up. +*) + +Section S. + + Variable D :Type. + + Definition Env := positive -> D. + + Definition jump (j:positive) (e:Env) := fun x => e (Pplus x j). + + Definition nth (n:positive) (e : Env ) := e n. + + Definition hd (x:D) (e: Env) := nth xH e. + + Definition tail (e: Env) := jump xH e. + + Lemma psucc : forall p, (match p with + | xI y' => xO (Psucc y') + | xO y' => xI y' + | 1%positive => 2%positive + end) = (p+1)%positive. + Proof. + destruct p. + auto with zarith. + rewrite xI_succ_xO. + auto with zarith. + reflexivity. + Qed. + + Lemma jump_Pplus : forall i j l, + forall x, jump (i + j) l x = jump i (jump j l) x. + Proof. + unfold jump. + intros. + rewrite Pplus_assoc. + reflexivity. + Qed. + + Lemma jump_simpl : forall p l, + forall x, jump p l x = + match p with + | xH => tail l x + | xO p => jump p (jump p l) x + | xI p => jump p (jump p (tail l)) x + end. + Proof. + destruct p ; unfold tail ; intros ; repeat rewrite <- jump_Pplus. + (* xI p = p + p + 1 *) + rewrite xI_succ_xO. + rewrite Pplus_diag. + rewrite <- Pplus_one_succ_r. + reflexivity. + (* xO p = p + p *) + rewrite Pplus_diag. + reflexivity. + reflexivity. + Qed. + + Ltac jump_s := + repeat + match goal with + | |- context [jump xH ?e] => rewrite (jump_simpl xH) + | |- context [jump (xO ?p) ?e] => rewrite (jump_simpl (xO p)) + | |- context [jump (xI ?p) ?e] => rewrite (jump_simpl (xI p)) + end. + + Lemma jump_tl : forall j l, forall x, tail (jump j l) x = jump j (tail l) x. + Proof. + unfold tail. + intros. + repeat rewrite <- jump_Pplus. + rewrite Pplus_comm. + reflexivity. + Qed. + + Lemma jump_Psucc : forall j l, + forall x, (jump (Psucc j) l x) = (jump 1 (jump j l) x). + Proof. + intros. + rewrite <- jump_Pplus. + rewrite Pplus_one_succ_r. + rewrite Pplus_comm. + reflexivity. + Qed. + + Lemma jump_Pdouble_minus_one : forall i l, + forall x, (jump (Pdouble_minus_one i) (tail l)) x = (jump i (jump i l)) x. + Proof. + unfold tail. + intros. + repeat rewrite <- jump_Pplus. + rewrite <- Pplus_one_succ_r. + rewrite Psucc_o_double_minus_one_eq_xO. + rewrite Pplus_diag. + reflexivity. + Qed. + + Lemma jump_x0_tail : forall p l, forall x, jump (xO p) (tail l) x = jump (xI p) l x. + Proof. + intros. + unfold jump. + unfold tail. + unfold jump. + rewrite <- Pplus_assoc. + simpl. + reflexivity. + Qed. + + Lemma nth_spec : forall p l x, + nth p l = + match p with + | xH => hd x l + | xO p => nth p (jump p l) + | xI p => nth p (jump p (tail l)) + end. + Proof. + unfold nth. + destruct p. + intros. + unfold jump, tail. + unfold jump. + rewrite Pplus_diag. + rewrite xI_succ_xO. + simpl. + reflexivity. + unfold jump. + rewrite Pplus_diag. + reflexivity. + unfold hd. + unfold nth. + reflexivity. + Qed. + + + Lemma nth_jump : forall p l x, nth p (tail l) = hd x (jump p l). + Proof. + unfold tail. + unfold hd. + unfold jump. + unfold nth. + intros. + rewrite Pplus_comm. + reflexivity. + Qed. + + Lemma nth_Pdouble_minus_one : + forall p l, nth (Pdouble_minus_one p) (tail l) = nth p (jump p l). + Proof. + intros. + unfold tail. + unfold nth, jump. + rewrite Pplus_diag. + rewrite <- Psucc_o_double_minus_one_eq_xO. + rewrite Pplus_one_succ_r. + reflexivity. + Qed. + +End S. + diff --git a/contrib/micromega/EnvRing.v b/contrib/micromega/EnvRing.v new file mode 100644 index 00000000..04e68272 --- /dev/null +++ b/contrib/micromega/EnvRing.v @@ -0,0 +1,1403 @@ +(************************************************************************) +(* V * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* F. Besson: to evaluate polynomials, the original code is using a list. + For big polynomials, this is inefficient -- linear access. + I have modified the code to use binary trees -- logarithmic access. *) + + +Set Implicit Arguments. +Require Import Setoid. +Require Import BinList. +Require Import Env. +Require Import BinPos. +Require Import BinNat. +Require Import BinInt. +Require Export Ring_theory. + +Open Local Scope positive_scope. +Import RingSyntax. + +Section MakeRingPol. + + (* Ring elements *) + Variable R:Type. + Variable (rO rI : R) (radd rmul rsub: R->R->R) (ropp : R->R). + Variable req : R -> R -> Prop. + + (* Ring properties *) + Variable Rsth : Setoid_Theory R req. + Variable Reqe : ring_eq_ext radd rmul ropp req. + Variable ARth : almost_ring_theory rO rI radd rmul rsub ropp req. + + (* Coefficients *) + Variable C: Type. + Variable (cO cI: C) (cadd cmul csub : C->C->C) (copp : C->C). + Variable ceqb : C->C->bool. + Variable phi : C -> R. + Variable CRmorph : ring_morph rO rI radd rmul rsub ropp req + cO cI cadd cmul csub copp ceqb phi. + + (* Power coefficients *) + Variable Cpow : Set. + Variable Cp_phi : N -> Cpow. + Variable rpow : R -> Cpow -> R. + Variable pow_th : power_theory rI rmul req Cp_phi rpow. + + + (* R notations *) + Notation "0" := rO. Notation "1" := rI. + Notation "x + y" := (radd x y). Notation "x * y " := (rmul x y). + Notation "x - y " := (rsub x y). Notation "- x" := (ropp x). + Notation "x == y" := (req x y). + + (* C notations *) + Notation "x +! y" := (cadd x y). Notation "x *! y " := (cmul x y). + Notation "x -! y " := (csub x y). Notation "-! x" := (copp x). + Notation " x ?=! y" := (ceqb x y). Notation "[ x ]" := (phi x). + + (* Usefull tactics *) + Add Setoid R req Rsth as R_set1. + Ltac rrefl := gen_reflexivity Rsth. + Add Morphism radd : radd_ext. exact (Radd_ext Reqe). Qed. + Add Morphism rmul : rmul_ext. exact (Rmul_ext Reqe). Qed. + Add Morphism ropp : ropp_ext. exact (Ropp_ext Reqe). Qed. + Add Morphism rsub : rsub_ext. exact (ARsub_ext Rsth Reqe ARth). Qed. + Ltac rsimpl := gen_srewrite Rsth Reqe ARth. + Ltac add_push := gen_add_push radd Rsth Reqe ARth. + Ltac mul_push := gen_mul_push rmul Rsth Reqe ARth. + + (* Definition of multivariable polynomials with coefficients in C : + Type [Pol] represents [X1 ... Xn]. + The representation is Horner's where a [n] variable polynomial + (C[X1..Xn]) is seen as a polynomial on [X1] which coefficients + are polynomials with [n-1] variables (C[X2..Xn]). + There are several optimisations to make the repr compacter: + - [Pc c] is the constant polynomial of value c + == c*X1^0*..*Xn^0 + - [Pinj j Q] is a polynomial constant w.r.t the [j] first variables. + variable indices are shifted of j in Q. + == X1^0 *..* Xj^0 * Q{X1 <- Xj+1;..; Xn-j <- Xn} + - [PX P i Q] is an optimised Horner form of P*X^i + Q + with P not the null polynomial + == P * X1^i + Q{X1 <- X2; ..; Xn-1 <- Xn} + + In addition: + - polynomials of the form (PX (PX P i (Pc 0)) j Q) are forbidden + since they can be represented by the simpler form (PX P (i+j) Q) + - (Pinj i (Pinj j P)) is (Pinj (i+j) P) + - (Pinj i (Pc c)) is (Pc c) + *) + + Inductive Pol : Type := + | Pc : C -> Pol + | Pinj : positive -> Pol -> Pol + | PX : Pol -> positive -> Pol -> Pol. + + Definition P0 := Pc cO. + Definition P1 := Pc cI. + + Fixpoint Peq (P P' : Pol) {struct P'} : bool := + match P, P' with + | Pc c, Pc c' => c ?=! c' + | Pinj j Q, Pinj j' Q' => + match Pcompare j j' Eq with + | Eq => Peq Q Q' + | _ => false + end + | PX P i Q, PX P' i' Q' => + match Pcompare i i' Eq with + | Eq => if Peq P P' then Peq Q Q' else false + | _ => false + end + | _, _ => false + end. + + Notation " P ?== P' " := (Peq P P'). + + Definition mkPinj j P := + match P with + | Pc _ => P + | Pinj j' Q => Pinj ((j + j'):positive) Q + | _ => Pinj j P + end. + + Definition mkPinj_pred j P:= + match j with + | xH => P + | xO j => Pinj (Pdouble_minus_one j) P + | xI j => Pinj (xO j) P + end. + + Definition mkPX P i Q := + match P with + | Pc c => if c ?=! cO then mkPinj xH Q else PX P i Q + | Pinj _ _ => PX P i Q + | PX P' i' Q' => if Q' ?== P0 then PX P' (i' + i) Q else PX P i Q + end. + + Definition mkXi i := PX P1 i P0. + + Definition mkX := mkXi 1. + + (** Opposite of addition *) + + Fixpoint Popp (P:Pol) : Pol := + match P with + | Pc c => Pc (-! c) + | Pinj j Q => Pinj j (Popp Q) + | PX P i Q => PX (Popp P) i (Popp Q) + end. + + Notation "-- P" := (Popp P). + + (** Addition et subtraction *) + + Fixpoint PaddC (P:Pol) (c:C) {struct P} : Pol := + match P with + | Pc c1 => Pc (c1 +! c) + | Pinj j Q => Pinj j (PaddC Q c) + | PX P i Q => PX P i (PaddC Q c) + end. + + Fixpoint PsubC (P:Pol) (c:C) {struct P} : Pol := + match P with + | Pc c1 => Pc (c1 -! c) + | Pinj j Q => Pinj j (PsubC Q c) + | PX P i Q => PX P i (PsubC Q c) + end. + + Section PopI. + + Variable Pop : Pol -> Pol -> Pol. + Variable Q : Pol. + + Fixpoint PaddI (j:positive) (P:Pol){struct P} : Pol := + match P with + | Pc c => mkPinj j (PaddC Q c) + | Pinj j' Q' => + match ZPminus j' j with + | Zpos k => mkPinj j (Pop (Pinj k Q') Q) + | Z0 => mkPinj j (Pop Q' Q) + | Zneg k => mkPinj j' (PaddI k Q') + end + | PX P i Q' => + match j with + | xH => PX P i (Pop Q' Q) + | xO j => PX P i (PaddI (Pdouble_minus_one j) Q') + | xI j => PX P i (PaddI (xO j) Q') + end + end. + + Fixpoint PsubI (j:positive) (P:Pol){struct P} : Pol := + match P with + | Pc c => mkPinj j (PaddC (--Q) c) + | Pinj j' Q' => + match ZPminus j' j with + | Zpos k => mkPinj j (Pop (Pinj k Q') Q) + | Z0 => mkPinj j (Pop Q' Q) + | Zneg k => mkPinj j' (PsubI k Q') + end + | PX P i Q' => + match j with + | xH => PX P i (Pop Q' Q) + | xO j => PX P i (PsubI (Pdouble_minus_one j) Q') + | xI j => PX P i (PsubI (xO j) Q') + end + end. + + Variable P' : Pol. + + Fixpoint PaddX (i':positive) (P:Pol) {struct P} : Pol := + match P with + | Pc c => PX P' i' P + | Pinj j Q' => + match j with + | xH => PX P' i' Q' + | xO j => PX P' i' (Pinj (Pdouble_minus_one j) Q') + | xI j => PX P' i' (Pinj (xO j) Q') + end + | PX P i Q' => + match ZPminus i i' with + | Zpos k => mkPX (Pop (PX P k P0) P') i' Q' + | Z0 => mkPX (Pop P P') i Q' + | Zneg k => mkPX (PaddX k P) i Q' + end + end. + + Fixpoint PsubX (i':positive) (P:Pol) {struct P} : Pol := + match P with + | Pc c => PX (--P') i' P + | Pinj j Q' => + match j with + | xH => PX (--P') i' Q' + | xO j => PX (--P') i' (Pinj (Pdouble_minus_one j) Q') + | xI j => PX (--P') i' (Pinj (xO j) Q') + end + | PX P i Q' => + match ZPminus i i' with + | Zpos k => mkPX (Pop (PX P k P0) P') i' Q' + | Z0 => mkPX (Pop P P') i Q' + | Zneg k => mkPX (PsubX k P) i Q' + end + end. + + + End PopI. + + Fixpoint Padd (P P': Pol) {struct P'} : Pol := + match P' with + | Pc c' => PaddC P c' + | Pinj j' Q' => PaddI Padd Q' j' P + | PX P' i' Q' => + match P with + | Pc c => PX P' i' (PaddC Q' c) + | Pinj j Q => + match j with + | xH => PX P' i' (Padd Q Q') + | xO j => PX P' i' (Padd (Pinj (Pdouble_minus_one j) Q) Q') + | xI j => PX P' i' (Padd (Pinj (xO j) Q) Q') + end + | PX P i Q => + match ZPminus i i' with + | Zpos k => mkPX (Padd (PX P k P0) P') i' (Padd Q Q') + | Z0 => mkPX (Padd P P') i (Padd Q Q') + | Zneg k => mkPX (PaddX Padd P' k P) i (Padd Q Q') + end + end + end. + Notation "P ++ P'" := (Padd P P'). + + Fixpoint Psub (P P': Pol) {struct P'} : Pol := + match P' with + | Pc c' => PsubC P c' + | Pinj j' Q' => PsubI Psub Q' j' P + | PX P' i' Q' => + match P with + | Pc c => PX (--P') i' (*(--(PsubC Q' c))*) (PaddC (--Q') c) + | Pinj j Q => + match j with + | xH => PX (--P') i' (Psub Q Q') + | xO j => PX (--P') i' (Psub (Pinj (Pdouble_minus_one j) Q) Q') + | xI j => PX (--P') i' (Psub (Pinj (xO j) Q) Q') + end + | PX P i Q => + match ZPminus i i' with + | Zpos k => mkPX (Psub (PX P k P0) P') i' (Psub Q Q') + | Z0 => mkPX (Psub P P') i (Psub Q Q') + | Zneg k => mkPX (PsubX Psub P' k P) i (Psub Q Q') + end + end + end. + Notation "P -- P'" := (Psub P P'). + + (** Multiplication *) + + Fixpoint PmulC_aux (P:Pol) (c:C) {struct P} : Pol := + match P with + | Pc c' => Pc (c' *! c) + | Pinj j Q => mkPinj j (PmulC_aux Q c) + | PX P i Q => mkPX (PmulC_aux P c) i (PmulC_aux Q c) + end. + + Definition PmulC P c := + if c ?=! cO then P0 else + if c ?=! cI then P else PmulC_aux P c. + + Section PmulI. + Variable Pmul : Pol -> Pol -> Pol. + Variable Q : Pol. + Fixpoint PmulI (j:positive) (P:Pol) {struct P} : Pol := + match P with + | Pc c => mkPinj j (PmulC Q c) + | Pinj j' Q' => + match ZPminus j' j with + | Zpos k => mkPinj j (Pmul (Pinj k Q') Q) + | Z0 => mkPinj j (Pmul Q' Q) + | Zneg k => mkPinj j' (PmulI k Q') + end + | PX P' i' Q' => + match j with + | xH => mkPX (PmulI xH P') i' (Pmul Q' Q) + | xO j' => mkPX (PmulI j P') i' (PmulI (Pdouble_minus_one j') Q') + | xI j' => mkPX (PmulI j P') i' (PmulI (xO j') Q') + end + end. + + End PmulI. +(* A symmetric version of the multiplication *) + + Fixpoint Pmul (P P'' : Pol) {struct P''} : Pol := + match P'' with + | Pc c => PmulC P c + | Pinj j' Q' => PmulI Pmul Q' j' P + | PX P' i' Q' => + match P with + | Pc c => PmulC P'' c + | Pinj j Q => + let QQ' := + match j with + | xH => Pmul Q Q' + | xO j => Pmul (Pinj (Pdouble_minus_one j) Q) Q' + | xI j => Pmul (Pinj (xO j) Q) Q' + end in + mkPX (Pmul P P') i' QQ' + | PX P i Q=> + let QQ' := Pmul Q Q' in + let PQ' := PmulI Pmul Q' xH P in + let QP' := Pmul (mkPinj xH Q) P' in + let PP' := Pmul P P' in + (mkPX (mkPX PP' i P0 ++ QP') i' P0) ++ mkPX PQ' i QQ' + end + end. + +(* Non symmetric *) +(* + Fixpoint Pmul_aux (P P' : Pol) {struct P'} : Pol := + match P' with + | Pc c' => PmulC P c' + | Pinj j' Q' => PmulI Pmul_aux Q' j' P + | PX P' i' Q' => + (mkPX (Pmul_aux P P') i' P0) ++ (PmulI Pmul_aux Q' xH P) + end. + + Definition Pmul P P' := + match P with + | Pc c => PmulC P' c + | Pinj j Q => PmulI Pmul_aux Q j P' + | PX P i Q => + (mkPX (Pmul_aux P P') i P0) ++ (PmulI Pmul_aux Q xH P') + end. +*) + Notation "P ** P'" := (Pmul P P'). + + Fixpoint Psquare (P:Pol) : Pol := + match P with + | Pc c => Pc (c *! c) + | Pinj j Q => Pinj j (Psquare Q) + | PX P i Q => + let twoPQ := Pmul P (mkPinj xH (PmulC Q (cI +! cI))) in + let Q2 := Psquare Q in + let P2 := Psquare P in + mkPX (mkPX P2 i P0 ++ twoPQ) i Q2 + end. + + (** Monomial **) + + Inductive Mon: Set := + mon0: Mon + | zmon: positive -> Mon -> Mon + | vmon: positive -> Mon -> Mon. + + Fixpoint Mphi(l:Env R) (M: Mon) {struct M} : R := + match M with + mon0 => rI + | zmon j M1 => Mphi (jump j l) M1 + | vmon i M1 => + let x := hd 0 l in + let xi := pow_pos rmul x i in + (Mphi (tail l) M1) * xi + end. + + Definition mkZmon j M := + match M with mon0 => mon0 | _ => zmon j M end. + + Definition zmon_pred j M := + match j with xH => M | _ => mkZmon (Ppred j) M end. + + Definition mkVmon i M := + match M with + | mon0 => vmon i mon0 + | zmon j m => vmon i (zmon_pred j m) + | vmon i' m => vmon (i+i') m + end. + + Fixpoint MFactor (P: Pol) (M: Mon) {struct P}: Pol * Pol := + match P, M with + _, mon0 => (Pc cO, P) + | Pc _, _ => (P, Pc cO) + | Pinj j1 P1, zmon j2 M1 => + match (j1 ?= j2) Eq with + Eq => let (R,S) := MFactor P1 M1 in + (mkPinj j1 R, mkPinj j1 S) + | Lt => let (R,S) := MFactor P1 (zmon (j2 - j1) M1) in + (mkPinj j1 R, mkPinj j1 S) + | Gt => (P, Pc cO) + end + | Pinj _ _, vmon _ _ => (P, Pc cO) + | PX P1 i Q1, zmon j M1 => + let M2 := zmon_pred j M1 in + let (R1, S1) := MFactor P1 M in + let (R2, S2) := MFactor Q1 M2 in + (mkPX R1 i R2, mkPX S1 i S2) + | PX P1 i Q1, vmon j M1 => + match (i ?= j) Eq with + Eq => let (R1,S1) := MFactor P1 (mkZmon xH M1) in + (mkPX R1 i Q1, S1) + | Lt => let (R1,S1) := MFactor P1 (vmon (j - i) M1) in + (mkPX R1 i Q1, S1) + | Gt => let (R1,S1) := MFactor P1 (mkZmon xH M1) in + (mkPX R1 i Q1, mkPX S1 (i-j) (Pc cO)) + end + end. + + Definition POneSubst (P1: Pol) (M1: Mon) (P2: Pol): option Pol := + let (Q1,R1) := MFactor P1 M1 in + match R1 with + (Pc c) => if c ?=! cO then None + else Some (Padd Q1 (Pmul P2 R1)) + | _ => Some (Padd Q1 (Pmul P2 R1)) + end. + + Fixpoint PNSubst1 (P1: Pol) (M1: Mon) (P2: Pol) (n: nat) {struct n}: Pol := + match POneSubst P1 M1 P2 with + Some P3 => match n with S n1 => PNSubst1 P3 M1 P2 n1 | _ => P3 end + | _ => P1 + end. + + Definition PNSubst (P1: Pol) (M1: Mon) (P2: Pol) (n: nat): option Pol := + match POneSubst P1 M1 P2 with + Some P3 => match n with S n1 => Some (PNSubst1 P3 M1 P2 n1) | _ => None end + | _ => None + end. + + Fixpoint PSubstL1 (P1: Pol) (LM1: list (Mon * Pol)) (n: nat) {struct LM1}: + Pol := + match LM1 with + cons (M1,P2) LM2 => PSubstL1 (PNSubst1 P1 M1 P2 n) LM2 n + | _ => P1 + end. + + Fixpoint PSubstL (P1: Pol) (LM1: list (Mon * Pol)) (n: nat) {struct LM1}: option Pol := + match LM1 with + cons (M1,P2) LM2 => + match PNSubst P1 M1 P2 n with + Some P3 => Some (PSubstL1 P3 LM2 n) + | None => PSubstL P1 LM2 n + end + | _ => None + end. + + Fixpoint PNSubstL (P1: Pol) (LM1: list (Mon * Pol)) (m n: nat) {struct m}: Pol := + match PSubstL P1 LM1 n with + Some P3 => match m with S m1 => PNSubstL P3 LM1 m1 n | _ => P3 end + | _ => P1 + end. + + (** Evaluation of a polynomial towards R *) + + Fixpoint Pphi(l:Env R) (P:Pol) {struct P} : R := + match P with + | Pc c => [c] + | Pinj j Q => Pphi (jump j l) Q + | PX P i Q => + let x := hd 0 l in + let xi := pow_pos rmul x i in + (Pphi l P) * xi + (Pphi (tail l) Q) + end. + + Reserved Notation "P @ l " (at level 10, no associativity). + Notation "P @ l " := (Pphi l P). + (** Proofs *) + Lemma ZPminus_spec : forall x y, + match ZPminus x y with + | Z0 => x = y + | Zpos k => x = (y + k)%positive + | Zneg k => y = (x + k)%positive + end. + Proof. + induction x;destruct y. + replace (ZPminus (xI x) (xI y)) with (Zdouble (ZPminus x y));trivial. + assert (H := IHx y);destruct (ZPminus x y);unfold Zdouble;rewrite H;trivial. + replace (ZPminus (xI x) (xO y)) with (Zdouble_plus_one (ZPminus x y));trivial. + assert (H := IHx y);destruct (ZPminus x y);unfold Zdouble_plus_one;rewrite H;trivial. + apply Pplus_xI_double_minus_one. + simpl;trivial. + replace (ZPminus (xO x) (xI y)) with (Zdouble_minus_one (ZPminus x y));trivial. + assert (H := IHx y);destruct (ZPminus x y);unfold Zdouble_minus_one;rewrite H;trivial. + apply Pplus_xI_double_minus_one. + replace (ZPminus (xO x) (xO y)) with (Zdouble (ZPminus x y));trivial. + assert (H := IHx y);destruct (ZPminus x y);unfold Zdouble;rewrite H;trivial. + replace (ZPminus (xO x) xH) with (Zpos (Pdouble_minus_one x));trivial. + rewrite <- Pplus_one_succ_l. + rewrite Psucc_o_double_minus_one_eq_xO;trivial. + replace (ZPminus xH (xI y)) with (Zneg (xO y));trivial. + replace (ZPminus xH (xO y)) with (Zneg (Pdouble_minus_one y));trivial. + rewrite <- Pplus_one_succ_l. + rewrite Psucc_o_double_minus_one_eq_xO;trivial. + simpl;trivial. + Qed. + + Lemma Peq_ok : forall P P', + (P ?== P') = true -> forall l, P@l == P'@ l. + Proof. + induction P;destruct P';simpl;intros;try discriminate;trivial. + apply (morph_eq CRmorph);trivial. + assert (H1 := Pcompare_Eq_eq p p0); destruct ((p ?= p0)%positive Eq); + try discriminate H. + rewrite (IHP P' H); rewrite H1;trivial;rrefl. + assert (H1 := Pcompare_Eq_eq p p0); destruct ((p ?= p0)%positive Eq); + try discriminate H. + rewrite H1;trivial. clear H1. + assert (H1 := IHP1 P'1);assert (H2 := IHP2 P'2); + destruct (P2 ?== P'1);[destruct (P3 ?== P'2); [idtac|discriminate H] + |discriminate H]. + rewrite (H1 H);rewrite (H2 H);rrefl. + Qed. + + Lemma Pphi0 : forall l, P0@l == 0. + Proof. + intros;simpl;apply (morph0 CRmorph). + Qed. + +Lemma env_morph : forall p e1 e2, (forall x, e1 x = e2 x) -> + p @ e1 = p @ e2. +Proof. + induction p ; simpl. + reflexivity. + intros. + apply IHp. + intros. + unfold jump. + apply H. + intros. + rewrite (IHp1 e1 e2) ; auto. + rewrite (IHp2 (tail e1) (tail e2)) ; auto. + unfold hd. unfold nth. rewrite H. reflexivity. + unfold tail. unfold jump. intros ; apply H. +Qed. + +Lemma Pjump_Pplus : forall P i j l, P @ (jump (i + j) l ) = P @ (jump j (jump i l)). +Proof. + intros. apply env_morph. intros. rewrite <- jump_Pplus. + rewrite Pplus_comm. + reflexivity. +Qed. + +Lemma Pjump_xO_tail : forall P p l, + P @ (jump (xO p) (tail l)) = P @ (jump (xI p) l). +Proof. + intros. + apply env_morph. + intros. + rewrite (@jump_simpl R (xI p)). + rewrite (@jump_simpl R (xO p)). + reflexivity. +Qed. + +Lemma Pjump_Pdouble_minus_one : forall P p l, + P @ (jump (Pdouble_minus_one p) (tail l)) = P @ (jump (xO p) l). +Proof. + intros. + apply env_morph. + intros. + rewrite jump_Pdouble_minus_one. + rewrite (@jump_simpl R (xO p)). + reflexivity. +Qed. + + + + Lemma Pphi1 : forall l, P1@l == 1. + Proof. + intros;simpl;apply (morph1 CRmorph). + Qed. + + Lemma mkPinj_ok : forall j l P, (mkPinj j P)@l == P@(jump j l). + Proof. + intros j l p;destruct p;simpl;rsimpl. + rewrite Pjump_Pplus. + reflexivity. + Qed. + + Let pow_pos_Pplus := + pow_pos_Pplus rmul Rsth Reqe.(Rmul_ext) ARth.(ARmul_comm) ARth.(ARmul_assoc). + + Lemma mkPX_ok : forall l P i Q, + (mkPX P i Q)@l == P@l*(pow_pos rmul (hd 0 l) i) + Q@(tail l). + Proof. + intros l P i Q;unfold mkPX. + destruct P;try (simpl;rrefl). + assert (H := morph_eq CRmorph c cO);destruct (c ?=! cO);simpl;try rrefl. + rewrite (H (refl_equal true));rewrite (morph0 CRmorph). + rewrite mkPinj_ok;rsimpl;simpl;rrefl. + assert (H := @Peq_ok P3 P0);destruct (P3 ?== P0);simpl;try rrefl. + rewrite (H (refl_equal true));trivial. + rewrite Pphi0. rewrite pow_pos_Pplus;rsimpl. + Qed. + + + Ltac Esimpl := + repeat (progress ( + match goal with + | |- context [P0@?l] => rewrite (Pphi0 l) + | |- context [P1@?l] => rewrite (Pphi1 l) + | |- context [(mkPinj ?j ?P)@?l] => rewrite (mkPinj_ok j l P) + | |- context [(mkPX ?P ?i ?Q)@?l] => rewrite (mkPX_ok l P i Q) + | |- context [[cO]] => rewrite (morph0 CRmorph) + | |- context [[cI]] => rewrite (morph1 CRmorph) + | |- context [[?x +! ?y]] => rewrite ((morph_add CRmorph) x y) + | |- context [[?x *! ?y]] => rewrite ((morph_mul CRmorph) x y) + | |- context [[?x -! ?y]] => rewrite ((morph_sub CRmorph) x y) + | |- context [[-! ?x]] => rewrite ((morph_opp CRmorph) x) + end)); + rsimpl; simpl. + + Lemma PaddC_ok : forall c P l, (PaddC P c)@l == P@l + [c]. + Proof. + induction P;simpl;intros;Esimpl;trivial. + rewrite IHP2;rsimpl. + Qed. + + Lemma PsubC_ok : forall c P l, (PsubC P c)@l == P@l - [c]. + Proof. + induction P;simpl;intros. + Esimpl. + rewrite IHP;rsimpl. + rewrite IHP2;rsimpl. + Qed. + + Lemma PmulC_aux_ok : forall c P l, (PmulC_aux P c)@l == P@l * [c]. + Proof. + induction P;simpl;intros;Esimpl;trivial. + rewrite IHP1;rewrite IHP2;rsimpl. + mul_push ([c]);rrefl. + Qed. + + Lemma PmulC_ok : forall c P l, (PmulC P c)@l == P@l * [c]. + Proof. + intros c P l; unfold PmulC. + assert (H:= morph_eq CRmorph c cO);destruct (c ?=! cO). + rewrite (H (refl_equal true));Esimpl. + assert (H1:= morph_eq CRmorph c cI);destruct (c ?=! cI). + rewrite (H1 (refl_equal true));Esimpl. + apply PmulC_aux_ok. + Qed. + + Lemma Popp_ok : forall P l, (--P)@l == - P@l. + Proof. + induction P;simpl;intros. + Esimpl. + apply IHP. + rewrite IHP1;rewrite IHP2;rsimpl. + Qed. + + Ltac Esimpl2 := + Esimpl; + repeat (progress ( + match goal with + | |- context [(PaddC ?P ?c)@?l] => rewrite (PaddC_ok c P l) + | |- context [(PsubC ?P ?c)@?l] => rewrite (PsubC_ok c P l) + | |- context [(PmulC ?P ?c)@?l] => rewrite (PmulC_ok c P l) + | |- context [(--?P)@?l] => rewrite (Popp_ok P l) + end)); Esimpl. + + + + + Lemma Padd_ok : forall P' P l, (P ++ P')@l == P@l + P'@l. + Proof. + induction P';simpl;intros;Esimpl2. + generalize P p l;clear P p l. + induction P;simpl;intros. + Esimpl2;apply (ARadd_comm ARth). + assert (H := ZPminus_spec p p0);destruct (ZPminus p p0). + rewrite H;Esimpl. rewrite IHP';rrefl. + rewrite H;Esimpl. rewrite IHP';Esimpl. + rewrite Pjump_Pplus. rrefl. + rewrite H;Esimpl. rewrite IHP. + rewrite Pjump_Pplus. rrefl. + destruct p0;simpl. + rewrite IHP2;simpl. rsimpl. + rewrite Pjump_xO_tail. Esimpl. + rewrite IHP2;simpl. + rewrite Pjump_Pdouble_minus_one. + rsimpl. + rewrite IHP'. + rsimpl. + destruct P;simpl. + Esimpl2;add_push [c];rrefl. + destruct p0;simpl;Esimpl2. + rewrite IHP'2;simpl. + rewrite Pjump_xO_tail. + rsimpl;add_push (P'1@l * (pow_pos rmul (hd 0 l) p));rrefl. + rewrite IHP'2;simpl. + rewrite Pjump_Pdouble_minus_one. rsimpl. + add_push (P'1@l * (pow_pos rmul (hd 0 l) p));rrefl. + rewrite IHP'2;rsimpl. + unfold tail. + add_push (P @ (jump 1 l));rrefl. + assert (H := ZPminus_spec p0 p);destruct (ZPminus p0 p);Esimpl2. + rewrite IHP'1;rewrite IHP'2;rsimpl. + add_push (P3 @ (tail l));rewrite H;rrefl. + rewrite IHP'1;rewrite IHP'2;simpl;Esimpl. + rewrite H;rewrite Pplus_comm. + rewrite pow_pos_Pplus;rsimpl. + add_push (P3 @ (tail l));rrefl. + assert (forall P k l, + (PaddX Padd P'1 k P) @ l == P@l + P'1@l * pow_pos rmul (hd 0 l) k). + induction P;simpl;intros;try apply (ARadd_comm ARth). + destruct p2; simpl; try apply (ARadd_comm ARth). + rewrite Pjump_xO_tail. + apply (ARadd_comm ARth). + rewrite Pjump_Pdouble_minus_one. + apply (ARadd_comm ARth). + assert (H1 := ZPminus_spec p2 k);destruct (ZPminus p2 k);Esimpl2. + rewrite IHP'1;rsimpl; rewrite H1;add_push (P5 @ (tail l0));rrefl. + rewrite IHP'1;simpl;Esimpl. + rewrite H1;rewrite Pplus_comm. + rewrite pow_pos_Pplus;simpl;Esimpl. + add_push (P5 @ (tail l0));rrefl. + rewrite IHP1;rewrite H1;rewrite Pplus_comm. + rewrite pow_pos_Pplus;simpl;rsimpl. + add_push (P5 @ (tail l0));rrefl. + rewrite H0;rsimpl. + add_push (P3 @ (tail l)). + rewrite H;rewrite Pplus_comm. + rewrite IHP'2;rewrite pow_pos_Pplus;rsimpl. + add_push (P3 @ (tail l));rrefl. + Qed. + + Lemma Psub_ok : forall P' P l, (P -- P')@l == P@l - P'@l. + Proof. + induction P';simpl;intros;Esimpl2;trivial. + generalize P p l;clear P p l. + induction P;simpl;intros. + Esimpl2;apply (ARadd_comm ARth). + assert (H := ZPminus_spec p p0);destruct (ZPminus p p0). + rewrite H;Esimpl. rewrite IHP';rsimpl. + rewrite H;Esimpl. rewrite IHP';Esimpl. + rewrite <- Pjump_Pplus;rewrite Pplus_comm;rrefl. + rewrite H;Esimpl. rewrite IHP. + rewrite <- Pjump_Pplus;rewrite Pplus_comm;rrefl. + destruct p0;simpl. + rewrite IHP2;simpl; try rewrite Pjump_xO_tail ; rsimpl. + rewrite IHP2;simpl. + rewrite Pjump_Pdouble_minus_one;rsimpl. + unfold tail ; rsimpl. + rewrite IHP';rsimpl. + destruct P;simpl. + repeat rewrite Popp_ok;Esimpl2;rsimpl;add_push [c];try rrefl. + destruct p0;simpl;Esimpl2. + rewrite IHP'2;simpl;rsimpl;add_push (P'1@l * (pow_pos rmul (hd 0 l) p));trivial. + rewrite Pjump_xO_tail. + add_push (P @ ((jump (xI p0) l)));rrefl. + rewrite IHP'2;simpl;rewrite Pjump_Pdouble_minus_one;rsimpl. + add_push (- (P'1 @ l * pow_pos rmul (hd 0 l) p));rrefl. + unfold tail. + rewrite IHP'2;rsimpl;add_push (P @ (jump 1 l));rrefl. + assert (H := ZPminus_spec p0 p);destruct (ZPminus p0 p);Esimpl2. + rewrite IHP'1; rewrite IHP'2;rsimpl. + add_push (P3 @ (tail l));rewrite H;rrefl. + rewrite IHP'1; rewrite IHP'2;rsimpl;simpl;Esimpl. + rewrite H;rewrite Pplus_comm. + rewrite pow_pos_Pplus;rsimpl. + add_push (P3 @ (tail l));rrefl. + assert (forall P k l, + (PsubX Psub P'1 k P) @ l == P@l + - P'1@l * pow_pos rmul (hd 0 l) k). + induction P;simpl;intros. + rewrite Popp_ok;rsimpl;apply (ARadd_comm ARth);trivial. + destruct p2;simpl; rewrite Popp_ok;rsimpl. + rewrite Pjump_xO_tail. + apply (ARadd_comm ARth);trivial. + rewrite Pjump_Pdouble_minus_one. + apply (ARadd_comm ARth);trivial. + apply (ARadd_comm ARth);trivial. + assert (H1 := ZPminus_spec p2 k);destruct (ZPminus p2 k);Esimpl2;rsimpl. + rewrite IHP'1;rsimpl;add_push (P5 @ (tail l0));rewrite H1;rrefl. + rewrite IHP'1;rewrite H1;rewrite Pplus_comm. + rewrite pow_pos_Pplus;simpl;Esimpl. + add_push (P5 @ (tail l0));rrefl. + rewrite IHP1;rewrite H1;rewrite Pplus_comm. + rewrite pow_pos_Pplus;simpl;rsimpl. + add_push (P5 @ (tail l0));rrefl. + rewrite H0;rsimpl. + rewrite IHP'2;rsimpl;add_push (P3 @ (tail l)). + rewrite H;rewrite Pplus_comm. + rewrite pow_pos_Pplus;rsimpl. + Qed. +(* Proof for the symmetric version *) + + Lemma PmulI_ok : + forall P', + (forall (P : Pol) (l : Env R), (Pmul P P') @ l == P @ l * P' @ l) -> + forall (P : Pol) (p : positive) (l : Env R), + (PmulI Pmul P' p P) @ l == P @ l * P' @ (jump p l). + Proof. + induction P;simpl;intros. + Esimpl2;apply (ARmul_comm ARth). + assert (H1 := ZPminus_spec p p0);destruct (ZPminus p p0);Esimpl2. + rewrite H1; rewrite H;rrefl. + rewrite H1; rewrite H. + rewrite Pjump_Pplus;simpl;rrefl. + rewrite H1. + rewrite Pjump_Pplus;rewrite IHP;rrefl. + destruct p0;Esimpl2. + rewrite IHP1;rewrite IHP2;rsimpl. + rewrite Pjump_xO_tail. + mul_push (pow_pos rmul (hd 0 l) p);rrefl. + rewrite IHP1;rewrite IHP2;simpl;rsimpl. + mul_push (pow_pos rmul (hd 0 l) p); rewrite Pjump_Pdouble_minus_one. + rrefl. + rewrite IHP1;simpl;rsimpl. + mul_push (pow_pos rmul (hd 0 l) p). + rewrite H;rrefl. + Qed. + +(* + Lemma PmulI_ok : + forall P', + (forall (P : Pol) (l : list R), (Pmul_aux P P') @ l == P @ l * P' @ l) -> + forall (P : Pol) (p : positive) (l : list R), + (PmulI Pmul_aux P' p P) @ l == P @ l * P' @ (jump p l). + Proof. + induction P;simpl;intros. + Esimpl2;apply (ARmul_comm ARth). + assert (H1 := ZPminus_spec p p0);destruct (ZPminus p p0);Esimpl2. + rewrite H1; rewrite H;rrefl. + rewrite H1; rewrite H. + rewrite Pplus_comm. + rewrite jump_Pplus;simpl;rrefl. + rewrite H1;rewrite Pplus_comm. + rewrite jump_Pplus;rewrite IHP;rrefl. + destruct p0;Esimpl2. + rewrite IHP1;rewrite IHP2;simpl;rsimpl. + mul_push (pow_pos rmul (hd 0 l) p);rrefl. + rewrite IHP1;rewrite IHP2;simpl;rsimpl. + mul_push (pow_pos rmul (hd 0 l) p); rewrite jump_Pdouble_minus_one;rrefl. + rewrite IHP1;simpl;rsimpl. + mul_push (pow_pos rmul (hd 0 l) p). + rewrite H;rrefl. + Qed. + + Lemma Pmul_aux_ok : forall P' P l,(Pmul_aux P P')@l == P@l * P'@l. + Proof. + induction P';simpl;intros. + Esimpl2;trivial. + apply PmulI_ok;trivial. + rewrite Padd_ok;Esimpl2. + rewrite (PmulI_ok P'2 IHP'2). rewrite IHP'1. rrefl. + Qed. +*) + +(* Proof for the symmetric version *) + Lemma Pmul_ok : forall P P' l, (P**P')@l == P@l * P'@l. + Proof. + intros P P';generalize P;clear P;induction P';simpl;intros. + apply PmulC_ok. apply PmulI_ok;trivial. + destruct P. + rewrite (ARmul_comm ARth);Esimpl2;Esimpl2. + Esimpl2. rewrite IHP'1;Esimpl2. + assert (match p0 with + | xI j => Pinj (xO j) P ** P'2 + | xO j => Pinj (Pdouble_minus_one j) P ** P'2 + | 1 => P ** P'2 + end @ (tail l) == P @ (jump p0 l) * P'2 @ (tail l)). + destruct p0;rewrite IHP'2;Esimpl. + rewrite Pjump_xO_tail. reflexivity. + rewrite Pjump_Pdouble_minus_one;Esimpl. + rewrite H;Esimpl. + rewrite Padd_ok; Esimpl2. rewrite Padd_ok; Esimpl2. + repeat (rewrite IHP'1 || rewrite IHP'2);simpl. + rewrite PmulI_ok;trivial. + unfold tail. + mul_push (P'1@l). simpl. mul_push (P'2 @ (jump 1 l)). Esimpl. + Qed. + +(* +Lemma Pmul_ok : forall P P' l, (P**P')@l == P@l * P'@l. + Proof. + destruct P;simpl;intros. + Esimpl2;apply (ARmul_comm ARth). + rewrite (PmulI_ok P (Pmul_aux_ok P)). + apply (ARmul_comm ARth). + rewrite Padd_ok; Esimpl2. + rewrite (PmulI_ok P3 (Pmul_aux_ok P3));trivial. + rewrite Pmul_aux_ok;mul_push (P' @ l). + rewrite (ARmul_comm ARth (P' @ l));rrefl. + Qed. +*) + + Lemma Psquare_ok : forall P l, (Psquare P)@l == P@l * P@l. + Proof. + induction P;simpl;intros;Esimpl2. + apply IHP. rewrite Padd_ok. rewrite Pmul_ok;Esimpl2. + rewrite IHP1;rewrite IHP2. + mul_push (pow_pos rmul (hd 0 l) p). mul_push (P2@l). + rrefl. + Qed. + + Lemma Mphi_morph : forall P env env', (forall x, env x = env' x ) -> + Mphi env P = Mphi env' P. + Proof. + induction P ; simpl. + reflexivity. + intros. + apply IHP. + intros. + unfold jump. + apply H. + (**) + intros. + replace (Mphi (tail env) P) with (Mphi (tail env') P). + unfold hd. unfold nth. + rewrite H. + reflexivity. + apply IHP. + unfold tail,jump. + intros. symmetry. apply H. + Qed. + +Lemma Mjump_xO_tail : forall M p l, + Mphi (jump (xO p) (tail l)) M = Mphi (jump (xI p) l) M. +Proof. + intros. + apply Mphi_morph. + intros. + rewrite (@jump_simpl R (xI p)). + rewrite (@jump_simpl R (xO p)). + reflexivity. +Qed. + +Lemma Mjump_Pdouble_minus_one : forall M p l, + Mphi (jump (Pdouble_minus_one p) (tail l)) M = Mphi (jump (xO p) l) M. +Proof. + intros. + apply Mphi_morph. + intros. + rewrite jump_Pdouble_minus_one. + rewrite (@jump_simpl R (xO p)). + reflexivity. +Qed. + +Lemma Mjump_Pplus : forall M i j l, Mphi (jump (i + j) l ) M = Mphi (jump j (jump i l)) M. +Proof. + intros. apply Mphi_morph. intros. rewrite <- jump_Pplus. + rewrite Pplus_comm. + reflexivity. +Qed. + + + + Lemma mkZmon_ok: forall M j l, + Mphi l (mkZmon j M) == Mphi l (zmon j M). + intros M j l; case M; simpl; intros; rsimpl. + Qed. + + Lemma zmon_pred_ok : forall M j l, + Mphi (tail l) (zmon_pred j M) == Mphi l (zmon j M). + Proof. + destruct j; simpl;intros l; rsimpl. + rewrite mkZmon_ok;rsimpl. + simpl. + rewrite Mjump_xO_tail. + reflexivity. + rewrite mkZmon_ok;simpl. + rewrite Mjump_Pdouble_minus_one; rsimpl. + Qed. + + Lemma mkVmon_ok : forall M i l, Mphi l (mkVmon i M) == Mphi l M*pow_pos rmul (hd 0 l) i. + Proof. + destruct M;simpl;intros;rsimpl. + rewrite zmon_pred_ok;simpl;rsimpl. + rewrite Pplus_comm;rewrite pow_pos_Pplus;rsimpl. + Qed. + + + Lemma Mphi_ok: forall P M l, + let (Q,R) := MFactor P M in + P@l == Q@l + (Mphi l M) * (R@l). + Proof. + intros P; elim P; simpl; auto; clear P. + intros c M l; case M; simpl; auto; try intro p; try intro m; + try rewrite (morph0 CRmorph); rsimpl. + + intros i P Hrec M l; case M; simpl; clear M. + rewrite (morph0 CRmorph); rsimpl. + intros j M. + case_eq ((i ?= j) Eq); intros He; simpl. + rewrite (Pcompare_Eq_eq _ _ He). + generalize (Hrec M (jump j l)); case (MFactor P M); + simpl; intros P2 Q2 H; repeat rewrite mkPinj_ok; auto. + generalize (Hrec (zmon (j -i) M) (jump i l)); + case (MFactor P (zmon (j -i) M)); simpl. + intros P2 Q2 H; repeat rewrite mkPinj_ok; auto. + rewrite <- (Pplus_minus _ _ (ZC2 _ _ He)). + rewrite Mjump_Pplus; auto. + rewrite (morph0 CRmorph); rsimpl. + intros P2 m; rewrite (morph0 CRmorph); rsimpl. + + intros P2 Hrec1 i Q2 Hrec2 M l; case M; simpl; auto. + rewrite (morph0 CRmorph); rsimpl. + intros j M1. + generalize (Hrec1 (zmon j M1) l); + case (MFactor P2 (zmon j M1)). + intros R1 S1 H1. + generalize (Hrec2 (zmon_pred j M1) (tail l)); + case (MFactor Q2 (zmon_pred j M1)); simpl. + intros R2 S2 H2; rewrite H1; rewrite H2. + repeat rewrite mkPX_ok; simpl. + rsimpl. + apply radd_ext; rsimpl. + rewrite (ARadd_comm ARth); rsimpl. + apply radd_ext; rsimpl. + rewrite (ARadd_comm ARth); rsimpl. + rewrite zmon_pred_ok;rsimpl. + intros j M1. + case_eq ((i ?= j) Eq); intros He; simpl. + rewrite (Pcompare_Eq_eq _ _ He). + generalize (Hrec1 (mkZmon xH M1) l); case (MFactor P2 (mkZmon xH M1)); + simpl; intros P3 Q3 H; repeat rewrite mkPinj_ok; auto. + rewrite H; rewrite mkPX_ok; rsimpl. + repeat (rewrite <-(ARadd_assoc ARth)). + apply radd_ext; rsimpl. + rewrite (ARadd_comm ARth); rsimpl. + apply radd_ext; rsimpl. + repeat (rewrite <-(ARmul_assoc ARth)). + rewrite mkZmon_ok. + apply rmul_ext; rsimpl. + rewrite (ARmul_comm ARth); rsimpl. + generalize (Hrec1 (vmon (j - i) M1) l); + case (MFactor P2 (vmon (j - i) M1)); + simpl; intros P3 Q3 H; repeat rewrite mkPinj_ok; auto. + rewrite H; rsimpl; repeat rewrite mkPinj_ok; auto. + rewrite mkPX_ok; rsimpl. + repeat (rewrite <-(ARadd_assoc ARth)). + apply radd_ext; rsimpl. + rewrite (ARadd_comm ARth); rsimpl. + apply radd_ext; rsimpl. + repeat (rewrite <-(ARmul_assoc ARth)). + apply rmul_ext; rsimpl. + rewrite (ARmul_comm ARth); rsimpl. + apply rmul_ext; rsimpl. + rewrite <- pow_pos_Pplus. + rewrite (Pplus_minus _ _ (ZC2 _ _ He)); rsimpl. + generalize (Hrec1 (mkZmon 1 M1) l); + case (MFactor P2 (mkZmon 1 M1)); + simpl; intros P3 Q3 H; repeat rewrite mkPinj_ok; auto. + rewrite H; rsimpl. + rewrite mkPX_ok; rsimpl. + repeat (rewrite <-(ARadd_assoc ARth)). + apply radd_ext; rsimpl. + rewrite (ARadd_comm ARth); rsimpl. + apply radd_ext; rsimpl. + rewrite mkZmon_ok. + repeat (rewrite <-(ARmul_assoc ARth)). + apply rmul_ext; rsimpl. + rewrite (ARmul_comm ARth); rsimpl. + rewrite mkPX_ok; simpl; rsimpl. + rewrite (morph0 CRmorph); rsimpl. + repeat (rewrite <-(ARmul_assoc ARth)). + rewrite (ARmul_comm ARth (Q3@l)); rsimpl. + apply rmul_ext; rsimpl. + rewrite <- pow_pos_Pplus. + rewrite (Pplus_minus _ _ He); rsimpl. + Qed. + +(* Proof for the symmetric version *) + + Lemma POneSubst_ok: forall P1 M1 P2 P3 l, + POneSubst P1 M1 P2 = Some P3 -> Mphi l M1 == P2@l -> P1@l == P3@l. + Proof. + intros P2 M1 P3 P4 l; unfold POneSubst. + generalize (Mphi_ok P2 M1 l); case (MFactor P2 M1); simpl; auto. + intros Q1 R1; case R1. + intros c H; rewrite H. + generalize (morph_eq CRmorph c cO); + case (c ?=! cO); simpl; auto. + intros H1 H2; rewrite H1; auto; rsimpl. + discriminate. + intros _ H1 H2; injection H1; intros; subst. + rewrite H2; rsimpl. + (* new version *) + rewrite Padd_ok; rewrite PmulC_ok; rsimpl. + intros i P5 H; rewrite H. + intros HH H1; injection HH; intros; subst; rsimpl. + rewrite Padd_ok; rewrite PmulI_ok by (intros;apply Pmul_ok). rewrite H1; rsimpl. + intros i P5 P6 H1 H2 H3; rewrite H1; rewrite H3. + assert (P4 = Q1 ++ P3 ** PX i P5 P6). + injection H2; intros; subst;trivial. + rewrite H;rewrite Padd_ok;rewrite Pmul_ok;rsimpl. +Qed. +(* + Lemma POneSubst_ok: forall P1 M1 P2 P3 l, + POneSubst P1 M1 P2 = Some P3 -> Mphi l M1 == P2@l -> P1@l == P3@l. +Proof. + intros P2 M1 P3 P4 l; unfold POneSubst. + generalize (Mphi_ok P2 M1 l); case (MFactor P2 M1); simpl; auto. + intros Q1 R1; case R1. + intros c H; rewrite H. + generalize (morph_eq CRmorph c cO); + case (c ?=! cO); simpl; auto. + intros H1 H2; rewrite H1; auto; rsimpl. + discriminate. + intros _ H1 H2; injection H1; intros; subst. + rewrite H2; rsimpl. + rewrite Padd_ok; rewrite Pmul_ok; rsimpl. + intros i P5 H; rewrite H. + intros HH H1; injection HH; intros; subst; rsimpl. + rewrite Padd_ok; rewrite Pmul_ok. rewrite H1; rsimpl. + intros i P5 P6 H1 H2 H3; rewrite H1; rewrite H3. + injection H2; intros; subst; rsimpl. + rewrite Padd_ok. + rewrite Pmul_ok; rsimpl. + Qed. +*) + Lemma PNSubst1_ok: forall n P1 M1 P2 l, + Mphi l M1 == P2@l -> P1@l == (PNSubst1 P1 M1 P2 n)@l. + Proof. + intros n; elim n; simpl; auto. + intros P2 M1 P3 l H. + generalize (fun P4 => @POneSubst_ok P2 M1 P3 P4 l); + case (POneSubst P2 M1 P3); [idtac | intros; rsimpl]. + intros P4 Hrec; rewrite (Hrec P4); auto; rsimpl. + intros n1 Hrec P2 M1 P3 l H. + generalize (fun P4 => @POneSubst_ok P2 M1 P3 P4 l); + case (POneSubst P2 M1 P3); [idtac | intros; rsimpl]. + intros P4 Hrec1; rewrite (Hrec1 P4); auto; rsimpl. + Qed. + + Lemma PNSubst_ok: forall n P1 M1 P2 l P3, + PNSubst P1 M1 P2 n = Some P3 -> Mphi l M1 == P2@l -> P1@l == P3@l. + Proof. + intros n P2 M1 P3 l P4; unfold PNSubst. + generalize (fun P4 => @POneSubst_ok P2 M1 P3 P4 l); + case (POneSubst P2 M1 P3); [idtac | intros; discriminate]. + intros P5 H1; case n; try (intros; discriminate). + intros n1 H2; injection H2; intros; subst. + rewrite <- PNSubst1_ok; auto. + Qed. + + Fixpoint MPcond (LM1: list (Mon * Pol)) (l: Env R) {struct LM1} : Prop := + match LM1 with + cons (M1,P2) LM2 => (Mphi l M1 == P2@l) /\ (MPcond LM2 l) + | _ => True + end. + + Lemma PSubstL1_ok: forall n LM1 P1 l, + MPcond LM1 l -> P1@l == (PSubstL1 P1 LM1 n)@l. + Proof. + intros n LM1; elim LM1; simpl; auto. + intros; rsimpl. + intros (M2,P2) LM2 Hrec P3 l [H H1]. + rewrite <- Hrec; auto. + apply PNSubst1_ok; auto. + Qed. + + Lemma PSubstL_ok: forall n LM1 P1 P2 l, + PSubstL P1 LM1 n = Some P2 -> MPcond LM1 l -> P1@l == P2@l. + Proof. + intros n LM1; elim LM1; simpl; auto. + intros; discriminate. + intros (M2,P2) LM2 Hrec P3 P4 l. + generalize (PNSubst_ok n P3 M2 P2); case (PNSubst P3 M2 P2 n). + intros P5 H0 H1 [H2 H3]; injection H1; intros; subst. + rewrite <- PSubstL1_ok; auto. + intros l1 H [H1 H2]; auto. + Qed. + + Lemma PNSubstL_ok: forall m n LM1 P1 l, + MPcond LM1 l -> P1@l == (PNSubstL P1 LM1 m n)@l. + Proof. + intros m; elim m; simpl; auto. + intros n LM1 P2 l H; generalize (fun P3 => @PSubstL_ok n LM1 P2 P3 l); + case (PSubstL P2 LM1 n); intros; rsimpl; auto. + intros m1 Hrec n LM1 P2 l H. + generalize (fun P3 => @PSubstL_ok n LM1 P2 P3 l); + case (PSubstL P2 LM1 n); intros; rsimpl; auto. + rewrite <- Hrec; auto. + Qed. + + (** Definition of polynomial expressions *) + + Inductive PExpr : Type := + | PEc : C -> PExpr + | PEX : positive -> PExpr + | PEadd : PExpr -> PExpr -> PExpr + | PEsub : PExpr -> PExpr -> PExpr + | PEmul : PExpr -> PExpr -> PExpr + | PEopp : PExpr -> PExpr + | PEpow : PExpr -> N -> PExpr. + + (** evaluation of polynomial expressions towards R *) + Definition mk_X j := mkPinj_pred j mkX. + + (** evaluation of polynomial expressions towards R *) + + Fixpoint PEeval (l:Env R) (pe:PExpr) {struct pe} : R := + match pe with + | PEc c => phi c + | PEX j => nth j l + | PEadd pe1 pe2 => (PEeval l pe1) + (PEeval l pe2) + | PEsub pe1 pe2 => (PEeval l pe1) - (PEeval l pe2) + | PEmul pe1 pe2 => (PEeval l pe1) * (PEeval l pe2) + | PEopp pe1 => - (PEeval l pe1) + | PEpow pe1 n => rpow (PEeval l pe1) (Cp_phi n) + end. + + (** Correctness proofs *) + + Lemma mkX_ok : forall p l, nth p l == (mk_X p) @ l. + Proof. + destruct p;simpl;intros;Esimpl;trivial. + rewrite nth_spec ; auto. + unfold hd. + rewrite <- nth_Pdouble_minus_one. + rewrite (nth_jump (Pdouble_minus_one p) l 1). + reflexivity. + Qed. + + Ltac Esimpl3 := + repeat match goal with + | |- context [(?P1 ++ ?P2)@?l] => rewrite (Padd_ok P2 P1 l) + | |- context [(?P1 -- ?P2)@?l] => rewrite (Psub_ok P2 P1 l) + end;Esimpl2;try rrefl;try apply (ARadd_comm ARth). + +(* Power using the chinise algorithm *) +(*Section POWER. + Variable subst_l : Pol -> Pol. + Fixpoint Ppow_pos (P:Pol) (p:positive){struct p} : Pol := + match p with + | xH => P + | xO p => subst_l (Psquare (Ppow_pos P p)) + | xI p => subst_l (Pmul P (Psquare (Ppow_pos P p))) + end. + + Definition Ppow_N P n := + match n with + | N0 => P1 + | Npos p => Ppow_pos P p + end. + + Lemma Ppow_pos_ok : forall l, (forall P, subst_l P@l == P@l) -> + forall P p, (Ppow_pos P p)@l == (pow_pos Pmul P p)@l. + Proof. + intros l subst_l_ok P. + induction p;simpl;intros;try rrefl;try rewrite subst_l_ok. + repeat rewrite Pmul_ok;rewrite Psquare_ok;rewrite IHp;rrefl. + repeat rewrite Pmul_ok;rewrite Psquare_ok;rewrite IHp;rrefl. + Qed. + + Lemma Ppow_N_ok : forall l, (forall P, subst_l P@l == P@l) -> + forall P n, (Ppow_N P n)@l == (pow_N P1 Pmul P n)@l. + Proof. destruct n;simpl. rrefl. apply Ppow_pos_ok. trivial. Qed. + + End POWER. *) + +Section POWER. + Variable subst_l : Pol -> Pol. + Fixpoint Ppow_pos (res P:Pol) (p:positive){struct p} : Pol := + match p with + | xH => subst_l (Pmul res P) + | xO p => Ppow_pos (Ppow_pos res P p) P p + | xI p => subst_l (Pmul (Ppow_pos (Ppow_pos res P p) P p) P) + end. + + Definition Ppow_N P n := + match n with + | N0 => P1 + | Npos p => Ppow_pos P1 P p + end. + + Lemma Ppow_pos_ok : forall l, (forall P, subst_l P@l == P@l) -> + forall res P p, (Ppow_pos res P p)@l == res@l * (pow_pos Pmul P p)@l. + Proof. + intros l subst_l_ok res P p. generalize res;clear res. + induction p;simpl;intros;try rewrite subst_l_ok; repeat rewrite Pmul_ok;repeat rewrite IHp. + rsimpl. mul_push (P@l);rsimpl. rsimpl. rrefl. + Qed. + + Lemma Ppow_N_ok : forall l, (forall P, subst_l P@l == P@l) -> + forall P n, (Ppow_N P n)@l == (pow_N P1 Pmul P n)@l. + Proof. destruct n;simpl. rrefl. rewrite Ppow_pos_ok. trivial. Esimpl. auto. Qed. + + End POWER. + + (** Normalization and rewriting *) + + Section NORM_SUBST_REC. + Variable n : nat. + Variable lmp:list (Mon*Pol). + Let subst_l P := PNSubstL P lmp n n. + Let Pmul_subst P1 P2 := subst_l (Pmul P1 P2). + Let Ppow_subst := Ppow_N subst_l. + + Fixpoint norm_aux (pe:PExpr) : Pol := + match pe with + | PEc c => Pc c + | PEX j => mk_X j + | PEadd (PEopp pe1) pe2 => Psub (norm_aux pe2) (norm_aux pe1) + | PEadd pe1 (PEopp pe2) => + Psub (norm_aux pe1) (norm_aux pe2) + | PEadd pe1 pe2 => Padd (norm_aux pe1) (norm_aux pe2) + | PEsub pe1 pe2 => Psub (norm_aux pe1) (norm_aux pe2) + | PEmul pe1 pe2 => Pmul (norm_aux pe1) (norm_aux pe2) + | PEopp pe1 => Popp (norm_aux pe1) + | PEpow pe1 n => Ppow_N (fun p => p) (norm_aux pe1) n + end. + + Definition norm_subst pe := subst_l (norm_aux pe). + + (* + Fixpoint norm_subst (pe:PExpr) : Pol := + match pe with + | PEc c => Pc c + | PEX j => subst_l (mk_X j) + | PEadd (PEopp pe1) pe2 => Psub (norm_subst pe2) (norm_subst pe1) + | PEadd pe1 (PEopp pe2) => + Psub (norm_subst pe1) (norm_subst pe2) + | PEadd pe1 pe2 => Padd (norm_subst pe1) (norm_subst pe2) + | PEsub pe1 pe2 => Psub (norm_subst pe1) (norm_subst pe2) + | PEmul pe1 pe2 => Pmul_subst (norm_subst pe1) (norm_subst pe2) + | PEopp pe1 => Popp (norm_subst pe1) + | PEpow pe1 n => Ppow_subst (norm_subst pe1) n + end. + + Lemma norm_subst_spec : + forall l pe, MPcond lmp l -> + PEeval l pe == (norm_subst pe)@l. + Proof. + intros;assert (subst_l_ok:forall P, (subst_l P)@l == P@l). + unfold subst_l;intros. + rewrite <- PNSubstL_ok;trivial. rrefl. + assert (Pms_ok:forall P1 P2, (Pmul_subst P1 P2)@l == P1@l*P2@l). + intros;unfold Pmul_subst;rewrite subst_l_ok;rewrite Pmul_ok;rrefl. + induction pe;simpl;Esimpl3. + rewrite subst_l_ok;apply mkX_ok. + rewrite IHpe1;rewrite IHpe2;destruct pe1;destruct pe2;Esimpl3. + rewrite IHpe1;rewrite IHpe2;rrefl. + rewrite Pms_ok;rewrite IHpe1;rewrite IHpe2;rrefl. + rewrite IHpe;rrefl. + unfold Ppow_subst. rewrite Ppow_N_ok. trivial. + rewrite pow_th.(rpow_pow_N). destruct n0;Esimpl3. + induction p;simpl;try rewrite IHp;try rewrite IHpe;repeat rewrite Pms_ok; + repeat rewrite Pmul_ok;rrefl. + Qed. +*) + Lemma norm_aux_spec : + forall l pe, (*MPcond lmp l ->*) + PEeval l pe == (norm_aux pe)@l. + Proof. + intros. + induction pe;simpl;Esimpl3. + apply mkX_ok. + rewrite IHpe1;rewrite IHpe2;destruct pe1;destruct pe2;Esimpl3. + rewrite IHpe1;rewrite IHpe2;rrefl. + rewrite IHpe1;rewrite IHpe2. rewrite Pmul_ok. rrefl. + rewrite IHpe;rrefl. + rewrite Ppow_N_ok by reflexivity. + rewrite pow_th.(rpow_pow_N). destruct n0;Esimpl3. + induction p;simpl;try rewrite IHp;try rewrite IHpe;repeat rewrite Pms_ok; + repeat rewrite Pmul_ok;rrefl. + Qed. + + + End NORM_SUBST_REC. + + +End MakeRingPol. + diff --git a/contrib/micromega/LICENSE.sos b/contrib/micromega/LICENSE.sos new file mode 100644 index 00000000..5aadfa2a --- /dev/null +++ b/contrib/micromega/LICENSE.sos @@ -0,0 +1,29 @@ + HOL Light copyright notice, licence and disclaimer + + (c) University of Cambridge 1998 + (c) Copyright, John Harrison 1998-2006 + +HOL Light version 2.20, hereinafter referred to as "the software", is a +computer theorem proving system written by John Harrison. Much of the +software was developed at the University of Cambridge Computer Laboratory, +New Museums Site, Pembroke Street, Cambridge, CB2 3QG, England. The +software is copyright, University of Cambridge 1998 and John Harrison +1998-2006. + +Permission to use, copy, modify, and distribute the software and its +documentation for any purpose and without fee is hereby granted. In the +case of further distribution of the software the present text, including +copyright notice, licence and disclaimer of warranty, must be included in +full and unmodified form in any release. Distribution of derivative +software obtained by modifying the software, or incorporating it into +other software, is permitted, provided the inclusion of the software is +acknowledged and that any changes made to the software are clearly +documented. + +John Harrison and the University of Cambridge disclaim all warranties +with regard to the software, including all implied warranties of +merchantability and fitness. In no event shall John Harrison or the +University of Cambridge be liable for any special, indirect, +incidental or consequential damages or any damages whatsoever, +including, but not limited to, those arising from computer failure or +malfunction, work stoppage, loss of profit or loss of contracts. diff --git a/contrib/micromega/MExtraction.v b/contrib/micromega/MExtraction.v new file mode 100644 index 00000000..a5ac92db --- /dev/null +++ b/contrib/micromega/MExtraction.v @@ -0,0 +1,23 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* *) +(* Micromega: A reflexive tactic using the Positivstellensatz *) +(* *) +(* Frédéric Besson (Irisa/Inria) 2006-2008 *) +(* *) +(************************************************************************) + +(* Used to generate micromega.ml *) + +Require Import ZMicromega. +Require Import QMicromega. +Require Import VarMap. +Require Import RingMicromega. +Require Import NArith. + +Extraction "micromega.ml" List.map simpl_cone map_cone indexes n_of_Z Nnat.N_of_nat ZTautoChecker QTautoChecker find. diff --git a/contrib/micromega/Micromegatac.v b/contrib/micromega/Micromegatac.v new file mode 100644 index 00000000..13c7eace --- /dev/null +++ b/contrib/micromega/Micromegatac.v @@ -0,0 +1,79 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* *) +(* Micromega: A reflexive tactic using the Positivstellensatz *) +(* *) +(* Frédéric Besson (Irisa/Inria) 2006-2008 *) +(* *) +(************************************************************************) + +Require Import ZMicromega. +Require Import QMicromega. +Require Import RMicromega. +Require Import QArith. +Require Export Ring_normalize. +Require Import ZArith. +Require Import Raxioms. +Require Export RingMicromega. +Require Import VarMap. +Require Tauto. + +Ltac micromegac dom d := + let tac := lazymatch dom with + | Z => + micromegap d ; + intros __wit __varmap __ff ; + change (Tauto.eval_f (Zeval_formula (@find Z Z0 __varmap)) __ff) ; + apply (ZTautoChecker_sound __ff __wit); vm_compute ; reflexivity + | R => + rmicromegap d ; + intros __wit __varmap __ff ; + change (Tauto.eval_f (Reval_formula (@find R 0%R __varmap)) __ff) ; + apply (RTautoChecker_sound __ff __wit); vm_compute ; reflexivity + | _ => fail "Unsupported domain" + end in tac. + +Tactic Notation "micromega" constr(dom) int_or_var(n) := micromegac dom n. +Tactic Notation "micromega" constr(dom) := micromegac dom ltac:-1. + +Ltac zfarkas := omicronp ; + intros __wit __varmap __ff ; + change (Tauto.eval_f (Zeval_formula (@find Z Z0 __varmap)) __ff) ; + apply (ZTautoChecker_sound __ff __wit); vm_compute ; reflexivity. + +Ltac omicron dom := + let tac := lazymatch dom with + | Z => + zomicronp ; + intros __wit __varmap __ff ; + change (Tauto.eval_f (Zeval_formula (@find Z Z0 __varmap)) __ff) ; + apply (ZTautoChecker_sound __ff __wit); vm_compute ; reflexivity + | Q => + qomicronp ; + intros __wit __varmap __ff ; + change (Tauto.eval_f (Qeval_formula (@find Q 0%Q __varmap)) __ff) ; + apply (QTautoChecker_sound __ff __wit); vm_compute ; reflexivity + | R => + romicronp ; + intros __wit __varmap __ff ; + change (Tauto.eval_f (Reval_formula (@find R 0%R __varmap)) __ff) ; + apply (RTautoChecker_sound __ff __wit); vm_compute ; reflexivity + | _ => fail "Unsupported domain" + end in tac. + +Ltac sos dom := + let tac := lazymatch dom with + | Z => + sosp ; + intros __wit __varmap __ff ; + change (Tauto.eval_f (Zeval_formula (@find Z Z0 __varmap)) __ff) ; + apply (ZTautoChecker_sound __ff __wit); vm_compute ; reflexivity + | _ => fail "Unsupported domain" + end in tac. + + diff --git a/contrib/micromega/OrderedRing.v b/contrib/micromega/OrderedRing.v new file mode 100644 index 00000000..149b7731 --- /dev/null +++ b/contrib/micromega/OrderedRing.v @@ -0,0 +1,458 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +Require Import Setoid. +Require Import Ring. + +(** Generic properties of ordered rings on a setoid equality *) + +Set Implicit Arguments. + +Module Import OrderedRingSyntax. +Export RingSyntax. + +Reserved Notation "x ~= y" (at level 70, no associativity). +Reserved Notation "x [=] y" (at level 70, no associativity). +Reserved Notation "x [~=] y" (at level 70, no associativity). +Reserved Notation "x [<] y" (at level 70, no associativity). +Reserved Notation "x [<=] y" (at level 70, no associativity). +End OrderedRingSyntax. + +Section DEFINITIONS. + +Variable R : Type. +Variable (rO rI : R) (rplus rtimes rminus: R -> R -> R) (ropp : R -> R). +Variable req rle rlt : R -> R -> Prop. +Notation "0" := rO. +Notation "1" := rI. +Notation "x + y" := (rplus x y). +Notation "x * y " := (rtimes x y). +Notation "x - y " := (rminus x y). +Notation "- x" := (ropp x). +Notation "x == y" := (req x y). +Notation "x ~= y" := (~ req x y). +Notation "x <= y" := (rle x y). +Notation "x < y" := (rlt x y). + +Record SOR : Type := mk_SOR_theory { + SORsetoid : Setoid_Theory R req; + SORplus_wd : forall x1 x2, x1 == x2 -> forall y1 y2, y1 == y2 -> x1 + y1 == x2 + y2; + SORtimes_wd : forall x1 x2, x1 == x2 -> forall y1 y2, y1 == y2 -> x1 * y1 == x2 * y2; + SORopp_wd : forall x1 x2, x1 == x2 -> -x1 == -x2; + SORle_wd : forall x1 x2, x1 == x2 -> forall y1 y2, y1 == y2 -> (x1 <= y1 <-> x2 <= y2); + SORlt_wd : forall x1 x2, x1 == x2 -> forall y1 y2, y1 == y2 -> (x1 < y1 <-> x2 < y2); + SORrt : ring_theory rO rI rplus rtimes rminus ropp req; + SORle_refl : forall n : R, n <= n; + SORle_antisymm : forall n m : R, n <= m -> m <= n -> n == m; + SORle_trans : forall n m p : R, n <= m -> m <= p -> n <= p; + SORlt_le_neq : forall n m : R, n < m <-> n <= m /\ n ~= m; + SORlt_trichotomy : forall n m : R, n < m \/ n == m \/ m < n; + SORplus_le_mono_l : forall n m p : R, n <= m -> p + n <= p + m; + SORtimes_pos_pos : forall n m : R, 0 < n -> 0 < m -> 0 < n * m; + SORneq_0_1 : 0 ~= 1 +}. + +(* We cannot use Relation_Definitions.order.ord_antisym and +Relations_1.Antisymmetric because they refer to Leibniz equality *) + +End DEFINITIONS. + +Section STRICT_ORDERED_RING. + +Variable R : Type. +Variable (rO rI : R) (rplus rtimes rminus: R -> R -> R) (ropp : R -> R). +Variable req rle rlt : R -> R -> Prop. + +Variable sor : SOR rO rI rplus rtimes rminus ropp req rle rlt. + +Notation "0" := rO. +Notation "1" := rI. +Notation "x + y" := (rplus x y). +Notation "x * y " := (rtimes x y). +Notation "x - y " := (rminus x y). +Notation "- x" := (ropp x). +Notation "x == y" := (req x y). +Notation "x ~= y" := (~ req x y). +Notation "x <= y" := (rle x y). +Notation "x < y" := (rlt x y). + + +Add Relation R req + reflexivity proved by sor.(SORsetoid).(@Equivalence_Reflexive _ _ ) + symmetry proved by sor.(SORsetoid).(@Equivalence_Symmetric _ _ ) + transitivity proved by sor.(SORsetoid).(@Equivalence_Transitive _ _ ) +as sor_setoid. + + +Add Morphism rplus with signature req ==> req ==> req as rplus_morph. +Proof. +exact sor.(SORplus_wd). +Qed. +Add Morphism rtimes with signature req ==> req ==> req as rtimes_morph. +Proof. +exact sor.(SORtimes_wd). +Qed. +Add Morphism ropp with signature req ==> req as ropp_morph. +Proof. +exact sor.(SORopp_wd). +Qed. +Add Morphism rle with signature req ==> req ==> iff as rle_morph. +Proof. +exact sor.(SORle_wd). +Qed. +Add Morphism rlt with signature req ==> req ==> iff as rlt_morph. +Proof. +exact sor.(SORlt_wd). +Qed. + +Add Ring SOR : sor.(SORrt). + +Add Morphism rminus with signature req ==> req ==> req as rminus_morph. +Proof. +intros x1 x2 H1 y1 y2 H2. +rewrite (sor.(SORrt).(Rsub_def) x1 y1). +rewrite (sor.(SORrt).(Rsub_def) x2 y2). +rewrite H1; now rewrite H2. +Qed. + +Theorem Rneq_symm : forall n m : R, n ~= m -> m ~= n. +Proof. +intros n m H1 H2; rewrite H2 in H1; now apply H1. +Qed. + +(* Propeties of plus, minus and opp *) + +Theorem Rplus_0_l : forall n : R, 0 + n == n. +Proof. +intro; ring. +Qed. + +Theorem Rplus_0_r : forall n : R, n + 0 == n. +Proof. +intro; ring. +Qed. + +Theorem Rtimes_0_r : forall n : R, n * 0 == 0. +Proof. +intro; ring. +Qed. + +Theorem Rplus_comm : forall n m : R, n + m == m + n. +Proof. +intros; ring. +Qed. + +Theorem Rtimes_0_l : forall n : R, 0 * n == 0. +Proof. +intro; ring. +Qed. + +Theorem Rtimes_comm : forall n m : R, n * m == m * n. +Proof. +intros; ring. +Qed. + +Theorem Rminus_eq_0 : forall n m : R, n - m == 0 <-> n == m. +Proof. +intros n m. +split; intro H. setoid_replace n with ((n - m) + m) by ring. rewrite H. +now rewrite Rplus_0_l. +rewrite H; ring. +Qed. + +Theorem Rplus_cancel_l : forall n m p : R, p + n == p + m <-> n == m. +Proof. +intros n m p; split; intro H. +setoid_replace n with (- p + (p + n)) by ring. +setoid_replace m with (- p + (p + m)) by ring. now rewrite H. +now rewrite H. +Qed. + +(* Relations *) + +Theorem Rle_refl : forall n : R, n <= n. +Proof sor.(SORle_refl). + +Theorem Rle_antisymm : forall n m : R, n <= m -> m <= n -> n == m. +Proof sor.(SORle_antisymm). + +Theorem Rle_trans : forall n m p : R, n <= m -> m <= p -> n <= p. +Proof sor.(SORle_trans). + +Theorem Rlt_trichotomy : forall n m : R, n < m \/ n == m \/ m < n. +Proof sor.(SORlt_trichotomy). + +Theorem Rlt_le_neq : forall n m : R, n < m <-> n <= m /\ n ~= m. +Proof sor.(SORlt_le_neq). + +Theorem Rneq_0_1 : 0 ~= 1. +Proof sor.(SORneq_0_1). + +Theorem Req_em : forall n m : R, n == m \/ n ~= m. +Proof. +intros n m. destruct (Rlt_trichotomy n m) as [H | [H | H]]; try rewrite Rlt_le_neq in H. +right; now destruct H. +now left. +right; apply Rneq_symm; now destruct H. +Qed. + +Theorem Req_dne : forall n m : R, ~ ~ n == m <-> n == m. +Proof. +intros n m; destruct (Req_em n m) as [H | H]. +split; auto. +split. intro H1; false_hyp H H1. auto. +Qed. + +Theorem Rle_lt_eq : forall n m : R, n <= m <-> n < m \/ n == m. +Proof. +intros n m; rewrite Rlt_le_neq. +split; [intro H | intros [[H1 H2] | H]]. +destruct (Req_em n m) as [H1 | H1]. now right. left; now split. +assumption. +rewrite H; apply Rle_refl. +Qed. + +Ltac le_less := rewrite Rle_lt_eq; left; try assumption. +Ltac le_equal := rewrite Rle_lt_eq; right; try reflexivity; try assumption. +Ltac le_elim H := rewrite Rle_lt_eq in H; destruct H as [H | H]. + +Theorem Rlt_trans : forall n m p : R, n < m -> m < p -> n < p. +Proof. +intros n m p; repeat rewrite Rlt_le_neq; intros [H1 H2] [H3 H4]; split. +now apply Rle_trans with m. +intro H. rewrite H in H1. pose proof (Rle_antisymm H3 H1). now apply H4. +Qed. + +Theorem Rle_lt_trans : forall n m p : R, n <= m -> m < p -> n < p. +Proof. +intros n m p H1 H2; le_elim H1. +now apply Rlt_trans with (m := m). now rewrite H1. +Qed. + +Theorem Rlt_le_trans : forall n m p : R, n < m -> m <= p -> n < p. +Proof. +intros n m p H1 H2; le_elim H2. +now apply Rlt_trans with (m := m). now rewrite <- H2. +Qed. + +Theorem Rle_gt_cases : forall n m : R, n <= m \/ m < n. +Proof. +intros n m; destruct (Rlt_trichotomy n m) as [H | [H | H]]. +left; now le_less. left; now le_equal. now right. +Qed. + +Theorem Rlt_neq : forall n m : R, n < m -> n ~= m. +Proof. +intros n m; rewrite Rlt_le_neq; now intros [_ H]. +Qed. + +Theorem Rle_ngt : forall n m : R, n <= m <-> ~ m < n. +Proof. +intros n m; split. +intros H H1; assert (H2 : n < n) by now apply Rle_lt_trans with m. now apply (Rlt_neq H2). +intro H. destruct (Rle_gt_cases n m) as [H1 | H1]. assumption. false_hyp H1 H. +Qed. + +Theorem Rlt_nge : forall n m : R, n < m <-> ~ m <= n. +Proof. +intros n m; split. +intros H H1; assert (H2 : n < n) by now apply Rlt_le_trans with m. now apply (Rlt_neq H2). +intro H. destruct (Rle_gt_cases m n) as [H1 | H1]. false_hyp H1 H. assumption. +Qed. + +(* Plus, minus and order *) + +Theorem Rplus_le_mono_l : forall n m p : R, n <= m <-> p + n <= p + m. +Proof. +intros n m p; split. +apply sor.(SORplus_le_mono_l). +intro H. apply (sor.(SORplus_le_mono_l) (p + n) (p + m) (- p)) in H. +setoid_replace (- p + (p + n)) with n in H by ring. +setoid_replace (- p + (p + m)) with m in H by ring. assumption. +Qed. + +Theorem Rplus_le_mono_r : forall n m p : R, n <= m <-> n + p <= m + p. +Proof. +intros n m p; rewrite (Rplus_comm n p); rewrite (Rplus_comm m p). +apply Rplus_le_mono_l. +Qed. + +Theorem Rplus_lt_mono_l : forall n m p : R, n < m <-> p + n < p + m. +Proof. +intros n m p; do 2 rewrite Rlt_le_neq. rewrite Rplus_cancel_l. +now rewrite <- Rplus_le_mono_l. +Qed. + +Theorem Rplus_lt_mono_r : forall n m p : R, n < m <-> n + p < m + p. +Proof. +intros n m p. +rewrite (Rplus_comm n p); rewrite (Rplus_comm m p); apply Rplus_lt_mono_l. +Qed. + +Theorem Rplus_lt_mono : forall n m p q : R, n < m -> p < q -> n + p < m + q. +Proof. +intros n m p q H1 H2. +apply Rlt_trans with (m + p); [now apply -> Rplus_lt_mono_r | now apply -> Rplus_lt_mono_l]. +Qed. + +Theorem Rplus_le_mono : forall n m p q : R, n <= m -> p <= q -> n + p <= m + q. +Proof. +intros n m p q H1 H2. +apply Rle_trans with (m + p); [now apply -> Rplus_le_mono_r | now apply -> Rplus_le_mono_l]. +Qed. + +Theorem Rplus_lt_le_mono : forall n m p q : R, n < m -> p <= q -> n + p < m + q. +Proof. +intros n m p q H1 H2. +apply Rlt_le_trans with (m + p); [now apply -> Rplus_lt_mono_r | now apply -> Rplus_le_mono_l]. +Qed. + +Theorem Rplus_le_lt_mono : forall n m p q : R, n <= m -> p < q -> n + p < m + q. +Proof. +intros n m p q H1 H2. +apply Rle_lt_trans with (m + p); [now apply -> Rplus_le_mono_r | now apply -> Rplus_lt_mono_l]. +Qed. + +Theorem Rplus_pos_pos : forall n m : R, 0 < n -> 0 < m -> 0 < n + m. +Proof. +intros n m H1 H2. rewrite <- (Rplus_0_l 0). now apply Rplus_lt_mono. +Qed. + +Theorem Rplus_pos_nonneg : forall n m : R, 0 < n -> 0 <= m -> 0 < n + m. +Proof. +intros n m H1 H2. rewrite <- (Rplus_0_l 0). now apply Rplus_lt_le_mono. +Qed. + +Theorem Rplus_nonneg_pos : forall n m : R, 0 <= n -> 0 < m -> 0 < n + m. +Proof. +intros n m H1 H2. rewrite <- (Rplus_0_l 0). now apply Rplus_le_lt_mono. +Qed. + +Theorem Rplus_nonneg_nonneg : forall n m : R, 0 <= n -> 0 <= m -> 0 <= n + m. +Proof. +intros n m H1 H2. rewrite <- (Rplus_0_l 0). now apply Rplus_le_mono. +Qed. + +Theorem Rle_le_minus : forall n m : R, n <= m <-> 0 <= m - n. +Proof. +intros n m. rewrite (@Rplus_le_mono_r n m (- n)). +setoid_replace (n + - n) with 0 by ring. +now setoid_replace (m + - n) with (m - n) by ring. +Qed. + +Theorem Rlt_lt_minus : forall n m : R, n < m <-> 0 < m - n. +Proof. +intros n m. rewrite (@Rplus_lt_mono_r n m (- n)). +setoid_replace (n + - n) with 0 by ring. +now setoid_replace (m + - n) with (m - n) by ring. +Qed. + +Theorem Ropp_lt_mono : forall n m : R, n < m <-> - m < - n. +Proof. +intros n m. split; intro H. +apply -> (@Rplus_lt_mono_l n m (- n - m)) in H. +setoid_replace (- n - m + n) with (- m) in H by ring. +now setoid_replace (- n - m + m) with (- n) in H by ring. +apply -> (@Rplus_lt_mono_l (- m) (- n) (n + m)) in H. +setoid_replace (n + m + - m) with n in H by ring. +now setoid_replace (n + m + - n) with m in H by ring. +Qed. + +Theorem Ropp_pos_neg : forall n : R, 0 < - n <-> n < 0. +Proof. +intro n; rewrite (Ropp_lt_mono n 0). now setoid_replace (- 0) with 0 by ring. +Qed. + +(* Times and order *) + +Theorem Rtimes_pos_pos : forall n m : R, 0 < n -> 0 < m -> 0 < n * m. +Proof sor.(SORtimes_pos_pos). + +Theorem Rtimes_nonneg_nonneg : forall n m : R, 0 <= n -> 0 <= m -> 0 <= n * m. +Proof. +intros n m H1 H2. +le_elim H1. le_elim H2. +le_less; now apply Rtimes_pos_pos. +rewrite <- H2; rewrite Rtimes_0_r; le_equal. +rewrite <- H1; rewrite Rtimes_0_l; le_equal. +Qed. + +Theorem Rtimes_pos_neg : forall n m : R, 0 < n -> m < 0 -> n * m < 0. +Proof. +intros n m H1 H2. apply -> Ropp_pos_neg. +setoid_replace (- (n * m)) with (n * (- m)) by ring. +apply Rtimes_pos_pos. assumption. now apply <- Ropp_pos_neg. +Qed. + +Theorem Rtimes_neg_neg : forall n m : R, n < 0 -> m < 0 -> 0 < n * m. +Proof. +intros n m H1 H2. +setoid_replace (n * m) with ((- n) * (- m)) by ring. +apply Rtimes_pos_pos; now apply <- Ropp_pos_neg. +Qed. + +Theorem Rtimes_square_nonneg : forall n : R, 0 <= n * n. +Proof. +intro n; destruct (Rlt_trichotomy 0 n) as [H | [H | H]]. +le_less; now apply Rtimes_pos_pos. +rewrite <- H, Rtimes_0_l; le_equal. +le_less; now apply Rtimes_neg_neg. +Qed. + +Theorem Rtimes_neq_0 : forall n m : R, n ~= 0 /\ m ~= 0 -> n * m ~= 0. +Proof. +intros n m [H1 H2]. +destruct (Rlt_trichotomy n 0) as [H3 | [H3 | H3]]; +destruct (Rlt_trichotomy m 0) as [H4 | [H4 | H4]]; +try (false_hyp H3 H1); try (false_hyp H4 H2). +apply Rneq_symm. apply Rlt_neq. now apply Rtimes_neg_neg. +apply Rlt_neq. rewrite Rtimes_comm. now apply Rtimes_pos_neg. +apply Rlt_neq. now apply Rtimes_pos_neg. +apply Rneq_symm. apply Rlt_neq. now apply Rtimes_pos_pos. +Qed. + +(* The following theorems are used to build a morphism from Z to R and +prove its properties in ZCoeff.v. They are not used in RingMicromega.v. *) + +(* Surprisingly, multilication is needed to prove the following theorem *) + +Theorem Ropp_neg_pos : forall n : R, - n < 0 <-> 0 < n. +Proof. +intro n; setoid_replace n with (- - n) by ring. rewrite Ropp_pos_neg. +now setoid_replace (- - n) with n by ring. +Qed. + +Theorem Rlt_0_1 : 0 < 1. +Proof. +apply <- Rlt_le_neq. split. +setoid_replace 1 with (1 * 1) by ring. apply Rtimes_square_nonneg. +apply Rneq_0_1. +Qed. + +Theorem Rlt_succ_r : forall n : R, n < 1 + n. +Proof. +intro n. rewrite <- (Rplus_0_l n); setoid_replace (1 + (0 + n)) with (1 + n) by ring. +apply -> Rplus_lt_mono_r. apply Rlt_0_1. +Qed. + +Theorem Rlt_lt_succ : forall n m : R, n < m -> n < 1 + m. +Proof. +intros n m H; apply Rlt_trans with m. assumption. apply Rlt_succ_r. +Qed. + +(*Theorem Rtimes_lt_mono_pos_l : forall n m p : R, 0 < p -> n < m -> p * n < p * m. +Proof. +intros n m p H1 H2. apply <- Rlt_lt_minus. +setoid_replace (p * m - p * n) with (p * (m - n)) by ring. +apply Rtimes_pos_pos. assumption. now apply -> Rlt_lt_minus. +Qed.*) + +End STRICT_ORDERED_RING. + diff --git a/contrib/micromega/QMicromega.v b/contrib/micromega/QMicromega.v new file mode 100644 index 00000000..9e95f6c4 --- /dev/null +++ b/contrib/micromega/QMicromega.v @@ -0,0 +1,259 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* *) +(* Micromega: A reflexive tactic using the Positivstellensatz *) +(* *) +(* Frédéric Besson (Irisa/Inria) 2006-2008 *) +(* *) +(************************************************************************) + +Require Import OrderedRing. +Require Import RingMicromega. +Require Import Refl. +Require Import QArith. +Require Import Qring. + +(* Qsrt has been removed from the library ? *) +Definition Qsrt : ring_theory 0 1 Qplus Qmult Qminus Qopp Qeq. +Proof. + constructor. + exact Qplus_0_l. + exact Qplus_comm. + exact Qplus_assoc. + exact Qmult_1_l. + exact Qmult_comm. + exact Qmult_assoc. + exact Qmult_plus_distr_l. + reflexivity. + exact Qplus_opp_r. +Qed. + + +Add Ring Qring : Qsrt. + +Lemma Qmult_neutral : forall x , 0 * x == 0. +Proof. + intros. + compute. + reflexivity. +Qed. + +(* Is there any qarith database ? *) + +Lemma Qsor : SOR 0 1 Qplus Qmult Qminus Qopp Qeq Qle Qlt. +Proof. + constructor; intros ; subst ; try (intuition (subst; auto with qarith)). + apply Q_Setoid. + rewrite H ; rewrite H0 ; reflexivity. + rewrite H ; rewrite H0 ; reflexivity. + rewrite H ; auto ; reflexivity. + rewrite <- H ; rewrite <- H0 ; auto. + rewrite H ; rewrite H0 ; auto. + rewrite <- H ; rewrite <- H0 ; auto. + rewrite H ; rewrite H0 ; auto. + apply Qsrt. + apply Qle_refl. + apply Qle_antisym ; auto. + eapply Qle_trans ; eauto. + apply Qlt_le_weak ; auto. + apply (Qlt_not_eq n m H H0) ; auto. + destruct (Qle_lt_or_eq _ _ H0) ; auto. + tauto. + destruct(Q_dec n m) as [[H1 |H1] | H1 ] ; tauto. + apply (Qplus_le_compat p p n m (Qle_refl p) H). + generalize (Qmult_lt_compat_r 0 n m H0 H). + rewrite Qmult_neutral. + auto. + compute in H. + discriminate. +Qed. + +Definition Qeq_bool (p q : Q) : bool := Zeq_bool (Qnum p * ' Qden q)%Z (Qnum q * ' Qden p)%Z. + +Definition Qle_bool (x y : Q) : bool := Zle_bool (Qnum x * ' Qden y)%Z (Qnum y * ' Qden x)%Z. + +Require ZMicromega. + +Lemma Qeq_bool_ok : forall x y, Qeq_bool x y = true -> x == y. +Proof. + intros. + unfold Qeq_bool in H. + unfold Qeq. + apply (Zeqb_ok _ _ H). +Qed. + + +Lemma Qeq_bool_neq : forall x y, Qeq_bool x y = false -> ~ x == y. +Proof. + unfold Qeq_bool,Qeq. + red ; intros ; subst. + rewrite H0 in H. + apply (ZMicromega.Zeq_bool_neq _ _ H). + reflexivity. +Qed. + +Lemma Qle_bool_imp_le : forall x y : Q, Qle_bool x y = true -> x <= y. +Proof. + unfold Qle_bool, Qle. + intros. + apply Zle_bool_imp_le ; auto. +Qed. + + + + +Lemma QSORaddon : + SORaddon 0 1 Qplus Qmult Qminus Qopp Qeq Qle (* ring elements *) + 0 1 Qplus Qmult Qminus Qopp (* coefficients *) + Qeq_bool Qle_bool + (fun x => x) (fun x => x) (pow_N 1 Qmult). +Proof. + constructor. + constructor ; intros ; try reflexivity. + apply Qeq_bool_ok ; auto. + constructor. + reflexivity. + intros x y. + apply Qeq_bool_neq ; auto. + apply Qle_bool_imp_le. +Qed. + + +(*Definition Zeval_expr := eval_pexpr 0 Zplus Zmult Zminus Zopp (fun x => x) (fun x => Z_of_N x) (Zpower).*) +Require Import EnvRing. + +Fixpoint Qeval_expr (env: PolEnv Q) (e: PExpr Q) : Q := + match e with + | PEc c => c + | PEX j => env j + | PEadd pe1 pe2 => (Qeval_expr env pe1) + (Qeval_expr env pe2) + | PEsub pe1 pe2 => (Qeval_expr env pe1) - (Qeval_expr env pe2) + | PEmul pe1 pe2 => (Qeval_expr env pe1) * (Qeval_expr env pe2) + | PEopp pe1 => - (Qeval_expr env pe1) + | PEpow pe1 n => Qpower (Qeval_expr env pe1) (Z_of_N n) + end. + +Lemma Qeval_expr_simpl : forall env e, + Qeval_expr env e = + match e with + | PEc c => c + | PEX j => env j + | PEadd pe1 pe2 => (Qeval_expr env pe1) + (Qeval_expr env pe2) + | PEsub pe1 pe2 => (Qeval_expr env pe1) - (Qeval_expr env pe2) + | PEmul pe1 pe2 => (Qeval_expr env pe1) * (Qeval_expr env pe2) + | PEopp pe1 => - (Qeval_expr env pe1) + | PEpow pe1 n => Qpower (Qeval_expr env pe1) (Z_of_N n) + end. +Proof. + destruct e ; reflexivity. +Qed. + +Definition Qeval_expr' := eval_pexpr Qplus Qmult Qminus Qopp (fun x => x) (fun x => x) (pow_N 1 Qmult). + +Lemma QNpower : forall r n, r ^ Z_of_N n = pow_N 1 Qmult r n. +Proof. + destruct n ; reflexivity. +Qed. + + +Lemma Qeval_expr_compat : forall env e, Qeval_expr env e = Qeval_expr' env e. +Proof. + induction e ; simpl ; subst ; try congruence. + rewrite IHe. + apply QNpower. +Qed. + +Definition Qeval_op2 (o : Op2) : Q -> Q -> Prop := +match o with +| OpEq => Qeq +| OpNEq => fun x y => ~ x == y +| OpLe => Qle +| OpGe => Qge +| OpLt => Qlt +| OpGt => Qgt +end. + +Definition Qeval_formula (e:PolEnv Q) (ff : Formula Q) := + let (lhs,o,rhs) := ff in Qeval_op2 o (Qeval_expr e lhs) (Qeval_expr e rhs). + +Definition Qeval_formula' := + eval_formula Qplus Qmult Qminus Qopp Qeq Qle Qlt (fun x => x) (fun x => x) (pow_N 1 Qmult). + +Lemma Qeval_formula_compat : forall env f, Qeval_formula env f <-> Qeval_formula' env f. +Proof. + intros. + unfold Qeval_formula. + destruct f. + repeat rewrite Qeval_expr_compat. + unfold Qeval_formula'. + unfold Qeval_expr'. + split ; destruct Fop ; simpl; auto. +Qed. + + + +Definition Qeval_nformula := + eval_nformula 0 Qplus Qmult Qminus Qopp Qeq Qle Qlt (fun x => x) (fun x => x) (pow_N 1 Qmult). + +Definition Qeval_op1 (o : Op1) : Q -> Prop := +match o with +| Equal => fun x : Q => x == 0 +| NonEqual => fun x : Q => ~ x == 0 +| Strict => fun x : Q => 0 < x +| NonStrict => fun x : Q => 0 <= x +end. + +Lemma Qeval_nformula_simpl : forall env f, Qeval_nformula env f = (let (p, op) := f in Qeval_op1 op (Qeval_expr env p)). +Proof. + intros. + destruct f. + rewrite Qeval_expr_compat. + reflexivity. +Qed. + +Lemma Qeval_nformula_dec : forall env d, (Qeval_nformula env d) \/ ~ (Qeval_nformula env d). +Proof. + exact (fun env d =>eval_nformula_dec Qsor (fun x => x) (fun x => x) (pow_N 1 Qmult) env d). +Qed. + +Definition QWitness := ConeMember Q. + +Definition QWeakChecker := check_normalised_formulas 0 1 Qplus Qmult Qminus Qopp Qeq_bool Qle_bool. + +Require Import List. + +Lemma QWeakChecker_sound : forall (l : list (NFormula Q)) (cm : QWitness), + QWeakChecker l cm = true -> + forall env, make_impl (Qeval_nformula env) l False. +Proof. + intros l cm H. + intro. + unfold Qeval_nformula. + apply (checker_nf_sound Qsor QSORaddon l cm). + unfold QWeakChecker in H. + exact H. +Qed. + +Require Import Tauto. + +Definition QTautoChecker (f : BFormula (Formula Q)) (w: list QWitness) : bool := + @tauto_checker (Formula Q) (NFormula Q) (@cnf_normalise Q) (@cnf_negate Q) QWitness QWeakChecker f w. + +Lemma QTautoChecker_sound : forall f w, QTautoChecker f w = true -> forall env, eval_f (Qeval_formula env) f. +Proof. + intros f w. + unfold QTautoChecker. + apply (tauto_checker_sound Qeval_formula Qeval_nformula). + apply Qeval_nformula_dec. + intros. rewrite Qeval_formula_compat. unfold Qeval_formula'. now apply (cnf_normalise_correct Qsor). + intros. rewrite Qeval_formula_compat. unfold Qeval_formula'. now apply (cnf_negate_correct Qsor). + intros t w0. + apply QWeakChecker_sound. +Qed. + + diff --git a/contrib/micromega/RMicromega.v b/contrib/micromega/RMicromega.v new file mode 100644 index 00000000..ef28db32 --- /dev/null +++ b/contrib/micromega/RMicromega.v @@ -0,0 +1,148 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* *) +(* Micromega: A reflexive tactic using the Positivstellensatz *) +(* *) +(* Frédéric Besson (Irisa/Inria) 2006-2008 *) +(* *) +(************************************************************************) + +Require Import OrderedRing. +Require Import RingMicromega. +Require Import Refl. +Require Import Raxioms RIneq Rpow_def DiscrR. +Require Setoid. + +Definition Rsrt : ring_theory R0 R1 Rplus Rmult Rminus Ropp (@eq R). +Proof. + constructor. + exact Rplus_0_l. + exact Rplus_comm. + intros. rewrite Rplus_assoc. auto. + exact Rmult_1_l. + exact Rmult_comm. + intros ; rewrite Rmult_assoc ; auto. + intros. rewrite Rmult_comm. rewrite Rmult_plus_distr_l. + rewrite (Rmult_comm z). rewrite (Rmult_comm z). auto. + reflexivity. + exact Rplus_opp_r. +Qed. + +Add Ring Rring : Rsrt. +Open Scope R_scope. + +Lemma Rmult_neutral : forall x:R , 0 * x = 0. +Proof. + intro ; ring. +Qed. + + +Lemma Rsor : SOR R0 R1 Rplus Rmult Rminus Ropp (@eq R) Rle Rlt. +Proof. + constructor; intros ; subst ; try (intuition (subst; try ring ; auto with real)). + constructor. + constructor. + unfold RelationClasses.Symmetric. auto. + unfold RelationClasses.Transitive. intros. subst. reflexivity. + apply Rsrt. + eapply Rle_trans ; eauto. + apply (Rlt_irrefl m) ; auto. + apply Rnot_le_lt. auto with real. + destruct (total_order_T n m) as [ [H1 | H1] | H1] ; auto. + intros. + rewrite <- (Rmult_neutral m). + apply (Rmult_lt_compat_r) ; auto. +Qed. + +Require ZMicromega. + +(* R with coeffs in Z *) + +Lemma RZSORaddon : + SORaddon R0 R1 Rplus Rmult Rminus Ropp (@eq R) Rle (* ring elements *) + 0%Z 1%Z Zplus Zmult Zminus Zopp (* coefficients *) + Zeq_bool Zle_bool + IZR Nnat.nat_of_N pow. +Proof. + constructor. + constructor ; intros ; try reflexivity. + apply plus_IZR. + symmetry. apply Z_R_minus. + apply mult_IZR. + apply Ropp_Ropp_IZR. + apply IZR_eq. + apply Zeqb_ok ; auto. + apply R_power_theory. + intros x y. + intro. + apply IZR_neq. + apply ZMicromega.Zeq_bool_neq ; auto. + intros. apply IZR_le. apply Zle_bool_imp_le. auto. +Qed. + + +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 IZR Nnat.nat_of_N pow. + + +Definition Reval_formula := + eval_formula Rplus Rmult Rminus Ropp (@eq R) Rle Rlt IZR Nnat.nat_of_N pow. + + +Definition Reval_nformula := + eval_nformula 0 Rplus Rmult Rminus Ropp (@eq R) Rle Rlt IZR Nnat.nat_of_N pow. + + +Lemma Reval_nformula_dec : forall env d, (Reval_nformula env d) \/ ~ (Reval_nformula env d). +Proof. + exact (fun env d =>eval_nformula_dec Rsor IZR Nnat.nat_of_N pow env d). +Qed. + +Definition RWitness := ConeMember Z. + +Definition RWeakChecker := check_normalised_formulas 0%Z 1%Z Zplus Zmult Zminus Zopp Zeq_bool Zle_bool. + +Require Import List. + +Lemma RWeakChecker_sound : forall (l : list (NFormula Z)) (cm : RWitness), + RWeakChecker l cm = true -> + forall env, make_impl (Reval_nformula env) l False. +Proof. + intros l cm H. + intro. + unfold Reval_nformula. + apply (checker_nf_sound Rsor RZSORaddon l cm). + unfold RWeakChecker in H. + exact H. +Qed. + +Require Import Tauto. + +Definition RTautoChecker (f : BFormula (Formula Z)) (w: list RWitness) : bool := + @tauto_checker (Formula Z) (NFormula Z) (@cnf_normalise Z) (@cnf_negate Z) RWitness RWeakChecker f w. + +Lemma RTautoChecker_sound : forall f w, RTautoChecker f w = true -> forall env, eval_f (Reval_formula env) f. +Proof. + intros f w. + unfold RTautoChecker. + apply (tauto_checker_sound Reval_formula Reval_nformula). + apply Reval_nformula_dec. + intros. unfold Reval_formula. now apply (cnf_normalise_correct Rsor). + intros. unfold Reval_formula. now apply (cnf_negate_correct Rsor). + intros t w0. + apply RWeakChecker_sound. +Qed. + + diff --git a/contrib/micromega/Refl.v b/contrib/micromega/Refl.v new file mode 100644 index 00000000..801d8b21 --- /dev/null +++ b/contrib/micromega/Refl.v @@ -0,0 +1,129 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* *) +(* Micromega: A reflexive tactic using the Positivstellensatz *) +(* *) +(* Frédéric Besson (Irisa/Inria) 2006-2008 *) +(* *) +(************************************************************************) + +Require Import List. +Require Setoid. + +Set Implicit Arguments. + +(* Refl of '->' '/\': basic properties *) + +Fixpoint make_impl (A : Type) (eval : A -> Prop) (l : list A) (goal : Prop) {struct l} : Prop := + match l with + | nil => goal + | cons e l => (eval e) -> (make_impl eval l goal) + end. + +Theorem make_impl_true : + forall (A : Type) (eval : A -> Prop) (l : list A), make_impl eval l True. +Proof. +induction l as [| a l IH]; simpl. +trivial. +intro; apply IH. +Qed. + +Fixpoint make_conj (A : Type) (eval : A -> Prop) (l : list A) {struct l} : Prop := + match l with + | nil => True + | cons e nil => (eval e) + | cons e l2 => ((eval e) /\ (make_conj eval l2)) + end. + +Theorem make_conj_cons : forall (A : Type) (eval : A -> Prop) (a : A) (l : list A), + make_conj eval (a :: l) <-> eval a /\ make_conj eval l. +Proof. +intros; destruct l; simpl; tauto. +Qed. + + +Lemma make_conj_impl : forall (A : Type) (eval : A -> Prop) (l : list A) (g : Prop), + (make_conj eval l -> g) <-> make_impl eval l g. +Proof. + induction l. + simpl. + tauto. + simpl. + intros. + destruct l. + simpl. + tauto. + generalize (IHl g). + tauto. +Qed. + +Lemma make_conj_in : forall (A : Type) (eval : A -> Prop) (l : list A), + make_conj eval l -> (forall p, In p l -> eval p). +Proof. + induction l. + simpl. + tauto. + simpl. + intros. + destruct l. + simpl in H0. + destruct H0. + subst; auto. + tauto. + destruct H. + destruct H0. + subst;auto. + apply IHl; auto. +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. + induction l1. + simpl. + tauto. + intros. + change ((a::l1) ++ l2) with (a :: (l1 ++ l2)). + rewrite make_conj_cons. + rewrite IHl1. + rewrite make_conj_cons. + tauto. +Qed. + +Lemma not_make_conj_cons : forall (A:Type) (t:A) a eval (no_middle_eval : (eval t) \/ ~ (eval t)), + ~ make_conj eval (t ::a) -> ~ (eval t) \/ (~ make_conj eval a). +Proof. + intros. + simpl in H. + destruct a. + tauto. + tauto. +Qed. + +Lemma not_make_conj_app : forall (A:Type) (t:list A) a eval + (no_middle_eval : forall d, eval d \/ ~ eval d) , + ~ make_conj eval (t ++ a) -> (~ make_conj eval t) \/ (~ make_conj eval a). +Proof. + induction t. + simpl. + tauto. + intros. + simpl ((a::t)++a0)in H. + destruct (@not_make_conj_cons _ _ _ _ (no_middle_eval a) H). + left ; red ; intros. + apply H0. + rewrite make_conj_cons in H1. + tauto. + destruct (IHt _ _ no_middle_eval H0). + left ; red ; intros. + apply H1. + rewrite make_conj_cons in H2. + tauto. + right ; auto. +Qed. diff --git a/contrib/micromega/RingMicromega.v b/contrib/micromega/RingMicromega.v new file mode 100644 index 00000000..6885b82c --- /dev/null +++ b/contrib/micromega/RingMicromega.v @@ -0,0 +1,779 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +Require Import NArith. +Require Import Relation_Definitions. +Require Import Setoid. +(*****) +Require Import Env. +Require Import EnvRing. +(*****) +Require Import List. +Require Import Bool. +Require Import OrderedRing. +Require Import Refl. + + +Set Implicit Arguments. + +Import OrderedRingSyntax. + +Section Micromega. + +(* Assume we have a strict(ly?) ordered ring *) + +Variable 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 sor : SOR rO rI rplus rtimes rminus ropp req rle rlt. + +Notation "0" := rO. +Notation "1" := rI. +Notation "x + y" := (rplus x y). +Notation "x * y " := (rtimes x y). +Notation "x - y " := (rminus x y). +Notation "- x" := (ropp x). +Notation "x == y" := (req x y). +Notation "x ~= y" := (~ req x y). +Notation "x <= y" := (rle x y). +Notation "x < y" := (rlt x y). + +(* Assume we have a type of coefficients C and a morphism from C to R *) + +Variable C : Type. +Variables cO cI : C. +Variables cplus ctimes cminus: C -> C -> C. +Variable copp : C -> C. +Variables ceqb cleb : C -> C -> bool. +Variable phi : C -> R. + +(* Power coefficients *) +Variable E : Set. (* the type of exponents *) +Variable pow_phi : N -> E. +Variable rpow : R -> E -> R. + +Notation "[ x ]" := (phi x). +Notation "x [=] y" := (ceqb x y). +Notation "x [<=] y" := (cleb x y). + +(* Let's collect all hypotheses in addition to the ordered ring axioms into +one structure *) + +Record SORaddon := mk_SOR_addon { + SORrm : ring_morph 0 1 rplus rtimes rminus ropp req cO cI cplus ctimes cminus copp ceqb phi; + SORpower : power_theory rI rtimes req pow_phi rpow; + SORcneqb_morph : forall x y : C, x [=] y = false -> [x] ~= [y]; + SORcleb_morph : forall x y : C, x [<=] y = true -> [x] <= [y] +}. + +Variable addon : SORaddon. + +Add Relation R req + reflexivity proved by sor.(SORsetoid).(@Equivalence_Reflexive _ _ ) + symmetry proved by sor.(SORsetoid).(@Equivalence_Symmetric _ _ ) + transitivity proved by sor.(SORsetoid).(@Equivalence_Transitive _ _ ) +as micomega_sor_setoid. + +Add Morphism rplus with signature req ==> req ==> req as rplus_morph. +Proof. +exact sor.(SORplus_wd). +Qed. +Add Morphism rtimes with signature req ==> req ==> req as rtimes_morph. +Proof. +exact sor.(SORtimes_wd). +Qed. +Add Morphism ropp with signature req ==> req as ropp_morph. +Proof. +exact sor.(SORopp_wd). +Qed. +Add Morphism rle with signature req ==> req ==> iff as rle_morph. +Proof. + exact sor.(SORle_wd). +Qed. +Add Morphism rlt with signature req ==> req ==> iff as rlt_morph. +Proof. + exact sor.(SORlt_wd). +Qed. + +Add Morphism rminus with signature req ==> req ==> req as rminus_morph. +Proof. + exact (rminus_morph sor). (* We already proved that minus is a morphism in OrderedRing.v *) +Qed. + +Definition cneqb (x y : C) := negb (ceqb x y). +Definition cltb (x y : C) := (cleb x y) && (cneqb x y). + +Notation "x [~=] y" := (cneqb x y). +Notation "x [<] y" := (cltb x y). + +Ltac le_less := rewrite (Rle_lt_eq sor); left; try assumption. +Ltac le_equal := rewrite (Rle_lt_eq sor); right; try reflexivity; try assumption. +Ltac le_elim H := rewrite (Rle_lt_eq sor) in H; destruct H as [H | H]. + +Lemma cleb_sound : forall x y : C, x [<=] y = true -> [x] <= [y]. +Proof. + exact addon.(SORcleb_morph). +Qed. + +Lemma cneqb_sound : forall x y : C, x [~=] y = true -> [x] ~= [y]. +Proof. +intros x y H1. apply addon.(SORcneqb_morph). unfold cneqb, negb in H1. +destruct (ceqb x y); now try discriminate. +Qed. + +Lemma cltb_sound : forall x y : C, x [<] y = true -> [x] < [y]. +Proof. +intros x y H. unfold cltb in H. apply andb_prop in H. destruct H as [H1 H2]. +apply cleb_sound in H1. apply cneqb_sound in H2. apply <- (Rlt_le_neq sor). now split. +Qed. + +(* Begin Micromega *) + +Definition PExprC := PExpr C. (* arbitrary expressions built from +, *, - *) +Definition PolC := Pol C. (* polynomials in generalized Horner form, defined in Ring_polynom or EnvRing *) +(*****) +(*Definition Env := Env R. (* For interpreting PExprC *)*) +Definition PolEnv := Env R. (* For interpreting PolC *) +(*****) +(*Definition Env := list R. +Definition PolEnv := list R.*) +(*****) + +(* What benefit do we get, in the case of EnvRing, from defining eval_pexpr +explicitely below and not through PEeval, as the following lemma says? The +function eval_pexpr seems to be a straightforward special case of PEeval +when the environment (i.e., the second last argument of PEeval) of type +off_map (which is (option positive * t)) is (None, env). *) + +(*****) +Fixpoint eval_pexpr (l : PolEnv) (pe : PExprC) {struct pe} : R := +match pe with +| PEc c => phi c +| PEX j => l j +| PEadd pe1 pe2 => (eval_pexpr l pe1) + (eval_pexpr l pe2) +| PEsub pe1 pe2 => (eval_pexpr l pe1) - (eval_pexpr l pe2) +| PEmul pe1 pe2 => (eval_pexpr l pe1) * (eval_pexpr l pe2) +| PEopp pe1 => - (eval_pexpr l pe1) +| PEpow pe1 n => rpow (eval_pexpr l pe1) (pow_phi n) +end. + + +Lemma eval_pexpr_simpl : forall (l : PolEnv) (pe : PExprC), + eval_pexpr l pe = + match pe with + | PEc c => phi c + | PEX j => l j + | PEadd pe1 pe2 => (eval_pexpr l pe1) + (eval_pexpr l pe2) + | PEsub pe1 pe2 => (eval_pexpr l pe1) - (eval_pexpr l pe2) + | PEmul pe1 pe2 => (eval_pexpr l pe1) * (eval_pexpr l pe2) + | PEopp pe1 => - (eval_pexpr l pe1) + | PEpow pe1 n => rpow (eval_pexpr l pe1) (pow_phi n) + end. +Proof. + intros ; destruct pe ; reflexivity. +Qed. + + + +Lemma eval_pexpr_PEeval : forall (env : PolEnv) (pe : PExprC), + eval_pexpr env pe = + PEeval rplus rtimes rminus ropp phi pow_phi rpow env pe. +Proof. +induction pe; simpl; intros. +reflexivity. +reflexivity. +rewrite <- IHpe1; rewrite <- IHpe2; reflexivity. +rewrite <- IHpe1; rewrite <- IHpe2; reflexivity. +rewrite <- IHpe1; rewrite <- IHpe2; reflexivity. +rewrite <- IHpe; reflexivity. +rewrite <- IHpe; reflexivity. +Qed. +(*****) +(*Definition eval_pexpr : Env -> PExprC -> R := + PEeval 0 rplus rtimes rminus ropp phi pow_phi rpow.*) +(*****) + +Inductive Op1 : Set := (* relations with 0 *) +| Equal (* == 0 *) +| NonEqual (* ~= 0 *) +| Strict (* > 0 *) +| NonStrict (* >= 0 *). + +Definition NFormula := (PExprC * Op1)%type. (* normalized formula *) + +Definition eval_op1 (o : Op1) : R -> Prop := +match o with +| Equal => fun x => x == 0 +| NonEqual => fun x : R => x ~= 0 +| Strict => fun x : R => 0 < x +| NonStrict => fun x : R => 0 <= x +end. + +Definition eval_nformula (env : PolEnv) (f : NFormula) : Prop := +let (p, op) := f in eval_op1 op (eval_pexpr env p). + + +Definition OpMult (o o' : Op1) : Op1 := +match o with +| Equal => Equal +| NonStrict => NonStrict (* (OpMult NonStrict Equal) could be defined as Equal *) +| Strict => o' +| NonEqual => NonEqual (* does not matter what we return here; see the following lemmas *) +end. + +Definition OpAdd (o o': Op1) : Op1 := +match o with +| Equal => o' +| NonStrict => + match o' with + | Strict => Strict + | _ => NonStrict + end +| Strict => Strict +| NonEqual => NonEqual (* does not matter what we return here *) +end. + +Lemma OpMultNonEqual : + forall o o' : Op1, o <> NonEqual -> o' <> NonEqual -> OpMult o o' <> NonEqual. +Proof. +intros o o' H1 H2; destruct o; destruct o'; simpl; try discriminate; +try (intro H; apply H1; reflexivity); +try (intro H; apply H2; reflexivity). +Qed. + +Lemma OpAdd_NonEqual : + forall o o' : Op1, o <> NonEqual -> o' <> NonEqual -> OpAdd o o' <> NonEqual. +Proof. +intros o o' H1 H2; destruct o; destruct o'; simpl; try discriminate; +try (intro H; apply H1; reflexivity); +try (intro H; apply H2; reflexivity). +Qed. + +Lemma OpMult_sound : + forall (o o' : Op1) (x y : R), o <> NonEqual -> o' <> NonEqual -> + eval_op1 o x -> eval_op1 o' y -> eval_op1 (OpMult o o') (x * y). +Proof. +unfold eval_op1; destruct o; simpl; intros o' x y H1 H2 H3 H4. +rewrite H3; now rewrite (Rtimes_0_l sor). +elimtype False; now apply H1. +destruct o'. +rewrite H4; now rewrite (Rtimes_0_r sor). +elimtype False; now apply H2. +now apply (Rtimes_pos_pos sor). +apply (Rtimes_nonneg_nonneg sor); [le_less | assumption]. +destruct o'. +rewrite H4, (Rtimes_0_r sor); le_equal. +elimtype False; now apply H2. +apply (Rtimes_nonneg_nonneg sor); [assumption | le_less]. +now apply (Rtimes_nonneg_nonneg sor). +Qed. + +Lemma OpAdd_sound : + forall (o o' : Op1) (e e' : R), o <> NonEqual -> o' <> NonEqual -> + eval_op1 o e -> eval_op1 o' e' -> eval_op1 (OpAdd o o') (e + e'). +Proof. +unfold eval_op1; destruct o; simpl; intros o' e e' H1 H2 H3 H4. +destruct o'. +now rewrite H3, H4, (Rplus_0_l sor). +elimtype False; now apply H2. +now rewrite H3, (Rplus_0_l sor). +now rewrite H3, (Rplus_0_l sor). +elimtype False; now apply H1. +destruct o'. +now rewrite H4, (Rplus_0_r sor). +elimtype False; now apply H2. +now apply (Rplus_pos_pos sor). +now apply (Rplus_pos_nonneg sor). +destruct o'. +now rewrite H4, (Rplus_0_r sor). +elimtype False; now apply H2. +now apply (Rplus_nonneg_pos sor). +now apply (Rplus_nonneg_nonneg sor). +Qed. + +(* We consider a monoid whose generators are polynomials from the +hypotheses of the form (p ~= 0). Thus it follows from the hypotheses that +every element of the monoid (i.e., arbitrary product of generators) is ~= +0. Therefore, the square of every element is > 0. *) + +Inductive Monoid (l : list NFormula) : PExprC -> Prop := +| M_One : Monoid l (PEc cI) +| M_In : forall p : PExprC, In (p, NonEqual) l -> Monoid l p +| M_Mult : forall (e1 e2 : PExprC), Monoid l e1 -> Monoid l e2 -> Monoid l (PEmul e1 e2). + +(* Do we really need to rely on the intermediate definition of monoid ? + InC why the restriction NonEqual ? + Could not we consider the IsIdeal as a IsMult ? + The same for IsSquare ? +*) + +Inductive Cone (l : list (NFormula)) : PExprC -> Op1 -> Prop := +| InC : forall p op, In (p, op) l -> op <> NonEqual -> Cone l p op +| IsIdeal : forall p, Cone l p Equal -> forall p', Cone l (PEmul p p') Equal +| IsSquare : forall p, Cone l (PEmul p p) NonStrict +| IsMonoid : forall p, Monoid l p -> Cone l (PEmul p p) Strict +| IsMult : forall p op q oq, Cone l p op -> Cone l q oq -> Cone l (PEmul p q) (OpMult op oq) +| IsAdd : forall p op q oq, Cone l p op -> Cone l q oq -> Cone l (PEadd p q) (OpAdd op oq) +| IsPos : forall c : C, cltb cO c = true -> Cone l (PEc c) Strict +| IsZ : Cone l (PEc cO) Equal. + +(* As promised, if all hypotheses are true in some environment, then every +member of the monoid is nonzero in this environment *) + +Lemma monoid_nonzero : forall (l : list NFormula) (env : PolEnv), + (forall f : NFormula, In f l -> eval_nformula env f) -> + forall p : PExprC, Monoid l p -> eval_pexpr env p ~= 0. +Proof. +intros l env H1 p H2. induction H2 as [| f H | e1 e2 H3 IH1 H4 IH2]; simpl. +rewrite addon.(SORrm).(morph1). apply (Rneq_symm sor). apply (Rneq_0_1 sor). +apply H1 in H. now simpl in H. +simpl in IH1, IH2. apply (Rtimes_neq_0 sor). now split. +Qed. + +(* If all members of a cone base are true in some environment, then every +member of the cone is true as well *) + +Lemma cone_true : + forall (l : list NFormula) (env : PolEnv), + (forall (f : NFormula), In f l -> eval_nformula env f) -> + forall (p : PExprC) (op : Op1), Cone l p op -> + op <> NonEqual /\ eval_nformula env (p, op). +Proof. +intros l env H1 p op H2. induction H2; simpl in *. +split. assumption. apply H1 in H. now unfold eval_nformula in H. +split. discriminate. destruct IHCone as [_ H3]. rewrite H3. now rewrite (Rtimes_0_l sor). +split. discriminate. apply (Rtimes_square_nonneg sor). +split. discriminate. apply <- (Rlt_le_neq sor). split. apply (Rtimes_square_nonneg sor). +apply (Rneq_symm sor). apply (Rtimes_neq_0 sor). split; now apply monoid_nonzero with l. +destruct IHCone1 as [IH1 IH2]; destruct IHCone2 as [IH3 IH4]. +split. now apply OpMultNonEqual. now apply OpMult_sound. +destruct IHCone1 as [IH1 IH2]; destruct IHCone2 as [IH3 IH4]. +split. now apply OpAdd_NonEqual. now apply OpAdd_sound. +split. discriminate. rewrite <- addon.(SORrm).(morph0). now apply cltb_sound. +split. discriminate. apply addon.(SORrm).(morph0). +Qed. + +(* Every element of a monoid is a product of some generators; therefore, +to determine an element we can give a list of generators' indices *) + +Definition MonoidMember : Set := list nat. + +Inductive ConeMember : Type := +| S_In : nat -> ConeMember +| S_Ideal : PExprC -> ConeMember -> ConeMember +| S_Square : PExprC -> ConeMember +| S_Monoid : MonoidMember -> ConeMember +| S_Mult : ConeMember -> ConeMember -> ConeMember +| S_Add : ConeMember -> ConeMember -> ConeMember +| S_Pos : C -> ConeMember +| S_Z : ConeMember. + +Definition nformula_times (f f' : NFormula) : NFormula := +let (p, op) := f in + let (p', op') := f' in + (PEmul p p', OpMult op op'). + +Definition nformula_plus (f f' : NFormula) : NFormula := +let (p, op) := f in + let (p', op') := f' in + (PEadd p p', OpAdd op op'). + +Definition nformula_times_0 (p : PExprC) (f : NFormula) : NFormula := +let (q, op) := f in + match op with + | Equal => (PEmul q p, Equal) + | _ => f + end. + +Fixpoint eval_monoid (l : list NFormula) (ns : MonoidMember) {struct ns} : PExprC := +match ns with +| nil => PEc cI +| n :: ns => + let p := match nth n l (PEc cI, NonEqual) with + | (q, NonEqual) => q + | _ => PEc cI + end in + PEmul p (eval_monoid l ns) +end. + +Theorem eval_monoid_in_monoid : + forall (l : list NFormula) (ns : MonoidMember), Monoid l (eval_monoid l ns). +Proof. +intro l; induction ns; simpl in *. +constructor. +apply M_Mult; [| assumption]. +destruct (nth_in_or_default a l (PEc cI, NonEqual)). +destruct (nth a l (PEc cI, NonEqual)). destruct o; try constructor. assumption. +rewrite e; simpl. constructor. +Qed. + +(* Provides the cone member from the witness, i.e., ConeMember *) +Fixpoint eval_cone (l : list NFormula) (cm : ConeMember) {struct cm} : NFormula := +match cm with +| S_In n => match nth n l (PEc cO, Equal) with + | (_, NonEqual) => (PEc cO, Equal) + | f => f + end +| S_Ideal p cm' => nformula_times_0 p (eval_cone l cm') +| S_Square p => (PEmul p p, NonStrict) +| S_Monoid m => let p := eval_monoid l m in (PEmul p p, Strict) +| S_Mult p q => nformula_times (eval_cone l p) (eval_cone l q) +| S_Add p q => nformula_plus (eval_cone l p) (eval_cone l q) +| S_Pos c => if cltb cO c then (PEc c, Strict) else (PEc cO, Equal) +| S_Z => (PEc cO, Equal) +end. + +Theorem eval_cone_in_cone : + forall (l : list NFormula) (cm : ConeMember), + let (p, op) := eval_cone l cm in Cone l p op. +Proof. +intros l cm; induction cm; simpl. +destruct (nth_in_or_default n l (PEc cO, Equal)). +destruct (nth n l (PEc cO, Equal)). destruct o; try (now apply InC). apply IsZ. +rewrite e. apply IsZ. +destruct (eval_cone l cm). destruct o; simpl; try assumption. now apply IsIdeal. +apply IsSquare. +apply IsMonoid. apply eval_monoid_in_monoid. +destruct (eval_cone l cm1). destruct (eval_cone l cm2). unfold nformula_times. now apply IsMult. +destruct (eval_cone l cm1). destruct (eval_cone l cm2). unfold nformula_plus. now apply IsAdd. +case_eq (cO [<] c) ; intros ; [apply IsPos ; auto| apply IsZ]. +apply IsZ. +Qed. + +(* (inconsistent_cone_member l p) means (p, op) is in the cone for some op +(> 0, >= 0, == 0, or ~= 0) and this formula is inconsistent. This fact +implies that l is inconsistent, as shown by the next lemma. Inconsistency +of a formula (p, op) can be established by normalizing p and showing that +it is a constant c for which (c, op) is false. (This is only a sufficient, +not necessary, condition, of course.) Membership in the cone can be +verified if we have a certificate. *) + +Definition inconsistent_cone_member (l : list NFormula) (p : PExprC) := + exists op : Op1, Cone l p op /\ + forall env : PolEnv, ~ eval_op1 op (eval_pexpr env p). + +(* If some element of a cone is inconsistent, then the base of the cone +is also inconsistent *) + +Lemma prove_inconsistent : + forall (l : list NFormula) (p : PExprC), + inconsistent_cone_member l p -> forall env, make_impl (eval_nformula env) l False. +Proof. +intros l p H env. +destruct H as [o [wit H]]. +apply -> make_conj_impl. +intro H1. apply H with env. +pose proof (@cone_true l env) as H2. +cut (forall f : NFormula, In f l -> eval_nformula env f). intro H3. +apply (proj2 (H2 H3 p o wit)). intro. now apply make_conj_in. +Qed. + +Definition normalise_pexpr : PExprC -> PolC := + norm_aux cO cI cplus ctimes cminus copp ceqb. + +(* The following definition we don't really need, hence it is commented *) +(*Definition eval_pol : PolEnv -> PolC -> R := Pphi 0 rplus rtimes phi.*) + +(* roughly speaking, normalise_pexpr_correct is a proof of + forall env p, eval_pexpr env p == eval_pol env (normalise_pexpr p) *) + +(*****) +Definition normalise_pexpr_correct := +let Rops_wd := mk_reqe rplus rtimes ropp req + sor.(SORplus_wd) + sor.(SORtimes_wd) + sor.(SORopp_wd) in + norm_aux_spec sor.(SORsetoid) Rops_wd (Rth_ARth (SORsetoid sor) Rops_wd sor.(SORrt)) + addon.(SORrm) addon.(SORpower). +(*****) +(*Definition normalise_pexpr_correct := +let Rops_wd := mk_reqe rplus rtimes ropp req + sor.(SORplus_wd) + sor.(SORtimes_wd) + sor.(SORopp_wd) in + norm_aux_spec sor.(SORsetoid) Rops_wd (Rth_ARth sor.(SORsetoid) Rops_wd sor.(SORrt)) + addon.(SORrm) addon.(SORpower) nil.*) +(*****) + +(* Check that a formula f is inconsistent by normalizing and comparing the +resulting constant with 0 *) + +Definition check_inconsistent (f : NFormula) : bool := +let (e, op) := f in + match normalise_pexpr e with + | Pc c => + match op with + | Equal => cneqb c cO + | NonStrict => c [<] cO + | Strict => c [<=] cO + | NonEqual => false (* eval_cone never returns (p, NonEqual) *) + end + | _ => false (* not a constant *) + end. + +Lemma check_inconsistent_sound : + forall (p : PExprC) (op : Op1), + check_inconsistent (p, op) = true -> forall env, ~ eval_op1 op (eval_pexpr env p). +Proof. +intros p op H1 env. unfold check_inconsistent, normalise_pexpr in H1. +destruct op; simpl; +(*****) +rewrite eval_pexpr_PEeval; +(*****) +(*unfold eval_pexpr;*) +(*****) +rewrite normalise_pexpr_correct; +destruct (norm_aux cO cI cplus ctimes cminus copp ceqb p); simpl; try discriminate H1; +try rewrite <- addon.(SORrm).(morph0); trivial. +now apply cneqb_sound. +apply cleb_sound in H1. now apply -> (Rle_ngt sor). +apply cltb_sound in H1. now apply -> (Rlt_nge sor). +Qed. + +Definition check_normalised_formulas : list NFormula -> ConeMember -> bool := + fun l cm => check_inconsistent (eval_cone l cm). + +Lemma checker_nf_sound : + forall (l : list NFormula) (cm : ConeMember), + check_normalised_formulas l cm = true -> + forall env : PolEnv, make_impl (eval_nformula env) l False. +Proof. +intros l cm H env. +unfold check_normalised_formulas in H. +case_eq (eval_cone l cm). intros p op H1. +apply prove_inconsistent with p. unfold inconsistent_cone_member. exists op. split. +pose proof (eval_cone_in_cone l cm) as H2. now rewrite H1 in H2. +apply check_inconsistent_sound. now rewrite <- H1. +Qed. + +(** Normalisation of formulae **) + +Inductive Op2 : Set := (* binary relations *) +| OpEq +| OpNEq +| OpLe +| OpGe +| OpLt +| OpGt. + +Definition eval_op2 (o : Op2) : R -> R -> Prop := +match o with +| OpEq => req +| OpNEq => fun x y : R => x ~= y +| OpLe => rle +| OpGe => fun x y : R => y <= x +| OpLt => fun x y : R => x < y +| OpGt => fun x y : R => y < x +end. + +Record Formula : Type := { + Flhs : PExprC; + Fop : Op2; + Frhs : PExprC +}. + +Definition eval_formula (env : PolEnv) (f : Formula) : Prop := + let (lhs, op, rhs) := f in + (eval_op2 op) (eval_pexpr env lhs) (eval_pexpr env rhs). + +(* We normalize Formulas by moving terms to one side *) + +Definition normalise (f : Formula) : NFormula := +let (lhs, op, rhs) := f in + match op with + | OpEq => (PEsub lhs rhs, Equal) + | OpNEq => (PEsub lhs rhs, NonEqual) + | OpLe => (PEsub rhs lhs, NonStrict) + | OpGe => (PEsub lhs rhs, NonStrict) + | OpGt => (PEsub lhs rhs, Strict) + | OpLt => (PEsub rhs lhs, Strict) + end. + +Definition negate (f : Formula) : NFormula := +let (lhs, op, rhs) := f in + match op with + | OpEq => (PEsub rhs lhs, NonEqual) + | OpNEq => (PEsub rhs lhs, Equal) + | OpLe => (PEsub lhs rhs, Strict) (* e <= e' == ~ e > e' *) + | OpGe => (PEsub rhs lhs, Strict) + | OpGt => (PEsub rhs lhs, NonStrict) + | OpLt => (PEsub lhs rhs, NonStrict) +end. + +Theorem normalise_sound : + forall (env : PolEnv) (f : Formula), + eval_formula env f -> eval_nformula env (normalise f). +Proof. +intros env f H; destruct f as [lhs op rhs]; simpl in *. +destruct op; simpl in *. +now apply <- (Rminus_eq_0 sor). +intros H1. apply -> (Rminus_eq_0 sor) in H1. now apply H. +now apply -> (Rle_le_minus sor). +now apply -> (Rle_le_minus sor). +now apply -> (Rlt_lt_minus sor). +now apply -> (Rlt_lt_minus sor). +Qed. + +Theorem negate_correct : + forall (env : PolEnv) (f : Formula), + eval_formula env f <-> ~ (eval_nformula env (negate f)). +Proof. +intros env f; destruct f as [lhs op rhs]; simpl. +destruct op; simpl. +symmetry. rewrite (Rminus_eq_0 sor). +split; intro H; [symmetry; now apply -> (Req_dne sor) | symmetry in H; now apply <- (Req_dne sor)]. +rewrite (Rminus_eq_0 sor). split; intro; now apply (Rneq_symm sor). +rewrite <- (Rlt_lt_minus sor). now rewrite <- (Rle_ngt sor). +rewrite <- (Rlt_lt_minus sor). now rewrite <- (Rle_ngt sor). +rewrite <- (Rle_le_minus sor). now rewrite <- (Rlt_nge sor). +rewrite <- (Rle_le_minus sor). now rewrite <- (Rlt_nge sor). +Qed. + +(** Another normalistion - this is used for cnf conversion **) + +Definition xnormalise (t:Formula) : list (NFormula) := + let (lhs,o,rhs) := t in + match o with + | OpEq => + (PEsub lhs rhs, Strict)::(PEsub rhs lhs , Strict)::nil + | OpNEq => (PEsub lhs rhs,Equal) :: nil + | OpGt => (PEsub rhs lhs,NonStrict) :: nil + | OpLt => (PEsub lhs rhs,NonStrict) :: nil + | OpGe => (PEsub rhs lhs , Strict) :: nil + | OpLe => (PEsub lhs rhs ,Strict) :: nil + end. + +Require Import Tauto. + +Definition cnf_normalise (t:Formula) : cnf (NFormula) := + List.map (fun x => x::nil) (xnormalise t). + + +Add Ring SORRing : sor.(SORrt). + +Lemma cnf_normalise_correct : forall env t, eval_cnf (eval_nformula env) (cnf_normalise t) -> eval_formula env t. +Proof. + unfold cnf_normalise, xnormalise ; simpl ; intros env t. + unfold eval_cnf. + destruct t as [lhs o rhs]; case_eq o ; simpl; + generalize (eval_pexpr env lhs); + generalize (eval_pexpr env rhs) ; intros z1 z2 ; intros. + (**) + apply sor.(SORle_antisymm). + rewrite (Rle_ngt sor). rewrite (Rlt_lt_minus sor). tauto. + rewrite (Rle_ngt sor). rewrite (Rlt_lt_minus sor). tauto. + now rewrite <- (Rminus_eq_0 sor). + rewrite (Rle_ngt sor). rewrite (Rlt_lt_minus sor). auto. + rewrite (Rle_ngt sor). rewrite (Rlt_lt_minus sor). auto. + rewrite (Rlt_nge sor). rewrite (Rle_le_minus sor). auto. + rewrite (Rlt_nge sor). rewrite (Rle_le_minus sor). auto. +Qed. + +Definition xnegate (t:Formula) : list (NFormula) := + let (lhs,o,rhs) := t in + match o with + | OpEq => (PEsub lhs rhs,Equal) :: nil + | OpNEq => (PEsub lhs rhs ,Strict)::(PEsub rhs lhs,Strict)::nil + | OpGt => (PEsub lhs rhs,Strict) :: nil + | OpLt => (PEsub rhs lhs,Strict) :: nil + | OpGe => (PEsub lhs rhs,NonStrict) :: nil + | OpLe => (PEsub rhs lhs,NonStrict) :: nil + end. + +Definition cnf_negate (t:Formula) : cnf (NFormula) := + List.map (fun x => x::nil) (xnegate t). + +Lemma cnf_negate_correct : forall env t, eval_cnf (eval_nformula env) (cnf_negate t) -> ~ eval_formula env t. +Proof. + unfold cnf_negate, xnegate ; simpl ; intros env t. + unfold eval_cnf. + destruct t as [lhs o rhs]; case_eq o ; simpl ; + generalize (eval_pexpr env lhs); + generalize (eval_pexpr env rhs) ; intros z1 z2 ; intros ; + intuition. + (**) + apply H0. + rewrite H1 ; ring. + (**) + apply H1. + apply sor.(SORle_antisymm). + rewrite (Rle_ngt sor). rewrite (Rlt_lt_minus sor). tauto. + rewrite (Rle_ngt sor). rewrite (Rlt_lt_minus sor). tauto. + (**) + apply H0. now rewrite (Rle_le_minus sor) in H1. + apply H0. now rewrite (Rle_le_minus sor) in H1. + apply H0. now rewrite (Rlt_lt_minus sor) in H1. + apply H0. now rewrite (Rlt_lt_minus sor) in H1. +Qed. + + +Lemma eval_nformula_dec : forall env d, (eval_nformula env d) \/ ~ (eval_nformula env d). +Proof. + intros. + destruct d ; simpl. + generalize (eval_pexpr env p); intros. + destruct o ; simpl. + apply (Req_em sor r 0). + destruct (Req_em sor r 0) ; tauto. + rewrite <- (Rle_ngt sor r 0). generalize (Rle_gt_cases sor r 0). tauto. + rewrite <- (Rlt_nge sor r 0). generalize (Rle_gt_cases sor 0 r). tauto. +Qed. + +(** Some syntactic simplifications of expressions and cone elements *) + + +Fixpoint simpl_expr (e:PExprC) : PExprC := + match e with + | PEmul y z => let y' := simpl_expr y in let z' := simpl_expr z in + match y' , z' with + | PEc c , z' => if ceqb c cI then z' else PEmul y' z' + | _ , _ => PEmul y' z' + end + | PEadd x y => PEadd (simpl_expr x) (simpl_expr y) + | _ => e + end. + + +Definition simpl_cone (e:ConeMember) : ConeMember := + match e with + | S_Square t => match simpl_expr t with + | PEc c => if ceqb cO c then S_Z else S_Pos (ctimes c c) + | x => S_Square x + end + | S_Mult t1 t2 => + match t1 , t2 with + | S_Z , x => S_Z + | x , S_Z => S_Z + | S_Pos c , S_Pos c' => S_Pos (ctimes c c') + | S_Pos p1 , S_Mult (S_Pos p2) x => S_Mult (S_Pos (ctimes p1 p2)) x + | S_Pos p1 , S_Mult x (S_Pos p2) => S_Mult (S_Pos (ctimes p1 p2)) x + | S_Mult (S_Pos p2) x , S_Pos p1 => S_Mult (S_Pos (ctimes p1 p2)) x + | S_Mult x (S_Pos p2) , S_Pos p1 => S_Mult (S_Pos (ctimes p1 p2)) x + | S_Pos x , S_Add y z => S_Add (S_Mult (S_Pos x) y) (S_Mult (S_Pos x) z) + | S_Pos c , _ => if ceqb cI c then t2 else S_Mult t1 t2 + | _ , S_Pos c => if ceqb cI c then t1 else S_Mult t1 t2 + | _ , _ => e + end + | S_Add t1 t2 => + match t1 , t2 with + | S_Z , x => x + | x , S_Z => x + | x , y => S_Add x y + end + | _ => e + end. + + + +End Micromega. + diff --git a/contrib/micromega/Tauto.v b/contrib/micromega/Tauto.v new file mode 100644 index 00000000..ef48efa6 --- /dev/null +++ b/contrib/micromega/Tauto.v @@ -0,0 +1,324 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* *) +(* Micromega: A reflexive tactic using the Positivstellensatz *) +(* *) +(* Frédéric Besson (Irisa/Inria) 2006-2008 *) +(* *) +(************************************************************************) + +Require Import List. +Require Import Refl. +Require Import Bool. + +Set Implicit Arguments. + + + Inductive BFormula (A:Type) : Type := + | TT : BFormula A + | FF : BFormula A + | X : Prop -> BFormula A + | A : A -> BFormula A + | Cj : BFormula A -> BFormula A -> BFormula A + | D : BFormula A-> BFormula A -> BFormula A + | N : BFormula A -> BFormula A + | I : BFormula A-> BFormula A-> BFormula A. + + Fixpoint eval_f (A:Type) (ev:A -> Prop ) (f:BFormula A) {struct f}: Prop := + match f with + | TT => True + | FF => False + | A a => ev a + | X p => p + | Cj e1 e2 => (eval_f ev e1) /\ (eval_f ev e2) + | D e1 e2 => (eval_f ev e1) \/ (eval_f ev e2) + | N e => ~ (eval_f ev e) + | I f1 f2 => (eval_f ev f1) -> (eval_f ev f2) + end. + + + Lemma map_simpl : forall A B f l, @map A B f l = match l with + | nil => nil + | a :: l=> (f a) :: (@map A B f l) + end. + Proof. + destruct l ; reflexivity. + Qed. + + + + Section S. + + Variable Env : Type. + Variable Term : Type. + Variable eval : Env -> Term -> Prop. + Variable Term' : Type. + Variable eval' : Env -> Term' -> Prop. + + + + Variable no_middle_eval' : forall env d, (eval' env d) \/ ~ (eval' env d). + + + Definition clause := list Term'. + Definition cnf := list clause. + + Variable normalise : Term -> cnf. + Variable negate : Term -> cnf. + + + Definition tt : cnf := @nil clause. + Definition ff : cnf := cons (@nil Term') nil. + + + Definition or_clause_cnf (t:clause) (f:cnf) : cnf := + List.map (fun x => (t++x)) f. + + Fixpoint or_cnf (f : cnf) (f' : cnf) {struct f}: cnf := + match f with + | nil => tt + | e :: rst => (or_cnf rst f') ++ (or_clause_cnf e f') + end. + + + Definition and_cnf (f1 : cnf) (f2 : cnf) : cnf := + f1 ++ f2. + + Fixpoint xcnf (pol : bool) (f : BFormula Term) {struct f}: cnf := + match f with + | TT => if pol then tt else ff + | FF => if pol then ff else tt + | X p => if pol then ff else ff (* This is not complete - cannot negate any proposition *) + | A x => if pol then normalise x else negate x + | N e => xcnf (negb pol) e + | Cj e1 e2 => + (if pol then and_cnf else or_cnf) (xcnf pol e1) (xcnf pol e2) + | D e1 e2 => (if pol then or_cnf else and_cnf) (xcnf pol e1) (xcnf pol e2) + | I e1 e2 => (if pol then or_cnf else and_cnf) (xcnf (negb pol) e1) (xcnf pol e2) + end. + + Definition eval_cnf (env : Term' -> Prop) (f:cnf) := make_conj (fun cl => ~ make_conj env cl) f. + + + Lemma eval_cnf_app : forall env x y, eval_cnf (eval' env) (x++y) -> eval_cnf (eval' env) x /\ eval_cnf (eval' env) y. + Proof. + unfold eval_cnf. + intros. + rewrite make_conj_app in H ; auto. + Qed. + + + Lemma or_clause_correct : forall env t f, eval_cnf (eval' env) (or_clause_cnf t f) -> (~ make_conj (eval' env) t) \/ (eval_cnf (eval' env) f). + Proof. + unfold eval_cnf. + unfold or_clause_cnf. + induction f. + simpl. + intros ; right;auto. + (**) + rewrite map_simpl. + intros. + rewrite make_conj_cons in H. + destruct H as [HH1 HH2]. + generalize (IHf HH2) ; clear IHf ; intro. + destruct H. + left ; auto. + rewrite make_conj_cons. + destruct (not_make_conj_app _ _ _ (no_middle_eval' env) HH1). + tauto. + tauto. + Qed. + + Lemma eval_cnf_cons : forall env a f, (~ make_conj (eval' env) a) -> eval_cnf (eval' env) f -> eval_cnf (eval' env) (a::f). + Proof. + intros. + unfold eval_cnf in *. + rewrite make_conj_cons ; eauto. + Qed. + + Lemma or_cnf_correct : forall env f f', eval_cnf (eval' env) (or_cnf f f') -> (eval_cnf (eval' env) f) \/ (eval_cnf (eval' env) f'). + Proof. + induction f. + unfold eval_cnf. + simpl. + tauto. + (**) + intros. + simpl in H. + destruct (eval_cnf_app _ _ _ H). + clear H. + destruct (IHf _ H0). + destruct (or_clause_correct _ _ _ H1). + left. + apply eval_cnf_cons ; auto. + right ; auto. + right ; auto. + Qed. + + Variable normalise_correct : forall env t, eval_cnf (eval' env) (normalise t) -> eval env t. + + Variable negate_correct : forall env t, eval_cnf (eval' env) (negate t) -> ~ eval env t. + + + Lemma xcnf_correct : forall f pol env, eval_cnf (eval' env) (xcnf pol f) -> eval_f (eval env) (if pol then f else N f). + Proof. + induction f. + (* TT *) + unfold eval_cnf. + simpl. + destruct pol ; simpl ; auto. + (* FF *) + unfold eval_cnf. + destruct pol; simpl ; auto. + (* P *) + simpl. + destruct pol ; intros ;simpl. + unfold eval_cnf in H. + (* Here I have to drop the proposition *) + simpl in H. + tauto. + (* Here, I could store P in the clause *) + unfold eval_cnf in H;simpl in H. + tauto. + (* A *) + simpl. + destruct pol ; simpl. + intros. + apply normalise_correct ; auto. + (* A 2 *) + intros. + apply negate_correct ; auto. + auto. + (* Cj *) + destruct pol ; simpl. + (* pol = true *) + intros. + unfold and_cnf in H. + destruct (eval_cnf_app _ _ _ H). + clear H. + split. + apply (IHf1 _ _ H0). + apply (IHf2 _ _ H1). + (* pol = false *) + intros. + destruct (or_cnf_correct _ _ _ H). + generalize (IHf1 false env H0). + simpl. + tauto. + generalize (IHf2 false env H0). + simpl. + tauto. + (* D *) + simpl. + destruct pol. + (* pol = true *) + intros. + destruct (or_cnf_correct _ _ _ H). + generalize (IHf1 _ env H0). + simpl. + tauto. + generalize (IHf2 _ env H0). + simpl. + tauto. + (* pol = true *) + unfold and_cnf. + intros. + destruct (eval_cnf_app _ _ _ H). + clear H. + simpl. + generalize (IHf1 _ _ H0). + generalize (IHf2 _ _ H1). + simpl. + tauto. + (**) + simpl. + destruct pol ; simpl. + intros. + apply (IHf false) ; auto. + intros. + generalize (IHf _ _ H). + tauto. + (* I *) + simpl; intros. + destruct pol. + simpl. + intro. + destruct (or_cnf_correct _ _ _ H). + generalize (IHf1 _ _ H1). + simpl in *. + tauto. + generalize (IHf2 _ _ H1). + auto. + (* pol = false *) + unfold and_cnf in H. + simpl in H. + destruct (eval_cnf_app _ _ _ H). + generalize (IHf1 _ _ H0). + generalize (IHf2 _ _ H1). + simpl. + tauto. + Qed. + + + Variable Witness : Type. + Variable checker : list Term' -> Witness -> bool. + + Variable checker_sound : forall t w, checker t w = true -> forall env, make_impl (eval' env) t False. + + Fixpoint cnf_checker (f : cnf) (l : list Witness) {struct f}: bool := + match f with + | nil => true + | e::f => match l with + | nil => false + | c::l => match checker e c with + | true => cnf_checker f l + | _ => false + end + end + end. + + Lemma cnf_checker_sound : forall t w, cnf_checker t w = true -> forall env, eval_cnf (eval' env) t. + Proof. + unfold eval_cnf. + induction t. + (* bc *) + simpl. + auto. + (* ic *) + simpl. + destruct w. + intros ; discriminate. + case_eq (checker a w) ; intros ; try discriminate. + generalize (@checker_sound _ _ H env). + generalize (IHt _ H0 env) ; intros. + destruct t. + red ; intro. + rewrite <- make_conj_impl in H2. + tauto. + rewrite <- make_conj_impl in H2. + tauto. + Qed. + + + Definition tauto_checker (f:BFormula Term) (w:list Witness) : bool := + cnf_checker (xcnf true f) w. + + Lemma tauto_checker_sound : forall t w, tauto_checker t w = true -> forall env, eval_f (eval env) t. + Proof. + unfold tauto_checker. + intros. + change (eval_f (eval env) t) with (eval_f (eval env) (if true then t else TT Term)). + apply (xcnf_correct t true). + eapply cnf_checker_sound ; eauto. + Qed. + + + + +End S. + diff --git a/contrib/micromega/VarMap.v b/contrib/micromega/VarMap.v new file mode 100644 index 00000000..240c0fb7 --- /dev/null +++ b/contrib/micromega/VarMap.v @@ -0,0 +1,258 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* *) +(* Micromega: A reflexive tactic using the Positivstellensatz *) +(* *) +(* Frédéric Besson (Irisa/Inria) 2006-2008 *) +(* *) +(************************************************************************) + +Require Import ZArith. +Require Import Coq.Arith.Max. +Require Import List. +Set Implicit Arguments. + +(* I have addded a Leaf constructor to the varmap data structure (/contrib/ring/Quote.v) + -- this is harmless and spares a lot of Empty. + This means smaller proof-terms. + BTW, by dropping the polymorphism, I get small (yet noticeable) speed-up. +*) + +Section MakeVarMap. + Variable A : Type. + Variable default : A. + + Inductive t : Type := + | Empty : t + | Leaf : A -> t + | Node : t -> A -> t -> t . + + Fixpoint find (vm : t ) (p:positive) {struct vm} : A := + match vm with + | Empty => default + | Leaf i => i + | Node l e r => match p with + | xH => e + | xO p => find l p + | xI p => find r p + end + end. + + (* an off_map (a map with offset) offers the same functionalites as /contrib/setoid_ring/BinList.v - it is used in EnvRing.v *) +(* + Definition off_map := (option positive *t )%type. + + + + Definition jump (j:positive) (l:off_map ) := + let (o,m) := l in + match o with + | None => (Some j,m) + | Some j0 => (Some (j+j0)%positive,m) + end. + + Definition nth (n:positive) (l: off_map ) := + let (o,m) := l in + let idx := match o with + | None => n + | Some i => i + n + end%positive in + find idx m. + + + Definition hd (l:off_map) := nth xH l. + + + Definition tail (l:off_map ) := jump xH l. + + + Lemma psucc : forall p, (match p with + | xI y' => xO (Psucc y') + | xO y' => xI y' + | 1%positive => 2%positive + end) = (p+1)%positive. + Proof. + destruct p. + auto with zarith. + rewrite xI_succ_xO. + auto with zarith. + reflexivity. + Qed. + + Lemma jump_Pplus : forall i j l, + (jump (i + j) l) = (jump i (jump j l)). + Proof. + unfold jump. + destruct l. + destruct o. + rewrite Pplus_assoc. + reflexivity. + reflexivity. + Qed. + + Lemma jump_simpl : forall p l, + jump p l = + match p with + | xH => tail l + | xO p => jump p (jump p l) + | xI p => jump p (jump p (tail l)) + end. + Proof. + destruct p ; unfold tail ; intros ; repeat rewrite <- jump_Pplus. + (* xI p = p + p + 1 *) + rewrite xI_succ_xO. + rewrite Pplus_diag. + rewrite <- Pplus_one_succ_r. + reflexivity. + (* xO p = p + p *) + rewrite Pplus_diag. + reflexivity. + reflexivity. + Qed. + + Ltac jump_s := + repeat + match goal with + | |- context [jump xH ?e] => rewrite (jump_simpl xH) + | |- context [jump (xO ?p) ?e] => rewrite (jump_simpl (xO p)) + | |- context [jump (xI ?p) ?e] => rewrite (jump_simpl (xI p)) + end. + + Lemma jump_tl : forall j l, tail (jump j l) = jump j (tail l). + Proof. + unfold tail. + intros. + repeat rewrite <- jump_Pplus. + rewrite Pplus_comm. + reflexivity. + Qed. + + Lemma jump_Psucc : forall j l, + (jump (Psucc j) l) = (jump 1 (jump j l)). + Proof. + intros. + rewrite <- jump_Pplus. + rewrite Pplus_one_succ_r. + rewrite Pplus_comm. + reflexivity. + Qed. + + Lemma jump_Pdouble_minus_one : forall i l, + (jump (Pdouble_minus_one i) (tail l)) = (jump i (jump i l)). + Proof. + unfold tail. + intros. + repeat rewrite <- jump_Pplus. + rewrite <- Pplus_one_succ_r. + rewrite Psucc_o_double_minus_one_eq_xO. + rewrite Pplus_diag. + reflexivity. + Qed. + + Lemma jump_x0_tail : forall p l, jump (xO p) (tail l) = jump (xI p) l. + Proof. + intros. + jump_s. + repeat rewrite <- jump_Pplus. + reflexivity. + Qed. + + + Lemma nth_spec : forall p l, + nth p l = + match p with + | xH => hd l + | xO p => nth p (jump p l) + | xI p => nth p (jump p (tail l)) + end. + Proof. + unfold nth. + destruct l. + destruct o. + simpl. + rewrite psucc. + destruct p. + replace (p0 + xI p)%positive with ((p + (p0 + 1) + p))%positive. + reflexivity. + rewrite xI_succ_xO. + rewrite Pplus_one_succ_r. + rewrite <- Pplus_diag. + rewrite Pplus_comm. + symmetry. + rewrite (Pplus_comm p0). + rewrite <- Pplus_assoc. + rewrite (Pplus_comm 1)%positive. + rewrite <- Pplus_assoc. + reflexivity. + (**) + replace ((p0 + xO p))%positive with (p + p0 + p)%positive. + reflexivity. + rewrite <- Pplus_diag. + rewrite <- Pplus_assoc. + rewrite Pplus_comm. + rewrite Pplus_assoc. + reflexivity. + reflexivity. + simpl. + destruct p. + rewrite xI_succ_xO. + rewrite Pplus_one_succ_r. + rewrite <- Pplus_diag. + symmetry. + rewrite Pplus_comm. + rewrite Pplus_assoc. + reflexivity. + rewrite Pplus_diag. + reflexivity. + reflexivity. + Qed. + + + Lemma nth_jump : forall p l, nth p (tail l) = hd (jump p l). + Proof. + destruct l. + unfold tail. + unfold hd. + unfold jump. + unfold nth. + destruct o. + symmetry. + rewrite Pplus_comm. + rewrite <- Pplus_assoc. + rewrite (Pplus_comm p0). + reflexivity. + rewrite Pplus_comm. + reflexivity. + Qed. + + Lemma nth_Pdouble_minus_one : + forall p l, nth (Pdouble_minus_one p) (tail l) = nth p (jump p l). + Proof. + destruct l. + unfold tail. + unfold nth, jump. + destruct o. + rewrite ((Pplus_comm p)). + rewrite <- (Pplus_assoc p0). + rewrite Pplus_diag. + rewrite <- Psucc_o_double_minus_one_eq_xO. + rewrite Pplus_one_succ_r. + rewrite (Pplus_comm (Pdouble_minus_one p)). + rewrite Pplus_assoc. + rewrite (Pplus_comm p0). + reflexivity. + rewrite <- Pplus_one_succ_l. + rewrite Psucc_o_double_minus_one_eq_xO. + rewrite Pplus_diag. + reflexivity. + Qed. + +*) + +End MakeVarMap. + diff --git a/contrib/micromega/ZCoeff.v b/contrib/micromega/ZCoeff.v new file mode 100644 index 00000000..ced67e39 --- /dev/null +++ b/contrib/micromega/ZCoeff.v @@ -0,0 +1,173 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +Require Import OrderedRing. +Require Import RingMicromega. +Require Import ZArith. +Require Import InitialRing. +Require Import Setoid. + +Import OrderedRingSyntax. + +Set Implicit Arguments. + +Section InitialMorphism. + +Variable 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 sor : SOR rO rI rplus rtimes rminus ropp req rle rlt. + +Notation "0" := rO. +Notation "1" := rI. +Notation "x + y" := (rplus x y). +Notation "x * y " := (rtimes x y). +Notation "x - y " := (rminus x y). +Notation "- x" := (ropp x). +Notation "x == y" := (req x y). +Notation "x ~= y" := (~ req x y). +Notation "x <= y" := (rle x y). +Notation "x < y" := (rlt x y). + +Lemma req_refl : forall x, req x x. +Proof. + destruct sor.(SORsetoid). + apply Equivalence_Reflexive. +Qed. + +Lemma req_sym : forall x y, req x y -> req y x. +Proof. + destruct sor.(SORsetoid). + apply Equivalence_Symmetric. +Qed. + +Lemma req_trans : forall x y z, req x y -> req y z -> req x z. +Proof. + destruct sor.(SORsetoid). + apply Equivalence_Transitive. +Qed. + + +Add Relation R req + reflexivity proved by sor.(SORsetoid).(@Equivalence_Reflexive _ _) + symmetry proved by sor.(SORsetoid).(@Equivalence_Symmetric _ _) + transitivity proved by sor.(SORsetoid).(@Equivalence_Transitive _ _) +as sor_setoid. + +Add Morphism rplus with signature req ==> req ==> req as rplus_morph. +Proof. +exact sor.(SORplus_wd). +Qed. +Add Morphism rtimes with signature req ==> req ==> req as rtimes_morph. +Proof. +exact sor.(SORtimes_wd). +Qed. +Add Morphism ropp with signature req ==> req as ropp_morph. +Proof. +exact sor.(SORopp_wd). +Qed. +Add Morphism rle with signature req ==> req ==> iff as rle_morph. +Proof. +exact sor.(SORle_wd). +Qed. +Add Morphism rlt with signature req ==> req ==> iff as rlt_morph. +Proof. +exact sor.(SORlt_wd). +Qed. +Add Morphism rminus with signature req ==> req ==> req as rminus_morph. +Proof. + exact (rminus_morph sor). +Qed. + +Ltac le_less := rewrite (Rle_lt_eq sor); left; try assumption. +Ltac le_equal := rewrite (Rle_lt_eq sor); right; try reflexivity; try assumption. + +Definition gen_order_phi_Z : Z -> R := gen_phiZ 0 1 rplus rtimes ropp. + +Notation phi_pos := (gen_phiPOS 1 rplus rtimes). +Notation phi_pos1 := (gen_phiPOS1 1 rplus rtimes). + +Notation "[ x ]" := (gen_order_phi_Z x). + +Lemma ring_ops_wd : ring_eq_ext rplus rtimes ropp req. +Proof. +constructor. +exact rplus_morph. +exact rtimes_morph. +exact ropp_morph. +Qed. + +Lemma Zring_morph : + ring_morph 0 1 rplus rtimes rminus ropp req + 0%Z 1%Z Zplus Zmult Zminus Zopp + Zeq_bool gen_order_phi_Z. +Proof. +exact (gen_phiZ_morph sor.(SORsetoid) ring_ops_wd sor.(SORrt)). +Qed. + +Lemma phi_pos1_pos : forall x : positive, 0 < phi_pos1 x. +Proof. +induction x as [x IH | x IH |]; simpl; +try apply (Rplus_pos_pos sor); try apply (Rtimes_pos_pos sor); try apply (Rplus_pos_pos sor); +try apply (Rlt_0_1 sor); assumption. +Qed. + +Lemma phi_pos1_succ : forall x : positive, phi_pos1 (Psucc x) == 1 + phi_pos1 x. +Proof. +exact (ARgen_phiPOS_Psucc sor.(SORsetoid) ring_ops_wd + (Rth_ARth sor.(SORsetoid) ring_ops_wd sor.(SORrt))). +Qed. + +Lemma clt_pos_morph : forall x y : positive, (x < y)%positive -> phi_pos1 x < phi_pos1 y. +Proof. +intros x y H. pattern y; apply Plt_ind with x. +rewrite phi_pos1_succ; apply (Rlt_succ_r sor). +clear y H; intros y _ H. rewrite phi_pos1_succ. now apply (Rlt_lt_succ sor). +assumption. +Qed. + +Lemma clt_morph : forall x y : Z, (x < y)%Z -> [x] < [y]. +Proof. +unfold Zlt; intros x y H; +do 2 rewrite (same_genZ sor.(SORsetoid) ring_ops_wd sor.(SORrt)); +destruct x; destruct y; simpl in *; try discriminate. +apply phi_pos1_pos. +now apply clt_pos_morph. +apply <- (Ropp_neg_pos sor); apply phi_pos1_pos. +apply (Rlt_trans sor) with 0. apply <- (Ropp_neg_pos sor); apply phi_pos1_pos. +apply phi_pos1_pos. +rewrite Pcompare_antisym in H; simpl in H. apply -> (Ropp_lt_mono sor). +now apply clt_pos_morph. +Qed. + +Lemma Zcleb_morph : forall x y : Z, Zle_bool x y = true -> [x] <= [y]. +Proof. +unfold Zle_bool; intros x y H. +case_eq (x ?= y)%Z; intro H1; rewrite H1 in H. +le_equal. apply Zring_morph.(morph_eq). unfold Zeq_bool; now rewrite H1. +le_less. now apply clt_morph. +discriminate. +Qed. + +Lemma Zcneqb_morph : forall x y : Z, Zeq_bool x y = false -> [x] ~= [y]. +Proof. +intros x y H. unfold Zeq_bool in H. +case_eq (Zcompare x y); intro H1; rewrite H1 in *; (discriminate || clear H). +apply (Rlt_neq sor). now apply clt_morph. +fold (x > y)%Z in H1. rewrite Zgt_iff_lt in H1. +apply (Rneq_symm sor). apply (Rlt_neq sor). now apply clt_morph. +Qed. + +End InitialMorphism. + + diff --git a/contrib/micromega/ZMicromega.v b/contrib/micromega/ZMicromega.v new file mode 100644 index 00000000..94c83f73 --- /dev/null +++ b/contrib/micromega/ZMicromega.v @@ -0,0 +1,714 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* *) +(* Micromega: A reflexive tactic using the Positivstellensatz *) +(* *) +(* Frédéric Besson (Irisa/Inria) 2006-2008 *) +(* *) +(************************************************************************) + +Require Import OrderedRing. +Require Import RingMicromega. +Require Import ZCoeff. +Require Import Refl. +Require Import ZArith. +Require Import List. +Require Import Bool. + +Ltac flatten_bool := + repeat match goal with + [ id : (_ && _)%bool = true |- _ ] => destruct (andb_prop _ _ id); clear id + | [ id : (_ || _)%bool = true |- _ ] => destruct (orb_prop _ _ id); clear id + end. + +Require Import EnvRing. + +Open Scope Z_scope. + +Lemma Zsor : SOR 0 1 Zplus Zmult Zminus Zopp (@eq Z) Zle Zlt. +Proof. + constructor ; intros ; subst ; try (intuition (auto with zarith)). + apply Zsth. + apply Zth. + destruct (Ztrichotomy n m) ; intuition (auto with zarith). + apply Zmult_lt_0_compat ; auto. +Qed. + +Lemma Zeq_bool_neq : forall x y, Zeq_bool x y = false -> x <> y. +Proof. + red ; intros. + subst. + unfold Zeq_bool in H. + rewrite Zcompare_refl in H. + discriminate. +Qed. + +Lemma ZSORaddon : + SORaddon 0 1 Zplus Zmult Zminus Zopp (@eq Z) Zle (* ring elements *) + 0%Z 1%Z Zplus Zmult Zminus Zopp (* coefficients *) + Zeq_bool Zle_bool + (fun x => x) (fun x => x) (pow_N 1 Zmult). +Proof. + constructor. + constructor ; intros ; try reflexivity. + apply Zeqb_ok ; auto. + constructor. + reflexivity. + intros x y. + apply Zeq_bool_neq ; auto. + apply Zle_bool_imp_le. +Qed. + + +(*Definition Zeval_expr := eval_pexpr 0 Zplus Zmult Zminus Zopp (fun x => x) (fun x => Z_of_N x) (Zpower).*) + +Fixpoint Zeval_expr (env: PolEnv Z) (e: PExpr Z) : Z := + match e with + | PEc c => c + | PEX j => env j + | PEadd pe1 pe2 => (Zeval_expr env pe1) + (Zeval_expr env pe2) + | PEsub pe1 pe2 => (Zeval_expr env pe1) - (Zeval_expr env pe2) + | PEmul pe1 pe2 => (Zeval_expr env pe1) * (Zeval_expr env pe2) + | PEopp pe1 => - (Zeval_expr env pe1) + | PEpow pe1 n => Zpower (Zeval_expr env pe1) (Z_of_N n) + end. + +Lemma Zeval_expr_simpl : forall env e, + Zeval_expr env e = + match e with + | PEc c => c + | PEX j => env j + | PEadd pe1 pe2 => (Zeval_expr env pe1) + (Zeval_expr env pe2) + | PEsub pe1 pe2 => (Zeval_expr env pe1) - (Zeval_expr env pe2) + | PEmul pe1 pe2 => (Zeval_expr env pe1) * (Zeval_expr env pe2) + | PEopp pe1 => - (Zeval_expr env pe1) + | PEpow pe1 n => Zpower (Zeval_expr env pe1) (Z_of_N n) + end. +Proof. + destruct e ; reflexivity. +Qed. + + +Definition Zeval_expr' := eval_pexpr Zplus Zmult Zminus Zopp (fun x => x) (fun x => x) (pow_N 1 Zmult). + +Lemma ZNpower : forall r n, r ^ Z_of_N n = pow_N 1 Zmult r n. +Proof. + destruct n. + reflexivity. + simpl. + unfold Zpower_pos. + replace (pow_pos Zmult r p) with (1 * (pow_pos Zmult r p)) by ring. + generalize 1. + induction p; simpl ; intros ; repeat rewrite IHp ; ring. +Qed. + + + +Lemma Zeval_expr_compat : forall env e, Zeval_expr env e = Zeval_expr' env e. +Proof. + induction e ; simpl ; subst ; try congruence. + rewrite IHe. + apply ZNpower. +Qed. + +Definition Zeval_op2 (o : Op2) : Z -> Z -> Prop := +match o with +| OpEq => @eq Z +| OpNEq => fun x y => ~ x = y +| OpLe => Zle +| OpGe => Zge +| OpLt => Zlt +| OpGt => Zgt +end. + +Definition Zeval_formula (e: PolEnv Z) (ff : Formula Z) := + let (lhs,o,rhs) := ff in Zeval_op2 o (Zeval_expr e lhs) (Zeval_expr e rhs). + +Definition Zeval_formula' := + eval_formula Zplus Zmult Zminus Zopp (@eq Z) Zle Zlt (fun x => x) (fun x => x) (pow_N 1 Zmult). + +Lemma Zeval_formula_compat : forall env f, Zeval_formula env f <-> Zeval_formula' env f. +Proof. + intros. + unfold Zeval_formula. + destruct f. + repeat rewrite Zeval_expr_compat. + unfold Zeval_formula'. + unfold Zeval_expr'. + split ; destruct Fop ; simpl; auto with zarith. +Qed. + + + +Definition Zeval_nformula := + eval_nformula 0 Zplus Zmult Zminus Zopp (@eq Z) Zle Zlt (fun x => x) (fun x => x) (pow_N 1 Zmult). + +Definition Zeval_op1 (o : Op1) : Z -> Prop := +match o with +| Equal => fun x : Z => x = 0 +| NonEqual => fun x : Z => x <> 0 +| Strict => fun x : Z => 0 < x +| NonStrict => fun x : Z => 0 <= x +end. + +Lemma Zeval_nformula_simpl : forall env f, Zeval_nformula env f = (let (p, op) := f in Zeval_op1 op (Zeval_expr env p)). +Proof. + intros. + destruct f. + rewrite Zeval_expr_compat. + reflexivity. +Qed. + +Lemma Zeval_nformula_dec : forall env d, (Zeval_nformula env d) \/ ~ (Zeval_nformula env d). +Proof. + exact (fun env d =>eval_nformula_dec Zsor (fun x => x) (fun x => x) (pow_N 1%Z Zmult) env d). +Qed. + +Definition ZWitness := ConeMember Z. + +Definition ZWeakChecker := check_normalised_formulas 0 1 Zplus Zmult Zminus Zopp Zeq_bool Zle_bool. + +Lemma ZWeakChecker_sound : forall (l : list (NFormula Z)) (cm : ZWitness), + ZWeakChecker l cm = true -> + forall env, make_impl (Zeval_nformula env) l False. +Proof. + intros l cm H. + intro. + unfold Zeval_nformula. + apply (checker_nf_sound Zsor ZSORaddon l cm). + unfold ZWeakChecker in H. + exact H. +Qed. + +Definition xnormalise (t:Formula Z) : list (NFormula Z) := + let (lhs,o,rhs) := t in + match o with + | OpEq => + ((PEsub lhs (PEadd rhs (PEc 1))),NonStrict)::((PEsub rhs (PEadd lhs (PEc 1))),NonStrict)::nil + | OpNEq => (PEsub lhs rhs,Equal) :: nil + | OpGt => (PEsub rhs lhs,NonStrict) :: nil + | OpLt => (PEsub lhs rhs,NonStrict) :: nil + | OpGe => (PEsub rhs (PEadd lhs (PEc 1)),NonStrict) :: nil + | OpLe => (PEsub lhs (PEadd rhs (PEc 1)),NonStrict) :: nil + end. + +Require Import Tauto. + +Definition normalise (t:Formula Z) : cnf (NFormula Z) := + List.map (fun x => x::nil) (xnormalise t). + + +Lemma normalise_correct : forall env t, eval_cnf (Zeval_nformula env) (normalise t) <-> Zeval_formula env t. +Proof. + unfold normalise, xnormalise ; simpl ; intros env t. + rewrite Zeval_formula_compat. + unfold eval_cnf. + destruct t as [lhs o rhs]; case_eq o ; simpl; + generalize ( eval_pexpr Zplus Zmult Zminus Zopp (fun x : Z => x) + (fun x : BinNat.N => x) (pow_N 1 Zmult) env lhs); + generalize (eval_pexpr Zplus Zmult Zminus Zopp (fun x : Z => x) + (fun x : BinNat.N => x) (pow_N 1 Zmult) env rhs) ; intros z1 z2 ; intros ; subst; + intuition (auto with zarith). +Qed. + +Definition xnegate (t:RingMicromega.Formula Z) : list (NFormula Z) := + let (lhs,o,rhs) := t in + match o with + | OpEq => (PEsub lhs rhs,Equal) :: nil + | OpNEq => ((PEsub lhs (PEadd rhs (PEc 1))),NonStrict)::((PEsub rhs (PEadd lhs (PEc 1))),NonStrict)::nil + | OpGt => (PEsub lhs (PEadd rhs (PEc 1)),NonStrict) :: nil + | OpLt => (PEsub rhs (PEadd lhs (PEc 1)),NonStrict) :: nil + | OpGe => (PEsub lhs rhs,NonStrict) :: nil + | OpLe => (PEsub rhs lhs,NonStrict) :: nil + end. + +Definition negate (t:RingMicromega.Formula Z) : cnf (NFormula Z) := + List.map (fun x => x::nil) (xnegate t). + +Lemma negate_correct : forall env t, eval_cnf (Zeval_nformula env) (negate t) <-> ~ Zeval_formula env t. +Proof. + unfold negate, xnegate ; simpl ; intros env t. + rewrite Zeval_formula_compat. + unfold eval_cnf. + destruct t as [lhs o rhs]; case_eq o ; simpl ; + generalize ( eval_pexpr Zplus Zmult Zminus Zopp (fun x : Z => x) + (fun x : BinNat.N => x) (pow_N 1 Zmult) env lhs); + generalize (eval_pexpr Zplus Zmult Zminus Zopp (fun x : Z => x) + (fun x : BinNat.N => x) (pow_N 1 Zmult) env rhs) ; intros z1 z2 ; intros ; + intuition (auto with zarith). +Qed. + + +Definition ZweakTautoChecker (w: list ZWitness) (f : BFormula (Formula Z)) : bool := + @tauto_checker (Formula Z) (NFormula Z) normalise negate ZWitness ZWeakChecker f w. + +(* To get a complete checker, the proof format has to be enriched *) + +Require Import Zdiv. +Open Scope Z_scope. + +Definition ceiling (a b:Z) : Z := + let (q,r) := Zdiv_eucl a b in + match r with + | Z0 => q + | _ => q + 1 + end. + +Lemma narrow_interval_lower_bound : forall a b x, a > 0 -> a * x >= b -> x >= ceiling b a. +Proof. + unfold ceiling. + intros. + generalize (Z_div_mod b a H). + destruct (Zdiv_eucl b a). + intros. + destruct H1. + destruct H2. + subst. + destruct (Ztrichotomy z0 0) as [ HH1 | [HH2 | HH3]]; destruct z0 ; try auto with zarith ; try discriminate. + assert (HH :x >= z \/ x < z) by (destruct (Ztrichotomy x z) ; auto with zarith). + destruct HH ;auto. + generalize (Zmult_lt_compat_l _ _ _ H3 H1). + auto with zarith. + clear H2. + assert (HH :x >= z +1 \/ x <= z) by (destruct (Ztrichotomy x z) ; intuition (auto with zarith)). + destruct HH ;auto. + assert (0 < a) by auto with zarith. + generalize (Zmult_lt_0_le_compat_r _ _ _ H2 H1). + intros. + rewrite Zmult_comm in H4. + rewrite (Zmult_comm z) in H4. + auto with zarith. +Qed. + +Lemma narrow_interval_upper_bound : forall a b x, a > 0 -> a * x <= b -> x <= Zdiv b a. +Proof. + unfold Zdiv. + intros. + generalize (Z_div_mod b a H). + destruct (Zdiv_eucl b a). + intros. + destruct H1. + destruct H2. + subst. + assert (HH :x <= z \/ z <= x -1) by (destruct (Ztrichotomy x z) ; intuition (auto with zarith)). + destruct HH ;auto. + assert (0 < a) by auto with zarith. + generalize (Zmult_lt_0_le_compat_r _ _ _ H4 H1). + intros. + ring_simplify in H5. + rewrite Zmult_comm in H5. + auto with zarith. +Qed. + + +(* In this case, a certificate is made of a pair of inequations, in 1 variable, + that do not have an integer solution. + => modify the fourier elimination + *) +Require Import QArith. + + +Inductive ProofTerm : Type := +| RatProof : ZWitness -> ProofTerm +| CutProof : PExprC Z -> Q -> ZWitness -> ProofTerm -> ProofTerm +| EnumProof : Q -> PExprC Z -> Q -> ZWitness -> ZWitness -> list ProofTerm -> ProofTerm. + +(* n/d <= x -> d*x - n >= 0 *) + +Definition makeLb (v:PExpr Z) (q:Q) : NFormula Z := + let (n,d) := q in (PEsub (PEmul (PEc (Zpos d)) v) (PEc n),NonStrict). + +(* x <= n/d -> d * x <= d *) +Definition makeUb (v:PExpr Z) (q:Q) : NFormula Z := + let (n,d) := q in + (PEsub (PEc n) (PEmul (PEc (Zpos d)) v), NonStrict). + +Definition qceiling (q:Q) : Z := + let (n,d) := q in ceiling n (Zpos d). + +Definition qfloor (q:Q) : Z := + let (n,d) := q in Zdiv n (Zpos d). + +Definition makeLbCut (v:PExprC Z) (q:Q) : NFormula Z := + (PEsub v (PEc (qceiling q)), NonStrict). + +Definition neg_nformula (f : NFormula Z) := + let (e,o) := f in + (PEopp (PEadd e (PEc 1%Z)), o). + +Lemma neg_nformula_sound : forall env f, snd f = NonStrict ->( ~ (Zeval_nformula env (neg_nformula f)) <-> Zeval_nformula env f). +Proof. + unfold neg_nformula. + destruct f. + simpl. + intros ; subst ; simpl in *. + split; auto with zarith. +Qed. + + +Definition cutChecker (l:list (NFormula Z)) (e: PExpr Z) (lb:Q) (pf : ZWitness) : option (NFormula Z) := + let (lb,lc) := (makeLb e lb,makeLbCut e lb) in + if ZWeakChecker (neg_nformula lb::l) pf then Some lc else None. + + +Fixpoint ZChecker (l:list (NFormula Z)) (pf : ProofTerm) {struct pf} : bool := + match pf with + | RatProof pf => ZWeakChecker l pf + | CutProof e q pf rst => + match cutChecker l e q pf with + | None => false + | Some c => ZChecker (c::l) rst + end + | EnumProof lb e ub pf1 pf2 rst => + match cutChecker l e lb pf1 , cutChecker l (PEopp e) (Qopp ub) pf2 with + | None , _ | _ , None => false + | Some _ , Some _ => let (lb',ub') := (qceiling lb, Zopp (qceiling (- ub))) in + (fix label (pfs:list ProofTerm) := + fun lb ub => + match pfs with + | nil => if Z_gt_dec lb ub then true else false + | pf::rsr => andb (ZChecker ((PEsub e (PEc lb), Equal) :: l) pf) (label rsr (Zplus lb 1%Z) ub) + end) + rst lb' ub' + end + end. + + +Lemma ZChecker_simpl : forall (pf : ProofTerm) (l:list (NFormula Z)), + ZChecker l pf = + match pf with + | RatProof pf => ZWeakChecker l pf + | CutProof e q pf rst => + match cutChecker l e q pf with + | None => false + | Some c => ZChecker (c::l) rst + end + | EnumProof lb e ub pf1 pf2 rst => + match cutChecker l e lb pf1 , cutChecker l (PEopp e) (Qopp ub) pf2 with + | None , _ | _ , None => false + | Some _ , Some _ => let (lb',ub') := (qceiling lb, Zopp (qceiling (- ub))) in + (fix label (pfs:list ProofTerm) := + fun lb ub => + match pfs with + | nil => if Z_gt_dec lb ub then true else false + | pf::rsr => andb (ZChecker ((PEsub e (PEc lb), Equal) :: l) pf) (label rsr (Zplus lb 1%Z) ub) + end) + rst lb' ub' + end + end. +Proof. + destruct pf ; reflexivity. +Qed. + +(* +Fixpoint depth (n:nat) : ProofTerm -> option nat := + match n with + | O => fun pf => None + | S n => + fun pf => + match pf with + | RatProof _ => Some O + | CutProof _ _ _ p => option_map S (depth n p) + | EnumProof _ _ _ _ _ l => + let f := fun pf x => + match x , depth n pf with + | None , _ | _ , None => None + | Some n1 , Some n2 => Some (Max.max n1 n2) + end in + List.fold_right f (Some O) l + end + end. +*) +Fixpoint bdepth (pf : ProofTerm) : nat := + match pf with + | RatProof _ => O + | CutProof _ _ _ p => S (bdepth p) + | EnumProof _ _ _ _ _ l => S (List.fold_right (fun pf x => Max.max (bdepth pf) x) O l) + end. + +Require Import Wf_nat. + +Lemma in_bdepth : forall l a b p c c0 y, In y l -> ltof ProofTerm bdepth y (EnumProof a b p c c0 l). +Proof. + induction l. + simpl. + tauto. + simpl. + intros. + destruct H. + subst. + unfold ltof. + simpl. + generalize ( (fold_right + (fun (pf : ProofTerm) (x : nat) => Max.max (bdepth pf) x) 0%nat l)). + intros. + generalize (bdepth y) ; intros. + generalize (Max.max_l n0 n) (Max.max_r n0 n). + omega. + generalize (IHl a0 b p c c0 y H). + unfold ltof. + simpl. + generalize ( (fold_right (fun (pf : ProofTerm) (x : nat) => Max.max (bdepth pf) x) 0%nat + l)). + intros. + generalize (Max.max_l (bdepth a) n) (Max.max_r (bdepth a) n). + omega. +Qed. + +Lemma lb_lbcut : forall env e q, Zeval_nformula env (makeLb e q) -> Zeval_nformula env (makeLbCut e q). +Proof. + unfold makeLb, makeLbCut. + destruct q. + rewrite Zeval_nformula_simpl. + rewrite Zeval_nformula_simpl. + unfold Zeval_op1. + rewrite Zeval_expr_simpl. + rewrite Zeval_expr_simpl. + rewrite Zeval_expr_simpl. + intro. + rewrite Zeval_expr_simpl. + revert H. + generalize (Zeval_expr env e). + rewrite Zeval_expr_simpl. + rewrite Zeval_expr_simpl. + unfold qceiling. + intros. + assert ( z >= ceiling Qnum (' Qden))%Z. + apply narrow_interval_lower_bound. + compute. + reflexivity. + destruct z ; auto with zarith. + auto with zarith. +Qed. + +Lemma cutChecker_sound : forall e lb pf l res, cutChecker l e lb pf = Some res -> + forall env, make_impl (Zeval_nformula env) l (Zeval_nformula env res). +Proof. + unfold cutChecker. + intros. + revert H. + case_eq (ZWeakChecker (neg_nformula (makeLb e lb) :: l) pf); intros ; [idtac | discriminate]. + generalize (ZWeakChecker_sound _ _ H env). + intros. + inversion H0 ; subst ; clear H0. + apply -> make_conj_impl. + simpl in H1. + rewrite <- make_conj_impl in H1. + intros. + apply -> neg_nformula_sound ; auto. + red ; intros. + apply H1 ; auto. + clear H H1 H0. + generalize (lb_lbcut env e lb). + intros. + destruct (Zeval_nformula_dec env ((neg_nformula (makeLb e lb)))). + auto. + rewrite -> neg_nformula_sound in H0. + assert (HH := H H0). + rewrite <- neg_nformula_sound in HH. + tauto. + reflexivity. + unfold makeLb. + destruct lb. + reflexivity. +Qed. + + +Lemma cutChecker_sound_bound : forall e lb pf l res, cutChecker l e lb pf = Some res -> + forall env, make_conj (Zeval_nformula env) l -> (Zeval_expr env e >= qceiling lb)%Z. +Proof. + intros. + generalize (cutChecker_sound _ _ _ _ _ H env). + intros. + rewrite <- (make_conj_impl) in H1. + generalize (H1 H0). + unfold cutChecker in H. + destruct (ZWeakChecker (neg_nformula (makeLb e lb) :: l) pf). + unfold makeLbCut in H. + inversion H ; subst. + clear H. + simpl. + rewrite Zeval_expr_compat. + unfold Zeval_expr'. + auto with zarith. + discriminate. +Qed. + + +Lemma ZChecker_sound : forall w l, ZChecker l w = true -> forall env, make_impl (Zeval_nformula env) l False. +Proof. + induction w using (well_founded_ind (well_founded_ltof _ bdepth)). + destruct w. + (* RatProof *) + simpl. + intros. + eapply ZWeakChecker_sound. + apply H0. + (* CutProof *) + simpl. + intro. + case_eq (cutChecker l p q z) ; intros. + generalize (cutChecker_sound _ _ _ _ _ H0 env). + intro. + assert (make_impl (Zeval_nformula env) (n::l) False). + eapply (H w) ; auto. + unfold ltof. + simpl. + auto with arith. + simpl in H3. + rewrite <- make_conj_impl in H2. + rewrite <- make_conj_impl in H3. + rewrite <- make_conj_impl. + tauto. + discriminate. + (* EnumProof *) + intro. + rewrite ZChecker_simpl. + case_eq (cutChecker l0 p q z). + rename q into llb. + case_eq (cutChecker l0 (PEopp p) (- q0) z0). + intros. + rename q0 into uub. + (* get the bounds of the enum *) + rewrite <- make_conj_impl. + intro. + assert (qceiling llb <= Zeval_expr env p <= - qceiling ( - uub))%Z. + generalize (cutChecker_sound_bound _ _ _ _ _ H0 env H3). + generalize (cutChecker_sound_bound _ _ _ _ _ H1 env H3). + intros. + rewrite Zeval_expr_simpl in H5. + auto with zarith. + clear H0 H1. + revert H2 H3 H4. + generalize (qceiling llb) (- qceiling (- uub))%Z. + set (FF := (fix label (pfs : list ProofTerm) (lb ub : Z) {struct pfs} : bool := + match pfs with + | nil => if Z_gt_dec lb ub then true else false + | pf :: rsr => + (ZChecker ((PEsub p (PEc lb), Equal) :: l0) pf && + label rsr (lb + 1)%Z ub)%bool + end)). + intros z1 z2. + intros. + assert (forall x, z1 <= x <= z2 -> exists pr, + (In pr l /\ + ZChecker ((PEsub p (PEc x),Equal) :: l0) pr = true))%Z. + clear H. + revert H2. + clear H4. + revert z1 z2. + induction l;simpl ;intros. + destruct (Z_gt_dec z1 z2). + intros. + apply False_ind ; omega. + discriminate. + intros. + simpl in H2. + flatten_bool. + assert (HH:(x = z1 \/ z1 +1 <=x)%Z) by omega. + destruct HH. + subst. + exists a ; auto. + assert (z1 + 1 <= x <= z2)%Z by omega. + destruct (IHl _ _ H1 _ H4). + destruct H5. + exists x0 ; split;auto. + (*/asser *) + destruct (H0 _ H4) as [pr [Hin Hcheker]]. + assert (make_impl (Zeval_nformula env) ((PEsub p (PEc (Zeval_expr env p)),Equal) :: l0) False). + apply (H pr);auto. + apply in_bdepth ; auto. + rewrite <- make_conj_impl in H1. + apply H1. + rewrite make_conj_cons. + split ;auto. + rewrite Zeval_nformula_simpl; + unfold Zeval_op1; + rewrite Zeval_expr_simpl. + generalize (Zeval_expr env p). + intros. + rewrite Zeval_expr_simpl. + auto with zarith. + intros ; discriminate. + intros ; discriminate. +Qed. + +Definition ZTautoChecker (f : BFormula (Formula Z)) (w: list ProofTerm): bool := + @tauto_checker (Formula Z) (NFormula Z) normalise negate ProofTerm ZChecker f w. + +Lemma ZTautoChecker_sound : forall f w, ZTautoChecker f w = true -> forall env, eval_f (Zeval_formula env) f. +Proof. + intros f w. + unfold ZTautoChecker. + apply (tauto_checker_sound Zeval_formula Zeval_nformula). + apply Zeval_nformula_dec. + intros env t. + rewrite normalise_correct ; auto. + intros env t. + rewrite negate_correct ; auto. + intros t w0. + apply ZChecker_sound. +Qed. + + +Open Scope Z_scope. + + +Fixpoint map_cone (f: nat -> nat) (e:ZWitness) : ZWitness := + match e with + | S_In n => S_In _ (f n) + | S_Ideal e cm => S_Ideal e (map_cone f cm) + | S_Square _ => e + | S_Monoid l => S_Monoid _ (List.map f l) + | S_Mult cm1 cm2 => S_Mult (map_cone f cm1) (map_cone f cm2) + | S_Add cm1 cm2 => S_Add (map_cone f cm1) (map_cone f cm2) + | _ => e + end. + +Fixpoint indexes (e:ZWitness) : list nat := + match e with + | S_In n => n::nil + | S_Ideal e cm => indexes cm + | S_Square e => nil + | S_Monoid l => l + | S_Mult cm1 cm2 => (indexes cm1)++ (indexes cm2) + | S_Add cm1 cm2 => (indexes cm1)++ (indexes cm2) + | _ => nil + end. + +(** To ease bindings from ml code **) +(*Definition varmap := Quote.varmap.*) +Definition make_impl := Refl.make_impl. +Definition make_conj := Refl.make_conj. + +Require VarMap. + +(*Definition varmap_type := VarMap.t Z. *) +Definition env := PolEnv Z. +Definition node := @VarMap.Node Z. +Definition empty := @VarMap.Empty Z. +Definition leaf := @VarMap.Leaf Z. + +Definition coneMember := ZWitness. + +Definition eval := Zeval_formula. + +Definition prod_pos_nat := prod positive nat. + +Require Import Int. + + +Definition n_of_Z (z:Z) : BinNat.N := + match z with + | Z0 => N0 + | Zpos p => Npos p + | Zneg p => N0 + end. + + + diff --git a/contrib/micromega/certificate.ml b/contrib/micromega/certificate.ml new file mode 100644 index 00000000..88e882e6 --- /dev/null +++ b/contrib/micromega/certificate.ml @@ -0,0 +1,618 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* *) +(* Micromega: A reflexive tactic using the Positivstellensatz *) +(* *) +(* Frédéric Besson (Irisa/Inria) 2006-2008 *) +(* *) +(************************************************************************) + +(* We take as input a list of polynomials [p1...pn] and return an unfeasibility + certificate polynomial. *) + +(*open Micromega.Polynomial*) +open Big_int +open Num + +module Mc = Micromega +module Ml2C = Mutils.CamlToCoq +module C2Ml = Mutils.CoqToCaml + +let (<+>) = add_num +let (<->) = minus_num +let (<*>) = mult_num + +type var = Mc.positive + +module Monomial : +sig + type t + val const : t + val var : var -> t + val find : var -> t -> int + val mult : var -> t -> t + val prod : t -> t -> t + val compare : t -> t -> int + val pp : out_channel -> t -> unit + val fold : (var -> int -> 'a -> 'a) -> t -> 'a -> 'a +end + = +struct + (* A monomial is represented by a multiset of variables *) + module Map = Map.Make(struct type t = var let compare = Pervasives.compare end) + open Map + + type t = int Map.t + + (* The monomial that corresponds to a constant *) + let const = Map.empty + + (* The monomial 'x' *) + let var x = Map.add x 1 Map.empty + + (* Get the degre of a variable in a monomial *) + let find x m = try find x m with Not_found -> 0 + + (* Multiply a monomial by a variable *) + let mult x m = add x ( (find x m) + 1) m + + (* Product of monomials *) + let prod m1 m2 = Map.fold (fun k d m -> add k ((find k m) + d) m) m1 m2 + + (* Total ordering of monomials *) + let compare m1 m2 = Map.compare Pervasives.compare m1 m2 + + let pp o m = Map.iter (fun k v -> + if v = 1 then Printf.fprintf o "x%i." (C2Ml.index k) + else Printf.fprintf o "x%i^%i." (C2Ml.index k) v) m + + let fold = fold + +end + + +module Poly : + (* A polynomial is a map of monomials *) + (* + This is probably a naive implementation + (expected to be fast enough - Coq is probably the bottleneck) + *The new ring contribution is using a sparse Horner representation. + *) +sig + type t + val get : Monomial.t -> t -> num + val variable : var -> t + val add : Monomial.t -> num -> t -> t + val constant : num -> t + val mult : Monomial.t -> num -> t -> t + val product : t -> t -> t + val addition : t -> t -> t + val uminus : t -> t + val fold : (Monomial.t -> num -> 'a -> 'a) -> t -> 'a -> 'a + val pp : out_channel -> t -> unit + val compare : t -> t -> int +end = +struct + (*normalisation bug : 0*x ... *) + module P = Map.Make(Monomial) + open P + + type t = num P.t + + let pp o p = P.iter (fun k v -> + if compare_num v (Int 0) <> 0 + then + if Monomial.compare Monomial.const k = 0 + then Printf.fprintf o "%s " (string_of_num v) + else Printf.fprintf o "%s*%a " (string_of_num v) Monomial.pp k) p + + (* Get the coefficient of monomial mn *) + let get : Monomial.t -> t -> num = + fun mn p -> try find mn p with Not_found -> (Int 0) + + + (* The polynomial 1.x *) + let variable : var -> t = + fun x -> add (Monomial.var x) (Int 1) empty + + (*The constant polynomial *) + let constant : num -> t = + fun c -> add (Monomial.const) c empty + + (* The addition of a monomial *) + + let add : Monomial.t -> num -> t -> t = + fun mn v p -> + let vl = (get mn p) <+> v in + add mn vl p + + + (** Design choice: empty is not a polynomial + I do not remember why .... + **) + + (* The product by a monomial *) + let mult : Monomial.t -> num -> t -> t = + fun mn v p -> + fold (fun mn' v' res -> P.add (Monomial.prod mn mn') (v<*>v') res) p empty + + + let addition : t -> t -> t = + fun p1 p2 -> fold (fun mn v p -> add mn v p) p1 p2 + + + let product : t -> t -> t = + fun p1 p2 -> + fold (fun mn v res -> addition (mult mn v p2) res ) p1 empty + + + let uminus : t -> t = + fun p -> map (fun v -> minus_num v) p + + let fold = P.fold + + let compare = compare compare_num +end + +open Mutils +type 'a number_spec = { + bigint_to_number : big_int -> 'a; + number_to_num : 'a -> num; + zero : 'a; + unit : 'a; + mult : 'a -> 'a -> 'a; + eqb : 'a -> 'a -> Mc.bool +} + +let z_spec = { + bigint_to_number = Ml2C.bigint ; + number_to_num = (fun x -> Big_int (C2Ml.z_big_int x)); + zero = Mc.Z0; + unit = Mc.Zpos Mc.XH; + mult = Mc.zmult; + eqb = Mc.zeq_bool +} + + +let q_spec = { + bigint_to_number = (fun x -> {Mc.qnum = Ml2C.bigint x; Mc.qden = Mc.XH}); + number_to_num = C2Ml.q_to_num; + zero = {Mc.qnum = Mc.Z0;Mc.qden = Mc.XH}; + unit = {Mc.qnum = (Mc.Zpos Mc.XH) ; Mc.qden = Mc.XH}; + mult = Mc.qmult; + eqb = Mc.qeq_bool +} + +let r_spec = z_spec + + + + +let dev_form n_spec p = + let rec dev_form p = + match p with + | Mc.PEc z -> Poly.constant (n_spec.number_to_num z) + | Mc.PEX v -> Poly.variable v + | Mc.PEmul(p1,p2) -> + let p1 = dev_form p1 in + let p2 = dev_form p2 in + Poly.product p1 p2 + | Mc.PEadd(p1,p2) -> Poly.addition (dev_form p1) (dev_form p2) + | Mc.PEopp p -> Poly.uminus (dev_form p) + | Mc.PEsub(p1,p2) -> Poly.addition (dev_form p1) (Poly.uminus (dev_form p2)) + | Mc.PEpow(p,n) -> + let p = dev_form p in + let n = C2Ml.n n in + let rec pow n = + if n = 0 + then Poly.constant (n_spec.number_to_num n_spec.unit) + else Poly.product p (pow (n-1)) in + pow n in + dev_form p + + +let monomial_to_polynomial mn = + Monomial.fold + (fun v i acc -> + let mn = if i = 1 then Mc.PEX v else Mc.PEpow (Mc.PEX v ,Ml2C.n i) in + if acc = Mc.PEc (Mc.Zpos Mc.XH) + then mn + else Mc.PEmul(mn,acc)) + mn + (Mc.PEc (Mc.Zpos Mc.XH)) + +let list_to_polynomial vars l = + assert (List.for_all (fun x -> ceiling_num x =/ x) l); + let var x = monomial_to_polynomial (List.nth vars x) in + let rec xtopoly p i = function + | [] -> p + | c::l -> if c =/ (Int 0) then xtopoly p (i+1) l + else let c = Mc.PEc (Ml2C.bigint (numerator c)) in + let mn = + if c = Mc.PEc (Mc.Zpos Mc.XH) + then var i + else Mc.PEmul (c,var i) in + let p' = if p = Mc.PEc Mc.Z0 then mn else + Mc.PEadd (mn, p) in + xtopoly p' (i+1) l in + + xtopoly (Mc.PEc Mc.Z0) 0 l + +let rec fixpoint f x = + let y' = f x in + if y' = x then y' + else fixpoint f y' + + + + + + + + +let rec_simpl_cone n_spec e = + let simpl_cone = + Mc.simpl_cone n_spec.zero n_spec.unit n_spec.mult n_spec.eqb in + + let rec rec_simpl_cone = function + | Mc.S_Mult(t1, t2) -> + simpl_cone (Mc.S_Mult (rec_simpl_cone t1, rec_simpl_cone t2)) + | Mc.S_Add(t1,t2) -> + simpl_cone (Mc.S_Add (rec_simpl_cone t1, rec_simpl_cone t2)) + | x -> simpl_cone x in + rec_simpl_cone e + + +let simplify_cone n_spec c = fixpoint (rec_simpl_cone n_spec) c + +type cone_prod = + Const of cone + | Ideal of cone *cone + | Mult of cone * cone + | Other of cone +and cone = Mc.zWitness + + + +let factorise_linear_cone c = + + let rec cone_list c l = + match c with + | Mc.S_Add (x,r) -> cone_list r (x::l) + | _ -> c :: l in + + let factorise c1 c2 = + match c1 , c2 with + | Mc.S_Ideal(x,y) , Mc.S_Ideal(x',y') -> + if x = x' then Some (Mc.S_Ideal(x, Mc.S_Add(y,y'))) else None + | Mc.S_Mult(x,y) , Mc.S_Mult(x',y') -> + if x = x' then Some (Mc.S_Mult(x, Mc.S_Add(y,y'))) else None + | _ -> None in + + let rec rebuild_cone l pending = + match l with + | [] -> (match pending with + | None -> Mc.S_Z + | Some p -> p + ) + | e::l -> + (match pending with + | None -> rebuild_cone l (Some e) + | Some p -> (match factorise p e with + | None -> Mc.S_Add(p, rebuild_cone l (Some e)) + | Some f -> rebuild_cone l (Some f) ) + ) in + + (rebuild_cone (List.sort Pervasives.compare (cone_list c [])) None) + + + +(* The binding with Fourier might be a bit obsolete + -- how does it handle equalities ? *) + +(* Certificates are elements of the cone such that P = 0 *) + +(* To begin with, we search for certificates of the form: + a1.p1 + ... an.pn + b1.q1 +... + bn.qn + c = 0 + where pi >= 0 qi > 0 + ai >= 0 + bi >= 0 + Sum bi + c >= 1 + This is a linear problem: each monomial is considered as a variable. + Hence, we can use fourier. + + The variable c is at index 0 +*) + +open Mfourier + (*module Fourier = Fourier(Vector.VList)(SysSet(Vector.VList))*) + (*module Fourier = Fourier(Vector.VSparse)(SysSetAlt(Vector.VSparse))*) +module Fourier = Mfourier.Fourier(Vector.VSparse)(*(SysSetAlt(Vector.VMap))*) + +module Vect = Fourier.Vect +open Fourier.Cstr + +(* fold_left followed by a rev ! *) + +let constrain_monomial mn l = + let coeffs = List.fold_left (fun acc p -> (Poly.get mn p)::acc) [] l in + if mn = Monomial.const + then + { coeffs = Vect.from_list ((Big_int unit_big_int):: (List.rev coeffs)) ; + op = Eq ; + cst = Big_int zero_big_int } + else + { coeffs = Vect.from_list ((Big_int zero_big_int):: (List.rev coeffs)) ; + op = Eq ; + cst = Big_int zero_big_int } + + +let positivity l = + let rec xpositivity i l = + match l with + | [] -> [] + | (_,Mc.Equal)::l -> xpositivity (i+1) l + | (_,_)::l -> + {coeffs = Vect.update (i+1) (fun _ -> Int 1) Vect.null ; + op = Ge ; + cst = Int 0 } :: (xpositivity (i+1) l) + in + xpositivity 0 l + + +let string_of_op = function + | Mc.Strict -> "> 0" + | Mc.NonStrict -> ">= 0" + | Mc.Equal -> "= 0" + | Mc.NonEqual -> "<> 0" + + + +(* If the certificate includes at least one strict inequality, + the obtained polynomial can also be 0 *) +let build_linear_system l = + + (* Gather the monomials: HINT add up of the polynomials *) + let l' = List.map fst l in + let monomials = + List.fold_left (fun acc p -> Poly.addition p acc) (Poly.constant (Int 0)) l' + in (* For each monomial, compute a constraint *) + let s0 = + Poly.fold (fun mn _ res -> (constrain_monomial mn l')::res) monomials [] in + (* I need at least something strictly positive *) + let strict = { + coeffs = Vect.from_list ((Big_int unit_big_int):: + (List.map (fun (x,y) -> + match y with Mc.Strict -> + Big_int unit_big_int + | _ -> Big_int zero_big_int) l)); + op = Ge ; cst = Big_int unit_big_int } in + (* Add the positivity constraint *) + {coeffs = Vect.from_list ([Big_int unit_big_int]) ; + op = Ge ; + cst = Big_int zero_big_int}::(strict::(positivity l)@s0) + + +let big_int_to_z = Ml2C.bigint + +(* For Q, this is a pity that the certificate has been scaled + -- at a lower layer, certificates are using nums... *) +let make_certificate n_spec cert li = + let bint_to_cst = n_spec.bigint_to_number in + match cert with + | [] -> None + | e::cert' -> + let cst = match compare_big_int e zero_big_int with + | 0 -> Mc.S_Z + | 1 -> Mc.S_Pos (bint_to_cst e) + | _ -> failwith "positivity error" + in + let rec scalar_product cert l = + match cert with + | [] -> Mc.S_Z + | c::cert -> match l with + | [] -> failwith "make_certificate(1)" + | i::l -> + let r = scalar_product cert l in + match compare_big_int c zero_big_int with + | -1 -> Mc.S_Add ( + Mc.S_Ideal (Mc.PEc ( bint_to_cst c), Mc.S_In (Ml2C.nat i)), + r) + | 0 -> r + | _ -> Mc.S_Add ( + Mc.S_Mult (Mc.S_Pos (bint_to_cst c), Mc.S_In (Ml2C.nat i)), + r) in + + Some ((factorise_linear_cone + (simplify_cone n_spec (Mc.S_Add (cst, scalar_product cert' li))))) + + +exception Found of Monomial.t + +let raw_certificate l = + let sys = build_linear_system l in + try + match Fourier.find_point sys with + | None -> None + | Some cert -> Some (rats_to_ints (Vect.to_list cert)) + (* should not use rats_to_ints *) + with x -> + if debug + then (Printf.printf "raw certificate %s" (Printexc.to_string x); + flush stdout) ; + None + + +let simple_linear_prover to_constant l = + let (lc,li) = List.split l in + match raw_certificate lc with + | None -> None (* No certificate *) + | Some cert -> make_certificate to_constant cert li + + + +let linear_prover n_spec l = + let li = List.combine l (interval 0 (List.length l -1)) in + let (l1,l') = List.partition + (fun (x,_) -> if snd' x = Mc.NonEqual then true else false) li in + let l' = List.map + (fun (c,i) -> let (Mc.Pair(x,y)) = c in + match y with + Mc.NonEqual -> failwith "cannot happen" + | y -> ((dev_form n_spec x, y),i)) l' in + + simple_linear_prover n_spec l' + + +let linear_prover n_spec l = + try linear_prover n_spec l with + x -> (print_string (Printexc.to_string x); None) + +(* zprover.... *) + +(* I need to gather the set of variables ---> + Then go for fold + Once I have an interval, I need a certificate : 2 other fourier elims. + (I could probably get the certificate directly + as it is done in the fourier contrib.) +*) + +let make_linear_system l = + let l' = List.map fst l in + let monomials = List.fold_left (fun acc p -> Poly.addition p acc) + (Poly.constant (Int 0)) l' in + let monomials = Poly.fold + (fun mn _ l -> if mn = Monomial.const then l else mn::l) monomials [] in + (List.map (fun (c,op) -> + {coeffs = Vect.from_list (List.map (fun mn -> (Poly.get mn c)) monomials) ; + op = op ; + cst = minus_num ( (Poly.get Monomial.const c))}) l + ,monomials) + + +open Interval +let pplus x y = Mc.PEadd(x,y) +let pmult x y = Mc.PEmul(x,y) +let pconst x = Mc.PEc x +let popp x = Mc.PEopp x + +let debug = false + +(* keep track of enumerated vectors *) +let rec mem p x l = + match l with [] -> false | e::l -> if p x e then true else mem p x l + +let rec remove_assoc p x l = + match l with [] -> [] | e::l -> if p x (fst e) then + remove_assoc p x l else e::(remove_assoc p x l) + +let eq x y = Vect.compare x y = 0 + +(* Beurk... this code is a shame *) + +let rec zlinear_prover sys = xzlinear_prover [] sys + +and xzlinear_prover enum l : (Mc.proofTerm option) = + match linear_prover z_spec l with + | Some prf -> Some (Mc.RatProof prf) + | None -> + let ll = List.fold_right (fun (Mc.Pair(e,k)) r -> match k with + Mc.NonEqual -> r + | k -> (dev_form z_spec e , + match k with + | Mc.Strict | Mc.NonStrict -> Ge + (* Loss of precision -- weakness of fourier*) + | Mc.Equal -> Eq + | Mc.NonEqual -> failwith "Cannot happen") :: r) l [] in + + let (sys,var) = make_linear_system ll in + let res = + match Fourier.find_Q_interval sys with + | Some(i,x,j) -> if i =/ j + then Some(i,Vect.set x (Int 1) Vect.null,i) else None + | None -> None in + let res = match res with + | None -> + begin + let candidates = List.fold_right + (fun cstr acc -> + let gcd = Big_int (Vect.gcd cstr.coeffs) in + let vect = Vect.mul (Int 1 // gcd) cstr.coeffs in + if mem eq vect enum then acc + else ((vect,Fourier.optimise vect sys)::acc)) sys [] in + let candidates = List.fold_left (fun l (x,i) -> + match i with + None -> (x,Empty)::l + | Some i -> (x,i)::l) [] (candidates) in + match List.fold_left (fun (x1,i1) (x2,i2) -> + if smaller_itv i1 i2 + then (x1,i1) else (x2,i2)) (Vect.null,Itv(None,None)) candidates + with + | (i,Empty) -> None + | (x,Itv(Some i, Some j)) -> Some(i,x,j) + | (x,Point n) -> Some(n,x,n) + | x -> match Fourier.find_Q_interval sys with + | None -> None + | Some(i,x,j) -> + if i =/ j + then Some(i,Vect.set x (Int 1) Vect.null,i) + else None + end + | _ -> res in + + match res with + | Some (lb,e,ub) -> + let (lbn,lbd) = + (Ml2C.bigint (sub_big_int (numerator lb) unit_big_int), + Ml2C.bigint (denominator lb)) in + let (ubn,ubd) = + (Ml2C.bigint (add_big_int unit_big_int (numerator ub)) , + Ml2C.bigint (denominator ub)) in + let expr = list_to_polynomial var (Vect.to_list e) in + (match + (*x <= ub -> x > ub *) + linear_prover z_spec + (Mc.Pair(pplus (pmult (pconst ubd) expr) (popp (pconst ubn)), + Mc.NonStrict) :: l), + (* lb <= x -> lb > x *) + linear_prover z_spec + (Mc.Pair( pplus (popp (pmult (pconst lbd) expr)) (pconst lbn) , + Mc.NonStrict)::l) + with + | Some cub , Some clb -> + (match zlinear_enum (e::enum) expr + (ceiling_num lb) (floor_num ub) l + with + | None -> None + | Some prf -> + Some (Mc.EnumProof(Ml2C.q lb,expr,Ml2C.q ub,clb,cub,prf))) + | _ -> None + ) + | _ -> None +and xzlinear_enum enum expr clb cub l = + if clb >/ cub + then Some Mc.Nil + else + let pexpr = pplus (popp (pconst (Ml2C.bigint (numerator clb)))) expr in + let sys' = (Mc.Pair(pexpr, Mc.Equal))::l in + match xzlinear_prover enum sys' with + | None -> if debug then print_string "zlp?"; None + | Some prf -> if debug then print_string "zlp!"; + match zlinear_enum enum expr (clb +/ (Int 1)) cub l with + | None -> None + | Some prfl -> Some (Mc.Cons(prf,prfl)) + +and zlinear_enum enum expr clb cub l = + let res = xzlinear_enum enum expr clb cub l in + if debug then Printf.printf "zlinear_enum %s %s -> %s\n" + (string_of_num clb) + (string_of_num cub) + (match res with + | None -> "None" + | Some r -> "Some") ; res + diff --git a/contrib/micromega/coq_micromega.ml b/contrib/micromega/coq_micromega.ml new file mode 100644 index 00000000..29e2a183 --- /dev/null +++ b/contrib/micromega/coq_micromega.ml @@ -0,0 +1,1290 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* *) +(* Micromega: A reflexive tactic using the Positivstellensatz *) +(* *) +(* Frédéric Besson (Irisa/Inria) 2006-2008 *) +(* *) +(************************************************************************) + +open Mutils +let debug = false + +let time str f x = + let t0 = (Unix.times()).Unix.tms_utime in + let res = f x in + let t1 = (Unix.times()).Unix.tms_utime in + (*if debug then*) (Printf.printf "time %s %f\n" str (t1 -. t0) ; + flush stdout); + res + +type ('a,'b) formula = + | TT + | FF + | X of 'b + | A of 'a * Names.name + | C of ('a,'b) formula * ('a,'b) formula * Names.name + | D of ('a,'b) formula * ('a,'b) formula * Names.name + | N of ('a,'b) formula * Names.name + | I of ('a,'b) formula * ('a,'b) formula * Names.name + +let none = Names.Anonymous + +let tag_formula t f = + match f with + | A(x,_) -> A(x,t) + | C(x,y,_) -> C(x,y,t) + | D(x,y,_) -> D(x,y,t) + | N(x,_) -> N(x,t) + | I(x,y,_) -> I(x,y,t) + | _ -> f + +let tt = [] +let ff = [ [] ] + + +type ('constant,'contr) sentence = + ('constant Micromega.formula, 'contr) formula + +let cnf negate normalise f = + let negate a = + CoqToCaml.list (fun cl -> CoqToCaml.list (fun x -> x) cl) (negate a) in + + let normalise a = + CoqToCaml.list (fun cl -> CoqToCaml.list (fun x -> x) cl) (normalise a) in + + let and_cnf x y = x @ y in + let or_clause_cnf t f = List.map (fun x -> t@x ) f in + + let rec or_cnf f f' = + match f with + | [] -> tt + | e :: rst -> (or_cnf rst f') @ (or_clause_cnf e f') in + + let rec xcnf (pol : bool) f = + match f with + | TT -> if pol then tt else ff (* ?? *) + | FF -> if pol then ff else tt (* ?? *) + | X p -> if pol then ff else ff (* ?? *) + | A(x,t) -> if pol then normalise x else negate x + | N(e,t) -> xcnf (not pol) e + | C(e1,e2,t) -> + (if pol then and_cnf else or_cnf) (xcnf pol e1) (xcnf pol e2) + | D(e1,e2,t) -> + (if pol then or_cnf else and_cnf) (xcnf pol e1) (xcnf pol e2) + | I(e1,e2,t) -> + (if pol then or_cnf else and_cnf) (xcnf (not pol) e1) (xcnf pol e2) in + + xcnf true f + + + +module M = +struct + open Coqlib + open Term + (* let constant = gen_constant_in_modules "Omicron" coq_modules*) + + + let logic_dir = ["Coq";"Logic";"Decidable"] + let coq_modules = + init_modules @ + [logic_dir] @ arith_modules @ zarith_base_modules @ + [ ["Coq";"Lists";"List"]; + ["ZMicromega"]; + ["Tauto"]; + ["RingMicromega"]; + ["EnvRing"]; + ["Coq"; "micromega"; "ZMicromega"]; + ["Coq" ; "micromega" ; "Tauto"]; + ["Coq" ; "micromega" ; "RingMicromega"]; + ["Coq" ; "micromega" ; "EnvRing"]; + ["Coq";"QArith"; "QArith_base"]; + ["Coq";"Reals" ; "Rdefinitions"]; + ["LRing_normalise"]] + + let constant = gen_constant_in_modules "ZMicromega" coq_modules + + let coq_and = lazy (constant "and") + let coq_or = lazy (constant "or") + let coq_not = lazy (constant "not") + let coq_iff = lazy (constant "iff") + let coq_True = lazy (constant "True") + let coq_False = lazy (constant "False") + + let coq_cons = lazy (constant "cons") + let coq_nil = lazy (constant "nil") + let coq_list = lazy (constant "list") + + let coq_O = lazy (constant "O") + let coq_S = lazy (constant "S") + let coq_nat = lazy (constant "nat") + + let coq_NO = lazy + (gen_constant_in_modules "N" [ ["Coq";"NArith";"BinNat" ]] "N0") + let coq_Npos = lazy + (gen_constant_in_modules "N" [ ["Coq";"NArith"; "BinNat"]] "Npos") + (* let coq_n = lazy (constant "N")*) + + let coq_pair = lazy (constant "pair") + let coq_None = lazy (constant "None") + let coq_option = lazy (constant "option") + let coq_positive = lazy (constant "positive") + let coq_xH = lazy (constant "xH") + let coq_xO = lazy (constant "xO") + let coq_xI = lazy (constant "xI") + + let coq_N0 = lazy (constant "N0") + let coq_N0 = lazy (constant "Npos") + + + let coq_Z = lazy (constant "Z") + let coq_Q = lazy (constant "Q") + let coq_R = lazy (constant "R") + + let coq_ZERO = lazy (constant "Z0") + let coq_POS = lazy (constant "Zpos") + let coq_NEG = lazy (constant "Zneg") + + let coq_QWitness = lazy + (gen_constant_in_modules "QMicromega" + [["Coq"; "micromega"; "QMicromega"]] "QWitness") + let coq_ZWitness = lazy + (gen_constant_in_modules "QMicromega" + [["Coq"; "micromega"; "ZMicromega"]] "ZWitness") + + + let coq_Build_Witness = lazy (constant "Build_Witness") + + + let coq_Qmake = lazy (constant "Qmake") + + let coq_proofTerm = lazy (constant "ProofTerm") + let coq_ratProof = lazy (constant "RatProof") + let coq_cutProof = lazy (constant "CutProof") + let coq_enumProof = lazy (constant "EnumProof") + + let coq_Zgt = lazy (constant "Zgt") + let coq_Zge = lazy (constant "Zge") + let coq_Zle = lazy (constant "Zle") + let coq_Zlt = lazy (constant "Zlt") + let coq_Eq = lazy (constant "eq") + + let coq_Zplus = lazy (constant "Zplus") + let coq_Zminus = lazy (constant "Zminus") + let coq_Zopp = lazy (constant "Zopp") + let coq_Zmult = lazy (constant "Zmult") + let coq_N_of_Z = lazy + (gen_constant_in_modules "ZArithRing" + [["Coq";"setoid_ring";"ZArithRing"]] "N_of_Z") + + + let coq_PEX = lazy (constant "PEX" ) + let coq_PEc = lazy (constant"PEc") + let coq_PEadd = lazy (constant "PEadd") + let coq_PEopp = lazy (constant "PEopp") + let coq_PEmul = lazy (constant "PEmul") + let coq_PEsub = lazy (constant "PEsub") + let coq_PEpow = lazy (constant "PEpow") + + + let coq_OpEq = lazy (constant "OpEq") + let coq_OpNEq = lazy (constant "OpNEq") + let coq_OpLe = lazy (constant "OpLe") + let coq_OpLt = lazy (constant "OpLt") + let coq_OpGe = lazy (constant "OpGe") + let coq_OpGt = lazy (constant "OpGt") + + + let coq_S_In = lazy (constant "S_In") + let coq_S_Square = lazy (constant "S_Square") + let coq_S_Monoid = lazy (constant "S_Monoid") + let coq_S_Ideal = lazy (constant "S_Ideal") + let coq_S_Mult = lazy (constant "S_Mult") + let coq_S_Add = lazy (constant "S_Add") + let coq_S_Pos = lazy (constant "S_Pos") + let coq_S_Z = lazy (constant "S_Z") + let coq_coneMember = lazy (constant "coneMember") + + + let coq_make_impl = lazy + (gen_constant_in_modules "Zmicromega" [["Refl"]] "make_impl") + let coq_make_conj = lazy + (gen_constant_in_modules "Zmicromega" [["Refl"]] "make_conj") + + let coq_Build = lazy + (gen_constant_in_modules "RingMicromega" + [["Coq" ; "micromega" ; "RingMicromega"] ; ["RingMicromega"] ] + "Build_Formula") + let coq_Cstr = lazy + (gen_constant_in_modules "RingMicromega" + [["Coq" ; "micromega" ; "RingMicromega"] ; ["RingMicromega"] ] "Formula") + + type parse_error = + | Ukn + | BadStr of string + | BadNum of int + | BadTerm of Term.constr + | Msg of string + | Goal of (Term.constr list ) * Term.constr * parse_error + + let string_of_error = function + | Ukn -> "ukn" + | BadStr s -> s + | BadNum i -> string_of_int i + | BadTerm _ -> "BadTerm" + | Msg s -> s + | Goal _ -> "Goal" + + + exception ParseError + + + + + let get_left_construct term = + match Term.kind_of_term term with + | Term.Construct(_,i) -> (i,[| |]) + | Term.App(l,rst) -> + (match Term.kind_of_term l with + | Term.Construct(_,i) -> (i,rst) + | _ -> raise ParseError + ) + | _ -> raise ParseError + + module Mc = Micromega + + let rec parse_nat term = + let (i,c) = get_left_construct term in + match i with + | 1 -> Mc.O + | 2 -> Mc.S (parse_nat (c.(0))) + | i -> raise ParseError + + + let pp_nat o n = Printf.fprintf o "%i" (CoqToCaml.nat n) + + + let rec dump_nat x = + match x with + | Mc.O -> Lazy.force coq_O + | Mc.S p -> Term.mkApp(Lazy.force coq_S,[| dump_nat p |]) + + + let rec parse_positive term = + let (i,c) = get_left_construct term in + match i with + | 1 -> Mc.XI (parse_positive c.(0)) + | 2 -> Mc.XO (parse_positive c.(0)) + | 3 -> Mc.XH + | i -> raise ParseError + + + let rec dump_positive x = + match x with + | Mc.XH -> Lazy.force coq_xH + | Mc.XO p -> Term.mkApp(Lazy.force coq_xO,[| dump_positive p |]) + | Mc.XI p -> Term.mkApp(Lazy.force coq_xI,[| dump_positive p |]) + + let pp_positive o x = Printf.fprintf o "%i" (CoqToCaml.positive x) + + + let rec dump_n x = + match x with + | Mc.N0 -> Lazy.force coq_N0 + | Mc.Npos p -> Term.mkApp(Lazy.force coq_Npos,[| dump_positive p|]) + + let rec dump_index x = + match x with + | Mc.XH -> Lazy.force coq_xH + | Mc.XO p -> Term.mkApp(Lazy.force coq_xO,[| dump_index p |]) + | Mc.XI p -> Term.mkApp(Lazy.force coq_xI,[| dump_index p |]) + + + let pp_index o x = Printf.fprintf o "%i" (CoqToCaml.index x) + + let rec dump_n x = + match x with + | Mc.N0 -> Lazy.force coq_NO + | Mc.Npos p -> Term.mkApp(Lazy.force coq_Npos,[| dump_positive p |]) + + let rec pp_n o x = output_string o (string_of_int (CoqToCaml.n x)) + + let dump_pair t1 t2 dump_t1 dump_t2 (Mc.Pair (x,y)) = + Term.mkApp(Lazy.force coq_pair,[| t1 ; t2 ; dump_t1 x ; dump_t2 y|]) + + + let rec parse_z term = + let (i,c) = get_left_construct term in + match i with + | 1 -> Mc.Z0 + | 2 -> Mc.Zpos (parse_positive c.(0)) + | 3 -> Mc.Zneg (parse_positive c.(0)) + | i -> raise ParseError + + let dump_z x = + match x with + | Mc.Z0 ->Lazy.force coq_ZERO + | Mc.Zpos p -> Term.mkApp(Lazy.force coq_POS,[| dump_positive p|]) + | Mc.Zneg p -> Term.mkApp(Lazy.force coq_NEG,[| dump_positive p|]) + + let pp_z o x = Printf.fprintf o "%i" (CoqToCaml.z x) + +let dump_num bd1 = + Term.mkApp(Lazy.force coq_Qmake, + [|dump_z (CamlToCoq.bigint (numerator bd1)) ; + dump_positive (CamlToCoq.positive_big_int (denominator bd1)) |]) + + +let dump_q q = + Term.mkApp(Lazy.force coq_Qmake, + [| dump_z q.Micromega.qnum ; dump_positive q.Micromega.qden|]) + +let parse_q term = + match Term.kind_of_term term with + | Term.App(c, args) -> + ( + match Term.kind_of_term c with + Term.Construct((n,j),i) -> + if Names.string_of_kn n = "Coq.QArith.QArith_base#<>#Q" + then {Mc.qnum = parse_z args.(0) ; Mc.qden = parse_positive args.(1) } + else raise ParseError + | _ -> raise ParseError + ) + | _ -> raise ParseError + + let rec parse_list parse_elt term = + let (i,c) = get_left_construct term in + match i with + | 1 -> Mc.Nil + | 2 -> Mc.Cons(parse_elt c.(1), parse_list parse_elt c.(2)) + | i -> raise ParseError + + + let rec dump_list typ dump_elt l = + match l with + | Mc.Nil -> Term.mkApp(Lazy.force coq_nil,[| typ |]) + | Mc.Cons(e,l) -> Term.mkApp(Lazy.force coq_cons, + [| typ; dump_elt e;dump_list typ dump_elt l|]) + + let rec dump_ml_list typ dump_elt l = + match l with + | [] -> Term.mkApp(Lazy.force coq_nil,[| typ |]) + | e::l -> Term.mkApp(Lazy.force coq_cons, + [| typ; dump_elt e;dump_ml_list typ dump_elt l|]) + + + + let pp_list op cl elt o l = + let rec _pp o l = + match l with + | Mc.Nil -> () + | Mc.Cons(e,Mc.Nil) -> Printf.fprintf o "%a" elt e + | Mc.Cons(e,l) -> Printf.fprintf o "%a ,%a" elt e _pp l in + Printf.fprintf o "%s%a%s" op _pp l cl + + + + let pp_var = pp_positive + let dump_var = dump_positive + + let rec pp_expr o e = + match e with + | Mc.PEX n -> Printf.fprintf o "V %a" pp_var n + | Mc.PEc z -> pp_z o z + | Mc.PEadd(e1,e2) -> Printf.fprintf o "(%a)+(%a)" pp_expr e1 pp_expr e2 + | Mc.PEmul(e1,e2) -> Printf.fprintf o "%a*(%a)" pp_expr e1 pp_expr e2 + | Mc.PEopp e -> Printf.fprintf o "-(%a)" pp_expr e + | Mc.PEsub(e1,e2) -> Printf.fprintf o "(%a)-(%a)" pp_expr e1 pp_expr e2 + | Mc.PEpow(e,n) -> Printf.fprintf o "(%a)^(%a)" pp_expr e pp_n n + + + let dump_expr typ dump_z e = + let rec dump_expr e = + match e with + | Mc.PEX n -> mkApp(Lazy.force coq_PEX,[| typ; dump_var n |]) + | Mc.PEc z -> mkApp(Lazy.force coq_PEc,[| typ ; dump_z z |]) + | Mc.PEadd(e1,e2) -> mkApp(Lazy.force coq_PEadd, + [| typ; dump_expr e1;dump_expr e2|]) + | Mc.PEsub(e1,e2) -> mkApp(Lazy.force coq_PEsub, + [| typ; dump_expr e1;dump_expr e2|]) + | Mc.PEopp e -> mkApp(Lazy.force coq_PEopp, + [| typ; dump_expr e|]) + | Mc.PEmul(e1,e2) -> mkApp(Lazy.force coq_PEmul, + [| typ; dump_expr e1;dump_expr e2|]) + | Mc.PEpow(e,n) -> mkApp(Lazy.force coq_PEpow, + [| typ; dump_expr e; dump_n n|]) + in + dump_expr e + + let rec dump_monoid l = dump_list (Lazy.force coq_nat) dump_nat l + + let rec dump_cone typ dump_z e = + let z = Lazy.force typ in + let rec dump_cone e = + match e with + | Mc.S_In n -> mkApp(Lazy.force coq_S_In,[| z; dump_nat n |]) + | Mc.S_Ideal(e,c) -> mkApp(Lazy.force coq_S_Ideal, + [| z; dump_expr z dump_z e ; dump_cone c |]) + | Mc.S_Square e -> mkApp(Lazy.force coq_S_Square, + [| z;dump_expr z dump_z e|]) + | Mc.S_Monoid l -> mkApp (Lazy.force coq_S_Monoid, + [|z; dump_monoid l|]) + | Mc.S_Add(e1,e2) -> mkApp(Lazy.force coq_S_Add, + [| z; dump_cone e1; dump_cone e2|]) + | Mc.S_Mult(e1,e2) -> mkApp(Lazy.force coq_S_Mult, + [| z; dump_cone e1; dump_cone e2|]) + | Mc.S_Pos p -> mkApp(Lazy.force coq_S_Pos,[| z; dump_z p|]) + | Mc.S_Z -> mkApp( Lazy.force coq_S_Z,[| z|]) in + dump_cone e + + + let pp_cone pp_z o e = + let rec pp_cone o e = + match e with + | Mc.S_In n -> + Printf.fprintf o "(S_In %a)%%nat" pp_nat n + | Mc.S_Ideal(e,c) -> + Printf.fprintf o "(S_Ideal %a %a)" pp_expr e pp_cone c + | Mc.S_Square e -> + Printf.fprintf o "(S_Square %a)" pp_expr e + | Mc.S_Monoid l -> + Printf.fprintf o "(S_Monoid %a)" (pp_list "[" "]" pp_nat) l + | Mc.S_Add(e1,e2) -> + Printf.fprintf o "(S_Add %a %a)" pp_cone e1 pp_cone e2 + | Mc.S_Mult(e1,e2) -> + Printf.fprintf o "(S_Mult %a %a)" pp_cone e1 pp_cone e2 + | Mc.S_Pos p -> + Printf.fprintf o "(S_Pos %a)%%positive" pp_z p + | Mc.S_Z -> + Printf.fprintf o "S_Z" in + pp_cone o e + + + + + let rec parse_op term = + let (i,c) = get_left_construct term in + match i with + | 1 -> Mc.OpEq + | 2 -> Mc.OpLe + | 3 -> Mc.OpGe + | 4 -> Mc.OpGt + | 5 -> Mc.OpLt + | i -> raise ParseError + + + let rec dump_op = function + | Mc.OpEq-> Lazy.force coq_OpEq + | Mc.OpNEq-> Lazy.force coq_OpNEq + | Mc.OpLe -> Lazy.force coq_OpLe + | Mc.OpGe -> Lazy.force coq_OpGe + | Mc.OpGt-> Lazy.force coq_OpGt + | Mc.OpLt-> Lazy.force coq_OpLt + + + + let pp_op o e= + match e with + | Mc.OpEq-> Printf.fprintf o "=" + | Mc.OpNEq-> Printf.fprintf o "<>" + | Mc.OpLe -> Printf.fprintf o "=<" + | Mc.OpGe -> Printf.fprintf o ">=" + | Mc.OpGt-> Printf.fprintf o ">" + | Mc.OpLt-> Printf.fprintf o "<" + + + + + let pp_cstr o {Mc.flhs = l ; Mc.fop = op ; Mc.frhs = r } = + Printf.fprintf o"(%a %a %a)" pp_expr l pp_op op pp_expr r + + let dump_cstr typ dump_constant {Mc.flhs = e1 ; Mc.fop = o ; Mc.frhs = e2} = + Term.mkApp(Lazy.force coq_Build, + [| typ; dump_expr typ dump_constant e1 ; + dump_op o ; + dump_expr typ dump_constant e2|]) + + + + let parse_zop (op,args) = + match kind_of_term op with + | Const x -> + (match Names.string_of_con x with + | "Coq.ZArith.BinInt#<>#Zgt" -> (Mc.OpGt, args.(0), args.(1)) + | "Coq.ZArith.BinInt#<>#Zge" -> (Mc.OpGe, args.(0), args.(1)) + | "Coq.ZArith.BinInt#<>#Zlt" -> (Mc.OpLt, args.(0), args.(1)) + | "Coq.ZArith.BinInt#<>#Zle" -> (Mc.OpLe, args.(0), args.(1)) + (*| "Coq.Init.Logic#<>#not" -> Mc.OpNEq (* for backward compat *)*) + | s -> raise ParseError + ) + | Ind(n,0) -> + (match Names.string_of_kn n with + | "Coq.Init.Logic#<>#eq" -> + if args.(0) <> Lazy.force coq_Z + then raise ParseError + else (Mc.OpEq, args.(1), args.(2)) + | _ -> raise ParseError) + | _ -> failwith "parse_zop" + + + let parse_rop (op,args) = + try + match kind_of_term op with + | Const x -> + (match Names.string_of_con x with + | "Coq.Reals.Rdefinitions#<>#Rgt" -> (Mc.OpGt, args.(0), args.(1)) + | "Coq.Reals.Rdefinitions#<>#Rge" -> (Mc.OpGe, args.(0), args.(1)) + | "Coq.Reals.Rdefinitions#<>#Rlt" -> (Mc.OpLt, args.(0), args.(1)) + | "Coq.Reals.Rdefinitions#<>#Rle" -> (Mc.OpLe, args.(0), args.(1)) + (*| "Coq.Init.Logic#<>#not"-> Mc.OpNEq (* for backward compat *)*) + | s -> raise ParseError + ) + | Ind(n,0) -> + (match Names.string_of_kn n with + | "Coq.Init.Logic#<>#eq" -> + (* if args.(0) <> Lazy.force coq_R + then raise ParseError + else*) (Mc.OpEq, args.(1), args.(2)) + | _ -> raise ParseError) + | _ -> failwith "parse_rop" + with x -> + (Pp.pp (Pp.str "parse_rop failure ") ; + Pp.pp (Printer.prterm op) ; Pp.pp_flush ()) + ; raise x + + + let parse_qop (op,args) = + ( + (match kind_of_term op with + | Const x -> + (match Names.string_of_con x with + | "Coq.QArith.QArith_base#<>#Qgt" -> Mc.OpGt + | "Coq.QArith.QArith_base#<>#Qge" -> Mc.OpGe + | "Coq.QArith.QArith_base#<>#Qlt" -> Mc.OpLt + | "Coq.QArith.QArith_base#<>#Qle" -> Mc.OpLe + | "Coq.QArith.QArith_base#<>#Qeq" -> Mc.OpEq + | s -> raise ParseError + ) + | _ -> failwith "parse_zop") , args.(0) , args.(1)) + + + module Env = + struct + type t = constr list + + let compute_rank_add env v = + let rec _add env n v = + match env with + | [] -> ([v],n) + | e::l -> + if eq_constr e v + then (env,n) + else + let (env,n) = _add l ( n+1) v in + (e::env,n) in + let (env, n) = _add env 1 v in + (env, CamlToCoq.idx n) + + + let empty = [] + + let elements env = env + + end + + + let is_constant t = (* This is an approx *) + match kind_of_term t with + | Construct(i,_) -> true + | _ -> false + + + type 'a op = + | Binop of ('a Mc.pExpr -> 'a Mc.pExpr -> 'a Mc.pExpr) + | Opp + | Power + | Ukn of string + + + let parse_expr parse_constant parse_exp ops_spec env term = + if debug + then (Pp.pp (Pp.str "parse_expr: "); + Pp.pp_flush ();Pp.pp (Printer.prterm term); Pp.pp_flush ()); + + let constant_or_variable env term = + try + ( Mc.PEc (parse_constant term) , env) + with ParseError -> + let (env,n) = Env.compute_rank_add env term in + (Mc.PEX n , env) in + + let rec parse_expr env term = + let combine env op (t1,t2) = + let (expr1,env) = parse_expr env t1 in + let (expr2,env) = parse_expr env t2 in + (op expr1 expr2,env) in + match kind_of_term term with + | App(t,args) -> + ( + match kind_of_term t with + | Const c -> + ( match ops_spec (Names.string_of_con c) with + | Binop f -> combine env f (args.(0),args.(1)) + | Opp -> let (expr,env) = parse_expr env args.(0) in + (Mc.PEopp expr, env) + | Power -> + let (expr,env) = parse_expr env args.(0) in + let exp = (parse_exp args.(1)) in + (Mc.PEpow(expr, exp) , env) + | Ukn s -> + if debug + then (Printf.printf "unknown op: %s\n" s; flush stdout;); + let (env,n) = Env.compute_rank_add env term in (Mc.PEX n, env) + ) + | _ -> constant_or_variable env term + ) + | _ -> constant_or_variable env term in + parse_expr env term + + +let zop_spec = function + | "Coq.ZArith.BinInt#<>#Zplus" -> Binop (fun x y -> Mc.PEadd(x,y)) + | "Coq.ZArith.BinInt#<>#Zminus" -> Binop (fun x y -> Mc.PEsub(x,y)) + | "Coq.ZArith.BinInt#<>#Zmult" -> Binop (fun x y -> Mc.PEmul (x,y)) + | "Coq.ZArith.BinInt#<>#Zopp" -> Opp + | "Coq.ZArith.Zpow_def#<>#Zpower" -> Power + | s -> Ukn s + +let qop_spec = function + | "Coq.QArith.QArith_base#<>#Qplus" -> Binop (fun x y -> Mc.PEadd(x,y)) + | "Coq.QArith.QArith_base#<>#Qminus" -> Binop (fun x y -> Mc.PEsub(x,y)) + | "Coq.QArith.QArith_base#<>#Qmult" -> Binop (fun x y -> Mc.PEmul (x,y)) + | "Coq.QArith.QArith_base#<>#Qopp" -> Opp + | "Coq.QArith.QArith_base#<>#Qpower" -> Power + | s -> Ukn s + +let rop_spec = function + | "Coq.Reals.Rdefinitions#<>#Rplus" -> Binop (fun x y -> Mc.PEadd(x,y)) + | "Coq.Reals.Rdefinitions#<>#Rminus" -> Binop (fun x y -> Mc.PEsub(x,y)) + | "Coq.Reals.Rdefinitions#<>#Rmult" -> Binop (fun x y -> Mc.PEmul (x,y)) + | "Coq.Reals.Rdefinitions#<>#Ropp" -> Opp + | "Coq.Reals.Rpow_def#<>#pow" -> Power + | s -> Ukn s + + + + + +let zconstant = parse_z +let qconstant = parse_q + + +let rconstant term = + if debug + then (Pp.pp_flush (); + Pp.pp (Pp.str "rconstant: "); + Pp.pp (Printer.prterm term); Pp.pp_flush ()); + match Term.kind_of_term term with + | Const x -> + (match Names.string_of_con x with + | "Coq.Reals.Rdefinitions#<>#R0" -> Mc.Z0 + | "Coq.Reals.Rdefinitions#<>#R1" -> Mc.Zpos Mc.XH + | _ -> raise ParseError + ) + | _ -> raise ParseError + + +let parse_zexpr = + parse_expr zconstant (fun x -> Mc.n_of_Z (parse_z x)) zop_spec +let parse_qexpr = + parse_expr qconstant (fun x -> Mc.n_of_Z (parse_z x)) qop_spec +let parse_rexpr = + parse_expr rconstant (fun x -> Mc.n_of_nat (parse_nat x)) rop_spec + + + let parse_arith parse_op parse_expr env cstr = + if debug + then (Pp.pp_flush (); + Pp.pp (Pp.str "parse_arith: "); + Pp.pp (Printer.prterm cstr); + Pp.pp_flush ()); + match kind_of_term cstr with + | App(op,args) -> + let (op,lhs,rhs) = parse_op (op,args) in + let (e1,env) = parse_expr env lhs in + let (e2,env) = parse_expr env rhs in + ({Mc.flhs = e1; Mc.fop = op;Mc.frhs = e2},env) + | _ -> failwith "error : parse_arith(2)" + + let parse_zarith = parse_arith parse_zop parse_zexpr + + let parse_qarith = parse_arith parse_qop parse_qexpr + + let parse_rarith = parse_arith parse_rop parse_rexpr + + + (* generic parsing of arithmetic expressions *) + + let rec parse_conj parse_arith env term = + match kind_of_term term with + | App(l,rst) -> + (match kind_of_term l with + | Ind (n,_) -> + ( match Names.string_of_kn n with + | "Coq.Init.Logic#<>#and" -> + let (e1,env) = parse_arith env rst.(0) in + let (e2,env) = parse_conj parse_arith env rst.(1) in + (Mc.Cons(e1,e2),env) + | _ -> (* This might be an equality *) + let (e,env) = parse_arith env term in + (Mc.Cons(e,Mc.Nil),env)) + | _ -> (* This is an arithmetic expression *) + let (e,env) = parse_arith env term in + (Mc.Cons(e,Mc.Nil),env)) + | _ -> failwith "parse_conj(2)" + + + + let rec f2f = function + | TT -> Mc.TT + | FF -> Mc.FF + | X _ -> Mc.X + | A (x,_) -> Mc.A x + | C (a,b,_) -> Mc.Cj(f2f a,f2f b) + | D (a,b,_) -> Mc.D(f2f a,f2f b) + | N (a,_) -> Mc.N(f2f a) + | I(a,b,_) -> Mc.I(f2f a,f2f b) + + let is_prop t = + match t with + | Names.Anonymous -> true (* Not quite right *) + | Names.Name x -> false + + let mkC f1 f2 = C(f1,f2,none) + let mkD f1 f2 = D(f1,f2,none) + let mkIff f1 f2 = C(I(f1,f2,none),I(f2,f2,none),none) + let mkI f1 f2 = I(f1,f2,none) + + let mkformula_binary g term f1 f2 = + match f1 , f2 with + | X _ , X _ -> X(term) + | _ -> g f1 f2 + + let parse_formula parse_atom env term = + let parse_atom env t = try let (at,env) = parse_atom env t in (A(at,none), env) with _ -> (X(t),env) in + + let rec xparse_formula env term = + match kind_of_term term with + | App(l,rst) -> + (match rst with + | [|a;b|] when l = Lazy.force coq_and -> + let f,env = xparse_formula env a in + let g,env = xparse_formula env b in + mkformula_binary mkC term f g,env + | [|a;b|] when l = Lazy.force coq_or -> + let f,env = xparse_formula env a in + let g,env = xparse_formula env b in + mkformula_binary mkD term f g,env + | [|a|] when l = Lazy.force coq_not -> + let (f,env) = xparse_formula env a in (N(f,none), env) + | [|a;b|] when l = Lazy.force coq_iff -> + let f,env = xparse_formula env a in + let g,env = xparse_formula env b in + mkformula_binary mkIff term f g,env + | _ -> parse_atom env term) + | Prod(typ,a,b) when not (Termops.dependent (mkRel 1) b) -> + let f,env = xparse_formula env a in + let g,env = xparse_formula env b in + mkformula_binary mkI term f g,env + | _ when term = Lazy.force coq_True -> (TT,env) + | _ when term = Lazy.force coq_False -> (FF,env) + | _ -> X(term),env in + xparse_formula env term + + let coq_TT = lazy + (gen_constant_in_modules "ZMicromega" + [["Coq" ; "micromega" ; "Tauto"];["Tauto"]] "TT") + let coq_FF = lazy + (gen_constant_in_modules "ZMicromega" + [["Coq" ; "micromega" ; "Tauto"];["Tauto"]] "FF") + let coq_And = lazy + (gen_constant_in_modules "ZMicromega" + [["Coq" ; "micromega" ; "Tauto"];["Tauto"]] "Cj") + let coq_Or = lazy + (gen_constant_in_modules "ZMicromega" + [["Coq" ; "micromega" ; "Tauto"];["Tauto"]] "D") + let coq_Neg = lazy + (gen_constant_in_modules "ZMicromega" + [["Coq" ; "micromega" ; "Tauto"];["Tauto"]] "N") + let coq_Atom = lazy + (gen_constant_in_modules "ZMicromega" + [["Coq" ; "micromega" ; "Tauto"];["Tauto"]] "A") + let coq_X = lazy + (gen_constant_in_modules "ZMicromega" + [["Coq" ; "micromega" ; "Tauto"];["Tauto"]] "X") + let coq_Impl = lazy + (gen_constant_in_modules "ZMicromega" + [["Coq" ; "micromega" ; "Tauto"];["Tauto"]] "I") + let coq_Formula = lazy + (gen_constant_in_modules "ZMicromega" + [["Coq" ; "micromega" ; "Tauto"];["Tauto"]] "BFormula") + + let dump_formula typ dump_atom f = + let rec xdump f = + match f with + | TT -> mkApp(Lazy.force coq_TT,[| typ|]) + | FF -> mkApp(Lazy.force coq_FF,[| typ|]) + | C(x,y,_) -> mkApp(Lazy.force coq_And,[| typ ; xdump x ; xdump y|]) + | D(x,y,_) -> mkApp(Lazy.force coq_Or,[| typ ; xdump x ; xdump y|]) + | I(x,y,_) -> mkApp(Lazy.force coq_Impl,[| typ ; xdump x ; xdump y|]) + | N(x,_) -> mkApp(Lazy.force coq_Neg,[| typ ; xdump x|]) + | A(x,_) -> mkApp(Lazy.force coq_Atom,[| typ ; dump_atom x|]) + | X(t) -> mkApp(Lazy.force coq_X,[| typ ; t|]) in + + xdump f + + + (* Backward compat *) + + let rec parse_concl parse_arith env term = + match kind_of_term term with + | Prod(_,expr,rst) -> (* a -> b *) + let (lhs,rhs,env) = parse_concl parse_arith env rst in + let (e,env) = parse_arith env expr in + (Mc.Cons(e,lhs),rhs,env) + | App(_,_) -> + let (conj, env) = parse_conj parse_arith env term in + (Mc.Nil,conj,env) + | Ind(n,_) -> + (match (Names.string_of_kn n) with + | "Coq.Init.Logic#<>#False" -> (Mc.Nil,Mc.Nil,env) + | s -> + print_string s ; flush stdout; + failwith "parse_concl") + | _ -> failwith "parse_concl" + + + let rec parse_hyps parse_arith env goal_hyps hyps = + match hyps with + | [] -> ([],goal_hyps,env) + | (i,t)::l -> + let (li,lt,env) = parse_hyps parse_arith env goal_hyps l in + try + let (c,env) = parse_arith env t in + (i::li, Mc.Cons(c,lt), env) + with x -> + (*(if debug then Printf.printf "parse_arith : %s\n" x);*) + (li,lt,env) + + + let parse_goal parse_arith env hyps term = + try + let (lhs,rhs,env) = parse_concl parse_arith env term in + let (li,lt,env) = parse_hyps parse_arith env lhs hyps in + (li,lt,rhs,env) + with Failure x -> raise ParseError + (* backward compat *) + + + (* ! reverse the list of bindings *) + let set l concl = + let rec _set acc = function + | [] -> acc + | (e::l) -> + let (name,expr,typ) = e in + _set (Term.mkNamedLetIn + (Names.id_of_string name) + expr typ acc) l in + _set concl l + + +end + +open M + + +let rec sig_of_cone = function + | Mc.S_In n -> [CoqToCaml.nat n] + | Mc.S_Ideal(e,w) -> sig_of_cone w + | Mc.S_Mult(w1,w2) -> + (sig_of_cone w1)@(sig_of_cone w2) + | Mc.S_Add(w1,w2) -> (sig_of_cone w1)@(sig_of_cone w2) + | _ -> [] + +let same_proof sg cl1 cl2 = + let cl1 = CoqToCaml.list (fun x -> x) cl1 in + let cl2 = CoqToCaml.list (fun x -> x) cl2 in + let rec xsame_proof sg = + match sg with + | [] -> true + | n::sg -> (try List.nth cl1 n = List.nth cl2 n with _ -> false) + && (xsame_proof sg ) in + xsame_proof sg + + + + +let tags_of_clause tgs wit clause = + let rec xtags tgs = function + | Mc.S_In n -> Names.Idset.union tgs + (snd (List.nth clause (CoqToCaml.nat n) )) + | Mc.S_Ideal(e,w) -> xtags tgs w + | Mc.S_Mult (w1,w2) | Mc.S_Add(w1,w2) -> xtags (xtags tgs w1) w2 + | _ -> tgs in + xtags tgs wit + +let tags_of_cnf wits cnf = + List.fold_left2 (fun acc w cl -> tags_of_clause acc w cl) + Names.Idset.empty wits cnf + + +let find_witness prover polys1 = + let l = CoqToCaml.list (fun x -> x) polys1 in + try_any prover l + +let rec witness prover l1 l2 = + match l2 with + | Micromega.Nil -> Some (Micromega.Nil) + | Micromega.Cons(e,l2) -> + match find_witness prover (Micromega.Cons( e,l1)) with + | None -> None + | Some w -> + (match witness prover l1 l2 with + | None -> None + | Some l -> Some (Micromega.Cons (w,l)) + ) + + +let rec apply_ids t ids = + match ids with + | [] -> t + | i::ids -> apply_ids (Term.mkApp(t,[| Term.mkVar i |])) ids + + +let coq_Node = lazy + (Coqlib.gen_constant_in_modules "VarMap" + [["Coq" ; "micromega" ; "VarMap"];["VarMap"]] "Node") +let coq_Leaf = lazy + (Coqlib.gen_constant_in_modules "VarMap" + [["Coq" ; "micromega" ; "VarMap"];["VarMap"]] "Leaf") +let coq_Empty = lazy + (Coqlib.gen_constant_in_modules "VarMap" + [["Coq" ; "micromega" ;"VarMap"];["VarMap"]] "Empty") + + +let btree_of_array typ a = + let size_of_a = Array.length a in + let semi_size_of_a = size_of_a lsr 1 in + let node = Lazy.force coq_Node + and leaf = Lazy.force coq_Leaf + and empty = Term.mkApp (Lazy.force coq_Empty, [| typ |]) in + let rec aux n = + if n > size_of_a + then empty + else if n > semi_size_of_a + then Term.mkApp (leaf, [| typ; a.(n-1) |]) + else Term.mkApp (node, [| typ; aux (2*n); a.(n-1); aux (2*n+1) |]) + in + aux 1 + +let btree_of_array typ a = + try + btree_of_array typ a + with x -> + failwith (Printf.sprintf "btree of array : %s" (Printexc.to_string x)) + +let dump_varmap typ env = + btree_of_array typ (Array.of_list env) + + +let rec pp_varmap o vm = + match vm with + | Mc.Empty -> output_string o "[]" + | Mc.Leaf z -> Printf.fprintf o "[%a]" pp_z z + | Mc.Node(l,z,r) -> Printf.fprintf o "[%a, %a, %a]" pp_varmap l pp_z z pp_varmap r + + + +let rec dump_proof_term = function + | Micromega.RatProof cone -> + Term.mkApp(Lazy.force coq_ratProof, [|dump_cone coq_Z dump_z cone|]) + | Micromega.CutProof(e,q,cone,prf) -> + Term.mkApp(Lazy.force coq_cutProof, + [| dump_expr (Lazy.force coq_Z) dump_z e ; + dump_q q ; + dump_cone coq_Z dump_z cone ; + dump_proof_term prf|]) + | Micromega.EnumProof( q1,e1,q2,c1,c2,prfs) -> + Term.mkApp (Lazy.force coq_enumProof, + [| dump_q q1 ; dump_expr (Lazy.force coq_Z) dump_z e1 ; dump_q q2; + dump_cone coq_Z dump_z c1 ; dump_cone coq_Z dump_z c2 ; + dump_list (Lazy.force coq_proofTerm) dump_proof_term prfs |]) + +let pp_q o q = Printf.fprintf o "%a/%a" pp_z q.Micromega.qnum pp_positive q.Micromega.qden + + +let rec pp_proof_term o = function + | Micromega.RatProof cone -> Printf.fprintf o "R[%a]" (pp_cone pp_z) cone + | Micromega.CutProof(e,q,_,p) -> failwith "not implemented" + | Micromega.EnumProof(q1,e1,q2,c1,c2,rst) -> + Printf.fprintf o "EP[%a,%a,%a,%a,%a,%a]" + pp_q q1 pp_expr e1 pp_q q2 (pp_cone pp_z) c1 (pp_cone pp_z) c2 + (pp_list "[" "]" pp_proof_term) rst + +let rec parse_hyps parse_arith env hyps = + match hyps with + | [] -> ([],env) + | (i,t)::l -> + let (lhyps,env) = parse_hyps parse_arith env l in + try + let (c,env) = parse_formula parse_arith env t in + ((i,c)::lhyps, env) + with _ -> (lhyps,env) + (*(if debug then Printf.printf "parse_arith : %s\n" x);*) + + +exception ParseError + +let parse_goal parse_arith env hyps term = + (* try*) + let (f,env) = parse_formula parse_arith env term in + let (lhyps,env) = parse_hyps parse_arith env hyps in + (lhyps,f,env) + (* with Failure x -> raise ParseError*) + + +type ('a, 'b) domain_spec = { + typ : Term.constr; (* Z, Q , R *) + coeff : Term.constr ; (* Z, Q *) + dump_coeff : 'a -> Term.constr ; + proof_typ : Term.constr ; + dump_proof : 'b -> Term.constr +} + +let zz_domain_spec = lazy { + typ = Lazy.force coq_Z; + coeff = Lazy.force coq_Z; + dump_coeff = dump_z ; + proof_typ = Lazy.force coq_proofTerm ; + dump_proof = dump_proof_term +} + +let qq_domain_spec = lazy { + typ = Lazy.force coq_Q; + coeff = Lazy.force coq_Q; + dump_coeff = dump_q ; + proof_typ = Lazy.force coq_QWitness ; + dump_proof = dump_cone coq_Q dump_q +} + +let rz_domain_spec = lazy { + typ = Lazy.force coq_R; + coeff = Lazy.force coq_Z; + dump_coeff = dump_z; + proof_typ = Lazy.force coq_ZWitness ; + dump_proof = dump_cone coq_Z dump_z +} + + + + +let micromega_order_change spec cert cert_typ env ff gl = + let formula_typ = (Term.mkApp( Lazy.force coq_Cstr,[| spec.coeff|])) in + + let ff = dump_formula formula_typ (dump_cstr spec.coeff spec.dump_coeff) ff in + let vm = dump_varmap ( spec.typ) env in + Tactics.change_in_concl None + (set + [ + ("__ff", ff, Term.mkApp(Lazy.force coq_Formula ,[| formula_typ |])); + ("__varmap", vm , Term.mkApp + (Coqlib.gen_constant_in_modules "VarMap" + [["Coq" ; "micromega" ; "VarMap"];["VarMap"]] "t", [| spec.typ|])); + ("__wit", cert,cert_typ) + ] + (Tacmach.pf_concl gl ) + + ) + gl + + +let detect_duplicates cnf wit = + let cnf = CoqToCaml.list (fun x -> x) cnf in + let wit = CoqToCaml.list (fun x -> x) wit in + + let rec xdup cnf wit = + match wit with + | [] -> [] + | w :: wit -> + let sg = sig_of_cone w in + match cnf with + | [] -> [] + | e::cnf -> + let (dups,cnf) = (List.partition (fun x -> same_proof sg e x) cnf) in + dups@(xdup cnf wit) in + xdup cnf wit + +let find_witness prover polys1 = + try_any prover polys1 + + +let witness_list_with_tags prover l = + + let rec xwitness_list l = + match l with + | [] -> Some([]) + | e::l -> + match find_witness prover (List.map fst e) with + | None -> None + | Some w -> + (match xwitness_list l with + | None -> None + | Some l -> Some (w::l) + ) in + xwitness_list l + +let witness_list_without_tags prover l = + + let rec xwitness_list l = + match l with + | [] -> Some([]) + | e::l -> + match find_witness prover e with + | None -> None + | Some w -> + (match xwitness_list l with + | None -> None + | Some l -> Some (w::l) + ) in + xwitness_list l + +let witness_list prover l = + let rec xwitness_list l = + match l with + | Micromega.Nil -> Some(Micromega.Nil) + | Micromega.Cons(e,l) -> + match find_witness prover e with + | None -> None + | Some w -> + (match xwitness_list l with + | None -> None + | Some l -> Some (Micromega.Cons(w,l)) + ) in + xwitness_list l + + + + +let is_singleton = function [] -> true | [e] -> true | _ -> false + + +let micromega_tauto negate normalise spec prover env polys1 polys2 gl = + let spec = Lazy.force spec in + let (ff,ids) = + List.fold_right + (fun (id,f) (cc,ids) -> + match f with + X _ -> (cc,ids) + | _ -> (I(tag_formula (Names.Name id) f,cc,none), id::ids)) + polys1 (polys2,[]) in + + let cnf_ff = cnf negate normalise ff in + + if debug then + (Pp.pp (Pp.str "Formula....\n") ; + let formula_typ = (Term.mkApp( Lazy.force coq_Cstr,[| spec.coeff|])) in + let ff = dump_formula formula_typ + (dump_cstr spec.typ spec.dump_coeff) ff in + Pp.pp (Printer.prterm ff) ; Pp.pp_flush ()) ; + + match witness_list_without_tags prover cnf_ff with + | None -> Tacticals.tclFAIL 0 (Pp.str "Cannot find witness") gl + | Some res -> (*Printf.printf "\nList %i" (List.length res); *) + let (ff,res,ids) = (ff,res,List.map Term.mkVar ids) in + let res' = dump_ml_list (spec.proof_typ) spec.dump_proof res in + (Tacticals.tclTHENSEQ + [ + Tactics.generalize ids; + micromega_order_change spec res' + (Term.mkApp(Lazy.force coq_list,[| spec.proof_typ|])) env ff ; + ]) gl + + +let micromega_gen parse_arith negate normalise spec prover gl = + let concl = Tacmach.pf_concl gl in + let hyps = Tacmach.pf_hyps_types gl in + try + let (hyps,concl,env) = parse_goal parse_arith Env.empty hyps concl in + let env = Env.elements env in + micromega_tauto negate normalise spec prover env hyps concl gl + with + | Failure x -> flush stdout ; Pp.pp_flush () ; + Tacticals.tclFAIL 0 (Pp.str x) gl + | ParseError -> Tacticals.tclFAIL 0 (Pp.str "Bad logical fragment") gl + + +let lift_ratproof prover l = + match prover l with + | None -> None + | Some c -> Some (Mc.RatProof c) + + +type csdpcert = Certificate.Mc.z Certificate.Mc.coneMember option +type micromega_polys = (Micromega.z Mc.pExpr, Mc.op1) Micromega.prod list +type provername = string * int option + +let call_csdpcert provername poly = + let tmp_to,ch_to = Filename.open_temp_file "csdpcert" ".in" in + let tmp_from = Filename.temp_file "csdpcert" ".out" in + output_value ch_to (provername,poly : provername * micromega_polys); + close_out ch_to; + let cmdname = + Filename.concat Coq_config.bindir + ("csdpcert" ^ Coq_config.exec_extension) in + let c = Sys.command (cmdname ^" "^ tmp_to ^" "^ tmp_from) in + (try Sys.remove tmp_to with _ -> ()); + if c <> 0 then Util.error ("Failed to call csdp certificate generator"); + let ch_from = open_in tmp_from in + let cert = (input_value ch_from : csdpcert) in + close_in ch_from; Sys.remove tmp_from; + cert + +let omicron gl = + micromega_gen parse_zarith Mc.negate Mc.normalise zz_domain_spec + [lift_ratproof + (Certificate.linear_prover Certificate.z_spec), "fourier refutation" ] gl + + +let qomicron gl = + micromega_gen parse_qarith Mc.cnf_negate Mc.cnf_normalise qq_domain_spec + [ Certificate.linear_prover Certificate.q_spec, "fourier refutation" ] gl + +let romicron gl = + micromega_gen parse_rarith Mc.cnf_negate Mc.cnf_normalise rz_domain_spec + [ Certificate.linear_prover Certificate.z_spec, "fourier refutation" ] gl + + +let rmicromega i gl = + micromega_gen parse_rarith Mc.negate Mc.normalise rz_domain_spec + [ call_csdpcert ("real_nonlinear_prover", Some i), "fourier refutation" ] gl + + +let micromega i gl = + micromega_gen parse_zarith Mc.negate Mc.normalise zz_domain_spec + [lift_ratproof (call_csdpcert ("real_nonlinear_prover",Some i)), + "fourier refutation" ] gl + + +let sos gl = + micromega_gen parse_zarith Mc.negate Mc.normalise zz_domain_spec + [lift_ratproof (call_csdpcert ("pure_sos", None)), "pure sos refutation"] gl + +let zomicron gl = + micromega_gen parse_zarith Mc.negate Mc.normalise zz_domain_spec + [Certificate.zlinear_prover, "zprover"] gl diff --git a/contrib/micromega/csdpcert.ml b/contrib/micromega/csdpcert.ml new file mode 100644 index 00000000..cfaf6ae1 --- /dev/null +++ b/contrib/micromega/csdpcert.ml @@ -0,0 +1,333 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* *) +(* Micromega: A reflexive tactic using the Positivstellensatz *) +(* *) +(* Frédéric Besson (Irisa/Inria) 2006-2008 *) +(* *) +(************************************************************************) + +open Big_int +open Num +open Sos + +module Mc = Micromega +module Ml2C = Mutils.CamlToCoq +module C2Ml = Mutils.CoqToCaml + +let debug = false + +module M = +struct + open Mc + + let rec expr_to_term = function + | PEc z -> Const (Big_int (C2Ml.z_big_int z)) + | PEX v -> Var ("x"^(string_of_int (C2Ml.index v))) + | PEmul(p1,p2) -> + let p1 = expr_to_term p1 in + let p2 = expr_to_term p2 in + let res = Mul(p1,p2) in res + + | PEadd(p1,p2) -> Add(expr_to_term p1, expr_to_term p2) + | PEsub(p1,p2) -> Sub(expr_to_term p1, expr_to_term p2) + | PEpow(p,n) -> Pow(expr_to_term p , C2Ml.n n) + | PEopp p -> Opp (expr_to_term p) + + + let rec term_to_expr = function + | Const n -> PEc (Ml2C.bigint (big_int_of_num n)) + | Zero -> PEc ( Z0) + | Var s -> PEX (Ml2C.index + (int_of_string (String.sub s 1 (String.length s - 1)))) + | Mul(p1,p2) -> PEmul(term_to_expr p1, term_to_expr p2) + | Add(p1,p2) -> PEadd(term_to_expr p1, term_to_expr p2) + | Opp p -> PEopp (term_to_expr p) + | Pow(t,n) -> PEpow (term_to_expr t,Ml2C.n n) + | Sub(t1,t2) -> PEsub (term_to_expr t1, term_to_expr t2) + | _ -> failwith "term_to_expr: not implemented" + + let term_to_expr e = + let e' = term_to_expr e in + if debug + then Printf.printf "term_to_expr : %s - %s\n" + (string_of_poly (poly_of_term e)) + (string_of_poly (poly_of_term (expr_to_term e'))); + e' + +end +open M + +open List +open Mutils + +let rec scale_term t = + match t with + | Zero -> unit_big_int , Zero + | Const n -> (denominator n) , Const (Big_int (numerator n)) + | Var n -> unit_big_int , Var n + | Inv _ -> failwith "scale_term : not implemented" + | Opp t -> let s, t = scale_term t in s, Opp t + | Add(t1,t2) -> let s1,y1 = scale_term t1 and s2,y2 = scale_term t2 in + let g = gcd_big_int s1 s2 in + let s1' = div_big_int s1 g in + let s2' = div_big_int s2 g in + let e = mult_big_int g (mult_big_int s1' s2') in + if (compare_big_int e unit_big_int) = 0 + then (unit_big_int, Add (y1,y2)) + else e, Add (Mul(Const (Big_int s2'), y1), + Mul (Const (Big_int s1'), y2)) + | Sub _ -> failwith "scale term: not implemented" + | Mul(y,z) -> let s1,y1 = scale_term y and s2,y2 = scale_term z in + mult_big_int s1 s2 , Mul (y1, y2) + | Pow(t,n) -> let s,t = scale_term t in + power_big_int_positive_int s n , Pow(t,n) + | _ -> failwith "scale_term : not implemented" + +let scale_term t = + let (s,t') = scale_term t in + s,t' + + + + +let rec scale_certificate pos = match pos with + | Axiom_eq i -> unit_big_int , Axiom_eq i + | Axiom_le i -> unit_big_int , Axiom_le i + | Axiom_lt i -> unit_big_int , Axiom_lt i + | Monoid l -> unit_big_int , Monoid l + | Rational_eq n -> (denominator n) , Rational_eq (Big_int (numerator n)) + | Rational_le n -> (denominator n) , Rational_le (Big_int (numerator n)) + | Rational_lt n -> (denominator n) , Rational_lt (Big_int (numerator n)) + | Square t -> let s,t' = scale_term t in + mult_big_int s s , Square t' + | Eqmul (t, y) -> let s1,y1 = scale_term t and s2,y2 = scale_certificate y in + mult_big_int s1 s2 , Eqmul (y1,y2) + | Sum (y, z) -> let s1,y1 = scale_certificate y + and s2,y2 = scale_certificate z in + let g = gcd_big_int s1 s2 in + let s1' = div_big_int s1 g in + let s2' = div_big_int s2 g in + mult_big_int g (mult_big_int s1' s2'), + Sum (Product(Rational_le (Big_int s2'), y1), + Product (Rational_le (Big_int s1'), y2)) + | Product (y, z) -> + let s1,y1 = scale_certificate y and s2,y2 = scale_certificate z in + mult_big_int s1 s2 , Product (y1,y2) + + +let is_eq = function Mc.Equal -> true | _ -> false +let is_le = function Mc.NonStrict -> true | _ -> false +let is_lt = function Mc.Strict -> true | _ -> false + +let get_index_of_ith_match f i l = + let rec get j res l = + match l with + | [] -> failwith "bad index" + | e::l -> if f e + then + (if j = i then res else get (j+1) (res+1) l ) + else get j (res+1) l in + get 0 0 l + + +let cert_of_pos eq le lt ll l pos = + let s,pos = (scale_certificate pos) in + let rec _cert_of_pos = function + Axiom_eq i -> let idx = get_index_of_ith_match is_eq i l in + Mc.S_In (Ml2C.nat idx) + | Axiom_le i -> let idx = get_index_of_ith_match is_le i l in + Mc.S_In (Ml2C.nat idx) + | Axiom_lt i -> let idx = get_index_of_ith_match is_lt i l in + Mc.S_In (Ml2C.nat idx) + | Monoid l -> Mc.S_Monoid (Ml2C.list Ml2C.nat l) + | Rational_eq n | Rational_le n | Rational_lt n -> + if compare_num n (Int 0) = 0 then Mc.S_Z else + Mc.S_Pos (Ml2C.bigint (big_int_of_num n)) + | Square t -> Mc.S_Square (term_to_expr t) + | Eqmul (t, y) -> Mc.S_Ideal(term_to_expr t, _cert_of_pos y) + | Sum (y, z) -> Mc.S_Add (_cert_of_pos y, _cert_of_pos z) + | Product (y, z) -> Mc.S_Mult (_cert_of_pos y, _cert_of_pos z) in + s, Certificate.simplify_cone Certificate.z_spec (_cert_of_pos pos) + + +let term_of_cert l pos = + let l = List.map fst' l in + let rec _cert_of_pos = function + | Mc.S_In i -> expr_to_term (List.nth l (C2Ml.nat i)) + | Mc.S_Pos p -> Const (C2Ml.num p) + | Mc.S_Z -> Const (Int 0) + | Mc.S_Square t -> Mul(expr_to_term t, expr_to_term t) + | Mc.S_Monoid m -> List.fold_right + (fun x m -> Mul (expr_to_term (List.nth l (C2Ml.nat x)),m)) + (C2Ml.list (fun x -> x) m) (Const (Int 1)) + | Mc.S_Ideal (t, y) -> Mul(expr_to_term t, _cert_of_pos y) + | Mc.S_Add (y, z) -> Add (_cert_of_pos y, _cert_of_pos z) + | Mc.S_Mult (y, z) -> Mul (_cert_of_pos y, _cert_of_pos z) in + (_cert_of_pos pos) + +let rec canonical_sum_to_string = function s -> failwith "not implemented" + +let print_canonical_sum m = Format.print_string (canonical_sum_to_string m) + +let print_list_term l = + print_string "print_list_term\n"; + List.iter (fun (Mc.Pair(e,k)) -> Printf.printf "q: %s %s ;" + (string_of_poly (poly_of_term (expr_to_term e))) + (match k with + Mc.Equal -> "= " + | Mc.Strict -> "> " + | Mc.NonStrict -> ">= " + | _ -> failwith "not_implemented")) l ; + print_string "\n" + + +let partition_expr l = + let rec f i = function + | [] -> ([],[],[]) + | Mc.Pair(e,k)::l -> + let (eq,ge,neq) = f (i+1) l in + match k with + | Mc.Equal -> ((e,i)::eq,ge,neq) + | Mc.NonStrict -> (eq,(e,Axiom_le i)::ge,neq) + | Mc.Strict -> (* e > 0 == e >= 0 /\ e <> 0 *) + (eq, (e,Axiom_lt i)::ge,(e,Axiom_lt i)::neq) + | Mc.NonEqual -> (eq,ge,(e,Axiom_eq i)::neq) + (* Not quite sure -- Coq interface has changed *) + in f 0 l + + +let rec sets_of_list l = + match l with + | [] -> [[]] + | e::l -> let s = sets_of_list l in + s@(List.map (fun s0 -> e::s0) s) + +let cert_of_pos pos = + let s,pos = (scale_certificate pos) in + let rec _cert_of_pos = function + Axiom_eq i -> Mc.S_In (Ml2C.nat i) + | Axiom_le i -> Mc.S_In (Ml2C.nat i) + | Axiom_lt i -> Mc.S_In (Ml2C.nat i) + | Monoid l -> Mc.S_Monoid (Ml2C.list Ml2C.nat l) + | Rational_eq n | Rational_le n | Rational_lt n -> + if compare_num n (Int 0) = 0 then Mc.S_Z else + Mc.S_Pos (Ml2C.bigint (big_int_of_num n)) + | Square t -> Mc.S_Square (term_to_expr t) + | Eqmul (t, y) -> Mc.S_Ideal(term_to_expr t, _cert_of_pos y) + | Sum (y, z) -> Mc.S_Add (_cert_of_pos y, _cert_of_pos z) + | Product (y, z) -> Mc.S_Mult (_cert_of_pos y, _cert_of_pos z) in + s, Certificate.simplify_cone Certificate.z_spec (_cert_of_pos pos) + +(* The exploration is probably not complete - for simple cases, it works... *) +let real_nonlinear_prover d l = + try + let (eq,ge,neq) = partition_expr l in + + let rec elim_const = function + [] -> [] + | (x,y)::l -> let p = poly_of_term (expr_to_term x) in + if poly_isconst p + then elim_const l + else (p,y)::(elim_const l) in + + let eq = elim_const eq in + let peq = List.map fst eq in + + let pge = List.map + (fun (e,psatz) -> poly_of_term (expr_to_term e),psatz) ge in + + let monoids = List.map (fun m -> (List.fold_right (fun (p,kd) y -> + let p = poly_of_term (expr_to_term p) in + match kd with + | Axiom_lt i -> poly_mul p y + | Axiom_eq i -> poly_mul (poly_pow p 2) y + | _ -> failwith "monoids") m (poly_const (Int 1)) , map snd m)) + (sets_of_list neq) in + + let (cert_ideal, cert_cone,monoid) = deepen_until d (fun d -> + list_try_find (fun m -> let (ci,cc) = + real_positivnullstellensatz_general false d peq pge (poly_neg (fst m) ) in + (ci,cc,snd m)) monoids) 0 in + + let proofs_ideal = map2 (fun q i -> Eqmul(term_of_poly q,Axiom_eq i)) + cert_ideal (List.map snd eq) in + + let proofs_cone = map term_of_sos cert_cone in + + let proof_ne = + let (neq , lt) = List.partition + (function Axiom_eq _ -> true | _ -> false ) monoid in + let sq = match + (List.map (function Axiom_eq i -> i | _ -> failwith "error") neq) + with + | [] -> Rational_lt (Int 1) + | l -> Monoid l in + List.fold_right (fun x y -> Product(x,y)) lt sq in + + let proof = list_fold_right_elements + (fun s t -> Sum(s,t)) (proof_ne :: proofs_ideal @ proofs_cone) in + + let s,proof' = scale_certificate proof in + let cert = snd (cert_of_pos proof') in + if debug + then Printf.printf "cert poly : %s\n" + (string_of_poly (poly_of_term (term_of_cert l cert))); + match Mc.zWeakChecker (Ml2C.list (fun x -> x) l) cert with + | Mc.True -> Some cert + | Mc.False -> (print_string "buggy certificate" ; flush stdout) ;None + with + | Sos.TooDeep -> None + + +(* This is somewhat buggy, over Z, strict inequality vanish... *) +let pure_sos l = + (* If there is no strict inequality, + I should nonetheless be able to try something - over Z > is equivalent to -1 >= *) + try + let l = List.combine l (interval 0 (length l -1)) in + let (lt,i) = try (List.find (fun (x,_) -> snd' x = Mc.Strict) l) + with Not_found -> List.hd l in + let plt = poly_neg (poly_of_term (expr_to_term (fst' lt))) in + let (n,polys) = sumofsquares plt in (* n * (ci * pi^2) *) + let pos = Product (Rational_lt n, + List.fold_right (fun (c,p) rst -> Sum (Product (Rational_lt c, Square + (term_of_poly p)), rst)) + polys (Rational_lt (Int 0))) in + let proof = Sum(Axiom_lt i, pos) in + let s,proof' = scale_certificate proof in + let cert = snd (cert_of_pos proof') in + Some cert + with + | Not_found -> (* This is no strict inequality *) None + | x -> None + + +type micromega_polys = (Micromega.z Mc.pExpr, Mc.op1) Micromega.prod list +type csdp_certificate = Certificate.Mc.z Certificate.Mc.coneMember option +type provername = string * int option + +let main () = + if Array.length Sys.argv <> 3 then + (Printf.printf "Usage: csdpcert inputfile outputfile\n"; exit 1); + let input_file = Sys.argv.(1) in + let output_file = Sys.argv.(2) in + let inch = open_in input_file in + let (prover,poly) = (input_value inch : provername * micromega_polys) in + close_in inch; + let cert = + match prover with + | "real_nonlinear_prover", Some d -> real_nonlinear_prover d poly + | "pure_sos", None -> pure_sos poly + | prover, _ -> (Printf.printf "unknown prover: %s\n" prover; exit 1) in + let outch = open_out output_file in + output_value outch (cert:csdp_certificate); + close_out outch; + exit 0;; + +let _ = main () in () diff --git a/contrib/micromega/g_micromega.ml4 b/contrib/micromega/g_micromega.ml4 new file mode 100644 index 00000000..259b5d4b --- /dev/null +++ b/contrib/micromega/g_micromega.ml4 @@ -0,0 +1,59 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* *) +(* Micromega: A reflexive tactic using the Positivstellensatz *) +(* *) +(* Frédéric Besson (Irisa/Inria) 2006-2008 *) +(* *) +(************************************************************************) + +(*i camlp4deps: "parsing/grammar.cma" i*) + +(* $Id: g_micromega.ml4 10947 2008-05-19 19:10:40Z herbelin $ *) + +open Quote +open Ring +open Mutils +open Rawterm +open Util + +let out_arg = function + | ArgVar _ -> anomaly "Unevaluated or_var variable" + | ArgArg x -> x + +TACTIC EXTEND Micromega +| [ "micromegap" int_or_var(i) ] -> [ Coq_micromega.micromega (out_arg i) ] +| [ "micromegap" ] -> [ Coq_micromega.micromega (-1) ] +END + +TACTIC EXTEND Sos +[ "sosp" ] -> [ Coq_micromega.sos] +END + + +TACTIC EXTEND Omicron +[ "omicronp" ] -> [ Coq_micromega.omicron] +END + +TACTIC EXTEND QOmicron +[ "qomicronp" ] -> [ Coq_micromega.qomicron] +END + + +TACTIC EXTEND ZOmicron +[ "zomicronp" ] -> [ Coq_micromega.zomicron] +END + +TACTIC EXTEND ROmicron +[ "romicronp" ] -> [ Coq_micromega.romicron] +END + +TACTIC EXTEND RMicromega +| [ "rmicromegap" int_or_var(i) ] -> [ Coq_micromega.rmicromega (out_arg i) ] +| [ "rmicromegap" ] -> [ Coq_micromega.rmicromega (-1) ] +END diff --git a/contrib/micromega/mfourier.ml b/contrib/micromega/mfourier.ml new file mode 100644 index 00000000..415d3a3e --- /dev/null +++ b/contrib/micromega/mfourier.ml @@ -0,0 +1,667 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* *) +(* Micromega: A reflexive tactic using the Positivstellensatz *) +(* *) +(* Frédéric Besson (Irisa/Inria) 2006-2008 *) +(* *) +(************************************************************************) + +(* Yet another implementation of Fourier *) +open Num + +module Cmp = + (* How to compare pairs, lists ... *) +struct + let rec compare_lexical l = + match l with + | [] -> 0 (* Equal *) + | f::l -> + let cmp = f () in + if cmp = 0 then compare_lexical l else cmp + + let rec compare_list cmp l1 l2 = + match l1 , l2 with + | [] , [] -> 0 + | [] , _ -> -1 + | _ , [] -> 1 + | e1::l1 , e2::l2 -> + let c = cmp e1 e2 in + if c = 0 then compare_list cmp l1 l2 else c + + let hash_list hash l = + let rec xhash res l = + match l with + | [] -> res + | e::l -> xhash ((hash e) lxor res) l in + xhash (Hashtbl.hash []) l + +end + +module Interval = +struct + (** The type of intervals. **) + type intrvl = Empty | Point of num | Itv of num option * num option + + (** + Different intervals can denote the same set of variables e.g., + Point n && Itv (Some n, Some n) + Itv (Some x) (Some y) && Empty if x > y + see the 'belongs_to' function. + **) + + (* The set of numerics that belong to an interval *) + let belongs_to n = function + | Empty -> false + | Point x -> n =/ x + | Itv(Some x, Some y) -> x <=/ n && n <=/ y + | Itv(None,Some y) -> n <=/ y + | Itv(Some x,None) -> x <=/ n + | Itv(None,None) -> true + + let string_of_bound = function + | None -> "oo" + | Some n -> Printf.sprintf "Bd(%s)" (string_of_num n) + + let string_of_intrvl = function + | Empty -> "[]" + | Point n -> Printf.sprintf "[%s]" (string_of_num n) + | Itv(bd1,bd2) -> + Printf.sprintf "[%s,%s]" (string_of_bound bd1) (string_of_bound bd2) + + let pick_closed_to_zero = function + | Empty -> None + | Point n -> Some n + | Itv(None,None) -> Some (Int 0) + | Itv(None,Some i) -> + Some (if (Int 0) <=/ (floor_num i) then Int 0 else floor_num i) + | Itv(Some i,None) -> + Some (if i <=/ (Int 0) then Int 0 else ceiling_num i) + | Itv(Some i,Some j) -> + Some ( + if i <=/ Int 0 && Int 0 <=/ j + then Int 0 + else if ceiling_num i <=/ floor_num j + then ceiling_num i (* why not *) else i) + + type status = + | O | Qonly | Z | Q + + let interval_kind = function + | Empty -> O + | Point n -> if ceiling_num n =/ n then Z else Qonly + | Itv(None,None) -> Z + | Itv(None,Some i) -> if ceiling_num i <>/ i then Q else Z + | Itv(Some i,None) -> if ceiling_num i <>/ i then Q else Z + | Itv(Some i,Some j) -> + if ceiling_num i <>/ i or floor_num j <>/ j then Q else Z + + let empty_z = function + | Empty -> true + | Point n -> ceiling_num n <>/ n + | Itv(None,None) | Itv(None,Some _) | Itv(Some _,None) -> false + | Itv(Some i,Some j) -> ceiling_num i >/ floor_num j + + + let normalise b1 b2 = + match b1 , b2 with + | Some i , Some j -> + (match compare_num i j with + | 1 -> Empty + | 0 -> Point i + | _ -> Itv(b1,b2) + ) + | _ -> Itv(b1,b2) + + + + let min x y = + match x , y with + | None , x | x , None -> x + | Some i , Some j -> Some (min_num i j) + + let max x y = + match x , y with + | None , x | x , None -> x + | Some i , Some j -> Some (max_num i j) + + let inter i1 i2 = + match i1,i2 with + | Empty , _ -> Empty + | _ , Empty -> Empty + | Point n , Point m -> if n =/ m then i1 else Empty + | Point n , Itv (mn,mx) | Itv (mn,mx) , Point n-> + if (match mn with + | None -> true + | Some mn -> mn <=/ n) && + (match mx with + | None -> true + | Some mx -> n <=/ mx) then Point n else Empty + | Itv (min1,max1) , Itv (min2,max2) -> + let bmin = max min1 min2 + and bmax = min max1 max2 in + normalise bmin bmax + + (* a.x >= b*) + let bound_of_constraint (a,b) = + match compare_num a (Int 0) with + | 0 -> + if compare_num b (Int 0) = 1 + then Empty + (*actually this is a contradiction failwith "bound_of_constraint" *) + else Itv (None,None) + | 1 -> Itv (Some (div_num b a),None) + | -1 -> Itv (None, Some (div_num b a)) + | x -> failwith "bound_of_constraint(2)" + + + let bounded x = + match x with + | Itv(None,_) | Itv(_,None) -> false + | _ -> true + + + let range = function + | Empty -> Some (Int 0) + | Point n -> Some (Int (if ceiling_num n =/ n then 1 else 0)) + | Itv(None,_) | Itv(_,None)-> None + | Itv(Some i,Some j) -> Some (floor_num j -/ceiling_num i +/ (Int 1)) + + (* Returns the interval of smallest range *) + let smaller_itv i1 i2 = + match range i1 , range i2 with + | None , _ -> false + | _ , None -> true + | Some i , Some j -> i <=/ j + +end +open Interval + +(* A set of constraints *) +module Sys(V:Vector.S) (* : Vector.SystemS with module Vect = V*) = +struct + + module Vect = V + + module Cstr = Vector.Cstr(V) + open Cstr + + + module CMap = Map.Make( + struct + type t = Vect.t + let compare = Vect.compare + end) + + module CstrBag = + struct + + type mut_itv = { mutable itv : intrvl} + + type t = mut_itv CMap.t + + exception Contradiction + + let cstr_to_itv cstr = + let (n,l) = V.normalise cstr.coeffs in + if n =/ (Int 0) + then (Vect.null, bound_of_constraint (Int 0,cstr.cst)) (* Might be empty *) + else + match cstr.op with + | Eq -> let n = cstr.cst // n in (l, Point n) + | Ge -> + match compare_num n (Int 0) with + | 0 -> failwith "intrvl_of_constraint" + | 1 -> (l,Itv (Some (cstr.cst // n), None)) + | -1 -> (l, Itv(None,Some (cstr.cst // n))) + | _ -> failwith "cstr_to_itv" + + + let empty = CMap.empty + + + + + let is_empty = CMap.is_empty + + let find_vect v bag = + try + (bag,CMap.find v bag) + with Not_found -> let x = { itv = Itv(None,None)} in (CMap.add v x bag ,x) + + + let add (v,b) bag = + match b with + | Empty -> raise Contradiction + | Itv(None,None) -> bag + | _ -> + let (bag,intrl) = find_vect v bag in + match inter b intrl.itv with + | Empty -> raise Contradiction + | itv -> intrl.itv <- itv ; bag + + exception Found of cstr + + let find_equation bag = + try + CMap.fold (fun v i () -> + match i.itv with + | Point n -> let e = {coeffs = v ; op = Eq ; cst = n} + in raise (Found e) + | _ -> () ) bag () ; None + with Found c -> Some c + + + let fold f bag acc = + CMap.fold (fun v itv acc -> + match itv.itv with + | Empty | Itv(None,None) -> failwith "fold Empty" + | Itv(None ,Some i) -> + f {coeffs = V.mul (Int (-1)) v ; op = Ge ; cst = minus_num i} acc + | Point n -> f {coeffs = v ; op = Eq ; cst = n} acc + | Itv(x,y) -> + (match x with + | None -> (fun x -> x) + | Some i -> f {coeffs = v ; op = Ge ; cst = i}) + (match y with + | None -> acc + | Some i -> + f {coeffs = V.mul (Int (-1)) v ; op = Ge ; cst = minus_num i} acc + ) ) bag acc + + + let remove l _ = failwith "remove:Not implemented" + + module Map = + Map.Make( + struct + type t = int + let compare : int -> int -> int = Pervasives.compare + end) + + let split f (t:t) = + let res = + fold (fun e m -> let i = f e in + Map.add i (add (cstr_to_itv e) + (try Map.find i m with + Not_found -> empty)) m) t Map.empty in + (fun i -> try Map.find i res with Not_found -> empty) + + type map = (int list * int list) Map.t + + + let status (b:t) = + let _ , map = fold (fun c ( (idx:int),(res: map)) -> + ( idx + 1, + List.fold_left (fun (res:map) (pos,s) -> + let (lp,ln) = try Map.find pos res with Not_found -> ([],[]) in + match s with + | Vect.Pos -> Map.add pos (idx::lp,ln) res + | Vect.Neg -> + Map.add pos (lp, idx::ln) res) res + (Vect.status c.coeffs))) b (0,Map.empty) in + Map.fold (fun k e res -> (k,e)::res) map [] + + + type it = num CMap.t + + let iterator x = x + + let element it = failwith "element:Not implemented" + + end +end + +module Fourier(Vect : Vector.S) = +struct + module Vect = Vect + module Sys = Sys( Vect) + module Cstr = Sys.Cstr + module Bag = Sys.CstrBag + + open Cstr + open Sys + + let debug = false + + let print_bag msg b = + print_endline msg; + CstrBag.fold (fun e () -> print_endline (Cstr.string_of_cstr e)) b () + + let print_bag_file file msg b = + let f = open_out file in + output_string f msg; + CstrBag.fold (fun e () -> + Printf.fprintf f "%s\n" (Cstr.string_of_cstr e)) b () + + + (* A system with only inequations -- + *) + let partition i m = + let splitter cstr = compare_num (Vect.get i cstr.coeffs ) (Int 0) in + let split = CstrBag.split splitter m in + (split (-1) , split 0, split 1) + + + (* op of the result is arbitrary Ge *) + let lin_comb n1 c1 n2 c2 = + { coeffs = Vect.lin_comb n1 c1.coeffs n2 c2.coeffs ; + op = Ge ; + cst = (n1 */ c1.cst) +/ (n2 */ c2.cst)} + + (* BUG? : operator of the result ? *) + + let combine_project i c1 c2 = + let p = Vect.get i c1.coeffs + and n = Vect.get i c2.coeffs in + assert (n </ Int 0 && p >/ Int 0) ; + let nopp = minus_num n in + let c =lin_comb nopp c1 p c2 in + let op = if c1.op = Ge || c2.op = Ge then Ge else Eq in + CstrBag.cstr_to_itv {coeffs = c.coeffs ; op = op ; cst= c.cst } + + + let project i m = + let (neg,zero,pos) = partition i m in + let project1 cpos acc = + CstrBag.fold (fun cneg res -> + CstrBag.add (combine_project i cpos cneg) res) neg acc in + (CstrBag.fold project1 pos zero) + + (* Given a vector [x1 -> v1; ... ; xn -> vn] + and a constraint {x1 ; .... xn >= c } + *) + let evaluate_constraint i map cstr = + let {coeffs = _coeffs ; op = _op ; cst = _cst} = cstr in + let vi = Vect.get i _coeffs in + let v = Vect.set i (Int 0) _coeffs in + (vi, _cst -/ Vect.dotp map v) + + + let rec bounds m itv = + match m with + | [] -> itv + | e::m -> bounds m (inter itv (bound_of_constraint e)) + + + + let compare_status (i,(lp,ln)) (i',(lp',ln')) = + let cmp = Pervasives.compare + ((List.length lp) * (List.length ln)) + ((List.length lp') * (List.length ln')) in + if cmp = 0 + then Pervasives.compare i i' + else cmp + + let cardinal m = CstrBag.fold (fun _ x -> x + 1) m 0 + + let lightest_projection l c m = + let bound = c in + if debug then (Printf.printf "l%i" bound; flush stdout) ; + let rec xlight best l = + match l with + | [] -> best + | i::l -> + let proj = (project i m) in + let cproj = cardinal proj in + (*Printf.printf " p %i " cproj; flush stdout;*) + match best with + | None -> + if cproj < bound + then Some(cproj,proj,i) + else xlight (Some(cproj,proj,i)) l + | Some (cbest,_,_) -> + if cproj < cbest + then + if cproj < bound then Some(cproj,proj,i) + else xlight (Some(cproj,proj,i)) l + else xlight best l in + match xlight None l with + | None -> None + | Some(_,p,i) -> Some (p,i) + + + + exception Equality of cstr + + let find_equality m = Bag.find_equation m + + + + let pivot (n,v) eq ge = + assert (eq.op = Eq) ; + let res = + match + compare_num v (Int 0), + compare_num (Vect.get n ge.coeffs) (Int 0) + with + | 0 , _ -> failwith "Buggy" + | _ ,0 -> (CstrBag.cstr_to_itv ge) + | 1 , -1 -> combine_project n eq ge + | -1 , 1 -> combine_project n ge eq + | 1 , 1 -> + combine_project n ge + {coeffs = Vect.mul (Int (-1)) eq.coeffs; + op = eq.op ; + cst = minus_num eq.cst} + | -1 , -1 -> + combine_project n + {coeffs = Vect.mul (Int (-1)) eq.coeffs; + op = eq.op ; cst = minus_num eq.cst} ge + | _ -> failwith "pivot" in + res + + let check_cstr v c = + let {coeffs = _coeffs ; op = _op ; cst = _cst} = c in + let vl = Vect.dotp v _coeffs in + match _op with + | Eq -> vl =/ _cst + | Ge -> vl >= _cst + + + let forall p sys = + try + CstrBag.fold (fun c () -> if p c then () else raise Not_found) sys (); true + with Not_found -> false + + + let check_sys v sys = forall (check_cstr v) sys + + let check_null_cstr c = + let {coeffs = _coeffs ; op = _op ; cst = _cst} = c in + match _op with + | Eq -> (Int 0) =/ _cst + | Ge -> (Int 0) >= _cst + + let check_null sys = forall check_null_cstr sys + + + let optimise_ge + quick_check choose choose_idx return_empty return_ge return_eq m = + let c = cardinal m in + let bound = 2 * c in + if debug then (Printf.printf "optimise_ge: %i\n" c; flush stdout); + + let rec xoptimise m = + if debug then (Printf.printf "x%i" (cardinal m) ; flush stdout); + if debug then (print_bag "xoptimise" m ; flush stdout); + if quick_check m + then return_empty m + else + match find_equality m with + | None -> xoptimise_ge m + | Some eq -> xoptimise_eq eq m + + and xoptimise_ge m = + begin + let c = cardinal m in + let l = List.map fst (List.sort compare_status (CstrBag.status m)) in + let idx = choose bound l c m in + match idx with + | None -> return_empty m + | Some (proj,i) -> + match xoptimise proj with + | None -> None + | Some mapping -> return_ge m i mapping + end + and xoptimise_eq eq m = + let l = List.map fst (Vect.status eq.coeffs) in + match choose_idx l with + | None -> (*if l = [] then None else*) return_empty m + | Some i -> + let p = (i,Vect.get i eq.coeffs) in + let m' = CstrBag.fold + (fun ge res -> CstrBag.add (pivot p eq ge) res) m CstrBag.empty in + match xoptimise ( m') with + | None -> None + | Some mapp -> return_eq m eq i mapp in + try + let res = xoptimise m in res + with CstrBag.Contradiction -> (*print_string "contradiction" ;*) None + + + + let minimise m = + let opt_zero_choose bound l c m = + if c > bound + then lightest_projection l c m + else match l with + | [] -> None + | i::_ -> Some (project i m, i) in + + let choose_idx = function [] -> None | x::l -> Some x in + + let opt_zero_return_empty m = Some Vect.null in + + + let opt_zero_return_ge m i mapping = + let (it:intrvl) = CstrBag.fold (fun cstr itv -> Interval.inter + (bound_of_constraint (evaluate_constraint i mapping cstr)) itv) m + (Itv (None, None)) in + match pick_closed_to_zero it with + | None -> print_endline "Cannot pick" ; None + | Some v -> + let res = (Vect.set i v mapping) in + if debug + then Printf.printf "xoptimise res %i [%s]" i (Vect.string res) ; + Some res in + + let opt_zero_return_eq m eq i mapp = + let (a,b) = evaluate_constraint i mapp eq in + Some (Vect.set i (div_num b a) mapp) in + + optimise_ge check_null opt_zero_choose + choose_idx opt_zero_return_empty opt_zero_return_ge opt_zero_return_eq m + + let normalise cstr = [CstrBag.cstr_to_itv cstr] + + let find_point l = + (* List.iter (fun e -> print_endline (Cstr.string_of_cstr e)) l;*) + try + let m = List.fold_left (fun sys e -> CstrBag.add (CstrBag.cstr_to_itv e) sys) + CstrBag.empty l in + match minimise m with + | None -> None + | Some res -> + if debug then Printf.printf "[%s]" (Vect.string res); + Some res + with CstrBag.Contradiction -> None + + + let find_q_interval_for x m = + if debug then Printf.printf "find_q_interval_for %i\n" x ; + + let choose bound l c m = + let rec xchoose l = + match l with + | [] -> None + | i::l -> if i = x then xchoose l else Some (project i m,i) in + xchoose l in + + let rec choose_idx = function + [] -> None + | e::l -> if e = x then choose_idx l else Some e in + + let return_empty m = (* Beurk *) + (* returns the interval of x *) + Some (CstrBag.fold (fun cstr itv -> + let i = if cstr.op = Eq + then Point (cstr.cst // Vect.get x cstr.coeffs) + else if Vect.is_null (Vect.set x (Int 0) cstr.coeffs) + then bound_of_constraint (Vect.get x cstr.coeffs , cstr.cst) + else itv + in + Interval.inter i itv) m (Itv (None, None))) in + + let return_ge m i res = Some res in + + let return_eq m eq i res = Some res in + + try + optimise_ge + (fun x -> false) choose choose_idx return_empty return_ge return_eq m + with CstrBag.Contradiction -> None + + + let find_q_intervals sys = + let variables = + List.map fst (List.sort compare_status (CstrBag.status sys)) in + List.map (fun x -> (x,find_q_interval_for x sys)) variables + + let pp_option f o = function + None -> Printf.fprintf o "None" + | Some x -> Printf.fprintf o "Some %a" f x + + let optimise vect sys = + (* we have to modify the system with a dummy variable *) + let fresh = + List.fold_left (fun fr c -> Pervasives.max fr (Vect.fresh c.coeffs)) 0 sys in + assert (List.for_all (fun x -> Vect.get fresh x.coeffs =/ Int 0) sys); + let cstr = { + coeffs = Vect.set fresh (Int (-1)) vect ; + op = Eq ; + cst = (Int 0)} in + try + find_q_interval_for fresh + (List.fold_left + (fun bg c -> CstrBag.add (CstrBag.cstr_to_itv c) bg) + CstrBag.empty (cstr::sys)) + with CstrBag.Contradiction -> None + + + let optimise vect sys = + let res = optimise vect sys in + if debug + then Printf.printf "optimise %s -> %a\n" + (Vect.string vect) (pp_option (fun o x -> Printf.printf "%s" (string_of_intrvl x))) res + ; res + + let find_Q_interval sys = + try + let sys = + (List.fold_left + (fun bg c -> CstrBag.add (CstrBag.cstr_to_itv c) bg) CstrBag.empty sys) in + let candidates = + List.fold_left + (fun l (x,i) -> match i with + None -> (x,Empty)::l + | Some i -> (x,i)::l) [] (find_q_intervals sys) in + match List.fold_left + (fun (x1,i1) (x2,i2) -> + if smaller_itv i1 i2 + then (x1,i1) else (x2,i2)) (-1,Itv(None,None)) candidates + with + | (i,Empty) -> None + | (x,Itv(Some i, Some j)) -> Some(i,x,j) + | (x,Point n) -> Some(n,x,n) + | _ -> None + with CstrBag.Contradiction -> None + + +end + diff --git a/contrib/micromega/micromega.ml b/contrib/micromega/micromega.ml new file mode 100644 index 00000000..e151e4e1 --- /dev/null +++ b/contrib/micromega/micromega.ml @@ -0,0 +1,1512 @@ +type __ = Obj.t +let __ = let rec f _ = Obj.repr f in Obj.repr f + +type bool = + | True + | False + +(** val negb : bool -> bool **) + +let negb = function + | True -> False + | False -> True + +type nat = + | O + | S of nat + +type 'a option = + | Some of 'a + | None + +type ('a, 'b) prod = + | Pair of 'a * 'b + +type comparison = + | Eq + | Lt + | Gt + +(** val compOpp : comparison -> comparison **) + +let compOpp = function + | Eq -> Eq + | Lt -> Gt + | Gt -> Lt + +type sumbool = + | Left + | Right + +type 'a sumor = + | Inleft of 'a + | Inright + +type 'a list = + | Nil + | Cons of 'a * 'a list + +(** val app : 'a1 list -> 'a1 list -> 'a1 list **) + +let rec app l m = + match l with + | Nil -> m + | Cons (a, l1) -> Cons (a, (app l1 m)) + +(** val nth : nat -> 'a1 list -> 'a1 -> 'a1 **) + +let rec nth n0 l default = + match n0 with + | O -> (match l with + | Nil -> default + | Cons (x, l') -> x) + | S m -> + (match l with + | Nil -> default + | Cons (x, t0) -> nth m t0 default) + +(** val map : ('a1 -> 'a2) -> 'a1 list -> 'a2 list **) + +let rec map f = function + | Nil -> Nil + | Cons (a, t0) -> Cons ((f a), (map f t0)) + +type positive = + | XI of positive + | XO of positive + | XH + +(** val psucc : positive -> positive **) + +let rec psucc = function + | XI p -> XO (psucc p) + | XO p -> XI p + | XH -> XO XH + +(** val pplus : positive -> positive -> positive **) + +let rec pplus x y = + match x with + | XI p -> + (match y with + | XI q0 -> XO (pplus_carry p q0) + | XO q0 -> XI (pplus p q0) + | XH -> XO (psucc p)) + | XO p -> + (match y with + | XI q0 -> XI (pplus p q0) + | XO q0 -> XO (pplus p q0) + | XH -> XI p) + | XH -> + (match y with + | XI q0 -> XO (psucc q0) + | XO q0 -> XI q0 + | XH -> XO XH) + +(** val pplus_carry : positive -> positive -> positive **) + +and pplus_carry x y = + match x with + | XI p -> + (match y with + | XI q0 -> XI (pplus_carry p q0) + | XO q0 -> XO (pplus_carry p q0) + | XH -> XI (psucc p)) + | XO p -> + (match y with + | XI q0 -> XO (pplus_carry p q0) + | XO q0 -> XI (pplus p q0) + | XH -> XO (psucc p)) + | XH -> + (match y with + | XI q0 -> XI (psucc q0) + | XO q0 -> XO (psucc q0) + | XH -> XI XH) + +(** val p_of_succ_nat : nat -> positive **) + +let rec p_of_succ_nat = function + | O -> XH + | S x -> psucc (p_of_succ_nat x) + +(** val pdouble_minus_one : positive -> positive **) + +let rec pdouble_minus_one = function + | XI p -> XI (XO p) + | XO p -> XI (pdouble_minus_one p) + | XH -> XH + +type positive_mask = + | IsNul + | IsPos of positive + | IsNeg + +(** val pdouble_plus_one_mask : positive_mask -> positive_mask **) + +let pdouble_plus_one_mask = function + | IsNul -> IsPos XH + | IsPos p -> IsPos (XI p) + | IsNeg -> IsNeg + +(** val pdouble_mask : positive_mask -> positive_mask **) + +let pdouble_mask = function + | IsNul -> IsNul + | IsPos p -> IsPos (XO p) + | IsNeg -> IsNeg + +(** val pdouble_minus_two : positive -> positive_mask **) + +let pdouble_minus_two = function + | XI p -> IsPos (XO (XO p)) + | XO p -> IsPos (XO (pdouble_minus_one p)) + | XH -> IsNul + +(** val pminus_mask : positive -> positive -> positive_mask **) + +let rec pminus_mask x y = + match x with + | XI p -> + (match y with + | XI q0 -> pdouble_mask (pminus_mask p q0) + | XO q0 -> pdouble_plus_one_mask (pminus_mask p q0) + | XH -> IsPos (XO p)) + | XO p -> + (match y with + | XI q0 -> pdouble_plus_one_mask (pminus_mask_carry p q0) + | XO q0 -> pdouble_mask (pminus_mask p q0) + | XH -> IsPos (pdouble_minus_one p)) + | XH -> (match y with + | XH -> IsNul + | _ -> IsNeg) + +(** val pminus_mask_carry : positive -> positive -> positive_mask **) + +and pminus_mask_carry x y = + match x with + | XI p -> + (match y with + | XI q0 -> pdouble_plus_one_mask (pminus_mask_carry p q0) + | XO q0 -> pdouble_mask (pminus_mask p q0) + | XH -> IsPos (pdouble_minus_one p)) + | XO p -> + (match y with + | XI q0 -> pdouble_mask (pminus_mask_carry p q0) + | XO q0 -> pdouble_plus_one_mask (pminus_mask_carry p q0) + | XH -> pdouble_minus_two p) + | XH -> IsNeg + +(** val pminus : positive -> positive -> positive **) + +let pminus x y = + match pminus_mask x y with + | IsPos z0 -> z0 + | _ -> XH + +(** val pmult : positive -> positive -> positive **) + +let rec pmult x y = + match x with + | XI p -> pplus y (XO (pmult p y)) + | XO p -> XO (pmult p y) + | XH -> y + +(** val pcompare : positive -> positive -> comparison -> comparison **) + +let rec pcompare x y r = + match x with + | XI p -> + (match y with + | XI q0 -> pcompare p q0 r + | XO q0 -> pcompare p q0 Gt + | XH -> Gt) + | XO p -> + (match y with + | XI q0 -> pcompare p q0 Lt + | XO q0 -> pcompare p q0 r + | XH -> Gt) + | XH -> (match y with + | XH -> r + | _ -> Lt) + +type n = + | N0 + | Npos of positive + +type z = + | Z0 + | Zpos of positive + | Zneg of positive + +(** val zdouble_plus_one : z -> z **) + +let zdouble_plus_one = function + | Z0 -> Zpos XH + | Zpos p -> Zpos (XI p) + | Zneg p -> Zneg (pdouble_minus_one p) + +(** val zdouble_minus_one : z -> z **) + +let zdouble_minus_one = function + | Z0 -> Zneg XH + | Zpos p -> Zpos (pdouble_minus_one p) + | Zneg p -> Zneg (XI p) + +(** val zdouble : z -> z **) + +let zdouble = function + | Z0 -> Z0 + | Zpos p -> Zpos (XO p) + | Zneg p -> Zneg (XO p) + +(** val zPminus : positive -> positive -> z **) + +let rec zPminus x y = + match x with + | XI p -> + (match y with + | XI q0 -> zdouble (zPminus p q0) + | XO q0 -> zdouble_plus_one (zPminus p q0) + | XH -> Zpos (XO p)) + | XO p -> + (match y with + | XI q0 -> zdouble_minus_one (zPminus p q0) + | XO q0 -> zdouble (zPminus p q0) + | XH -> Zpos (pdouble_minus_one p)) + | XH -> + (match y with + | XI q0 -> Zneg (XO q0) + | XO q0 -> Zneg (pdouble_minus_one q0) + | XH -> Z0) + +(** val zplus : z -> z -> z **) + +let zplus x y = + match x with + | Z0 -> y + | Zpos x' -> + (match y with + | Z0 -> Zpos x' + | Zpos y' -> Zpos (pplus x' y') + | Zneg y' -> + (match pcompare x' y' Eq with + | Eq -> Z0 + | Lt -> Zneg (pminus y' x') + | Gt -> Zpos (pminus x' y'))) + | Zneg x' -> + (match y with + | Z0 -> Zneg x' + | Zpos y' -> + (match pcompare x' y' Eq with + | Eq -> Z0 + | Lt -> Zpos (pminus y' x') + | Gt -> Zneg (pminus x' y')) + | Zneg y' -> Zneg (pplus x' y')) + +(** val zopp : z -> z **) + +let zopp = function + | Z0 -> Z0 + | Zpos x0 -> Zneg x0 + | Zneg x0 -> Zpos x0 + +(** val zminus : z -> z -> z **) + +let zminus m n0 = + zplus m (zopp n0) + +(** val zmult : z -> z -> z **) + +let zmult x y = + match x with + | Z0 -> Z0 + | Zpos x' -> + (match y with + | Z0 -> Z0 + | Zpos y' -> Zpos (pmult x' y') + | Zneg y' -> Zneg (pmult x' y')) + | Zneg x' -> + (match y with + | Z0 -> Z0 + | Zpos y' -> Zneg (pmult x' y') + | Zneg y' -> Zpos (pmult x' y')) + +(** val zcompare : z -> z -> comparison **) + +let zcompare x y = + match x with + | Z0 -> (match y with + | Z0 -> Eq + | Zpos y' -> Lt + | Zneg y' -> Gt) + | Zpos x' -> (match y with + | Zpos y' -> pcompare x' y' Eq + | _ -> Gt) + | Zneg x' -> + (match y with + | Zneg y' -> compOpp (pcompare x' y' Eq) + | _ -> Lt) + +(** val dcompare_inf : comparison -> sumbool sumor **) + +let dcompare_inf = function + | Eq -> Inleft Left + | Lt -> Inleft Right + | Gt -> Inright + +(** val zcompare_rec : + z -> z -> (__ -> 'a1) -> (__ -> 'a1) -> (__ -> 'a1) -> 'a1 **) + +let zcompare_rec x y h1 h2 h3 = + match dcompare_inf (zcompare x y) with + | Inleft x0 -> (match x0 with + | Left -> h1 __ + | Right -> h2 __) + | Inright -> h3 __ + +(** val z_gt_dec : z -> z -> sumbool **) + +let z_gt_dec x y = + zcompare_rec x y (fun _ -> Right) (fun _ -> Right) (fun _ -> Left) + +(** val zle_bool : z -> z -> bool **) + +let zle_bool x y = + match zcompare x y with + | Gt -> False + | _ -> True + +(** val zge_bool : z -> z -> bool **) + +let zge_bool x y = + match zcompare x y with + | Lt -> False + | _ -> True + +(** val zgt_bool : z -> z -> bool **) + +let zgt_bool x y = + match zcompare x y with + | Gt -> True + | _ -> False + +(** val zeq_bool : z -> z -> bool **) + +let zeq_bool x y = + match zcompare x y with + | Eq -> True + | _ -> False + +(** val n_of_nat : nat -> n **) + +let n_of_nat = function + | O -> N0 + | S n' -> Npos (p_of_succ_nat n') + +(** val zdiv_eucl_POS : positive -> z -> (z, z) prod **) + +let rec zdiv_eucl_POS a b = + match a with + | XI a' -> + let Pair (q0, r) = zdiv_eucl_POS a' b in + let r' = zplus (zmult (Zpos (XO XH)) r) (Zpos XH) in + (match zgt_bool b r' with + | True -> Pair ((zmult (Zpos (XO XH)) q0), r') + | False -> Pair ((zplus (zmult (Zpos (XO XH)) q0) (Zpos XH)), + (zminus r' b))) + | XO a' -> + let Pair (q0, r) = zdiv_eucl_POS a' b in + let r' = zmult (Zpos (XO XH)) r in + (match zgt_bool b r' with + | True -> Pair ((zmult (Zpos (XO XH)) q0), r') + | False -> Pair ((zplus (zmult (Zpos (XO XH)) q0) (Zpos XH)), + (zminus r' b))) + | XH -> + (match zge_bool b (Zpos (XO XH)) with + | True -> Pair (Z0, (Zpos XH)) + | False -> Pair ((Zpos XH), Z0)) + +(** val zdiv_eucl : z -> z -> (z, z) prod **) + +let zdiv_eucl a b = + match a with + | Z0 -> Pair (Z0, Z0) + | Zpos a' -> + (match b with + | Z0 -> Pair (Z0, Z0) + | Zpos p -> zdiv_eucl_POS a' b + | Zneg b' -> + let Pair (q0, r) = zdiv_eucl_POS a' (Zpos b') in + (match r with + | Z0 -> Pair ((zopp q0), Z0) + | _ -> Pair ((zopp (zplus q0 (Zpos XH))), (zplus b r)))) + | Zneg a' -> + (match b with + | Z0 -> Pair (Z0, Z0) + | Zpos p -> + let Pair (q0, r) = zdiv_eucl_POS a' b in + (match r with + | Z0 -> Pair ((zopp q0), Z0) + | _ -> Pair ((zopp (zplus q0 (Zpos XH))), (zminus b r))) + | Zneg b' -> + let Pair (q0, r) = zdiv_eucl_POS a' (Zpos b') in + Pair (q0, (zopp r))) + +type 'c pol = + | Pc of 'c + | Pinj of positive * 'c pol + | PX of 'c pol * positive * 'c pol + +(** val p0 : 'a1 -> 'a1 pol **) + +let p0 cO = + Pc cO + +(** val p1 : 'a1 -> 'a1 pol **) + +let p1 cI = + Pc cI + +(** val peq : ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> bool **) + +let rec peq ceqb p p' = + match p with + | Pc c -> (match p' with + | Pc c' -> ceqb c c' + | _ -> False) + | Pinj (j, q0) -> + (match p' with + | Pinj (j', q') -> + (match pcompare j j' Eq with + | Eq -> peq ceqb q0 q' + | _ -> False) + | _ -> False) + | PX (p2, i, q0) -> + (match p' with + | PX (p'0, i', q') -> + (match pcompare i i' Eq with + | Eq -> + (match peq ceqb p2 p'0 with + | True -> peq ceqb q0 q' + | False -> False) + | _ -> False) + | _ -> False) + +(** val mkPinj_pred : positive -> 'a1 pol -> 'a1 pol **) + +let mkPinj_pred j p = + match j with + | XI j0 -> Pinj ((XO j0), p) + | XO j0 -> Pinj ((pdouble_minus_one j0), p) + | XH -> p + +(** val mkPX : + 'a1 -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol **) + +let mkPX cO ceqb p i q0 = + match p with + | Pc c -> + (match ceqb c cO with + | True -> + (match q0 with + | Pc c0 -> q0 + | Pinj (j', q1) -> Pinj ((pplus XH j'), q1) + | PX (p2, p3, p4) -> Pinj (XH, q0)) + | False -> PX (p, i, q0)) + | Pinj (p2, p3) -> PX (p, i, q0) + | PX (p', i', q') -> + (match peq ceqb q' (p0 cO) with + | True -> PX (p', (pplus i' i), q0) + | False -> PX (p, i, q0)) + +(** val mkXi : 'a1 -> 'a1 -> positive -> 'a1 pol **) + +let mkXi cO cI i = + PX ((p1 cI), i, (p0 cO)) + +(** val mkX : 'a1 -> 'a1 -> 'a1 pol **) + +let mkX cO cI = + mkXi cO cI XH + +(** val popp : ('a1 -> 'a1) -> 'a1 pol -> 'a1 pol **) + +let rec popp copp = function + | Pc c -> Pc (copp c) + | Pinj (j, q0) -> Pinj (j, (popp copp q0)) + | PX (p2, i, q0) -> PX ((popp copp p2), i, (popp copp q0)) + +(** val paddC : ('a1 -> 'a1 -> 'a1) -> 'a1 pol -> 'a1 -> 'a1 pol **) + +let rec paddC cadd p c = + match p with + | Pc c1 -> Pc (cadd c1 c) + | Pinj (j, q0) -> Pinj (j, (paddC cadd q0 c)) + | PX (p2, i, q0) -> PX (p2, i, (paddC cadd q0 c)) + +(** val psubC : ('a1 -> 'a1 -> 'a1) -> 'a1 pol -> 'a1 -> 'a1 pol **) + +let rec psubC csub p c = + match p with + | Pc c1 -> Pc (csub c1 c) + | Pinj (j, q0) -> Pinj (j, (psubC csub q0 c)) + | PX (p2, i, q0) -> PX (p2, i, (psubC csub q0 c)) + +(** val paddI : + ('a1 -> 'a1 -> 'a1) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1 pol -> + positive -> 'a1 pol -> 'a1 pol **) + +let rec paddI cadd pop q0 j = function + | Pc c -> + let p2 = paddC cadd q0 c in + (match p2 with + | Pc c0 -> p2 + | Pinj (j', q1) -> Pinj ((pplus j j'), q1) + | PX (p3, p4, p5) -> Pinj (j, p2)) + | Pinj (j', q') -> + (match zPminus j' j with + | Z0 -> + let p2 = pop q' q0 in + (match p2 with + | Pc c -> p2 + | Pinj (j'0, q1) -> Pinj ((pplus j j'0), q1) + | PX (p3, p4, p5) -> Pinj (j, p2)) + | Zpos k -> + let p2 = pop (Pinj (k, q')) q0 in + (match p2 with + | Pc c -> p2 + | Pinj (j'0, q1) -> Pinj ((pplus j j'0), q1) + | PX (p3, p4, p5) -> Pinj (j, p2)) + | Zneg k -> + let p2 = paddI cadd pop q0 k q' in + (match p2 with + | Pc c -> p2 + | Pinj (j'0, q1) -> Pinj ((pplus j' j'0), q1) + | PX (p3, p4, p5) -> Pinj (j', p2))) + | PX (p2, i, q') -> + (match j with + | XI j0 -> PX (p2, i, (paddI cadd pop q0 (XO j0) q')) + | XO j0 -> PX (p2, i, (paddI cadd pop q0 (pdouble_minus_one j0) q')) + | XH -> PX (p2, i, (pop q' q0))) + +(** val psubI : + ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> + 'a1 pol -> positive -> 'a1 pol -> 'a1 pol **) + +let rec psubI cadd copp pop q0 j = function + | Pc c -> + let p2 = paddC cadd (popp copp q0) c in + (match p2 with + | Pc c0 -> p2 + | Pinj (j', q1) -> Pinj ((pplus j j'), q1) + | PX (p3, p4, p5) -> Pinj (j, p2)) + | Pinj (j', q') -> + (match zPminus j' j with + | Z0 -> + let p2 = pop q' q0 in + (match p2 with + | Pc c -> p2 + | Pinj (j'0, q1) -> Pinj ((pplus j j'0), q1) + | PX (p3, p4, p5) -> Pinj (j, p2)) + | Zpos k -> + let p2 = pop (Pinj (k, q')) q0 in + (match p2 with + | Pc c -> p2 + | Pinj (j'0, q1) -> Pinj ((pplus j j'0), q1) + | PX (p3, p4, p5) -> Pinj (j, p2)) + | Zneg k -> + let p2 = psubI cadd copp pop q0 k q' in + (match p2 with + | Pc c -> p2 + | Pinj (j'0, q1) -> Pinj ((pplus j' j'0), q1) + | PX (p3, p4, p5) -> Pinj (j', p2))) + | PX (p2, i, q') -> + (match j with + | XI j0 -> PX (p2, i, (psubI cadd copp pop q0 (XO j0) q')) + | XO j0 -> PX (p2, i, + (psubI cadd copp pop q0 (pdouble_minus_one j0) q')) + | XH -> PX (p2, i, (pop q' q0))) + +(** val paddX : + 'a1 -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1 pol + -> positive -> 'a1 pol -> 'a1 pol **) + +let rec paddX cO ceqb pop p' i' p = match p with + | Pc c -> PX (p', i', p) + | Pinj (j, q') -> + (match j with + | XI j0 -> PX (p', i', (Pinj ((XO j0), q'))) + | XO j0 -> PX (p', i', (Pinj ((pdouble_minus_one j0), q'))) + | XH -> PX (p', i', q')) + | PX (p2, i, q') -> + (match zPminus i i' with + | Z0 -> mkPX cO ceqb (pop p2 p') i q' + | Zpos k -> mkPX cO ceqb (pop (PX (p2, k, (p0 cO))) p') i' q' + | Zneg k -> mkPX cO ceqb (paddX cO ceqb pop p' k p2) i q') + +(** val psubX : + 'a1 -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol -> 'a1 + pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol **) + +let rec psubX cO copp ceqb pop p' i' p = match p with + | Pc c -> PX ((popp copp p'), i', p) + | Pinj (j, q') -> + (match j with + | XI j0 -> PX ((popp copp p'), i', (Pinj ((XO j0), q'))) + | XO j0 -> PX ((popp copp p'), i', (Pinj ( + (pdouble_minus_one j0), q'))) + | XH -> PX ((popp copp p'), i', q')) + | PX (p2, i, q') -> + (match zPminus i i' with + | Z0 -> mkPX cO ceqb (pop p2 p') i q' + | Zpos k -> mkPX cO ceqb (pop (PX (p2, k, (p0 cO))) p') i' q' + | Zneg k -> mkPX cO ceqb (psubX cO copp ceqb pop p' k p2) i q') + +(** val padd : + 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol + -> 'a1 pol **) + +let rec padd cO cadd ceqb p = function + | Pc c' -> paddC cadd p c' + | Pinj (j', q') -> paddI cadd (fun x x0 -> padd cO cadd ceqb x x0) q' j' p + | PX (p'0, i', q') -> + (match p with + | Pc c -> PX (p'0, i', (paddC cadd q' c)) + | Pinj (j, q0) -> + (match j with + | XI j0 -> PX (p'0, i', + (padd cO cadd ceqb (Pinj ((XO j0), q0)) q')) + | XO j0 -> PX (p'0, i', + (padd cO cadd ceqb (Pinj ((pdouble_minus_one j0), q0)) + q')) + | XH -> PX (p'0, i', (padd cO cadd ceqb q0 q'))) + | PX (p2, i, q0) -> + (match zPminus i i' with + | Z0 -> + mkPX cO ceqb (padd cO cadd ceqb p2 p'0) i + (padd cO cadd ceqb q0 q') + | Zpos k -> + mkPX cO ceqb + (padd cO cadd ceqb (PX (p2, k, (p0 cO))) p'0) i' + (padd cO cadd ceqb q0 q') + | Zneg k -> + mkPX cO ceqb + (paddX cO ceqb (fun x x0 -> padd cO cadd ceqb x x0) p'0 + k p2) i (padd cO cadd ceqb q0 q'))) + +(** val psub : + 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 + -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol **) + +let rec psub cO cadd csub copp ceqb p = function + | Pc c' -> psubC csub p c' + | Pinj (j', q') -> + psubI cadd copp (fun x x0 -> psub cO cadd csub copp ceqb x x0) q' j' p + | PX (p'0, i', q') -> + (match p with + | Pc c -> PX ((popp copp p'0), i', (paddC cadd (popp copp q') c)) + | Pinj (j, q0) -> + (match j with + | XI j0 -> PX ((popp copp p'0), i', + (psub cO cadd csub copp ceqb (Pinj ((XO j0), q0)) q')) + | XO j0 -> PX ((popp copp p'0), i', + (psub cO cadd csub copp ceqb (Pinj + ((pdouble_minus_one j0), q0)) q')) + | XH -> PX ((popp copp p'0), i', + (psub cO cadd csub copp ceqb q0 q'))) + | PX (p2, i, q0) -> + (match zPminus i i' with + | Z0 -> + mkPX cO ceqb (psub cO cadd csub copp ceqb p2 p'0) i + (psub cO cadd csub copp ceqb q0 q') + | Zpos k -> + mkPX cO ceqb + (psub cO cadd csub copp ceqb (PX (p2, k, (p0 cO))) p'0) + i' (psub cO cadd csub copp ceqb q0 q') + | Zneg k -> + mkPX cO ceqb + (psubX cO copp ceqb (fun x x0 -> + psub cO cadd csub copp ceqb x x0) p'0 k p2) i + (psub cO cadd csub copp ceqb q0 q'))) + +(** val pmulC_aux : + 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 -> + 'a1 pol **) + +let rec pmulC_aux cO cmul ceqb p c = + match p with + | Pc c' -> Pc (cmul c' c) + | Pinj (j, q0) -> + let p2 = pmulC_aux cO cmul ceqb q0 c in + (match p2 with + | Pc c0 -> p2 + | Pinj (j', q1) -> Pinj ((pplus j j'), q1) + | PX (p3, p4, p5) -> Pinj (j, p2)) + | PX (p2, i, q0) -> + mkPX cO ceqb (pmulC_aux cO cmul ceqb p2 c) i + (pmulC_aux cO cmul ceqb q0 c) + +(** val pmulC : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> + 'a1 -> 'a1 pol **) + +let pmulC cO cI cmul ceqb p c = + match ceqb c cO with + | True -> p0 cO + | False -> + (match ceqb c cI with + | True -> p + | False -> pmulC_aux cO cmul ceqb p c) + +(** val pmulI : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> + 'a1 pol -> 'a1 pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol **) + +let rec pmulI cO cI cmul ceqb pmul0 q0 j = function + | Pc c -> + let p2 = pmulC cO cI cmul ceqb q0 c in + (match p2 with + | Pc c0 -> p2 + | Pinj (j', q1) -> Pinj ((pplus j j'), q1) + | PX (p3, p4, p5) -> Pinj (j, p2)) + | Pinj (j', q') -> + (match zPminus j' j with + | Z0 -> + let p2 = pmul0 q' q0 in + (match p2 with + | Pc c -> p2 + | Pinj (j'0, q1) -> Pinj ((pplus j j'0), q1) + | PX (p3, p4, p5) -> Pinj (j, p2)) + | Zpos k -> + let p2 = pmul0 (Pinj (k, q')) q0 in + (match p2 with + | Pc c -> p2 + | Pinj (j'0, q1) -> Pinj ((pplus j j'0), q1) + | PX (p3, p4, p5) -> Pinj (j, p2)) + | Zneg k -> + let p2 = pmulI cO cI cmul ceqb pmul0 q0 k q' in + (match p2 with + | Pc c -> p2 + | Pinj (j'0, q1) -> Pinj ((pplus j' j'0), q1) + | PX (p3, p4, p5) -> Pinj (j', p2))) + | PX (p', i', q') -> + (match j with + | XI j' -> + mkPX cO ceqb (pmulI cO cI cmul ceqb pmul0 q0 j p') i' + (pmulI cO cI cmul ceqb pmul0 q0 (XO j') q') + | XO j' -> + mkPX cO ceqb (pmulI cO cI cmul ceqb pmul0 q0 j p') i' + (pmulI cO cI cmul ceqb pmul0 q0 (pdouble_minus_one j') q') + | XH -> + mkPX cO ceqb (pmulI cO cI cmul ceqb pmul0 q0 XH p') i' + (pmul0 q' q0)) + +(** val pmul : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 + -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol **) + +let rec pmul cO cI cadd cmul ceqb p p'' = match p'' with + | Pc c -> pmulC cO cI cmul ceqb p c + | Pinj (j', q') -> + pmulI cO cI cmul ceqb (fun x x0 -> pmul cO cI cadd cmul ceqb x x0) q' + j' p + | PX (p', i', q') -> + (match p with + | Pc c -> pmulC cO cI cmul ceqb p'' c + | Pinj (j, q0) -> + mkPX cO ceqb (pmul cO cI cadd cmul ceqb p p') i' + (match j with + | XI j0 -> + pmul cO cI cadd cmul ceqb (Pinj ((XO j0), q0)) q' + | XO j0 -> + pmul cO cI cadd cmul ceqb (Pinj + ((pdouble_minus_one j0), q0)) q' + | XH -> pmul cO cI cadd cmul ceqb q0 q') + | PX (p2, i, q0) -> + padd cO cadd ceqb + (mkPX cO ceqb + (padd cO cadd ceqb + (mkPX cO ceqb (pmul cO cI cadd cmul ceqb p2 p') i (p0 cO)) + (pmul cO cI cadd cmul ceqb + (match q0 with + | Pc c -> q0 + | Pinj (j', q1) -> Pinj ((pplus XH j'), q1) + | PX (p3, p4, p5) -> Pinj (XH, q0)) p')) i' + (p0 cO)) + (mkPX cO ceqb + (pmulI cO cI cmul ceqb (fun x x0 -> + pmul cO cI cadd cmul ceqb x x0) q' XH p2) i + (pmul cO cI cadd cmul ceqb q0 q'))) + +type 'c pExpr = + | PEc of 'c + | PEX of positive + | PEadd of 'c pExpr * 'c pExpr + | PEsub of 'c pExpr * 'c pExpr + | PEmul of 'c pExpr * 'c pExpr + | PEopp of 'c pExpr + | PEpow of 'c pExpr * n + +(** val mk_X : 'a1 -> 'a1 -> positive -> 'a1 pol **) + +let mk_X cO cI j = + mkPinj_pred j (mkX cO cI) + +(** val ppow_pos : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 + -> bool) -> ('a1 pol -> 'a1 pol) -> 'a1 pol -> 'a1 pol -> positive -> 'a1 + pol **) + +let rec ppow_pos cO cI cadd cmul ceqb subst_l res p = function + | XI p3 -> + subst_l + (pmul cO cI cadd cmul ceqb + (ppow_pos cO cI cadd cmul ceqb subst_l + (ppow_pos cO cI cadd cmul ceqb subst_l res p p3) p p3) p) + | XO p3 -> + ppow_pos cO cI cadd cmul ceqb subst_l + (ppow_pos cO cI cadd cmul ceqb subst_l res p p3) p p3 + | XH -> subst_l (pmul cO cI cadd cmul ceqb res p) + +(** val ppow_N : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 + -> bool) -> ('a1 pol -> 'a1 pol) -> 'a1 pol -> n -> 'a1 pol **) + +let ppow_N cO cI cadd cmul ceqb subst_l p = function + | N0 -> p1 cI + | Npos p2 -> ppow_pos cO cI cadd cmul ceqb subst_l (p1 cI) p p2 + +(** val norm_aux : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 + -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExpr -> 'a1 pol **) + +let rec norm_aux cO cI cadd cmul csub copp ceqb = function + | PEc c -> Pc c + | PEX j -> mk_X cO cI j + | PEadd (pe1, pe2) -> + (match pe1 with + | PEopp pe3 -> + psub cO cadd csub copp ceqb + (norm_aux cO cI cadd cmul csub copp ceqb pe2) + (norm_aux cO cI cadd cmul csub copp ceqb pe3) + | _ -> + (match pe2 with + | PEopp pe3 -> + psub cO cadd csub copp ceqb + (norm_aux cO cI cadd cmul csub copp ceqb pe1) + (norm_aux cO cI cadd cmul csub copp ceqb pe3) + | _ -> + padd cO cadd ceqb + (norm_aux cO cI cadd cmul csub copp ceqb pe1) + (norm_aux cO cI cadd cmul csub copp ceqb pe2))) + | PEsub (pe1, pe2) -> + psub cO cadd csub copp ceqb + (norm_aux cO cI cadd cmul csub copp ceqb pe1) + (norm_aux cO cI cadd cmul csub copp ceqb pe2) + | PEmul (pe1, pe2) -> + pmul cO cI cadd cmul ceqb (norm_aux cO cI cadd cmul csub copp ceqb pe1) + (norm_aux cO cI cadd cmul csub copp ceqb pe2) + | PEopp pe1 -> popp copp (norm_aux cO cI cadd cmul csub copp ceqb pe1) + | PEpow (pe1, n0) -> + ppow_N cO cI cadd cmul ceqb (fun p -> p) + (norm_aux cO cI cadd cmul csub copp ceqb pe1) n0 + +type 'a bFormula = + | TT + | FF + | X + | A of 'a + | Cj of 'a bFormula * 'a bFormula + | D of 'a bFormula * 'a bFormula + | N of 'a bFormula + | I of 'a bFormula * 'a bFormula + +type 'term' clause = 'term' list + +type 'term' cnf = 'term' clause list + +(** val tt : 'a1 cnf **) + +let tt = + Nil + +(** val ff : 'a1 cnf **) + +let ff = + Cons (Nil, Nil) + +(** val or_clause_cnf : 'a1 clause -> 'a1 cnf -> 'a1 cnf **) + +let or_clause_cnf t0 f = + map (fun x -> app t0 x) f + +(** val or_cnf : 'a1 cnf -> 'a1 cnf -> 'a1 cnf **) + +let rec or_cnf f f' = + match f with + | Nil -> tt + | Cons (e, rst) -> app (or_cnf rst f') (or_clause_cnf e f') + +(** val and_cnf : 'a1 cnf -> 'a1 cnf -> 'a1 cnf **) + +let and_cnf f1 f2 = + app f1 f2 + +(** val xcnf : + ('a1 -> 'a2 cnf) -> ('a1 -> 'a2 cnf) -> bool -> 'a1 bFormula -> 'a2 cnf **) + +let rec xcnf normalise0 negate0 pol0 = function + | TT -> (match pol0 with + | True -> tt + | False -> ff) + | FF -> (match pol0 with + | True -> ff + | False -> tt) + | X -> ff + | A x -> (match pol0 with + | True -> normalise0 x + | False -> negate0 x) + | Cj (e1, e2) -> + (match pol0 with + | True -> + and_cnf (xcnf normalise0 negate0 pol0 e1) + (xcnf normalise0 negate0 pol0 e2) + | False -> + or_cnf (xcnf normalise0 negate0 pol0 e1) + (xcnf normalise0 negate0 pol0 e2)) + | D (e1, e2) -> + (match pol0 with + | True -> + or_cnf (xcnf normalise0 negate0 pol0 e1) + (xcnf normalise0 negate0 pol0 e2) + | False -> + and_cnf (xcnf normalise0 negate0 pol0 e1) + (xcnf normalise0 negate0 pol0 e2)) + | N e -> xcnf normalise0 negate0 (negb pol0) e + | I (e1, e2) -> + (match pol0 with + | True -> + or_cnf (xcnf normalise0 negate0 (negb pol0) e1) + (xcnf normalise0 negate0 pol0 e2) + | False -> + and_cnf (xcnf normalise0 negate0 (negb pol0) e1) + (xcnf normalise0 negate0 pol0 e2)) + +(** val cnf_checker : + ('a1 list -> 'a2 -> bool) -> 'a1 cnf -> 'a2 list -> bool **) + +let rec cnf_checker checker f l = + match f with + | Nil -> True + | Cons (e, f0) -> + (match l with + | Nil -> False + | Cons (c, l0) -> + (match checker e c with + | True -> cnf_checker checker f0 l0 + | False -> False)) + +(** val tauto_checker : + ('a1 -> 'a2 cnf) -> ('a1 -> 'a2 cnf) -> ('a2 list -> 'a3 -> bool) -> 'a1 + bFormula -> 'a3 list -> bool **) + +let tauto_checker normalise0 negate0 checker f w = + cnf_checker checker (xcnf normalise0 negate0 True f) w + +type 'c pExprC = 'c pExpr + +type 'c polC = 'c pol + +type op1 = + | Equal + | NonEqual + | Strict + | NonStrict + +type 'c nFormula = ('c pExprC, op1) prod + +type monoidMember = nat list + +type 'c coneMember = + | S_In of nat + | S_Ideal of 'c pExprC * 'c coneMember + | S_Square of 'c pExprC + | S_Monoid of monoidMember + | S_Mult of 'c coneMember * 'c coneMember + | S_Add of 'c coneMember * 'c coneMember + | S_Pos of 'c + | S_Z + +(** val nformula_times : 'a1 nFormula -> 'a1 nFormula -> 'a1 nFormula **) + +let nformula_times f f' = + let Pair (p, op) = f in + let Pair (p', op') = f' in + Pair ((PEmul (p, p')), + (match op with + | Equal -> Equal + | NonEqual -> NonEqual + | Strict -> op' + | NonStrict -> NonStrict)) + +(** val nformula_plus : 'a1 nFormula -> 'a1 nFormula -> 'a1 nFormula **) + +let nformula_plus f f' = + let Pair (p, op) = f in + let Pair (p', op') = f' in + Pair ((PEadd (p, p')), + (match op with + | Equal -> op' + | NonEqual -> NonEqual + | Strict -> Strict + | NonStrict -> (match op' with + | Strict -> Strict + | _ -> NonStrict))) + +(** val eval_monoid : + 'a1 -> 'a1 nFormula list -> monoidMember -> 'a1 pExprC **) + +let rec eval_monoid cI l = function + | Nil -> PEc cI + | Cons (n0, ns0) -> PEmul + ((let Pair (q0, o) = nth n0 l (Pair ((PEc cI), NonEqual)) in + (match o with + | NonEqual -> q0 + | _ -> PEc cI)), (eval_monoid cI l ns0)) + +(** val eval_cone : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 + nFormula list -> 'a1 coneMember -> 'a1 nFormula **) + +let rec eval_cone cO cI ceqb cleb l = function + | S_In n0 -> + let Pair (p, o) = nth n0 l (Pair ((PEc cO), Equal)) in + (match o with + | NonEqual -> Pair ((PEc cO), Equal) + | _ -> nth n0 l (Pair ((PEc cO), Equal))) + | S_Ideal (p, cm') -> + let f = eval_cone cO cI ceqb cleb l cm' in + let Pair (q0, op) = f in + (match op with + | Equal -> Pair ((PEmul (q0, p)), Equal) + | _ -> f) + | S_Square p -> Pair ((PEmul (p, p)), NonStrict) + | S_Monoid m -> let p = eval_monoid cI l m in Pair ((PEmul (p, p)), Strict) + | S_Mult (p, q0) -> + nformula_times (eval_cone cO cI ceqb cleb l p) + (eval_cone cO cI ceqb cleb l q0) + | S_Add (p, q0) -> + nformula_plus (eval_cone cO cI ceqb cleb l p) + (eval_cone cO cI ceqb cleb l q0) + | S_Pos c -> + (match match cleb cO c with + | True -> negb (ceqb cO c) + | False -> False with + | True -> Pair ((PEc c), Strict) + | False -> Pair ((PEc cO), Equal)) + | S_Z -> Pair ((PEc cO), Equal) + +(** val normalise_pexpr : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 + -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExprC -> 'a1 polC **) + +let normalise_pexpr cO cI cplus ctimes cminus copp ceqb x = + norm_aux cO cI cplus ctimes cminus copp ceqb x + +(** val check_inconsistent : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 + -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) + -> 'a1 nFormula -> bool **) + +let check_inconsistent cO cI cplus ctimes cminus copp ceqb cleb = function + | Pair (e, op) -> + (match normalise_pexpr cO cI cplus ctimes cminus copp ceqb e with + | Pc c -> + (match op with + | Equal -> negb (ceqb c cO) + | NonEqual -> False + | Strict -> cleb c cO + | NonStrict -> + (match cleb c cO with + | True -> negb (ceqb c cO) + | False -> False)) + | _ -> False) + +(** val check_normalised_formulas : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 + -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) + -> 'a1 nFormula list -> 'a1 coneMember -> bool **) + +let check_normalised_formulas cO cI cplus ctimes cminus copp ceqb cleb l cm = + check_inconsistent cO cI cplus ctimes cminus copp ceqb cleb + (eval_cone cO cI ceqb cleb l cm) + +type op2 = + | OpEq + | OpNEq + | OpLe + | OpGe + | OpLt + | OpGt + +type 'c formula = { flhs : 'c pExprC; fop : op2; frhs : 'c pExprC } + +(** val flhs : 'a1 formula -> 'a1 pExprC **) + +let flhs x = x.flhs + +(** val fop : 'a1 formula -> op2 **) + +let fop x = x.fop + +(** val frhs : 'a1 formula -> 'a1 pExprC **) + +let frhs x = x.frhs + +(** val xnormalise : 'a1 formula -> 'a1 nFormula list **) + +let xnormalise t0 = + let { flhs = lhs; fop = o; frhs = rhs } = t0 in + (match o with + | OpEq -> Cons ((Pair ((PEsub (lhs, rhs)), Strict)), (Cons ((Pair + ((PEsub (rhs, lhs)), Strict)), Nil))) + | OpNEq -> Cons ((Pair ((PEsub (lhs, rhs)), Equal)), Nil) + | OpLe -> Cons ((Pair ((PEsub (lhs, rhs)), Strict)), Nil) + | OpGe -> Cons ((Pair ((PEsub (rhs, lhs)), Strict)), Nil) + | OpLt -> Cons ((Pair ((PEsub (lhs, rhs)), NonStrict)), Nil) + | OpGt -> Cons ((Pair ((PEsub (rhs, lhs)), NonStrict)), Nil)) + +(** val cnf_normalise : 'a1 formula -> 'a1 nFormula cnf **) + +let cnf_normalise t0 = + map (fun x -> Cons (x, Nil)) (xnormalise t0) + +(** val xnegate : 'a1 formula -> 'a1 nFormula list **) + +let xnegate t0 = + let { flhs = lhs; fop = o; frhs = rhs } = t0 in + (match o with + | OpEq -> Cons ((Pair ((PEsub (lhs, rhs)), Equal)), Nil) + | OpNEq -> Cons ((Pair ((PEsub (lhs, rhs)), Strict)), (Cons ((Pair + ((PEsub (rhs, lhs)), Strict)), Nil))) + | OpLe -> Cons ((Pair ((PEsub (rhs, lhs)), NonStrict)), Nil) + | OpGe -> Cons ((Pair ((PEsub (lhs, rhs)), NonStrict)), Nil) + | OpLt -> Cons ((Pair ((PEsub (rhs, lhs)), Strict)), Nil) + | OpGt -> Cons ((Pair ((PEsub (lhs, rhs)), Strict)), Nil)) + +(** val cnf_negate : 'a1 formula -> 'a1 nFormula cnf **) + +let cnf_negate t0 = + map (fun x -> Cons (x, Nil)) (xnegate t0) + +(** val simpl_expr : + 'a1 -> ('a1 -> 'a1 -> bool) -> 'a1 pExprC -> 'a1 pExprC **) + +let rec simpl_expr cI ceqb e = match e with + | PEadd (x, y) -> PEadd ((simpl_expr cI ceqb x), (simpl_expr cI ceqb y)) + | PEmul (y, z0) -> + let y' = simpl_expr cI ceqb y in + (match y' with + | PEc c -> + (match ceqb c cI with + | True -> simpl_expr cI ceqb z0 + | False -> PEmul (y', (simpl_expr cI ceqb z0))) + | _ -> PEmul (y', (simpl_expr cI ceqb z0))) + | _ -> e + +(** val simpl_cone : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 + coneMember -> 'a1 coneMember **) + +let simpl_cone cO cI ctimes ceqb e = match e with + | S_Square t0 -> + (match simpl_expr cI ceqb t0 with + | PEc c -> + (match ceqb cO c with + | True -> S_Z + | False -> S_Pos (ctimes c c)) + | _ -> S_Square (simpl_expr cI ceqb t0)) + | S_Mult (t1, t2) -> + (match t1 with + | S_Mult (x, x0) -> + (match x with + | S_Pos p2 -> + (match t2 with + | S_Pos c -> S_Mult ((S_Pos (ctimes c p2)), x0) + | S_Z -> S_Z + | _ -> e) + | _ -> + (match x0 with + | S_Pos p2 -> + (match t2 with + | S_Pos c -> S_Mult ((S_Pos (ctimes c p2)), x) + | S_Z -> S_Z + | _ -> e) + | _ -> + (match t2 with + | S_Pos c -> + (match ceqb cI c with + | True -> t1 + | False -> S_Mult (t1, t2)) + | S_Z -> S_Z + | _ -> e))) + | S_Pos c -> + (match t2 with + | S_Mult (x, x0) -> + (match x with + | S_Pos p2 -> S_Mult ((S_Pos (ctimes c p2)), x0) + | _ -> + (match x0 with + | S_Pos p2 -> S_Mult ((S_Pos (ctimes c p2)), x) + | _ -> + (match ceqb cI c with + | True -> t2 + | False -> S_Mult (t1, t2)))) + | S_Add (y, z0) -> S_Add ((S_Mult ((S_Pos c), y)), (S_Mult + ((S_Pos c), z0))) + | S_Pos c0 -> S_Pos (ctimes c c0) + | S_Z -> S_Z + | _ -> + (match ceqb cI c with + | True -> t2 + | False -> S_Mult (t1, t2))) + | S_Z -> S_Z + | _ -> + (match t2 with + | S_Pos c -> + (match ceqb cI c with + | True -> t1 + | False -> S_Mult (t1, t2)) + | S_Z -> S_Z + | _ -> e)) + | S_Add (t1, t2) -> + (match t1 with + | S_Z -> t2 + | _ -> (match t2 with + | S_Z -> t1 + | _ -> S_Add (t1, t2))) + | _ -> e + +type q = { qnum : z; qden : positive } + +(** val qnum : q -> z **) + +let qnum x = x.qnum + +(** val qden : q -> positive **) + +let qden x = x.qden + +(** val qplus : q -> q -> q **) + +let qplus x y = + { qnum = (zplus (zmult x.qnum (Zpos y.qden)) (zmult y.qnum (Zpos x.qden))); + qden = (pmult x.qden y.qden) } + +(** val qmult : q -> q -> q **) + +let qmult x y = + { qnum = (zmult x.qnum y.qnum); qden = (pmult x.qden y.qden) } + +(** val qopp : q -> q **) + +let qopp x = + { qnum = (zopp x.qnum); qden = x.qden } + +(** val qminus : q -> q -> q **) + +let qminus x y = + qplus x (qopp y) + +type 'a t = + | Empty + | Leaf of 'a + | Node of 'a t * 'a * 'a t + +(** val find : 'a1 -> 'a1 t -> positive -> 'a1 **) + +let rec find default vm p = + match vm with + | Empty -> default + | Leaf i -> i + | Node (l, e, r) -> + (match p with + | XI p2 -> find default r p2 + | XO p2 -> find default l p2 + | XH -> e) + +type zWitness = z coneMember + +(** val zWeakChecker : z nFormula list -> z coneMember -> bool **) + +let zWeakChecker x x0 = + check_normalised_formulas Z0 (Zpos XH) zplus zmult zminus zopp zeq_bool + zle_bool x x0 + +(** val xnormalise0 : z formula -> z nFormula list **) + +let xnormalise0 t0 = + let { flhs = lhs; fop = o; frhs = rhs } = t0 in + (match o with + | OpEq -> Cons ((Pair ((PEsub (lhs, (PEadd (rhs, (PEc (Zpos XH)))))), + NonStrict)), (Cons ((Pair ((PEsub (rhs, (PEadd (lhs, (PEc (Zpos + XH)))))), NonStrict)), Nil))) + | OpNEq -> Cons ((Pair ((PEsub (lhs, rhs)), Equal)), Nil) + | OpLe -> Cons ((Pair ((PEsub (lhs, (PEadd (rhs, (PEc (Zpos XH)))))), + NonStrict)), Nil) + | OpGe -> Cons ((Pair ((PEsub (rhs, (PEadd (lhs, (PEc (Zpos XH)))))), + NonStrict)), Nil) + | OpLt -> Cons ((Pair ((PEsub (lhs, rhs)), NonStrict)), Nil) + | OpGt -> Cons ((Pair ((PEsub (rhs, lhs)), NonStrict)), Nil)) + +(** val normalise : z formula -> z nFormula cnf **) + +let normalise t0 = + map (fun x -> Cons (x, Nil)) (xnormalise0 t0) + +(** val xnegate0 : z formula -> z nFormula list **) + +let xnegate0 t0 = + let { flhs = lhs; fop = o; frhs = rhs } = t0 in + (match o with + | OpEq -> Cons ((Pair ((PEsub (lhs, rhs)), Equal)), Nil) + | OpNEq -> Cons ((Pair ((PEsub (lhs, (PEadd (rhs, (PEc (Zpos XH)))))), + NonStrict)), (Cons ((Pair ((PEsub (rhs, (PEadd (lhs, (PEc (Zpos + XH)))))), NonStrict)), Nil))) + | OpLe -> Cons ((Pair ((PEsub (rhs, lhs)), NonStrict)), Nil) + | OpGe -> Cons ((Pair ((PEsub (lhs, rhs)), NonStrict)), Nil) + | OpLt -> Cons ((Pair ((PEsub (rhs, (PEadd (lhs, (PEc (Zpos XH)))))), + NonStrict)), Nil) + | OpGt -> Cons ((Pair ((PEsub (lhs, (PEadd (rhs, (PEc (Zpos XH)))))), + NonStrict)), Nil)) + +(** val negate : z formula -> z nFormula cnf **) + +let negate t0 = + map (fun x -> Cons (x, Nil)) (xnegate0 t0) + +(** val ceiling : z -> z -> z **) + +let ceiling a b = + let Pair (q0, r) = zdiv_eucl a b in + (match r with + | Z0 -> q0 + | _ -> zplus q0 (Zpos XH)) + +type proofTerm = + | RatProof of zWitness + | CutProof of z pExprC * q * zWitness * proofTerm + | EnumProof of q * z pExprC * q * zWitness * zWitness * proofTerm list + +(** val makeLb : z pExpr -> q -> z nFormula **) + +let makeLb v q0 = + let { qnum = n0; qden = d } = q0 in + Pair ((PEsub ((PEmul ((PEc (Zpos d)), v)), (PEc n0))), NonStrict) + +(** val qceiling : q -> z **) + +let qceiling q0 = + let { qnum = n0; qden = d } = q0 in ceiling n0 (Zpos d) + +(** val makeLbCut : z pExprC -> q -> z nFormula **) + +let makeLbCut v q0 = + Pair ((PEsub (v, (PEc (qceiling q0)))), NonStrict) + +(** val neg_nformula : z nFormula -> (z pExpr, op1) prod **) + +let neg_nformula = function + | Pair (e, o) -> Pair ((PEopp (PEadd (e, (PEc (Zpos XH))))), o) + +(** val cutChecker : + z nFormula list -> z pExpr -> q -> zWitness -> z nFormula option **) + +let cutChecker l e lb pf = + match zWeakChecker (Cons ((neg_nformula (makeLb e lb)), l)) pf with + | True -> Some (makeLbCut e lb) + | False -> None + +(** val zChecker : z nFormula list -> proofTerm -> bool **) + +let rec zChecker l = function + | RatProof pf0 -> zWeakChecker l pf0 + | CutProof (e, q0, pf0, rst) -> + (match cutChecker l e q0 pf0 with + | Some c -> zChecker (Cons (c, l)) rst + | None -> False) + | EnumProof (lb, e, ub, pf1, pf2, rst) -> + (match cutChecker l e lb pf1 with + | Some n0 -> + (match cutChecker l (PEopp e) (qopp ub) pf2 with + | Some n1 -> + let rec label pfs lb0 ub0 = + match pfs with + | Nil -> + (match z_gt_dec lb0 ub0 with + | Left -> True + | Right -> False) + | Cons (pf0, rsr) -> + (match zChecker (Cons ((Pair ((PEsub (e, (PEc + lb0))), Equal)), l)) pf0 with + | True -> label rsr (zplus lb0 (Zpos XH)) ub0 + | False -> False) + in label rst (qceiling lb) (zopp (qceiling (qopp ub))) + | None -> False) + | None -> False) + +(** val zTautoChecker : z formula bFormula -> proofTerm list -> bool **) + +let zTautoChecker f w = + tauto_checker normalise negate zChecker f w + +(** val map_cone : (nat -> nat) -> zWitness -> zWitness **) + +let rec map_cone f e = match e with + | S_In n0 -> S_In (f n0) + | S_Ideal (e0, cm) -> S_Ideal (e0, (map_cone f cm)) + | S_Monoid l -> S_Monoid (map f l) + | S_Mult (cm1, cm2) -> S_Mult ((map_cone f cm1), (map_cone f cm2)) + | S_Add (cm1, cm2) -> S_Add ((map_cone f cm1), (map_cone f cm2)) + | _ -> e + +(** val indexes : zWitness -> nat list **) + +let rec indexes = function + | S_In n0 -> Cons (n0, Nil) + | S_Ideal (e0, cm) -> indexes cm + | S_Monoid l -> l + | S_Mult (cm1, cm2) -> app (indexes cm1) (indexes cm2) + | S_Add (cm1, cm2) -> app (indexes cm1) (indexes cm2) + | _ -> Nil + +(** val n_of_Z : z -> n **) + +let n_of_Z = function + | Zpos p -> Npos p + | _ -> N0 + +(** val qeq_bool : q -> q -> bool **) + +let qeq_bool p q0 = + zeq_bool (zmult p.qnum (Zpos q0.qden)) (zmult q0.qnum (Zpos p.qden)) + +(** val qle_bool : q -> q -> bool **) + +let qle_bool x y = + zle_bool (zmult x.qnum (Zpos y.qden)) (zmult y.qnum (Zpos x.qden)) + +type qWitness = q coneMember + +(** val qWeakChecker : q nFormula list -> q coneMember -> bool **) + +let qWeakChecker x x0 = + check_normalised_formulas { qnum = Z0; qden = XH } { qnum = (Zpos XH); + qden = XH } qplus qmult qminus qopp qeq_bool qle_bool x x0 + +(** val qTautoChecker : q formula bFormula -> qWitness list -> bool **) + +let qTautoChecker f w = + tauto_checker (fun x -> cnf_normalise x) (fun x -> + cnf_negate x) qWeakChecker f w + diff --git a/contrib/micromega/micromega.mli b/contrib/micromega/micromega.mli new file mode 100644 index 00000000..f94f091e --- /dev/null +++ b/contrib/micromega/micromega.mli @@ -0,0 +1,398 @@ +type __ = Obj.t + +type bool = + | True + | False + +val negb : bool -> bool + +type nat = + | O + | S of nat + +type 'a option = + | Some of 'a + | None + +type ('a, 'b) prod = + | Pair of 'a * 'b + +type comparison = + | Eq + | Lt + | Gt + +val compOpp : comparison -> comparison + +type sumbool = + | Left + | Right + +type 'a sumor = + | Inleft of 'a + | Inright + +type 'a list = + | Nil + | Cons of 'a * 'a list + +val app : 'a1 list -> 'a1 list -> 'a1 list + +val nth : nat -> 'a1 list -> 'a1 -> 'a1 + +val map : ('a1 -> 'a2) -> 'a1 list -> 'a2 list + +type positive = + | XI of positive + | XO of positive + | XH + +val psucc : positive -> positive + +val pplus : positive -> positive -> positive + +val pplus_carry : positive -> positive -> positive + +val p_of_succ_nat : nat -> positive + +val pdouble_minus_one : positive -> positive + +type positive_mask = + | IsNul + | IsPos of positive + | IsNeg + +val pdouble_plus_one_mask : positive_mask -> positive_mask + +val pdouble_mask : positive_mask -> positive_mask + +val pdouble_minus_two : positive -> positive_mask + +val pminus_mask : positive -> positive -> positive_mask + +val pminus_mask_carry : positive -> positive -> positive_mask + +val pminus : positive -> positive -> positive + +val pmult : positive -> positive -> positive + +val pcompare : positive -> positive -> comparison -> comparison + +type n = + | N0 + | Npos of positive + +type z = + | Z0 + | Zpos of positive + | Zneg of positive + +val zdouble_plus_one : z -> z + +val zdouble_minus_one : z -> z + +val zdouble : z -> z + +val zPminus : positive -> positive -> z + +val zplus : z -> z -> z + +val zopp : z -> z + +val zminus : z -> z -> z + +val zmult : z -> z -> z + +val zcompare : z -> z -> comparison + +val dcompare_inf : comparison -> sumbool sumor + +val zcompare_rec : z -> z -> (__ -> 'a1) -> (__ -> 'a1) -> (__ -> 'a1) -> 'a1 + +val z_gt_dec : z -> z -> sumbool + +val zle_bool : z -> z -> bool + +val zge_bool : z -> z -> bool + +val zgt_bool : z -> z -> bool + +val zeq_bool : z -> z -> bool + +val n_of_nat : nat -> n + +val zdiv_eucl_POS : positive -> z -> (z, z) prod + +val zdiv_eucl : z -> z -> (z, z) prod + +type 'c pol = + | Pc of 'c + | Pinj of positive * 'c pol + | PX of 'c pol * positive * 'c pol + +val p0 : 'a1 -> 'a1 pol + +val p1 : 'a1 -> 'a1 pol + +val peq : ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> bool + +val mkPinj_pred : positive -> 'a1 pol -> 'a1 pol + +val mkPX : + 'a1 -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol + +val mkXi : 'a1 -> 'a1 -> positive -> 'a1 pol + +val mkX : 'a1 -> 'a1 -> 'a1 pol + +val popp : ('a1 -> 'a1) -> 'a1 pol -> 'a1 pol + +val paddC : ('a1 -> 'a1 -> 'a1) -> 'a1 pol -> 'a1 -> 'a1 pol + +val psubC : ('a1 -> 'a1 -> 'a1) -> 'a1 pol -> 'a1 -> 'a1 pol + +val paddI : + ('a1 -> 'a1 -> 'a1) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1 pol -> + positive -> 'a1 pol -> 'a1 pol + +val psubI : + ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> + 'a1 pol -> positive -> 'a1 pol -> 'a1 pol + +val paddX : + 'a1 -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1 pol + -> positive -> 'a1 pol -> 'a1 pol + +val psubX : + 'a1 -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol -> 'a1 + pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol + +val padd : + 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> + 'a1 pol + +val psub : + 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 + -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol + +val pmulC_aux : + 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 -> 'a1 + pol + +val pmulC : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 + -> 'a1 pol + +val pmulI : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> + 'a1 pol -> 'a1 pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol + +val pmul : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> + bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol + +type 'c pExpr = + | PEc of 'c + | PEX of positive + | PEadd of 'c pExpr * 'c pExpr + | PEsub of 'c pExpr * 'c pExpr + | PEmul of 'c pExpr * 'c pExpr + | PEopp of 'c pExpr + | PEpow of 'c pExpr * n + +val mk_X : 'a1 -> 'a1 -> positive -> 'a1 pol + +val ppow_pos : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> + bool) -> ('a1 pol -> 'a1 pol) -> 'a1 pol -> 'a1 pol -> positive -> 'a1 pol + +val ppow_N : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> + bool) -> ('a1 pol -> 'a1 pol) -> 'a1 pol -> n -> 'a1 pol + +val norm_aux : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> + 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExpr -> 'a1 pol + +type 'a bFormula = + | TT + | FF + | X + | A of 'a + | Cj of 'a bFormula * 'a bFormula + | D of 'a bFormula * 'a bFormula + | N of 'a bFormula + | I of 'a bFormula * 'a bFormula + +type 'term' clause = 'term' list + +type 'term' cnf = 'term' clause list + +val tt : 'a1 cnf + +val ff : 'a1 cnf + +val or_clause_cnf : 'a1 clause -> 'a1 cnf -> 'a1 cnf + +val or_cnf : 'a1 cnf -> 'a1 cnf -> 'a1 cnf + +val and_cnf : 'a1 cnf -> 'a1 cnf -> 'a1 cnf + +val xcnf : + ('a1 -> 'a2 cnf) -> ('a1 -> 'a2 cnf) -> bool -> 'a1 bFormula -> 'a2 cnf + +val cnf_checker : ('a1 list -> 'a2 -> bool) -> 'a1 cnf -> 'a2 list -> bool + +val tauto_checker : + ('a1 -> 'a2 cnf) -> ('a1 -> 'a2 cnf) -> ('a2 list -> 'a3 -> bool) -> 'a1 + bFormula -> 'a3 list -> bool + +type 'c pExprC = 'c pExpr + +type 'c polC = 'c pol + +type op1 = + | Equal + | NonEqual + | Strict + | NonStrict + +type 'c nFormula = ('c pExprC, op1) prod + +type monoidMember = nat list + +type 'c coneMember = + | S_In of nat + | S_Ideal of 'c pExprC * 'c coneMember + | S_Square of 'c pExprC + | S_Monoid of monoidMember + | S_Mult of 'c coneMember * 'c coneMember + | S_Add of 'c coneMember * 'c coneMember + | S_Pos of 'c + | S_Z + +val nformula_times : 'a1 nFormula -> 'a1 nFormula -> 'a1 nFormula + +val nformula_plus : 'a1 nFormula -> 'a1 nFormula -> 'a1 nFormula + +val eval_monoid : 'a1 -> 'a1 nFormula list -> monoidMember -> 'a1 pExprC + +val eval_cone : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula + list -> 'a1 coneMember -> 'a1 nFormula + +val normalise_pexpr : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> + 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExprC -> 'a1 polC + +val check_inconsistent : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> + 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 + nFormula -> bool + +val check_normalised_formulas : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> + 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 + nFormula list -> 'a1 coneMember -> bool + +type op2 = + | OpEq + | OpNEq + | OpLe + | OpGe + | OpLt + | OpGt + +type 'c formula = { flhs : 'c pExprC; fop : op2; frhs : 'c pExprC } + +val flhs : 'a1 formula -> 'a1 pExprC + +val fop : 'a1 formula -> op2 + +val frhs : 'a1 formula -> 'a1 pExprC + +val xnormalise : 'a1 formula -> 'a1 nFormula list + +val cnf_normalise : 'a1 formula -> 'a1 nFormula cnf + +val xnegate : 'a1 formula -> 'a1 nFormula list + +val cnf_negate : 'a1 formula -> 'a1 nFormula cnf + +val simpl_expr : 'a1 -> ('a1 -> 'a1 -> bool) -> 'a1 pExprC -> 'a1 pExprC + +val simpl_cone : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 coneMember + -> 'a1 coneMember + +type q = { qnum : z; qden : positive } + +val qnum : q -> z + +val qden : q -> positive + +val qplus : q -> q -> q + +val qmult : q -> q -> q + +val qopp : q -> q + +val qminus : q -> q -> q + +type 'a t = + | Empty + | Leaf of 'a + | Node of 'a t * 'a * 'a t + +val find : 'a1 -> 'a1 t -> positive -> 'a1 + +type zWitness = z coneMember + +val zWeakChecker : z nFormula list -> z coneMember -> bool + +val xnormalise0 : z formula -> z nFormula list + +val normalise : z formula -> z nFormula cnf + +val xnegate0 : z formula -> z nFormula list + +val negate : z formula -> z nFormula cnf + +val ceiling : z -> z -> z + +type proofTerm = + | RatProof of zWitness + | CutProof of z pExprC * q * zWitness * proofTerm + | EnumProof of q * z pExprC * q * zWitness * zWitness * proofTerm list + +val makeLb : z pExpr -> q -> z nFormula + +val qceiling : q -> z + +val makeLbCut : z pExprC -> q -> z nFormula + +val neg_nformula : z nFormula -> (z pExpr, op1) prod + +val cutChecker : + z nFormula list -> z pExpr -> q -> zWitness -> z nFormula option + +val zChecker : z nFormula list -> proofTerm -> bool + +val zTautoChecker : z formula bFormula -> proofTerm list -> bool + +val map_cone : (nat -> nat) -> zWitness -> zWitness + +val indexes : zWitness -> nat list + +val n_of_Z : z -> n + +val qeq_bool : q -> q -> bool + +val qle_bool : q -> q -> bool + +type qWitness = q coneMember + +val qWeakChecker : q nFormula list -> q coneMember -> bool + +val qTautoChecker : q formula bFormula -> qWitness list -> bool + diff --git a/contrib/micromega/mutils.ml b/contrib/micromega/mutils.ml new file mode 100644 index 00000000..2473608f --- /dev/null +++ b/contrib/micromega/mutils.ml @@ -0,0 +1,305 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* *) +(* Micromega: A reflexive tactic using the Positivstellensatz *) +(* *) +(* Frédéric Besson (Irisa/Inria) 2006-2008 *) +(* *) +(************************************************************************) + +let debug = false + +let fst' (Micromega.Pair(x,y)) = x +let snd' (Micromega.Pair(x,y)) = y + +let rec try_any l x = + match l with + | [] -> None + | (f,s)::l -> match f x with + | None -> try_any l x + | x -> x + +let list_try_find f = + let rec try_find_f = function + | [] -> failwith "try_find" + | h::t -> try f h with Failure _ -> try_find_f t + in + try_find_f + +let rec list_fold_right_elements f l = + let rec aux = function + | [] -> invalid_arg "list_fold_right_elements" + | [x] -> x + | x::l -> f x (aux l) in + aux l + +let interval n m = + let rec interval_n (l,m) = + if n > m then l else interval_n (m::l,pred m) + in + interval_n ([],m) + +open Num +open Big_int + +let ppcm x y = + let g = gcd_big_int x y in + let x' = div_big_int x g in + let y' = div_big_int y g in + mult_big_int g (mult_big_int x' y') + + +let denominator = function + | Int _ | Big_int _ -> unit_big_int + | Ratio r -> Ratio.denominator_ratio r + +let numerator = function + | Ratio r -> Ratio.numerator_ratio r + | Int i -> Big_int.big_int_of_int i + | Big_int i -> i + +let rec ppcm_list c l = + match l with + | [] -> c + | e::l -> ppcm_list (ppcm c (denominator e)) l + +let rec rec_gcd_list c l = + match l with + | [] -> c + | e::l -> rec_gcd_list (gcd_big_int c (numerator e)) l + +let rec gcd_list l = + let res = rec_gcd_list zero_big_int l in + if compare_big_int res zero_big_int = 0 + then unit_big_int else res + + + +let rats_to_ints l = + let c = ppcm_list unit_big_int l in + List.map (fun x -> (div_big_int (mult_big_int (numerator x) c) + (denominator x))) l + +(* Nasty reordering of lists - useful to trim certificate down *) +let mapi f l = + let rec xmapi i l = + match l with + | [] -> [] + | e::l -> (f e i)::(xmapi (i+1) l) in + xmapi 0 l + + +let concatMapi f l = List.rev (mapi (fun e i -> (i,f e)) l) + +(* assoc_pos j [a0...an] = [j,a0....an,j+n],j+n+1 *) +let assoc_pos j l = (mapi (fun e i -> e,i+j) l, j + (List.length l)) + +let assoc_pos_assoc l = + let rec xpos i l = + match l with + | [] -> [] + | (x,l) ::rst -> let (l',j) = assoc_pos i l in + (x,l')::(xpos j rst) in + xpos 0 l + +let filter_pos f l = + (* Could sort ... take care of duplicates... *) + let rec xfilter l = + match l with + | [] -> [] + | (x,e)::l -> + if List.exists (fun ee -> List.mem ee f) (List.map snd e) + then (x,e)::(xfilter l) + else xfilter l in + xfilter l + +let select_pos lpos l = + let rec xselect i lpos l = + match lpos with + | [] -> [] + | j::rpos -> + match l with + | [] -> failwith "select_pos" + | e::l -> + if i = j + then e:: (xselect (i+1) rpos l) + else xselect (i+1) lpos l in + xselect 0 lpos l + + +module CoqToCaml = +struct + open Micromega + + let rec nat = function + | O -> 0 + | S n -> (nat n) + 1 + + + let rec positive p = + match p with + | XH -> 1 + | XI p -> 1+ 2*(positive p) + | XO p -> 2*(positive p) + + + let n nt = + match nt with + | N0 -> 0 + | Npos p -> positive p + + + let rec index i = (* Swap left-right ? *) + match i with + | XH -> 1 + | XI i -> 1+(2*(index i)) + | XO i -> 2*(index i) + + + let z x = + match x with + | Z0 -> 0 + | Zpos p -> (positive p) + | Zneg p -> - (positive p) + + open Big_int + + let rec positive_big_int p = + match p with + | XH -> unit_big_int + | XI p -> add_int_big_int 1 (mult_int_big_int 2 (positive_big_int p)) + | XO p -> (mult_int_big_int 2 (positive_big_int p)) + + + let z_big_int x = + match x with + | Z0 -> zero_big_int + | Zpos p -> (positive_big_int p) + | Zneg p -> minus_big_int (positive_big_int p) + + + let num x = Num.Big_int (z_big_int x) + + let rec list elt l = + match l with + | Nil -> [] + | Cons(e,l) -> (elt e)::(list elt l) + + let q_to_num {qnum = x ; qden = y} = + Big_int (z_big_int x) // (Big_int (z_big_int (Zpos y))) + +end + + +module CamlToCoq = +struct + open Micromega + + let rec nat = function + | 0 -> O + | n -> S (nat (n-1)) + + + let rec positive n = + if n=1 then XH + else if n land 1 = 1 then XI (positive (n lsr 1)) + else XO (positive (n lsr 1)) + + let n nt = + if nt < 0 + then assert false + else if nt = 0 then N0 + else Npos (positive nt) + + + + + + let rec index n = + if n=1 then XH + else if n land 1 = 1 then XI (index (n lsr 1)) + else XO (index (n lsr 1)) + + + let idx n = + (*a.k.a path_of_int *) + (* returns the list of digits of n in reverse order with + initial 1 removed *) + let rec digits_of_int n = + if n=1 then [] + else (n mod 2 = 1)::(digits_of_int (n lsr 1)) + in + List.fold_right + (fun b c -> (if b then XI c else XO c)) + (List.rev (digits_of_int n)) + (XH) + + + + let z x = + match compare x 0 with + | 0 -> Z0 + | 1 -> Zpos (positive x) + | _ -> (* this should be -1 *) + Zneg (positive (-x)) + + open Big_int + + let positive_big_int n = + let two = big_int_of_int 2 in + let rec _pos n = + if eq_big_int n unit_big_int then XH + else + let (q,m) = quomod_big_int n two in + if eq_big_int unit_big_int m + then XI (_pos q) + else XO (_pos q) in + _pos n + + let bigint x = + match sign_big_int x with + | 0 -> Z0 + | 1 -> Zpos (positive_big_int x) + | _ -> Zneg (positive_big_int (minus_big_int x)) + + let q n = + {Micromega.qnum = bigint (numerator n) ; + Micromega.qden = positive_big_int (denominator n)} + + + let list elt l = List.fold_right (fun x l -> Cons(elt x, l)) l Nil + +end + +module Cmp = +struct + + let rec compare_lexical l = + match l with + | [] -> 0 (* Equal *) + | f::l -> + let cmp = f () in + if cmp = 0 then compare_lexical l else cmp + + let rec compare_list cmp l1 l2 = + match l1 , l2 with + | [] , [] -> 0 + | [] , _ -> -1 + | _ , [] -> 1 + | e1::l1 , e2::l2 -> + let c = cmp e1 e2 in + if c = 0 then compare_list cmp l1 l2 else c + + let hash_list hash l = + let rec _hash_list l h = + match l with + | [] -> h lxor (Hashtbl.hash []) + | e::l -> _hash_list l ((hash e) lxor h) in + + _hash_list l 0 +end diff --git a/contrib/micromega/sos.ml b/contrib/micromega/sos.ml new file mode 100644 index 00000000..e3d72ed9 --- /dev/null +++ b/contrib/micromega/sos.ml @@ -0,0 +1,1919 @@ +(* ========================================================================= *) +(* - This code originates from John Harrison's HOL LIGHT 2.20 *) +(* (see file LICENSE.sos for license, copyright and disclaimer) *) +(* - Laurent Théry (thery@sophia.inria.fr) has isolated the HOL *) +(* independent bits *) +(* - Frédéric Besson (fbesson@irisa.fr) is using it to feed micromega *) +(* - Addition of a csdp cache by the Coq development team *) +(* ========================================================================= *) + +(* ========================================================================= *) +(* Nonlinear universal reals procedure using SOS decomposition. *) +(* ========================================================================= *) + +open Num;; +open List;; + +let debugging = ref false;; + +exception Sanity;; + +exception Unsolvable;; + +(* ------------------------------------------------------------------------- *) +(* Comparisons that are reflexive on NaN and also short-circuiting. *) +(* ------------------------------------------------------------------------- *) + +let (=?) = fun x y -> Pervasives.compare x y = 0;; +let (<?) = fun x y -> Pervasives.compare x y < 0;; +let (<=?) = fun x y -> Pervasives.compare x y <= 0;; +let (>?) = fun x y -> Pervasives.compare x y > 0;; +let (>=?) = fun x y -> Pervasives.compare x y >= 0;; + +(* ------------------------------------------------------------------------- *) +(* Combinators. *) +(* ------------------------------------------------------------------------- *) + +let (o) = fun f g x -> f(g x);; + +(* ------------------------------------------------------------------------- *) +(* Some useful functions on "num" type. *) +(* ------------------------------------------------------------------------- *) + + +let num_0 = Int 0 +and num_1 = Int 1 +and num_2 = Int 2 +and num_10 = Int 10;; + +let pow2 n = power_num num_2 (Int n);; +let pow10 n = power_num num_10 (Int n);; + +let numdom r = + let r' = Ratio.normalize_ratio (ratio_of_num r) in + num_of_big_int(Ratio.numerator_ratio r'), + num_of_big_int(Ratio.denominator_ratio r');; + +let numerator = (o) fst numdom +and denominator = (o) snd numdom;; + +let gcd_num n1 n2 = + num_of_big_int(Big_int.gcd_big_int (big_int_of_num n1) (big_int_of_num n2));; + +let lcm_num x y = + if x =/ num_0 & y =/ num_0 then num_0 + else abs_num((x */ y) // gcd_num x y);; + + +(* ------------------------------------------------------------------------- *) +(* List basics. *) +(* ------------------------------------------------------------------------- *) + +let rec el n l = + if n = 0 then hd l else el (n - 1) (tl l);; + + +(* ------------------------------------------------------------------------- *) +(* Various versions of list iteration. *) +(* ------------------------------------------------------------------------- *) + +let rec itlist f l b = + match l with + [] -> b + | (h::t) -> f h (itlist f t b);; + +let rec end_itlist f l = + match l with + [] -> failwith "end_itlist" + | [x] -> x + | (h::t) -> f h (end_itlist f t);; + +let rec itlist2 f l1 l2 b = + match (l1,l2) with + ([],[]) -> b + | (h1::t1,h2::t2) -> f h1 h2 (itlist2 f t1 t2 b) + | _ -> failwith "itlist2";; + +(* ------------------------------------------------------------------------- *) +(* All pairs arising from applying a function over two lists. *) +(* ------------------------------------------------------------------------- *) + +let rec allpairs f l1 l2 = + match l1 with + h1::t1 -> itlist (fun x a -> f h1 x :: a) l2 (allpairs f t1 l2) + | [] -> [];; + +(* ------------------------------------------------------------------------- *) +(* String operations (surely there is a better way...) *) +(* ------------------------------------------------------------------------- *) + +let implode l = itlist (^) l "";; + +let explode s = + let rec exap n l = + if n < 0 then l else + exap (n - 1) ((String.sub s n 1)::l) in + exap (String.length s - 1) [];; + + +(* ------------------------------------------------------------------------- *) +(* Attempting function or predicate applications. *) +(* ------------------------------------------------------------------------- *) + +let can f x = try (f x; true) with Failure _ -> false;; + + +(* ------------------------------------------------------------------------- *) +(* Repetition of a function. *) +(* ------------------------------------------------------------------------- *) + +let rec funpow n f x = + if n < 1 then x else funpow (n-1) f (f x);; + + +(* ------------------------------------------------------------------------- *) +(* term?? *) +(* ------------------------------------------------------------------------- *) + +type vname = string;; + +type term = +| Zero +| Const of Num.num +| Var of vname +| Inv of term +| Opp of term +| Add of (term * term) +| Sub of (term * term) +| Mul of (term * term) +| Div of (term * term) +| Pow of (term * int);; + + +(* ------------------------------------------------------------------------- *) +(* Data structure for Positivstellensatz refutations. *) +(* ------------------------------------------------------------------------- *) + +type positivstellensatz = + Axiom_eq of int + | Axiom_le of int + | Axiom_lt of int + | Rational_eq of num + | Rational_le of num + | Rational_lt of num + | Square of term + | Monoid of int list + | Eqmul of term * positivstellensatz + | Sum of positivstellensatz * positivstellensatz + | Product of positivstellensatz * positivstellensatz;; + + + +(* ------------------------------------------------------------------------- *) +(* Replication and sequences. *) +(* ------------------------------------------------------------------------- *) + +let rec replicate x n = + if n < 1 then [] + else x::(replicate x (n - 1));; + +let rec (--) = fun m n -> if m > n then [] else m::((m + 1) -- n);; + +(* ------------------------------------------------------------------------- *) +(* Various useful list operations. *) +(* ------------------------------------------------------------------------- *) + +let rec forall p l = + match l with + [] -> true + | h::t -> p(h) & forall p t;; + +let rec tryfind f l = + match l with + [] -> failwith "tryfind" + | (h::t) -> try f h with Failure _ -> tryfind f t;; + +let index x = + let rec ind n l = + match l with + [] -> failwith "index" + | (h::t) -> if x =? h then n else ind (n + 1) t in + ind 0;; + +(* ------------------------------------------------------------------------- *) +(* "Set" operations on lists. *) +(* ------------------------------------------------------------------------- *) + +let rec mem x lis = + match lis with + [] -> false + | (h::t) -> x =? h or mem x t;; + +let insert x l = + if mem x l then l else x::l;; + +let union l1 l2 = itlist insert l1 l2;; + +let subtract l1 l2 = filter (fun x -> not (mem x l2)) l1;; + +(* ------------------------------------------------------------------------- *) +(* Merging and bottom-up mergesort. *) +(* ------------------------------------------------------------------------- *) + +let rec merge ord l1 l2 = + match l1 with + [] -> l2 + | h1::t1 -> match l2 with + [] -> l1 + | h2::t2 -> if ord h1 h2 then h1::(merge ord t1 l2) + else h2::(merge ord l1 t2);; + + +(* ------------------------------------------------------------------------- *) +(* Common measure predicates to use with "sort". *) +(* ------------------------------------------------------------------------- *) + +let increasing f x y = f x <? f y;; + +let decreasing f x y = f x >? f y;; + +(* ------------------------------------------------------------------------- *) +(* Zipping, unzipping etc. *) +(* ------------------------------------------------------------------------- *) + +let rec zip l1 l2 = + match (l1,l2) with + ([],[]) -> [] + | (h1::t1,h2::t2) -> (h1,h2)::(zip t1 t2) + | _ -> failwith "zip";; + +let rec unzip = + function [] -> [],[] + | ((a,b)::rest) -> let alist,blist = unzip rest in + (a::alist,b::blist);; + +(* ------------------------------------------------------------------------- *) +(* Iterating functions over lists. *) +(* ------------------------------------------------------------------------- *) + +let rec do_list f l = + match l with + [] -> () + | (h::t) -> (f h; do_list f t);; + +(* ------------------------------------------------------------------------- *) +(* Sorting. *) +(* ------------------------------------------------------------------------- *) + +let rec sort cmp lis = + match lis with + [] -> [] + | piv::rest -> + let r,l = partition (cmp piv) rest in + (sort cmp l) @ (piv::(sort cmp r));; + +(* ------------------------------------------------------------------------- *) +(* Removing adjacent (NB!) equal elements from list. *) +(* ------------------------------------------------------------------------- *) + +let rec uniq l = + match l with + x::(y::_ as t) -> let t' = uniq t in + if x =? y then t' else + if t'==t then l else x::t' + | _ -> l;; + +(* ------------------------------------------------------------------------- *) +(* Convert list into set by eliminating duplicates. *) +(* ------------------------------------------------------------------------- *) + +let setify s = uniq (sort (<=?) s);; + +(* ------------------------------------------------------------------------- *) +(* Polymorphic finite partial functions via Patricia trees. *) +(* *) +(* The point of this strange representation is that it is canonical (equal *) +(* functions have the same encoding) yet reasonably efficient on average. *) +(* *) +(* Idea due to Diego Olivier Fernandez Pons (OCaml list, 2003/11/10). *) +(* ------------------------------------------------------------------------- *) + +type ('a,'b)func = + Empty + | Leaf of int * ('a*'b)list + | Branch of int * int * ('a,'b)func * ('a,'b)func;; + +(* ------------------------------------------------------------------------- *) +(* Undefined function. *) +(* ------------------------------------------------------------------------- *) + +let undefined = Empty;; + +(* ------------------------------------------------------------------------- *) +(* In case of equality comparison worries, better use this. *) +(* ------------------------------------------------------------------------- *) + +let is_undefined f = + match f with + Empty -> true + | _ -> false;; + +(* ------------------------------------------------------------------------- *) +(* Operation analagous to "map" for lists. *) +(* ------------------------------------------------------------------------- *) + +let mapf = + let rec map_list f l = + match l with + [] -> [] + | (x,y)::t -> (x,f(y))::(map_list f t) in + let rec mapf f t = + match t with + Empty -> Empty + | Leaf(h,l) -> Leaf(h,map_list f l) + | Branch(p,b,l,r) -> Branch(p,b,mapf f l,mapf f r) in + mapf;; + +(* ------------------------------------------------------------------------- *) +(* Operations analogous to "fold" for lists. *) +(* ------------------------------------------------------------------------- *) + +let foldl = + let rec foldl_list f a l = + match l with + [] -> a + | (x,y)::t -> foldl_list f (f a x y) t in + let rec foldl f a t = + match t with + Empty -> a + | Leaf(h,l) -> foldl_list f a l + | Branch(p,b,l,r) -> foldl f (foldl f a l) r in + foldl;; + +let foldr = + let rec foldr_list f l a = + match l with + [] -> a + | (x,y)::t -> f x y (foldr_list f t a) in + let rec foldr f t a = + match t with + Empty -> a + | Leaf(h,l) -> foldr_list f l a + | Branch(p,b,l,r) -> foldr f l (foldr f r a) in + foldr;; + +(* ------------------------------------------------------------------------- *) +(* Redefinition and combination. *) +(* ------------------------------------------------------------------------- *) + +let (|->),combine = + let ldb x y = let z = x lxor y in z land (-z) in + let newbranch p1 t1 p2 t2 = + let b = ldb p1 p2 in + let p = p1 land (b - 1) in + if p1 land b = 0 then Branch(p,b,t1,t2) + else Branch(p,b,t2,t1) in + let rec define_list (x,y as xy) l = + match l with + (a,b as ab)::t -> + if x =? a then xy::t + else if x <? a then xy::l + else ab::(define_list xy t) + | [] -> [xy] + and combine_list op z l1 l2 = + match (l1,l2) with + [],_ -> l2 + | _,[] -> l1 + | ((x1,y1 as xy1)::t1,(x2,y2 as xy2)::t2) -> + if x1 <? x2 then xy1::(combine_list op z t1 l2) + else if x2 <? x1 then xy2::(combine_list op z l1 t2) else + let y = op y1 y2 and l = combine_list op z t1 t2 in + if z(y) then l else (x1,y)::l in + let (|->) x y = + let k = Hashtbl.hash x in + let rec upd t = + match t with + Empty -> Leaf (k,[x,y]) + | Leaf(h,l) -> + if h = k then Leaf(h,define_list (x,y) l) + else newbranch h t k (Leaf(k,[x,y])) + | Branch(p,b,l,r) -> + if k land (b - 1) <> p then newbranch p t k (Leaf(k,[x,y])) + else if k land b = 0 then Branch(p,b,upd l,r) + else Branch(p,b,l,upd r) in + upd in + let rec combine op z t1 t2 = + match (t1,t2) with + Empty,_ -> t2 + | _,Empty -> t1 + | Leaf(h1,l1),Leaf(h2,l2) -> + if h1 = h2 then + let l = combine_list op z l1 l2 in + if l = [] then Empty else Leaf(h1,l) + else newbranch h1 t1 h2 t2 + | (Leaf(k,lis) as lf),(Branch(p,b,l,r) as br) | + (Branch(p,b,l,r) as br),(Leaf(k,lis) as lf) -> + if k land (b - 1) = p then + if k land b = 0 then + let l' = combine op z lf l in + if is_undefined l' then r else Branch(p,b,l',r) + else + let r' = combine op z lf r in + if is_undefined r' then l else Branch(p,b,l,r') + else + newbranch k lf p br + | Branch(p1,b1,l1,r1),Branch(p2,b2,l2,r2) -> + if b1 < b2 then + if p2 land (b1 - 1) <> p1 then newbranch p1 t1 p2 t2 + else if p2 land b1 = 0 then + let l = combine op z l1 t2 in + if is_undefined l then r1 else Branch(p1,b1,l,r1) + else + let r = combine op z r1 t2 in + if is_undefined r then l1 else Branch(p1,b1,l1,r) + else if b2 < b1 then + if p1 land (b2 - 1) <> p2 then newbranch p1 t1 p2 t2 + else if p1 land b2 = 0 then + let l = combine op z t1 l2 in + if is_undefined l then r2 else Branch(p2,b2,l,r2) + else + let r = combine op z t1 r2 in + if is_undefined r then l2 else Branch(p2,b2,l2,r) + else if p1 = p2 then + let l = combine op z l1 l2 and r = combine op z r1 r2 in + if is_undefined l then r + else if is_undefined r then l else Branch(p1,b1,l,r) + else + newbranch p1 t1 p2 t2 in + (|->),combine;; + +(* ------------------------------------------------------------------------- *) +(* Special case of point function. *) +(* ------------------------------------------------------------------------- *) + +let (|=>) = fun x y -> (x |-> y) undefined;; + + +(* ------------------------------------------------------------------------- *) +(* Grab an arbitrary element. *) +(* ------------------------------------------------------------------------- *) + +let rec choose t = + match t with + Empty -> failwith "choose: completely undefined function" + | Leaf(h,l) -> hd l + | Branch(b,p,t1,t2) -> choose t1;; + +(* ------------------------------------------------------------------------- *) +(* Application. *) +(* ------------------------------------------------------------------------- *) + +let applyd = + let rec apply_listd l d x = + match l with + (a,b)::t -> if x =? a then b + else if x >? a then apply_listd t d x else d x + | [] -> d x in + fun f d x -> + let k = Hashtbl.hash x in + let rec look t = + match t with + Leaf(h,l) when h = k -> apply_listd l d x + | Branch(p,b,l,r) -> look (if k land b = 0 then l else r) + | _ -> d x in + look f;; + +let apply f = applyd f (fun x -> failwith "apply");; + +let tryapplyd f a d = applyd f (fun x -> d) a;; + +let defined f x = try apply f x; true with Failure _ -> false;; + +(* ------------------------------------------------------------------------- *) +(* Undefinition. *) +(* ------------------------------------------------------------------------- *) + +let undefine = + let rec undefine_list x l = + match l with + (a,b as ab)::t -> + if x =? a then t + else if x <? a then l else + let t' = undefine_list x t in + if t' == t then l else ab::t' + | [] -> [] in + fun x -> + let k = Hashtbl.hash x in + let rec und t = + match t with + Leaf(h,l) when h = k -> + let l' = undefine_list x l in + if l' == l then t + else if l' = [] then Empty + else Leaf(h,l') + | Branch(p,b,l,r) when k land (b - 1) = p -> + if k land b = 0 then + let l' = und l in + if l' == l then t + else if is_undefined l' then r + else Branch(p,b,l',r) + else + let r' = und r in + if r' == r then t + else if is_undefined r' then l + else Branch(p,b,l,r') + | _ -> t in + und;; + + +(* ------------------------------------------------------------------------- *) +(* Mapping to sorted-list representation of the graph, domain and range. *) +(* ------------------------------------------------------------------------- *) + +let graph f = setify (foldl (fun a x y -> (x,y)::a) [] f);; + +let dom f = setify(foldl (fun a x y -> x::a) [] f);; + +let ran f = setify(foldl (fun a x y -> y::a) [] f);; + +(* ------------------------------------------------------------------------- *) +(* Turn a rational into a decimal string with d sig digits. *) +(* ------------------------------------------------------------------------- *) + +let decimalize = + let rec normalize y = + if abs_num y </ Int 1 // Int 10 then normalize (Int 10 */ y) - 1 + else if abs_num y >=/ Int 1 then normalize (y // Int 10) + 1 + else 0 in + fun d x -> + if x =/ Int 0 then "0.0" else + let y = abs_num x in + let e = normalize y in + let z = pow10(-e) */ y +/ Int 1 in + let k = round_num(pow10 d */ z) in + (if x </ Int 0 then "-0." else "0.") ^ + implode(tl(explode(string_of_num k))) ^ + (if e = 0 then "" else "e"^string_of_int e);; + + +(* ------------------------------------------------------------------------- *) +(* Iterations over numbers, and lists indexed by numbers. *) +(* ------------------------------------------------------------------------- *) + +let rec itern k l f a = + match l with + [] -> a + | h::t -> itern (k + 1) t f (f h k a);; + +let rec iter (m,n) f a = + if n < m then a + else iter (m+1,n) f (f m a);; + +(* ------------------------------------------------------------------------- *) +(* The main types. *) +(* ------------------------------------------------------------------------- *) + +type vector = int*(int,num)func;; + +type matrix = (int*int)*(int*int,num)func;; + +type monomial = (vname,int)func;; + +type poly = (monomial,num)func;; + +(* ------------------------------------------------------------------------- *) +(* Assignment avoiding zeros. *) +(* ------------------------------------------------------------------------- *) + +let (|-->) x y a = if y =/ Int 0 then a else (x |-> y) a;; + +(* ------------------------------------------------------------------------- *) +(* This can be generic. *) +(* ------------------------------------------------------------------------- *) + +let element (d,v) i = tryapplyd v i (Int 0);; + +let mapa f (d,v) = + d,foldl (fun a i c -> (i |--> f(c)) a) undefined v;; + +let is_zero (d,v) = + match v with + Empty -> true + | _ -> false;; + +(* ------------------------------------------------------------------------- *) +(* Vectors. Conventionally indexed 1..n. *) +(* ------------------------------------------------------------------------- *) + +let vector_0 n = (n,undefined:vector);; + +let dim (v:vector) = fst v;; + +let vector_const c n = + if c =/ Int 0 then vector_0 n + else (n,itlist (fun k -> k |-> c) (1--n) undefined :vector);; + +let vector_1 = vector_const (Int 1);; + +let vector_cmul c (v:vector) = + let n = dim v in + if c =/ Int 0 then vector_0 n + else n,mapf (fun x -> c */ x) (snd v);; + +let vector_neg (v:vector) = (fst v,mapf minus_num (snd v) :vector);; + +let vector_add (v1:vector) (v2:vector) = + let m = dim v1 and n = dim v2 in + if m <> n then failwith "vector_add: incompatible dimensions" else + (n,combine (+/) (fun x -> x =/ Int 0) (snd v1) (snd v2) :vector);; + +let vector_sub v1 v2 = vector_add v1 (vector_neg v2);; + +let vector_of_list l = + let n = length l in + (n,itlist2 (|->) (1--n) l undefined :vector);; + +(* ------------------------------------------------------------------------- *) +(* Matrices; again rows and columns indexed from 1. *) +(* ------------------------------------------------------------------------- *) + +let matrix_0 (m,n) = ((m,n),undefined:matrix);; + +let dimensions (m:matrix) = fst m;; + +let matrix_const c (m,n as mn) = + if m <> n then failwith "matrix_const: needs to be square" + else if c =/ Int 0 then matrix_0 mn + else (mn,itlist (fun k -> (k,k) |-> c) (1--n) undefined :matrix);; + +let matrix_1 = matrix_const (Int 1);; + +let matrix_cmul c (m:matrix) = + let (i,j) = dimensions m in + if c =/ Int 0 then matrix_0 (i,j) + else (i,j),mapf (fun x -> c */ x) (snd m);; + +let matrix_neg (m:matrix) = (dimensions m,mapf minus_num (snd m) :matrix);; + +let matrix_add (m1:matrix) (m2:matrix) = + let d1 = dimensions m1 and d2 = dimensions m2 in + if d1 <> d2 then failwith "matrix_add: incompatible dimensions" + else (d1,combine (+/) (fun x -> x =/ Int 0) (snd m1) (snd m2) :matrix);; + +let matrix_sub m1 m2 = matrix_add m1 (matrix_neg m2);; + +let row k (m:matrix) = + let i,j = dimensions m in + (j, + foldl (fun a (i,j) c -> if i = k then (j |-> c) a else a) undefined (snd m) + : vector);; + +let column k (m:matrix) = + let i,j = dimensions m in + (i, + foldl (fun a (i,j) c -> if j = k then (i |-> c) a else a) undefined (snd m) + : vector);; + +let transp (m:matrix) = + let i,j = dimensions m in + ((j,i),foldl (fun a (i,j) c -> ((j,i) |-> c) a) undefined (snd m) :matrix);; + +let diagonal (v:vector) = + let n = dim v in + ((n,n),foldl (fun a i c -> ((i,i) |-> c) a) undefined (snd v) : matrix);; + +let matrix_of_list l = + let m = length l in + if m = 0 then matrix_0 (0,0) else + let n = length (hd l) in + (m,n),itern 1 l (fun v i -> itern 1 v (fun c j -> (i,j) |-> c)) undefined;; + +(* ------------------------------------------------------------------------- *) +(* Monomials. *) +(* ------------------------------------------------------------------------- *) + +let monomial_eval assig (m:monomial) = + foldl (fun a x k -> a */ power_num (apply assig x) (Int k)) + (Int 1) m;; + +let monomial_1 = (undefined:monomial);; + +let monomial_var x = (x |=> 1 :monomial);; + +let (monomial_mul:monomial->monomial->monomial) = + combine (+) (fun x -> false);; + +let monomial_pow (m:monomial) k = + if k = 0 then monomial_1 + else mapf (fun x -> k * x) m;; + +let monomial_divides (m1:monomial) (m2:monomial) = + foldl (fun a x k -> tryapplyd m2 x 0 >= k & a) true m1;; + +let monomial_div (m1:monomial) (m2:monomial) = + let m = combine (+) (fun x -> x = 0) m1 (mapf (fun x -> -x) m2) in + if foldl (fun a x k -> k >= 0 & a) true m then m + else failwith "monomial_div: non-divisible";; + +let monomial_degree x (m:monomial) = tryapplyd m x 0;; + +let monomial_lcm (m1:monomial) (m2:monomial) = + (itlist (fun x -> x |-> max (monomial_degree x m1) (monomial_degree x m2)) + (union (dom m1) (dom m2)) undefined :monomial);; + +let monomial_multidegree (m:monomial) = foldl (fun a x k -> k + a) 0 m;; + +let monomial_variables m = dom m;; + +(* ------------------------------------------------------------------------- *) +(* Polynomials. *) +(* ------------------------------------------------------------------------- *) + +let eval assig (p:poly) = + foldl (fun a m c -> a +/ c */ monomial_eval assig m) (Int 0) p;; + +let poly_0 = (undefined:poly);; + +let poly_isconst (p:poly) = foldl (fun a m c -> m = monomial_1 & a) true p;; + +let poly_var x = ((monomial_var x) |=> Int 1 :poly);; + +let poly_const c = + if c =/ Int 0 then poly_0 else (monomial_1 |=> c);; + +let poly_cmul c (p:poly) = + if c =/ Int 0 then poly_0 + else mapf (fun x -> c */ x) p;; + +let poly_neg (p:poly) = (mapf minus_num p :poly);; + +let poly_add (p1:poly) (p2:poly) = + (combine (+/) (fun x -> x =/ Int 0) p1 p2 :poly);; + +let poly_sub p1 p2 = poly_add p1 (poly_neg p2);; + +let poly_cmmul (c,m) (p:poly) = + if c =/ Int 0 then poly_0 + else if m = monomial_1 then mapf (fun d -> c */ d) p + else foldl (fun a m' d -> (monomial_mul m m' |-> c */ d) a) poly_0 p;; + +let poly_mul (p1:poly) (p2:poly) = + foldl (fun a m c -> poly_add (poly_cmmul (c,m) p2) a) poly_0 p1;; + +let poly_div (p1:poly) (p2:poly) = + if not(poly_isconst p2) then failwith "poly_div: non-constant" else + let c = eval undefined p2 in + if c =/ Int 0 then failwith "poly_div: division by zero" + else poly_cmul (Int 1 // c) p1;; + +let poly_square p = poly_mul p p;; + +let rec poly_pow p k = + if k = 0 then poly_const (Int 1) + else if k = 1 then p + else let q = poly_square(poly_pow p (k / 2)) in + if k mod 2 = 1 then poly_mul p q else q;; + +let poly_exp p1 p2 = + if not(poly_isconst p2) then failwith "poly_exp: not a constant" else + poly_pow p1 (Num.int_of_num (eval undefined p2));; + +let degree x (p:poly) = foldl (fun a m c -> max (monomial_degree x m) a) 0 p;; + +let multidegree (p:poly) = + foldl (fun a m c -> max (monomial_multidegree m) a) 0 p;; + +let poly_variables (p:poly) = + foldr (fun m c -> union (monomial_variables m)) p [];; + +(* ------------------------------------------------------------------------- *) +(* Order monomials for human presentation. *) +(* ------------------------------------------------------------------------- *) + +let humanorder_varpow (x1,k1) (x2,k2) = x1 < x2 or (x1 = x2 & k1 > k2);; + +let humanorder_monomial = + let rec ord l1 l2 = match (l1,l2) with + _,[] -> true + | [],_ -> false + | h1::t1,h2::t2 -> humanorder_varpow h1 h2 or (h1 = h2 & ord t1 t2) in + fun m1 m2 -> m1 = m2 or + ord (sort humanorder_varpow (graph m1)) + (sort humanorder_varpow (graph m2));; + +(* ------------------------------------------------------------------------- *) +(* Conversions to strings. *) +(* ------------------------------------------------------------------------- *) + +let string_of_vector min_size max_size (v:vector) = + let n_raw = dim v in + if n_raw = 0 then "[]" else + let n = max min_size (min n_raw max_size) in + let xs = map ((o) string_of_num (element v)) (1--n) in + "[" ^ end_itlist (fun s t -> s ^ ", " ^ t) xs ^ + (if n_raw > max_size then ", ...]" else "]");; + +let string_of_matrix max_size (m:matrix) = + let i_raw,j_raw = dimensions m in + let i = min max_size i_raw and j = min max_size j_raw in + let rstr = map (fun k -> string_of_vector j j (row k m)) (1--i) in + "["^end_itlist(fun s t -> s^";\n "^t) rstr ^ + (if j > max_size then "\n ...]" else "]");; + +let string_of_vname (v:vname): string = (v: string);; + +let rec string_of_term t = + match t with + Opp t1 -> "(- " ^ string_of_term t1 ^ ")" +| Add (t1, t2) -> + "(" ^ (string_of_term t1) ^ " + " ^ (string_of_term t2) ^ ")" +| Sub (t1, t2) -> + "(" ^ (string_of_term t1) ^ " - " ^ (string_of_term t2) ^ ")" +| Mul (t1, t2) -> + "(" ^ (string_of_term t1) ^ " * " ^ (string_of_term t2) ^ ")" +| Inv t1 -> "(/ " ^ string_of_term t1 ^ ")" +| Div (t1, t2) -> + "(" ^ (string_of_term t1) ^ " / " ^ (string_of_term t2) ^ ")" +| Pow (t1, n1) -> + "(" ^ (string_of_term t1) ^ " ^ " ^ (string_of_int n1) ^ ")" +| Zero -> "0" +| Var v -> "x" ^ (string_of_vname v) +| Const x -> string_of_num x;; + + +let string_of_varpow x k = + if k = 1 then string_of_vname x else string_of_vname x^"^"^string_of_int k;; + +let string_of_monomial m = + if m = monomial_1 then "1" else + let vps = List.fold_right (fun (x,k) a -> string_of_varpow x k :: a) + (sort humanorder_varpow (graph m)) [] in + end_itlist (fun s t -> s^"*"^t) vps;; + +let string_of_cmonomial (c,m) = + if m = monomial_1 then string_of_num c + else if c =/ Int 1 then string_of_monomial m + else string_of_num c ^ "*" ^ string_of_monomial m;; + +let string_of_poly (p:poly) = + if p = poly_0 then "<<0>>" else + let cms = sort (fun (m1,_) (m2,_) -> humanorder_monomial m1 m2) (graph p) in + let s = + List.fold_left (fun a (m,c) -> + if c </ Int 0 then a ^ " - " ^ string_of_cmonomial(minus_num c,m) + else a ^ " + " ^ string_of_cmonomial(c,m)) + "" cms in + let s1 = String.sub s 0 3 + and s2 = String.sub s 3 (String.length s - 3) in + "<<" ^(if s1 = " + " then s2 else "-"^s2)^">>";; + +(* ------------------------------------------------------------------------- *) +(* Printers. *) +(* ------------------------------------------------------------------------- *) + +let print_vector v = Format.print_string(string_of_vector 0 20 v);; + +let print_matrix m = Format.print_string(string_of_matrix 20 m);; + +let print_monomial m = Format.print_string(string_of_monomial m);; + +let print_poly m = Format.print_string(string_of_poly m);; + +(* +#install_printer print_vector;; +#install_printer print_matrix;; +#install_printer print_monomial;; +#install_printer print_poly;; +*) + +(* ------------------------------------------------------------------------- *) +(* Conversion from term. *) +(* ------------------------------------------------------------------------- *) + +let rec poly_of_term t = match t with + Zero -> poly_0 +| Const n -> poly_const n +| Var x -> poly_var x +| Opp t1 -> poly_neg (poly_of_term t1) +| Inv t1 -> + let p = poly_of_term t1 in + if poly_isconst p then poly_const(Int 1 // eval undefined p) + else failwith "poly_of_term: inverse of non-constant polyomial" +| Add (l, r) -> poly_add (poly_of_term l) (poly_of_term r) +| Sub (l, r) -> poly_sub (poly_of_term l) (poly_of_term r) +| Mul (l, r) -> poly_mul (poly_of_term l) (poly_of_term r) +| Div (l, r) -> + let p = poly_of_term l and q = poly_of_term r in + if poly_isconst q then poly_cmul (Int 1 // eval undefined q) p + else failwith "poly_of_term: division by non-constant polynomial" +| Pow (t, n) -> + poly_pow (poly_of_term t) n;; + +(* ------------------------------------------------------------------------- *) +(* String of vector (just a list of space-separated numbers). *) +(* ------------------------------------------------------------------------- *) + +let sdpa_of_vector (v:vector) = + let n = dim v in + let strs = map (o (decimalize 20) (element v)) (1--n) in + end_itlist (fun x y -> x ^ " " ^ y) strs ^ "\n";; + +(* ------------------------------------------------------------------------- *) +(* String for block diagonal matrix numbered k. *) +(* ------------------------------------------------------------------------- *) + +let sdpa_of_blockdiagonal k m = + let pfx = string_of_int k ^" " in + let ents = + foldl (fun a (b,i,j) c -> if i > j then a else ((b,i,j),c)::a) [] m in + let entss = sort (increasing fst) ents in + itlist (fun ((b,i,j),c) a -> + pfx ^ string_of_int b ^ " " ^ string_of_int i ^ " " ^ string_of_int j ^ + " " ^ decimalize 20 c ^ "\n" ^ a) entss "";; + +(* ------------------------------------------------------------------------- *) +(* String for a matrix numbered k, in SDPA sparse format. *) +(* ------------------------------------------------------------------------- *) + +let sdpa_of_matrix k (m:matrix) = + let pfx = string_of_int k ^ " 1 " in + let ms = foldr (fun (i,j) c a -> if i > j then a else ((i,j),c)::a) + (snd m) [] in + let mss = sort (increasing fst) ms in + itlist (fun ((i,j),c) a -> + pfx ^ string_of_int i ^ " " ^ string_of_int j ^ + " " ^ decimalize 20 c ^ "\n" ^ a) mss "";; + +(* ------------------------------------------------------------------------- *) +(* String in SDPA sparse format for standard SDP problem: *) +(* *) +(* X = v_1 * [M_1] + ... + v_m * [M_m] - [M_0] must be PSD *) +(* Minimize obj_1 * v_1 + ... obj_m * v_m *) +(* ------------------------------------------------------------------------- *) + +let sdpa_of_problem comment obj mats = + let m = length mats - 1 + and n,_ = dimensions (hd mats) in + "\"" ^ comment ^ "\"\n" ^ + string_of_int m ^ "\n" ^ + "1\n" ^ + string_of_int n ^ "\n" ^ + sdpa_of_vector obj ^ + itlist2 (fun k m a -> sdpa_of_matrix (k - 1) m ^ a) + (1--length mats) mats "";; + +(* ------------------------------------------------------------------------- *) +(* More parser basics. *) +(* ------------------------------------------------------------------------- *) + +exception Noparse;; + + +let isspace,issep,isbra,issymb,isalpha,isnum,isalnum = + let charcode s = Char.code(String.get s 0) in + let spaces = " \t\n\r" + and separators = ",;" + and brackets = "()[]{}" + and symbs = "\\!@#$%^&*-+|\\<=>/?~.:" + and alphas = "'abcdefghijklmnopqrstuvwxyz_ABCDEFGHIJKLMNOPQRSTUVWXYZ" + and nums = "0123456789" in + let allchars = spaces^separators^brackets^symbs^alphas^nums in + let csetsize = itlist ((o) max charcode) (explode allchars) 256 in + let ctable = Array.make csetsize 0 in + do_list (fun c -> Array.set ctable (charcode c) 1) (explode spaces); + do_list (fun c -> Array.set ctable (charcode c) 2) (explode separators); + do_list (fun c -> Array.set ctable (charcode c) 4) (explode brackets); + do_list (fun c -> Array.set ctable (charcode c) 8) (explode symbs); + do_list (fun c -> Array.set ctable (charcode c) 16) (explode alphas); + do_list (fun c -> Array.set ctable (charcode c) 32) (explode nums); + let isspace c = Array.get ctable (charcode c) = 1 + and issep c = Array.get ctable (charcode c) = 2 + and isbra c = Array.get ctable (charcode c) = 4 + and issymb c = Array.get ctable (charcode c) = 8 + and isalpha c = Array.get ctable (charcode c) = 16 + and isnum c = Array.get ctable (charcode c) = 32 + and isalnum c = Array.get ctable (charcode c) >= 16 in + isspace,issep,isbra,issymb,isalpha,isnum,isalnum;; + +let (||) parser1 parser2 input = + try parser1 input + with Noparse -> parser2 input;; + +let (++) parser1 parser2 input = + let result1,rest1 = parser1 input in + let result2,rest2 = parser2 rest1 in + (result1,result2),rest2;; + +let rec many prs input = + try let result,next = prs input in + let results,rest = many prs next in + (result::results),rest + with Noparse -> [],input;; + +let (>>) prs treatment input = + let result,rest = prs input in + treatment(result),rest;; + +let fix err prs input = + try prs input + with Noparse -> failwith (err ^ " expected");; + +let rec listof prs sep err = + prs ++ many (sep ++ fix err prs >> snd) >> (fun (h,t) -> h::t);; + +let possibly prs input = + try let x,rest = prs input in [x],rest + with Noparse -> [],input;; + +let some p = + function + [] -> raise Noparse + | (h::t) -> if p h then (h,t) else raise Noparse;; + +let a tok = some (fun item -> item = tok);; + +let rec atleast n prs i = + (if n <= 0 then many prs + else prs ++ atleast (n - 1) prs >> (fun (h,t) -> h::t)) i;; + +let finished input = + if input = [] then 0,input else failwith "Unparsed input";; + +let word s = + end_itlist (fun p1 p2 -> (p1 ++ p2) >> (fun (s,t) -> s^t)) + (map a (explode s));; + +let token s = + many (some isspace) ++ word s ++ many (some isspace) + >> (fun ((_,t),_) -> t);; + +let decimal = + let numeral = some isnum in + let decimalint = atleast 1 numeral >> ((o) Num.num_of_string implode) in + let decimalfrac = atleast 1 numeral + >> (fun s -> Num.num_of_string(implode s) // pow10 (length s)) in + let decimalsig = + decimalint ++ possibly (a "." ++ decimalfrac >> snd) + >> (function (h,[]) -> h | (h,[x]) -> h +/ x | _ -> failwith "decimalsig") in + let signed prs = + a "-" ++ prs >> ((o) minus_num snd) + || a "+" ++ prs >> snd + || prs in + let exponent = (a "e" || a "E") ++ signed decimalint >> snd in + signed decimalsig ++ possibly exponent + >> (function (h,[]) -> h | (h,[x]) -> h */ power_num (Int 10) x | _ -> + failwith "exponent");; + +let mkparser p s = + let x,rst = p(explode s) in + if rst = [] then x else failwith "mkparser: unparsed input";; + +let parse_decimal = mkparser decimal;; + +(* ------------------------------------------------------------------------- *) +(* Parse back a vector. *) +(* ------------------------------------------------------------------------- *) + +let parse_csdpoutput = + let rec skipupto dscr prs inp = + (dscr ++ prs >> snd + || some (fun c -> true) ++ skipupto dscr prs >> snd) inp in + let ignore inp = (),[] in + let csdpoutput = + (decimal ++ many(a " " ++ decimal >> snd) >> (fun (h,t) -> h::t)) ++ + (a " " ++ a "\n" ++ ignore) >> ((o) vector_of_list fst) in + mkparser csdpoutput;; + +(* ------------------------------------------------------------------------- *) +(* CSDP parameters; so far I'm sticking with the defaults. *) +(* ------------------------------------------------------------------------- *) + +let csdp_default_parameters = +"axtol=1.0e-8 +atytol=1.0e-8 +objtol=1.0e-8 +pinftol=1.0e8 +dinftol=1.0e8 +maxiter=100 +minstepfrac=0.9 +maxstepfrac=0.97 +minstepp=1.0e-8 +minstepd=1.0e-8 +usexzgap=1 +tweakgap=0 +affine=0 +printlevel=1 +";; + +let csdp_params = csdp_default_parameters;; + +(* ------------------------------------------------------------------------- *) +(* The same thing with CSDP. *) +(* Modified by the Coq development team to use a cache *) +(* ------------------------------------------------------------------------- *) + +let buffer_add_line buff line = + Buffer.add_string buff line; Buffer.add_char buff '\n' + +let string_of_file filename = + let fd = open_in filename in + let buff = Buffer.create 16 in + try while true do buffer_add_line buff (input_line fd) done; failwith "" + with End_of_file -> (close_in fd; Buffer.contents buff) + +let file_of_string filename s = + let fd = Pervasives.open_out filename in + output_string fd s; close_out fd + +let request_mark = "*** REQUEST ***" +let answer_mark = "*** ANSWER ***" +let end_mark = "*** END ***" +let infeasible_mark = "Infeasible\n" +let failure_mark = "Failure\n" + +let cache_name = "csdp.cache" + +let look_in_cache string_problem = + let n = String.length string_problem in + try + let inch = open_in cache_name in + let rec search () = + while input_line inch <> request_mark do () done; + let i = ref 0 in + while !i < n & string_problem.[!i] = input_char inch do incr i done; + if !i < n or input_line inch <> answer_mark then + search () + else begin + let buff = Buffer.create 16 in + let line = ref (input_line inch) in + while (!line <> end_mark) do + buffer_add_line buff !line; line := input_line inch + done; + close_in inch; + Buffer.contents buff + end in + try search () with End_of_file -> close_in inch; raise Not_found + with Sys_error _ -> raise Not_found + +let flush_to_cache string_problem string_result = + try + let flags = [Open_append;Open_text;Open_creat] in + let outch = open_out_gen flags 0o666 cache_name in + begin + try + Printf.fprintf outch "%s\n" request_mark; + Printf.fprintf outch "%s" string_problem; + Printf.fprintf outch "%s\n" answer_mark; + Printf.fprintf outch "%s" string_result; + Printf.fprintf outch "%s\n" end_mark; + with Sys_error _ as e -> close_out outch; raise e + end; + close_out outch + with Sys_error _ -> + print_endline "Warning: Could not open or write to csdp cache" + +exception CsdpInfeasible + +let run_csdp dbg string_problem = + try + let res = look_in_cache string_problem in + if res = infeasible_mark then raise CsdpInfeasible; + if res = failure_mark then failwith "csdp error"; + res + with Not_found -> + let input_file = Filename.temp_file "sos" ".dat-s" in + let output_file = Filename.temp_file "sos" ".dat-s" in + let temp_path = Filename.dirname input_file in + let params_file = Filename.concat temp_path "param.csdp" in + file_of_string input_file string_problem; + file_of_string params_file csdp_params; + let rv = Sys.command("cd "^temp_path^"; csdp "^input_file^" "^output_file^ + (if dbg then "" else "> /dev/null")) in + if rv = 1 or rv = 2 then + (flush_to_cache string_problem infeasible_mark; raise CsdpInfeasible); + if rv = 127 then + (print_string "csdp not found, exiting..."; exit 1); + if rv <> 0 & rv <> 3 (* reduced accuracy *) then + (flush_to_cache string_problem failure_mark; + failwith("csdp: error "^string_of_int rv)); + let string_result = string_of_file output_file in + flush_to_cache string_problem string_result; + if not dbg then + (Sys.remove input_file; Sys.remove output_file; Sys.remove params_file); + string_result + +let csdp obj mats = + try parse_csdpoutput (run_csdp !debugging (sdpa_of_problem "" obj mats)) + with CsdpInfeasible -> failwith "csdp: Problem is infeasible" + +(* ------------------------------------------------------------------------- *) +(* Try some apparently sensible scaling first. Note that this is purely to *) +(* get a cleaner translation to floating-point, and doesn't affect any of *) +(* the results, in principle. In practice it seems a lot better when there *) +(* are extreme numbers in the original problem. *) +(* ------------------------------------------------------------------------- *) + +let scale_then = + let common_denominator amat acc = + foldl (fun a m c -> lcm_num (denominator c) a) acc amat + and maximal_element amat acc = + foldl (fun maxa m c -> max_num maxa (abs_num c)) acc amat in + fun solver obj mats -> + let cd1 = itlist common_denominator mats (Int 1) + and cd2 = common_denominator (snd obj) (Int 1) in + let mats' = map (mapf (fun x -> cd1 */ x)) mats + and obj' = vector_cmul cd2 obj in + let max1 = itlist maximal_element mats' (Int 0) + and max2 = maximal_element (snd obj') (Int 0) in + let scal1 = pow2 (20-int_of_float(log(float_of_num max1) /. log 2.0)) + and scal2 = pow2 (20-int_of_float(log(float_of_num max2) /. log 2.0)) in + let mats'' = map (mapf (fun x -> x */ scal1)) mats' + and obj'' = vector_cmul scal2 obj' in + solver obj'' mats'';; + +(* ------------------------------------------------------------------------- *) +(* Round a vector to "nice" rationals. *) +(* ------------------------------------------------------------------------- *) + +let nice_rational n x = round_num (n */ x) // n;; + +let nice_vector n = mapa (nice_rational n);; + +(* ------------------------------------------------------------------------- *) +(* Reduce linear program to SDP (diagonal matrices) and test with CSDP. This *) +(* one tests A [-1;x1;..;xn] >= 0 (i.e. left column is negated constants). *) +(* ------------------------------------------------------------------------- *) + +let linear_program_basic a = + let m,n = dimensions a in + let mats = map (fun j -> diagonal (column j a)) (1--n) + and obj = vector_const (Int 1) m in + try ignore (run_csdp false (sdpa_of_problem "" obj mats)); true + with CsdpInfeasible -> false + +(* ------------------------------------------------------------------------- *) +(* Test whether a point is in the convex hull of others. Rather than use *) +(* computational geometry, express as linear inequalities and call CSDP. *) +(* This is a bit lazy of me, but it's easy and not such a bottleneck so far. *) +(* ------------------------------------------------------------------------- *) + +let in_convex_hull pts pt = + let pts1 = (1::pt) :: map (fun x -> 1::x) pts in + let pts2 = map (fun p -> map (fun x -> -x) p @ p) pts1 in + let n = length pts + 1 + and v = 2 * (length pt + 1) in + let m = v + n - 1 in + let mat = + (m,n), + itern 1 pts2 (fun pts j -> itern 1 pts (fun x i -> (i,j) |-> Int x)) + (iter (1,n) (fun i -> (v + i,i+1) |-> Int 1) undefined) in + linear_program_basic mat;; + +(* ------------------------------------------------------------------------- *) +(* Filter down a set of points to a minimal set with the same convex hull. *) +(* ------------------------------------------------------------------------- *) + +let minimal_convex_hull = + let augment1 = function (m::ms) -> if in_convex_hull ms m then ms else ms@[m] + | _ -> failwith "augment1" + in + let augment m ms = funpow 3 augment1 (m::ms) in + fun mons -> + let mons' = itlist augment (tl mons) [hd mons] in + funpow (length mons') augment1 mons';; + +(* ------------------------------------------------------------------------- *) +(* Stuff for "equations" (generic A->num functions). *) +(* ------------------------------------------------------------------------- *) + +let equation_cmul c eq = + if c =/ Int 0 then Empty else mapf (fun d -> c */ d) eq;; + +let equation_add eq1 eq2 = combine (+/) (fun x -> x =/ Int 0) eq1 eq2;; + +let equation_eval assig eq = + let value v = apply assig v in + foldl (fun a v c -> a +/ value(v) */ c) (Int 0) eq;; + +(* ------------------------------------------------------------------------- *) +(* Eliminate among linear equations: return unconstrained variables and *) +(* assignments for the others in terms of them. We give one pseudo-variable *) +(* "one" that's used for a constant term. *) +(* ------------------------------------------------------------------------- *) + + +let eliminate_equations = + let rec extract_first p l = + match l with + [] -> failwith "extract_first" + | h::t -> if p(h) then h,t else + let k,s = extract_first p t in + k,h::s in + let rec eliminate vars dun eqs = + match vars with + [] -> if forall is_undefined eqs then dun + else (raise Unsolvable) + | v::vs -> + try let eq,oeqs = extract_first (fun e -> defined e v) eqs in + let a = apply eq v in + let eq' = equation_cmul (Int(-1) // a) (undefine v eq) in + let elim e = + let b = tryapplyd e v (Int 0) in + if b =/ Int 0 then e else + equation_add e (equation_cmul (minus_num b // a) eq) in + eliminate vs ((v |-> eq') (mapf elim dun)) (map elim oeqs) + with Failure _ -> eliminate vs dun eqs in + fun one vars eqs -> + let assig = eliminate vars undefined eqs in + let vs = foldl (fun a x f -> subtract (dom f) [one] @ a) [] assig in + setify vs,assig;; + +(* ------------------------------------------------------------------------- *) +(* Eliminate all variables, in an essentially arbitrary order. *) +(* ------------------------------------------------------------------------- *) + +let eliminate_all_equations one = + let choose_variable eq = + let (v,_) = choose eq in + if v = one then + let eq' = undefine v eq in + if is_undefined eq' then failwith "choose_variable" else + let (w,_) = choose eq' in w + else v in + let rec eliminate dun eqs = + match eqs with + [] -> dun + | eq::oeqs -> + if is_undefined eq then eliminate dun oeqs else + let v = choose_variable eq in + let a = apply eq v in + let eq' = equation_cmul (Int(-1) // a) (undefine v eq) in + let elim e = + let b = tryapplyd e v (Int 0) in + if b =/ Int 0 then e else + equation_add e (equation_cmul (minus_num b // a) eq) in + eliminate ((v |-> eq') (mapf elim dun)) (map elim oeqs) in + fun eqs -> + let assig = eliminate undefined eqs in + let vs = foldl (fun a x f -> subtract (dom f) [one] @ a) [] assig in + setify vs,assig;; + +(* ------------------------------------------------------------------------- *) +(* Solve equations by assigning arbitrary numbers. *) +(* ------------------------------------------------------------------------- *) + +let solve_equations one eqs = + let vars,assigs = eliminate_all_equations one eqs in + let vfn = itlist (fun v -> (v |-> Int 0)) vars (one |=> Int(-1)) in + let ass = + combine (+/) (fun c -> false) (mapf (equation_eval vfn) assigs) vfn in + if forall (fun e -> equation_eval ass e =/ Int 0) eqs + then undefine one ass else raise Sanity;; + +(* ------------------------------------------------------------------------- *) +(* Hence produce the "relevant" monomials: those whose squares lie in the *) +(* Newton polytope of the monomials in the input. (This is enough according *) +(* to Reznik: "Extremal PSD forms with few terms", Duke Math. Journal, *) +(* vol 45, pp. 363--374, 1978. *) +(* *) +(* These are ordered in sort of decreasing degree. In particular the *) +(* constant monomial is last; this gives an order in diagonalization of the *) +(* quadratic form that will tend to display constants. *) +(* ------------------------------------------------------------------------- *) + +let newton_polytope pol = + let vars = poly_variables pol in + let mons = map (fun m -> map (fun x -> monomial_degree x m) vars) (dom pol) + and ds = map (fun x -> (degree x pol + 1) / 2) vars in + let all = itlist (fun n -> allpairs (fun h t -> h::t) (0--n)) ds [[]] + and mons' = minimal_convex_hull mons in + let all' = + filter (fun m -> in_convex_hull mons' (map (fun x -> 2 * x) m)) all in + map (fun m -> itlist2 (fun v i a -> if i = 0 then a else (v |-> i) a) + vars m monomial_1) (rev all');; + +(* ------------------------------------------------------------------------- *) +(* Diagonalize (Cholesky/LDU) the matrix corresponding to a quadratic form. *) +(* ------------------------------------------------------------------------- *) + +let diag m = + let nn = dimensions m in + let n = fst nn in + if snd nn <> n then failwith "diagonalize: non-square matrix" else + let rec diagonalize i m = + if is_zero m then [] else + let a11 = element m (i,i) in + if a11 </ Int 0 then failwith "diagonalize: not PSD" + else if a11 =/ Int 0 then + if is_zero(row i m) then diagonalize (i + 1) m + else failwith "diagonalize: not PSD" + else + let v = row i m in + let v' = mapa (fun a1k -> a1k // a11) v in + let m' = + (n,n), + iter (i+1,n) (fun j -> + iter (i+1,n) (fun k -> + ((j,k) |--> (element m (j,k) -/ element v j */ element v' k)))) + undefined in + (a11,v')::diagonalize (i + 1) m' in + diagonalize 1 m;; + +(* ------------------------------------------------------------------------- *) +(* Adjust a diagonalization to collect rationals at the start. *) +(* ------------------------------------------------------------------------- *) + +let deration d = + if d = [] then Int 0,d else + let adj(c,l) = + let a = foldl (fun a i c -> lcm_num a (denominator c)) (Int 1) (snd l) // + foldl (fun a i c -> gcd_num a (numerator c)) (Int 0) (snd l) in + (c // (a */ a)),mapa (fun x -> a */ x) l in + let d' = map adj d in + let a = itlist ((o) lcm_num ((o) denominator fst)) d' (Int 1) // + itlist ((o) gcd_num ((o) numerator fst)) d' (Int 0) in + (Int 1 // a),map (fun (c,l) -> (a */ c,l)) d';; + +(* ------------------------------------------------------------------------- *) +(* Enumeration of monomials with given multidegree bound. *) +(* ------------------------------------------------------------------------- *) + +let rec enumerate_monomials d vars = + if d < 0 then [] + else if d = 0 then [undefined] + else if vars = [] then [monomial_1] else + let alts = + map (fun k -> let oths = enumerate_monomials (d - k) (tl vars) in + map (fun ks -> if k = 0 then ks else (hd vars |-> k) ks) oths) + (0--d) in + end_itlist (@) alts;; + +(* ------------------------------------------------------------------------- *) +(* Enumerate products of distinct input polys with degree <= d. *) +(* We ignore any constant input polynomials. *) +(* Give the output polynomial and a record of how it was derived. *) +(* ------------------------------------------------------------------------- *) + +let rec enumerate_products d pols = + if d = 0 then [poly_const num_1,Rational_lt num_1] else if d < 0 then [] else + match pols with + [] -> [poly_const num_1,Rational_lt num_1] + | (p,b)::ps -> let e = multidegree p in + if e = 0 then enumerate_products d ps else + enumerate_products d ps @ + map (fun (q,c) -> poly_mul p q,Product(b,c)) + (enumerate_products (d - e) ps);; + +(* ------------------------------------------------------------------------- *) +(* Multiply equation-parametrized poly by regular poly and add accumulator. *) +(* ------------------------------------------------------------------------- *) + +let epoly_pmul p q acc = + foldl (fun a m1 c -> + foldl (fun b m2 e -> + let m = monomial_mul m1 m2 in + let es = tryapplyd b m undefined in + (m |-> equation_add (equation_cmul c e) es) b) + a q) acc p;; + +(* ------------------------------------------------------------------------- *) +(* Usual operations on equation-parametrized poly. *) +(* ------------------------------------------------------------------------- *) + +let epoly_cmul c l = + if c =/ Int 0 then undefined else mapf (equation_cmul c) l;; + +let epoly_neg x = epoly_cmul (Int(-1)) x;; + +let epoly_add x = combine equation_add is_undefined x;; + +let epoly_sub p q = epoly_add p (epoly_neg q);; + +(* ------------------------------------------------------------------------- *) +(* Convert regular polynomial. Note that we treat (0,0,0) as -1. *) +(* ------------------------------------------------------------------------- *) + +let epoly_of_poly p = + foldl (fun a m c -> (m |-> ((0,0,0) |=> minus_num c)) a) undefined p;; + +(* ------------------------------------------------------------------------- *) +(* String for block diagonal matrix numbered k. *) +(* ------------------------------------------------------------------------- *) + +let sdpa_of_blockdiagonal k m = + let pfx = string_of_int k ^" " in + let ents = + foldl (fun a (b,i,j) c -> if i > j then a else ((b,i,j),c)::a) [] m in + let entss = sort (increasing fst) ents in + itlist (fun ((b,i,j),c) a -> + pfx ^ string_of_int b ^ " " ^ string_of_int i ^ " " ^ string_of_int j ^ + " " ^ decimalize 20 c ^ "\n" ^ a) entss "";; + +(* ------------------------------------------------------------------------- *) +(* SDPA for problem using block diagonal (i.e. multiple SDPs) *) +(* ------------------------------------------------------------------------- *) + +let sdpa_of_blockproblem comment nblocks blocksizes obj mats = + let m = length mats - 1 in + "\"" ^ comment ^ "\"\n" ^ + string_of_int m ^ "\n" ^ + string_of_int nblocks ^ "\n" ^ + (end_itlist (fun s t -> s^" "^t) (map string_of_int blocksizes)) ^ + "\n" ^ + sdpa_of_vector obj ^ + itlist2 (fun k m a -> sdpa_of_blockdiagonal (k - 1) m ^ a) + (1--length mats) mats "";; + +(* ------------------------------------------------------------------------- *) +(* Hence run CSDP on a problem in block diagonal form. *) +(* ------------------------------------------------------------------------- *) + +let csdp_blocks nblocks blocksizes obj mats = + let string_problem = sdpa_of_blockproblem "" nblocks blocksizes obj mats in + try parse_csdpoutput (run_csdp !debugging string_problem) + with CsdpInfeasible -> failwith "csdp: Problem is infeasible" + +(* ------------------------------------------------------------------------- *) +(* 3D versions of matrix operations to consider blocks separately. *) +(* ------------------------------------------------------------------------- *) + +let bmatrix_add = combine (+/) (fun x -> x =/ Int 0);; + +let bmatrix_cmul c bm = + if c =/ Int 0 then undefined + else mapf (fun x -> c */ x) bm;; + +let bmatrix_neg = bmatrix_cmul (Int(-1));; + +let bmatrix_sub m1 m2 = bmatrix_add m1 (bmatrix_neg m2);; + +(* ------------------------------------------------------------------------- *) +(* Smash a block matrix into components. *) +(* ------------------------------------------------------------------------- *) + +let blocks blocksizes bm = + map (fun (bs,b0) -> + let m = foldl + (fun a (b,i,j) c -> if b = b0 then ((i,j) |-> c) a else a) + undefined bm in + (*let d = foldl (fun a (i,j) c -> max a (max i j)) 0 m in*) + (((bs,bs),m):matrix)) + (zip blocksizes (1--length blocksizes));; + +(* ------------------------------------------------------------------------- *) +(* Positiv- and Nullstellensatz. Flag "linf" forces a linear representation. *) +(* ------------------------------------------------------------------------- *) + +let real_positivnullstellensatz_general linf d eqs leqs pol + : poly list * (positivstellensatz * (num * poly) list) list = + + let vars = itlist ((o) union poly_variables) (pol::eqs @ map fst leqs) [] in + let monoid = + if linf then + (poly_const num_1,Rational_lt num_1):: + (filter (fun (p,c) -> multidegree p <= d) leqs) + else enumerate_products d leqs in + let nblocks = length monoid in + let mk_idmultiplier k p = + let e = d - multidegree p in + let mons = enumerate_monomials e vars in + let nons = zip mons (1--length mons) in + mons, + itlist (fun (m,n) -> (m |-> ((-k,-n,n) |=> Int 1))) nons undefined in + let mk_sqmultiplier k (p,c) = + let e = (d - multidegree p) / 2 in + let mons = enumerate_monomials e vars in + let nons = zip mons (1--length mons) in + mons, + itlist (fun (m1,n1) -> + itlist (fun (m2,n2) a -> + let m = monomial_mul m1 m2 in + if n1 > n2 then a else + let c = if n1 = n2 then Int 1 else Int 2 in + let e = tryapplyd a m undefined in + (m |-> equation_add ((k,n1,n2) |=> c) e) a) + nons) + nons undefined in + let sqmonlist,sqs = unzip(map2 mk_sqmultiplier (1--length monoid) monoid) + and idmonlist,ids = unzip(map2 mk_idmultiplier (1--length eqs) eqs) in + let blocksizes = map length sqmonlist in + let bigsum = + itlist2 (fun p q a -> epoly_pmul p q a) eqs ids + (itlist2 (fun (p,c) s a -> epoly_pmul p s a) monoid sqs + (epoly_of_poly(poly_neg pol))) in + let eqns = foldl (fun a m e -> e::a) [] bigsum in + let pvs,assig = eliminate_all_equations (0,0,0) eqns in + let qvars = (0,0,0)::pvs in + let allassig = itlist (fun v -> (v |-> (v |=> Int 1))) pvs assig in + let mk_matrix v = + foldl (fun m (b,i,j) ass -> if b < 0 then m else + let c = tryapplyd ass v (Int 0) in + if c =/ Int 0 then m else + ((b,j,i) |-> c) (((b,i,j) |-> c) m)) + undefined allassig in + let diagents = foldl + (fun a (b,i,j) e -> if b > 0 & i = j then equation_add e a else a) + undefined allassig in + let mats = map mk_matrix qvars + and obj = length pvs, + itern 1 pvs (fun v i -> (i |--> tryapplyd diagents v (Int 0))) + undefined in + let raw_vec = if pvs = [] then vector_0 0 + else scale_then (csdp_blocks nblocks blocksizes) obj mats in + let find_rounding d = + (if !debugging then + (Format.print_string("Trying rounding with limit "^string_of_num d); + Format.print_newline()) + else ()); + let vec = nice_vector d raw_vec in + let blockmat = iter (1,dim vec) + (fun i a -> bmatrix_add (bmatrix_cmul (element vec i) (el i mats)) a) + (bmatrix_neg (el 0 mats)) in + let allmats = blocks blocksizes blockmat in + vec,map diag allmats in + let vec,ratdias = + if pvs = [] then find_rounding num_1 + else tryfind find_rounding (map Num.num_of_int (1--31) @ + map pow2 (5--66)) in + let newassigs = + itlist (fun k -> el (k - 1) pvs |-> element vec k) + (1--dim vec) ((0,0,0) |=> Int(-1)) in + let finalassigs = + foldl (fun a v e -> (v |-> equation_eval newassigs e) a) newassigs + allassig in + let poly_of_epoly p = + foldl (fun a v e -> (v |--> equation_eval finalassigs e) a) + undefined p in + let mk_sos mons = + let mk_sq (c,m) = + c,itlist (fun k a -> (el (k - 1) mons |--> element m k) a) + (1--length mons) undefined in + map mk_sq in + let sqs = map2 mk_sos sqmonlist ratdias + and cfs = map poly_of_epoly ids in + let msq = filter (fun (a,b) -> b <> []) (map2 (fun a b -> a,b) monoid sqs) in + let eval_sq sqs = itlist + (fun (c,q) -> poly_add (poly_cmul c (poly_mul q q))) sqs poly_0 in + let sanity = + itlist (fun ((p,c),s) -> poly_add (poly_mul p (eval_sq s))) msq + (itlist2 (fun p q -> poly_add (poly_mul p q)) cfs eqs + (poly_neg pol)) in + if not(is_undefined sanity) then raise Sanity else + cfs,map (fun (a,b) -> snd a,b) msq;; + + +let term_of_monoid l1 m = itlist (fun i m -> Mul (nth l1 i,m)) m (Const num_1) + +let rec term_of_pos l1 x = match x with + Axiom_eq i -> failwith "term_of_pos" + | Axiom_le i -> nth l1 i + | Axiom_lt i -> nth l1 i + | Monoid m -> term_of_monoid l1 m + | Rational_eq n -> Const n + | Rational_le n -> Const n + | Rational_lt n -> Const n + | Square t -> Pow (t, 2) + | Eqmul (t, y) -> Mul (t, term_of_pos l1 y) + | Sum (y, z) -> Add (term_of_pos l1 y, term_of_pos l1 z) + | Product (y, z) -> Mul (term_of_pos l1 y, term_of_pos l1 z);; + + +let dest_monomial mon = sort (increasing fst) (graph mon);; + +let monomial_order = + let rec lexorder l1 l2 = + match (l1,l2) with + [],[] -> true + | vps,[] -> false + | [],vps -> true + | ((x1,n1)::vs1),((x2,n2)::vs2) -> + if x1 < x2 then true + else if x2 < x1 then false + else if n1 < n2 then false + else if n2 < n1 then true + else lexorder vs1 vs2 in + fun m1 m2 -> + if m2 = monomial_1 then true else if m1 = monomial_1 then false else + let mon1 = dest_monomial m1 and mon2 = dest_monomial m2 in + let deg1 = itlist ((o) (+) snd) mon1 0 + and deg2 = itlist ((o) (+) snd) mon2 0 in + if deg1 < deg2 then false else if deg1 > deg2 then true + else lexorder mon1 mon2;; + +let dest_poly p = + map (fun (m,c) -> c,dest_monomial m) + (sort (fun (m1,_) (m2,_) -> monomial_order m1 m2) (graph p));; + +(* ------------------------------------------------------------------------- *) +(* Map back polynomials and their composites to term. *) +(* ------------------------------------------------------------------------- *) + +let term_of_varpow = + fun x k -> + if k = 1 then Var x else Pow (Var x, k);; + +let term_of_monomial = + fun m -> if m = monomial_1 then Const num_1 else + let m' = dest_monomial m in + let vps = itlist (fun (x,k) a -> term_of_varpow x k :: a) m' [] in + end_itlist (fun s t -> Mul (s,t)) vps;; + +let term_of_cmonomial = + fun (m,c) -> + if m = monomial_1 then Const c + else if c =/ num_1 then term_of_monomial m + else Mul (Const c,term_of_monomial m);; + +let term_of_poly = + fun p -> + if p = poly_0 then Zero else + let cms = map term_of_cmonomial + (sort (fun (m1,_) (m2,_) -> monomial_order m1 m2) (graph p)) in + end_itlist (fun t1 t2 -> Add (t1,t2)) cms;; + +let term_of_sqterm (c,p) = + Product(Rational_lt c,Square(term_of_poly p));; + +let term_of_sos (pr,sqs) = + if sqs = [] then pr + else Product(pr,end_itlist (fun a b -> Sum(a,b)) (map term_of_sqterm sqs));; + +let rec deepen f n = + try (*print_string "Searching with depth limit "; + print_int n; print_newline();*) f n + with Failure _ -> deepen f (n + 1);; + + + + + +exception TooDeep + +let deepen_until limit f n = + match compare limit 0 with + | 0 -> raise TooDeep + | -1 -> deepen f n + | _ -> + let rec d_until f n = + try if !debugging + then (print_string "Searching with depth limit "; + print_int n; print_newline()) ; f n + with Failure x -> + if !debugging then (Printf.printf "solver error : %s\n" x) ; + if n = limit then raise TooDeep else d_until f (n + 1) in + d_until f n + + +(* patch to remove zero polynomials from equalities. + In this case, hol light loops *) + +let real_nonlinear_prover depthmax eqs les lts = + let eq = map poly_of_term eqs + and le = map poly_of_term les + and lt = map poly_of_term lts in + let pol = itlist poly_mul lt (poly_const num_1) + and lep = map (fun (t,i) -> t,Axiom_le i) (zip le (0--(length le - 1))) + and ltp = map (fun (t,i) -> t,Axiom_lt i) (zip lt (0--(length lt - 1))) + and eqp = itlist2 (fun t i res -> + if t = undefined then res else (t,Axiom_eq i)::res) eq (0--(length eq - 1)) [] + in + + let proof = + let leq = lep @ ltp in + let eq = List.map fst eqp in + let tryall d = + let e = multidegree pol (*and pol' = poly_neg pol*) in + let k = if e = 0 then 1 else d / e in + tryfind (fun i -> d,i, + real_positivnullstellensatz_general false d eq leq + (poly_neg(poly_pow pol i))) + (0--k) in + let d,i,(cert_ideal,cert_cone) = deepen_until depthmax tryall 0 in + let proofs_ideal = + map2 (fun q i -> Eqmul(term_of_poly q,i)) + cert_ideal (List.map snd eqp) + and proofs_cone = map term_of_sos cert_cone + and proof_ne = + if lt = [] then Rational_lt num_1 else + let p = end_itlist (fun s t -> Product(s,t)) (map snd ltp) in + funpow i (fun q -> Product(p,q)) (Rational_lt num_1) in + end_itlist (fun s t -> Sum(s,t)) (proof_ne :: proofs_ideal @ proofs_cone) in + if !debugging then (print_string("Translating proof certificate to Coq"); print_newline()); + proof;; + + +(* ------------------------------------------------------------------------- *) +(* Now pure SOS stuff. *) +(* ------------------------------------------------------------------------- *) + +(* ------------------------------------------------------------------------- *) +(* Some combinatorial helper functions. *) +(* ------------------------------------------------------------------------- *) + +let rec allpermutations l = + if l = [] then [[]] else + itlist (fun h acc -> map (fun t -> h::t) + (allpermutations (subtract l [h])) @ acc) l [];; + +let allvarorders l = + map (fun vlis x -> index x vlis) (allpermutations l);; + +let changevariables_monomial zoln (m:monomial) = + foldl (fun a x k -> (assoc x zoln |-> k) a) monomial_1 m;; + +let changevariables zoln pol = + foldl (fun a m c -> (changevariables_monomial zoln m |-> c) a) + poly_0 pol;; + +(* ------------------------------------------------------------------------- *) +(* Sum-of-squares function with some lowbrow symmetry reductions. *) +(* ------------------------------------------------------------------------- *) + +let sumofsquares_general_symmetry tool pol = + let vars = poly_variables pol + and lpps = newton_polytope pol in + let n = length lpps in + let sym_eqs = + let invariants = filter + (fun vars' -> + is_undefined(poly_sub pol (changevariables (zip vars vars') pol))) + (allpermutations vars) in +(* let lpps2 = allpairs monomial_mul lpps lpps in*) +(* let lpp2_classes = + setify(map (fun m -> + setify(map (fun vars' -> changevariables_monomial (zip vars vars') m) + invariants)) lpps2) in *) + let lpns = zip lpps (1--length lpps) in + let lppcs = + filter (fun (m,(n1,n2)) -> n1 <= n2) + (allpairs + (fun (m1,n1) (m2,n2) -> (m1,m2),(n1,n2)) lpns lpns) in + let clppcs = end_itlist (@) + (map (fun ((m1,m2),(n1,n2)) -> + map (fun vars' -> + (changevariables_monomial (zip vars vars') m1, + changevariables_monomial (zip vars vars') m2),(n1,n2)) + invariants) + lppcs) in + let clppcs_dom = setify(map fst clppcs) in + let clppcs_cls = map (fun d -> filter (fun (e,_) -> e = d) clppcs) + clppcs_dom in + let eqvcls = map (o setify (map snd)) clppcs_cls in + let mk_eq cls acc = + match cls with + [] -> raise Sanity + | [h] -> acc + | h::t -> map (fun k -> (k |-> Int(-1)) (h |=> Int 1)) t @ acc in + itlist mk_eq eqvcls [] in + let eqs = foldl (fun a x y -> y::a) [] + (itern 1 lpps (fun m1 n1 -> + itern 1 lpps (fun m2 n2 f -> + let m = monomial_mul m1 m2 in + if n1 > n2 then f else + let c = if n1 = n2 then Int 1 else Int 2 in + (m |-> ((n1,n2) |-> c) (tryapplyd f m undefined)) f)) + (foldl (fun a m c -> (m |-> ((0,0)|=>c)) a) + undefined pol)) @ + sym_eqs in + let pvs,assig = eliminate_all_equations (0,0) eqs in + let allassig = itlist (fun v -> (v |-> (v |=> Int 1))) pvs assig in + let qvars = (0,0)::pvs in + let diagents = + end_itlist equation_add (map (fun i -> apply allassig (i,i)) (1--n)) in + let mk_matrix v = + ((n,n), + foldl (fun m (i,j) ass -> let c = tryapplyd ass v (Int 0) in + if c =/ Int 0 then m else + ((j,i) |-> c) (((i,j) |-> c) m)) + undefined allassig :matrix) in + let mats = map mk_matrix qvars + and obj = length pvs, + itern 1 pvs (fun v i -> (i |--> tryapplyd diagents v (Int 0))) + undefined in + let raw_vec = if pvs = [] then vector_0 0 else tool obj mats in + let find_rounding d = + (if !debugging then + (Format.print_string("Trying rounding with limit "^string_of_num d); + Format.print_newline()) + else ()); + let vec = nice_vector d raw_vec in + let mat = iter (1,dim vec) + (fun i a -> matrix_add (matrix_cmul (element vec i) (el i mats)) a) + (matrix_neg (el 0 mats)) in + deration(diag mat) in + let rat,dia = + if pvs = [] then + let mat = matrix_neg (el 0 mats) in + deration(diag mat) + else + tryfind find_rounding (map Num.num_of_int (1--31) @ + map pow2 (5--66)) in + let poly_of_lin(d,v) = + d,foldl(fun a i c -> (el (i - 1) lpps |-> c) a) undefined (snd v) in + let lins = map poly_of_lin dia in + let sqs = map (fun (d,l) -> poly_mul (poly_const d) (poly_pow l 2)) lins in + let sos = poly_cmul rat (end_itlist poly_add sqs) in + if is_undefined(poly_sub sos pol) then rat,lins else raise Sanity;; + +let (sumofsquares: poly -> Num.num * (( Num.num * poly) list)) = +sumofsquares_general_symmetry csdp;; diff --git a/contrib/micromega/sos.mli b/contrib/micromega/sos.mli new file mode 100644 index 00000000..31c9518c --- /dev/null +++ b/contrib/micromega/sos.mli @@ -0,0 +1,66 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + + +type vname = string;; + +type term = +| Zero +| Const of Num.num +| Var of vname +| Inv of term +| Opp of term +| Add of (term * term) +| Sub of (term * term) +| Mul of (term * term) +| Div of (term * term) +| Pow of (term * int) + +type positivstellensatz = + Axiom_eq of int + | Axiom_le of int + | Axiom_lt of int + | Rational_eq of Num.num + | Rational_le of Num.num + | Rational_lt of Num.num + | Square of term + | Monoid of int list + | Eqmul of term * positivstellensatz + | Sum of positivstellensatz * positivstellensatz + | Product of positivstellensatz * positivstellensatz + +type poly + +val poly_isconst : poly -> bool + +val poly_neg : poly -> poly + +val poly_mul : poly -> poly -> poly + +val poly_pow : poly -> int -> poly + +val poly_const : Num.num -> poly + +val poly_of_term : term -> poly + +val term_of_poly : poly -> term + +val term_of_sos : positivstellensatz * (Num.num * poly) list -> + positivstellensatz + +val string_of_poly : poly -> string + +exception TooDeep + +val deepen_until : int -> (int -> 'a) -> int -> 'a + +val real_positivnullstellensatz_general : bool -> int -> poly list -> + (poly * positivstellensatz) list -> + poly -> poly list * (positivstellensatz * (Num.num * poly) list) list + +val sumofsquares : poly -> Num.num * ( Num.num * poly) list diff --git a/contrib/micromega/vector.ml b/contrib/micromega/vector.ml new file mode 100644 index 00000000..fee4ebfc --- /dev/null +++ b/contrib/micromega/vector.ml @@ -0,0 +1,674 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* *) +(* Micromega: A reflexive tactic using the Positivstellensatz *) +(* *) +(* Frédéric Besson (Irisa/Inria) 2006-2008 *) +(* *) +(************************************************************************) + +open Num + +module type S = +sig + type t + + val fresh : t -> int + + val null : t + + val is_null : t -> bool + + val get : int -> t -> num + + val update : int -> (num -> num) -> t -> t + (* behaviour is undef if index < 0 -- might loop*) + + val set : int -> num -> t -> t + + (* + For efficiency... + + val get_update : int -> (num -> num) -> t -> num * t + *) + + val mul : num -> t -> t + + val uminus : t -> t + + val add : t -> t -> t + + val dotp : t -> t -> num + + val lin_comb : num -> t -> num -> t -> t + (* lin_comb n1 t1 n2 t2 = (n1 * t1) + (n2 * t2) *) + + val gcd : t -> Big_int.big_int + + val normalise : t -> num * t + + val hash : t -> int + + val compare : t -> t -> int + + type it + + val iterator : t -> it + val element : it -> (num*it) option + + val string : t -> string + + type status = Pos | Neg + + (* the result list is ordered by fst *) + val status : t -> (int * status) list + + val from_list : num list -> t + val to_list : t -> num list + +end + + +module type SystemS = +sig + + module Vect : S + + module Cstr : + sig + type kind = Eq | Ge + val string_of_kind : kind -> string + type cstr = {coeffs : Vect.t ; op : kind ; cst : num} + val string_of_cstr : cstr -> string + val compare : cstr -> cstr -> int + end + open Cstr + + + module CstrBag : + sig + type t + exception Contradiction + + val empty : t + + val is_empty : t -> bool + + val add : cstr -> t -> t + (* c can be deduced from add c t *) + + val find : (cstr -> bool) -> t -> cstr option + + val fold : (cstr -> 'a -> 'a) -> t -> 'a -> 'a + + val status : t -> (int * (int list * int list)) list + (* aggregate of vector statuses *) + + val remove : cstr -> t -> t + + (* remove_list the ith element -- it is the ith element visited by 'fold' *) + + val split : (cstr -> int) -> t -> (int -> t) + + type it + val iterator : t -> it + val element : it -> (cstr*it) option + + end + +end + +let zero_num = Int 0 +let unit_num = Int 1 + + + + +module Cstr(V:S) = +struct + type kind = Eq | Ge + let string_of_kind = function Eq -> "Eq" | Ge -> "Ge" + + type cstr = {coeffs : V.t ; op : kind ; cst : num} + + let string_of_cstr {coeffs =a ; op = b ; cst =c} = + Printf.sprintf "{coeffs = %s;op=%s;cst=%s}" (V.string a) (string_of_kind b) (string_of_num c) + + type t = cstr + let compare + {coeffs = v1 ; op = op1 ; cst = c1} + {coeffs = v2 ; op = op2 ; cst = c2} = + Mutils.Cmp.compare_lexical [ + (fun () -> V.compare v1 v2); + (fun () -> Pervasives.compare op1 op2); + (fun () -> compare_num c1 c2) + ] + + +end + + + +module VList : S with type t = num list = +struct + type t = num list + + let fresh l = failwith "not implemented" + + let null = [] + + let is_null = List.for_all ((=/) zero_num) + + let normalise l = failwith "Not implemented" + (*match l with (* Buggy : What if the first num is zero! *) + | [] -> (Int 0,[]) + | [n] -> (n,[Int 1]) + | n::l -> (n, (Int 1) :: List.map (fun x -> x // n) l) + *) + + + let get i l = try List.nth l i with _ -> zero_num + + (* This is not tail-recursive *) + let rec update i f t = + match t with + | [] -> if i = 0 then [f zero_num] else (zero_num)::(update (i-1) f []) + | e::t -> if i = 0 then (f e)::t else e::(update (i-1) f t) + + let rec set i n t = + match t with + | [] -> if i = 0 then [n] else (zero_num)::(set (i-1) n []) + | e::t -> if i = 0 then (n)::t else e::(set (i-1) n t) + + + + + let rec mul z t = + match z with + | Int 0 -> null + | Int 1 -> t + | _ -> List.map (mult_num z) t + + let uminus t = mul (Int (-1)) t + + let rec add t1 t2 = + match t1,t2 with + | [], _ -> t2 + | _ , [] -> t1 + | e1::t1,e2::t2 -> (e1 +/ e2 )::(add t1 t2) + + let dotp t1 t2 = + let rec _dotp t1 t2 acc = + match t1, t2 with + | [] , _ -> acc + | _ , [] -> acc + | e1::t1,e2::t2 -> _dotp t1 t2 (acc +/ (e1 */ e2)) in + _dotp t1 t2 zero_num + + let add_mul n t1 t2 = + match n with + | Int 0 -> t2 + | Int 1 -> add t1 t2 + | _ -> + let rec _add_mul t1 t2 = + match t1,t2 with + | [], _ -> t2 + | _ , [] -> mul n t1 + | e1::t1,e2::t2 -> ( (n */e1) +/ e2 )::(_add_mul t1 t2) in + _add_mul t1 t2 + + let lin_comb n1 t1 n2 t2 = + match n1,n2 with + | Int 0 , _ -> mul n2 t2 + | Int 1 , _ -> add_mul n2 t2 t1 + | _ , Int 0 -> mul n1 t1 + | _ , Int 1 -> add_mul n1 t1 t2 + | _ -> + let rec _lin_comb t1 t2 = + match t1,t2 with + | [], _ -> mul n2 t2 + | _ , [] -> mul n1 t1 + | e1::t1,e2::t2 -> ( (n1 */e1) +/ (n2 */ e2 ))::(_lin_comb t1 t2) in + _lin_comb t1 t2 + + (* could be computed on the fly *) + let gcd t =Mutils.gcd_list t + + + + + let hash = Mutils.Cmp.hash_list int_of_num + + let compare = Mutils.Cmp.compare_list compare_num + + type it = t + let iterator (x:t) : it = x + let element it = + match it with + | [] -> None + | e::l -> Some (e,l) + + (* TODO: Buffer! *) + let string l = List.fold_right (fun n s -> (string_of_num n)^";"^s) l "" + + type status = Pos | Neg + + let status l = + let rec xstatus i l = + match l with + | [] -> [] + | e::l -> + begin + match compare_num e (Int 0) with + | 1 -> (i,Pos):: (xstatus (i+1) l) + | 0 -> xstatus (i+1) l + | -1 -> (i,Neg) :: (xstatus (i+1) l) + | _ -> assert false + end in + xstatus 0 l + + let from_list l = l + let to_list l = l + +end + +module VMap : S = +struct + module Map = Map.Make(struct type t = int let compare (x:int) (y:int) = Pervasives.compare x y end) + + type t = num Map.t + + let null = Map.empty + + let fresh m = failwith "not implemented" + + let is_null = Map.is_empty + + let normalise m = failwith "Not implemented" + + + + let get i l = try Map.find i l with _ -> zero_num + + let update i f t = + try + let res = f (Map.find i t) in + if res =/ zero_num + then Map.remove i t + else Map.add i res t + with + Not_found -> + let res = f zero_num in + if res =/ zero_num then t else Map.add i res t + + let set i n t = + if n =/ zero_num then Map.remove i t + else Map.add i n t + + + let rec mul z t = + match z with + | Int 0 -> null + | Int 1 -> t + | _ -> Map.map (mult_num z) t + + let uminus t = mul (Int (-1)) t + + + let map2 f m1 m2 = + let res,m2' = + Map.fold (fun k e (res,m2) -> + let v = f e (get k m2) in + if v =/ zero_num + then (res,Map.remove k m2) + else (Map.add k v res,Map.remove k m2)) m1 (Map.empty,m2) in + Map.fold (fun k e res -> + let v = f zero_num e in + if v =/ zero_num + then res else Map.add k v res) m2' res + + let add t1 t2 = map2 (+/) t1 t2 + + + let dotp t1 t2 = + Map.fold (fun k e res -> + res +/ (e */ get k t2)) t1 zero_num + + + + let add_mul n t1 t2 = + match n with + | Int 0 -> t2 + | Int 1 -> add t1 t2 + | _ -> map2 (fun x y -> (n */ x) +/ y) t1 t2 + + let lin_comb n1 t1 n2 t2 = + match n1,n2 with + | Int 0 , _ -> mul n2 t2 + | Int 1 , _ -> add_mul n2 t2 t1 + | _ , Int 0 -> mul n1 t1 + | _ , Int 1 -> add_mul n1 t1 t2 + | _ -> map2 (fun x y -> (n1 */ x) +/ (n2 */ y)) t1 t2 + + + let hash map = Map.fold (fun k e res -> k lxor (int_of_num e) lxor res) map 0 + + let compare = Map.compare compare_num + + type it = t * int + + let iterator (x:t) : it = (x,0) + + let element (mp,id) = + try + Some (Map.find id mp, (mp, id+1)) + with + Not_found -> None + + (* TODO: Buffer! *) + type status = Pos | Neg + + let status l = Map.fold (fun k e l -> + match compare_num e (Int 0) with + | 1 -> (k,Pos)::l + | 0 -> l + | -1 -> (k,Neg) :: l + | _ -> assert false) l [] + let from_list l = + let rec from_list i l map = + match l with + | [] -> map + | e::l -> from_list (i+1) l (if e <>/ Int 0 then Map.add i e map else map) in + from_list 0 l Map.empty + + let gcd m = + let res = Map.fold (fun _ e x -> Big_int.gcd_big_int x (Mutils.numerator e)) m Big_int.zero_big_int in + if Big_int.compare_big_int res Big_int.zero_big_int = 0 + then Big_int.unit_big_int else res + + + let to_list m = + let l = List.rev (Map.fold (fun k e l -> (k,e)::l) m []) in + let rec xto_list i l = + match l with + | [] -> [] + | (x,v)::l' -> if i = x then v::(xto_list (i+1) l') else zero_num ::(xto_list (i+1) l) in + xto_list 0 l + + let string l = VList.string (to_list l) + + +end + + +module VSparse : S = +struct + + type t = (int*num) list + + let null = [] + + let fresh l = List.fold_left (fun acc (i,_) -> max (i+1) acc) 0 l + + let is_null l = l = [] + + let rec is_sorted l = + match l with + | [] -> true + | [e] -> true + | (i,_)::(j,x)::l -> i < j && is_sorted ((j,x)::l) + + + let check l = (List.for_all (fun (_,n) -> compare_num n (Int 0) <> 0) l) && (is_sorted l) + + (* let get i t = + assert (check t); + try List.assoc i t with Not_found -> zero_num *) + + let rec get (i:int) t = + match t with + | [] -> zero_num + | (j,n)::t -> + match compare i j with + | 0 -> n + | 1 -> get i t + | _ -> zero_num + + let cons i v rst = if v =/ Int 0 then rst else (i,v)::rst + + let rec update i f t = + match t with + | [] -> cons i (f zero_num) [] + | (k,v)::l -> + match Pervasives.compare i k with + | 0 -> cons k (f v) l + | -1 -> cons i (f zero_num) t + | 1 -> (k,v) ::(update i f l) + | _ -> failwith "compare_num" + + let update i f t = + assert (check t); + let res = update i f t in + assert (check t) ; res + + + let rec set i n t = + match t with + | [] -> cons i n [] + | (k,v)::l -> + match Pervasives.compare i k with + | 0 -> cons k n l + | -1 -> cons i n t + | 1 -> (k,v) :: (set i n l) + | _ -> failwith "compare_num" + + + let rec map f l = + match l with + | [] -> [] + | (i,e)::l -> cons i (f e) (map f l) + + let rec mul z t = + match z with + | Int 0 -> null + | Int 1 -> t + | _ -> List.map (fun (i,n) -> (i, mult_num z n)) t + + let mul z t = + assert (check t) ; + let res = mul z t in + assert (check res) ; + res + + let uminus t = mul (Int (-1)) t + + + let normalise l = + match l with + | [] -> (Int 0,[]) + | (i,n)::_ -> (n, mul ((Int 1) // n) l) + + + let rec map2 f m1 m2 = + match m1, m2 with + | [] , [] -> [] + | l , [] -> map (fun x -> f x zero_num) l + | [] ,l -> map (f zero_num) l + | (i,e)::l1,(i',e')::l2 -> + match Pervasives.compare i i' with + | 0 -> cons i (f e e') (map2 f l1 l2) + | -1 -> cons i (f e zero_num) (map2 f l1 m2) + | 1 -> cons i' (f zero_num e') (map2 f m1 l2) + | _ -> assert false + + (* let add t1 t2 = map2 (+/) t1 t2*) + + let rec add (m1:t) (m2:t) = + match m1, m2 with + | [] , [] -> [] + | l , [] -> l + | [] ,l -> l + | (i,e)::l1,(i',e')::l2 -> + match Pervasives.compare i i' with + | 0 -> cons i ( e +/ e') (add l1 l2) + | -1 -> (i,e) :: (add l1 m2) + | 1 -> (i', e') :: (add m1 l2) + | _ -> assert false + + + + + let add t1 t2 = + assert (check t1 && check t2); + let res = add t1 t2 in + assert (check res); + res + + + let rec dotp (t1:t) (t2:t) = + match t1, t2 with + | [] , _ -> zero_num + | _ , [] -> zero_num + | (i,e)::l1 , (i',e')::l2 -> + match Pervasives.compare i i' with + | 0 -> (e */ e') +/ (dotp l1 l2) + | -1 -> dotp l1 t2 + | 1 -> dotp t1 l2 + | _ -> assert false + + let dotp t1 t2 = + assert (check t1 && check t2) ; dotp t1 t2 + + let add_mul n t1 t2 = + match n with + | Int 0 -> t2 + | Int 1 -> add t1 t2 + | _ -> map2 (fun x y -> (n */ x) +/ y) t1 t2 + + let add_mul n (t1:t) (t2:t) = + match n with + | Int 0 -> t2 + | Int 1 -> add t1 t2 + | _ -> + let rec xadd_mul m1 m2 = + match m1, m2 with + | [] , [] -> [] + | _ , [] -> mul n m1 + | [] , _ -> m2 + | (i,e)::l1,(i',e')::l2 -> + match Pervasives.compare i i' with + | 0 -> cons i ( n */ e +/ e') (xadd_mul l1 l2) + | -1 -> (i,n */ e) :: (xadd_mul l1 m2) + | 1 -> (i', e') :: (xadd_mul m1 l2) + | _ -> assert false in + xadd_mul t1 t2 + + + + + let lin_comb n1 t1 n2 t2 = + match n1,n2 with + | Int 0 , _ -> mul n2 t2 + | Int 1 , _ -> add_mul n2 t2 t1 + | _ , Int 0 -> mul n1 t1 + | _ , Int 1 -> add_mul n1 t1 t2 + | _ -> (*map2 (fun x y -> (n1 */ x) +/ (n2 */ y)) t1 t2*) + let rec xlin_comb m1 m2 = + match m1, m2 with + | [] , [] -> [] + | _ , [] -> mul n1 m1 + | [] , _ -> mul n2 m2 + | (i,e)::l1,(i',e')::l2 -> + match Pervasives.compare i i' with + | 0 -> cons i ( n1 */ e +/ n2 */ e') (xlin_comb l1 l2) + | -1 -> (i,n1 */ e) :: (xlin_comb l1 m2) + | 1 -> (i', n2 */ e') :: (xlin_comb m1 l2) + | _ -> assert false in + xlin_comb t1 t2 + + + + + + let lin_comb n1 t1 n2 t2 = + assert (check t1 && check t2); + let res = lin_comb n1 t1 n2 t2 in + assert (check res); res + + let hash = Mutils.Cmp.hash_list (fun (x,y) -> (Hashtbl.hash x) lxor (int_of_num y)) + + + let compare = Mutils.Cmp.compare_list (fun x y -> Mutils.Cmp.compare_lexical + [ + (fun () -> Pervasives.compare (fst x) (fst y)); + (fun () -> compare_num (snd x) (snd y))]) + + (* + let compare (x:t) (y:t) = + let rec xcompare acc1 acc2 x y = + match x , y with + | [] , [] -> xcomp acc1 acc2 + | [] , _ -> -1 + | _ , [] -> 1 + | (i,n1)::l1 , (j,n2)::l2 -> + match Pervasives.compare i j with + | 0 -> xcompare (n1::acc1) (n2::acc2) l1 l2 + | c -> c + and xcomp acc1 acc2 = Mutils.Cmp.compare_list compare_num acc1 acc2 in + xcompare [] [] x y + *) + + type it = t + + let iterator (x:t) : it = x + + let element l = failwith "Not_implemented" + + (* TODO: Buffer! *) + type status = Pos | Neg + + let status l = List.map (fun (i,e) -> + match compare_num e (Int 0) with + | 1 -> i,Pos + | -1 -> i,Neg + | _ -> assert false) l + + let from_list (l: num list) = + let rec xfrom_list i l = + match l with + | [] -> [] + | e::l -> + if e <>/ Int 0 + then (i,e)::(xfrom_list (i+1) l) + else xfrom_list (i+1) l in + + let res = xfrom_list 0 l in + assert (check res) ; res + + + let gcd m = + let res = List.fold_left (fun x (i,e) -> Big_int.gcd_big_int x (Mutils.numerator e)) Big_int.zero_big_int m in + if Big_int.compare_big_int res Big_int.zero_big_int = 0 + then Big_int.unit_big_int else res + + let to_list m = + let rec xto_list i l = + match l with + | [] -> [] + | (x,v)::l' -> + if i = x then v::(xto_list (i+1) l') else zero_num ::(xto_list (i+1) l) in + xto_list 0 m + + let to_list l = + assert (check l); + to_list l + + + let string l = VList.string (to_list l) + +end |