summaryrefslogtreecommitdiff
path: root/plugins/micromega
diff options
context:
space:
mode:
Diffstat (limited to 'plugins/micromega')
-rw-r--r--plugins/micromega/Env.v6
-rw-r--r--plugins/micromega/EnvRing.v2
-rw-r--r--plugins/micromega/Lia.v34
-rw-r--r--plugins/micromega/Lqa.v51
-rw-r--r--plugins/micromega/Lra.v52
-rw-r--r--plugins/micromega/MExtraction.v2
-rw-r--r--plugins/micromega/Psatz.v71
-rw-r--r--plugins/micromega/QMicromega.v2
-rw-r--r--plugins/micromega/RMicromega.v2
-rw-r--r--plugins/micromega/RingMicromega.v2
-rw-r--r--plugins/micromega/VarMap.v29
-rw-r--r--plugins/micromega/ZMicromega.v2
-rw-r--r--plugins/micromega/certificate.ml1565
-rw-r--r--plugins/micromega/coq_micromega.ml887
-rw-r--r--plugins/micromega/g_micromega.ml453
-rw-r--r--plugins/micromega/mfourier.ml26
-rw-r--r--plugins/micromega/micromega.ml2827
-rw-r--r--plugins/micromega/micromega.mli950
-rw-r--r--plugins/micromega/micromega_plugin.mlpack (renamed from plugins/micromega/micromega_plugin.mllib)1
-rw-r--r--plugins/micromega/mutils.ml21
-rw-r--r--plugins/micromega/persistent_cache.ml10
-rw-r--r--plugins/micromega/vo.itarget4
22 files changed, 2112 insertions, 4487 deletions
diff --git a/plugins/micromega/Env.v b/plugins/micromega/Env.v
index a19e9df9..7e3ef892 100644
--- a/plugins/micromega/Env.v
+++ b/plugins/micromega/Env.v
@@ -93,7 +93,7 @@ End S.
Ltac jump_simpl :=
repeat
match goal with
- | |- appcontext [jump xH] => rewrite (jump_simpl xH)
- | |- appcontext [jump (xO ?p)] => rewrite (jump_simpl (xO p))
- | |- appcontext [jump (xI ?p)] => rewrite (jump_simpl (xI p))
+ | |- context [jump xH] => rewrite (jump_simpl xH)
+ | |- context [jump (xO ?p)] => rewrite (jump_simpl (xO p))
+ | |- context [jump (xI ?p)] => rewrite (jump_simpl (xI p))
end.
diff --git a/plugins/micromega/EnvRing.v b/plugins/micromega/EnvRing.v
index fd4bb248..904ee4da 100644
--- a/plugins/micromega/EnvRing.v
+++ b/plugins/micromega/EnvRing.v
@@ -914,7 +914,7 @@ Qed.
revert P1. induction LM1 as [|(M2,P2') LM2 IH]; simpl; intros.
- discriminate.
- assert (H':=PNSubst_ok n P3 M2 P2'). destruct PNSubst.
- * injection H; intros <-. rewrite <- PSubstL1_ok; intuition.
+ * injection H as <-. rewrite <- PSubstL1_ok; intuition.
* now apply IH.
Qed.
diff --git a/plugins/micromega/Lia.v b/plugins/micromega/Lia.v
index 3e58e81a..47b6f7c7 100644
--- a/plugins/micromega/Lia.v
+++ b/plugins/micromega/Lia.v
@@ -8,7 +8,7 @@
(* *)
(* Micromega: A reflexive tactic using the Positivstellensatz *)
(* *)
-(* Frédéric Besson (Irisa/Inria) 2013 *)
+(* Frédéric Besson (Irisa/Inria) 2013-2016 *)
(* *)
(************************************************************************)
@@ -16,27 +16,27 @@ Require Import ZMicromega.
Require Import ZArith.
Require Import RingMicromega.
Require Import VarMap.
-Require Tauto.
+Require Coq.micromega.Tauto.
Declare ML Module "micromega_plugin".
+
Ltac preprocess :=
zify ; unfold Z.succ in * ; unfold Z.pred in *.
-Ltac lia :=
- preprocess;
- xlia ;
- abstract (
- intros __wit __varmap __ff ;
- change (Tauto.eval_f (Zeval_formula (@find Z Z0 __varmap)) __ff) ;
- apply (ZTautoChecker_sound __ff __wit); vm_cast_no_check (eq_refl true)).
-
-Ltac nia :=
- preprocess;
- xnlia ;
- abstract (
- intros __wit __varmap __ff ;
- change (Tauto.eval_f (Zeval_formula (@find Z Z0 __varmap)) __ff) ;
- apply (ZTautoChecker_sound __ff __wit); vm_cast_no_check (eq_refl true)).
+Ltac zchange :=
+ intros __wit __varmap __ff ;
+ change (Tauto.eval_f (Zeval_formula (@find Z Z0 __varmap)) __ff) ;
+ apply (ZTautoChecker_sound __ff __wit).
+
+Ltac zchecker_no_abstract := zchange ; vm_compute ; reflexivity.
+
+Ltac zchecker_abstract := zchange ; vm_cast_no_check (eq_refl true).
+
+Ltac zchecker := zchecker_no_abstract.
+
+Ltac lia := preprocess; xlia zchecker.
+
+Ltac nia := preprocess; xnlia zchecker.
(* Local Variables: *)
diff --git a/plugins/micromega/Lqa.v b/plugins/micromega/Lqa.v
new file mode 100644
index 00000000..acd2751a
--- /dev/null
+++ b/plugins/micromega/Lqa.v
@@ -0,0 +1,51 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+(* *)
+(* Micromega: A reflexive tactic using the Positivstellensatz *)
+(* *)
+(* Frédéric Besson (Irisa/Inria) 2016 *)
+(* *)
+(************************************************************************)
+
+Require Import QMicromega.
+Require Import QArith.
+Require Import RingMicromega.
+Require Import VarMap.
+Require Coq.micromega.Tauto.
+Declare ML Module "micromega_plugin".
+
+Ltac rchange :=
+ intros __wit __varmap __ff ;
+ change (Tauto.eval_f (Qeval_formula (@find Q 0%Q __varmap)) __ff) ;
+ apply (QTautoChecker_sound __ff __wit).
+
+Ltac rchecker_no_abstract := rchange ; vm_compute ; reflexivity.
+Ltac rchecker_abstract := rchange ; vm_cast_no_check (eq_refl true).
+
+Ltac rchecker := rchecker_no_abstract.
+
+(** Here, lra stands for linear rational arithmetic *)
+Ltac lra := lra_Q rchecker.
+
+(** Here, nra stands for non-linear rational arithmetic *)
+Ltac nra := xnqa rchecker.
+
+Ltac xpsatz dom d :=
+ let tac := lazymatch dom with
+ | Q =>
+ ((sos_Q rchecker) || (psatz_Q d rchecker))
+ | _ => fail "Unsupported domain"
+ end in tac.
+
+Tactic Notation "psatz" constr(dom) int_or_var(n) := xpsatz dom n.
+Tactic Notation "psatz" constr(dom) := xpsatz dom ltac:(-1).
+
+
+(* Local Variables: *)
+(* coding: utf-8 *)
+(* End: *)
diff --git a/plugins/micromega/Lra.v b/plugins/micromega/Lra.v
new file mode 100644
index 00000000..5b97d8ed
--- /dev/null
+++ b/plugins/micromega/Lra.v
@@ -0,0 +1,52 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+(* *)
+(* Micromega: A reflexive tactic using the Positivstellensatz *)
+(* *)
+(* Frédéric Besson (Irisa/Inria) 2016 *)
+(* *)
+(************************************************************************)
+
+Require Import RMicromega.
+Require Import QMicromega.
+Require Import Rdefinitions.
+Require Import RingMicromega.
+Require Import VarMap.
+Require Coq.micromega.Tauto.
+Declare ML Module "micromega_plugin".
+
+Ltac rchange :=
+ intros __wit __varmap __ff ;
+ change (Tauto.eval_f (Reval_formula (@find R 0%R __varmap)) __ff) ;
+ apply (RTautoChecker_sound __ff __wit).
+
+Ltac rchecker_no_abstract := rchange ; vm_compute ; reflexivity.
+Ltac rchecker_abstract := rchange ; vm_cast_no_check (eq_refl true).
+
+Ltac rchecker := rchecker_no_abstract.
+
+(** Here, lra stands for linear real arithmetic *)
+Ltac lra := unfold Rdiv in * ; lra_R rchecker.
+
+(** Here, nra stands for non-linear real arithmetic *)
+Ltac nra := unfold Rdiv in * ; xnra rchecker.
+
+Ltac xpsatz dom d :=
+ let tac := lazymatch dom with
+ | R =>
+ (sos_R rchecker) || (psatz_R d rchecker)
+ | _ => fail "Unsupported domain"
+ end in tac.
+
+Tactic Notation "psatz" constr(dom) int_or_var(n) := xpsatz dom n.
+Tactic Notation "psatz" constr(dom) := xpsatz dom ltac:(-1).
+
+
+(* Local Variables: *)
+(* coding: utf-8 *)
+(* End: *)
diff --git a/plugins/micromega/MExtraction.v b/plugins/micromega/MExtraction.v
index 0a41af45..d28bb828 100644
--- a/plugins/micromega/MExtraction.v
+++ b/plugins/micromega/MExtraction.v
@@ -50,7 +50,7 @@ Extract Constant Rinv => "fun x -> 1 / x".
Extraction "micromega.ml"
List.map simpl_cone (*map_cone indexes*)
- denorm Qpower
+ denorm Qpower vm_add
n_of_Z N.of_nat ZTautoChecker ZWeakChecker QTautoChecker RTautoChecker find.
diff --git a/plugins/micromega/Psatz.v b/plugins/micromega/Psatz.v
index a461b26a..8acf0ff8 100644
--- a/plugins/micromega/Psatz.v
+++ b/plugins/micromega/Psatz.v
@@ -8,7 +8,7 @@
(* *)
(* Micromega: A reflexive tactic using the Positivstellensatz *)
(* *)
-(* Frédéric Besson (Irisa/Inria) 2006-2008 *)
+(* Frédéric Besson (Irisa/Inria) 2006-2016 *)
(* *)
(************************************************************************)
@@ -20,75 +20,36 @@ Require Import ZArith.
Require Import Rdefinitions.
Require Import RingMicromega.
Require Import VarMap.
-Require Tauto.
-Declare ML Module "micromega_plugin".
+Require Coq.micromega.Tauto.
+Require Lia.
+Require Lra.
+Require Lqa.
-Ltac preprocess :=
- zify ; unfold Z.succ in * ; unfold Z.pred in *.
+Declare ML Module "micromega_plugin".
-Ltac lia :=
- preprocess;
- xlia ;
- abstract (
- intros __wit __varmap __ff ;
- change (Tauto.eval_f (Zeval_formula (@find Z Z0 __varmap)) __ff) ;
- apply (ZTautoChecker_sound __ff __wit); vm_cast_no_check (eq_refl true)).
+Ltac lia := Lia.lia.
-Ltac nia :=
- preprocess;
- xnlia ;
- abstract (
- intros __wit __varmap __ff ;
- change (Tauto.eval_f (Zeval_formula (@find Z Z0 __varmap)) __ff) ;
- apply (ZTautoChecker_sound __ff __wit); vm_cast_no_check (eq_refl true)).
+Ltac nia := Lia.nia.
Ltac xpsatz dom d :=
let tac := lazymatch dom with
| Z =>
- (sos_Z || psatz_Z d) ;
- abstract(
- intros __wit __varmap __ff ;
- change (Tauto.eval_f (Zeval_formula (@find Z Z0 __varmap)) __ff) ;
- apply (ZTautoChecker_sound __ff __wit); vm_cast_no_check (eq_refl true))
+ (sos_Z Lia.zchecker) || (psatz_Z d Lia.zchecker)
| R =>
- (sos_R || psatz_R d) ;
- (* If csdp is not installed, the previous step might not produce any
- progress: the rest of the tactical will then fail. Hence the 'try'. *)
- try (abstract(intros __wit __varmap __ff ;
- change (Tauto.eval_f (Reval_formula (@find R 0%R __varmap)) __ff) ;
- apply (RTautoChecker_sound __ff __wit); vm_cast_no_check (eq_refl true)))
- | Q =>
- (sos_Q || psatz_Q d) ;
- (* If csdp is not installed, the previous step might not produce any
- progress: the rest of the tactical will then fail. Hence the 'try'. *)
- try (abstract(intros __wit __varmap __ff ;
- change (Tauto.eval_f (Qeval_formula (@find Q 0%Q __varmap)) __ff) ;
- apply (QTautoChecker_sound __ff __wit); vm_cast_no_check (eq_refl true)))
+ (sos_R Lra.rchecker) || (psatz_R d Lra.rchecker)
+ | Q => (sos_Q Lqa.rchecker) || (psatz_Q d Lqa.rchecker)
| _ => fail "Unsupported domain"
end in tac.
Tactic Notation "psatz" constr(dom) int_or_var(n) := xpsatz dom n.
-Tactic Notation "psatz" constr(dom) := xpsatz dom ltac:-1.
+Tactic Notation "psatz" constr(dom) := xpsatz dom ltac:(-1).
Ltac psatzl dom :=
let tac := lazymatch dom with
- | Z => lia
- | Q =>
- psatzl_Q ;
- (* If csdp is not installed, the previous step might not produce any
- progress: the rest of the tactical will then fail. Hence the 'try'. *)
- try (abstract(intros __wit __varmap __ff ;
- change (Tauto.eval_f (Qeval_formula (@find Q 0%Q __varmap)) __ff) ;
- apply (QTautoChecker_sound __ff __wit); vm_cast_no_check (eq_refl true)))
- | R =>
- unfold Rdiv in * ;
- psatzl_R ;
- (* If csdp is not installed, the previous step might not produce any
- progress: the rest of the tactical will then fail. Hence the 'try'. *)
- try abstract((intros __wit __varmap __ff ;
- change (Tauto.eval_f (Reval_formula (@find R 0%R __varmap)) __ff) ;
- apply (RTautoChecker_sound __ff __wit); vm_cast_no_check (eq_refl true)))
+ | Z => Lia.lia
+ | Q => Lqa.lra
+ | R => Lra.lra
| _ => fail "Unsupported domain"
end in tac.
@@ -96,6 +57,8 @@ Ltac psatzl dom :=
Ltac lra :=
first [ psatzl R | psatzl Q ].
+Ltac nra :=
+ first [ Lra.nra | Lqa.nra ].
(* Local Variables: *)
diff --git a/plugins/micromega/QMicromega.v b/plugins/micromega/QMicromega.v
index 43268363..b13285f5 100644
--- a/plugins/micromega/QMicromega.v
+++ b/plugins/micromega/QMicromega.v
@@ -168,7 +168,7 @@ Proof.
exact H.
Qed.
-Require Import Tauto.
+Require Import Coq.micromega.Tauto.
Definition Qnormalise := @cnf_normalise Q 0 1 Qplus Qmult Qminus Qopp Qeq_bool.
Definition Qnegate := @cnf_negate Q 0 1 Qplus Qmult Qminus Qopp Qeq_bool.
diff --git a/plugins/micromega/RMicromega.v b/plugins/micromega/RMicromega.v
index 72353a99..2352d78d 100644
--- a/plugins/micromega/RMicromega.v
+++ b/plugins/micromega/RMicromega.v
@@ -533,7 +533,7 @@ Proof.
exact H.
Qed.
-Require Import Tauto.
+Require Import Coq.micromega.Tauto.
Definition Rnormalise := @cnf_normalise Q 0%Q 1%Q Qplus Qmult Qminus Qopp Qeq_bool.
Definition Rnegate := @cnf_negate Q 0%Q 1%Q Qplus Qmult Qminus Qopp Qeq_bool.
diff --git a/plugins/micromega/RingMicromega.v b/plugins/micromega/RingMicromega.v
index 751a81df..ed49c3df 100644
--- a/plugins/micromega/RingMicromega.v
+++ b/plugins/micromega/RingMicromega.v
@@ -794,7 +794,7 @@ Definition xnormalise (t:Formula C) : list (NFormula) :=
| OpLe => (psub lhs rhs ,Strict) :: nil
end.
-Require Import Tauto.
+Require Import Coq.micromega.Tauto.
Definition cnf_normalise (t:Formula C) : cnf (NFormula) :=
List.map (fun x => x::nil) (xnormalise t).
diff --git a/plugins/micromega/VarMap.v b/plugins/micromega/VarMap.v
index 4981ddb3..2d2c0bc7 100644
--- a/plugins/micromega/VarMap.v
+++ b/plugins/micromega/VarMap.v
@@ -1,7 +1,7 @@
(* -*- coding: utf-8 -*- *)
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -26,6 +26,7 @@ Set Implicit Arguments.
*)
Section MakeVarMap.
+
Variable A : Type.
Variable default : A.
@@ -46,5 +47,29 @@ Section MakeVarMap.
end.
-End MakeVarMap.
+ Fixpoint singleton (x:positive) (v : A) : t :=
+ match x with
+ | xH => Leaf v
+ | xO p => Node (singleton p v) default Empty
+ | xI p => Node Empty default (singleton p v)
+ end.
+
+ Fixpoint vm_add (x: positive) (v : A) (m : t) {struct m} : t :=
+ match m with
+ | Empty => singleton x v
+ | Leaf vl =>
+ match x with
+ | xH => Leaf v
+ | xO p => Node (singleton p v) vl Empty
+ | xI p => Node Empty vl (singleton p v)
+ end
+ | Node l o r =>
+ match x with
+ | xH => Node l v r
+ | xI p => Node l o (vm_add p v r)
+ | xO p => Node (vm_add p v l) o r
+ end
+ end.
+
+End MakeVarMap.
diff --git a/plugins/micromega/ZMicromega.v b/plugins/micromega/ZMicromega.v
index d7ddef2b..5aa8d03f 100644
--- a/plugins/micromega/ZMicromega.v
+++ b/plugins/micromega/ZMicromega.v
@@ -198,7 +198,7 @@ Definition xnormalise (t:Formula Z) : list (NFormula Z) :=
| OpLe => (psub lhs (padd rhs (Pc 1)),NonStrict) :: nil
end.
-Require Import Tauto BinNums.
+Require Import Coq.micromega.Tauto BinNums.
Definition normalise (t:Formula Z) : cnf (NFormula Z) :=
List.map (fun x => x::nil) (xnormalise t).
diff --git a/plugins/micromega/certificate.ml b/plugins/micromega/certificate.ml
index a5fceb56..459c72f9 100644
--- a/plugins/micromega/certificate.ml
+++ b/plugins/micromega/certificate.ml
@@ -63,82 +63,82 @@ 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 (C2Ml.positive 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 Int.equal 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
+ | Mc.PEc z -> Poly.constant (n_spec.number_to_num z)
+ | Mc.PEX v -> Poly.variable (C2Ml.positive 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 Int.equal 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 v = Ml2C.positive v in
- let mn = if Int.equal i 1 then Mc.PEX v else Mc.PEpow (Mc.PEX v ,Ml2C.n i) in
- if Pervasives.(=) acc (Mc.PEc (Mc.Zpos Mc.XH)) (** FIXME *)
- then mn
- else Mc.PEmul(mn,acc))
- mn
- (Mc.PEc (Mc.Zpos Mc.XH))
+ let v = Ml2C.positive v in
+ let mn = if Int.equal i 1 then Mc.PEX v else Mc.PEpow (Mc.PEX v ,Ml2C.n i) in
+ if Pervasives.(=) acc (Mc.PEc (Mc.Zpos Mc.XH)) (** FIXME *)
+ 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 Pervasives.(=) c (Mc.PEc (Mc.Zpos Mc.XH))
- then var i
- else Mc.PEmul (c,var i) in
- let p' = if Pervasives.(=) 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
+ else let c = Mc.PEc (Ml2C.bigint (numerator c)) in
+ let mn =
+ if Pervasives.(=) c (Mc.PEc (Mc.Zpos Mc.XH))
+ then var i
+ else Mc.PEmul (c,var i) in
+ let p' = if Pervasives.(=) 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 Pervasives.(=) y' x then y'
- else fixpoint f y'
+ if Pervasives.(=) 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.PsatzMulE(t1, t2) ->
- simpl_cone (Mc.PsatzMulE (rec_simpl_cone t1, rec_simpl_cone t2))
- | Mc.PsatzAdd(t1,t2) ->
- simpl_cone (Mc.PsatzAdd (rec_simpl_cone t1, rec_simpl_cone t2))
- | x -> simpl_cone x in
- rec_simpl_cone e
-
-
+ | Mc.PsatzMulE(t1, t2) ->
+ simpl_cone (Mc.PsatzMulE (rec_simpl_cone t1, rec_simpl_cone t2))
+ | Mc.PsatzAdd(t1,t2) ->
+ simpl_cone (Mc.PsatzAdd (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
+ Const of cone
+| Ideal of cone *cone
+| Mult of cone * cone
+| Other of cone
and cone = Mc.zWitness
@@ -147,32 +147,32 @@ let factorise_linear_cone c =
let rec cone_list c l =
match c with
- | Mc.PsatzAdd (x,r) -> cone_list r (x::l)
- | _ -> c :: l in
-
+ | Mc.PsatzAdd (x,r) -> cone_list r (x::l)
+ | _ -> c :: l in
+
let factorise c1 c2 =
match c1 , c2 with
- | Mc.PsatzMulC(x,y) , Mc.PsatzMulC(x',y') ->
- if Pervasives.(=) x x' then Some (Mc.PsatzMulC(x, Mc.PsatzAdd(y,y'))) else None
- | Mc.PsatzMulE(x,y) , Mc.PsatzMulE(x',y') ->
- if Pervasives.(=) x x' then Some (Mc.PsatzMulE(x, Mc.PsatzAdd(y,y'))) else None
- | _ -> None in
-
+ | Mc.PsatzMulC(x,y) , Mc.PsatzMulC(x',y') ->
+ if Pervasives.(=) x x' then Some (Mc.PsatzMulC(x, Mc.PsatzAdd(y,y'))) else None
+ | Mc.PsatzMulE(x,y) , Mc.PsatzMulE(x',y') ->
+ if Pervasives.(=) x x' then Some (Mc.PsatzMulE(x, Mc.PsatzAdd(y,y'))) else None
+ | _ -> None in
+
let rec rebuild_cone l pending =
match l with
- | [] -> (match pending with
- | None -> Mc.PsatzZ
- | Some p -> p
- )
- | e::l ->
- (match pending with
- | None -> rebuild_cone l (Some e)
- | Some p -> (match factorise p e with
- | None -> Mc.PsatzAdd(p, rebuild_cone l (Some e))
- | Some f -> rebuild_cone l (Some f) )
- ) in
+ | [] -> (match pending with
+ | None -> Mc.PsatzZ
+ | Some p -> p
+ )
+ | e::l ->
+ (match pending with
+ | None -> rebuild_cone l (Some e)
+ | Some p -> (match factorise p e with
+ | None -> Mc.PsatzAdd(p, rebuild_cone l (Some e))
+ | Some f -> rebuild_cone l (Some f) )
+ ) in
- (rebuild_cone (List.sort Pervasives.compare (cone_list c [])) None)
+ (rebuild_cone (List.sort Pervasives.compare (cone_list c [])) None)
@@ -199,28 +199,28 @@ open Mfourier
let constrain_monomial mn l =
let coeffs = List.fold_left (fun acc p -> (Poly.get mn p)::acc) [] l in
- if Pervasives.(=) 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 }
+ if Pervasives.(=) 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)
+ | [] -> []
+ | (_,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
+ xpositivity 0 l
let string_of_op = function
@@ -241,23 +241,23 @@ let build_linear_system l =
let monomials =
List.fold_left (fun acc p ->
- Poly.fold (fun m _ acc -> MonSet.add m acc) p acc)
- (MonSet.singleton Monomial.const) l'
+ Poly.fold (fun m _ acc -> MonSet.add m acc) p acc)
+ (MonSet.singleton Monomial.const) l'
in (* For each monomial, compute a constraint *)
let s0 =
MonSet.fold (fun mn res -> (constrain_monomial mn l')::res) monomials [] in
- (* I need at least something strictly positive *)
+ (* 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));
+ 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)
+ {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
@@ -266,32 +266,32 @@ let big_int_to_z = Ml2C.bigint
-- 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
- | [] -> failwith "empty_certificate"
- | e::cert' ->
-(* let cst = match compare_big_int e zero_big_int with
- | 0 -> Mc.PsatzZ
- | 1 -> Mc.PsatzC (bint_to_cst e)
- | _ -> failwith "positivity error"
- in *)
- let rec scalar_product cert l =
- match cert with
- | [] -> Mc.PsatzZ
- | 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.PsatzAdd (
- Mc.PsatzMulC (Mc.Pc ( bint_to_cst c), Mc.PsatzIn (Ml2C.nat i)),
- r)
- | 0 -> r
- | _ -> Mc.PsatzAdd (
- Mc.PsatzMulE (Mc.PsatzC (bint_to_cst c), Mc.PsatzIn (Ml2C.nat i)),
- r) in
- (factorise_linear_cone
- (simplify_cone n_spec (scalar_product cert' li)))
+ match cert with
+ | [] -> failwith "empty_certificate"
+ | e::cert' ->
+ (* let cst = match compare_big_int e zero_big_int with
+ | 0 -> Mc.PsatzZ
+ | 1 -> Mc.PsatzC (bint_to_cst e)
+ | _ -> failwith "positivity error"
+ in *)
+ let rec scalar_product cert l =
+ match cert with
+ | [] -> Mc.PsatzZ
+ | 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.PsatzAdd (
+ Mc.PsatzMulC (Mc.Pc ( bint_to_cst c), Mc.PsatzIn (Ml2C.nat i)),
+ r)
+ | 0 -> r
+ | _ -> Mc.PsatzAdd (
+ Mc.PsatzMulE (Mc.PsatzC (bint_to_cst c), Mc.PsatzIn (Ml2C.nat i)),
+ r) in
+ (factorise_linear_cone
+ (simplify_cone n_spec (scalar_product cert' li)))
exception Found of Monomial.t
@@ -301,92 +301,155 @@ exception Strict
module MonMap = Map.Make(Monomial)
let primal l =
- let vr = ref 0 in
-
- let vect_of_poly map p =
- Poly.fold (fun mn vl (map,vect) ->
- if Pervasives.(=) mn Monomial.const
- then (map,vect)
- else
- let (mn,m) = try (MonMap.find mn map,map) with Not_found -> let res = (!vr, MonMap.add mn !vr map) in incr vr ; res in
- (m,if Int.equal (sign_num vl) 0 then vect else (mn,vl)::vect)) p (map,[]) in
-
- let op_op = function Mc.NonStrict -> Ge |Mc.Equal -> Eq | _ -> raise Strict in
+ let vr = ref 0 in
+
+ let vect_of_poly map p =
+ Poly.fold (fun mn vl (map,vect) ->
+ if Pervasives.(=) mn Monomial.const
+ then (map,vect)
+ else
+ let (mn,m) = try (MonMap.find mn map,map) with Not_found -> let res = (!vr, MonMap.add mn !vr map) in incr vr ; res in
+ (m,if Int.equal (sign_num vl) 0 then vect else (mn,vl)::vect)) p (map,[]) in
+
+ let op_op = function Mc.NonStrict -> Ge |Mc.Equal -> Eq | _ -> raise Strict in
- let cmp x y = Int.compare (fst x) (fst y) in
+ let cmp x y = Int.compare (fst x) (fst y) in
- snd (List.fold_right (fun (p,op) (map,l) ->
- let (mp,vect) = vect_of_poly map p in
- let cstr = {coeffs = List.sort cmp vect; op = op_op op ; cst = minus_num (Poly.get Monomial.const p)} in
+ snd (List.fold_right (fun (p,op) (map,l) ->
+ let (mp,vect) = vect_of_poly map p in
+ let cstr = {coeffs = List.sort cmp vect; op = op_op op ; cst = minus_num (Poly.get Monomial.const p)} in
- (mp,cstr::l)) l (MonMap.empty,[]))
+ (mp,cstr::l)) l (MonMap.empty,[]))
let dual_raw_certificate (l: (Poly.t * Mc.op1) list) =
-(* List.iter (fun (p,op) -> Printf.fprintf stdout "%a %s 0\n" Poly.pp p (string_of_op op) ) l ; *)
-
+ (* List.iter (fun (p,op) -> Printf.fprintf stdout "%a %s 0\n" Poly.pp p (string_of_op op) ) l ; *)
+
let sys = build_linear_system l in
- try
- match Fourier.find_point sys with
- | Inr _ -> None
- | Inl cert -> Some (rats_to_ints (Vect.to_list cert))
- (* should not use rats_to_ints *)
- with x when Errors.noncritical x ->
- if debug
- then (Printf.printf "raw certificate %s" (Printexc.to_string x);
- flush stdout) ;
- None
+ try
+ match Fourier.find_point sys with
+ | Inr _ -> None
+ | Inl cert -> Some (rats_to_ints (Vect.to_list cert))
+ (* should not use rats_to_ints *)
+ with x when CErrors.noncritical x ->
+ if debug
+ then (Printf.printf "raw certificate %s" (Printexc.to_string x);
+ flush stdout) ;
+ None
let raw_certificate l =
- try
- let p = primal l in
- match Fourier.find_point p with
- | Inr prf ->
- if debug then Printf.printf "AProof : %a\n" pp_proof prf ;
- let cert = List.map (fun (x,n) -> x+1,n) (fst (List.hd (Proof.mk_proof p prf))) in
- if debug then Printf.printf "CProof : %a" Vect.pp_vect cert ;
- Some (rats_to_ints (Vect.to_list cert))
- | Inl _ -> None
- with Strict ->
+ try
+ let p = primal l in
+ match Fourier.find_point p with
+ | Inr prf ->
+ if debug then Printf.printf "AProof : %a\n" pp_proof prf ;
+ let cert = List.map (fun (x,n) -> x+1,n) (fst (List.hd (Proof.mk_proof p prf))) in
+ if debug then Printf.printf "CProof : %a" Vect.pp_vect cert ;
+ Some (rats_to_ints (Vect.to_list cert))
+ | Inl _ -> None
+ with Strict ->
(* Fourier elimination should handle > *)
- dual_raw_certificate l
+ dual_raw_certificate l
let simple_linear_prover l =
let (lc,li) = List.split l in
- match raw_certificate lc with
- | None -> None (* No certificate *)
- | Some cert -> Some (cert,li)
-
+ match raw_certificate lc with
+ | None -> None (* No certificate *)
+ | Some cert -> Some (cert,li)
+
let linear_prover n_spec l =
- let build_system n_spec l =
- let li = List.combine l (interval 0 (List.length l -1)) in
- let (l1,l') = List.partition
- (fun (x,_) -> if Pervasives.(=) (snd x) Mc.NonEqual then true else false) li in
- List.map
- (fun ((x,y),i) -> match y with
- Mc.NonEqual -> failwith "cannot happen"
- | y -> ((dev_form n_spec x, y),i)) l' in
- let l' = build_system n_spec l in
- simple_linear_prover (*n_spec*) l'
+ let build_system n_spec l =
+ let li = List.combine l (interval 0 (List.length l -1)) in
+ let (l1,l') = List.partition
+ (fun (x,_) -> if Pervasives.(=) (snd x) Mc.NonEqual then true else false) li in
+ List.map
+ (fun ((x,y),i) -> match y with
+ Mc.NonEqual -> failwith "cannot happen"
+ | y -> ((dev_form n_spec x, y),i)) l' in
+ let l' = build_system n_spec l in
+ simple_linear_prover (*n_spec*) l'
let linear_prover n_spec l =
try linear_prover n_spec l
- with x when Errors.noncritical x ->
- (print_string (Printexc.to_string x); None)
+ with x when CErrors.noncritical x ->
+ (print_string (Printexc.to_string x); None)
+
+let compute_max_nb_cstr l d =
+ let len = List.length l in
+ max len (max d (len * d))
+
+let linear_prover_with_cert prfdepth spec l =
+ max_nb_cstr := compute_max_nb_cstr l prfdepth ;
+ match linear_prover spec l with
+ | None -> None
+ | Some cert -> Some (make_certificate spec cert)
+
+let nlinear_prover prfdepth (sys: (Mc.q Mc.pExpr * Mc.op1) list) =
+ LinPoly.MonT.clear ();
+ max_nb_cstr := compute_max_nb_cstr sys prfdepth ;
+ (* Assign a proof to the initial hypotheses *)
+ let sys = mapi (fun c i -> (c,Mc.PsatzIn (Ml2C.nat i))) sys in
+
+
+ (* Add all the product of hypotheses *)
+ let prod = all_pairs (fun ((c,o),p) ((c',o'),p') ->
+ ((Mc.PEmul(c,c') , Mc.opMult o o') , Mc.PsatzMulE(p,p'))) sys in
+
+ (* Only filter those have a meaning *)
+ let prod = List.fold_left (fun l ((c,o),p) ->
+ match o with
+ | None -> l
+ | Some o -> ((c,o),p) :: l) [] prod in
+
+ let sys = sys @ prod in
+
+ let square =
+ (* Collect the squares and state that they are positive *)
+ let pols = List.map (fun ((p,_),_) -> dev_form q_spec p) sys in
+ let square =
+ List.fold_left (fun acc p ->
+ Poly.fold
+ (fun m _ acc ->
+ match Monomial.sqrt m with
+ | None -> acc
+ | Some s -> MonMap.add s m acc) p acc) MonMap.empty pols in
-let linear_prover_with_cert spec l =
- match linear_prover spec l with
- | None -> None
- | Some cert -> Some (make_certificate spec cert)
+ let pol_of_mon m =
+ Monomial.fold (fun x v p -> Mc.PEmul(Mc.PEpow(Mc.PEX(Ml2C.positive x),Ml2C.n v),p)) m (Mc.PEc q_spec.unit) in
+
+ let norm0 =
+ Mc.norm q_spec.zero q_spec.unit Mc.qplus Mc.qmult Mc.qminus Mc.qopp Mc.qeq_bool in
+
+
+ MonMap.fold (fun s m acc -> ((pol_of_mon m , Mc.NonStrict), Mc.PsatzSquare(norm0 (pol_of_mon s)))::acc) square [] in
+
+ let sys = sys @ square in
+ (* Call the linear prover without the proofs *)
+ let sys_no_prf = List.map fst sys in
+
+ match linear_prover q_spec sys_no_prf with
+ | None -> None
+ | Some cert ->
+ let cert = make_certificate q_spec cert in
+ let rec map_psatz = function
+ | Mc.PsatzIn n -> snd (List.nth sys (C2Ml.nat n))
+ | Mc.PsatzSquare c -> Mc.PsatzSquare c
+ | Mc.PsatzMulC(c,p) -> Mc.PsatzMulC(c, map_psatz p)
+ | Mc.PsatzMulE(p1,p2) -> Mc.PsatzMulE(map_psatz p1,map_psatz p2)
+ | Mc.PsatzAdd(p1,p2) -> Mc.PsatzAdd(map_psatz p1,map_psatz p2)
+ | Mc.PsatzC c -> Mc.PsatzC c
+ | Mc.PsatzZ -> Mc.PsatzZ in
+ Some (map_psatz cert)
+
let make_linear_system l =
@@ -395,11 +458,11 @@ let make_linear_system l =
(Poly.constant (Int 0)) l' in
let monomials = Poly.fold
(fun mn _ l -> if Pervasives.(=) 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)
+ (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)
let pplus x y = Mc.PEadd(x,y)
@@ -413,7 +476,7 @@ let rec 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)
+ remove_assoc p x l else e::(remove_assoc p x l)
let eq x y = Int.equal (Vect.compare x y) 0
@@ -424,39 +487,39 @@ let remove e l = List.fold_left (fun l x -> if eq x e then l else x::l) [] l
only searching for naive cutting planes *)
let develop_constraint z_spec (e,k) =
- match k with
- | Mc.NonStrict -> (dev_form z_spec e , Ge)
- | Mc.Equal -> (dev_form z_spec e , Eq)
- | _ -> assert false
+ match k with
+ | Mc.NonStrict -> (dev_form z_spec e , Ge)
+ | Mc.Equal -> (dev_form z_spec e , Eq)
+ | _ -> assert false
let op_of_op_compat = function
- | Ge -> Mc.NonStrict
- | Eq -> Mc.Equal
+ | Ge -> Mc.NonStrict
+ | Eq -> Mc.Equal
let integer_vector coeffs =
- let vars , coeffs = List.split coeffs in
- List.combine vars (List.map (fun x -> Big_int x) (rats_to_ints coeffs))
+ let vars , coeffs = List.split coeffs in
+ List.combine vars (List.map (fun x -> Big_int x) (rats_to_ints coeffs))
let integer_cstr {coeffs = coeffs ; op = op ; cst = cst } =
- let vars , coeffs = List.split coeffs in
- match rats_to_ints (cst::coeffs) with
- | cst :: coeffs ->
- {
- coeffs = List.combine vars (List.map (fun x -> Big_int x) coeffs) ;
- op = op ; cst = Big_int cst}
- | _ -> assert false
-
+ let vars , coeffs = List.split coeffs in
+ match rats_to_ints (cst::coeffs) with
+ | cst :: coeffs ->
+ {
+ coeffs = List.combine vars (List.map (fun x -> Big_int x) coeffs) ;
+ op = op ; cst = Big_int cst}
+ | _ -> assert false
+
let pexpr_of_cstr_compat var cstr =
- let {coeffs = coeffs ; op = op ; cst = cst } = integer_cstr cstr in
- try
- let expr = list_to_polynomial var (Vect.to_list coeffs) in
- let d = Ml2C.bigint (denominator cst) in
- let n = Ml2C.bigint (numerator cst) in
- (pplus (pmult (pconst d) expr) (popp (pconst n)), op_of_op_compat op)
- with Failure _ -> failwith "pexpr_of_cstr_compat"
+ let {coeffs = coeffs ; op = op ; cst = cst } = integer_cstr cstr in
+ try
+ let expr = list_to_polynomial var (Vect.to_list coeffs) in
+ let d = Ml2C.bigint (denominator cst) in
+ let n = Ml2C.bigint (numerator cst) in
+ (pplus (pmult (pconst d) expr) (popp (pconst n)), op_of_op_compat op)
+ with Failure _ -> failwith "pexpr_of_cstr_compat"
@@ -465,41 +528,41 @@ open Sos_types
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 Int.equal (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"
+ | 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 Int.equal (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'
+ s,t'
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 Int.equal j i then res else get (j+1) (res+1) l )
- else get j (res+1) l in
- get 0 0 l
+ | [] -> failwith "bad index"
+ | e::l -> if f e
+ then
+ (if Int.equal j i then res else get (j+1) (res+1) l )
+ else get j (res+1) l in
+ get 0 0 l
let rec scale_certificate pos = match pos with
@@ -511,97 +574,97 @@ let rec scale_certificate pos = match pos with
| 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'
+ 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)
+ 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))
+ 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 s1,y1 = scale_certificate y and s2,y2 = scale_certificate z in
+ mult_big_int s1 s2 , Product (y1,y2)
open Micromega
- let rec term_to_q_expr = function
- | Const n -> PEc (Ml2C.q n)
- | Zero -> PEc ( Ml2C.q (Int 0))
- | Var s -> PEX (Ml2C.index
- (int_of_string (String.sub s 1 (String.length s - 1))))
- | Mul(p1,p2) -> PEmul(term_to_q_expr p1, term_to_q_expr p2)
- | Add(p1,p2) -> PEadd(term_to_q_expr p1, term_to_q_expr p2)
- | Opp p -> PEopp (term_to_q_expr p)
- | Pow(t,n) -> PEpow (term_to_q_expr t,Ml2C.n n)
- | Sub(t1,t2) -> PEsub (term_to_q_expr t1, term_to_q_expr t2)
- | _ -> failwith "term_to_q_expr: not implemented"
-
- let term_to_q_pol e = Mc.norm_aux (Ml2C.q (Int 0)) (Ml2C.q (Int 1)) Mc.qplus Mc.qmult Mc.qminus Mc.qopp Mc.qeq_bool (term_to_q_expr e)
-
-
- let rec product l =
- match l with
- | [] -> Mc.PsatzZ
- | [i] -> Mc.PsatzIn (Ml2C.nat i)
- | i ::l -> Mc.PsatzMulE(Mc.PsatzIn (Ml2C.nat i), product l)
+let rec term_to_q_expr = function
+ | Const n -> PEc (Ml2C.q n)
+ | Zero -> PEc ( Ml2C.q (Int 0))
+ | Var s -> PEX (Ml2C.index
+ (int_of_string (String.sub s 1 (String.length s - 1))))
+ | Mul(p1,p2) -> PEmul(term_to_q_expr p1, term_to_q_expr p2)
+ | Add(p1,p2) -> PEadd(term_to_q_expr p1, term_to_q_expr p2)
+ | Opp p -> PEopp (term_to_q_expr p)
+ | Pow(t,n) -> PEpow (term_to_q_expr t,Ml2C.n n)
+ | Sub(t1,t2) -> PEsub (term_to_q_expr t1, term_to_q_expr t2)
+ | _ -> failwith "term_to_q_expr: not implemented"
+
+let term_to_q_pol e = Mc.norm_aux (Ml2C.q (Int 0)) (Ml2C.q (Int 1)) Mc.qplus Mc.qmult Mc.qminus Mc.qopp Mc.qeq_bool (term_to_q_expr e)
+
+
+let rec product l =
+ match l with
+ | [] -> Mc.PsatzZ
+ | [i] -> Mc.PsatzIn (Ml2C.nat i)
+ | i ::l -> Mc.PsatzMulE(Mc.PsatzIn (Ml2C.nat i), product l)
let q_cert_of_pos pos =
let rec _cert_of_pos = function
- Axiom_eq i -> Mc.PsatzIn (Ml2C.nat i)
+ Axiom_eq i -> Mc.PsatzIn (Ml2C.nat i)
| Axiom_le i -> Mc.PsatzIn (Ml2C.nat i)
| Axiom_lt i -> Mc.PsatzIn (Ml2C.nat i)
| Monoid l -> product l
| Rational_eq n | Rational_le n | Rational_lt n ->
- if Int.equal (compare_num n (Int 0)) 0 then Mc.PsatzZ else
- Mc.PsatzC (Ml2C.q n)
+ if Int.equal (compare_num n (Int 0)) 0 then Mc.PsatzZ else
+ Mc.PsatzC (Ml2C.q n)
| Square t -> Mc.PsatzSquare (term_to_q_pol t)
| Eqmul (t, y) -> Mc.PsatzMulC(term_to_q_pol t, _cert_of_pos y)
| Sum (y, z) -> Mc.PsatzAdd (_cert_of_pos y, _cert_of_pos z)
| Product (y, z) -> Mc.PsatzMulE (_cert_of_pos y, _cert_of_pos z) in
- simplify_cone q_spec (_cert_of_pos pos)
+ simplify_cone q_spec (_cert_of_pos pos)
- let rec term_to_z_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_z_expr p1, term_to_z_expr p2)
- | Add(p1,p2) -> PEadd(term_to_z_expr p1, term_to_z_expr p2)
- | Opp p -> PEopp (term_to_z_expr p)
- | Pow(t,n) -> PEpow (term_to_z_expr t,Ml2C.n n)
- | Sub(t1,t2) -> PEsub (term_to_z_expr t1, term_to_z_expr t2)
- | _ -> failwith "term_to_z_expr: not implemented"
+let rec term_to_z_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_z_expr p1, term_to_z_expr p2)
+ | Add(p1,p2) -> PEadd(term_to_z_expr p1, term_to_z_expr p2)
+ | Opp p -> PEopp (term_to_z_expr p)
+ | Pow(t,n) -> PEpow (term_to_z_expr t,Ml2C.n n)
+ | Sub(t1,t2) -> PEsub (term_to_z_expr t1, term_to_z_expr t2)
+ | _ -> failwith "term_to_z_expr: not implemented"
- let term_to_z_pol e = Mc.norm_aux (Ml2C.z 0) (Ml2C.z 1) Mc.Z.add Mc.Z.mul Mc.Z.sub Mc.Z.opp Mc.zeq_bool (term_to_z_expr e)
+let term_to_z_pol e = Mc.norm_aux (Ml2C.z 0) (Ml2C.z 1) Mc.Z.add Mc.Z.mul Mc.Z.sub Mc.Z.opp Mc.zeq_bool (term_to_z_expr e)
let z_cert_of_pos pos =
let s,pos = (scale_certificate pos) in
let rec _cert_of_pos = function
- Axiom_eq i -> Mc.PsatzIn (Ml2C.nat i)
+ Axiom_eq i -> Mc.PsatzIn (Ml2C.nat i)
| Axiom_le i -> Mc.PsatzIn (Ml2C.nat i)
| Axiom_lt i -> Mc.PsatzIn (Ml2C.nat i)
| Monoid l -> product l
| Rational_eq n | Rational_le n | Rational_lt n ->
- if Int.equal (compare_num n (Int 0)) 0 then Mc.PsatzZ else
- Mc.PsatzC (Ml2C.bigint (big_int_of_num n))
+ if Int.equal (compare_num n (Int 0)) 0 then Mc.PsatzZ else
+ Mc.PsatzC (Ml2C.bigint (big_int_of_num n))
| Square t -> Mc.PsatzSquare (term_to_z_pol t)
| Eqmul (t, y) ->
- let is_unit =
- match t with
- | Const n -> n =/ Int 1
- | _ -> false in
- if is_unit
- then _cert_of_pos y
- else Mc.PsatzMulC(term_to_z_pol t, _cert_of_pos y)
+ let is_unit =
+ match t with
+ | Const n -> n =/ Int 1
+ | _ -> false in
+ if is_unit
+ then _cert_of_pos y
+ else Mc.PsatzMulC(term_to_z_pol t, _cert_of_pos y)
| Sum (y, z) -> Mc.PsatzAdd (_cert_of_pos y, _cert_of_pos z)
| Product (y, z) -> Mc.PsatzMulE (_cert_of_pos y, _cert_of_pos z) in
- simplify_cone z_spec (_cert_of_pos pos)
+ simplify_cone z_spec (_cert_of_pos pos)
(** All constraints (initial or derived) have an index and have a justification i.e., proof.
Given a constraint, all the coefficients are always integers.
@@ -612,116 +675,109 @@ open Num
open Big_int
open Polynomial
-(*module Mc = Micromega*)
-(*module Ml2C = Mutils.CamlToCoq
-module C2Ml = Mutils.CoqToCaml
-*)
-let debug = false
-
-
module Env =
struct
- type t = int list
+ type t = int list
- let id_of_hyp hyp l =
- let rec xid_of_hyp i l =
- match l with
- | [] -> failwith "id_of_hyp"
- | hyp'::l -> if Pervasives.(=) hyp hyp' then i else xid_of_hyp (i+1) l in
- xid_of_hyp 0 l
+ let id_of_hyp hyp l =
+ let rec xid_of_hyp i l =
+ match l with
+ | [] -> failwith "id_of_hyp"
+ | hyp'::l -> if Pervasives.(=) hyp hyp' then i else xid_of_hyp (i+1) l in
+ xid_of_hyp 0 l
end
let coq_poly_of_linpol (p,c) =
- let pol_of_mon m =
- Monomial.fold (fun x v p -> Mc.PEmul(Mc.PEpow(Mc.PEX(Ml2C.positive x),Ml2C.n v),p)) m (Mc.PEc (Mc.Zpos Mc.XH)) in
+ let pol_of_mon m =
+ Monomial.fold (fun x v p -> Mc.PEmul(Mc.PEpow(Mc.PEX(Ml2C.positive x),Ml2C.n v),p)) m (Mc.PEc (Mc.Zpos Mc.XH)) in
- List.fold_left (fun acc (x,v) ->
- let mn = LinPoly.MonT.retrieve x in
- Mc.PEadd(Mc.PEmul(Mc.PEc (Ml2C.bigint (numerator v)), pol_of_mon mn),acc)) (Mc.PEc (Ml2C.bigint (numerator c))) p
-
+ List.fold_left (fun acc (x,v) ->
+ let mn = LinPoly.MonT.retrieve x in
+ Mc.PEadd(Mc.PEmul(Mc.PEc (Ml2C.bigint (numerator v)), pol_of_mon mn),acc)) (Mc.PEc (Ml2C.bigint (numerator c))) p
+
let rec cmpl_prf_rule env = function
- | Hyp i | Def i -> Mc.PsatzIn (Ml2C.nat (Env.id_of_hyp i env))
- | Cst i -> Mc.PsatzC (Ml2C.bigint i)
- | Zero -> Mc.PsatzZ
- | MulPrf(p1,p2) -> Mc.PsatzMulE(cmpl_prf_rule env p1, cmpl_prf_rule env p2)
- | AddPrf(p1,p2) -> Mc.PsatzAdd(cmpl_prf_rule env p1 , cmpl_prf_rule env p2)
- | MulC(lp,p) -> let lp = Mc.norm0 (coq_poly_of_linpol lp) in
- Mc.PsatzMulC(lp,cmpl_prf_rule env p)
- | Square lp -> Mc.PsatzSquare (Mc.norm0 (coq_poly_of_linpol lp))
- | _ -> failwith "Cuts should already be compiled"
-
+ | Hyp i | Def i -> Mc.PsatzIn (Ml2C.nat (Env.id_of_hyp i env))
+ | Cst i -> Mc.PsatzC (Ml2C.bigint i)
+ | Zero -> Mc.PsatzZ
+ | MulPrf(p1,p2) -> Mc.PsatzMulE(cmpl_prf_rule env p1, cmpl_prf_rule env p2)
+ | AddPrf(p1,p2) -> Mc.PsatzAdd(cmpl_prf_rule env p1 , cmpl_prf_rule env p2)
+ | MulC(lp,p) -> let lp = Mc.norm0 (coq_poly_of_linpol lp) in
+ Mc.PsatzMulC(lp,cmpl_prf_rule env p)
+ | Square lp -> Mc.PsatzSquare (Mc.norm0 (coq_poly_of_linpol lp))
+ | _ -> failwith "Cuts should already be compiled"
+
let rec cmpl_proof env = function
- | Done -> Mc.DoneProof
- | Step(i,p,prf) ->
- begin
- match p with
- | CutPrf p' ->
- Mc.CutProof(cmpl_prf_rule env p', cmpl_proof (i::env) prf)
- | _ -> Mc.RatProof(cmpl_prf_rule env p,cmpl_proof (i::env) prf)
- end
- | Enum(i,p1,_,p2,l) ->
- Mc.EnumProof(cmpl_prf_rule env p1,cmpl_prf_rule env p2,List.map (cmpl_proof (i::env)) l)
+ | Done -> Mc.DoneProof
+ | Step(i,p,prf) ->
+ begin
+ match p with
+ | CutPrf p' ->
+ Mc.CutProof(cmpl_prf_rule env p', cmpl_proof (i::env) prf)
+ | _ -> Mc.RatProof(cmpl_prf_rule env p,cmpl_proof (i::env) prf)
+ end
+ | Enum(i,p1,_,p2,l) ->
+ Mc.EnumProof(cmpl_prf_rule env p1,cmpl_prf_rule env p2,List.map (cmpl_proof (i::env)) l)
let compile_proof env prf =
- let id = 1 + proof_max_id prf in
- let _,prf = normalise_proof id prf in
- if debug then Printf.fprintf stdout "compiled proof %a\n" output_proof prf;
- cmpl_proof env prf
+ let id = 1 + proof_max_id prf in
+ let _,prf = normalise_proof id prf in
+ if debug then Printf.fprintf stdout "compiled proof %a\n" output_proof prf;
+ cmpl_proof env prf
type prf_sys = (cstr_compat * prf_rule) list
let xlinear_prover sys =
- match Fourier.find_point sys with
- | Inr prf ->
- if debug then Printf.printf "AProof : %a\n" pp_proof prf ;
- let cert = (*List.map (fun (x,n) -> x+1,n)*) (fst (List.hd (Proof.mk_proof sys prf))) in
- if debug then Printf.printf "CProof : %a" Vect.pp_vect cert ;
- Some (rats_to_ints (Vect.to_list cert))
- | Inl _ -> None
+ match Fourier.find_point sys with
+ | Inr prf ->
+ if debug then Printf.printf "AProof : %a\n" pp_proof prf ;
+ let cert = (*List.map (fun (x,n) -> x+1,n)*) (fst (List.hd (Proof.mk_proof sys prf))) in
+ if debug then Printf.printf "CProof : %a" Vect.pp_vect cert ;
+ Some (rats_to_ints (Vect.to_list cert))
+ | Inl _ -> None
let output_num o n = output_string o (string_of_num n)
let output_bigint o n = output_string o (string_of_big_int n)
let proof_of_farkas prf cert =
-(* Printf.printf "\nproof_of_farkas %a , %a \n" (pp_list output_prf_rule) prf (pp_list output_bigint) cert ; *)
- let rec mk_farkas acc prf cert =
- match prf, cert with
- | _ , [] -> acc
- | [] , _ -> failwith "proof_of_farkas : not enough hyps"
- | p::prf,c::cert ->
- mk_farkas (add_proof (mul_proof c p) acc) prf cert in
- let res = mk_farkas Zero prf cert in
+ (* Printf.printf "\nproof_of_farkas %a , %a \n" (pp_list output_prf_rule) prf (pp_list output_bigint) cert ; *)
+ let rec mk_farkas acc prf cert =
+ match prf, cert with
+ | _ , [] -> acc
+ | [] , _ -> failwith "proof_of_farkas : not enough hyps"
+ | p::prf,c::cert ->
+ mk_farkas (add_proof (mul_proof c p) acc) prf cert in
+ let res = mk_farkas Zero prf cert in
(*Printf.printf "==> %a" output_prf_rule res ; *)
- res
+ res
let linear_prover sys =
- let (sysi,prfi) = List.split sys in
- match xlinear_prover sysi with
- | None -> None
- | Some cert -> Some (proof_of_farkas prfi cert)
+ let (sysi,prfi) = List.split sys in
+ match xlinear_prover sysi with
+ | None -> None
+ | Some cert -> Some (proof_of_farkas prfi cert)
let linear_prover =
- if debug
- then
- fun sys ->
- Printf.printf "<linear_prover"; flush stdout ;
- let res = linear_prover sys in
- Printf.printf ">"; flush stdout ;
- res
- else linear_prover
+ if debug
+ then
+ fun sys ->
+ Printf.printf "<linear_prover"; flush stdout ;
+ let res = linear_prover sys in
+ Printf.printf ">"; flush stdout ;
+ res
+ else linear_prover
@@ -733,11 +789,11 @@ let linear_prover =
*)
type checksat =
- | Tauto (* Tautology *)
- | Unsat of prf_rule (* Unsatisfiable *)
- | Cut of cstr_compat * prf_rule (* Cutting plane *)
- | Normalise of cstr_compat * prf_rule (* coefficients are relatively prime *)
-
+| Tauto (* Tautology *)
+| Unsat of prf_rule (* Unsatisfiable *)
+| Cut of cstr_compat * prf_rule (* Cutting plane *)
+| Normalise of cstr_compat * prf_rule (* coefficients are relatively prime *)
+
(** [check_sat]
- detects constraints that are not satisfiable;
@@ -745,83 +801,83 @@ type checksat =
*)
let check_sat (cstr,prf) =
- let {coeffs=coeffs ; op=op ; cst=cst} = cstr in
- match coeffs with
- | [] ->
- if eval_op op (Int 0) cst then Tauto else Unsat prf
- | _ ->
- let gcdi = (gcd_list (List.map snd coeffs)) in
- let gcd = Big_int gcdi in
- if eq_num gcd (Int 1)
- then Normalise(cstr,prf)
- else
- if Int.equal (sign_num (mod_num cst gcd)) 0
- then (* We can really normalise *)
- begin
- assert (sign_num gcd >=1 ) ;
- let cstr = {
- coeffs = List.map (fun (x,v) -> (x, v // gcd)) coeffs;
- op = op ; cst = cst // gcd
- } in
- Normalise(cstr,Gcd(gcdi,prf))
- (* Normalise(cstr,CutPrf prf)*)
- end
- else
- match op with
- | Eq -> Unsat (CutPrf prf)
- | Ge ->
- let cstr = {
- coeffs = List.map (fun (x,v) -> (x, v // gcd)) coeffs;
- op = op ; cst = ceiling_num (cst // gcd)
- } in Cut(cstr,CutPrf prf)
+ let {coeffs=coeffs ; op=op ; cst=cst} = cstr in
+ match coeffs with
+ | [] ->
+ if eval_op op (Int 0) cst then Tauto else Unsat prf
+ | _ ->
+ let gcdi = (gcd_list (List.map snd coeffs)) in
+ let gcd = Big_int gcdi in
+ if eq_num gcd (Int 1)
+ then Normalise(cstr,prf)
+ else
+ if Int.equal (sign_num (mod_num cst gcd)) 0
+ then (* We can really normalise *)
+ begin
+ assert (sign_num gcd >=1 ) ;
+ let cstr = {
+ coeffs = List.map (fun (x,v) -> (x, v // gcd)) coeffs;
+ op = op ; cst = cst // gcd
+ } in
+ Normalise(cstr,Gcd(gcdi,prf))
+ (* Normalise(cstr,CutPrf prf)*)
+ end
+ else
+ match op with
+ | Eq -> Unsat (CutPrf prf)
+ | Ge ->
+ let cstr = {
+ coeffs = List.map (fun (x,v) -> (x, v // gcd)) coeffs;
+ op = op ; cst = ceiling_num (cst // gcd)
+ } in Cut(cstr,CutPrf prf)
(** Proof generating pivoting over variable v *)
let pivot v (c1,p1) (c2,p2) =
- let {coeffs = v1 ; op = op1 ; cst = n1} = c1
- and {coeffs = v2 ; op = op2 ; cst = n2} = c2 in
+ let {coeffs = v1 ; op = op1 ; cst = n1} = c1
+ and {coeffs = v2 ; op = op2 ; cst = n2} = c2 in
(* Could factorise gcd... *)
- let xpivot cv1 cv2 =
- (
- {coeffs = Vect.add (Vect.mul cv1 v1) (Vect.mul cv2 v2) ;
- op = Proof.add_op op1 op2 ;
- cst = n1 */ cv1 +/ n2 */ cv2 },
+ let xpivot cv1 cv2 =
+ (
+ {coeffs = Vect.add (Vect.mul cv1 v1) (Vect.mul cv2 v2) ;
+ op = Proof.add_op op1 op2 ;
+ cst = n1 */ cv1 +/ n2 */ cv2 },
- AddPrf(mul_proof (numerator cv1) p1,mul_proof (numerator cv2) p2)) in
+ AddPrf(mul_proof (numerator cv1) p1,mul_proof (numerator cv2) p2)) in
+
+ match Vect.get v v1 , Vect.get v v2 with
+ | None , _ | _ , None -> None
+ | Some a , Some b ->
+ if Int.equal ((sign_num a) * (sign_num b)) (-1)
+ then
+ let cv1 = abs_num b
+ and cv2 = abs_num a in
+ Some (xpivot cv1 cv2)
+ else
+ if op1 == Eq
+ then
+ let cv1 = minus_num (b */ (Int (sign_num a)))
+ and cv2 = abs_num a in
+ Some (xpivot cv1 cv2)
+ else if op2 == Eq
+ then
+ let cv1 = abs_num b
+ and cv2 = minus_num (a */ (Int (sign_num b))) in
+ Some (xpivot cv1 cv2)
+ else None (* op2 could be Eq ... this might happen *)
- match Vect.get v v1 , Vect.get v v2 with
- | None , _ | _ , None -> None
- | Some a , Some b ->
- if Int.equal ((sign_num a) * (sign_num b)) (-1)
- then
- let cv1 = abs_num b
- and cv2 = abs_num a in
- Some (xpivot cv1 cv2)
- else
- if op1 == Eq
- then
- let cv1 = minus_num (b */ (Int (sign_num a)))
- and cv2 = abs_num a in
- Some (xpivot cv1 cv2)
- else if op2 == Eq
- then
- let cv1 = abs_num b
- and cv2 = minus_num (a */ (Int (sign_num b))) in
- Some (xpivot cv1 cv2)
- else None (* op2 could be Eq ... this might happen *)
-
exception FoundProof of prf_rule
let simpl_sys sys =
- List.fold_left (fun acc (c,p) ->
- match check_sat (c,p) with
- | Tauto -> acc
- | Unsat prf -> raise (FoundProof prf)
- | Cut(c,p) -> (c,p)::acc
- | Normalise (c,p) -> (c,p)::acc) [] sys
+ List.fold_left (fun acc (c,p) ->
+ match check_sat (c,p) with
+ | Tauto -> acc
+ | Unsat prf -> raise (FoundProof prf)
+ | Cut(c,p) -> (c,p)::acc
+ | Normalise (c,p) -> (c,p)::acc) [] sys
(** [ext_gcd a b] is the extended Euclid algorithm.
@@ -829,77 +885,77 @@ let simpl_sys sys =
Source: http://en.wikipedia.org/wiki/Extended_Euclidean_algorithm
*)
let rec ext_gcd a b =
- if Int.equal (sign_big_int b) 0
- then (unit_big_int,zero_big_int)
- else
- let (q,r) = quomod_big_int a b in
- let (s,t) = ext_gcd b r in
- (t, sub_big_int s (mult_big_int q t))
+ if Int.equal (sign_big_int b) 0
+ then (unit_big_int,zero_big_int)
+ else
+ let (q,r) = quomod_big_int a b in
+ let (s,t) = ext_gcd b r in
+ (t, sub_big_int s (mult_big_int q t))
let pp_ext_gcd a b =
- let a' = big_int_of_int a in
- let b' = big_int_of_int b in
-
- let (x,y) = ext_gcd a' b' in
- Printf.fprintf stdout "%s * %s + %s * %s = %s\n"
- (string_of_big_int x) (string_of_big_int a')
- (string_of_big_int y) (string_of_big_int b')
- (string_of_big_int (add_big_int (mult_big_int x a') (mult_big_int y b')))
+ let a' = big_int_of_int a in
+ let b' = big_int_of_int b in
+
+ let (x,y) = ext_gcd a' b' in
+ Printf.fprintf stdout "%s * %s + %s * %s = %s\n"
+ (string_of_big_int x) (string_of_big_int a')
+ (string_of_big_int y) (string_of_big_int b')
+ (string_of_big_int (add_big_int (mult_big_int x a') (mult_big_int y b')))
exception Result of (int * (proof * cstr_compat))
let split_equations psys =
- List.partition (fun (c,p) -> c.op == Eq)
+ List.partition (fun (c,p) -> c.op == Eq)
let extract_coprime (c1,p1) (c2,p2) =
- let rec exist2 vect1 vect2 =
- match vect1 , vect2 with
- | _ , [] | [], _ -> None
- | (v1,n1)::vect1' , (v2, n2) :: vect2' ->
- if Pervasives.(=) v1 v2
- then
- if Int.equal (compare_big_int (gcd_big_int (numerator n1) (numerator n2)) unit_big_int) 0
- then Some (v1,n1,n2)
- else
- exist2 vect1' vect2'
- else
- if v1 < v2
- then exist2 vect1' vect2
- else exist2 vect1 vect2' in
-
- if c1.op == Eq && c2.op == Eq
- then exist2 c1.coeffs c2.coeffs
- else None
+ let rec exist2 vect1 vect2 =
+ match vect1 , vect2 with
+ | _ , [] | [], _ -> None
+ | (v1,n1)::vect1' , (v2, n2) :: vect2' ->
+ if Pervasives.(=) v1 v2
+ then
+ if Int.equal (compare_big_int (gcd_big_int (numerator n1) (numerator n2)) unit_big_int) 0
+ then Some (v1,n1,n2)
+ else
+ exist2 vect1' vect2'
+ else
+ if v1 < v2
+ then exist2 vect1' vect2
+ else exist2 vect1 vect2' in
+
+ if c1.op == Eq && c2.op == Eq
+ then exist2 c1.coeffs c2.coeffs
+ else None
let extract2 pred l =
- let rec xextract2 rl l =
- match l with
- | [] -> (None,rl) (* Did not find *)
- | e::l ->
- match extract (pred e) l with
- | None,_ -> xextract2 (e::rl) l
- | Some (r,e'),l' -> Some (r,e,e'), List.rev_append rl l' in
-
- xextract2 [] l
+ let rec xextract2 rl l =
+ match l with
+ | [] -> (None,rl) (* Did not find *)
+ | e::l ->
+ match extract (pred e) l with
+ | None,_ -> xextract2 (e::rl) l
+ | Some (r,e'),l' -> Some (r,e,e'), List.rev_append rl l' in
+
+ xextract2 [] l
let extract_coprime_equation psys =
- extract2 extract_coprime psys
+ extract2 extract_coprime psys
let apply_and_normalise f psys =
- List.fold_left (fun acc pc' ->
- match f pc' with
- | None -> pc'::acc
- | Some pc' ->
- match check_sat pc' with
- | Tauto -> acc
- | Unsat prf -> raise (FoundProof prf)
- | Cut(c,p) -> (c,p)::acc
- | Normalise (c,p) -> (c,p)::acc
- ) [] psys
+ List.fold_left (fun acc pc' ->
+ match f pc' with
+ | None -> pc'::acc
+ | Some pc' ->
+ match check_sat pc' with
+ | Tauto -> acc
+ | Unsat prf -> raise (FoundProof prf)
+ | Cut(c,p) -> (c,p)::acc
+ | Normalise (c,p) -> (c,p)::acc
+ ) [] psys
@@ -908,314 +964,317 @@ let pivot_sys v pc psys = apply_and_normalise (pivot v pc) psys
let reduce_coprime psys =
- let oeq,sys = extract_coprime_equation psys in
- match oeq with
- | None -> None (* Nothing to do *)
- | Some((v,n1,n2),(c1,p1),(c2,p2) ) ->
- let (l1,l2) = ext_gcd (numerator n1) (numerator n2) in
- let l1' = Big_int l1 and l2' = Big_int l2 in
- let cstr =
- {coeffs = Vect.add (Vect.mul l1' c1.coeffs) (Vect.mul l2' c2.coeffs);
- op = Eq ;
- cst = (l1' */ c1.cst) +/ (l2' */ c2.cst)
- } in
- let prf = add_proof (mul_proof (numerator l1') p1) (mul_proof (numerator l2') p2) in
-
- Some (pivot_sys v (cstr,prf) ((c1,p1)::sys))
+ let oeq,sys = extract_coprime_equation psys in
+ match oeq with
+ | None -> None (* Nothing to do *)
+ | Some((v,n1,n2),(c1,p1),(c2,p2) ) ->
+ let (l1,l2) = ext_gcd (numerator n1) (numerator n2) in
+ let l1' = Big_int l1 and l2' = Big_int l2 in
+ let cstr =
+ {coeffs = Vect.add (Vect.mul l1' c1.coeffs) (Vect.mul l2' c2.coeffs);
+ op = Eq ;
+ cst = (l1' */ c1.cst) +/ (l2' */ c2.cst)
+ } in
+ let prf = add_proof (mul_proof (numerator l1') p1) (mul_proof (numerator l2') p2) in
+
+ Some (pivot_sys v (cstr,prf) ((c1,p1)::sys))
(** If there is an equation [eq] of the form 1.x + e = c, do a pivot over x with equation [eq] *)
let reduce_unary psys =
- let is_unary_equation (cstr,prf) =
- if cstr.op == Eq
- then
- try
- Some (fst (List.find (fun (_,n) -> n =/ (Int 1) || n=/ (Int (-1))) cstr.coeffs))
- with Not_found -> None
- else None in
-
- let (oeq,sys) = extract is_unary_equation psys in
- match oeq with
- | None -> None (* Nothing to do *)
- | Some(v,pc) ->
- Some(pivot_sys v pc sys)
+ let is_unary_equation (cstr,prf) =
+ if cstr.op == Eq
+ then
+ try
+ Some (fst (List.find (fun (_,n) -> n =/ (Int 1) || n=/ (Int (-1))) cstr.coeffs))
+ with Not_found -> None
+ else None in
+
+ let (oeq,sys) = extract is_unary_equation psys in
+ match oeq with
+ | None -> None (* Nothing to do *)
+ | Some(v,pc) ->
+ Some(pivot_sys v pc sys)
let reduce_non_lin_unary psys =
- let is_unary_equation (cstr,prf) =
- if cstr.op == Eq
- then
- try
- let x = fst (List.find (fun (x,n) -> (n =/ (Int 1) || n=/ (Int (-1))) && Monomial.is_var (LinPoly.MonT.retrieve x) ) cstr.coeffs) in
- let x' = LinPoly.MonT.retrieve x in
- if List.for_all (fun (y,_) -> Pervasives.(=) y x || Int.equal (snd (Monomial.div (LinPoly.MonT.retrieve y) x')) 0) cstr.coeffs
- then Some x
- else None
- with Not_found -> None
- else None in
-
-
- let (oeq,sys) = extract is_unary_equation psys in
- match oeq with
- | None -> None (* Nothing to do *)
- | Some(v,pc) ->
- Some(apply_and_normalise (LinPoly.pivot_eq v pc) sys)
+ let is_unary_equation (cstr,prf) =
+ if cstr.op == Eq
+ then
+ try
+ let x = fst (List.find (fun (x,n) -> (n =/ (Int 1) || n=/ (Int (-1))) && Monomial.is_var (LinPoly.MonT.retrieve x) ) cstr.coeffs) in
+ let x' = LinPoly.MonT.retrieve x in
+ if List.for_all (fun (y,_) -> Pervasives.(=) y x || Int.equal (snd (Monomial.div (LinPoly.MonT.retrieve y) x')) 0) cstr.coeffs
+ then Some x
+ else None
+ with Not_found -> None
+ else None in
+
+
+ let (oeq,sys) = extract is_unary_equation psys in
+ match oeq with
+ | None -> None (* Nothing to do *)
+ | Some(v,pc) ->
+ Some(apply_and_normalise (LinPoly.pivot_eq v pc) sys)
let reduce_var_change psys =
- let rec rel_prime vect =
- match vect with
- | [] -> None
- | (x,v)::vect ->
- let v = numerator v in
- try
- let (x',v') = List.find (fun (_,v') ->
- let v' = numerator v' in
- eq_big_int (gcd_big_int v v') unit_big_int) vect in
- Some ((x,v),(x',numerator v'))
- with Not_found -> rel_prime vect in
-
- let rel_prime (cstr,prf) = if cstr.op == Eq then rel_prime cstr.coeffs else None in
-
- let (oeq,sys) = extract rel_prime psys in
-
- match oeq with
- | None -> None
- | Some(((x,v),(x',v')),(c,p)) ->
- let (l1,l2) = ext_gcd v v' in
- let l1,l2 = Big_int l1 , Big_int l2 in
+ let rec rel_prime vect =
+ match vect with
+ | [] -> None
+ | (x,v)::vect ->
+ let v = numerator v in
+ try
+ let (x',v') = List.find (fun (_,v') ->
+ let v' = numerator v' in
+ eq_big_int (gcd_big_int v v') unit_big_int) vect in
+ Some ((x,v),(x',numerator v'))
+ with Not_found -> rel_prime vect in
+
+ let rel_prime (cstr,prf) = if cstr.op == Eq then rel_prime cstr.coeffs else None in
- let get v vect =
- match Vect.get v vect with
- | None -> Int 0
- | Some n -> n in
+ let (oeq,sys) = extract rel_prime psys in
+
+ match oeq with
+ | None -> None
+ | Some(((x,v),(x',v')),(c,p)) ->
+ let (l1,l2) = ext_gcd v v' in
+ let l1,l2 = Big_int l1 , Big_int l2 in
- let pivot_eq (c',p') =
- let {coeffs = coeffs ; op = op ; cst = cst} = c' in
- let vx = get x coeffs in
- let vx' = get x' coeffs in
- let m = minus_num (vx */ l1 +/ vx' */ l2) in
- Some ({coeffs =
- Vect.add (Vect.mul m c.coeffs) coeffs ; op = op ; cst = m */ c.cst +/ cst} ,
- AddPrf(MulC(([], m),p),p')) in
+ let get v vect =
+ match Vect.get v vect with
+ | None -> Int 0
+ | Some n -> n in
- Some (apply_and_normalise pivot_eq sys)
+ let pivot_eq (c',p') =
+ let {coeffs = coeffs ; op = op ; cst = cst} = c' in
+ let vx = get x coeffs in
+ let vx' = get x' coeffs in
+ let m = minus_num (vx */ l1 +/ vx' */ l2) in
+ Some ({coeffs =
+ Vect.add (Vect.mul m c.coeffs) coeffs ; op = op ; cst = m */ c.cst +/ cst} ,
+ AddPrf(MulC(([], m),p),p')) in
+ Some (apply_and_normalise pivot_eq sys)
- let reduce_pivot psys =
- let is_equation (cstr,prf) =
- if cstr.op == Eq
- then
- try
- Some (fst (List.hd cstr.coeffs))
- with Not_found -> None
- else None in
- let (oeq,sys) = extract is_equation psys in
- match oeq with
- | None -> None (* Nothing to do *)
- | Some(v,pc) ->
- if debug then
- Printf.printf "Bad news : loss of completeness %a=%s" Vect.pp_vect (fst pc).coeffs (string_of_num (fst pc).cst);
- Some(pivot_sys v pc sys)
+let reduce_pivot psys =
+ let is_equation (cstr,prf) =
+ if cstr.op == Eq
+ then
+ try
+ Some (fst (List.hd cstr.coeffs))
+ with Not_found -> None
+ else None in
+ let (oeq,sys) = extract is_equation psys in
+ match oeq with
+ | None -> None (* Nothing to do *)
+ | Some(v,pc) ->
+ if debug then
+ Printf.printf "Bad news : loss of completeness %a=%s" Vect.pp_vect (fst pc).coeffs (string_of_num (fst pc).cst);
+ Some(pivot_sys v pc sys)
- let iterate_until_stable f x =
- let rec iter x =
- match f x with
- | None -> x
- | Some x' -> iter x' in
- iter x
- let rec app_funs l x =
- match l with
- | [] -> None
- | f::fl ->
- match f x with
- | None -> app_funs fl x
- | Some x' -> Some x'
+let iterate_until_stable f x =
+ let rec iter x =
+ match f x with
+ | None -> x
+ | Some x' -> iter x' in
+ iter x
- let reduction_equations psys =
- iterate_until_stable (app_funs
- [reduce_unary ; reduce_coprime ;
- reduce_var_change (*; reduce_pivot*)]) psys
+let rec app_funs l x =
+ match l with
+ | [] -> None
+ | f::fl ->
+ match f x with
+ | None -> app_funs fl x
+ | Some x' -> Some x'
- let reduction_non_lin_equations psys =
- iterate_until_stable (app_funs
- [reduce_non_lin_unary (*; reduce_coprime ;
- reduce_var_change ; reduce_pivot *)]) psys
+let reduction_equations psys =
+ iterate_until_stable (app_funs
+ [reduce_unary ; reduce_coprime ;
+ reduce_var_change (*; reduce_pivot*)]) psys
+
+let reduction_non_lin_equations psys =
+ iterate_until_stable (app_funs
+ [reduce_non_lin_unary (*; reduce_coprime ;
+ reduce_var_change ; reduce_pivot *)]) psys
(** [get_bound sys] returns upon success an interval (lb,e,ub) with proofs *)
- let get_bound sys =
- let is_small (v,i) =
- match Itv.range i with
- | None -> false
- | Some i -> i <=/ (Int 1) in
-
- let select_best (x1,i1) (x2,i2) =
- if Itv.smaller_itv i1 i2
- then (x1,i1) else (x2,i2) in
+let get_bound sys =
+ let is_small (v,i) =
+ match Itv.range i with
+ | None -> false
+ | Some i -> i <=/ (Int 1) in
+
+ let select_best (x1,i1) (x2,i2) =
+ if Itv.smaller_itv i1 i2
+ then (x1,i1) else (x2,i2) in
(* For lia, there are no equations => these precautions are not needed *)
(* For nlia, there are equations => do not enumerate over equations! *)
- let all_planes sys =
- let (eq,ineq) = List.partition (fun c -> c.op == Eq) sys in
- match eq with
- | [] -> List.rev_map (fun c -> c.coeffs) ineq
- | _ ->
- List.fold_left (fun acc c ->
- if List.exists (fun c' -> Vect.equal c.coeffs c'.coeffs) eq
- then acc else c.coeffs ::acc) [] ineq in
-
- let smallest_interval =
- List.fold_left
- (fun acc vect ->
- if is_small acc
- then acc
- else
- match Fourier.optimise vect sys with
- | None -> acc
- | Some i ->
- if debug then Printf.printf "Found a new bound %a" Vect.pp_vect vect ;
- select_best (vect,i) acc) (Vect.null, (None,None)) (all_planes sys) in
- let smallest_interval =
- match smallest_interval
- with
- | (x,(Some i, Some j)) -> Some(i,x,j)
- | x -> None (* This should not be possible *)
- in
- match smallest_interval with
- | Some (lb,e,ub) ->
- let (lbn,lbd) = (sub_big_int (numerator lb) unit_big_int, denominator lb) in
- let (ubn,ubd) = (add_big_int unit_big_int (numerator ub) , denominator ub) in
- (match
+ let all_planes sys =
+ let (eq,ineq) = List.partition (fun c -> c.op == Eq) sys in
+ match eq with
+ | [] -> List.rev_map (fun c -> c.coeffs) ineq
+ | _ ->
+ List.fold_left (fun acc c ->
+ if List.exists (fun c' -> Vect.equal c.coeffs c'.coeffs) eq
+ then acc else c.coeffs ::acc) [] ineq in
+
+ let smallest_interval =
+ List.fold_left
+ (fun acc vect ->
+ if is_small acc
+ then acc
+ else
+ match Fourier.optimise vect sys with
+ | None -> acc
+ | Some i ->
+ if debug then Printf.printf "Found a new bound %a" Vect.pp_vect vect ;
+ select_best (vect,i) acc) (Vect.null, (None,None)) (all_planes sys) in
+ let smallest_interval =
+ match smallest_interval
+ with
+ | (x,(Some i, Some j)) -> Some(i,x,j)
+ | x -> None (* This should not be possible *)
+ in
+ match smallest_interval with
+ | Some (lb,e,ub) ->
+ let (lbn,lbd) = (sub_big_int (numerator lb) unit_big_int, denominator lb) in
+ let (ubn,ubd) = (add_big_int unit_big_int (numerator ub) , denominator ub) in
+ (match
(* x <= ub -> x > ub *)
- xlinear_prover ({coeffs = Vect.mul (Big_int ubd) e ; op = Ge ; cst = Big_int ubn} :: sys),
+ xlinear_prover ({coeffs = Vect.mul (Big_int ubd) e ; op = Ge ; cst = Big_int ubn} :: sys),
(* lb <= x -> lb > x *)
- xlinear_prover
- ({coeffs = Vect.mul (minus_num (Big_int lbd)) e ; op = Ge ; cst = minus_num (Big_int lbn)} :: sys)
- with
- | Some cub , Some clb -> Some(List.tl clb,(lb,e,ub), List.tl cub)
- | _ -> failwith "Interval without proof"
- )
- | None -> None
-
-
- let check_sys sys =
- List.for_all (fun (c,p) -> List.for_all (fun (_,n) -> sign_num n <> 0) c.coeffs) sys
-
-
- let xlia reduction_equations sys =
-
- let rec enum_proof (id:int) (sys:prf_sys) : proof option =
- if debug then (Printf.printf "enum_proof\n" ; flush stdout) ;
- assert (check_sys sys) ;
-
- let nsys,prf = List.split sys in
- match get_bound nsys with
- | None -> None (* Is the systeme really unbounded ? *)
- | Some(prf1,(lb,e,ub),prf2) ->
- if debug then Printf.printf "Found interval: %a in [%s;%s] -> " Vect.pp_vect e (string_of_num lb) (string_of_num ub) ;
- (match start_enum id e (ceiling_num lb) (floor_num ub) sys
- with
- | Some prfl ->
- Some(Enum(id,proof_of_farkas prf prf1,e, proof_of_farkas prf prf2,prfl))
- | None -> None
- )
-
- and start_enum id e clb cub sys =
- if clb >/ cub
- then Some []
- else
- let eq = {coeffs = e ; op = Eq ; cst = clb} in
- match aux_lia (id+1) ((eq, Def id) :: sys) with
- | None -> None
- | Some prf ->
- match start_enum id e (clb +/ (Int 1)) cub sys with
- | None -> None
- | Some l -> Some (prf::l)
-
- and aux_lia (id:int) (sys:prf_sys) : proof option =
- assert (check_sys sys) ;
- if debug then Printf.printf "xlia: %a \n" (pp_list (fun o (c,_) -> output_cstr o c)) sys ;
- try
- let sys = reduction_equations sys in
- if debug then
+ xlinear_prover
+ ({coeffs = Vect.mul (minus_num (Big_int lbd)) e ; op = Ge ; cst = minus_num (Big_int lbn)} :: sys)
+ with
+ | Some cub , Some clb -> Some(List.tl clb,(lb,e,ub), List.tl cub)
+ | _ -> failwith "Interval without proof"
+ )
+ | None -> None
+
+
+let check_sys sys =
+ List.for_all (fun (c,p) -> List.for_all (fun (_,n) -> sign_num n <> 0) c.coeffs) sys
+
+
+let xlia (can_enum:bool) reduction_equations sys =
+
+
+ let rec enum_proof (id:int) (sys:prf_sys) : proof option =
+ if debug then (Printf.printf "enum_proof\n" ; flush stdout) ;
+ assert (check_sys sys) ;
+
+ let nsys,prf = List.split sys in
+ match get_bound nsys with
+ | None -> None (* Is the systeme really unbounded ? *)
+ | Some(prf1,(lb,e,ub),prf2) ->
+ if debug then Printf.printf "Found interval: %a in [%s;%s] -> " Vect.pp_vect e (string_of_num lb) (string_of_num ub) ;
+ (match start_enum id e (ceiling_num lb) (floor_num ub) sys
+ with
+ | Some prfl ->
+ Some(Enum(id,proof_of_farkas prf prf1,e, proof_of_farkas prf prf2,prfl))
+ | None -> None
+ )
+
+ and start_enum id e clb cub sys =
+ if clb >/ cub
+ then Some []
+ else
+ let eq = {coeffs = e ; op = Eq ; cst = clb} in
+ match aux_lia (id+1) ((eq, Def id) :: sys) with
+ | None -> None
+ | Some prf ->
+ match start_enum id e (clb +/ (Int 1)) cub sys with
+ | None -> None
+ | Some l -> Some (prf::l)
+
+ and aux_lia (id:int) (sys:prf_sys) : proof option =
+ assert (check_sys sys) ;
+ if debug then Printf.printf "xlia: %a \n" (pp_list (fun o (c,_) -> output_cstr o c)) sys ;
+ try
+ let sys = reduction_equations sys in
+ if debug then
Printf.printf "after reduction: %a \n" (pp_list (fun o (c,_) -> output_cstr o c)) sys ;
- match linear_prover sys with
- | Some prf -> Some (Step(id,prf,Done))
- | None -> enum_proof id sys
- with FoundProof prf ->
+ match linear_prover sys with
+ | Some prf -> Some (Step(id,prf,Done))
+ | None -> if can_enum then enum_proof id sys else None
+ with FoundProof prf ->
(* [reduction_equations] can find a proof *)
- Some(Step(id,prf,Done)) in
+ Some(Step(id,prf,Done)) in
(* let sys' = List.map (fun (p,o) -> Mc.norm0 p , o) sys in*)
- let id = List.length sys in
- let orpf =
- try
- let sys = simpl_sys sys in
- aux_lia id sys
- with FoundProof pr -> Some(Step(id,pr,Done)) in
- match orpf with
- | None -> None
- | Some prf ->
+ let id = List.length sys in
+ let orpf =
+ try
+ let sys = simpl_sys sys in
+ aux_lia id sys
+ with FoundProof pr -> Some(Step(id,pr,Done)) in
+ match orpf with
+ | None -> None
+ | Some prf ->
(*Printf.printf "direct proof %a\n" output_proof prf ; *)
- let env = mapi (fun _ i -> i) sys in
- let prf = compile_proof env prf in
+ let env = mapi (fun _ i -> i) sys in
+ let prf = compile_proof env prf in
(*try
if Mc.zChecker sys' prf then Some prf else
raise Certificate.BadCertificate
with Failure s -> (Printf.printf "%s" s ; Some prf)
*) Some prf
-
-
- let cstr_compat_of_poly (p,o) =
- let (v,c) = LinPoly.linpol_of_pol p in
- {coeffs = v ; op = o ; cst = minus_num c }
-
-
- let lia sys =
- LinPoly.MonT.clear ();
- let sys = List.map (develop_constraint z_spec) sys in
- let (sys:cstr_compat list) = List.map cstr_compat_of_poly sys in
- let sys = mapi (fun c i -> (c,Hyp i)) sys in
- xlia reduction_equations sys
-
-
- let nlia sys =
- LinPoly.MonT.clear ();
- let sys = List.map (develop_constraint z_spec) sys in
- let sys = mapi (fun c i -> (c,Hyp i)) sys in
-
- let is_linear = List.for_all (fun ((p,_),_) -> Poly.is_linear p) sys in
-
- let collect_square =
- List.fold_left (fun acc ((p,_),_) -> Poly.fold
- (fun m _ acc ->
- match Monomial.sqrt m with
- | None -> acc
- | Some s -> MonMap.add s m acc) p acc) MonMap.empty sys in
- let sys = MonMap.fold (fun s m acc ->
- let s = LinPoly.linpol_of_pol (Poly.add s (Int 1) (Poly.constant (Int 0))) in
- let m = Poly.add m (Int 1) (Poly.constant (Int 0)) in
- ((m, Ge), (Square s))::acc) collect_square sys in
-
-(* List.iter (fun ((p,_),_) -> Printf.printf "square %a\n" Poly.pp p) gen_square*)
-
- let sys =
- if is_linear then sys
- else sys @ (all_sym_pairs (fun ((c,o),p) ((c',o'),p') ->
- ((Poly.product c c',opMult o o'), MulPrf(p,p'))) sys) in
+
- let sys = List.map (fun (c,p) -> cstr_compat_of_poly c,p) sys in
- assert (check_sys sys) ;
- xlia (if is_linear then reduction_equations else reduction_non_lin_equations) sys
+let cstr_compat_of_poly (p,o) =
+ let (v,c) = LinPoly.linpol_of_pol p in
+ {coeffs = v ; op = o ; cst = minus_num c }
+
+
+let lia (can_enum:bool) (prfdepth:int) sys =
+ LinPoly.MonT.clear ();
+ max_nb_cstr := compute_max_nb_cstr sys prfdepth ;
+ let sys = List.map (develop_constraint z_spec) sys in
+ let (sys:cstr_compat list) = List.map cstr_compat_of_poly sys in
+ let sys = mapi (fun c i -> (c,Hyp i)) sys in
+ xlia can_enum reduction_equations sys
+
+
+let nlia enum prfdepth sys =
+ LinPoly.MonT.clear ();
+ max_nb_cstr := compute_max_nb_cstr sys prfdepth;
+ let sys = List.map (develop_constraint z_spec) sys in
+ let sys = mapi (fun c i -> (c,Hyp i)) sys in
+
+ let is_linear = List.for_all (fun ((p,_),_) -> Poly.is_linear p) sys in
+
+ let collect_square =
+ List.fold_left (fun acc ((p,_),_) -> Poly.fold
+ (fun m _ acc ->
+ match Monomial.sqrt m with
+ | None -> acc
+ | Some s -> MonMap.add s m acc) p acc) MonMap.empty sys in
+ let sys = MonMap.fold (fun s m acc ->
+ let s = LinPoly.linpol_of_pol (Poly.add s (Int 1) (Poly.constant (Int 0))) in
+ let m = Poly.add m (Int 1) (Poly.constant (Int 0)) in
+ ((m, Ge), (Square s))::acc) collect_square sys in
+
+ (* List.iter (fun ((p,_),_) -> Printf.printf "square %a\n" Poly.pp p) gen_square*)
+
+ let sys =
+ if is_linear then sys
+ else sys @ (all_sym_pairs (fun ((c,o),p) ((c',o'),p') ->
+ ((Poly.product c c',opMult o o'), MulPrf(p,p'))) sys) in
+
+ let sys = List.map (fun (c,p) -> cstr_compat_of_poly c,p) sys in
+ assert (check_sys sys) ;
+ xlia enum (if is_linear then reduction_equations else reduction_non_lin_equations) sys
diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml
index cce0a728..e4b58a56 100644
--- a/plugins/micromega/coq_micromega.ml
+++ b/plugins/micromega/coq_micromega.ml
@@ -18,6 +18,7 @@
open Pp
open Mutils
+open Goptions
(**
* Debug flag
@@ -37,6 +38,53 @@ let time str f x =
flush stdout);
res
+
+(* Limit the proof search *)
+
+let max_depth = max_int
+
+(* Search limit for provers over Q R *)
+let lra_proof_depth = ref max_depth
+
+
+(* Search limit for provers over Z *)
+let lia_enum = ref true
+let lia_proof_depth = ref max_depth
+
+let get_lia_option () =
+ (!lia_enum,!lia_proof_depth)
+
+let get_lra_option () =
+ !lra_proof_depth
+
+
+
+let _ =
+
+ let int_opt l vref =
+ {
+ optsync = true;
+ optdepr = false;
+ optname = List.fold_right (^) l "";
+ optkey = l ;
+ optread = (fun () -> Some !vref);
+ optwrite = (fun x -> vref := (match x with None -> max_depth | Some v -> v))
+ } in
+
+ let lia_enum_opt =
+ {
+ optsync = true;
+ optdepr = false;
+ optname = "Lia Enum";
+ optkey = ["Lia";"Enum"];
+ optread = (fun () -> !lia_enum);
+ optwrite = (fun x -> lia_enum := x)
+ } in
+ let _ = declare_int_option (int_opt ["Lra"; "Depth"] lra_proof_depth) in
+ let _ = declare_int_option (int_opt ["Lia"; "Depth"] lia_proof_depth) in
+ let _ = declare_bool_option lia_enum_opt in
+ ()
+
(**
* Initialize a tag type to the Tag module declaration (see Mutils).
*)
@@ -99,6 +147,17 @@ let rec map_atoms fct f =
| N f -> N(map_atoms fct f)
| I(f1,o,f2) -> I(map_atoms fct f1, o , map_atoms fct f2)
+let rec map_prop fct f =
+ match f with
+ | TT -> TT
+ | FF -> FF
+ | X x -> X (fct x)
+ | A (at,tg,cstr) -> A(at,tg,cstr)
+ | C (f1,f2) -> C(map_prop fct f1, map_prop fct f2)
+ | D (f1,f2) -> D(map_prop fct f1, map_prop fct f2)
+ | N f -> N(map_prop fct f)
+ | I(f1,o,f2) -> I(map_prop fct f1, o , map_prop fct f2)
+
(**
* Collect the identifiers of a (string of) implications. Implication labels
* are inherited from Coq/CoC's higher order dependent type constructor (Pi).
@@ -244,7 +303,8 @@ let rec add_term t0 = function
*)
module ISet = Set.Make(Int)
-
+module IMap = Map.Make(Int)
+
(**
* Given a set of integers s=\{i0,...,iN\} and a list m, return the list of
* elements of m that are at position i0,...,iN.
@@ -303,7 +363,8 @@ struct
let r_modules =
[["Coq";"Reals" ; "Rdefinitions"];
["Coq";"Reals" ; "Rpow_def"] ;
-]
+ ["Coq";"Reals" ; "Raxioms"] ;
+ ]
let z_modules = [["Coq";"ZArith";"BinInt"]]
@@ -322,6 +383,8 @@ struct
let coq_and = lazy (init_constant "and")
let coq_or = lazy (init_constant "or")
let coq_not = lazy (init_constant "not")
+ let coq_not_gl_ref = (Nametab.locate ( Libnames.qualid_of_string "Coq.Init.Logic.not"))
+
let coq_iff = lazy (init_constant "iff")
let coq_True = lazy (init_constant "True")
let coq_False = lazy (init_constant "False")
@@ -359,6 +422,7 @@ struct
let coq_Qmake = lazy (constant "Qmake")
let coq_Rcst = lazy (constant "Rcst")
+
let coq_C0 = lazy (m_constant "C0")
let coq_C1 = lazy (m_constant "C1")
let coq_CQ = lazy (m_constant "CQ")
@@ -415,8 +479,9 @@ struct
let coq_Rdiv = lazy (r_constant "Rdiv")
let coq_Rinv = lazy (r_constant "Rinv")
let coq_Rpower = lazy (r_constant "pow")
+ let coq_IZR = lazy (r_constant "IZR")
let coq_IQR = lazy (constant "IQR")
- let coq_IZR = lazy (constant "IZR")
+
let coq_PEX = lazy (constant "PEX" )
let coq_PEc = lazy (constant"PEc")
@@ -896,8 +961,20 @@ struct
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)
+ (env, CamlToCoq.positive n)
+
+ let get_rank env v =
+ let rec _get_rank env n =
+ match env with
+ | [] -> raise (Invalid_argument "get_rank")
+ | e::l ->
+ if eq_constr e v
+ then n
+ else _get_rank l (n+1) in
+ _get_rank env 1
+
+
let empty = []
let elements env = env
@@ -910,7 +987,7 @@ struct
let parse_expr parse_constant parse_exp ops_spec env term =
if debug
- then Pp.msg_debug (Pp.str "parse_expr: " ++ Printer.prterm term);
+ then Feedback.msg_debug (Pp.str "parse_expr: " ++ Printer.prterm term);
(*
let constant_or_variable env term =
@@ -947,7 +1024,7 @@ struct
let (expr,env) = parse_expr env args.(0) in
let power = (parse_exp expr args.(1)) in
(power , env)
- with e when Errors.noncritical e ->
+ with e when CErrors.noncritical e ->
(* if the exponent is a variable *)
let (env,n) = Env.compute_rank_add env term in (Mc.PEX n, env)
end
@@ -994,7 +1071,7 @@ struct
coq_Rplus , (fun x y -> Mc.CPlus(x,y)) ;
coq_Rminus , (fun x y -> Mc.CMinus(x,y)) ;
coq_Rmult , (fun x y -> Mc.CMult(x,y)) ;
- coq_Rdiv , (fun x y -> Mc.CMult(x,Mc.CInv y)) ;
+ (* coq_Rdiv , (fun x y -> Mc.CMult(x,Mc.CInv y)) ;*)
]
let rec rconstant term =
@@ -1016,10 +1093,14 @@ struct
with
ParseError ->
match op with
- | op when Constr.equal op (Lazy.force coq_Rinv) -> Mc.CInv(rconstant args.(0))
- | op when Constr.equal op (Lazy.force coq_IQR) -> Mc.CQ (parse_q args.(0))
-(* | op when op = Lazy.force coq_IZR -> Mc.CZ (parse_z args.(0))*)
- | _ -> raise ParseError
+ | op when Constr.equal op (Lazy.force coq_Rinv) ->
+ let arg = rconstant args.(0) in
+ if Mc.qeq_bool (Mc.q_of_Rcst arg) {Mc.qnum = Mc.Z0 ; Mc.qden = Mc.XH}
+ then raise ParseError (* This is a division by zero -- no semantics *)
+ else Mc.CInv(arg)
+ | op when Constr.equal op (Lazy.force coq_IQR) -> Mc.CQ (parse_q args.(0))
+ | op when Constr.equal op (Lazy.force coq_IZR) -> Mc.CZ (parse_z args.(0))
+ | _ -> raise ParseError
end
| _ -> raise ParseError
@@ -1027,7 +1108,7 @@ struct
let rconstant term =
if debug
- then Pp.msg_debug (Pp.str "rconstant: " ++ Printer.prterm term ++ fnl ());
+ then Feedback.msg_debug (Pp.str "rconstant: " ++ Printer.prterm term ++ fnl ());
let res = rconstant term in
if debug then
(Printf.printf "rconstant -> %a\n" pp_Rcst res ; flush stdout) ;
@@ -1067,7 +1148,7 @@ struct
let parse_arith parse_op parse_expr env cstr gl =
if debug
- then Pp.msg_debug (Pp.str "parse_arith: " ++ Printer.prterm cstr ++ fnl ());
+ then Feedback.msg_debug (Pp.str "parse_arith: " ++ Printer.prterm cstr ++ fnl ());
match kind_of_term cstr with
| App(op,args) ->
let (op,lhs,rhs) = parse_op gl (op,args) in
@@ -1094,10 +1175,6 @@ struct
| 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)
let mkD f1 f2 = D(f1,f2)
@@ -1119,34 +1196,40 @@ struct
try
let (at,env) = parse_atom env t gl in
(A(at,tg,t), env,Tag.next tg)
- with e when Errors.noncritical e -> (X(t),env,tg) in
+ with e when CErrors.noncritical e -> (X(t),env,tg) in
+ let is_prop term =
+ let sort = Retyping.get_sort_of (Tacmach.pf_env gl) (Tacmach.project gl) term in
+ Term.is_prop_sort sort in
+
let rec xparse_formula env tg term =
- match kind_of_term term with
+ match kind_of_term term with
| App(l,rst) ->
- (match rst with
- | [|a;b|] when eq_constr l (Lazy.force coq_and) ->
- let f,env,tg = xparse_formula env tg a in
- let g,env, tg = xparse_formula env tg b in
- mkformula_binary mkC term f g,env,tg
- | [|a;b|] when eq_constr l (Lazy.force coq_or) ->
- let f,env,tg = xparse_formula env tg a in
- let g,env,tg = xparse_formula env tg b in
- mkformula_binary mkD term f g,env,tg
- | [|a|] when eq_constr l (Lazy.force coq_not) ->
- let (f,env,tg) = xparse_formula env tg a in (N(f), env,tg)
- | [|a;b|] when eq_constr l (Lazy.force coq_iff) ->
- let f,env,tg = xparse_formula env tg a in
- let g,env,tg = xparse_formula env tg b in
- mkformula_binary mkIff term f g,env,tg
- | _ -> parse_atom env tg term)
- | Prod(typ,a,b) when not (Termops.dependent (mkRel 1) b) ->
+ (match rst with
+ | [|a;b|] when eq_constr l (Lazy.force coq_and) ->
+ let f,env,tg = xparse_formula env tg a in
+ let g,env, tg = xparse_formula env tg b in
+ mkformula_binary mkC term f g,env,tg
+ | [|a;b|] when eq_constr l (Lazy.force coq_or) ->
+ let f,env,tg = xparse_formula env tg a in
+ let g,env,tg = xparse_formula env tg b in
+ mkformula_binary mkD term f g,env,tg
+ | [|a|] when eq_constr l (Lazy.force coq_not) ->
+ let (f,env,tg) = xparse_formula env tg a in (N(f), env,tg)
+ | [|a;b|] when eq_constr l (Lazy.force coq_iff) ->
let f,env,tg = xparse_formula env tg a in
let g,env,tg = xparse_formula env tg b in
- mkformula_binary mkI term f g,env,tg
+ mkformula_binary mkIff term f g,env,tg
+ | _ -> parse_atom env tg term)
+ | Prod(typ,a,b) when not (Termops.dependent (mkRel 1) b)->
+ let f,env,tg = xparse_formula env tg a in
+ let g,env,tg = xparse_formula env tg b in
+ mkformula_binary mkI term f g,env,tg
| _ when eq_constr term (Lazy.force coq_True) -> (TT,env,tg)
| _ when eq_constr term (Lazy.force coq_False) -> (FF,env,tg)
- | _ -> X(term),env,tg in
+ | _ when is_prop term -> X(term),env,tg
+ | _ -> raise ParseError
+ in
xparse_formula env tg ((*Reductionops.whd_zeta*) term)
let dump_formula typ dump_atom f =
@@ -1162,12 +1245,214 @@ struct
| X(t) -> mkApp(Lazy.force coq_X,[|typ ; t|]) in
xdump f
- (**
- * Given a conclusion and a list of affectations, rebuild a term prefixed by
- * the appropriate letins.
- * TODO: reverse the list of bindings!
- *)
+ let prop_env_of_formula form =
+ let rec doit env = function
+ | TT | FF | A(_,_,_) -> env
+ | X t -> fst (Env.compute_rank_add env t)
+ | C(f1,f2) | D(f1,f2) | I(f1,_,f2) ->
+ doit (doit env f1) f2
+ | N f -> doit env f in
+
+ doit [] form
+
+ let var_env_of_formula form =
+
+ let rec vars_of_expr = function
+ | Mc.PEX n -> ISet.singleton (CoqToCaml.positive n)
+ | Mc.PEc z -> ISet.empty
+ | Mc.PEadd(e1,e2) | Mc.PEmul(e1,e2) | Mc.PEsub(e1,e2) ->
+ ISet.union (vars_of_expr e1) (vars_of_expr e2)
+ | Mc.PEopp e | Mc.PEpow(e,_)-> vars_of_expr e
+ in
+
+ let vars_of_atom {Mc.flhs ; Mc.fop; Mc.frhs} =
+ ISet.union (vars_of_expr flhs) (vars_of_expr frhs) in
+
+ let rec doit = function
+ | TT | FF | X _ -> ISet.empty
+ | A (a,t,c) -> vars_of_atom a
+ | C(f1,f2) | D(f1,f2) |I (f1,_,f2) -> ISet.union (doit f1) (doit f2)
+ | N f -> doit f in
+
+ doit form
+
+
+
+
+ type 'cst dump_expr = (* 'cst is the type of the syntactic constants *)
+ {
+ interp_typ : constr;
+ dump_cst : 'cst -> constr;
+ dump_add : constr;
+ dump_sub : constr;
+ dump_opp : constr;
+ dump_mul : constr;
+ dump_pow : constr;
+ dump_pow_arg : Mc.n -> constr;
+ dump_op : (Mc.op2 * Term.constr) list
+ }
+
+let dump_zexpr = lazy
+ {
+ interp_typ = Lazy.force coq_Z;
+ dump_cst = dump_z;
+ dump_add = Lazy.force coq_Zplus;
+ dump_sub = Lazy.force coq_Zminus;
+ dump_opp = Lazy.force coq_Zopp;
+ dump_mul = Lazy.force coq_Zmult;
+ dump_pow = Lazy.force coq_Zpower;
+ dump_pow_arg = (fun n -> dump_z (CamlToCoq.z (CoqToCaml.n n)));
+ dump_op = List.map (fun (x,y) -> (y,Lazy.force x)) zop_table
+ }
+
+let dump_qexpr = lazy
+ {
+ interp_typ = Lazy.force coq_Q;
+ dump_cst = dump_q;
+ dump_add = Lazy.force coq_Qplus;
+ dump_sub = Lazy.force coq_Qminus;
+ dump_opp = Lazy.force coq_Qopp;
+ dump_mul = Lazy.force coq_Qmult;
+ dump_pow = Lazy.force coq_Qpower;
+ dump_pow_arg = (fun n -> dump_z (CamlToCoq.z (CoqToCaml.n n)));
+ dump_op = List.map (fun (x,y) -> (y,Lazy.force x)) qop_table
+ }
+
+ let dump_positive_as_R p =
+ let mult = Lazy.force coq_Rmult in
+ let add = Lazy.force coq_Rplus in
+
+ let one = Lazy.force coq_R1 in
+ let mk_add x y = mkApp(add,[|x;y|]) in
+ let mk_mult x y = mkApp(mult,[|x;y|]) in
+
+ let two = mk_add one one in
+
+ let rec dump_positive p =
+ match p with
+ | Mc.XH -> one
+ | Mc.XO p -> mk_mult two (dump_positive p)
+ | Mc.XI p -> mk_add one (mk_mult two (dump_positive p)) in
+
+ dump_positive p
+
+let dump_n_as_R n =
+ let z = CoqToCaml.n n in
+ if z = 0
+ then Lazy.force coq_R0
+ else dump_positive_as_R (CamlToCoq.positive z)
+
+
+let rec dump_Rcst_as_R cst =
+ match cst with
+ | Mc.C0 -> Lazy.force coq_R0
+ | Mc.C1 -> Lazy.force coq_R1
+ | Mc.CQ q -> Term.mkApp(Lazy.force coq_IQR, [| dump_q q |])
+ | Mc.CZ z -> Term.mkApp(Lazy.force coq_IZR, [| dump_z z |])
+ | Mc.CPlus(x,y) -> Term.mkApp(Lazy.force coq_Rplus, [| dump_Rcst_as_R x ; dump_Rcst_as_R y |])
+ | Mc.CMinus(x,y) -> Term.mkApp(Lazy.force coq_Rminus, [| dump_Rcst_as_R x ; dump_Rcst_as_R y |])
+ | Mc.CMult(x,y) -> Term.mkApp(Lazy.force coq_Rmult, [| dump_Rcst_as_R x ; dump_Rcst_as_R y |])
+ | Mc.CInv t -> Term.mkApp(Lazy.force coq_Rinv, [| dump_Rcst_as_R t |])
+ | Mc.COpp t -> Term.mkApp(Lazy.force coq_Ropp, [| dump_Rcst_as_R t |])
+
+
+let dump_rexpr = lazy
+ {
+ interp_typ = Lazy.force coq_R;
+ dump_cst = dump_Rcst_as_R;
+ dump_add = Lazy.force coq_Rplus;
+ dump_sub = Lazy.force coq_Rminus;
+ dump_opp = Lazy.force coq_Ropp;
+ dump_mul = Lazy.force coq_Rmult;
+ dump_pow = Lazy.force coq_Rpower;
+ dump_pow_arg = (fun n -> dump_nat (CamlToCoq.nat (CoqToCaml.n n)));
+ dump_op = List.map (fun (x,y) -> (y,Lazy.force x)) rop_table
+ }
+
+
+
+
+(** [make_goal_of_formula depxr vars props form] where
+ - vars is an environment for the arithmetic variables occuring in form
+ - props is an environment for the propositions occuring in form
+ @return a goal where all the variables and propositions of the formula are quantified
+
+*)
+
+let rec make_goal_of_formula dexpr form =
+
+ let vars_idx =
+ List.mapi (fun i v -> (v, i+1)) (ISet.elements (var_env_of_formula form)) in
+
+ (* List.iter (fun (v,i) -> Printf.fprintf stdout "var %i has index %i\n" v i) vars_idx ;*)
+
+ let props = prop_env_of_formula form in
+
+ let vars_n = List.map (fun (_,i) -> (Names.id_of_string (Printf.sprintf "__x%i" i)) , dexpr.interp_typ) vars_idx in
+ let props_n = List.mapi (fun i _ -> (Names.id_of_string (Printf.sprintf "__p%i" (i+1))) , Term.mkProp) props in
+
+ let var_name_pos = List.map2 (fun (idx,_) (id,_) -> id,idx) vars_idx vars_n in
+
+ let dump_expr i e =
+ let rec dump_expr = function
+ | Mc.PEX n -> mkRel (i+(List.assoc (CoqToCaml.positive n) vars_idx))
+ | Mc.PEc z -> dexpr.dump_cst z
+ | Mc.PEadd(e1,e2) -> mkApp(dexpr.dump_add,
+ [| dump_expr e1;dump_expr e2|])
+ | Mc.PEsub(e1,e2) -> mkApp(dexpr.dump_sub,
+ [| dump_expr e1;dump_expr e2|])
+ | Mc.PEopp e -> mkApp(dexpr.dump_opp,
+ [| dump_expr e|])
+ | Mc.PEmul(e1,e2) -> mkApp(dexpr.dump_mul,
+ [| dump_expr e1;dump_expr e2|])
+ | Mc.PEpow(e,n) -> mkApp(dexpr.dump_pow,
+ [| dump_expr e; dexpr.dump_pow_arg n|])
+ in dump_expr e in
+
+ let mkop op e1 e2 =
+ try
+ Term.mkApp(List.assoc op dexpr.dump_op, [| e1; e2|])
+ with Not_found ->
+ Term.mkApp(Lazy.force coq_Eq,[|dexpr.interp_typ ; e1 ;e2|]) in
+
+ let dump_cstr i { Mc.flhs ; Mc.fop ; Mc.frhs } =
+ mkop fop (dump_expr i flhs) (dump_expr i frhs) in
+
+ let rec xdump pi xi f =
+ match f with
+ | TT -> Lazy.force coq_True
+ | FF -> Lazy.force coq_False
+ | C(x,y) -> mkApp(Lazy.force coq_and,[|xdump pi xi x ; xdump pi xi y|])
+ | D(x,y) -> mkApp(Lazy.force coq_or,[| xdump pi xi x ; xdump pi xi y|])
+ | I(x,_,y) -> mkArrow (xdump pi xi x) (xdump (pi+1) (xi+1) y)
+ | N(x) -> mkArrow (xdump pi xi x) (Lazy.force coq_False)
+ | A(x,_,_) -> dump_cstr xi x
+ | X(t) -> let idx = Env.get_rank props t in
+ mkRel (pi+idx) in
+
+ let nb_vars = List.length vars_n in
+ let nb_props = List.length props_n in
+
+ (* Printf.fprintf stdout "NBProps : %i\n" nb_props ;*)
+
+ let subst_prop p =
+ let idx = Env.get_rank props p in
+ mkVar (Names.id_of_string (Printf.sprintf "__p%i" idx)) in
+
+ let form' = map_prop subst_prop form in
+
+ (Term.prodn nb_props (List.map (fun (x,y) -> Names.Name x,y) props_n)
+ (Term.prodn nb_vars (List.map (fun (x,y) -> Names.Name x,y) vars_n)
+ (xdump (List.length vars_n) 0 form)),
+ List.rev props_n, List.rev var_name_pos,form')
+
+ (**
+ * Given a conclusion and a list of affectations, rebuild a term prefixed by
+ * the appropriate letins.
+ * TODO: reverse the list of bindings!
+ *)
+
let set l concl =
let rec xset acc = function
| [] -> acc
@@ -1242,29 +1527,25 @@ 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 when Errors.noncritical 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 coq_VarMap = lazy
+ (Coqlib.gen_constant_in_modules "VarMap"
+ [["Coq" ; "micromega" ; "VarMap"] ; ["VarMap"]] "t")
+
+
+let rec dump_varmap typ m =
+ match m with
+ | Mc.Empty -> Term.mkApp(Lazy.force coq_Empty,[| typ |])
+ | Mc.Leaf v -> Term.mkApp(Lazy.force coq_Leaf,[| typ; v|])
+ | Mc.Node(l,o,r) ->
+ Term.mkApp (Lazy.force coq_Node, [| typ; dump_varmap typ l; o ; dump_varmap typ r |])
+
+
+let vm_of_list env =
+ match env with
+ | [] -> Mc.Empty
+ | (d,_)::_ ->
+ List.fold_left (fun vm (c,i) ->
+ Mc.vm_add d (CamlToCoq.positive i) c vm) Mc.Empty env
let rec pp_varmap o vm =
@@ -1329,7 +1610,7 @@ let rec parse_hyps gl parse_arith env tg hyps =
try
let (c,env,tg) = parse_formula gl parse_arith env tg t in
((i,c)::lhyps, env,tg)
- with e when Errors.noncritical e -> (lhyps,env,tg)
+ with e when CErrors.noncritical e -> (lhyps,env,tg)
(*(if debug then Printf.printf "parse_arith : %s\n" x);*)
@@ -1377,50 +1658,81 @@ let rcst_domain_spec = lazy {
dump_proof = dump_psatz coq_Q dump_q
}
+open Proofview.Notations
+
+(** Naive topological sort of constr according to the subterm-ordering *)
+
+(* An element is minimal x is minimal w.r.t y if
+ x <= y or (x and y are incomparable) *)
+
+let is_min le x y =
+ if le x y then true
+ else if le y x then false else true
+
+let is_minimal le l c = List.for_all (is_min le c) l
+
+let find_rem p l =
+ let rec xfind_rem acc l =
+ match l with
+ | [] -> (None, acc)
+ | x :: l -> if p x then (Some x, acc @ l)
+ else xfind_rem (x::acc) l in
+ xfind_rem [] l
+
+let find_minimal le l = find_rem (is_minimal le l) l
+
+let rec mk_topo_order le l =
+ match find_minimal le l with
+ | (None , _) -> []
+ | (Some v,l') -> v :: (mk_topo_order le l')
+
+
+let topo_sort_constr l = mk_topo_order Termops.dependent l
+
+
(**
* Instanciate the current Coq goal with a Micromega formula, a varmap, and a
* witness.
*)
-
-
-let micromega_order_change spec cert cert_typ env ff : Tacmach.tactic =
- let ids = Util.List.map_i (fun i _ -> (Names.Id.of_string ("__z"^(string_of_int i)))) 0 env in
+let micromega_order_change spec cert cert_typ env ff (*: unit Proofview.tactic*) =
+ (* let ids = Util.List.map_i (fun i _ -> (Names.Id.of_string ("__v"^(string_of_int i)))) 0 env in *)
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
- (* todo : directly generate the proof term - or generalize befor conversion? *)
- Tacticals.tclTHENSEQ [
- (fun gl ->
- Proofview.V82.of_tactic (Tactics.change_concl
- (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);
- Tactics.generalize env ;
- Tacticals.tclTHENSEQ (List.map (fun id -> Proofview.V82.of_tactic (Tactics.introduction id)) ids) ;
- ]
-
+ let vm = dump_varmap (spec.typ) (vm_of_list env) in
+ (* todo : directly generate the proof term - or generalize before conversion? *)
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
+ let gl = Tacmach.New.of_old (fun x -> x) gl in
+ Tacticals.New.tclTHENLIST
+ [
+ Tactics.change_concl
+ (set
+ [
+ ("__ff", ff, Term.mkApp(Lazy.force coq_Formula, [|formula_typ |]));
+ ("__varmap", vm, Term.mkApp(Lazy.force coq_VarMap, [|spec.typ|]));
+ ("__wit", cert, cert_typ)
+ ]
+ (Tacmach.pf_concl gl))
+ ]
+ end }
(**
* The datastructures that aggregate prover attributes.
*)
-type ('a,'prf) prover = {
+type ('option,'a,'prf) prover = {
name : string ; (* name of the prover *)
- prover : 'a list -> 'prf option ; (* the prover itself *)
+ get_option : unit ->'option ; (* find the options of the prover *)
+ prover : 'option * 'a list -> 'prf option ; (* the prover itself *)
hyps : 'prf -> ISet.t ; (* extract the indexes of the hypotheses really used in the proof *)
compact : 'prf -> (int -> int) -> 'prf ; (* remap the hyp indexes according to function *)
pp_prf : out_channel -> 'prf -> unit ;(* pretting printing of proof *)
pp_f : out_channel -> 'a -> unit (* pretty printing of the formulas (polynomials)*)
}
+
+
(**
* Given a list of provers and a disjunction of atoms, find a proof of any of
* the atoms. Returns an (optional) pair of a proof and a prover
@@ -1430,7 +1742,7 @@ type ('a,'prf) prover = {
let find_witness provers polys1 =
let provers = List.map (fun p ->
(fun l ->
- match p.prover l with
+ match p.prover (p.get_option (),l) with
| None -> None
| Some prf -> Some(prf,p)) , p.name) provers in
try_any provers (List.map fst polys1)
@@ -1482,10 +1794,10 @@ let compact_proofs (cnf_ff: 'cst cnf) res (cnf_ff': 'cst cnf) =
(pp_ml_list prover.pp_f) (List.map fst new_cl) ;
flush stdout
end ; *)
- let res = try prover.compact prf remap with x when Errors.noncritical x ->
+ let res = try prover.compact prf remap with x when CErrors.noncritical x ->
if debug then Printf.fprintf stdout "Proof compaction %s" (Printexc.to_string x) ;
(* This should not happen -- this is the recovery plan... *)
- match prover.prover (List.map fst new_cl) with
+ match prover.prover (prover.get_option () ,List.map fst new_cl) with
| None -> failwith "proof compaction error"
| Some p -> p
in
@@ -1562,6 +1874,7 @@ let rec abstract_wrt_formula f1 f2 =
exception CsdpNotFound
+
(**
* This is the core of Micromega: apply the prover, analyze the result and
* prune unused fomulas, and finally modify the proof state.
@@ -1586,12 +1899,12 @@ let micromega_tauto negate normalise unsat deduce spec prover env polys1 polys2
if debug then
begin
- Pp.pp (Pp.str "Formula....\n") ;
+ Feedback.msg_notice (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 ();
- Printf.fprintf stdout "cnf : %a\n" (pp_cnf (fun o _ -> ())) cnf_ff
+ Feedback.msg_notice (Printer.prterm ff);
+ Printf.fprintf stdout "cnf : %a\n" (pp_cnf (fun o _ -> ())) cnf_ff
end;
match witness_list_tags prover cnf_ff with
@@ -1611,11 +1924,11 @@ let micromega_tauto negate normalise unsat deduce spec prover env polys1 polys2
if debug then
begin
- Pp.pp (Pp.str "\nAFormula\n") ;
+ Feedback.msg_notice (Pp.str "\nAFormula\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 ();
+ Feedback.msg_notice (Printer.prterm ff');
Printf.fprintf stdout "cnf : %a\n" (pp_cnf (fun o _ -> ())) cnf_ff'
end;
@@ -1646,58 +1959,103 @@ let micromega_gen
(negate:'cst atom -> 'cst mc_cnf)
(normalise:'cst atom -> 'cst mc_cnf)
unsat deduce
- 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 gl parse_arith Env.empty hyps concl in
- let env = Env.elements env in
- let spec = Lazy.force spec in
-
+ spec dumpexpr prover tac =
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
+ let gl = Tacmach.New.of_old (fun x -> x) gl in
+ let concl = Tacmach.pf_concl gl in
+ let hyps = Tacmach.pf_hyps_types gl in
+ try
+ let (hyps,concl,env) = parse_goal gl parse_arith Env.empty hyps concl in
+ let env = Env.elements env in
+ let spec = Lazy.force spec in
+ let dumpexpr = Lazy.force dumpexpr in
+
match micromega_tauto negate normalise unsat deduce spec prover env hyps concl gl with
- | None -> Tacticals.tclFAIL 0 (Pp.str " Cannot find witness") gl
- | Some (ids,ff',res') ->
- (Tacticals.tclTHENSEQ
- [
- Tactics.generalize (List.map Term.mkVar ids) ;
- micromega_order_change spec res'
- (Term.mkApp(Lazy.force coq_list, [|spec.proof_typ|])) env ff'
- ]) gl
- with
- | ParseError -> Tacticals.tclFAIL 0 (Pp.str "Bad logical fragment") gl
- | CsdpNotFound -> flush stdout ; Pp.pp_flush () ;
- Tacticals.tclFAIL 0 (Pp.str
- (" Skipping what remains of this tactic: the complexity of the goal requires "
- ^ "the use of a specialized external tool called csdp. \n\n"
- ^ "Unfortunately Coq isn't aware of the presence of any \"csdp\" executable in the path. \n\n"
- ^ "Csdp packages are provided by some OS distributions; binaries and source code can be downloaded from https://projects.coin-or.org/Csdp")) gl
-
-
-
-let micromega_order_changer cert env ff gl =
+ | None -> Tacticals.New.tclFAIL 0 (Pp.str " Cannot find witness")
+ | Some (ids,ff',res') ->
+ let (arith_goal,props,vars,ff_arith) = make_goal_of_formula dumpexpr ff' in
+ let intro (id,_) = Tactics.introduction id in
+
+ let intro_vars = Tacticals.New.tclTHENLIST (List.map intro vars) in
+ let intro_props = Tacticals.New.tclTHENLIST (List.map intro props) in
+ let ipat_of_name id = Some (Loc.ghost, Misctypes.IntroNaming (Misctypes.IntroIdentifier id)) in
+ let goal_name = Tactics.fresh_id [] (Names.Id.of_string "__arith") gl in
+ let env' = List.map (fun (id,i) -> Term.mkVar id,i) vars in
+
+ let tac_arith = Tacticals.New.tclTHENLIST [ intro_props ; intro_vars ;
+ micromega_order_change spec res'
+ (Term.mkApp(Lazy.force coq_list, [|spec.proof_typ|])) env' ff_arith ] in
+
+ let goal_props = List.rev (prop_env_of_formula ff') in
+
+ let goal_vars = List.map (fun (_,i) -> List.nth env (i-1)) vars in
+
+ let arith_args = goal_props @ goal_vars in
+
+ let kill_arith =
+ Tacticals.New.tclTHEN
+ (Tactics.keep [])
+ ((*Tactics.tclABSTRACT None*)
+ (Tacticals.New.tclTHEN tac_arith tac)) in
+
+ Tacticals.New.tclTHENS
+ (Tactics.forward true (Some None) (ipat_of_name goal_name) arith_goal)
+ [
+ kill_arith;
+ (Tacticals.New.tclTHENLIST
+ [(Tactics.generalize (List.map Term.mkVar ids));
+ Tactics.exact_check (Term.applist (Term.mkVar goal_name, arith_args))
+ ] )
+ ]
+ with
+ | ParseError -> Tacticals.New.tclFAIL 0 (Pp.str "Bad logical fragment")
+ | Mfourier.TimeOut -> Tacticals.New.tclFAIL 0 (Pp.str "Timeout")
+ | CsdpNotFound -> flush stdout ;
+ Tacticals.New.tclFAIL 0 (Pp.str
+ (" Skipping what remains of this tactic: the complexity of the goal requires "
+ ^ "the use of a specialized external tool called csdp. \n\n"
+ ^ "Unfortunately Coq isn't aware of the presence of any \"csdp\" executable in the path. \n\n"
+ ^ "Csdp packages are provided by some OS distributions; binaries and source code can be downloaded from https://projects.coin-or.org/Csdp"))
+ end }
+
+let micromega_gen parse_arith
+ (negate:'cst atom -> 'cst mc_cnf)
+ (normalise:'cst atom -> 'cst mc_cnf)
+ unsat deduce
+ spec prover =
+ (micromega_gen parse_arith negate normalise unsat deduce spec prover)
+
+
+
+let micromega_order_changer cert env ff =
+ (*let ids = Util.List.map_i (fun i _ -> (Names.Id.of_string ("__v"^(string_of_int i)))) 0 env in *)
let coeff = Lazy.force coq_Rcst in
let dump_coeff = dump_Rcst in
let typ = Lazy.force coq_R in
let cert_typ = (Term.mkApp(Lazy.force coq_list, [|Lazy.force coq_QWitness |])) in
-
- let formula_typ = (Term.mkApp (Lazy.force coq_Cstr,[| coeff|])) in
- let ff = dump_formula formula_typ (dump_cstr coeff dump_coeff) ff in
- let vm = dump_varmap (typ) env in
- Proofview.V82.of_tactic (Tactics.change_concl
- (set
+
+ let formula_typ = (Term.mkApp (Lazy.force coq_Cstr,[| coeff|])) in
+ let ff = dump_formula formula_typ (dump_cstr coeff dump_coeff) ff in
+ let vm = dump_varmap (typ) (vm_of_list env) in
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
+ let gl = Tacmach.New.of_old (fun x -> x) gl in
+ Tacticals.New.tclTHENLIST
[
- ("__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", [|typ|]));
- ("__wit", cert, cert_typ)
+ (Tactics.change_concl
+ (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", [|typ|]));
+ ("__wit", cert, cert_typ)
+ ]
+ (Tacmach.pf_concl gl)));
+ (* Tacticals.New.tclTHENLIST (List.map (fun id -> (Tactics.introduction id)) ids)*)
]
- (Tacmach.pf_concl gl)
- ))
- gl
-
+ end }
-let micromega_genr prover gl =
+let micromega_genr prover tac =
let parse_arith = parse_rarith in
let negate = Mc.rnegate in
let normalise = Mc.rnormalise in
@@ -1710,41 +2068,77 @@ let micromega_genr prover gl =
proof_typ = Lazy.force coq_QWitness ;
dump_proof = dump_psatz coq_Q dump_q
} in
-
- let concl = Tacmach.pf_concl gl in
- let hyps = Tacmach.pf_hyps_types gl in
- try
- let (hyps,concl,env) = parse_goal gl parse_arith Env.empty hyps concl in
- let env = Env.elements env in
- let spec = Lazy.force spec in
-
- let hyps' = List.map (fun (n,f) -> (n, map_atoms (Micromega.map_Formula Micromega.q_of_Rcst) f)) hyps in
- let concl' = map_atoms (Micromega.map_Formula Micromega.q_of_Rcst) concl in
-
- match micromega_tauto negate normalise unsat deduce spec prover env hyps' concl' gl with
- | None -> Tacticals.tclFAIL 0 (Pp.str " Cannot find witness") gl
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
+ let gl = Tacmach.New.of_old (fun x -> x) gl in
+ let concl = Tacmach.pf_concl gl in
+ let hyps = Tacmach.pf_hyps_types gl in
+
+ try
+ let (hyps,concl,env) = parse_goal gl parse_arith Env.empty hyps concl in
+ let env = Env.elements env in
+ let spec = Lazy.force spec in
+
+ let hyps' = List.map (fun (n,f) -> (n, map_atoms (Micromega.map_Formula Micromega.q_of_Rcst) f)) hyps in
+ let concl' = map_atoms (Micromega.map_Formula Micromega.q_of_Rcst) concl in
+
+ match micromega_tauto negate normalise unsat deduce spec prover env hyps' concl' gl with
+ | None -> Tacticals.New.tclFAIL 0 (Pp.str " Cannot find witness")
| Some (ids,ff',res') ->
- let (ff,ids') = formula_hyps_concl
+ let (ff,ids) = formula_hyps_concl
(List.filter (fun (n,_) -> List.mem n ids) hyps) concl in
+ let ff' = abstract_wrt_formula ff' ff in
+
+ let (arith_goal,props,vars,ff_arith) = make_goal_of_formula (Lazy.force dump_rexpr) ff' in
+ let intro (id,_) = Tactics.introduction id in
+
+ let intro_vars = Tacticals.New.tclTHENLIST (List.map intro vars) in
+ let intro_props = Tacticals.New.tclTHENLIST (List.map intro props) in
+ let ipat_of_name id = Some (Loc.ghost, Misctypes.IntroNaming (Misctypes.IntroIdentifier id)) in
+ let goal_name = Tactics.fresh_id [] (Names.Id.of_string "__arith") gl in
+ let env' = List.map (fun (id,i) -> Term.mkVar id,i) vars in
+
+ let tac_arith = Tacticals.New.tclTHENLIST [ intro_props ; intro_vars ;
+ micromega_order_changer res' env' ff_arith ] in
+
+ let goal_props = List.rev (prop_env_of_formula ff') in
+
+ let goal_vars = List.map (fun (_,i) -> List.nth env (i-1)) vars in
+
+ let arith_args = goal_props @ goal_vars in
+
+ let kill_arith =
+ Tacticals.New.tclTHEN
+ (Tactics.keep [])
+ ((*Tactics.tclABSTRACT None*)
+ (Tacticals.New.tclTHEN tac_arith tac)) in
- (Tacticals.tclTHENSEQ
- [
- Tactics.generalize (List.map Term.mkVar ids) ;
- micromega_order_changer res' env (abstract_wrt_formula ff' ff)
- ]) gl
- with
- | ParseError -> Tacticals.tclFAIL 0 (Pp.str "Bad logical fragment") gl
- | CsdpNotFound -> flush stdout ; Pp.pp_flush () ;
- Tacticals.tclFAIL 0 (Pp.str
- (" Skipping what remains of this tactic: the complexity of the goal requires "
- ^ "the use of a specialized external tool called csdp. \n\n"
- ^ "Unfortunately Coq isn't aware of the presence of any \"csdp\" executable in the path. \n\n"
- ^ "Csdp packages are provided by some OS distributions; binaries and source code can be downloaded from https://projects.coin-or.org/Csdp")) gl
+ Tacticals.New.tclTHENS
+ (Tactics.forward true (Some None) (ipat_of_name goal_name) arith_goal)
+ [
+ kill_arith;
+ (Tacticals.New.tclTHENLIST
+ [(Tactics.generalize (List.map Term.mkVar ids));
+ Tactics.exact_check (Term.applist (Term.mkVar goal_name, arith_args))
+ ] )
+ ]
+ with
+ | ParseError -> Tacticals.New.tclFAIL 0 (Pp.str "Bad logical fragment")
+ | Mfourier.TimeOut -> Tacticals.New.tclFAIL 0 (Pp.str "Timeout")
+ | CsdpNotFound -> flush stdout ;
+ Tacticals.New.tclFAIL 0 (Pp.str
+ (" Skipping what remains of this tactic: the complexity of the goal requires "
+ ^ "the use of a specialized external tool called csdp. \n\n"
+ ^ "Unfortunately Coq isn't aware of the presence of any \"csdp\" executable in the path. \n\n"
+ ^ "Csdp packages are provided by some OS distributions; binaries and source code can be downloaded from https://projects.coin-or.org/Csdp"))
+ end }
+let micromega_genr prover = (micromega_genr prover)
+
+
let lift_ratproof prover l =
match prover l with
| None -> None
@@ -1766,7 +2160,7 @@ module Cache = PHashtable(struct
let hash = Hashtbl.hash
end)
-let csdp_cache = "csdp.cache"
+let csdp_cache = ".csdp.cache"
(**
* Build the command to call csdpcert, and launch it. This in turn will call
@@ -1818,7 +2212,7 @@ let call_csdpcert_q provername poly =
let cert = Certificate.q_cert_of_pos cert in
if Mc.qWeakChecker poly cert
then Some cert
- else ((print_string "buggy certificate" ; flush stdout) ;None)
+ else ((print_string "buggy certificate") ;None)
let call_csdpcert_z provername poly =
let l = List.map (fun (e,o) -> (z_to_q_pol e,o)) poly in
@@ -1898,38 +2292,61 @@ let compact_pt pt f =
let lift_pexpr_prover p l = p (List.map (fun (e,o) -> Mc.denorm e , o) l)
-let linear_prover_Z = {
- name = "linear prover" ;
- prover = lift_ratproof (lift_pexpr_prover (Certificate.linear_prover_with_cert Certificate.z_spec)) ;
- hyps = hyps_of_pt ;
- compact = compact_pt ;
- pp_prf = pp_proof_term;
- pp_f = fun o x -> pp_pol pp_z o (fst x)
-}
+module CacheZ = PHashtable(struct
+ type prover_option = bool * int
+
+ type t = prover_option * ((Mc.z Mc.pol * Mc.op1) list)
+ let equal = (=)
+ let hash = Hashtbl.hash
+end)
+module CacheQ = PHashtable(struct
+ type t = int * ((Mc.q Mc.pol * Mc.op1) list)
+ let equal = (=)
+ let hash = Hashtbl.hash
+end)
+
+let memo_zlinear_prover = CacheZ.memo ".lia.cache" (fun ((ce,b),s) -> lift_pexpr_prover (Certificate.lia ce b) s)
+let memo_nlia = CacheZ.memo ".nia.cache" (fun ((ce,b),s) -> lift_pexpr_prover (Certificate.nlia ce b) s)
+let memo_nra = CacheQ.memo ".nra.cache" (fun (o,s) -> lift_pexpr_prover (Certificate.nlinear_prover o) s)
+
+
+
let linear_prover_Q = {
- name = "linear prover";
- prover = lift_pexpr_prover (Certificate.linear_prover_with_cert Certificate.q_spec) ;
- hyps = hyps_of_cone ;
- compact = compact_cone ;
- pp_prf = pp_psatz pp_q ;
- pp_f = fun o x -> pp_pol pp_q o (fst x)
+ name = "linear prover";
+ get_option = get_lra_option ;
+ prover = (fun (o,l) -> lift_pexpr_prover (Certificate.linear_prover_with_cert o Certificate.q_spec) l) ;
+ hyps = hyps_of_cone ;
+ compact = compact_cone ;
+ pp_prf = pp_psatz pp_q ;
+ pp_f = fun o x -> pp_pol pp_q o (fst x)
}
let linear_prover_R = {
name = "linear prover";
- prover = lift_pexpr_prover (Certificate.linear_prover_with_cert Certificate.q_spec) ;
+ get_option = get_lra_option ;
+ prover = (fun (o,l) -> lift_pexpr_prover (Certificate.linear_prover_with_cert o Certificate.q_spec) l) ;
hyps = hyps_of_cone ;
compact = compact_cone ;
pp_prf = pp_psatz pp_q ;
pp_f = fun o x -> pp_pol pp_q o (fst x)
}
+let nlinear_prover_R = {
+ name = "nra";
+ get_option = get_lra_option;
+ prover = memo_nra ;
+ hyps = hyps_of_cone ;
+ compact = compact_cone ;
+ pp_prf = pp_psatz pp_q ;
+ pp_f = fun o x -> pp_pol pp_q o (fst x)
+}
let non_linear_prover_Q str o = {
name = "real nonlinear prover";
- prover = call_csdpcert_q (str, o);
+ get_option = (fun () -> (str,o));
+ prover = (fun (o,l) -> call_csdpcert_q o l);
hyps = hyps_of_cone;
compact = compact_cone ;
pp_prf = pp_psatz pp_q ;
@@ -1938,7 +2355,8 @@ let non_linear_prover_Q str o = {
let non_linear_prover_R str o = {
name = "real nonlinear prover";
- prover = call_csdpcert_q (str, o);
+ get_option = (fun () -> (str,o));
+ prover = (fun (o,l) -> call_csdpcert_q o l);
hyps = hyps_of_cone;
compact = compact_cone;
pp_prf = pp_psatz pp_q;
@@ -1947,30 +2365,18 @@ let non_linear_prover_R str o = {
let non_linear_prover_Z str o = {
name = "real nonlinear prover";
- prover = lift_ratproof (call_csdpcert_z (str, o));
+ get_option = (fun () -> (str,o));
+ prover = (fun (o,l) -> lift_ratproof (call_csdpcert_z o) l);
hyps = hyps_of_pt;
compact = compact_pt;
pp_prf = pp_proof_term;
pp_f = fun o x -> pp_pol pp_z o (fst x)
}
-module CacheZ = PHashtable(struct
- type t = (Mc.z Mc.pol * Mc.op1) list
- let equal = Pervasives.(=)
- let hash = Hashtbl.hash
-end)
-
-let memo_zlinear_prover = CacheZ.memo "lia.cache" (lift_pexpr_prover Certificate.lia)
-let memo_nlia = CacheZ.memo "nlia.cache" (lift_pexpr_prover Certificate.nlia)
-
-(*let memo_zlinear_prover = (lift_pexpr_prover Lia.lia)*)
-(*let memo_zlinear_prover = CacheZ.memo "lia.cache" (lift_pexpr_prover Certificate.zlinear_prover)*)
-
-
-
let linear_Z = {
name = "lia";
- prover = memo_zlinear_prover ;
+ get_option = get_lia_option;
+ prover = memo_zlinear_prover ;
hyps = hyps_of_pt;
compact = compact_pt;
pp_prf = pp_proof_term;
@@ -1979,7 +2385,8 @@ let linear_Z = {
let nlinear_Z = {
name = "nlia";
- prover = memo_nlia ;
+ get_option = get_lia_option;
+ prover = memo_nlia ;
hyps = hyps_of_pt;
compact = compact_pt;
pp_prf = pp_proof_term;
@@ -2001,57 +2408,53 @@ let tauto_lia ff =
* solvers
*)
-let psatzl_Z gl =
- micromega_gen parse_zarith Mc.negate Mc.normalise Mc.zunsat Mc.zdeduce zz_domain_spec
- [ linear_prover_Z ] gl
-
-let psatzl_Q gl =
- micromega_gen parse_qarith Mc.qnegate Mc.qnormalise Mc.qunsat Mc.qdeduce qq_domain_spec
- [ linear_prover_Q ] gl
-
-let psatz_Q i gl =
- micromega_gen parse_qarith Mc.qnegate Mc.qnormalise Mc.qunsat Mc.qdeduce qq_domain_spec
- [ non_linear_prover_Q "real_nonlinear_prover" (Some i) ] gl
+let lra_Q =
+ micromega_gen parse_qarith Mc.qnegate Mc.qnormalise Mc.qunsat Mc.qdeduce qq_domain_spec dump_qexpr
+ [ linear_prover_Q ]
+let psatz_Q i =
+ micromega_gen parse_qarith Mc.qnegate Mc.qnormalise Mc.qunsat Mc.qdeduce qq_domain_spec dump_qexpr
+ [ non_linear_prover_Q "real_nonlinear_prover" (Some i) ]
-let psatzl_R gl =
- micromega_genr [ linear_prover_R ] gl
+let lra_R =
+ micromega_genr [ linear_prover_R ]
+let psatz_R i =
+ micromega_genr [ non_linear_prover_R "real_nonlinear_prover" (Some i) ]
-let psatz_R i gl =
- micromega_genr [ non_linear_prover_R "real_nonlinear_prover" (Some i) ] gl
+let psatz_Z i =
+ micromega_gen parse_zarith Mc.negate Mc.normalise Mc.zunsat Mc.zdeduce zz_domain_spec dump_zexpr
+ [ non_linear_prover_Z "real_nonlinear_prover" (Some i) ]
-let psatz_Z i gl =
- micromega_gen parse_zarith Mc.negate Mc.normalise Mc.zunsat Mc.zdeduce zz_domain_spec
- [ non_linear_prover_Z "real_nonlinear_prover" (Some i) ] gl
+let sos_Z =
+ micromega_gen parse_zarith Mc.negate Mc.normalise Mc.zunsat Mc.zdeduce zz_domain_spec dump_zexpr
+ [ non_linear_prover_Z "pure_sos" None ]
-let sos_Z gl =
- micromega_gen parse_zarith Mc.negate Mc.normalise Mc.zunsat Mc.zdeduce zz_domain_spec
- [ non_linear_prover_Z "pure_sos" None ] gl
+let sos_Q =
+ micromega_gen parse_qarith Mc.qnegate Mc.qnormalise Mc.qunsat Mc.qdeduce qq_domain_spec dump_qexpr
+ [ non_linear_prover_Q "pure_sos" None ]
-let sos_Q gl =
- micromega_gen parse_qarith Mc.qnegate Mc.qnormalise Mc.qunsat Mc.qdeduce qq_domain_spec
- [ non_linear_prover_Q "pure_sos" None ] gl
+let sos_R =
+ micromega_genr [ non_linear_prover_R "pure_sos" None ]
-let sos_R gl =
- micromega_genr [ non_linear_prover_R "pure_sos" None ] gl
+let xlia = micromega_gen parse_zarith Mc.negate Mc.normalise Mc.zunsat Mc.zdeduce zz_domain_spec dump_zexpr
+ [ linear_Z ]
-let xlia gl =
- try
- micromega_gen parse_zarith Mc.negate Mc.normalise Mc.zunsat Mc.zdeduce zz_domain_spec
- [ linear_Z ] gl
- with reraise -> (*Printexc.print_backtrace stdout ;*) raise reraise
+let xnlia =
+ micromega_gen parse_zarith Mc.negate Mc.normalise Mc.zunsat Mc.zdeduce zz_domain_spec dump_zexpr
+ [ nlinear_Z ]
-let xnlia gl =
- try
- micromega_gen parse_zarith Mc.negate Mc.normalise Mc.zunsat Mc.zdeduce zz_domain_spec
- [ nlinear_Z ] gl
- with reraise -> (*Printexc.print_backtrace stdout ;*) raise reraise
+let nra =
+ micromega_genr [ nlinear_prover_R ]
+let nqa =
+ micromega_gen parse_qarith Mc.qnegate Mc.qnormalise Mc.qunsat Mc.qdeduce qq_domain_spec dump_qexpr
+ [ nlinear_prover_R ]
+
(* Local Variables: *)
(* coding: utf-8 *)
diff --git a/plugins/micromega/g_micromega.ml4 b/plugins/micromega/g_micromega.ml4
index 75237aaa..027f690f 100644
--- a/plugins/micromega/g_micromega.ml4
+++ b/plugins/micromega/g_micromega.ml4
@@ -16,63 +16,68 @@
(*i camlp4deps: "grammar/grammar.cma" i*)
-open Errors
-open Misctypes
+open Constrarg
DECLARE PLUGIN "micromega_plugin"
-let out_arg = function
- | ArgVar _ -> anomaly (Pp.str "Unevaluated or_var variable")
- | ArgArg x -> x
+TACTIC EXTEND RED
+| [ "myred" ] -> [ Tactics.red_in_concl ]
+END
+
+
TACTIC EXTEND PsatzZ
-| [ "psatz_Z" int_or_var(i) ] -> [ Proofview.V82.tactic (Coq_micromega.psatz_Z (out_arg i)) ]
-| [ "psatz_Z" ] -> [ Proofview.V82.tactic (Coq_micromega.psatz_Z (-1)) ]
+| [ "psatz_Z" int_or_var(i) tactic(t) ] -> [ (Coq_micromega.psatz_Z i
+ (Tacinterp.tactic_of_value ist t))
+ ]
+| [ "psatz_Z" tactic(t)] -> [ (Coq_micromega.psatz_Z (-1)) (Tacinterp.tactic_of_value ist t) ]
END
TACTIC EXTEND Lia
-[ "xlia" ] -> [ Proofview.V82.tactic (Coq_micromega.xlia) ]
+[ "xlia" tactic(t) ] -> [ (Coq_micromega.xlia (Tacinterp.tactic_of_value ist t)) ]
END
TACTIC EXTEND Nia
-[ "xnlia" ] -> [ Proofview.V82.tactic (Coq_micromega.xnlia) ]
+[ "xnlia" tactic(t) ] -> [ (Coq_micromega.xnlia (Tacinterp.tactic_of_value ist t)) ]
+END
+
+TACTIC EXTEND NRA
+[ "xnra" tactic(t) ] -> [ (Coq_micromega.nra (Tacinterp.tactic_of_value ist t))]
END
+TACTIC EXTEND NQA
+[ "xnqa" tactic(t) ] -> [ (Coq_micromega.nqa (Tacinterp.tactic_of_value ist t))]
+END
+
TACTIC EXTEND Sos_Z
-| [ "sos_Z" ] -> [ Proofview.V82.tactic (Coq_micromega.sos_Z) ]
+| [ "sos_Z" tactic(t) ] -> [ (Coq_micromega.sos_Z (Tacinterp.tactic_of_value ist t)) ]
END
TACTIC EXTEND Sos_Q
-| [ "sos_Q" ] -> [ Proofview.V82.tactic (Coq_micromega.sos_Q) ]
+| [ "sos_Q" tactic(t) ] -> [ (Coq_micromega.sos_Q (Tacinterp.tactic_of_value ist t)) ]
END
TACTIC EXTEND Sos_R
-| [ "sos_R" ] -> [ Proofview.V82.tactic (Coq_micromega.sos_R) ]
-END
-
-(*
-TACTIC EXTEND Omicron
-[ "psatzl_Z" ] -> [ Proofview.V82.tactic (Coq_micromega.psatzl_Z) ]
+| [ "sos_R" tactic(t) ] -> [ (Coq_micromega.sos_R (Tacinterp.tactic_of_value ist t)) ]
END
-*)
TACTIC EXTEND LRA_Q
-[ "psatzl_Q" ] -> [ Proofview.V82.tactic (Coq_micromega.psatzl_Q) ]
+[ "lra_Q" tactic(t) ] -> [ (Coq_micromega.lra_Q (Tacinterp.tactic_of_value ist t)) ]
END
TACTIC EXTEND LRA_R
-[ "psatzl_R" ] -> [ Proofview.V82.tactic (Coq_micromega.psatzl_R) ]
+[ "lra_R" tactic(t) ] -> [ (Coq_micromega.lra_R (Tacinterp.tactic_of_value ist t)) ]
END
TACTIC EXTEND PsatzR
-| [ "psatz_R" int_or_var(i) ] -> [ Proofview.V82.tactic (Coq_micromega.psatz_R (out_arg i)) ]
-| [ "psatz_R" ] -> [ Proofview.V82.tactic (Coq_micromega.psatz_R (-1)) ]
+| [ "psatz_R" int_or_var(i) tactic(t) ] -> [ (Coq_micromega.psatz_R i (Tacinterp.tactic_of_value ist t)) ]
+| [ "psatz_R" tactic(t) ] -> [ (Coq_micromega.psatz_R (-1) (Tacinterp.tactic_of_value ist t)) ]
END
TACTIC EXTEND PsatzQ
-| [ "psatz_Q" int_or_var(i) ] -> [ Proofview.V82.tactic (Coq_micromega.psatz_Q (out_arg i)) ]
-| [ "psatz_Q" ] -> [ Proofview.V82.tactic (Coq_micromega.psatz_Q (-1)) ]
+| [ "psatz_Q" int_or_var(i) tactic(t) ] -> [ (Coq_micromega.psatz_Q i (Tacinterp.tactic_of_value ist t)) ]
+| [ "psatz_Q" tactic(t) ] -> [ (Coq_micromega.psatz_Q (-1) (Tacinterp.tactic_of_value ist t)) ]
END
diff --git a/plugins/micromega/mfourier.ml b/plugins/micromega/mfourier.ml
index a36369d2..f4f9b3c2 100644
--- a/plugins/micromega/mfourier.ml
+++ b/plugins/micromega/mfourier.ml
@@ -98,12 +98,12 @@ module PSet = ISet
module System = Hashtbl.Make(Vect)
- type proof =
- | Hyp of int
- | Elim of var * proof * proof
- | And of proof * proof
-
+type proof =
+| Hyp of int
+| Elim of var * proof * proof
+| And of proof * proof
+let max_nb_cstr = ref max_int
type system = {
sys : cstr_info ref System.t ;
@@ -120,7 +120,7 @@ and cstr_info = {
(** A system of constraints has the form [\{sys = s ; vars = v\}].
[s] is a hashtable mapping a normalised vector to a [cstr_info] record where
- [bound] is an interval
- - [prf_idx] is the set of hypothese indexes (i.e. constraints in the initial system) used to obtain the current constraint.
+ - [prf_idx] is the set of hypothesis indexes (i.e. constraints in the initial system) used to obtain the current constraint.
In the initial system, each constraint is given an unique singleton proof_idx.
When a new constraint c is computed by a function f(c1,...,cn), its proof_idx is ISet.fold union (List.map (fun x -> x.proof_idx) [c1;...;cn]
- [pos] is the number of positive values of the vector
@@ -208,8 +208,7 @@ let merge_cstr_info i1 i2 =
*)
let xadd_cstr vect cstr_info sys =
- if debug && Int.equal (System.length sys mod 1000) 0 then (print_string "*" ; flush stdout) ;
- try
+ try
let info = System.find sys vect in
match merge_cstr_info cstr_info !info with
| None -> raise (SystemContradiction (And(cstr_info.prf, (!info).prf)))
@@ -217,6 +216,13 @@ let xadd_cstr vect cstr_info sys =
with
| Not_found -> System.replace sys vect (ref cstr_info)
+exception TimeOut
+
+let xadd_cstr vect cstr_info sys =
+ if debug && Int.equal (System.length sys mod 1000) 0 then (print_string "*" ; flush stdout) ;
+ if System.length sys < !max_nb_cstr
+ then xadd_cstr vect cstr_info sys
+ else raise TimeOut
type cstr_ext =
| Contradiction (** The constraint is contradictory.
@@ -728,7 +734,7 @@ struct
| Inl (s,_) ->
try
Some (bound_of_variable IMap.empty fresh s.sys)
- with x when Errors.noncritical x ->
+ with x when CErrors.noncritical x ->
Printf.printf "optimise Exception : %s" (Printexc.to_string x);
None
@@ -866,7 +872,7 @@ let mk_proof hyps prf =
| Elim(v,prf1,prf2) ->
let prfsl = mk_proof prf1
and prfsr = mk_proof prf2 in
- (* I take only the pairs for which the elimination is meaningfull *)
+ (* I take only the pairs for which the elimination is meaningful *)
forall_pairs (pivot v) prfsl prfsr
| And(prf1,prf2) ->
let prfsl1 = mk_proof prf1
diff --git a/plugins/micromega/micromega.ml b/plugins/micromega/micromega.ml
index 0537cdbe..5cf1da8e 100644
--- a/plugins/micromega/micromega.ml
+++ b/plugins/micromega/micromega.ml
@@ -1,6 +1,3 @@
-type __ = Obj.t
-let __ = let rec f _ = Obj.repr f in Obj.repr f
-
(** val negb : bool -> bool **)
let negb = function
@@ -11,16 +8,6 @@ type nat =
| O
| S of nat
-(** val fst : ('a1 * 'a2) -> 'a1 **)
-
-let fst = function
-| x,y -> x
-
-(** val snd : ('a1 * 'a2) -> 'a2 **)
-
-let snd = function
-| x,y -> y
-
(** val app : 'a1 list -> 'a1 list -> 'a1 list **)
let rec app l m =
@@ -40,42 +27,15 @@ let compOpp = function
| Lt -> Gt
| Gt -> Lt
-type compareSpecT =
-| CompEqT
-| CompLtT
-| CompGtT
-
-(** val compareSpec2Type : comparison -> compareSpecT **)
-
-let compareSpec2Type = function
-| Eq -> CompEqT
-| Lt -> CompLtT
-| Gt -> CompGtT
-
-type 'a compSpecT = compareSpecT
-
-(** val compSpec2Type : 'a1 -> 'a1 -> comparison -> 'a1 compSpecT **)
-
-let compSpec2Type x y c =
- compareSpec2Type c
-
-type 'a sig0 =
- 'a
- (* singleton inductive, whose constructor was exist *)
-
-(** val plus : nat -> nat -> nat **)
-
-let rec plus n0 m =
- match n0 with
- | O -> m
- | S p -> S (plus p m)
-
-(** val nat_iter : nat -> ('a1 -> 'a1) -> 'a1 -> 'a1 **)
+module Coq__1 = struct
+ (** val add : nat -> nat -> nat **)
+ let rec add n0 m =
+ match n0 with
+ | O -> m
+ | S p -> S (add p m)
+end
+let add = Coq__1.add
-let rec nat_iter n0 f x =
- match n0 with
- | O -> x
- | S n' -> f (nat_iter n' f x)
type positive =
| XI of positive
@@ -91,592 +51,25 @@ type z =
| Zpos of positive
| Zneg of positive
-module type TotalOrder' =
- sig
- type t
- end
-
-module MakeOrderTac =
- functor (O:TotalOrder') ->
- struct
-
- end
-
-module MaxLogicalProperties =
- functor (O:TotalOrder') ->
- functor (M:sig
- val max : O.t -> O.t -> O.t
- end) ->
- struct
- module T = MakeOrderTac(O)
- end
-
-module Pos =
- struct
- type t = positive
-
- (** val succ : positive -> positive **)
-
- let rec succ = function
- | XI p -> XO (succ p)
- | XO p -> XI p
- | XH -> XO XH
-
- (** val add : positive -> positive -> positive **)
-
- let rec add x y =
- match x with
- | XI p ->
- (match y with
- | XI q0 -> XO (add_carry p q0)
- | XO q0 -> XI (add p q0)
- | XH -> XO (succ p))
- | XO p ->
- (match y with
- | XI q0 -> XI (add p q0)
- | XO q0 -> XO (add p q0)
- | XH -> XI p)
- | XH ->
- (match y with
- | XI q0 -> XO (succ q0)
- | XO q0 -> XI q0
- | XH -> XO XH)
-
- (** val add_carry : positive -> positive -> positive **)
-
- and add_carry x y =
- match x with
- | XI p ->
- (match y with
- | XI q0 -> XI (add_carry p q0)
- | XO q0 -> XO (add_carry p q0)
- | XH -> XI (succ p))
- | XO p ->
- (match y with
- | XI q0 -> XO (add_carry p q0)
- | XO q0 -> XI (add p q0)
- | XH -> XO (succ p))
- | XH ->
- (match y with
- | XI q0 -> XI (succ q0)
- | XO q0 -> XO (succ q0)
- | XH -> XI XH)
-
- (** val pred_double : positive -> positive **)
-
- let rec pred_double = function
- | XI p -> XI (XO p)
- | XO p -> XI (pred_double p)
- | XH -> XH
-
- (** val pred : positive -> positive **)
-
- let pred = function
- | XI p -> XO p
- | XO p -> pred_double p
- | XH -> XH
-
- (** val pred_N : positive -> n **)
-
- let pred_N = function
- | XI p -> Npos (XO p)
- | XO p -> Npos (pred_double p)
- | XH -> N0
-
+module Pos =
+ struct
type mask =
| IsNul
| IsPos of positive
| IsNeg
-
- (** val mask_rect : 'a1 -> (positive -> 'a1) -> 'a1 -> mask -> 'a1 **)
-
- let mask_rect f f0 f1 = function
- | IsNul -> f
- | IsPos x -> f0 x
- | IsNeg -> f1
-
- (** val mask_rec : 'a1 -> (positive -> 'a1) -> 'a1 -> mask -> 'a1 **)
-
- let mask_rec f f0 f1 = function
- | IsNul -> f
- | IsPos x -> f0 x
- | IsNeg -> f1
-
- (** val succ_double_mask : mask -> mask **)
-
- let succ_double_mask = function
- | IsNul -> IsPos XH
- | IsPos p -> IsPos (XI p)
- | IsNeg -> IsNeg
-
- (** val double_mask : mask -> mask **)
-
- let double_mask = function
- | IsPos p -> IsPos (XO p)
- | x0 -> x0
-
- (** val double_pred_mask : positive -> mask **)
-
- let double_pred_mask = function
- | XI p -> IsPos (XO (XO p))
- | XO p -> IsPos (XO (pred_double p))
- | XH -> IsNul
-
- (** val pred_mask : mask -> mask **)
-
- let pred_mask = function
- | IsPos q0 ->
- (match q0 with
- | XH -> IsNul
- | _ -> IsPos (pred q0))
- | _ -> IsNeg
-
- (** val sub_mask : positive -> positive -> mask **)
-
- let rec sub_mask x y =
- match x with
- | XI p ->
- (match y with
- | XI q0 -> double_mask (sub_mask p q0)
- | XO q0 -> succ_double_mask (sub_mask p q0)
- | XH -> IsPos (XO p))
- | XO p ->
- (match y with
- | XI q0 -> succ_double_mask (sub_mask_carry p q0)
- | XO q0 -> double_mask (sub_mask p q0)
- | XH -> IsPos (pred_double p))
- | XH ->
- (match y with
- | XH -> IsNul
- | _ -> IsNeg)
-
- (** val sub_mask_carry : positive -> positive -> mask **)
-
- and sub_mask_carry x y =
- match x with
- | XI p ->
- (match y with
- | XI q0 -> succ_double_mask (sub_mask_carry p q0)
- | XO q0 -> double_mask (sub_mask p q0)
- | XH -> IsPos (pred_double p))
- | XO p ->
- (match y with
- | XI q0 -> double_mask (sub_mask_carry p q0)
- | XO q0 -> succ_double_mask (sub_mask_carry p q0)
- | XH -> double_pred_mask p)
- | XH -> IsNeg
-
- (** val sub : positive -> positive -> positive **)
-
- let sub x y =
- match sub_mask x y with
- | IsPos z0 -> z0
- | _ -> XH
-
- (** val mul : positive -> positive -> positive **)
-
- let rec mul x y =
- match x with
- | XI p -> add y (XO (mul p y))
- | XO p -> XO (mul p y)
- | XH -> y
-
- (** val iter : positive -> ('a1 -> 'a1) -> 'a1 -> 'a1 **)
-
- let rec iter n0 f x =
- match n0 with
- | XI n' -> f (iter n' f (iter n' f x))
- | XO n' -> iter n' f (iter n' f x)
- | XH -> f x
-
- (** val pow : positive -> positive -> positive **)
-
- let pow x y =
- iter y (mul x) XH
-
- (** val div2 : positive -> positive **)
-
- let div2 = function
- | XI p2 -> p2
- | XO p2 -> p2
- | XH -> XH
-
- (** val div2_up : positive -> positive **)
-
- let div2_up = function
- | XI p2 -> succ p2
- | XO p2 -> p2
- | XH -> XH
-
- (** val size_nat : positive -> nat **)
-
- let rec size_nat = function
- | XI p2 -> S (size_nat p2)
- | XO p2 -> S (size_nat p2)
- | XH -> S O
-
- (** val size : positive -> positive **)
-
- let rec size = function
- | XI p2 -> succ (size p2)
- | XO p2 -> succ (size p2)
- | XH -> XH
-
- (** val compare_cont : positive -> positive -> comparison -> comparison **)
-
- let rec compare_cont x y r =
- match x with
- | XI p ->
- (match y with
- | XI q0 -> compare_cont p q0 r
- | XO q0 -> compare_cont p q0 Gt
- | XH -> Gt)
- | XO p ->
- (match y with
- | XI q0 -> compare_cont p q0 Lt
- | XO q0 -> compare_cont p q0 r
- | XH -> Gt)
- | XH ->
- (match y with
- | XH -> r
- | _ -> Lt)
-
- (** val compare : positive -> positive -> comparison **)
-
- let compare x y =
- compare_cont x y Eq
-
- (** val min : positive -> positive -> positive **)
-
- let min p p' =
- match compare p p' with
- | Gt -> p'
- | _ -> p
-
- (** val max : positive -> positive -> positive **)
-
- let max p p' =
- match compare p p' with
- | Gt -> p
- | _ -> p'
-
- (** val eqb : positive -> positive -> bool **)
-
- let rec eqb p q0 =
- match p with
- | XI p2 ->
- (match q0 with
- | XI q1 -> eqb p2 q1
- | _ -> false)
- | XO p2 ->
- (match q0 with
- | XO q1 -> eqb p2 q1
- | _ -> false)
- | XH ->
- (match q0 with
- | XH -> true
- | _ -> false)
-
- (** val leb : positive -> positive -> bool **)
-
- let leb x y =
- match compare x y with
- | Gt -> false
- | _ -> true
-
- (** val ltb : positive -> positive -> bool **)
-
- let ltb x y =
- match compare x y with
- | Lt -> true
- | _ -> false
-
- (** val sqrtrem_step :
- (positive -> positive) -> (positive -> positive) -> (positive * mask)
- -> positive * mask **)
-
- let sqrtrem_step f g = function
- | s,y ->
- (match y with
- | IsPos r ->
- let s' = XI (XO s) in
- let r' = g (f r) in
- if leb s' r' then (XI s),(sub_mask r' s') else (XO s),(IsPos r')
- | _ -> (XO s),(sub_mask (g (f XH)) (XO (XO XH))))
-
- (** val sqrtrem : positive -> positive * mask **)
-
- let rec sqrtrem = function
- | XI p2 ->
- (match p2 with
- | XI p3 -> sqrtrem_step (fun x -> XI x) (fun x -> XI x) (sqrtrem p3)
- | XO p3 -> sqrtrem_step (fun x -> XO x) (fun x -> XI x) (sqrtrem p3)
- | XH -> XH,(IsPos (XO XH)))
- | XO p2 ->
- (match p2 with
- | XI p3 -> sqrtrem_step (fun x -> XI x) (fun x -> XO x) (sqrtrem p3)
- | XO p3 -> sqrtrem_step (fun x -> XO x) (fun x -> XO x) (sqrtrem p3)
- | XH -> XH,(IsPos XH))
- | XH -> XH,IsNul
-
- (** val sqrt : positive -> positive **)
-
- let sqrt p =
- fst (sqrtrem p)
-
- (** val gcdn : nat -> positive -> positive -> positive **)
-
- let rec gcdn n0 a b =
- match n0 with
- | O -> XH
- | S n1 ->
- (match a with
- | XI a' ->
- (match b with
- | XI b' ->
- (match compare a' b' with
- | Eq -> a
- | Lt -> gcdn n1 (sub b' a') a
- | Gt -> gcdn n1 (sub a' b') b)
- | XO b0 -> gcdn n1 a b0
- | XH -> XH)
- | XO a0 ->
- (match b with
- | XI p -> gcdn n1 a0 b
- | XO b0 -> XO (gcdn n1 a0 b0)
- | XH -> XH)
- | XH -> XH)
-
- (** val gcd : positive -> positive -> positive **)
-
- let gcd a b =
- gcdn (plus (size_nat a) (size_nat b)) a b
-
- (** val ggcdn :
- nat -> positive -> positive -> positive * (positive * positive) **)
-
- let rec ggcdn n0 a b =
- match n0 with
- | O -> XH,(a,b)
- | S n1 ->
- (match a with
- | XI a' ->
- (match b with
- | XI b' ->
- (match compare a' b' with
- | Eq -> a,(XH,XH)
- | Lt ->
- let g,p = ggcdn n1 (sub b' a') a in
- let ba,aa = p in g,(aa,(add aa (XO ba)))
- | Gt ->
- let g,p = ggcdn n1 (sub a' b') b in
- let ab,bb = p in g,((add bb (XO ab)),bb))
- | XO b0 ->
- let g,p = ggcdn n1 a b0 in let aa,bb = p in g,(aa,(XO bb))
- | XH -> XH,(a,XH))
- | XO a0 ->
- (match b with
- | XI p ->
- let g,p2 = ggcdn n1 a0 b in let aa,bb = p2 in g,((XO aa),bb)
- | XO b0 -> let g,p = ggcdn n1 a0 b0 in (XO g),p
- | XH -> XH,(a,XH))
- | XH -> XH,(XH,b))
-
- (** val ggcd : positive -> positive -> positive * (positive * positive) **)
-
- let ggcd a b =
- ggcdn (plus (size_nat a) (size_nat b)) a b
-
- (** val coq_Nsucc_double : n -> n **)
-
- let coq_Nsucc_double = function
- | N0 -> Npos XH
- | Npos p -> Npos (XI p)
-
- (** val coq_Ndouble : n -> n **)
-
- let coq_Ndouble = function
- | N0 -> N0
- | Npos p -> Npos (XO p)
-
- (** val coq_lor : positive -> positive -> positive **)
-
- let rec coq_lor p q0 =
- match p with
- | XI p2 ->
- (match q0 with
- | XI q1 -> XI (coq_lor p2 q1)
- | XO q1 -> XI (coq_lor p2 q1)
- | XH -> p)
- | XO p2 ->
- (match q0 with
- | XI q1 -> XI (coq_lor p2 q1)
- | XO q1 -> XO (coq_lor p2 q1)
- | XH -> XI p2)
- | XH ->
- (match q0 with
- | XO q1 -> XI q1
- | _ -> q0)
-
- (** val coq_land : positive -> positive -> n **)
-
- let rec coq_land p q0 =
- match p with
- | XI p2 ->
- (match q0 with
- | XI q1 -> coq_Nsucc_double (coq_land p2 q1)
- | XO q1 -> coq_Ndouble (coq_land p2 q1)
- | XH -> Npos XH)
- | XO p2 ->
- (match q0 with
- | XI q1 -> coq_Ndouble (coq_land p2 q1)
- | XO q1 -> coq_Ndouble (coq_land p2 q1)
- | XH -> N0)
- | XH ->
- (match q0 with
- | XO q1 -> N0
- | _ -> Npos XH)
-
- (** val ldiff : positive -> positive -> n **)
-
- let rec ldiff p q0 =
- match p with
- | XI p2 ->
- (match q0 with
- | XI q1 -> coq_Ndouble (ldiff p2 q1)
- | XO q1 -> coq_Nsucc_double (ldiff p2 q1)
- | XH -> Npos (XO p2))
- | XO p2 ->
- (match q0 with
- | XI q1 -> coq_Ndouble (ldiff p2 q1)
- | XO q1 -> coq_Ndouble (ldiff p2 q1)
- | XH -> Npos p)
- | XH ->
- (match q0 with
- | XO q1 -> Npos XH
- | _ -> N0)
-
- (** val coq_lxor : positive -> positive -> n **)
-
- let rec coq_lxor p q0 =
- match p with
- | XI p2 ->
- (match q0 with
- | XI q1 -> coq_Ndouble (coq_lxor p2 q1)
- | XO q1 -> coq_Nsucc_double (coq_lxor p2 q1)
- | XH -> Npos (XO p2))
- | XO p2 ->
- (match q0 with
- | XI q1 -> coq_Nsucc_double (coq_lxor p2 q1)
- | XO q1 -> coq_Ndouble (coq_lxor p2 q1)
- | XH -> Npos (XI p2))
- | XH ->
- (match q0 with
- | XI q1 -> Npos (XO q1)
- | XO q1 -> Npos (XI q1)
- | XH -> N0)
-
- (** val shiftl_nat : positive -> nat -> positive **)
-
- let shiftl_nat p n0 =
- nat_iter n0 (fun x -> XO x) p
-
- (** val shiftr_nat : positive -> nat -> positive **)
-
- let shiftr_nat p n0 =
- nat_iter n0 div2 p
-
- (** val shiftl : positive -> n -> positive **)
-
- let shiftl p = function
- | N0 -> p
- | Npos n1 -> iter n1 (fun x -> XO x) p
-
- (** val shiftr : positive -> n -> positive **)
-
- let shiftr p = function
- | N0 -> p
- | Npos n1 -> iter n1 div2 p
-
- (** val testbit_nat : positive -> nat -> bool **)
-
- let rec testbit_nat p n0 =
- match p with
- | XI p2 ->
- (match n0 with
- | O -> true
- | S n' -> testbit_nat p2 n')
- | XO p2 ->
- (match n0 with
- | O -> false
- | S n' -> testbit_nat p2 n')
- | XH ->
- (match n0 with
- | O -> true
- | S n1 -> false)
-
- (** val testbit : positive -> n -> bool **)
-
- let rec testbit p n0 =
- match p with
- | XI p2 ->
- (match n0 with
- | N0 -> true
- | Npos n1 -> testbit p2 (pred_N n1))
- | XO p2 ->
- (match n0 with
- | N0 -> false
- | Npos n1 -> testbit p2 (pred_N n1))
- | XH ->
- (match n0 with
- | N0 -> true
- | Npos p2 -> false)
-
- (** val iter_op : ('a1 -> 'a1 -> 'a1) -> positive -> 'a1 -> 'a1 **)
-
- let rec iter_op op p a =
- match p with
- | XI p2 -> op a (iter_op op p2 (op a a))
- | XO p2 -> iter_op op p2 (op a a)
- | XH -> a
-
- (** val to_nat : positive -> nat **)
-
- let to_nat x =
- iter_op plus x (S O)
-
- (** val of_nat : nat -> positive **)
-
- let rec of_nat = function
- | O -> XH
- | S x ->
- (match x with
- | O -> XH
- | S n1 -> succ (of_nat x))
-
- (** val of_succ_nat : nat -> positive **)
-
- let rec of_succ_nat = function
- | O -> XH
- | S x -> succ (of_succ_nat x)
end
-module Coq_Pos =
- struct
- module Coq__1 = struct
- type t = positive
- end
- type t = Coq__1.t
-
+module Coq_Pos =
+ struct
(** val succ : positive -> positive **)
-
+
let rec succ = function
| XI p -> XO (succ p)
| XO p -> XI p
| XH -> XO XH
-
+
(** val add : positive -> positive -> positive **)
-
+
let rec add x y =
match x with
| XI p ->
@@ -694,9 +87,9 @@ module Coq_Pos =
| XI q0 -> XO (succ q0)
| XO q0 -> XI q0
| XH -> XO XH)
-
+
(** val add_carry : positive -> positive -> positive **)
-
+
and add_carry x y =
match x with
| XI p ->
@@ -714,78 +107,41 @@ module Coq_Pos =
| XI q0 -> XI (succ q0)
| XO q0 -> XO (succ q0)
| XH -> XI XH)
-
+
(** val pred_double : positive -> positive **)
-
+
let rec pred_double = function
| XI p -> XI (XO p)
| XO p -> XI (pred_double p)
| XH -> XH
-
- (** val pred : positive -> positive **)
-
- let pred = function
- | XI p -> XO p
- | XO p -> pred_double p
- | XH -> XH
-
- (** val pred_N : positive -> n **)
-
- let pred_N = function
- | XI p -> Npos (XO p)
- | XO p -> Npos (pred_double p)
- | XH -> N0
-
+
type mask = Pos.mask =
| IsNul
| IsPos of positive
| IsNeg
-
- (** val mask_rect : 'a1 -> (positive -> 'a1) -> 'a1 -> mask -> 'a1 **)
-
- let mask_rect f f0 f1 = function
- | IsNul -> f
- | IsPos x -> f0 x
- | IsNeg -> f1
-
- (** val mask_rec : 'a1 -> (positive -> 'a1) -> 'a1 -> mask -> 'a1 **)
-
- let mask_rec f f0 f1 = function
- | IsNul -> f
- | IsPos x -> f0 x
- | IsNeg -> f1
-
+
(** val succ_double_mask : mask -> mask **)
-
+
let succ_double_mask = function
| IsNul -> IsPos XH
| IsPos p -> IsPos (XI p)
| IsNeg -> IsNeg
-
+
(** val double_mask : mask -> mask **)
-
+
let double_mask = function
| IsPos p -> IsPos (XO p)
| x0 -> x0
-
+
(** val double_pred_mask : positive -> mask **)
-
+
let double_pred_mask = function
| XI p -> IsPos (XO (XO p))
| XO p -> IsPos (XO (pred_double p))
| XH -> IsNul
-
- (** val pred_mask : mask -> mask **)
-
- let pred_mask = function
- | IsPos q0 ->
- (match q0 with
- | XH -> IsNul
- | _ -> IsPos (pred q0))
- | _ -> IsNeg
-
+
(** val sub_mask : positive -> positive -> mask **)
-
+
let rec sub_mask x y =
match x with
| XI p ->
@@ -802,9 +158,9 @@ module Coq_Pos =
(match y with
| XH -> IsNul
| _ -> IsNeg)
-
+
(** val sub_mask_carry : positive -> positive -> mask **)
-
+
and sub_mask_carry x y =
match x with
| XI p ->
@@ -818,167 +174,56 @@ module Coq_Pos =
| XO q0 -> succ_double_mask (sub_mask_carry p q0)
| XH -> double_pred_mask p)
| XH -> IsNeg
-
+
(** val sub : positive -> positive -> positive **)
-
+
let sub x y =
match sub_mask x y with
| IsPos z0 -> z0
| _ -> XH
-
+
(** val mul : positive -> positive -> positive **)
-
+
let rec mul x y =
match x with
| XI p -> add y (XO (mul p y))
| XO p -> XO (mul p y)
| XH -> y
-
- (** val iter : positive -> ('a1 -> 'a1) -> 'a1 -> 'a1 **)
-
- let rec iter n0 f x =
- match n0 with
- | XI n' -> f (iter n' f (iter n' f x))
- | XO n' -> iter n' f (iter n' f x)
- | XH -> f x
-
- (** val pow : positive -> positive -> positive **)
-
- let pow x y =
- iter y (mul x) XH
-
- (** val div2 : positive -> positive **)
-
- let div2 = function
- | XI p2 -> p2
- | XO p2 -> p2
- | XH -> XH
-
- (** val div2_up : positive -> positive **)
-
- let div2_up = function
- | XI p2 -> succ p2
- | XO p2 -> p2
- | XH -> XH
-
+
(** val size_nat : positive -> nat **)
-
+
let rec size_nat = function
| XI p2 -> S (size_nat p2)
| XO p2 -> S (size_nat p2)
| XH -> S O
-
- (** val size : positive -> positive **)
-
- let rec size = function
- | XI p2 -> succ (size p2)
- | XO p2 -> succ (size p2)
- | XH -> XH
-
- (** val compare_cont : positive -> positive -> comparison -> comparison **)
-
- let rec compare_cont x y r =
+
+ (** val compare_cont :
+ comparison -> positive -> positive -> comparison **)
+
+ let rec compare_cont r x y =
match x with
| XI p ->
(match y with
- | XI q0 -> compare_cont p q0 r
- | XO q0 -> compare_cont p q0 Gt
+ | XI q0 -> compare_cont r p q0
+ | XO q0 -> compare_cont Gt p q0
| XH -> Gt)
| XO p ->
(match y with
- | XI q0 -> compare_cont p q0 Lt
- | XO q0 -> compare_cont p q0 r
+ | XI q0 -> compare_cont Lt p q0
+ | XO q0 -> compare_cont r p q0
| XH -> Gt)
| XH ->
(match y with
| XH -> r
| _ -> Lt)
-
+
(** val compare : positive -> positive -> comparison **)
-
- let compare x y =
- compare_cont x y Eq
-
- (** val min : positive -> positive -> positive **)
-
- let min p p' =
- match compare p p' with
- | Gt -> p'
- | _ -> p
-
- (** val max : positive -> positive -> positive **)
-
- let max p p' =
- match compare p p' with
- | Gt -> p
- | _ -> p'
-
- (** val eqb : positive -> positive -> bool **)
-
- let rec eqb p q0 =
- match p with
- | XI p2 ->
- (match q0 with
- | XI q1 -> eqb p2 q1
- | _ -> false)
- | XO p2 ->
- (match q0 with
- | XO q1 -> eqb p2 q1
- | _ -> false)
- | XH ->
- (match q0 with
- | XH -> true
- | _ -> false)
-
- (** val leb : positive -> positive -> bool **)
-
- let leb x y =
- match compare x y with
- | Gt -> false
- | _ -> true
-
- (** val ltb : positive -> positive -> bool **)
-
- let ltb x y =
- match compare x y with
- | Lt -> true
- | _ -> false
-
- (** val sqrtrem_step :
- (positive -> positive) -> (positive -> positive) -> (positive * mask)
- -> positive * mask **)
-
- let sqrtrem_step f g = function
- | s,y ->
- (match y with
- | IsPos r ->
- let s' = XI (XO s) in
- let r' = g (f r) in
- if leb s' r' then (XI s),(sub_mask r' s') else (XO s),(IsPos r')
- | _ -> (XO s),(sub_mask (g (f XH)) (XO (XO XH))))
-
- (** val sqrtrem : positive -> positive * mask **)
-
- let rec sqrtrem = function
- | XI p2 ->
- (match p2 with
- | XI p3 -> sqrtrem_step (fun x -> XI x) (fun x -> XI x) (sqrtrem p3)
- | XO p3 -> sqrtrem_step (fun x -> XO x) (fun x -> XI x) (sqrtrem p3)
- | XH -> XH,(IsPos (XO XH)))
- | XO p2 ->
- (match p2 with
- | XI p3 -> sqrtrem_step (fun x -> XI x) (fun x -> XO x) (sqrtrem p3)
- | XO p3 -> sqrtrem_step (fun x -> XO x) (fun x -> XO x) (sqrtrem p3)
- | XH -> XH,(IsPos XH))
- | XH -> XH,IsNul
-
- (** val sqrt : positive -> positive **)
-
- let sqrt p =
- fst (sqrtrem p)
-
+
+ let compare =
+ compare_cont Eq
+
(** val gcdn : nat -> positive -> positive -> positive **)
-
+
let rec gcdn n0 a b =
match n0 with
| O -> XH
@@ -995,1001 +240,30 @@ module Coq_Pos =
| XH -> XH)
| XO a0 ->
(match b with
- | XI p -> gcdn n1 a0 b
+ | XI _ -> gcdn n1 a0 b
| XO b0 -> XO (gcdn n1 a0 b0)
| XH -> XH)
| XH -> XH)
-
+
(** val gcd : positive -> positive -> positive **)
-
+
let gcd a b =
- gcdn (plus (size_nat a) (size_nat b)) a b
-
- (** val ggcdn :
- nat -> positive -> positive -> positive * (positive * positive) **)
-
- let rec ggcdn n0 a b =
- match n0 with
- | O -> XH,(a,b)
- | S n1 ->
- (match a with
- | XI a' ->
- (match b with
- | XI b' ->
- (match compare a' b' with
- | Eq -> a,(XH,XH)
- | Lt ->
- let g,p = ggcdn n1 (sub b' a') a in
- let ba,aa = p in g,(aa,(add aa (XO ba)))
- | Gt ->
- let g,p = ggcdn n1 (sub a' b') b in
- let ab,bb = p in g,((add bb (XO ab)),bb))
- | XO b0 ->
- let g,p = ggcdn n1 a b0 in let aa,bb = p in g,(aa,(XO bb))
- | XH -> XH,(a,XH))
- | XO a0 ->
- (match b with
- | XI p ->
- let g,p2 = ggcdn n1 a0 b in let aa,bb = p2 in g,((XO aa),bb)
- | XO b0 -> let g,p = ggcdn n1 a0 b0 in (XO g),p
- | XH -> XH,(a,XH))
- | XH -> XH,(XH,b))
-
- (** val ggcd : positive -> positive -> positive * (positive * positive) **)
-
- let ggcd a b =
- ggcdn (plus (size_nat a) (size_nat b)) a b
-
- (** val coq_Nsucc_double : n -> n **)
-
- let coq_Nsucc_double = function
- | N0 -> Npos XH
- | Npos p -> Npos (XI p)
-
- (** val coq_Ndouble : n -> n **)
-
- let coq_Ndouble = function
- | N0 -> N0
- | Npos p -> Npos (XO p)
-
- (** val coq_lor : positive -> positive -> positive **)
-
- let rec coq_lor p q0 =
- match p with
- | XI p2 ->
- (match q0 with
- | XI q1 -> XI (coq_lor p2 q1)
- | XO q1 -> XI (coq_lor p2 q1)
- | XH -> p)
- | XO p2 ->
- (match q0 with
- | XI q1 -> XI (coq_lor p2 q1)
- | XO q1 -> XO (coq_lor p2 q1)
- | XH -> XI p2)
- | XH ->
- (match q0 with
- | XO q1 -> XI q1
- | _ -> q0)
-
- (** val coq_land : positive -> positive -> n **)
-
- let rec coq_land p q0 =
- match p with
- | XI p2 ->
- (match q0 with
- | XI q1 -> coq_Nsucc_double (coq_land p2 q1)
- | XO q1 -> coq_Ndouble (coq_land p2 q1)
- | XH -> Npos XH)
- | XO p2 ->
- (match q0 with
- | XI q1 -> coq_Ndouble (coq_land p2 q1)
- | XO q1 -> coq_Ndouble (coq_land p2 q1)
- | XH -> N0)
- | XH ->
- (match q0 with
- | XO q1 -> N0
- | _ -> Npos XH)
-
- (** val ldiff : positive -> positive -> n **)
-
- let rec ldiff p q0 =
- match p with
- | XI p2 ->
- (match q0 with
- | XI q1 -> coq_Ndouble (ldiff p2 q1)
- | XO q1 -> coq_Nsucc_double (ldiff p2 q1)
- | XH -> Npos (XO p2))
- | XO p2 ->
- (match q0 with
- | XI q1 -> coq_Ndouble (ldiff p2 q1)
- | XO q1 -> coq_Ndouble (ldiff p2 q1)
- | XH -> Npos p)
- | XH ->
- (match q0 with
- | XO q1 -> Npos XH
- | _ -> N0)
-
- (** val coq_lxor : positive -> positive -> n **)
-
- let rec coq_lxor p q0 =
- match p with
- | XI p2 ->
- (match q0 with
- | XI q1 -> coq_Ndouble (coq_lxor p2 q1)
- | XO q1 -> coq_Nsucc_double (coq_lxor p2 q1)
- | XH -> Npos (XO p2))
- | XO p2 ->
- (match q0 with
- | XI q1 -> coq_Nsucc_double (coq_lxor p2 q1)
- | XO q1 -> coq_Ndouble (coq_lxor p2 q1)
- | XH -> Npos (XI p2))
- | XH ->
- (match q0 with
- | XI q1 -> Npos (XO q1)
- | XO q1 -> Npos (XI q1)
- | XH -> N0)
-
- (** val shiftl_nat : positive -> nat -> positive **)
-
- let shiftl_nat p n0 =
- nat_iter n0 (fun x -> XO x) p
-
- (** val shiftr_nat : positive -> nat -> positive **)
-
- let shiftr_nat p n0 =
- nat_iter n0 div2 p
-
- (** val shiftl : positive -> n -> positive **)
-
- let shiftl p = function
- | N0 -> p
- | Npos n1 -> iter n1 (fun x -> XO x) p
-
- (** val shiftr : positive -> n -> positive **)
-
- let shiftr p = function
- | N0 -> p
- | Npos n1 -> iter n1 div2 p
-
- (** val testbit_nat : positive -> nat -> bool **)
-
- let rec testbit_nat p n0 =
- match p with
- | XI p2 ->
- (match n0 with
- | O -> true
- | S n' -> testbit_nat p2 n')
- | XO p2 ->
- (match n0 with
- | O -> false
- | S n' -> testbit_nat p2 n')
- | XH ->
- (match n0 with
- | O -> true
- | S n1 -> false)
-
- (** val testbit : positive -> n -> bool **)
-
- let rec testbit p n0 =
- match p with
- | XI p2 ->
- (match n0 with
- | N0 -> true
- | Npos n1 -> testbit p2 (pred_N n1))
- | XO p2 ->
- (match n0 with
- | N0 -> false
- | Npos n1 -> testbit p2 (pred_N n1))
- | XH ->
- (match n0 with
- | N0 -> true
- | Npos p2 -> false)
-
- (** val iter_op : ('a1 -> 'a1 -> 'a1) -> positive -> 'a1 -> 'a1 **)
-
- let rec iter_op op p a =
- match p with
- | XI p2 -> op a (iter_op op p2 (op a a))
- | XO p2 -> iter_op op p2 (op a a)
- | XH -> a
-
- (** val to_nat : positive -> nat **)
-
- let to_nat x =
- iter_op plus x (S O)
-
- (** val of_nat : nat -> positive **)
-
- let rec of_nat = function
- | O -> XH
- | S x ->
- (match x with
- | O -> XH
- | S n1 -> succ (of_nat x))
-
+ gcdn (Coq__1.add (size_nat a) (size_nat b)) a b
+
(** val of_succ_nat : nat -> positive **)
-
+
let rec of_succ_nat = function
| O -> XH
| S x -> succ (of_succ_nat x)
-
- (** val eq_dec : positive -> positive -> bool **)
-
- let rec eq_dec p y0 =
- match p with
- | XI p2 ->
- (match y0 with
- | XI p3 -> eq_dec p2 p3
- | _ -> false)
- | XO p2 ->
- (match y0 with
- | XO p3 -> eq_dec p2 p3
- | _ -> false)
- | XH ->
- (match y0 with
- | XH -> true
- | _ -> false)
-
- (** val peano_rect : 'a1 -> (positive -> 'a1 -> 'a1) -> positive -> 'a1 **)
-
- let rec peano_rect a f p =
- let f2 = peano_rect (f XH a) (fun p2 x -> f (succ (XO p2)) (f (XO p2) x))
- in
- (match p with
- | XI q0 -> f (XO q0) (f2 q0)
- | XO q0 -> f2 q0
- | XH -> a)
-
- (** val peano_rec : 'a1 -> (positive -> 'a1 -> 'a1) -> positive -> 'a1 **)
-
- let peano_rec =
- peano_rect
-
- type coq_PeanoView =
- | PeanoOne
- | PeanoSucc of positive * coq_PeanoView
-
- (** val coq_PeanoView_rect :
- 'a1 -> (positive -> coq_PeanoView -> 'a1 -> 'a1) -> positive ->
- coq_PeanoView -> 'a1 **)
-
- let rec coq_PeanoView_rect f f0 p = function
- | PeanoOne -> f
- | PeanoSucc (p3, p4) -> f0 p3 p4 (coq_PeanoView_rect f f0 p3 p4)
-
- (** val coq_PeanoView_rec :
- 'a1 -> (positive -> coq_PeanoView -> 'a1 -> 'a1) -> positive ->
- coq_PeanoView -> 'a1 **)
-
- let rec coq_PeanoView_rec f f0 p = function
- | PeanoOne -> f
- | PeanoSucc (p3, p4) -> f0 p3 p4 (coq_PeanoView_rec f f0 p3 p4)
-
- (** val peanoView_xO : positive -> coq_PeanoView -> coq_PeanoView **)
-
- let rec peanoView_xO p = function
- | PeanoOne -> PeanoSucc (XH, PeanoOne)
- | PeanoSucc (p2, q1) ->
- PeanoSucc ((succ (XO p2)), (PeanoSucc ((XO p2), (peanoView_xO p2 q1))))
-
- (** val peanoView_xI : positive -> coq_PeanoView -> coq_PeanoView **)
-
- let rec peanoView_xI p = function
- | PeanoOne -> PeanoSucc ((succ XH), (PeanoSucc (XH, PeanoOne)))
- | PeanoSucc (p2, q1) ->
- PeanoSucc ((succ (XI p2)), (PeanoSucc ((XI p2), (peanoView_xI p2 q1))))
-
- (** val peanoView : positive -> coq_PeanoView **)
-
- let rec peanoView = function
- | XI p2 -> peanoView_xI p2 (peanoView p2)
- | XO p2 -> peanoView_xO p2 (peanoView p2)
- | XH -> PeanoOne
-
- (** val coq_PeanoView_iter :
- 'a1 -> (positive -> 'a1 -> 'a1) -> positive -> coq_PeanoView -> 'a1 **)
-
- let rec coq_PeanoView_iter a f p = function
- | PeanoOne -> a
- | PeanoSucc (p2, q1) -> f p2 (coq_PeanoView_iter a f p2 q1)
-
- (** val switch_Eq : comparison -> comparison -> comparison **)
-
- let switch_Eq c = function
- | Eq -> c
- | x -> x
-
- (** val mask2cmp : mask -> comparison **)
-
- let mask2cmp = function
- | IsNul -> Eq
- | IsPos p2 -> Gt
- | IsNeg -> Lt
-
- module T =
- struct
-
- end
-
- module ORev =
- struct
- type t = Coq__1.t
- end
-
- module MRev =
- struct
- (** val max : t -> t -> t **)
-
- let max x y =
- min y x
- end
-
- module MPRev = MaxLogicalProperties(ORev)(MRev)
-
- module P =
- struct
- (** val max_case_strong :
- t -> t -> (t -> t -> __ -> 'a1 -> 'a1) -> (__ -> 'a1) -> (__ -> 'a1)
- -> 'a1 **)
-
- let max_case_strong n0 m compat hl hr =
- let c = compSpec2Type n0 m (compare n0 m) in
- (match c with
- | CompGtT -> compat n0 (max n0 m) __ (hl __)
- | _ -> compat m (max n0 m) __ (hr __))
-
- (** val max_case :
- t -> t -> (t -> t -> __ -> 'a1 -> 'a1) -> 'a1 -> 'a1 -> 'a1 **)
-
- let max_case n0 m x x0 x1 =
- max_case_strong n0 m x (fun _ -> x0) (fun _ -> x1)
-
- (** val max_dec : t -> t -> bool **)
-
- let max_dec n0 m =
- max_case n0 m (fun x y _ h0 -> h0) true false
-
- (** val min_case_strong :
- t -> t -> (t -> t -> __ -> 'a1 -> 'a1) -> (__ -> 'a1) -> (__ -> 'a1)
- -> 'a1 **)
-
- let min_case_strong n0 m compat hl hr =
- let c = compSpec2Type n0 m (compare n0 m) in
- (match c with
- | CompGtT -> compat m (min n0 m) __ (hr __)
- | _ -> compat n0 (min n0 m) __ (hl __))
-
- (** val min_case :
- t -> t -> (t -> t -> __ -> 'a1 -> 'a1) -> 'a1 -> 'a1 -> 'a1 **)
-
- let min_case n0 m x x0 x1 =
- min_case_strong n0 m x (fun _ -> x0) (fun _ -> x1)
-
- (** val min_dec : t -> t -> bool **)
-
- let min_dec n0 m =
- min_case n0 m (fun x y _ h0 -> h0) true false
- end
-
- (** val max_case_strong : t -> t -> (__ -> 'a1) -> (__ -> 'a1) -> 'a1 **)
-
- let max_case_strong n0 m x x0 =
- P.max_case_strong n0 m (fun x1 y _ x2 -> x2) x x0
-
- (** val max_case : t -> t -> 'a1 -> 'a1 -> 'a1 **)
-
- let max_case n0 m x x0 =
- max_case_strong n0 m (fun _ -> x) (fun _ -> x0)
-
- (** val max_dec : t -> t -> bool **)
-
- let max_dec =
- P.max_dec
-
- (** val min_case_strong : t -> t -> (__ -> 'a1) -> (__ -> 'a1) -> 'a1 **)
-
- let min_case_strong n0 m x x0 =
- P.min_case_strong n0 m (fun x1 y _ x2 -> x2) x x0
-
- (** val min_case : t -> t -> 'a1 -> 'a1 -> 'a1 **)
-
- let min_case n0 m x x0 =
- min_case_strong n0 m (fun _ -> x) (fun _ -> x0)
-
- (** val min_dec : t -> t -> bool **)
-
- let min_dec =
- P.min_dec
end
-module N =
- struct
- type t = n
-
- (** val zero : n **)
-
- let zero =
- N0
-
- (** val one : n **)
-
- let one =
- Npos XH
-
- (** val two : n **)
-
- let two =
- Npos (XO XH)
-
- (** val succ_double : n -> n **)
-
- let succ_double = function
- | N0 -> Npos XH
- | Npos p -> Npos (XI p)
-
- (** val double : n -> n **)
-
- let double = function
- | N0 -> N0
- | Npos p -> Npos (XO p)
-
- (** val succ : n -> n **)
-
- let succ = function
- | N0 -> Npos XH
- | Npos p -> Npos (Coq_Pos.succ p)
-
- (** val pred : n -> n **)
-
- let pred = function
- | N0 -> N0
- | Npos p -> Coq_Pos.pred_N p
-
- (** val succ_pos : n -> positive **)
-
- let succ_pos = function
- | N0 -> XH
- | Npos p -> Coq_Pos.succ p
-
- (** val add : n -> n -> n **)
-
- let add n0 m =
- match n0 with
- | N0 -> m
- | Npos p ->
- (match m with
- | N0 -> n0
- | Npos q0 -> Npos (Coq_Pos.add p q0))
-
- (** val sub : n -> n -> n **)
-
- let sub n0 m =
- match n0 with
- | N0 -> N0
- | Npos n' ->
- (match m with
- | N0 -> n0
- | Npos m' ->
- (match Coq_Pos.sub_mask n' m' with
- | Coq_Pos.IsPos p -> Npos p
- | _ -> N0))
-
- (** val mul : n -> n -> n **)
-
- let mul n0 m =
- match n0 with
- | N0 -> N0
- | Npos p ->
- (match m with
- | N0 -> N0
- | Npos q0 -> Npos (Coq_Pos.mul p q0))
-
- (** val compare : n -> n -> comparison **)
-
- let compare n0 m =
- match n0 with
- | N0 ->
- (match m with
- | N0 -> Eq
- | Npos m' -> Lt)
- | Npos n' ->
- (match m with
- | N0 -> Gt
- | Npos m' -> Coq_Pos.compare n' m')
-
- (** val eqb : n -> n -> bool **)
-
- let eqb n0 m =
- match n0 with
- | N0 ->
- (match m with
- | N0 -> true
- | Npos p -> false)
- | Npos p ->
- (match m with
- | N0 -> false
- | Npos q0 -> Coq_Pos.eqb p q0)
-
- (** val leb : n -> n -> bool **)
-
- let leb x y =
- match compare x y with
- | Gt -> false
- | _ -> true
-
- (** val ltb : n -> n -> bool **)
-
- let ltb x y =
- match compare x y with
- | Lt -> true
- | _ -> false
-
- (** val min : n -> n -> n **)
-
- let min n0 n' =
- match compare n0 n' with
- | Gt -> n'
- | _ -> n0
-
- (** val max : n -> n -> n **)
-
- let max n0 n' =
- match compare n0 n' with
- | Gt -> n0
- | _ -> n'
-
- (** val div2 : n -> n **)
-
- let div2 = function
- | N0 -> N0
- | Npos p2 ->
- (match p2 with
- | XI p -> Npos p
- | XO p -> Npos p
- | XH -> N0)
-
- (** val even : n -> bool **)
-
- let even = function
- | N0 -> true
- | Npos p ->
- (match p with
- | XO p2 -> true
- | _ -> false)
-
- (** val odd : n -> bool **)
-
- let odd n0 =
- negb (even n0)
-
- (** val pow : n -> n -> n **)
-
- let pow n0 = function
- | N0 -> Npos XH
- | Npos p2 ->
- (match n0 with
- | N0 -> N0
- | Npos q0 -> Npos (Coq_Pos.pow q0 p2))
-
- (** val log2 : n -> n **)
-
- let log2 = function
- | N0 -> N0
- | Npos p2 ->
- (match p2 with
- | XI p -> Npos (Coq_Pos.size p)
- | XO p -> Npos (Coq_Pos.size p)
- | XH -> N0)
-
- (** val size : n -> n **)
-
- let size = function
- | N0 -> N0
- | Npos p -> Npos (Coq_Pos.size p)
-
- (** val size_nat : n -> nat **)
-
- let size_nat = function
- | N0 -> O
- | Npos p -> Coq_Pos.size_nat p
-
- (** val pos_div_eucl : positive -> n -> n * n **)
-
- let rec pos_div_eucl a b =
- match a with
- | XI a' ->
- let q0,r = pos_div_eucl a' b in
- let r' = succ_double r in
- if leb b r' then (succ_double q0),(sub r' b) else (double q0),r'
- | XO a' ->
- let q0,r = pos_div_eucl a' b in
- let r' = double r in
- if leb b r' then (succ_double q0),(sub r' b) else (double q0),r'
- | XH ->
- (match b with
- | N0 -> N0,(Npos XH)
- | Npos p ->
- (match p with
- | XH -> (Npos XH),N0
- | _ -> N0,(Npos XH)))
-
- (** val div_eucl : n -> n -> n * n **)
-
- let div_eucl a b =
- match a with
- | N0 -> N0,N0
- | Npos na ->
- (match b with
- | N0 -> N0,a
- | Npos p -> pos_div_eucl na b)
-
- (** val div : n -> n -> n **)
-
- let div a b =
- fst (div_eucl a b)
-
- (** val modulo : n -> n -> n **)
-
- let modulo a b =
- snd (div_eucl a b)
-
- (** val gcd : n -> n -> n **)
-
- let gcd a b =
- match a with
- | N0 -> b
- | Npos p ->
- (match b with
- | N0 -> a
- | Npos q0 -> Npos (Coq_Pos.gcd p q0))
-
- (** val ggcd : n -> n -> n * (n * n) **)
-
- let ggcd a b =
- match a with
- | N0 -> b,(N0,(Npos XH))
- | Npos p ->
- (match b with
- | N0 -> a,((Npos XH),N0)
- | Npos q0 ->
- let g,p2 = Coq_Pos.ggcd p q0 in
- let aa,bb = p2 in (Npos g),((Npos aa),(Npos bb)))
-
- (** val sqrtrem : n -> n * n **)
-
- let sqrtrem = function
- | N0 -> N0,N0
- | Npos p ->
- let s,m = Coq_Pos.sqrtrem p in
- (match m with
- | Coq_Pos.IsPos r -> (Npos s),(Npos r)
- | _ -> (Npos s),N0)
-
- (** val sqrt : n -> n **)
-
- let sqrt = function
- | N0 -> N0
- | Npos p -> Npos (Coq_Pos.sqrt p)
-
- (** val coq_lor : n -> n -> n **)
-
- let coq_lor n0 m =
- match n0 with
- | N0 -> m
- | Npos p ->
- (match m with
- | N0 -> n0
- | Npos q0 -> Npos (Coq_Pos.coq_lor p q0))
-
- (** val coq_land : n -> n -> n **)
-
- let coq_land n0 m =
- match n0 with
- | N0 -> N0
- | Npos p ->
- (match m with
- | N0 -> N0
- | Npos q0 -> Coq_Pos.coq_land p q0)
-
- (** val ldiff : n -> n -> n **)
-
- let ldiff n0 m =
- match n0 with
- | N0 -> N0
- | Npos p ->
- (match m with
- | N0 -> n0
- | Npos q0 -> Coq_Pos.ldiff p q0)
-
- (** val coq_lxor : n -> n -> n **)
-
- let coq_lxor n0 m =
- match n0 with
- | N0 -> m
- | Npos p ->
- (match m with
- | N0 -> n0
- | Npos q0 -> Coq_Pos.coq_lxor p q0)
-
- (** val shiftl_nat : n -> nat -> n **)
-
- let shiftl_nat a n0 =
- nat_iter n0 double a
-
- (** val shiftr_nat : n -> nat -> n **)
-
- let shiftr_nat a n0 =
- nat_iter n0 div2 a
-
- (** val shiftl : n -> n -> n **)
-
- let shiftl a n0 =
- match a with
- | N0 -> N0
- | Npos a0 -> Npos (Coq_Pos.shiftl a0 n0)
-
- (** val shiftr : n -> n -> n **)
-
- let shiftr a = function
- | N0 -> a
- | Npos p -> Coq_Pos.iter p div2 a
-
- (** val testbit_nat : n -> nat -> bool **)
-
- let testbit_nat = function
- | N0 -> (fun x -> false)
- | Npos p -> Coq_Pos.testbit_nat p
-
- (** val testbit : n -> n -> bool **)
-
- let testbit a n0 =
- match a with
- | N0 -> false
- | Npos p -> Coq_Pos.testbit p n0
-
- (** val to_nat : n -> nat **)
-
- let to_nat = function
- | N0 -> O
- | Npos p -> Coq_Pos.to_nat p
-
+module N =
+ struct
(** val of_nat : nat -> n **)
-
+
let of_nat = function
| O -> N0
| S n' -> Npos (Coq_Pos.of_succ_nat n')
-
- (** val iter : n -> ('a1 -> 'a1) -> 'a1 -> 'a1 **)
-
- let iter n0 f x =
- match n0 with
- | N0 -> x
- | Npos p -> Coq_Pos.iter p f x
-
- (** val eq_dec : n -> n -> bool **)
-
- let eq_dec n0 m =
- match n0 with
- | N0 ->
- (match m with
- | N0 -> true
- | Npos p -> false)
- | Npos x ->
- (match m with
- | N0 -> false
- | Npos p2 -> Coq_Pos.eq_dec x p2)
-
- (** val discr : n -> positive option **)
-
- let discr = function
- | N0 -> None
- | Npos p -> Some p
-
- (** val binary_rect :
- 'a1 -> (n -> 'a1 -> 'a1) -> (n -> 'a1 -> 'a1) -> n -> 'a1 **)
-
- let binary_rect f0 f2 fS2 n0 =
- let f2' = fun p -> f2 (Npos p) in
- let fS2' = fun p -> fS2 (Npos p) in
- (match n0 with
- | N0 -> f0
- | Npos p ->
- let rec f = function
- | XI p3 -> fS2' p3 (f p3)
- | XO p3 -> f2' p3 (f p3)
- | XH -> fS2 N0 f0
- in f p)
-
- (** val binary_rec :
- 'a1 -> (n -> 'a1 -> 'a1) -> (n -> 'a1 -> 'a1) -> n -> 'a1 **)
-
- let binary_rec =
- binary_rect
-
- (** val peano_rect : 'a1 -> (n -> 'a1 -> 'a1) -> n -> 'a1 **)
-
- let peano_rect f0 f n0 =
- let f' = fun p -> f (Npos p) in
- (match n0 with
- | N0 -> f0
- | Npos p -> Coq_Pos.peano_rect (f N0 f0) f' p)
-
- (** val peano_rec : 'a1 -> (n -> 'a1 -> 'a1) -> n -> 'a1 **)
-
- let peano_rec =
- peano_rect
-
- module BootStrap =
- struct
-
- end
-
- (** val recursion : 'a1 -> (n -> 'a1 -> 'a1) -> n -> 'a1 **)
-
- let recursion x =
- peano_rect x
-
- module OrderElts =
- struct
- type t = n
- end
-
- module OrderTac = MakeOrderTac(OrderElts)
-
- module NZPowP =
- struct
-
- end
-
- module NZSqrtP =
- struct
-
- end
-
- (** val sqrt_up : n -> n **)
-
- let sqrt_up a =
- match compare N0 a with
- | Lt -> succ (sqrt (pred a))
- | _ -> N0
-
- (** val log2_up : n -> n **)
-
- let log2_up a =
- match compare (Npos XH) a with
- | Lt -> succ (log2 (pred a))
- | _ -> N0
-
- module NZDivP =
- struct
-
- end
-
- (** val lcm : n -> n -> n **)
-
- let lcm a b =
- mul a (div b (gcd a b))
-
- (** val b2n : bool -> n **)
-
- let b2n = function
- | true -> Npos XH
- | false -> N0
-
- (** val setbit : n -> n -> n **)
-
- let setbit a n0 =
- coq_lor a (shiftl (Npos XH) n0)
-
- (** val clearbit : n -> n -> n **)
-
- let clearbit a n0 =
- ldiff a (shiftl (Npos XH) n0)
-
- (** val ones : n -> n **)
-
- let ones n0 =
- pred (shiftl (Npos XH) n0)
-
- (** val lnot : n -> n -> n **)
-
- let lnot a n0 =
- coq_lxor a (ones n0)
-
- module T =
- struct
-
- end
-
- module ORev =
- struct
- type t = n
- end
-
- module MRev =
- struct
- (** val max : n -> n -> n **)
-
- let max x y =
- min y x
- end
-
- module MPRev = MaxLogicalProperties(ORev)(MRev)
-
- module P =
- struct
- (** val max_case_strong :
- n -> n -> (n -> n -> __ -> 'a1 -> 'a1) -> (__ -> 'a1) -> (__ -> 'a1)
- -> 'a1 **)
-
- let max_case_strong n0 m compat hl hr =
- let c = compSpec2Type n0 m (compare n0 m) in
- (match c with
- | CompGtT -> compat n0 (max n0 m) __ (hl __)
- | _ -> compat m (max n0 m) __ (hr __))
-
- (** val max_case :
- n -> n -> (n -> n -> __ -> 'a1 -> 'a1) -> 'a1 -> 'a1 -> 'a1 **)
-
- let max_case n0 m x x0 x1 =
- max_case_strong n0 m x (fun _ -> x0) (fun _ -> x1)
-
- (** val max_dec : n -> n -> bool **)
-
- let max_dec n0 m =
- max_case n0 m (fun x y _ h0 -> h0) true false
-
- (** val min_case_strong :
- n -> n -> (n -> n -> __ -> 'a1 -> 'a1) -> (__ -> 'a1) -> (__ -> 'a1)
- -> 'a1 **)
-
- let min_case_strong n0 m compat hl hr =
- let c = compSpec2Type n0 m (compare n0 m) in
- (match c with
- | CompGtT -> compat m (min n0 m) __ (hr __)
- | _ -> compat n0 (min n0 m) __ (hl __))
-
- (** val min_case :
- n -> n -> (n -> n -> __ -> 'a1 -> 'a1) -> 'a1 -> 'a1 -> 'a1 **)
-
- let min_case n0 m x x0 x1 =
- min_case_strong n0 m x (fun _ -> x0) (fun _ -> x1)
-
- (** val min_dec : n -> n -> bool **)
-
- let min_dec n0 m =
- min_case n0 m (fun x y _ h0 -> h0) true false
- end
-
- (** val max_case_strong : n -> n -> (__ -> 'a1) -> (__ -> 'a1) -> 'a1 **)
-
- let max_case_strong n0 m x x0 =
- P.max_case_strong n0 m (fun x1 y _ x2 -> x2) x x0
-
- (** val max_case : n -> n -> 'a1 -> 'a1 -> 'a1 **)
-
- let max_case n0 m x x0 =
- max_case_strong n0 m (fun _ -> x) (fun _ -> x0)
-
- (** val max_dec : n -> n -> bool **)
-
- let max_dec =
- P.max_dec
-
- (** val min_case_strong : n -> n -> (__ -> 'a1) -> (__ -> 'a1) -> 'a1 **)
-
- let min_case_strong n0 m x x0 =
- P.min_case_strong n0 m (fun x1 y _ x2 -> x2) x x0
-
- (** val min_case : n -> n -> 'a1 -> 'a1 -> 'a1 **)
-
- let min_case n0 m x x0 =
- min_case_strong n0 m (fun _ -> x) (fun _ -> x0)
-
- (** val min_dec : n -> n -> bool **)
-
- let min_dec =
- P.min_dec
end
(** val pow_pos : ('a1 -> 'a1 -> 'a1) -> 'a1 -> positive -> 'a1 **)
@@ -2006,66 +280,49 @@ let rec nth n0 l default =
| O ->
(match l with
| [] -> default
- | x::l' -> x)
+ | x::_ -> x)
| S m ->
(match l with
| [] -> default
- | x::t1 -> nth m t1 default)
+ | _::t0 -> nth m t0 default)
(** val map : ('a1 -> 'a2) -> 'a1 list -> 'a2 list **)
let rec map f = function
| [] -> []
-| a::t1 -> (f a)::(map f t1)
+| a::t0 -> (f a)::(map f t0)
(** val fold_right : ('a2 -> 'a1 -> 'a1) -> 'a1 -> 'a2 list -> 'a1 **)
let rec fold_right f a0 = function
| [] -> a0
-| b::t1 -> f b (fold_right f a0 t1)
-
-module Z =
- struct
- type t = z
-
- (** val zero : z **)
-
- let zero =
- Z0
-
- (** val one : z **)
-
- let one =
- Zpos XH
-
- (** val two : z **)
-
- let two =
- Zpos (XO XH)
-
+| b::t0 -> f b (fold_right f a0 t0)
+
+module Z =
+ struct
(** val double : z -> z **)
-
+
let double = function
| Z0 -> Z0
| Zpos p -> Zpos (XO p)
| Zneg p -> Zneg (XO p)
-
+
(** val succ_double : z -> z **)
-
+
let succ_double = function
| Z0 -> Zpos XH
| Zpos p -> Zpos (XI p)
| Zneg p -> Zneg (Coq_Pos.pred_double p)
-
+
(** val pred_double : z -> z **)
-
+
let pred_double = function
| Z0 -> Zneg XH
| Zpos p -> Zpos (Coq_Pos.pred_double p)
| Zneg p -> Zneg (XI p)
-
+
(** val pos_sub : positive -> positive -> z **)
-
+
let rec pos_sub x y =
match x with
| XI p ->
@@ -2083,9 +340,9 @@ module Z =
| XI q0 -> Zneg (XO q0)
| XO q0 -> Zneg (Coq_Pos.pred_double q0)
| XH -> Z0)
-
+
(** val add : z -> z -> z **)
-
+
let add x y =
match x with
| Z0 -> y
@@ -2099,31 +356,21 @@ module Z =
| Z0 -> x
| Zpos y' -> pos_sub y' x'
| Zneg y' -> Zneg (Coq_Pos.add x' y'))
-
+
(** val opp : z -> z **)
-
+
let opp = function
| Z0 -> Z0
| Zpos x0 -> Zneg x0
| Zneg x0 -> Zpos x0
-
- (** val succ : z -> z **)
-
- let succ x =
- add x (Zpos XH)
-
- (** val pred : z -> z **)
-
- let pred x =
- add x (Zneg XH)
-
+
(** val sub : z -> z -> z **)
-
+
let sub m n0 =
add m (opp n0)
-
+
(** val mul : z -> z -> z **)
-
+
let mul x y =
match x with
| Z0 -> Z0
@@ -2137,28 +384,16 @@ module Z =
| Z0 -> Z0
| Zpos y' -> Zneg (Coq_Pos.mul x' y')
| Zneg y' -> Zpos (Coq_Pos.mul x' y'))
-
- (** val pow_pos : z -> positive -> z **)
-
- let pow_pos z0 n0 =
- Coq_Pos.iter n0 (mul z0) (Zpos XH)
-
- (** val pow : z -> z -> z **)
-
- let pow x = function
- | Z0 -> Zpos XH
- | Zpos p -> pow_pos x p
- | Zneg p -> Z0
-
+
(** val compare : z -> z -> comparison **)
-
+
let compare x y =
match x with
| Z0 ->
(match y with
| Z0 -> Eq
- | Zpos y' -> Lt
- | Zneg y' -> Gt)
+ | Zpos _ -> Lt
+ | Zneg _ -> Gt)
| Zpos x' ->
(match y with
| Zpos y' -> Coq_Pos.compare x' y'
@@ -2167,151 +402,74 @@ module Z =
(match y with
| Zneg y' -> compOpp (Coq_Pos.compare x' y')
| _ -> Lt)
-
- (** val sgn : z -> z **)
-
- let sgn = function
- | Z0 -> Z0
- | Zpos p -> Zpos XH
- | Zneg p -> Zneg XH
-
+
(** val leb : z -> z -> bool **)
-
+
let leb x y =
match compare x y with
| Gt -> false
| _ -> true
-
- (** val geb : z -> z -> bool **)
-
- let geb x y =
- match compare x y with
- | Lt -> false
- | _ -> true
-
+
(** val ltb : z -> z -> bool **)
-
+
let ltb x y =
match compare x y with
| Lt -> true
| _ -> false
-
+
(** val gtb : z -> z -> bool **)
-
+
let gtb x y =
match compare x y with
| Gt -> true
| _ -> false
-
- (** val eqb : z -> z -> bool **)
-
- let eqb x y =
- match x with
- | Z0 ->
- (match y with
- | Z0 -> true
- | _ -> false)
- | Zpos p ->
- (match y with
- | Zpos q0 -> Coq_Pos.eqb p q0
- | _ -> false)
- | Zneg p ->
- (match y with
- | Zneg q0 -> Coq_Pos.eqb p q0
- | _ -> false)
-
+
(** val max : z -> z -> z **)
-
+
let max n0 m =
match compare n0 m with
| Lt -> m
| _ -> n0
-
- (** val min : z -> z -> z **)
-
- let min n0 m =
- match compare n0 m with
- | Gt -> m
- | _ -> n0
-
+
(** val abs : z -> z **)
-
+
let abs = function
| Zneg p -> Zpos p
| x -> x
-
- (** val abs_nat : z -> nat **)
-
- let abs_nat = function
- | Z0 -> O
- | Zpos p -> Coq_Pos.to_nat p
- | Zneg p -> Coq_Pos.to_nat p
-
- (** val abs_N : z -> n **)
-
- let abs_N = function
- | Z0 -> N0
- | Zpos p -> Npos p
- | Zneg p -> Npos p
-
- (** val to_nat : z -> nat **)
-
- let to_nat = function
- | Zpos p -> Coq_Pos.to_nat p
- | _ -> O
-
+
(** val to_N : z -> n **)
-
+
let to_N = function
| Zpos p -> Npos p
| _ -> N0
-
- (** val of_nat : nat -> z **)
-
- let of_nat = function
- | O -> Z0
- | S n1 -> Zpos (Coq_Pos.of_succ_nat n1)
-
- (** val of_N : n -> z **)
-
- let of_N = function
- | N0 -> Z0
- | Npos p -> Zpos p
-
- (** val iter : z -> ('a1 -> 'a1) -> 'a1 -> 'a1 **)
-
- let iter n0 f x =
- match n0 with
- | Zpos p -> Coq_Pos.iter p f x
- | _ -> x
-
+
(** val pos_div_eucl : positive -> z -> z * z **)
-
+
let rec pos_div_eucl a b =
match a with
| XI a' ->
let q0,r = pos_div_eucl a' b in
let r' = add (mul (Zpos (XO XH)) r) (Zpos XH) in
- if gtb b r'
+ if ltb r' b
then (mul (Zpos (XO XH)) q0),r'
else (add (mul (Zpos (XO XH)) q0) (Zpos XH)),(sub r' b)
| XO a' ->
let q0,r = pos_div_eucl a' b in
let r' = mul (Zpos (XO XH)) r in
- if gtb b r'
+ if ltb r' b
then (mul (Zpos (XO XH)) q0),r'
else (add (mul (Zpos (XO XH)) q0) (Zpos XH)),(sub r' b)
- | XH -> if geb b (Zpos (XO XH)) then Z0,(Zpos XH) else (Zpos XH),Z0
-
+ | XH -> if leb (Zpos (XO XH)) b then Z0,(Zpos XH) else (Zpos XH),Z0
+
(** val div_eucl : z -> z -> z * z **)
-
+
let div_eucl a b =
match a with
| Z0 -> Z0,Z0
| Zpos a' ->
(match b with
| Z0 -> Z0,Z0
- | Zpos p -> pos_div_eucl a' b
+ | Zpos _ -> pos_div_eucl a' b
| Zneg b' ->
let q0,r = pos_div_eucl a' (Zpos b') in
(match r with
@@ -2320,131 +478,20 @@ module Z =
| Zneg a' ->
(match b with
| Z0 -> Z0,Z0
- | Zpos p ->
+ | Zpos _ ->
let q0,r = pos_div_eucl a' b in
(match r with
| Z0 -> (opp q0),Z0
| _ -> (opp (add q0 (Zpos XH))),(sub b r))
| Zneg b' -> let q0,r = pos_div_eucl a' (Zpos b') in q0,(opp r))
-
+
(** val div : z -> z -> z **)
-
+
let div a b =
- let q0,x = div_eucl a b in q0
-
- (** val modulo : z -> z -> z **)
-
- let modulo a b =
- let x,r = div_eucl a b in r
-
- (** val quotrem : z -> z -> z * z **)
-
- let quotrem a b =
- match a with
- | Z0 -> Z0,Z0
- | Zpos a0 ->
- (match b with
- | Z0 -> Z0,a
- | Zpos b0 ->
- let q0,r = N.pos_div_eucl a0 (Npos b0) in (of_N q0),(of_N r)
- | Zneg b0 ->
- let q0,r = N.pos_div_eucl a0 (Npos b0) in (opp (of_N q0)),(of_N r))
- | Zneg a0 ->
- (match b with
- | Z0 -> Z0,a
- | Zpos b0 ->
- let q0,r = N.pos_div_eucl a0 (Npos b0) in
- (opp (of_N q0)),(opp (of_N r))
- | Zneg b0 ->
- let q0,r = N.pos_div_eucl a0 (Npos b0) in (of_N q0),(opp (of_N r)))
-
- (** val quot : z -> z -> z **)
-
- let quot a b =
- fst (quotrem a b)
-
- (** val rem : z -> z -> z **)
-
- let rem a b =
- snd (quotrem a b)
-
- (** val even : z -> bool **)
-
- let even = function
- | Z0 -> true
- | Zpos p ->
- (match p with
- | XO p2 -> true
- | _ -> false)
- | Zneg p ->
- (match p with
- | XO p2 -> true
- | _ -> false)
-
- (** val odd : z -> bool **)
-
- let odd = function
- | Z0 -> false
- | Zpos p ->
- (match p with
- | XO p2 -> false
- | _ -> true)
- | Zneg p ->
- (match p with
- | XO p2 -> false
- | _ -> true)
-
- (** val div2 : z -> z **)
-
- let div2 = function
- | Z0 -> Z0
- | Zpos p ->
- (match p with
- | XH -> Z0
- | _ -> Zpos (Coq_Pos.div2 p))
- | Zneg p -> Zneg (Coq_Pos.div2_up p)
-
- (** val quot2 : z -> z **)
-
- let quot2 = function
- | Z0 -> Z0
- | Zpos p ->
- (match p with
- | XH -> Z0
- | _ -> Zpos (Coq_Pos.div2 p))
- | Zneg p ->
- (match p with
- | XH -> Z0
- | _ -> Zneg (Coq_Pos.div2 p))
-
- (** val log2 : z -> z **)
-
- let log2 = function
- | Zpos p2 ->
- (match p2 with
- | XI p -> Zpos (Coq_Pos.size p)
- | XO p -> Zpos (Coq_Pos.size p)
- | XH -> Z0)
- | _ -> Z0
-
- (** val sqrtrem : z -> z * z **)
-
- let sqrtrem = function
- | Zpos p ->
- let s,m = Coq_Pos.sqrtrem p in
- (match m with
- | Coq_Pos.IsPos r -> (Zpos s),(Zpos r)
- | _ -> (Zpos s),Z0)
- | _ -> Z0,Z0
-
- (** val sqrt : z -> z **)
-
- let sqrt = function
- | Zpos p -> Zpos (Coq_Pos.sqrt p)
- | _ -> Z0
-
+ let q0,_ = div_eucl a b in q0
+
(** val gcd : z -> z -> z **)
-
+
let gcd a b =
match a with
| Z0 -> abs b
@@ -2458,316 +505,6 @@ module Z =
| Z0 -> abs a
| Zpos b0 -> Zpos (Coq_Pos.gcd a0 b0)
| Zneg b0 -> Zpos (Coq_Pos.gcd a0 b0))
-
- (** val ggcd : z -> z -> z * (z * z) **)
-
- let ggcd a b =
- match a with
- | Z0 -> (abs b),(Z0,(sgn b))
- | Zpos a0 ->
- (match b with
- | Z0 -> (abs a),((sgn a),Z0)
- | Zpos b0 ->
- let g,p = Coq_Pos.ggcd a0 b0 in
- let aa,bb = p in (Zpos g),((Zpos aa),(Zpos bb))
- | Zneg b0 ->
- let g,p = Coq_Pos.ggcd a0 b0 in
- let aa,bb = p in (Zpos g),((Zpos aa),(Zneg bb)))
- | Zneg a0 ->
- (match b with
- | Z0 -> (abs a),((sgn a),Z0)
- | Zpos b0 ->
- let g,p = Coq_Pos.ggcd a0 b0 in
- let aa,bb = p in (Zpos g),((Zneg aa),(Zpos bb))
- | Zneg b0 ->
- let g,p = Coq_Pos.ggcd a0 b0 in
- let aa,bb = p in (Zpos g),((Zneg aa),(Zneg bb)))
-
- (** val testbit : z -> z -> bool **)
-
- let testbit a = function
- | Z0 -> odd a
- | Zpos p ->
- (match a with
- | Z0 -> false
- | Zpos a0 -> Coq_Pos.testbit a0 (Npos p)
- | Zneg a0 -> negb (N.testbit (Coq_Pos.pred_N a0) (Npos p)))
- | Zneg p -> false
-
- (** val shiftl : z -> z -> z **)
-
- let shiftl a = function
- | Z0 -> a
- | Zpos p -> Coq_Pos.iter p (mul (Zpos (XO XH))) a
- | Zneg p -> Coq_Pos.iter p div2 a
-
- (** val shiftr : z -> z -> z **)
-
- let shiftr a n0 =
- shiftl a (opp n0)
-
- (** val coq_lor : z -> z -> z **)
-
- let coq_lor a b =
- match a with
- | Z0 -> b
- | Zpos a0 ->
- (match b with
- | Z0 -> a
- | Zpos b0 -> Zpos (Coq_Pos.coq_lor a0 b0)
- | Zneg b0 -> Zneg (N.succ_pos (N.ldiff (Coq_Pos.pred_N b0) (Npos a0))))
- | Zneg a0 ->
- (match b with
- | Z0 -> a
- | Zpos b0 -> Zneg (N.succ_pos (N.ldiff (Coq_Pos.pred_N a0) (Npos b0)))
- | Zneg b0 ->
- Zneg
- (N.succ_pos (N.coq_land (Coq_Pos.pred_N a0) (Coq_Pos.pred_N b0))))
-
- (** val coq_land : z -> z -> z **)
-
- let coq_land a b =
- match a with
- | Z0 -> Z0
- | Zpos a0 ->
- (match b with
- | Z0 -> Z0
- | Zpos b0 -> of_N (Coq_Pos.coq_land a0 b0)
- | Zneg b0 -> of_N (N.ldiff (Npos a0) (Coq_Pos.pred_N b0)))
- | Zneg a0 ->
- (match b with
- | Z0 -> Z0
- | Zpos b0 -> of_N (N.ldiff (Npos b0) (Coq_Pos.pred_N a0))
- | Zneg b0 ->
- Zneg
- (N.succ_pos (N.coq_lor (Coq_Pos.pred_N a0) (Coq_Pos.pred_N b0))))
-
- (** val ldiff : z -> z -> z **)
-
- let ldiff a b =
- match a with
- | Z0 -> Z0
- | Zpos a0 ->
- (match b with
- | Z0 -> a
- | Zpos b0 -> of_N (Coq_Pos.ldiff a0 b0)
- | Zneg b0 -> of_N (N.coq_land (Npos a0) (Coq_Pos.pred_N b0)))
- | Zneg a0 ->
- (match b with
- | Z0 -> a
- | Zpos b0 ->
- Zneg (N.succ_pos (N.coq_lor (Coq_Pos.pred_N a0) (Npos b0)))
- | Zneg b0 -> of_N (N.ldiff (Coq_Pos.pred_N b0) (Coq_Pos.pred_N a0)))
-
- (** val coq_lxor : z -> z -> z **)
-
- let coq_lxor a b =
- match a with
- | Z0 -> b
- | Zpos a0 ->
- (match b with
- | Z0 -> a
- | Zpos b0 -> of_N (Coq_Pos.coq_lxor a0 b0)
- | Zneg b0 ->
- Zneg (N.succ_pos (N.coq_lxor (Npos a0) (Coq_Pos.pred_N b0))))
- | Zneg a0 ->
- (match b with
- | Z0 -> a
- | Zpos b0 ->
- Zneg (N.succ_pos (N.coq_lxor (Coq_Pos.pred_N a0) (Npos b0)))
- | Zneg b0 -> of_N (N.coq_lxor (Coq_Pos.pred_N a0) (Coq_Pos.pred_N b0)))
-
- (** val eq_dec : z -> z -> bool **)
-
- let eq_dec x y =
- match x with
- | Z0 ->
- (match y with
- | Z0 -> true
- | _ -> false)
- | Zpos x0 ->
- (match y with
- | Zpos p2 -> Coq_Pos.eq_dec x0 p2
- | _ -> false)
- | Zneg x0 ->
- (match y with
- | Zneg p2 -> Coq_Pos.eq_dec x0 p2
- | _ -> false)
-
- module BootStrap =
- struct
-
- end
-
- module OrderElts =
- struct
- type t = z
- end
-
- module OrderTac = MakeOrderTac(OrderElts)
-
- (** val sqrt_up : z -> z **)
-
- let sqrt_up a =
- match compare Z0 a with
- | Lt -> succ (sqrt (pred a))
- | _ -> Z0
-
- (** val log2_up : z -> z **)
-
- let log2_up a =
- match compare (Zpos XH) a with
- | Lt -> succ (log2 (pred a))
- | _ -> Z0
-
- module NZDivP =
- struct
-
- end
-
- module Quot2Div =
- struct
- (** val div : z -> z -> z **)
-
- let div =
- quot
-
- (** val modulo : z -> z -> z **)
-
- let modulo =
- rem
- end
-
- module NZQuot =
- struct
-
- end
-
- (** val lcm : z -> z -> z **)
-
- let lcm a b =
- abs (mul a (div b (gcd a b)))
-
- (** val b2z : bool -> z **)
-
- let b2z = function
- | true -> Zpos XH
- | false -> Z0
-
- (** val setbit : z -> z -> z **)
-
- let setbit a n0 =
- coq_lor a (shiftl (Zpos XH) n0)
-
- (** val clearbit : z -> z -> z **)
-
- let clearbit a n0 =
- ldiff a (shiftl (Zpos XH) n0)
-
- (** val lnot : z -> z **)
-
- let lnot a =
- pred (opp a)
-
- (** val ones : z -> z **)
-
- let ones n0 =
- pred (shiftl (Zpos XH) n0)
-
- module T =
- struct
-
- end
-
- module ORev =
- struct
- type t = z
- end
-
- module MRev =
- struct
- (** val max : z -> z -> z **)
-
- let max x y =
- min y x
- end
-
- module MPRev = MaxLogicalProperties(ORev)(MRev)
-
- module P =
- struct
- (** val max_case_strong :
- z -> z -> (z -> z -> __ -> 'a1 -> 'a1) -> (__ -> 'a1) -> (__ -> 'a1)
- -> 'a1 **)
-
- let max_case_strong n0 m compat hl hr =
- let c = compSpec2Type n0 m (compare n0 m) in
- (match c with
- | CompGtT -> compat n0 (max n0 m) __ (hl __)
- | _ -> compat m (max n0 m) __ (hr __))
-
- (** val max_case :
- z -> z -> (z -> z -> __ -> 'a1 -> 'a1) -> 'a1 -> 'a1 -> 'a1 **)
-
- let max_case n0 m x x0 x1 =
- max_case_strong n0 m x (fun _ -> x0) (fun _ -> x1)
-
- (** val max_dec : z -> z -> bool **)
-
- let max_dec n0 m =
- max_case n0 m (fun x y _ h0 -> h0) true false
-
- (** val min_case_strong :
- z -> z -> (z -> z -> __ -> 'a1 -> 'a1) -> (__ -> 'a1) -> (__ -> 'a1)
- -> 'a1 **)
-
- let min_case_strong n0 m compat hl hr =
- let c = compSpec2Type n0 m (compare n0 m) in
- (match c with
- | CompGtT -> compat m (min n0 m) __ (hr __)
- | _ -> compat n0 (min n0 m) __ (hl __))
-
- (** val min_case :
- z -> z -> (z -> z -> __ -> 'a1 -> 'a1) -> 'a1 -> 'a1 -> 'a1 **)
-
- let min_case n0 m x x0 x1 =
- min_case_strong n0 m x (fun _ -> x0) (fun _ -> x1)
-
- (** val min_dec : z -> z -> bool **)
-
- let min_dec n0 m =
- min_case n0 m (fun x y _ h0 -> h0) true false
- end
-
- (** val max_case_strong : z -> z -> (__ -> 'a1) -> (__ -> 'a1) -> 'a1 **)
-
- let max_case_strong n0 m x x0 =
- P.max_case_strong n0 m (fun x1 y _ x2 -> x2) x x0
-
- (** val max_case : z -> z -> 'a1 -> 'a1 -> 'a1 **)
-
- let max_case n0 m x x0 =
- max_case_strong n0 m (fun _ -> x) (fun _ -> x0)
-
- (** val max_dec : z -> z -> bool **)
-
- let max_dec =
- P.max_dec
-
- (** val min_case_strong : z -> z -> (__ -> 'a1) -> (__ -> 'a1) -> 'a1 **)
-
- let min_case_strong n0 m x x0 =
- P.min_case_strong n0 m (fun x1 y _ x2 -> x2) x x0
-
- (** val min_case : z -> z -> 'a1 -> 'a1 -> 'a1 **)
-
- let min_case n0 m x x0 =
- min_case_strong n0 m (fun _ -> x) (fun _ -> x0)
-
- (** val min_dec : z -> z -> bool **)
-
- let min_dec =
- P.min_dec
end
(** val zeq_bool : z -> z -> bool **)
@@ -2818,9 +555,9 @@ let rec peq ceqb p p' =
(** val mkPinj : positive -> 'a1 pol -> 'a1 pol **)
let mkPinj j p = match p with
-| Pc c -> p
+| Pc _ -> p
| Pinj (j', q0) -> Pinj ((Coq_Pos.add j j'), q0)
-| PX (p2, p3, p4) -> Pinj (j, p)
+| PX (_, _, _) -> Pinj (j, p)
(** val mkPinj_pred : positive -> 'a1 pol -> 'a1 pol **)
@@ -2831,12 +568,13 @@ let mkPinj_pred j p =
| XH -> p
(** val mkPX :
- 'a1 -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol **)
+ 'a1 -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> positive -> 'a1 pol -> 'a1
+ pol **)
let mkPX cO ceqb p i q0 =
match p with
| Pc c -> if ceqb c cO then mkPinj XH q0 else PX (p, i, q0)
- | Pinj (p2, p3) -> PX (p, i, q0)
+ | Pinj (_, _) -> PX (p, i, q0)
| PX (p', i', q') ->
if peq ceqb q' (p0 cO)
then PX (p', (Coq_Pos.add i' i), q0)
@@ -2893,8 +631,8 @@ let rec paddI cadd pop q0 j = function
| 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 **)
+ ('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 -> mkPinj j (paddC cadd (popp copp q0) c)
@@ -2911,11 +649,11 @@ let rec psubI cadd copp pop q0 j = function
| 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 **)
+ '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)
+| Pc _ -> PX (p', i', p)
| Pinj (j, q') ->
(match j with
| XI j0 -> PX (p', i', (Pinj ((XO j0), q')))
@@ -2928,15 +666,16 @@ let rec paddX cO ceqb pop p' i' p = match p with
| 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 **)
+ '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)
+| Pc _ -> 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 ((Coq_Pos.pred_double j0), q')))
+ | XO j0 ->
+ PX ((popp copp p'), i', (Pinj ((Coq_Pos.pred_double j0), q')))
| XH -> PX ((popp copp p'), i', q'))
| PX (p2, i, q') ->
(match Z.pos_sub i i' with
@@ -2945,8 +684,8 @@ let rec psubX cO copp ceqb pop p' i' p = match p with
| 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 **)
+ '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'
@@ -2964,7 +703,8 @@ let rec padd cO cadd ceqb p = function
| PX (p2, i, q0) ->
(match Z.pos_sub i i' with
| Z0 ->
- mkPX cO ceqb (padd cO cadd ceqb p2 p'0) i (padd cO cadd ceqb q0 q')
+ 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')
@@ -2973,8 +713,8 @@ let rec padd cO cadd ceqb p = function
(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 **)
+ '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'
@@ -2989,25 +729,27 @@ let rec psub cO cadd csub copp ceqb p = function
(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 ((Coq_Pos.pred_double j0), q0))
- q'))
- | XH -> PX ((popp copp p'0), i', (psub cO cadd csub copp ceqb q0 q')))
+ (psub cO cadd csub copp ceqb (Pinj ((Coq_Pos.pred_double j0),
+ q0)) q'))
+ | XH ->
+ PX ((popp copp p'0), i', (psub cO cadd csub copp ceqb q0 q')))
| PX (p2, i, q0) ->
(match Z.pos_sub 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')
+ 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 (psub cO cadd csub copp ceqb) 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 **)
+ 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1
+ -> 'a1 pol **)
let rec pmulC_aux cO cmul ceqb p c =
match p with
@@ -3018,8 +760,8 @@ let rec pmulC_aux cO cmul ceqb p c =
(pmulC_aux cO cmul ceqb q0 c)
(** val pmulC :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol ->
- 'a1 -> 'a1 pol **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol
+ -> 'a1 -> 'a1 pol **)
let pmulC cO cI cmul ceqb p c =
if ceqb c cO
@@ -3027,8 +769,8 @@ let pmulC cO cI cmul ceqb p c =
else if ceqb c cI then p else 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 **)
+ '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 -> mkPinj j (pmulC cO cI cmul ceqb q0 c)
@@ -3049,12 +791,13 @@ let rec pmulI cO cI cmul ceqb pmul0 q0 j = function
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 **)
+ '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 (pmul cO cI cadd cmul ceqb) q' j' p
+| Pinj (j', q') ->
+ pmulI cO cI cmul ceqb (pmul cO cI cadd cmul ceqb) q' j' p
| PX (p', i', q') ->
(match p with
| Pc c -> pmulC cO cI cmul ceqb p'' c
@@ -3063,22 +806,24 @@ let rec pmul cO cI cadd cmul ceqb p p'' = match p'' with
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 ((Coq_Pos.pred_double j0), q0)) q'
+ pmul cO cI cadd cmul ceqb (Pinj ((Coq_Pos.pred_double j0), q0))
+ q'
| XH -> pmul cO cI cadd cmul ceqb q0 q'
in
mkPX cO ceqb (pmul cO cI cadd cmul ceqb p p') i' qQ'
| PX (p2, i, q0) ->
let qQ' = pmul cO cI cadd cmul ceqb q0 q' in
- let pQ' = pmulI cO cI cmul ceqb (pmul cO cI cadd cmul ceqb) q' XH p2 in
+ let pQ' = pmulI cO cI cmul ceqb (pmul cO cI cadd cmul ceqb) q' XH p2
+ in
let qP' = pmul cO cI cadd cmul ceqb (mkPinj XH q0) p' in
let pP' = pmul cO cI cadd cmul ceqb p2 p' in
padd cO cadd ceqb
- (mkPX cO ceqb (padd cO cadd ceqb (mkPX cO ceqb pP' i (p0 cO)) qP') i'
- (p0 cO)) (mkPX cO ceqb pQ' i qQ'))
+ (mkPX cO ceqb (padd cO cadd ceqb (mkPX cO ceqb pP' i (p0 cO)) qP')
+ i' (p0 cO)) (mkPX cO ceqb pQ' i qQ'))
(** val psquare :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> bool) -> 'a1 pol -> 'a1 pol **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
+ 'a1 -> bool) -> 'a1 pol -> 'a1 pol **)
let rec psquare cO cI cadd cmul ceqb = function
| Pc c -> Pc (cmul c c)
@@ -3107,9 +852,9 @@ 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 **)
+ '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 ->
@@ -3123,16 +868,17 @@ let rec ppow_pos cO cI cadd cmul ceqb subst_l res p = function
| 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 **)
+ '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 **)
+ '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
@@ -3153,7 +899,8 @@ let rec norm_aux cO cI cadd cmul csub copp ceqb = function
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)
+ 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)
@@ -3185,9 +932,9 @@ let rec map_bformula fct = function
| N f0 -> N (map_bformula fct f0)
| I (f1, f2) -> I ((map_bformula fct f1), (map_bformula fct f2))
-type 'term' clause = 'term' list
+type 'x clause = 'x list
-type 'term' cnf = 'term' clause list
+type 'x cnf = 'x clause list
(** val tt : 'a1 cnf **)
@@ -3200,52 +947,52 @@ let ff =
[]::[]
(** val add_term :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 -> 'a1 clause -> 'a1
- clause option **)
+ ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 -> 'a1 clause ->
+ 'a1 clause option **)
-let rec add_term unsat deduce t1 = function
+let rec add_term unsat deduce t0 = function
| [] ->
- (match deduce t1 t1 with
- | Some u -> if unsat u then None else Some (t1::[])
- | None -> Some (t1::[]))
+ (match deduce t0 t0 with
+ | Some u -> if unsat u then None else Some (t0::[])
+ | None -> Some (t0::[]))
| t'::cl0 ->
- (match deduce t1 t' with
+ (match deduce t0 t' with
| Some u ->
if unsat u
then None
- else (match add_term unsat deduce t1 cl0 with
+ else (match add_term unsat deduce t0 cl0 with
| Some cl' -> Some (t'::cl')
| None -> None)
| None ->
- (match add_term unsat deduce t1 cl0 with
+ (match add_term unsat deduce t0 cl0 with
| Some cl' -> Some (t'::cl')
| None -> None))
(** val or_clause :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 clause -> 'a1 clause
- -> 'a1 clause option **)
+ ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 clause -> 'a1
+ clause -> 'a1 clause option **)
let rec or_clause unsat deduce cl1 cl2 =
match cl1 with
| [] -> Some cl2
- | t1::cl ->
- (match add_term unsat deduce t1 cl2 with
+ | t0::cl ->
+ (match add_term unsat deduce t0 cl2 with
| Some cl' -> or_clause unsat deduce cl cl'
| None -> None)
(** val or_clause_cnf :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 clause -> 'a1 cnf ->
- 'a1 cnf **)
+ ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 clause -> 'a1 cnf
+ -> 'a1 cnf **)
-let or_clause_cnf unsat deduce t1 f =
+let or_clause_cnf unsat deduce t0 f =
fold_right (fun e acc ->
- match or_clause unsat deduce t1 e with
+ match or_clause unsat deduce t0 e with
| Some cl -> cl::acc
| None -> acc) [] f
(** val or_cnf :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 cnf -> 'a1 cnf -> 'a1
- cnf **)
+ ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 cnf -> 'a1 cnf ->
+ 'a1 cnf **)
let rec or_cnf unsat deduce f f' =
match f with
@@ -3259,8 +1006,8 @@ let and_cnf f1 f2 =
app f1 f2
(** val xcnf :
- ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a2 cnf) -> ('a1
- -> 'a2 cnf) -> bool -> 'a1 bFormula -> 'a2 cnf **)
+ ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a2 cnf) ->
+ ('a1 -> 'a2 cnf) -> bool -> 'a1 bFormula -> 'a2 cnf **)
let rec xcnf unsat deduce normalise0 negate0 pol0 = function
| TT -> if pol0 then tt else ff
@@ -3300,9 +1047,9 @@ let rec cnf_checker checker f l =
| c::l0 -> if checker e c then cnf_checker checker f0 l0 else false)
(** val tauto_checker :
- ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a2 cnf) -> ('a1
- -> 'a2 cnf) -> ('a2 list -> 'a3 -> bool) -> 'a1 bFormula -> 'a3 list ->
- bool **)
+ ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a2 cnf) ->
+ ('a1 -> 'a2 cnf) -> ('a2 list -> 'a3 -> bool) -> 'a1 bFormula -> 'a3
+ list -> bool **)
let tauto_checker unsat deduce normalise0 negate0 checker f w =
cnf_checker checker (xcnf unsat deduce normalise0 negate0 true f) w
@@ -3335,18 +1082,18 @@ let opMult o o' =
| Equal -> Some Equal
| NonEqual ->
(match o' with
- | Strict -> None
- | NonStrict -> None
- | x -> Some x)
+ | Equal -> Some Equal
+ | NonEqual -> Some NonEqual
+ | _ -> None)
| Strict ->
(match o' with
| NonEqual -> None
| _ -> Some o')
| NonStrict ->
(match o' with
+ | Equal -> Some Equal
| NonEqual -> None
- | Strict -> Some NonStrict
- | x -> Some x)
+ | _ -> Some NonStrict)
(** val opAdd : op1 -> op1 -> op1 option **)
@@ -3394,8 +1141,8 @@ let map_option2 f o o' =
| None -> None
(** val pexpr_times_nformula :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> bool) -> 'a1 polC -> 'a1 nFormula -> 'a1 nFormula option **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
+ 'a1 -> bool) -> 'a1 polC -> 'a1 nFormula -> 'a1 nFormula option **)
let pexpr_times_nformula cO cI cplus ctimes ceqb e = function
| ef,o ->
@@ -3404,8 +1151,8 @@ let pexpr_times_nformula cO cI cplus ctimes ceqb e = function
| _ -> None)
(** val nformula_times_nformula :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> bool) -> 'a1 nFormula -> 'a1 nFormula -> 'a1 nFormula option **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
+ 'a1 -> bool) -> 'a1 nFormula -> 'a1 nFormula -> 'a1 nFormula option **)
let nformula_times_nformula cO cI cplus ctimes ceqb f1 f2 =
let e1,o1 = f1 in
@@ -3414,8 +1161,8 @@ let nformula_times_nformula cO cI cplus ctimes ceqb f1 f2 =
(opMult o1 o2)
(** val nformula_plus_nformula :
- 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula -> 'a1
- nFormula -> 'a1 nFormula option **)
+ 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula ->
+ 'a1 nFormula -> 'a1 nFormula option **)
let nformula_plus_nformula cO cplus ceqb f1 f2 =
let e1,o1 = f1 in
@@ -3423,9 +1170,9 @@ let nformula_plus_nformula cO cplus ceqb f1 f2 =
map_option (fun x -> Some ((padd cO cplus ceqb e1 e2),x)) (opAdd o1 o2)
(** val eval_Psatz :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz -> 'a1
- nFormula option **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
+ 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz
+ -> 'a1 nFormula option **)
let rec eval_Psatz cO cI cplus ctimes ceqb cleb l = function
| PsatzIn n0 -> Some (nth n0 l ((Pc cO),Equal))
@@ -3460,9 +1207,9 @@ let check_inconsistent cO ceqb cleb = function
| _ -> false)
(** val check_normalised_formulas :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz ->
- bool **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
+ 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz
+ -> bool **)
let check_normalised_formulas cO cI cplus ctimes ceqb cleb l cm =
match eval_Psatz cO cI cplus ctimes ceqb cleb l cm with
@@ -3479,51 +1226,41 @@ type op2 =
type 't formula = { flhs : 't pExpr; fop : op2; frhs : 't pExpr }
-(** val flhs : 'a1 formula -> 'a1 pExpr **)
-
-let flhs x = x.flhs
-
-(** val fop : 'a1 formula -> op2 **)
-
-let fop x = x.fop
-
-(** val frhs : 'a1 formula -> 'a1 pExpr **)
-
-let frhs x = x.frhs
-
(** val norm :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExpr -> 'a1 pol **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
+ 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExpr ->
+ 'a1 pol **)
let norm cO cI cplus ctimes cminus copp ceqb =
norm_aux cO cI cplus ctimes cminus copp ceqb
(** val psub0 :
- 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1
- -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol **)
+ 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) ->
+ ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol **)
let psub0 cO cplus cminus copp ceqb =
psub cO cplus cminus copp ceqb
(** val padd0 :
- 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol
- -> 'a1 pol **)
+ 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1
+ pol -> 'a1 pol **)
let padd0 cO cplus ceqb =
padd cO cplus ceqb
(** val xnormalise :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1
- nFormula list **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
+ 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula ->
+ 'a1 nFormula list **)
-let xnormalise cO cI cplus ctimes cminus copp ceqb t1 =
- let { flhs = lhs; fop = o; frhs = rhs } = t1 in
+let xnormalise cO cI cplus ctimes cminus copp ceqb t0 =
+ let { flhs = lhs; fop = o; frhs = rhs } = t0 in
let lhs0 = norm cO cI cplus ctimes cminus copp ceqb lhs in
let rhs0 = norm cO cI cplus ctimes cminus copp ceqb rhs in
(match o with
| OpEq ->
- ((psub0 cO cplus cminus copp ceqb lhs0 rhs0),Strict)::(((psub0 cO cplus
+ ((psub0 cO cplus cminus copp ceqb lhs0 rhs0),Strict)::(((psub0 cO
+ cplus
cminus copp
ceqb rhs0
lhs0),Strict)::[])
@@ -3534,26 +1271,27 @@ let xnormalise cO cI cplus ctimes cminus copp ceqb t1 =
| OpGt -> ((psub0 cO cplus cminus copp ceqb rhs0 lhs0),NonStrict)::[])
(** val cnf_normalise :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1
- nFormula cnf **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
+ 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula ->
+ 'a1 nFormula cnf **)
-let cnf_normalise cO cI cplus ctimes cminus copp ceqb t1 =
- map (fun x -> x::[]) (xnormalise cO cI cplus ctimes cminus copp ceqb t1)
+let cnf_normalise cO cI cplus ctimes cminus copp ceqb t0 =
+ map (fun x -> x::[]) (xnormalise cO cI cplus ctimes cminus copp ceqb t0)
(** val xnegate :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1
- nFormula list **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
+ 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula ->
+ 'a1 nFormula list **)
-let xnegate cO cI cplus ctimes cminus copp ceqb t1 =
- let { flhs = lhs; fop = o; frhs = rhs } = t1 in
+let xnegate cO cI cplus ctimes cminus copp ceqb t0 =
+ let { flhs = lhs; fop = o; frhs = rhs } = t0 in
let lhs0 = norm cO cI cplus ctimes cminus copp ceqb lhs in
let rhs0 = norm cO cI cplus ctimes cminus copp ceqb rhs in
(match o with
| OpEq -> ((psub0 cO cplus cminus copp ceqb lhs0 rhs0),Equal)::[]
| OpNEq ->
- ((psub0 cO cplus cminus copp ceqb lhs0 rhs0),Strict)::(((psub0 cO cplus
+ ((psub0 cO cplus cminus copp ceqb lhs0 rhs0),Strict)::(((psub0 cO
+ cplus
cminus copp
ceqb rhs0
lhs0),Strict)::[])
@@ -3563,12 +1301,12 @@ let xnegate cO cI cplus ctimes cminus copp ceqb t1 =
| OpGt -> ((psub0 cO cplus cminus copp ceqb lhs0 rhs0),Strict)::[])
(** val cnf_negate :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
- -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1
- nFormula cnf **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
+ 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula ->
+ 'a1 nFormula cnf **)
-let cnf_negate cO cI cplus ctimes cminus copp ceqb t1 =
- map (fun x -> x::[]) (xnegate cO cI cplus ctimes cminus copp ceqb t1)
+let cnf_negate cO cI cplus ctimes cminus copp ceqb t0 =
+ map (fun x -> x::[]) (xnegate cO cI cplus ctimes cminus copp ceqb t0)
(** val xdenorm : positive -> 'a1 pol -> 'a1 pExpr **)
@@ -3602,14 +1340,14 @@ let map_Formula c_of_S f =
{ flhs = (map_PExpr c_of_S l); fop = o; frhs = (map_PExpr c_of_S r) }
(** val simpl_cone :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 psatz ->
- 'a1 psatz **)
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 psatz
+ -> 'a1 psatz **)
let simpl_cone cO cI ctimes ceqb e = match e with
-| PsatzSquare t1 ->
- (match t1 with
+| PsatzSquare t0 ->
+ (match t0 with
| Pc c -> if ceqb cO c then PsatzZ else PsatzC (ctimes c c)
- | _ -> PsatzSquare t1)
+ | _ -> PsatzSquare t0)
| PsatzMulE (t1, t2) ->
(match t1 with
| PsatzMulE (x, x0) ->
@@ -3641,7 +1379,8 @@ let simpl_cone cO cI ctimes ceqb e = match e with
| PsatzC p2 -> PsatzMulE ((PsatzC (ctimes c p2)), x)
| _ -> if ceqb cI c then t2 else PsatzMulE (t1, t2)))
| PsatzAdd (y, z0) ->
- PsatzAdd ((PsatzMulE ((PsatzC c), y)), (PsatzMulE ((PsatzC c), z0)))
+ PsatzAdd ((PsatzMulE ((PsatzC c), y)), (PsatzMulE ((PsatzC c),
+ z0)))
| PsatzC c0 -> PsatzC (ctimes c c0)
| PsatzZ -> PsatzZ
| _ -> if ceqb cI c then t2 else PsatzMulE (t1, t2))
@@ -3683,7 +1422,8 @@ let qle_bool x y =
(** val qplus : q -> q -> q **)
let qplus x y =
- { qnum = (Z.add (Z.mul x.qnum (Zpos y.qden)) (Z.mul y.qnum (Zpos x.qden)));
+ { qnum =
+ (Z.add (Z.mul x.qnum (Zpos y.qden)) (Z.mul y.qnum (Zpos x.qden)));
qden = (Coq_Pos.mul x.qden y.qden) }
(** val qmult : q -> q -> q **)
@@ -3711,8 +1451,8 @@ let qinv x =
(** val qpower_positive : q -> positive -> q **)
-let qpower_positive q0 p =
- pow_pos qmult q0 p
+let qpower_positive =
+ pow_pos qmult
(** val qpower : q -> z -> q **)
@@ -3721,12 +1461,12 @@ let qpower q0 = function
| Zpos p -> qpower_positive q0 p
| Zneg p -> qinv (qpower_positive q0 p)
-type 'a t0 =
+type 'a t =
| Empty
| Leaf of 'a
-| Node of 'a t0 * 'a * 'a t0
+| Node of 'a t * 'a * 'a t
-(** val find : 'a1 -> 'a1 t0 -> positive -> 'a1 **)
+(** val find : 'a1 -> 'a1 t -> positive -> 'a1 **)
let rec find default vm p =
match vm with
@@ -3738,6 +1478,29 @@ let rec find default vm p =
| XO p2 -> find default l p2
| XH -> e)
+(** val singleton : 'a1 -> positive -> 'a1 -> 'a1 t **)
+
+let rec singleton default x v =
+ match x with
+ | XI p -> Node (Empty, default, (singleton default p v))
+ | XO p -> Node ((singleton default p v), default, Empty)
+ | XH -> Leaf v
+
+(** val vm_add : 'a1 -> positive -> 'a1 -> 'a1 t -> 'a1 t **)
+
+let rec vm_add default x v = function
+| Empty -> singleton default x v
+| Leaf vl ->
+ (match x with
+ | XI p -> Node (Empty, vl, (singleton default p v))
+ | XO p -> Node ((singleton default p v), vl, Empty)
+ | XH -> Leaf v)
+| Node (l, o, r) ->
+ (match x with
+ | XI p -> Node (l, o, (vm_add default p v r))
+ | XO p -> Node ((vm_add default p v l), o, r)
+ | XH -> Node (l, v, r))
+
type zWitness = z psatz
(** val zWeakChecker : z nFormula list -> z psatz -> bool **)
@@ -3762,8 +1525,8 @@ let norm0 =
(** val xnormalise0 : z formula -> z nFormula list **)
-let xnormalise0 t1 =
- let { flhs = lhs; fop = o; frhs = rhs } = t1 in
+let xnormalise0 t0 =
+ let { flhs = lhs; fop = o; frhs = rhs } = t0 in
let lhs0 = norm0 lhs in
let rhs0 = norm0 rhs in
(match o with
@@ -3780,13 +1543,13 @@ let xnormalise0 t1 =
(** val normalise : z formula -> z nFormula cnf **)
-let normalise t1 =
- map (fun x -> x::[]) (xnormalise0 t1)
+let normalise t0 =
+ map (fun x -> x::[]) (xnormalise0 t0)
(** val xnegate0 : z formula -> z nFormula list **)
-let xnegate0 t1 =
- let { flhs = lhs; fop = o; frhs = rhs } = t1 in
+let xnegate0 t0 =
+ let { flhs = lhs; fop = o; frhs = rhs } = t0 in
let lhs0 = norm0 lhs in
let rhs0 = norm0 rhs in
(match o with
@@ -3803,8 +1566,8 @@ let xnegate0 t1 =
(** val negate : z formula -> z nFormula cnf **)
-let negate t1 =
- map (fun x -> x::[]) (xnegate0 t1)
+let negate t0 =
+ map (fun x -> x::[]) (xnegate0 t0)
(** val zunsat : z nFormula -> bool **)
@@ -3839,8 +1602,8 @@ let zgcdM x y =
let rec zgcd_pol = function
| Pc c -> Z0,c
-| Pinj (p2, p3) -> zgcd_pol p3
-| PX (p2, p3, q0) ->
+| Pinj (_, p2) -> zgcd_pol p2
+| PX (p2, _, q0) ->
let g1,c1 = zgcd_pol p2 in
let g2,c2 = zgcd_pol q0 in (zgcdM (zgcdM g1 c1) g2),c2
@@ -3872,7 +1635,8 @@ let genCuttingPlane = function
then None
else Some ((makeCuttingPlane e),Equal)
| NonEqual -> Some ((e,Z0),op)
- | Strict -> Some ((makeCuttingPlane (psubC Z.sub e (Zpos XH))),NonStrict)
+ | Strict ->
+ Some ((makeCuttingPlane (psubC Z.sub e (Zpos XH))),NonStrict)
| NonStrict -> Some ((makeCuttingPlane e),NonStrict))
(** val nformula_of_cutting_plane : ((z polC * z) * op1) -> z nFormula **)
@@ -3966,8 +1730,8 @@ let qnormalise =
(** val qnegate : q formula -> q nFormula cnf **)
let qnegate =
- cnf_negate { qnum = Z0; qden = XH } { qnum = (Zpos XH); qden = XH } qplus
- qmult qminus qopp qeq_bool
+ cnf_negate { qnum = Z0; qden = XH } { qnum = (Zpos XH); qden = XH }
+ qplus qmult qminus qopp qeq_bool
(** val qunsat : q nFormula -> bool **)
@@ -4025,8 +1789,8 @@ let rnormalise =
(** val rnegate : q formula -> q nFormula cnf **)
let rnegate =
- cnf_negate { qnum = Z0; qden = XH } { qnum = (Zpos XH); qden = XH } qplus
- qmult qminus qopp qeq_bool
+ cnf_negate { qnum = Z0; qden = XH } { qnum = (Zpos XH); qden = XH }
+ qplus qmult qminus qopp qeq_bool
(** val runsat : q nFormula -> bool **)
@@ -4043,4 +1807,3 @@ let rdeduce =
let rTautoChecker f w =
tauto_checker runsat rdeduce rnormalise rnegate rWeakChecker
(map_bformula (map_Formula q_of_Rcst) f) w
-
diff --git a/plugins/micromega/micromega.mli b/plugins/micromega/micromega.mli
index bcd61f39..beb042f4 100644
--- a/plugins/micromega/micromega.mli
+++ b/plugins/micromega/micromega.mli
@@ -1,15 +1,9 @@
-type __ = Obj.t
-
val negb : bool -> bool
type nat =
| O
| S of nat
-val fst : ('a1 * 'a2) -> 'a1
-
-val snd : ('a1 * 'a2) -> 'a2
-
val app : 'a1 list -> 'a1 list -> 'a1 list
type comparison =
@@ -19,24 +13,7 @@ type comparison =
val compOpp : comparison -> comparison
-type compareSpecT =
-| CompEqT
-| CompLtT
-| CompGtT
-
-val compareSpec2Type : comparison -> compareSpecT
-
-type 'a compSpecT = compareSpecT
-
-val compSpec2Type : 'a1 -> 'a1 -> comparison -> 'a1 compSpecT
-
-type 'a sig0 =
- 'a
- (* singleton inductive, whose constructor was exist *)
-
-val plus : nat -> nat -> nat
-
-val nat_iter : nat -> ('a1 -> 'a1) -> 'a1 -> 'a1
+val add : nat -> nat -> nat
type positive =
| XI of positive
@@ -52,560 +29,59 @@ type z =
| Zpos of positive
| Zneg of positive
-module type TotalOrder' =
- sig
- type t
- end
-
-module MakeOrderTac :
- functor (O:TotalOrder') ->
- sig
-
- end
-
-module MaxLogicalProperties :
- functor (O:TotalOrder') ->
- functor (M:sig
- val max : O.t -> O.t -> O.t
- end) ->
- sig
- module T :
- sig
-
- end
- end
-
-module Pos :
- sig
- type t = positive
-
- val succ : positive -> positive
-
- val add : positive -> positive -> positive
-
- val add_carry : positive -> positive -> positive
-
- val pred_double : positive -> positive
-
- val pred : positive -> positive
-
- val pred_N : positive -> n
-
+module Pos :
+ sig
type mask =
| IsNul
| IsPos of positive
| IsNeg
-
- val mask_rect : 'a1 -> (positive -> 'a1) -> 'a1 -> mask -> 'a1
-
- val mask_rec : 'a1 -> (positive -> 'a1) -> 'a1 -> mask -> 'a1
-
- val succ_double_mask : mask -> mask
-
- val double_mask : mask -> mask
-
- val double_pred_mask : positive -> mask
-
- val pred_mask : mask -> mask
-
- val sub_mask : positive -> positive -> mask
-
- val sub_mask_carry : positive -> positive -> mask
-
- val sub : positive -> positive -> positive
-
- val mul : positive -> positive -> positive
-
- val iter : positive -> ('a1 -> 'a1) -> 'a1 -> 'a1
-
- val pow : positive -> positive -> positive
-
- val div2 : positive -> positive
-
- val div2_up : positive -> positive
-
- val size_nat : positive -> nat
-
- val size : positive -> positive
-
- val compare_cont : positive -> positive -> comparison -> comparison
-
- val compare : positive -> positive -> comparison
-
- val min : positive -> positive -> positive
-
- val max : positive -> positive -> positive
-
- val eqb : positive -> positive -> bool
-
- val leb : positive -> positive -> bool
-
- val ltb : positive -> positive -> bool
-
- val sqrtrem_step :
- (positive -> positive) -> (positive -> positive) -> (positive * mask) ->
- positive * mask
-
- val sqrtrem : positive -> positive * mask
-
- val sqrt : positive -> positive
-
- val gcdn : nat -> positive -> positive -> positive
-
- val gcd : positive -> positive -> positive
-
- val ggcdn : nat -> positive -> positive -> positive * (positive * positive)
-
- val ggcd : positive -> positive -> positive * (positive * positive)
-
- val coq_Nsucc_double : n -> n
-
- val coq_Ndouble : n -> n
-
- val coq_lor : positive -> positive -> positive
-
- val coq_land : positive -> positive -> n
-
- val ldiff : positive -> positive -> n
-
- val coq_lxor : positive -> positive -> n
-
- val shiftl_nat : positive -> nat -> positive
-
- val shiftr_nat : positive -> nat -> positive
-
- val shiftl : positive -> n -> positive
-
- val shiftr : positive -> n -> positive
-
- val testbit_nat : positive -> nat -> bool
-
- val testbit : positive -> n -> bool
-
- val iter_op : ('a1 -> 'a1 -> 'a1) -> positive -> 'a1 -> 'a1
-
- val to_nat : positive -> nat
-
- val of_nat : nat -> positive
-
- val of_succ_nat : nat -> positive
end
-module Coq_Pos :
- sig
- module Coq__1 : sig
- type t = positive
- end
- type t = Coq__1.t
-
+module Coq_Pos :
+ sig
val succ : positive -> positive
-
+
val add : positive -> positive -> positive
-
+
val add_carry : positive -> positive -> positive
-
+
val pred_double : positive -> positive
-
- val pred : positive -> positive
-
- val pred_N : positive -> n
-
+
type mask = Pos.mask =
| IsNul
| IsPos of positive
| IsNeg
-
- val mask_rect : 'a1 -> (positive -> 'a1) -> 'a1 -> mask -> 'a1
-
- val mask_rec : 'a1 -> (positive -> 'a1) -> 'a1 -> mask -> 'a1
-
+
val succ_double_mask : mask -> mask
-
+
val double_mask : mask -> mask
-
+
val double_pred_mask : positive -> mask
-
- val pred_mask : mask -> mask
-
+
val sub_mask : positive -> positive -> mask
-
+
val sub_mask_carry : positive -> positive -> mask
-
+
val sub : positive -> positive -> positive
-
+
val mul : positive -> positive -> positive
-
- val iter : positive -> ('a1 -> 'a1) -> 'a1 -> 'a1
-
- val pow : positive -> positive -> positive
-
- val div2 : positive -> positive
-
- val div2_up : positive -> positive
-
+
val size_nat : positive -> nat
-
- val size : positive -> positive
-
- val compare_cont : positive -> positive -> comparison -> comparison
-
+
+ val compare_cont : comparison -> positive -> positive -> comparison
+
val compare : positive -> positive -> comparison
-
- val min : positive -> positive -> positive
-
- val max : positive -> positive -> positive
-
- val eqb : positive -> positive -> bool
-
- val leb : positive -> positive -> bool
-
- val ltb : positive -> positive -> bool
-
- val sqrtrem_step :
- (positive -> positive) -> (positive -> positive) -> (positive * mask) ->
- positive * mask
-
- val sqrtrem : positive -> positive * mask
-
- val sqrt : positive -> positive
-
+
val gcdn : nat -> positive -> positive -> positive
-
+
val gcd : positive -> positive -> positive
-
- val ggcdn : nat -> positive -> positive -> positive * (positive * positive)
-
- val ggcd : positive -> positive -> positive * (positive * positive)
-
- val coq_Nsucc_double : n -> n
-
- val coq_Ndouble : n -> n
-
- val coq_lor : positive -> positive -> positive
-
- val coq_land : positive -> positive -> n
-
- val ldiff : positive -> positive -> n
-
- val coq_lxor : positive -> positive -> n
-
- val shiftl_nat : positive -> nat -> positive
-
- val shiftr_nat : positive -> nat -> positive
-
- val shiftl : positive -> n -> positive
-
- val shiftr : positive -> n -> positive
-
- val testbit_nat : positive -> nat -> bool
-
- val testbit : positive -> n -> bool
-
- val iter_op : ('a1 -> 'a1 -> 'a1) -> positive -> 'a1 -> 'a1
-
- val to_nat : positive -> nat
-
- val of_nat : nat -> positive
-
+
val of_succ_nat : nat -> positive
-
- val eq_dec : positive -> positive -> bool
-
- val peano_rect : 'a1 -> (positive -> 'a1 -> 'a1) -> positive -> 'a1
-
- val peano_rec : 'a1 -> (positive -> 'a1 -> 'a1) -> positive -> 'a1
-
- type coq_PeanoView =
- | PeanoOne
- | PeanoSucc of positive * coq_PeanoView
-
- val coq_PeanoView_rect :
- 'a1 -> (positive -> coq_PeanoView -> 'a1 -> 'a1) -> positive ->
- coq_PeanoView -> 'a1
-
- val coq_PeanoView_rec :
- 'a1 -> (positive -> coq_PeanoView -> 'a1 -> 'a1) -> positive ->
- coq_PeanoView -> 'a1
-
- val peanoView_xO : positive -> coq_PeanoView -> coq_PeanoView
-
- val peanoView_xI : positive -> coq_PeanoView -> coq_PeanoView
-
- val peanoView : positive -> coq_PeanoView
-
- val coq_PeanoView_iter :
- 'a1 -> (positive -> 'a1 -> 'a1) -> positive -> coq_PeanoView -> 'a1
-
- val switch_Eq : comparison -> comparison -> comparison
-
- val mask2cmp : mask -> comparison
-
- module T :
- sig
-
- end
-
- module ORev :
- sig
- type t = Coq__1.t
- end
-
- module MRev :
- sig
- val max : t -> t -> t
- end
-
- module MPRev :
- sig
- module T :
- sig
-
- end
- end
-
- module P :
- sig
- val max_case_strong :
- t -> t -> (t -> t -> __ -> 'a1 -> 'a1) -> (__ -> 'a1) -> (__ -> 'a1) ->
- 'a1
-
- val max_case :
- t -> t -> (t -> t -> __ -> 'a1 -> 'a1) -> 'a1 -> 'a1 -> 'a1
-
- val max_dec : t -> t -> bool
-
- val min_case_strong :
- t -> t -> (t -> t -> __ -> 'a1 -> 'a1) -> (__ -> 'a1) -> (__ -> 'a1) ->
- 'a1
-
- val min_case :
- t -> t -> (t -> t -> __ -> 'a1 -> 'a1) -> 'a1 -> 'a1 -> 'a1
-
- val min_dec : t -> t -> bool
- end
-
- val max_case_strong : t -> t -> (__ -> 'a1) -> (__ -> 'a1) -> 'a1
-
- val max_case : t -> t -> 'a1 -> 'a1 -> 'a1
-
- val max_dec : t -> t -> bool
-
- val min_case_strong : t -> t -> (__ -> 'a1) -> (__ -> 'a1) -> 'a1
-
- val min_case : t -> t -> 'a1 -> 'a1 -> 'a1
-
- val min_dec : t -> t -> bool
end
-module N :
- sig
- type t = n
-
- val zero : n
-
- val one : n
-
- val two : n
-
- val succ_double : n -> n
-
- val double : n -> n
-
- val succ : n -> n
-
- val pred : n -> n
-
- val succ_pos : n -> positive
-
- val add : n -> n -> n
-
- val sub : n -> n -> n
-
- val mul : n -> n -> n
-
- val compare : n -> n -> comparison
-
- val eqb : n -> n -> bool
-
- val leb : n -> n -> bool
-
- val ltb : n -> n -> bool
-
- val min : n -> n -> n
-
- val max : n -> n -> n
-
- val div2 : n -> n
-
- val even : n -> bool
-
- val odd : n -> bool
-
- val pow : n -> n -> n
-
- val log2 : n -> n
-
- val size : n -> n
-
- val size_nat : n -> nat
-
- val pos_div_eucl : positive -> n -> n * n
-
- val div_eucl : n -> n -> n * n
-
- val div : n -> n -> n
-
- val modulo : n -> n -> n
-
- val gcd : n -> n -> n
-
- val ggcd : n -> n -> n * (n * n)
-
- val sqrtrem : n -> n * n
-
- val sqrt : n -> n
-
- val coq_lor : n -> n -> n
-
- val coq_land : n -> n -> n
-
- val ldiff : n -> n -> n
-
- val coq_lxor : n -> n -> n
-
- val shiftl_nat : n -> nat -> n
-
- val shiftr_nat : n -> nat -> n
-
- val shiftl : n -> n -> n
-
- val shiftr : n -> n -> n
-
- val testbit_nat : n -> nat -> bool
-
- val testbit : n -> n -> bool
-
- val to_nat : n -> nat
-
+module N :
+ sig
val of_nat : nat -> n
-
- val iter : n -> ('a1 -> 'a1) -> 'a1 -> 'a1
-
- val eq_dec : n -> n -> bool
-
- val discr : n -> positive option
-
- val binary_rect : 'a1 -> (n -> 'a1 -> 'a1) -> (n -> 'a1 -> 'a1) -> n -> 'a1
-
- val binary_rec : 'a1 -> (n -> 'a1 -> 'a1) -> (n -> 'a1 -> 'a1) -> n -> 'a1
-
- val peano_rect : 'a1 -> (n -> 'a1 -> 'a1) -> n -> 'a1
-
- val peano_rec : 'a1 -> (n -> 'a1 -> 'a1) -> n -> 'a1
-
- module BootStrap :
- sig
-
- end
-
- val recursion : 'a1 -> (n -> 'a1 -> 'a1) -> n -> 'a1
-
- module OrderElts :
- sig
- type t = n
- end
-
- module OrderTac :
- sig
-
- end
-
- module NZPowP :
- sig
-
- end
-
- module NZSqrtP :
- sig
-
- end
-
- val sqrt_up : n -> n
-
- val log2_up : n -> n
-
- module NZDivP :
- sig
-
- end
-
- val lcm : n -> n -> n
-
- val b2n : bool -> n
-
- val setbit : n -> n -> n
-
- val clearbit : n -> n -> n
-
- val ones : n -> n
-
- val lnot : n -> n -> n
-
- module T :
- sig
-
- end
-
- module ORev :
- sig
- type t = n
- end
-
- module MRev :
- sig
- val max : n -> n -> n
- end
-
- module MPRev :
- sig
- module T :
- sig
-
- end
- end
-
- module P :
- sig
- val max_case_strong :
- n -> n -> (n -> n -> __ -> 'a1 -> 'a1) -> (__ -> 'a1) -> (__ -> 'a1) ->
- 'a1
-
- val max_case :
- n -> n -> (n -> n -> __ -> 'a1 -> 'a1) -> 'a1 -> 'a1 -> 'a1
-
- val max_dec : n -> n -> bool
-
- val min_case_strong :
- n -> n -> (n -> n -> __ -> 'a1 -> 'a1) -> (__ -> 'a1) -> (__ -> 'a1) ->
- 'a1
-
- val min_case :
- n -> n -> (n -> n -> __ -> 'a1 -> 'a1) -> 'a1 -> 'a1 -> 'a1
-
- val min_dec : n -> n -> bool
- end
-
- val max_case_strong : n -> n -> (__ -> 'a1) -> (__ -> 'a1) -> 'a1
-
- val max_case : n -> n -> 'a1 -> 'a1 -> 'a1
-
- val max_dec : n -> n -> bool
-
- val min_case_strong : n -> n -> (__ -> 'a1) -> (__ -> 'a1) -> 'a1
-
- val min_case : n -> n -> 'a1 -> 'a1 -> 'a1
-
- val min_dec : n -> n -> bool
end
val pow_pos : ('a1 -> 'a1 -> 'a1) -> 'a1 -> positive -> 'a1
@@ -616,225 +92,45 @@ val map : ('a1 -> 'a2) -> 'a1 list -> 'a2 list
val fold_right : ('a2 -> 'a1 -> 'a1) -> 'a1 -> 'a2 list -> 'a1
-module Z :
- sig
- type t = z
-
- val zero : z
-
- val one : z
-
- val two : z
-
+module Z :
+ sig
val double : z -> z
-
+
val succ_double : z -> z
-
+
val pred_double : z -> z
-
+
val pos_sub : positive -> positive -> z
-
+
val add : z -> z -> z
-
+
val opp : z -> z
-
- val succ : z -> z
-
- val pred : z -> z
-
+
val sub : z -> z -> z
-
+
val mul : z -> z -> z
-
- val pow_pos : z -> positive -> z
-
- val pow : z -> z -> z
-
+
val compare : z -> z -> comparison
-
- val sgn : z -> z
-
+
val leb : z -> z -> bool
-
- val geb : z -> z -> bool
-
+
val ltb : z -> z -> bool
-
+
val gtb : z -> z -> bool
-
- val eqb : z -> z -> bool
-
+
val max : z -> z -> z
-
- val min : z -> z -> z
-
+
val abs : z -> z
-
- val abs_nat : z -> nat
-
- val abs_N : z -> n
-
- val to_nat : z -> nat
-
+
val to_N : z -> n
-
- val of_nat : nat -> z
-
- val of_N : n -> z
-
- val iter : z -> ('a1 -> 'a1) -> 'a1 -> 'a1
-
+
val pos_div_eucl : positive -> z -> z * z
-
+
val div_eucl : z -> z -> z * z
-
+
val div : z -> z -> z
-
- val modulo : z -> z -> z
-
- val quotrem : z -> z -> z * z
-
- val quot : z -> z -> z
-
- val rem : z -> z -> z
-
- val even : z -> bool
-
- val odd : z -> bool
-
- val div2 : z -> z
-
- val quot2 : z -> z
-
- val log2 : z -> z
-
- val sqrtrem : z -> z * z
-
- val sqrt : z -> z
-
+
val gcd : z -> z -> z
-
- val ggcd : z -> z -> z * (z * z)
-
- val testbit : z -> z -> bool
-
- val shiftl : z -> z -> z
-
- val shiftr : z -> z -> z
-
- val coq_lor : z -> z -> z
-
- val coq_land : z -> z -> z
-
- val ldiff : z -> z -> z
-
- val coq_lxor : z -> z -> z
-
- val eq_dec : z -> z -> bool
-
- module BootStrap :
- sig
-
- end
-
- module OrderElts :
- sig
- type t = z
- end
-
- module OrderTac :
- sig
-
- end
-
- val sqrt_up : z -> z
-
- val log2_up : z -> z
-
- module NZDivP :
- sig
-
- end
-
- module Quot2Div :
- sig
- val div : z -> z -> z
-
- val modulo : z -> z -> z
- end
-
- module NZQuot :
- sig
-
- end
-
- val lcm : z -> z -> z
-
- val b2z : bool -> z
-
- val setbit : z -> z -> z
-
- val clearbit : z -> z -> z
-
- val lnot : z -> z
-
- val ones : z -> z
-
- module T :
- sig
-
- end
-
- module ORev :
- sig
- type t = z
- end
-
- module MRev :
- sig
- val max : z -> z -> z
- end
-
- module MPRev :
- sig
- module T :
- sig
-
- end
- end
-
- module P :
- sig
- val max_case_strong :
- z -> z -> (z -> z -> __ -> 'a1 -> 'a1) -> (__ -> 'a1) -> (__ -> 'a1) ->
- 'a1
-
- val max_case :
- z -> z -> (z -> z -> __ -> 'a1 -> 'a1) -> 'a1 -> 'a1 -> 'a1
-
- val max_dec : z -> z -> bool
-
- val min_case_strong :
- z -> z -> (z -> z -> __ -> 'a1 -> 'a1) -> (__ -> 'a1) -> (__ -> 'a1) ->
- 'a1
-
- val min_case :
- z -> z -> (z -> z -> __ -> 'a1 -> 'a1) -> 'a1 -> 'a1 -> 'a1
-
- val min_dec : z -> z -> bool
- end
-
- val max_case_strong : z -> z -> (__ -> 'a1) -> (__ -> 'a1) -> 'a1
-
- val max_case : z -> z -> 'a1 -> 'a1 -> 'a1
-
- val max_dec : z -> z -> bool
-
- val min_case_strong : z -> z -> (__ -> 'a1) -> (__ -> 'a1) -> 'a1
-
- val min_case : z -> z -> 'a1 -> 'a1 -> 'a1
-
- val min_dec : z -> z -> bool
end
val zeq_bool : z -> z -> bool
@@ -872,44 +168,44 @@ val paddI :
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
+ ('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
+ '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
+ '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
+ '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
+ '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
+ '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
+ '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
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol
val psquare :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
- bool) -> 'a1 pol -> 'a1 pol
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> bool) -> 'a1 pol -> 'a1 pol
type 'c pExpr =
| PEc of 'c
@@ -923,16 +219,17 @@ type 'c pExpr =
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
+ '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
+ '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
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExpr -> 'a1 pol
type 'a bFormula =
| TT
@@ -946,9 +243,9 @@ type 'a bFormula =
val map_bformula : ('a1 -> 'a2) -> 'a1 bFormula -> 'a2 bFormula
-type 'term' clause = 'term' list
+type 'x clause = 'x list
-type 'term' cnf = 'term' clause list
+type 'x cnf = 'x clause list
val tt : 'a1 cnf
@@ -959,12 +256,12 @@ val add_term :
clause option
val or_clause :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 clause -> 'a1 clause ->
- 'a1 clause option
+ ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 clause -> 'a1 clause
+ -> 'a1 clause option
val or_clause_cnf :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 clause -> 'a1 cnf -> 'a1
- cnf
+ ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 clause -> 'a1 cnf ->
+ 'a1 cnf
val or_cnf :
('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 cnf -> 'a1 cnf -> 'a1
@@ -973,18 +270,20 @@ val or_cnf :
val and_cnf : 'a1 cnf -> 'a1 cnf -> 'a1 cnf
val xcnf :
- ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a2 cnf) -> ('a1 ->
- 'a2 cnf) -> bool -> 'a1 bFormula -> 'a2 cnf
+ ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('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 :
- ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a2 cnf) -> ('a1 ->
- 'a2 cnf) -> ('a2 list -> 'a3 -> bool) -> 'a1 bFormula -> 'a3 list -> bool
+ ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a2 cnf) -> ('a1
+ -> 'a2 cnf) -> ('a2 list -> 'a3 -> bool) -> 'a1 bFormula -> 'a3 list ->
+ bool
val cneqb : ('a1 -> 'a1 -> bool) -> 'a1 -> 'a1 -> bool
-val cltb : ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 -> 'a1 -> bool
+val cltb :
+ ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 -> 'a1 -> bool
type 'c polC = 'c pol
@@ -1015,28 +314,30 @@ val map_option2 :
('a1 -> 'a2 -> 'a3 option) -> 'a1 option -> 'a2 option -> 'a3 option
val pexpr_times_nformula :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
- bool) -> 'a1 polC -> 'a1 nFormula -> 'a1 nFormula option
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> bool) -> 'a1 polC -> 'a1 nFormula -> 'a1 nFormula option
val nformula_times_nformula :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
- bool) -> 'a1 nFormula -> 'a1 nFormula -> 'a1 nFormula option
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> bool) -> 'a1 nFormula -> 'a1 nFormula -> 'a1 nFormula option
val nformula_plus_nformula :
- 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula -> 'a1
- nFormula -> 'a1 nFormula option
+ 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula ->
+ 'a1 nFormula -> 'a1 nFormula option
val eval_Psatz :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
- bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz -> 'a1
- nFormula option
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz ->
+ 'a1 nFormula option
val check_inconsistent :
- 'a1 -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula -> bool
+ 'a1 -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula ->
+ bool
val check_normalised_formulas :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
- bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz -> bool
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz ->
+ bool
type op2 =
| OpEq
@@ -1048,43 +349,37 @@ type op2 =
type 't formula = { flhs : 't pExpr; fop : op2; frhs : 't pExpr }
-val flhs : 'a1 formula -> 'a1 pExpr
-
-val fop : 'a1 formula -> op2
-
-val frhs : 'a1 formula -> 'a1 pExpr
-
val norm :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
- 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExpr -> 'a1 pol
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExpr -> 'a1 pol
val psub0 :
- 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1
- -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol
+ 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) ->
+ ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol
val padd0 :
- 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol ->
- 'a1 pol
+ 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol
+ -> 'a1 pol
val xnormalise :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
- 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1 nFormula
- list
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1
+ nFormula list
val cnf_normalise :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
- 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1 nFormula
- cnf
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1
+ nFormula cnf
val xnegate :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
- 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1 nFormula
- list
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1
+ nFormula list
val cnf_negate :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 ->
- 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1 nFormula
- cnf
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1
+ -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1
+ nFormula cnf
val xdenorm : positive -> 'a1 pol -> 'a1 pExpr
@@ -1095,8 +390,8 @@ val map_PExpr : ('a2 -> 'a1) -> 'a2 pExpr -> 'a1 pExpr
val map_Formula : ('a2 -> 'a1) -> 'a2 formula -> 'a1 formula
val simpl_cone :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 psatz ->
- 'a1 psatz
+ 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 psatz
+ -> 'a1 psatz
type q = { qnum : z; qden : positive }
@@ -1122,12 +417,16 @@ val qpower_positive : q -> positive -> q
val qpower : q -> z -> q
-type 'a t0 =
+type 'a t =
| Empty
| Leaf of 'a
-| Node of 'a t0 * 'a * 'a t0
+| Node of 'a t * 'a * 'a t
+
+val find : 'a1 -> 'a1 t -> positive -> 'a1
-val find : 'a1 -> 'a1 t0 -> positive -> 'a1
+val singleton : 'a1 -> positive -> 'a1 -> 'a1 t
+
+val vm_add : 'a1 -> positive -> 'a1 -> 'a1 t -> 'a1 t
type zWitness = z psatz
@@ -1221,4 +520,3 @@ val runsat : q nFormula -> bool
val rdeduce : q nFormula -> q nFormula -> q nFormula option
val rTautoChecker : rcst formula bFormula -> rWitness list -> bool
-
diff --git a/plugins/micromega/micromega_plugin.mllib b/plugins/micromega/micromega_plugin.mlpack
index f53a9e37..ed253da3 100644
--- a/plugins/micromega/micromega_plugin.mllib
+++ b/plugins/micromega/micromega_plugin.mlpack
@@ -7,4 +7,3 @@ Certificate
Persistent_cache
Coq_micromega
G_micromega
-Micromega_plugin_mod
diff --git a/plugins/micromega/mutils.ml b/plugins/micromega/mutils.ml
index 2dd443f0..b4c6d032 100644
--- a/plugins/micromega/mutils.ml
+++ b/plugins/micromega/mutils.ml
@@ -66,6 +66,15 @@ let all_sym_pairs f l =
| e::l -> xpairs (pair_with acc e l) l in
xpairs [] l
+let all_pairs f l =
+ let pair_with acc e l = List.fold_left (fun acc x -> (f e x) ::acc) acc l in
+
+ let rec xpairs acc l =
+ match l with
+ | [] -> acc
+ | e::lx -> xpairs (pair_with acc e l) lx in
+ xpairs [] l
+
let rec map3 f l1 l2 l3 =
@@ -285,18 +294,6 @@ struct
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 Int.equal n 1 then []
- else (Int.equal (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
diff --git a/plugins/micromega/persistent_cache.ml b/plugins/micromega/persistent_cache.ml
index 6a03e2d6..0e6d346a 100644
--- a/plugins/micromega/persistent_cache.ml
+++ b/plugins/micromega/persistent_cache.ml
@@ -92,7 +92,7 @@ let read_key_elem inch =
Some (Marshal.from_channel inch)
with
| End_of_file -> None
- | e when Errors.noncritical e -> raise InvalidTableFormat
+ | e when CErrors.noncritical e -> raise InvalidTableFormat
(**
We used to only lock/unlock regions.
@@ -212,9 +212,11 @@ let find t k =
res
let memo cache f =
- let tbl = lazy (open_in cache) in
- fun x ->
- let tbl = Lazy.force tbl in
+ let tbl = lazy (try Some (open_in cache) with _ -> None) in
+ fun x ->
+ match Lazy.force tbl with
+ | None -> f x
+ | Some tbl ->
try
find tbl x
with
diff --git a/plugins/micromega/vo.itarget b/plugins/micromega/vo.itarget
index bf6a1a7d..c9009ea4 100644
--- a/plugins/micromega/vo.itarget
+++ b/plugins/micromega/vo.itarget
@@ -10,4 +10,6 @@ Tauto.vo
VarMap.vo
ZCoeff.vo
ZMicromega.vo
-Lia.vo \ No newline at end of file
+Lia.vo
+Lqa.vo
+Lra.vo