aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
-rw-r--r--.ocamlformat1
-rw-r--r--.ocamlformat-ignore5
-rw-r--r--plugins/micromega/certificate.ml1344
-rw-r--r--plugins/micromega/certificate.mli28
-rw-r--r--plugins/micromega/coq_micromega.ml3268
-rw-r--r--plugins/micromega/coq_micromega.mli3
-rw-r--r--plugins/micromega/csdpcert.ml261
-rw-r--r--plugins/micromega/itv.ml92
-rw-r--r--plugins/micromega/itv.mli1
-rw-r--r--plugins/micromega/mfourier.ml1032
-rw-r--r--plugins/micromega/mfourier.mli17
-rw-r--r--plugins/micromega/micromega.mli784
-rw-r--r--plugins/micromega/mutils.ml594
-rw-r--r--plugins/micromega/mutils.mli83
-rw-r--r--plugins/micromega/persistent_cache.ml277
-rw-r--r--plugins/micromega/persistent_cache.mli33
-rw-r--r--plugins/micromega/polynomial.ml1271
-rw-r--r--plugins/micromega/polynomial.mli151
-rw-r--r--plugins/micromega/simplex.ml781
-rw-r--r--plugins/micromega/simplex.mli5
-rw-r--r--plugins/micromega/sos.ml1297
-rw-r--r--plugins/micromega/sos.mli22
-rw-r--r--plugins/micromega/sos_lib.ml527
-rw-r--r--plugins/micromega/sos_lib.mli39
-rw-r--r--plugins/micromega/sos_types.ml74
-rw-r--r--plugins/micromega/sos_types.mli38
-rw-r--r--plugins/micromega/vect.ml361
-rw-r--r--plugins/micromega/vect.mli77
-rw-r--r--plugins/micromega/zify.ml821
-rw-r--r--plugins/micromega/zify.mli17
30 files changed, 6529 insertions, 6775 deletions
diff --git a/.ocamlformat b/.ocamlformat
index 3c6f577afd..59883180e5 100644
--- a/.ocamlformat
+++ b/.ocamlformat
@@ -4,3 +4,4 @@ sequence-style=terminator
cases-exp-indent=2
field-space=loose
exp-grouping=preserve
+break-cases=fit
diff --git a/.ocamlformat-ignore b/.ocamlformat-ignore
index af9d896319..b1f6597140 100644
--- a/.ocamlformat-ignore
+++ b/.ocamlformat-ignore
@@ -39,7 +39,6 @@ plugins/firstorder/*
plugins/fourier/*
plugins/funind/*
plugins/ltac/*
-plugins/micromega/*
plugins/nsatz/*
plugins/omega/*
plugins/rtauto/*
@@ -49,4 +48,6 @@ plugins/setoid_ring/*
plugins/ssr/*
plugins/ssrmatching/*
plugins/syntax/*
-# Enabled: none for now
+# Enabled: micromega
+# plugins/micromega/*
+plugins/micromega/micromega.ml
diff --git a/plugins/micromega/certificate.ml b/plugins/micromega/certificate.ml
index 82c2be582b..cb15274736 100644
--- a/plugins/micromega/certificate.ml
+++ b/plugins/micromega/certificate.ml
@@ -22,97 +22,85 @@ let debug = false
open Big_int
open Num
open Polynomial
-
module Mc = Micromega
module Ml2C = Mutils.CamlToCoq
module C2Ml = Mutils.CoqToCaml
let use_simplex = ref true
-
-type ('prf,'model) res =
- | Prf of 'prf
- | Model of 'model
- | Unknown
-
-type zres = (Mc.zArithProof , (int * Mc.z list)) res
-
-type qres = (Mc.q Mc.psatz , (int * Mc.q list)) res
-
+type ('prf, 'model) res = Prf of 'prf | Model of 'model | Unknown
+type zres = (Mc.zArithProof, int * Mc.z list) res
+type qres = (Mc.q Mc.psatz, int * Mc.q list) res
open Mutils
-type 'a number_spec = {
- bigint_to_number : big_int -> 'a;
- number_to_num : 'a -> num;
- zero : 'a;
- unit : 'a;
- mult : 'a -> 'a -> 'a;
- eqb : 'a -> 'a -> bool
- }
-
-let z_spec = {
- bigint_to_number = Ml2C.bigint ;
- number_to_num = (fun x -> Big_int (C2Ml.z_big_int x));
- zero = Mc.Z0;
- unit = Mc.Zpos Mc.XH;
- mult = Mc.Z.mul;
- eqb = Mc.zeq_bool
- }
-
-
-let q_spec = {
- bigint_to_number = (fun x -> {Mc.qnum = Ml2C.bigint x; Mc.qden = Mc.XH});
- number_to_num = C2Ml.q_to_num;
- zero = {Mc.qnum = Mc.Z0;Mc.qden = Mc.XH};
- unit = {Mc.qnum = (Mc.Zpos Mc.XH) ; Mc.qden = Mc.XH};
- mult = Mc.qmult;
- eqb = Mc.qeq_bool
- }
-
-let dev_form n_spec p =
+
+type 'a number_spec =
+ { bigint_to_number : big_int -> 'a
+ ; number_to_num : 'a -> num
+ ; zero : 'a
+ ; unit : 'a
+ ; mult : 'a -> 'a -> 'a
+ ; eqb : 'a -> 'a -> bool }
+
+let z_spec =
+ { bigint_to_number = Ml2C.bigint
+ ; number_to_num = (fun x -> Big_int (C2Ml.z_big_int x))
+ ; zero = Mc.Z0
+ ; unit = Mc.Zpos Mc.XH
+ ; mult = Mc.Z.mul
+ ; eqb = Mc.zeq_bool }
+
+let q_spec =
+ { bigint_to_number = (fun x -> {Mc.qnum = Ml2C.bigint x; Mc.qden = Mc.XH})
+ ; number_to_num = C2Ml.q_to_num
+ ; zero = {Mc.qnum = Mc.Z0; Mc.qden = Mc.XH}
+ ; unit = {Mc.qnum = Mc.Zpos Mc.XH; Mc.qden = Mc.XH}
+ ; mult = Mc.qmult
+ ; eqb = Mc.qeq_bool }
+
+let 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
+ | 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 rec fixpoint f x =
let y' = f x in
- if (=) y' x then y'
- else fixpoint f y'
+ if y' = x then y' else fixpoint f y'
-let rec_simpl_cone n_spec e =
+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
+ 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
-
let simplify_cone n_spec c = fixpoint (rec_simpl_cone n_spec) c
-
-
(* The binding with Fourier might be a bit obsolete
-- how does it handle equalities ? *)
@@ -133,174 +121,166 @@ let simplify_cone n_spec c = fixpoint (rec_simpl_cone n_spec) c
(* fold_left followed by a rev ! *)
let constrain_variable v l =
- let coeffs = List.fold_left (fun acc p -> (Vect.get v p.coeffs)::acc) [] l in
- { coeffs = Vect.from_list ((Big_int zero_big_int):: (Big_int zero_big_int):: (List.rev coeffs)) ;
- op = Eq ;
- cst = Big_int zero_big_int }
-
-
+ let coeffs = List.fold_left (fun acc p -> Vect.get v p.coeffs :: acc) [] l in
+ { coeffs =
+ Vect.from_list
+ (Big_int zero_big_int :: Big_int zero_big_int :: List.rev coeffs)
+ ; op = Eq
+ ; cst = Big_int zero_big_int }
let constrain_constant l =
- let coeffs = List.fold_left (fun acc p -> minus_num p.cst ::acc) [] l in
- { coeffs = Vect.from_list ((Big_int zero_big_int):: (Big_int unit_big_int):: (List.rev coeffs)) ;
- op = Eq ;
- cst = Big_int zero_big_int }
+ let coeffs = List.fold_left (fun acc p -> minus_num p.cst :: acc) [] l in
+ { coeffs =
+ Vect.from_list
+ (Big_int zero_big_int :: Big_int unit_big_int :: List.rev coeffs)
+ ; op = Eq
+ ; cst = Big_int zero_big_int }
let positivity l =
let rec xpositivity i l =
match l with
| [] -> []
- | c::l -> match c.op with
- | Eq -> xpositivity (i+1) l
- | _ ->
- {coeffs = Vect.update (i+1) (fun _ -> Int 1) Vect.null ;
- op = Ge ;
- cst = Int 0 } :: (xpositivity (i+1) l)
+ | c :: l -> (
+ match c.op with
+ | Eq -> xpositivity (i + 1) l
+ | _ ->
+ { coeffs = Vect.update (i + 1) (fun _ -> Int 1) Vect.null
+ ; op = Ge
+ ; cst = Int 0 }
+ :: xpositivity (i + 1) l )
in
xpositivity 1 l
-
-let cstr_of_poly (p,o) =
- let (c,l) = Vect.decomp_cst p in
- {coeffs = l; op = o ; cst = minus_num c}
-
-
+let cstr_of_poly (p, o) =
+ let c, l = Vect.decomp_cst p in
+ {coeffs = l; op = o; cst = minus_num c}
let variables_of_cstr c = Vect.variables c.coeffs
-
(* If the certificate includes at least one strict inequality,
the obtained polynomial can also be 0 *)
let build_dual_linear_system l =
-
let variables =
- List.fold_left (fun acc p -> ISet.union acc (variables_of_cstr p)) ISet.empty l in
+ List.fold_left
+ (fun acc p -> ISet.union acc (variables_of_cstr p))
+ ISet.empty l
+ in
(* For each monomial, compute a constraint *)
let s0 =
- ISet.fold (fun mn res -> (constrain_variable mn l)::res) variables [] in
- let c = constrain_constant l in
-
+ ISet.fold (fun mn res -> constrain_variable mn l :: res) variables []
+ in
+ let c = constrain_constant l in
(* I need at least something strictly positive *)
- let strict = {
- coeffs = Vect.from_list ((Big_int zero_big_int) :: (Big_int unit_big_int)::
- (List.map (fun c -> if is_strict c then Big_int unit_big_int else Big_int zero_big_int) l));
- op = Ge ; cst = Big_int unit_big_int } in
+ let strict =
+ { coeffs =
+ Vect.from_list
+ ( Big_int zero_big_int :: Big_int unit_big_int
+ :: List.map
+ (fun c ->
+ if is_strict c then Big_int unit_big_int
+ else 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 zero_big_int ;Big_int unit_big_int]) ;
- op = Ge ;
- cst = Big_int zero_big_int}::(strict::(positivity l)@c::s0)
+ { coeffs = Vect.from_list [Big_int zero_big_int; Big_int unit_big_int]
+ ; op = Ge
+ ; cst = Big_int zero_big_int }
+ :: ((strict :: positivity l) @ (c :: s0))
+
open Util
(** [direct_linear_prover l] does not handle strict inegalities *)
let fourier_linear_prover l =
match Mfourier.Fourier.find_point l with
| Inr prf ->
- if debug then Printf.printf "AProof : %a\n" Mfourier.pp_proof prf ;
- let cert = (*List.map (fun (x,n) -> x+1,n)*) (fst (List.hd (Mfourier.Proof.mk_proof l prf))) in
- if debug then Printf.printf "CProof : %a" Vect.pp cert ;
- (*Some (rats_to_ints (Vect.to_list cert))*)
- Some (Vect.normalise cert)
- | Inl _ -> None
-
+ if debug then Printf.printf "AProof : %a\n" Mfourier.pp_proof prf;
+ let cert =
+ (*List.map (fun (x,n) -> x+1,n)*)
+ fst (List.hd (Mfourier.Proof.mk_proof l prf))
+ in
+ if debug then Printf.printf "CProof : %a" Vect.pp cert;
+ (*Some (rats_to_ints (Vect.to_list cert))*)
+ Some (Vect.normalise cert)
+ | Inl _ -> None
let direct_linear_prover l =
- if !use_simplex
- then Simplex.find_unsat_certificate l
+ if !use_simplex then Simplex.find_unsat_certificate l
else fourier_linear_prover l
let find_point l =
- if !use_simplex
- then Simplex.find_point l
- else match Mfourier.Fourier.find_point l with
- | Inr _ -> None
- | Inl cert -> Some cert
+ if !use_simplex then Simplex.find_point l
+ else
+ match Mfourier.Fourier.find_point l with
+ | Inr _ -> None
+ | Inl cert -> Some cert
let optimise v l =
- if !use_simplex
- then Simplex.optimise v l
- else Mfourier.Fourier.optimise v l
-
-
+ if !use_simplex then Simplex.optimise v l else Mfourier.Fourier.optimise v l
let dual_raw_certificate l =
- if debug
- then begin
- Printf.printf "dual_raw_certificate\n";
- List.iter (fun c -> Printf.fprintf stdout "%a\n" output_cstr c) l
- end;
-
+ if debug then begin
+ Printf.printf "dual_raw_certificate\n";
+ List.iter (fun c -> Printf.fprintf stdout "%a\n" output_cstr c) l
+ end;
let sys = build_dual_linear_system l in
-
if debug then begin
- Printf.printf "dual_system\n";
- List.iter (fun c -> Printf.fprintf stdout "%a\n" output_cstr c) sys
- end;
-
+ Printf.printf "dual_system\n";
+ List.iter (fun c -> Printf.fprintf stdout "%a\n" output_cstr c) sys
+ end;
try
match find_point sys with
| None -> None
- | Some cert ->
- match Vect.choose cert with
- | None -> failwith "dual_raw_certificate: empty_certificate"
- | Some _ ->
- (*Some (rats_to_ints (Vect.to_list (Vect.decr_var 2 (Vect.set 1 (Int 0) cert))))*)
- Some (Vect.normalise (Vect.decr_var 2 (Vect.set 1 (Int 0) cert)))
- (* should not use rats_to_ints *)
+ | Some cert -> (
+ match Vect.choose cert with
+ | None -> failwith "dual_raw_certificate: empty_certificate"
+ | Some _ ->
+ (*Some (rats_to_ints (Vect.to_list (Vect.decr_var 2 (Vect.set 1 (Int 0) cert))))*)
+ Some (Vect.normalise (Vect.decr_var 2 (Vect.set 1 (Int 0) cert))) )
+ (* should not use rats_to_ints *)
with x when CErrors.noncritical x ->
- if debug
- then (Printf.printf "dual raw certificate %s" (Printexc.to_string x);
- flush stdout) ;
- None
-
-
+ if debug then (
+ Printf.printf "dual raw certificate %s" (Printexc.to_string x);
+ flush stdout );
+ None
let simple_linear_prover l =
- try
- direct_linear_prover l
+ try direct_linear_prover l
with Strict ->
(* Fourier elimination should handle > *)
dual_raw_certificate l
let env_of_list l =
- snd (List.fold_left (fun (i,m) p -> (i+1,IMap.add i p m)) (0,IMap.empty) l)
-
-
-
+ snd
+ (List.fold_left (fun (i, m) p -> (i + 1, IMap.add i p m)) (0, IMap.empty) l)
let linear_prover_cstr sys =
- let (sysi,prfi) = List.split sys in
-
-
+ let sysi, prfi = List.split sys in
match simple_linear_prover sysi with
| None -> None
| Some cert -> Some (ProofFormat.proof_of_farkas (env_of_list prfi) cert)
-let linear_prover_cstr =
- if debug
- then
- fun sys ->
- Printf.printf "<linear_prover"; flush stdout ;
+let linear_prover_cstr =
+ if debug then ( fun sys ->
+ Printf.printf "<linear_prover";
+ flush stdout;
let res = linear_prover_cstr sys in
- Printf.printf ">"; flush stdout ;
- res
+ Printf.printf ">"; flush stdout; res )
else linear_prover_cstr
-
-
let compute_max_nb_cstr l d =
let len = List.length l in
max len (max d (len * d))
-
-let develop_constraint z_spec (e,k) =
- (dev_form z_spec e,
- match k with
- | Mc.NonStrict -> Ge
- | Mc.Equal -> Eq
- | Mc.Strict -> Gt
- | _ -> assert false
- )
+let develop_constraint z_spec (e, k) =
+ ( dev_form z_spec e
+ , match k with
+ | Mc.NonStrict -> Ge
+ | Mc.Equal -> Eq
+ | Mc.Strict -> Gt
+ | _ -> assert false )
(** A single constraint can be unsat for the following reasons:
- 0 >= c for c a negative constant
@@ -312,125 +292,109 @@ type checksat =
| Tauto (* Tautology *)
| Unsat of ProofFormat.prf_rule (* Unsatisfiable *)
| Cut of cstr * ProofFormat.prf_rule (* Cutting plane *)
- | Normalise of cstr * ProofFormat.prf_rule (* Coefficients may be normalised i.e relatively prime *)
+ | Normalise of cstr * ProofFormat.prf_rule
-exception FoundProof of ProofFormat.prf_rule
+(* Coefficients may be normalised i.e relatively prime *)
+exception FoundProof of ProofFormat.prf_rule
(** [check_sat]
- detects constraints that are not satisfiable;
- normalises constraints and generate cuts.
*)
-let check_int_sat (cstr,prf) =
- let {coeffs=coeffs ; op=op ; cst=cst} = cstr in
+let check_int_sat (cstr, prf) =
+ let {coeffs; op; cst} = cstr in
match Vect.choose coeffs with
- | None ->
- if eval_op op (Int 0) cst then Tauto else Unsat prf
- | _ ->
- let gcdi = Vect.gcd 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 = Vect.div gcd coeffs;
- op = op ; cst = cst // gcd
- } in
- Normalise(cstr,ProofFormat.Gcd(gcdi,prf))
- (* Normalise(cstr,CutPrf prf)*)
- end
- else
- match op with
- | Eq -> Unsat (ProofFormat.CutPrf prf)
- | Ge ->
- let cstr = {
- coeffs = Vect.div gcd coeffs;
- op = op ; cst = ceiling_num (cst // gcd)
- } in Cut(cstr,ProofFormat.CutPrf prf)
- | Gt -> failwith "check_sat : Unexpected operator"
-
+ | None -> if eval_op op (Int 0) cst then Tauto else Unsat prf
+ | _ -> (
+ let gcdi = Vect.gcd 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 begin
+ (* We can really normalise *)
+ assert (sign_num gcd >= 1);
+ let cstr = {coeffs = Vect.div gcd coeffs; op; cst = cst // gcd} in
+ Normalise (cstr, ProofFormat.Gcd (gcdi, prf))
+ (* Normalise(cstr,CutPrf prf)*)
+ end
+ else
+ match op with
+ | Eq -> Unsat (ProofFormat.CutPrf prf)
+ | Ge ->
+ let cstr =
+ {coeffs = Vect.div gcd coeffs; op; cst = ceiling_num (cst // gcd)}
+ in
+ Cut (cstr, ProofFormat.CutPrf prf)
+ | Gt -> failwith "check_sat : Unexpected operator" )
let apply_and_normalise check f psys =
- List.fold_left (fun acc pc' ->
+ List.fold_left
+ (fun acc pc' ->
match f pc' with
- | None -> pc'::acc
- | Some pc' ->
- match check pc' with
- | Tauto -> acc
- | Unsat prf -> raise (FoundProof prf)
- | Cut(c,p) -> (c,p)::acc
- | Normalise (c,p) -> (c,p)::acc
- ) [] psys
-
-
+ | None -> pc' :: acc
+ | Some pc' -> (
+ match check pc' with
+ | Tauto -> acc
+ | Unsat prf -> raise (FoundProof prf)
+ | Cut (c, p) -> (c, p) :: acc
+ | Normalise (c, p) -> (c, p) :: acc ))
+ [] psys
let is_linear_for v pc =
LinPoly.is_linear (fst (fst pc)) || LinPoly.is_linear_for v (fst (fst pc))
-
-
-
(*let non_linear_pivot sys pc v pc' =
if LinPoly.is_linear (fst (fst pc'))
then None (* There are other ways to deal with those *)
else WithProof.linear_pivot sys pc v pc'
*)
-let is_linear_substitution sys ((p,o),prf) =
- let pred v = v =/ Int 1 || v =/ Int (-1) in
+let is_linear_substitution sys ((p, o), prf) =
+ let pred v = v =/ Int 1 || v =/ Int (-1) in
match o with
- | Eq -> begin
- match
- List.filter (fun v -> List.for_all (is_linear_for v) sys) (LinPoly.search_all_linear pred p)
- with
- | [] -> None
- | v::_ -> Some v (* make a choice *)
- end
- | _ -> None
-
+ | Eq -> (
+ match
+ List.filter
+ (fun v -> List.for_all (is_linear_for v) sys)
+ (LinPoly.search_all_linear pred p)
+ with
+ | [] -> None
+ | v :: _ -> Some v (* make a choice *) )
+ | _ -> None
let elim_simple_linear_equality sys0 =
-
let elim sys =
- let (oeq,sys') = extract (is_linear_substitution sys) sys in
+ let oeq, sys' = extract (is_linear_substitution sys) sys in
match oeq with
| None -> None
- | Some(v,pc) -> simplify (WithProof.linear_pivot sys0 pc v) sys' in
-
+ | Some (v, pc) -> simplify (WithProof.linear_pivot sys0 pc v) sys'
+ in
iterate_until_stable elim sys0
-
-
let output_sys o sys =
List.iter (fun s -> Printf.fprintf o "%a\n" WithProof.output s) sys
let subst sys =
let sys' = WithProof.subst sys in
- if debug then Printf.fprintf stdout "[subst:\n%a\n==>\n%a\n]" output_sys sys output_sys sys' ;
+ if debug then
+ Printf.fprintf stdout "[subst:\n%a\n==>\n%a\n]" output_sys sys output_sys
+ sys';
sys'
-
-
(** [saturate_linear_equality sys] generate new constraints
obtained by eliminating linear equalities by pivoting.
For integers, the obtained constraints are sound but not complete.
*)
- let saturate_by_linear_equalities sys0 =
- WithProof.saturate_subst false sys0
-
+let saturate_by_linear_equalities sys0 = WithProof.saturate_subst false sys0
let saturate_by_linear_equalities sys =
let sys' = saturate_by_linear_equalities sys in
- if debug then Printf.fprintf stdout "[saturate_by_linear_equalities:\n%a\n==>\n%a\n]" output_sys sys output_sys sys' ;
+ if debug then
+ Printf.fprintf stdout "[saturate_by_linear_equalities:\n%a\n==>\n%a\n]"
+ output_sys sys output_sys sys';
sys'
-
-
(* let saturate_linear_equality_non_linear sys0 =
let (l,_) = extract_all (is_substitution false) sys0 in
let rec elim l acc =
@@ -442,108 +406,117 @@ let saturate_by_linear_equalities sys =
elim l []
*)
-let bounded_vars (sys: WithProof.t list) =
- let l = (fst (extract_all (fun ((p,o),prf) ->
- LinPoly.is_variable p
- ) sys)) in
- List.fold_left (fun acc (i,wp) -> IMap.add i wp acc) IMap.empty l
+let bounded_vars (sys : WithProof.t list) =
+ let l = fst (extract_all (fun ((p, o), prf) -> LinPoly.is_variable p) sys) in
+ List.fold_left (fun acc (i, wp) -> IMap.add i wp acc) IMap.empty l
-let rec power n p =
- if n = 1 then p
- else WithProof.product p (power (n-1) p)
+let rec power n p = if n = 1 then p else WithProof.product p (power (n - 1) p)
let bound_monomial mp m =
- if Monomial.is_var m || Monomial.is_const m
- then None
+ if Monomial.is_var m || Monomial.is_const m then None
else
- try
- Some (Monomial.fold
- (fun v i acc ->
- let wp = IMap.find v mp in
- WithProof.product (power i wp) acc) m (WithProof.const (Int 1))
- )
- with Not_found -> None
-
-
-let bound_monomials (sys:WithProof.t list) =
+ try
+ Some
+ (Monomial.fold
+ (fun v i acc ->
+ let wp = IMap.find v mp in
+ WithProof.product (power i wp) acc)
+ m (WithProof.const (Int 1)))
+ with Not_found -> None
+
+let bound_monomials (sys : WithProof.t list) =
let mp = bounded_vars sys in
- let m =
- List.fold_left (fun acc ((p,_),_) ->
- Vect.fold (fun acc v _ -> let m = LinPoly.MonT.retrieve v in
- match bound_monomial mp m with
- | None -> acc
- | Some r -> IMap.add v r acc) acc p) IMap.empty sys in
- IMap.fold (fun _ e acc -> e::acc) m []
-
+ let m =
+ List.fold_left
+ (fun acc ((p, _), _) ->
+ Vect.fold
+ (fun acc v _ ->
+ let m = LinPoly.MonT.retrieve v in
+ match bound_monomial mp m with
+ | None -> acc
+ | Some r -> IMap.add v r acc)
+ acc p)
+ IMap.empty sys
+ in
+ IMap.fold (fun _ e acc -> e :: acc) m []
let develop_constraints prfdepth n_spec sys =
LinPoly.MonT.clear ();
- max_nb_cstr := compute_max_nb_cstr sys prfdepth ;
+ max_nb_cstr := compute_max_nb_cstr sys prfdepth;
let sys = List.map (develop_constraint n_spec) sys in
- List.mapi (fun i (p,o) -> ((LinPoly.linpol_of_pol p,o),ProofFormat.Hyp i)) sys
+ List.mapi
+ (fun i (p, o) -> ((LinPoly.linpol_of_pol p, o), ProofFormat.Hyp i))
+ sys
let square_of_var i =
let x = LinPoly.var i in
- ((LinPoly.product x x,Ge),(ProofFormat.Square x))
-
+ ((LinPoly.product x x, Ge), ProofFormat.Square x)
(** [nlinear_preprocess sys] augments the system [sys] by performing some limited non-linear reasoning.
For instance, it asserts that the x² ≥0 but also that if c₁ ≥ 0 ∈ sys and c₂ ≥ 0 ∈ sys then c₁ × c₂ ≥ 0.
The resulting system is linearised.
*)
-let nlinear_preprocess (sys:WithProof.t list) =
-
- let is_linear = List.for_all (fun ((p,_),_) -> LinPoly.is_linear p) sys in
-
+let nlinear_preprocess (sys : WithProof.t list) =
+ let is_linear = List.for_all (fun ((p, _), _) -> LinPoly.is_linear p) sys in
if is_linear then sys
else
let collect_square =
- List.fold_left (fun acc ((p,_),_) -> MonMap.union (fun k e1 e2 -> Some e1) acc (LinPoly.collect_square p)) MonMap.empty sys in
- let sys = MonMap.fold (fun s m acc ->
- let s = LinPoly.of_monomial s in
- let m = LinPoly.of_monomial m in
- ((m, Ge), (ProofFormat.Square s))::acc) collect_square sys in
-
- let collect_vars = List.fold_left (fun acc p -> ISet.union acc (LinPoly.variables (fst (fst p)))) ISet.empty sys in
-
- let sys = ISet.fold (fun i acc -> square_of_var i :: acc) collect_vars sys in
-
- let sys = sys @ (all_pairs WithProof.product sys) in
-
+ List.fold_left
+ (fun acc ((p, _), _) ->
+ MonMap.union (fun k e1 e2 -> Some e1) acc (LinPoly.collect_square p))
+ MonMap.empty sys
+ in
+ let sys =
+ MonMap.fold
+ (fun s m acc ->
+ let s = LinPoly.of_monomial s in
+ let m = LinPoly.of_monomial m in
+ ((m, Ge), ProofFormat.Square s) :: acc)
+ collect_square sys
+ in
+ let collect_vars =
+ List.fold_left
+ (fun acc p -> ISet.union acc (LinPoly.variables (fst (fst p))))
+ ISet.empty sys
+ in
+ let sys =
+ ISet.fold (fun i acc -> square_of_var i :: acc) collect_vars sys
+ in
+ let sys = sys @ all_pairs WithProof.product sys in
if debug then begin
- Printf.fprintf stdout "Preprocessed\n";
- List.iter (fun s -> Printf.fprintf stdout "%a\n" WithProof.output s) sys;
- end ;
-
+ Printf.fprintf stdout "Preprocessed\n";
+ List.iter (fun s -> Printf.fprintf stdout "%a\n" WithProof.output s) sys
+ end;
List.map (WithProof.annot "P") sys
-
-
let nlinear_prover prfdepth sys =
let sys = develop_constraints prfdepth q_spec sys in
let sys1 = elim_simple_linear_equality sys in
let sys2 = saturate_by_linear_equalities sys1 in
- let sys = nlinear_preprocess sys1@sys2 in
- let sys = List.map (fun ((p,o),prf) -> (cstr_of_poly (p,o), prf)) sys in
- let id = (List.fold_left
- (fun acc (_,r) -> max acc (ProofFormat.pr_rule_max_id r)) 0 sys) in
+ let sys = nlinear_preprocess sys1 @ sys2 in
+ let sys = List.map (fun ((p, o), prf) -> (cstr_of_poly (p, o), prf)) sys in
+ let id =
+ List.fold_left
+ (fun acc (_, r) -> max acc (ProofFormat.pr_rule_max_id r))
+ 0 sys
+ in
let env = CList.interval 0 id in
match linear_prover_cstr sys with
| None -> Unknown
- | Some cert ->
- Prf (ProofFormat.cmpl_prf_rule Mc.normQ CamlToCoq.q env cert)
-
+ | Some cert -> Prf (ProofFormat.cmpl_prf_rule Mc.normQ CamlToCoq.q env cert)
let linear_prover_with_cert prfdepth sys =
let sys = develop_constraints prfdepth q_spec sys in
(* let sys = nlinear_preprocess sys in *)
- let sys = List.map (fun (c,p) -> cstr_of_poly c,p) sys in
-
+ let sys = List.map (fun (c, p) -> (cstr_of_poly c, p)) sys in
match linear_prover_cstr sys with
| None -> Unknown
| Some cert ->
- Prf (ProofFormat.cmpl_prf_rule Mc.normQ CamlToCoq.q (List.mapi (fun i e -> i) sys) cert)
+ Prf
+ (ProofFormat.cmpl_prf_rule Mc.normQ CamlToCoq.q
+ (List.mapi (fun i e -> i) sys)
+ cert)
(* The prover is (probably) incomplete --
only searching for naive cutting planes *)
@@ -552,514 +525,525 @@ 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
- | 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))
+ | Zero -> (unit_big_int, Zero)
+ | Const n -> (denominator n, Const (Big_int (numerator n)))
+ | Var n -> (unit_big_int, Var n)
+ | 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)
+ | 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))
let scale_term t =
- let (s,t') = scale_term t in
- s,t'
-
-let rec scale_certificate pos = match pos with
- | Axiom_eq i -> unit_big_int , Axiom_eq i
- | Axiom_le i -> unit_big_int , Axiom_le i
- | Axiom_lt i -> unit_big_int , Axiom_lt i
- | Monoid l -> unit_big_int , Monoid l
- | Rational_eq n -> (denominator n) , Rational_eq (Big_int (numerator n))
- | Rational_le n -> (denominator n) , Rational_le (Big_int (numerator n))
- | Rational_lt n -> (denominator n) , Rational_lt (Big_int (numerator n))
- | Square t -> let s,t' = scale_term t in
- mult_big_int s s , Square t'
- | Eqmul (t, y) -> let s1,y1 = scale_term t and s2,y2 = scale_certificate y in
- mult_big_int s1 s2 , Eqmul (y1,y2)
- | Sum (y, z) -> let s1,y1 = scale_certificate y
- and s2,y2 = scale_certificate z in
- let g = gcd_big_int s1 s2 in
- let s1' = div_big_int s1 g in
- let s2' = div_big_int s2 g in
- mult_big_int g (mult_big_int s1' s2'),
- Sum (Product(Rational_le (Big_int s2'), y1),
- Product (Rational_le (Big_int s1'), y2))
+ let s, t' = scale_term t in
+ (s, t')
+
+let rec scale_certificate pos =
+ match pos with
+ | Axiom_eq i -> (unit_big_int, Axiom_eq i)
+ | Axiom_le i -> (unit_big_int, Axiom_le i)
+ | Axiom_lt i -> (unit_big_int, Axiom_lt i)
+ | Monoid l -> (unit_big_int, Monoid l)
+ | Rational_eq n -> (denominator n, Rational_eq (Big_int (numerator n)))
+ | Rational_le n -> (denominator n, Rational_le (Big_int (numerator n)))
+ | Rational_lt n -> (denominator n, Rational_lt (Big_int (numerator n)))
+ | Square t ->
+ let s, t' = scale_term t in
+ (mult_big_int s s, Square t')
+ | Eqmul (t, y) ->
+ let s1, y1 = scale_term t and s2, y2 = scale_certificate y in
+ (mult_big_int s1 s2, Eqmul (y1, y2))
+ | Sum (y, z) ->
+ let s1, y1 = scale_certificate y and s2, y2 = scale_certificate z in
+ let g = gcd_big_int s1 s2 in
+ let s1' = div_big_int s1 g in
+ let s2' = div_big_int s2 g in
+ ( mult_big_int g (mult_big_int s1' s2')
+ , Sum
+ ( Product (Rational_le (Big_int s2'), y1)
+ , Product (Rational_le (Big_int s1'), y2) ) )
| Product (y, z) ->
- let s1,y1 = scale_certificate y and s2,y2 = scale_certificate z in
- mult_big_int s1 s2 , Product (y1,y2)
-
+ let 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)
-
-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 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)
+
+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)
-
+ | i :: l -> Mc.PsatzMulE (Mc.PsatzIn (Ml2C.nat i), product l)
-let q_cert_of_pos pos =
+let q_cert_of_pos pos =
let rec _cert_of_pos = function
- 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
+ | 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)
- | 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
+ 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)
-
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)
-
-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
+ | 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)
+
+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_le i -> Mc.PsatzIn (Ml2C.nat i)
- | Axiom_lt i -> Mc.PsatzIn (Ml2C.nat i)
- | Monoid l -> product l
+ | 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))
- | Square t -> Mc.PsatzSquare (term_to_z_pol t)
+ 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)
- | 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
+ 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)
+open Mutils
(** All constraints (initial or derived) have an index and have a justification i.e., proof.
Given a constraint, all the coefficients are always integers.
*)
-open Mutils
+
open Num
open Big_int
open Polynomial
-
-
type prf_sys = (cstr * ProofFormat.prf_rule) list
-
-
(** 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 pivot v (c1, p1) (c2, p2) =
+ 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 = opAdd op1 op2 ;
- cst = n1 */ cv1 +/ n2 */ cv2 },
-
- ProofFormat.add_proof (ProofFormat.mul_cst_proof cv1 p1) (ProofFormat.mul_cst_proof cv2 p2)) in
-
- match Vect.get v v1 , Vect.get v v2 with
- | Int 0 , _ | _ , Int 0 -> None
- | a , 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 *)
-
+ ( { coeffs = Vect.add (Vect.mul cv1 v1) (Vect.mul cv2 v2)
+ ; op = opAdd op1 op2
+ ; cst = (n1 */ cv1) +/ (n2 */ cv2) }
+ , ProofFormat.add_proof
+ (ProofFormat.mul_cst_proof cv1 p1)
+ (ProofFormat.mul_cst_proof cv2 p2) )
+ in
+ match (Vect.get v v1, Vect.get v v2) with
+ | Int 0, _ | _, Int 0 -> None
+ | a, 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 *)
let simpl_sys sys =
- List.fold_left (fun acc (c,p) ->
- match check_int_sat (c,p) with
+ List.fold_left
+ (fun acc (c, p) ->
+ match check_int_sat (c, p) with
| Tauto -> acc
| Unsat prf -> raise (FoundProof prf)
- | Cut(c,p) -> (c,p)::acc
- | Normalise (c,p) -> (c,p)::acc) [] sys
-
+ | Cut (c, p) -> (c, p) :: acc
+ | Normalise (c, p) -> (c, p) :: acc)
+ [] sys
(** [ext_gcd a b] is the extended Euclid algorithm.
[ext_gcd a b = (x,y,g)] iff [ax+by=g]
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)
+ 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
+ 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 extract_coprime (c1,p1) (c2,p2) =
- if c1.op == Eq && c2.op == Eq
- then Vect.exists2 (fun n1 n2 ->
- Int.equal (compare_big_int (gcd_big_int (numerator n1) (numerator n2)) unit_big_int) 0)
- c1.coeffs c2.coeffs
+let extract_coprime (c1, p1) (c2, p2) =
+ if c1.op == Eq && c2.op == Eq then
+ Vect.exists2
+ (fun n1 n2 ->
+ Int.equal
+ (compare_big_int
+ (gcd_big_int (numerator n1) (numerator n2))
+ unit_big_int)
+ 0)
+ 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
-
+ | [] -> (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
-
-
-
-
-
-
+let extract_coprime_equation psys = extract2 extract_coprime psys
let pivot_sys v pc psys = apply_and_normalise check_int_sat (pivot v pc) psys
let reduce_coprime psys =
- let oeq,sys = extract_coprime_equation psys in
+ 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 = ProofFormat.add_proof (ProofFormat.mul_cst_proof l1' p1) (ProofFormat.mul_cst_proof l2' p2) in
-
- Some (pivot_sys v (cstr,prf) ((c1,p1)::sys))
+ | 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 =
+ ProofFormat.add_proof
+ (ProofFormat.mul_cst_proof l1' p1)
+ (ProofFormat.mul_cst_proof 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
- Vect.find (fun v n -> if n =/ (Int 1) || n=/ (Int (-1)) then Some v else None) cstr.coeffs
- else None in
-
- let (oeq,sys) = extract is_unary_equation psys in
+ let is_unary_equation (cstr, prf) =
+ if cstr.op == Eq then
+ Vect.find
+ (fun v n -> if n =/ Int 1 || n =/ Int (-1) then Some v else None)
+ cstr.coeffs
+ 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)
-
+ | Some (v, pc) -> Some (pivot_sys v pc sys)
let reduce_var_change psys =
-
let rec rel_prime vect =
match Vect.choose vect with
| None -> None
- | Some(x,v,vect) ->
- let v = numerator v in
- match Vect.find (fun x' v' ->
- let v' = numerator v' in
- if eq_big_int (gcd_big_int v v') unit_big_int
- then Some(x',v') else None) vect with
- | Some(x',v') -> Some ((x,v),(x', v'))
- | None -> 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
-
+ | Some (x, v, vect) -> (
+ let v = numerator v in
+ match
+ Vect.find
+ (fun x' v' ->
+ let v' = numerator v' in
+ if eq_big_int (gcd_big_int v v') unit_big_int then Some (x', v')
+ else None)
+ vect
+ with
+ | Some (x', v') -> Some ((x, v), (x', v'))
+ | None -> 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 pivot_eq (c',p') =
- let {coeffs = coeffs ; op = op ; cst = cst} = c' in
- let vx = Vect.get x coeffs in
- let vx' = Vect.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} ,
- ProofFormat.add_proof (ProofFormat.mul_cst_proof m p) p') in
-
- Some (apply_and_normalise check_int_sat pivot_eq sys)
-
+ | 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; op; cst} = c' in
+ let vx = Vect.get x coeffs in
+ let vx' = Vect.get x' coeffs in
+ let m = minus_num ((vx */ l1) +/ (vx' */ l2)) in
+ Some
+ ( { coeffs = Vect.add (Vect.mul m c.coeffs) coeffs
+ ; op
+ ; cst = (m */ c.cst) +/ cst }
+ , ProofFormat.add_proof (ProofFormat.mul_cst_proof m p) p' )
+ in
+ Some (apply_and_normalise check_int_sat pivot_eq sys)
let reduction_equations psys =
- iterate_until_stable (app_funs
- [reduce_unary ; reduce_coprime ;
- reduce_var_change (*; reduce_pivot*)]) psys
-
-
-
-
+ iterate_until_stable
+ (app_funs
+ [reduce_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 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
+ 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
-
+ | _ ->
+ 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
+ if is_small acc then acc
else
match optimise vect sys with
| None -> acc
| Some i ->
- if debug then Printf.printf "Found a new bound %a in %a" Vect.pp vect Itv.pp i;
- select_best (vect,i) acc) (Vect.null, (None,None)) (all_planes sys) in
+ if debug then
+ Printf.printf "Found a new bound %a in %a" Vect.pp vect Itv.pp i;
+ 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 *)
+ 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 *)
- direct_linear_prover ({coeffs = Vect.mul (Big_int ubd) e ; op = Ge ; cst = Big_int ubn} :: sys),
- (* lb <= x -> lb > x *)
+ | 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
+ (* x <= ub -> x > ub *)
+ match
+ ( direct_linear_prover
+ ( {coeffs = Vect.mul (Big_int ubd) e; op = Ge; cst = Big_int ubn}
+ :: sys )
+ , (* lb <= x -> lb > x *)
direct_linear_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 (Vect.to_list clb),(lb,e,ub), List.tl (Vect.to_list cub))
- | _ -> failwith "Interval without proof"
- )
+ ( { 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 (Vect.to_list clb), (lb, e, ub), List.tl (Vect.to_list cub))
+ | _ -> failwith "Interval without proof" )
| None -> None
-
let check_sys sys =
- List.for_all (fun (c,p) -> Vect.for_all (fun _ n -> sign_num n <> 0) c.coeffs) sys
+ List.for_all
+ (fun (c, p) -> Vect.for_all (fun _ n -> sign_num n <> 0) c.coeffs)
+ sys
open ProofFormat
-let xlia (can_enum:bool) reduction_equations sys =
-
-
- let rec enum_proof (id:int) (sys:prf_sys) =
- if debug then (Printf.printf "enum_proof\n" ; flush stdout) ;
- assert (check_sys sys) ;
-
- let nsys,prf = List.split sys in
+let xlia (can_enum : bool) reduction_equations sys =
+ let rec enum_proof (id : int) (sys : prf_sys) =
+ 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 -> Unknown (* Is the systeme really unbounded ? *)
- | Some(prf1,(lb,e,ub),prf2) ->
- if debug then Printf.printf "Found interval: %a in [%s;%s] -> " Vect.pp e (string_of_num lb) (string_of_num ub) ;
- (match start_enum id e (ceiling_num lb) (floor_num ub) sys
- with
- | Prf prfl ->
- Prf(ProofFormat.Enum(id,ProofFormat.proof_of_farkas (env_of_list prf) (Vect.from_list prf1),e,
- ProofFormat.proof_of_farkas (env_of_list prf) (Vect.from_list prf2),prfl))
- | _ -> Unknown
- )
-
- and start_enum id e clb cub sys =
- if clb >/ cub
- then Prf []
+ | Some (prf1, (lb, e, ub), prf2) -> (
+ if debug then
+ Printf.printf "Found interval: %a in [%s;%s] -> " Vect.pp e
+ (string_of_num lb) (string_of_num ub);
+ match start_enum id e (ceiling_num lb) (floor_num ub) sys with
+ | Prf prfl ->
+ Prf
+ (ProofFormat.Enum
+ ( id
+ , ProofFormat.proof_of_farkas (env_of_list prf)
+ (Vect.from_list prf1)
+ , e
+ , ProofFormat.proof_of_farkas (env_of_list prf)
+ (Vect.from_list prf2)
+ , prfl ))
+ | _ -> Unknown )
+ and start_enum id e clb cub sys =
+ if clb >/ cub then Prf []
else
- let eq = {coeffs = e ; op = Eq ; cst = clb} in
- match aux_lia (id+1) ((eq, ProofFormat.Def id) :: sys) with
+ let eq = {coeffs = e; op = Eq; cst = clb} in
+ match aux_lia (id + 1) ((eq, ProofFormat.Def id) :: sys) with
| Unknown | Model _ -> Unknown
- | Prf prf ->
- match start_enum id e (clb +/ (Int 1)) cub sys with
- | Prf l -> Prf (prf::l)
- | _ -> Unknown
-
-
- and aux_lia (id:int) (sys:prf_sys) =
- assert (check_sys sys) ;
- if debug then Printf.printf "xlia: %a \n" (pp_list ";" (fun o (c,_) -> output_cstr o c)) sys ;
+ | Prf prf -> (
+ match start_enum id e (clb +/ Int 1) cub sys with
+ | Prf l -> Prf (prf :: l)
+ | _ -> Unknown )
+ and aux_lia (id : int) (sys : prf_sys) =
+ 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 ;
+ Printf.printf "after reduction: %a \n"
+ (pp_list ";" (fun o (c, _) -> output_cstr o c))
+ sys;
match linear_prover_cstr sys with
- | Some prf -> Prf (Step(id,prf,Done))
- | None -> if can_enum then enum_proof id sys else Unknown
+ | Some prf -> Prf (Step (id, prf, Done))
+ | None -> if can_enum then enum_proof id sys else Unknown
with FoundProof prf ->
(* [reduction_equations] can find a proof *)
- Prf(Step(id,prf,Done)) in
-
+ Prf (Step (id, prf, Done))
+ in
(* let sys' = List.map (fun (p,o) -> Mc.norm0 p , o) sys in*)
- let id = 1 + (List.fold_left
- (fun acc (_,r) -> max acc (ProofFormat.pr_rule_max_id r)) 0 sys) in
+ let id =
+ 1
+ + List.fold_left
+ (fun acc (_, r) -> max acc (ProofFormat.pr_rule_max_id r))
+ 0 sys
+ in
let orpf =
try
let sys = simpl_sys sys in
aux_lia id sys
- with FoundProof pr -> Prf(Step(id,pr,Done)) in
+ with FoundProof pr -> Prf (Step (id, pr, Done))
+ in
match orpf with
| Unknown | Model _ -> Unknown
| Prf prf ->
- let env = CList.interval 0 (id - 1) in
- if debug then begin
- Printf.fprintf stdout "direct proof %a\n" output_proof prf;
- flush stdout;
- end;
- let prf = compile_proof env prf in
- (*try
+ let env = CList.interval 0 (id - 1) in
+ if debug then begin
+ Printf.fprintf stdout "direct proof %a\n" output_proof prf;
+ flush stdout
+ end;
+ 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)
- *) Prf prf
+ *)
+ Prf prf
let xlia_simplex env red sys =
let compile_prf sys prf =
- let id = 1 + (List.fold_left
- (fun acc (_,r) -> max acc (ProofFormat.pr_rule_max_id r)) 0 sys) in
+ let id =
+ 1
+ + List.fold_left
+ (fun acc (_, r) -> max acc (ProofFormat.pr_rule_max_id r))
+ 0 sys
+ in
let env = CList.interval 0 (id - 1) in
- Prf (compile_proof env prf) in
-
+ Prf (compile_proof env prf)
+ in
try
let sys = red sys in
-
match Simplex.integer_solver sys with
| None -> Unknown
| Some prf -> compile_prf sys prf
- with FoundProof prf -> compile_prf sys (Step(0,prf,Done))
+ with FoundProof prf -> compile_prf sys (Step (0, prf, Done))
let xlia env0 en red sys =
- if !use_simplex then xlia_simplex env0 red sys
- else xlia en red sys
-
+ if !use_simplex then xlia_simplex env0 red sys else xlia en red sys
let dump_file = ref None
let gen_bench (tac, prover) can_enum prfdepth sys =
let res = prover can_enum prfdepth sys in
- (match !dump_file with
+ ( match !dump_file with
| None -> ()
| Some file ->
- begin
- let o = open_out (Filename.temp_file ~temp_dir:(Sys.getcwd ()) file ".v") in
- let sys = develop_constraints prfdepth z_spec sys in
- Printf.fprintf o "Require Import ZArith Lia. Open Scope Z_scope.\n";
- Printf.fprintf o "Goal %a.\n" (LinPoly.pp_goal "Z") (List.map fst sys) ;
- begin
- match res with
- | Unknown | Model _ ->
- Printf.fprintf o "Proof.\n intros. Fail %s.\nAbort.\n" tac
- | Prf res ->
- Printf.fprintf o "Proof.\n intros. %s.\nQed.\n" tac
- end
- ;
- flush o ;
- close_out o ;
- end);
+ let o = open_out (Filename.temp_file ~temp_dir:(Sys.getcwd ()) file ".v") in
+ let sys = develop_constraints prfdepth z_spec sys in
+ Printf.fprintf o "Require Import ZArith Lia. Open Scope Z_scope.\n";
+ Printf.fprintf o "Goal %a.\n" (LinPoly.pp_goal "Z") (List.map fst sys);
+ begin
+ match res with
+ | Unknown | Model _ ->
+ Printf.fprintf o "Proof.\n intros. Fail %s.\nAbort.\n" tac
+ | Prf res -> Printf.fprintf o "Proof.\n intros. %s.\nQed.\n" tac
+ end;
+ flush o; close_out o );
res
-let lia (can_enum:bool) (prfdepth:int) sys =
+let lia (can_enum : bool) (prfdepth : int) sys =
let sys = develop_constraints prfdepth z_spec sys in
if debug then begin
- Printf.fprintf stdout "Input problem\n";
- List.iter (fun s -> Printf.fprintf stdout "%a\n" WithProof.output s) sys;
- Printf.fprintf stdout "Input problem\n";
- let string_of_op = function Eq -> "=" | Ge -> ">=" | Gt -> ">" in
- List.iter (fun ((p,op),_) -> Printf.fprintf stdout "(assert (%s %a))\n" (string_of_op op) Vect.pp_smt p) sys;
- end;
+ Printf.fprintf stdout "Input problem\n";
+ List.iter (fun s -> Printf.fprintf stdout "%a\n" WithProof.output s) sys;
+ Printf.fprintf stdout "Input problem\n";
+ let string_of_op = function Eq -> "=" | Ge -> ">=" | Gt -> ">" in
+ List.iter
+ (fun ((p, op), _) ->
+ Printf.fprintf stdout "(assert (%s %a))\n" (string_of_op op) Vect.pp_smt
+ p)
+ sys
+ end;
let sys = subst sys in
- let bnd = bound_monomials sys in (* To deal with non-linear monomials *)
- let sys = bnd@(saturate_by_linear_equalities sys)@sys in
-
-
- let sys' = List.map (fun ((p,o),prf) -> (cstr_of_poly (p,o), prf)) sys in
+ let bnd = bound_monomials sys in
+ (* To deal with non-linear monomials *)
+ let sys = bnd @ saturate_by_linear_equalities sys @ sys in
+ let sys' = List.map (fun ((p, o), prf) -> (cstr_of_poly (p, o), prf)) sys in
xlia (List.map fst sys) can_enum reduction_equations sys'
let make_cstr_system sys =
- List.map (fun ((p,o),prf) -> (cstr_of_poly (p,o), prf)) sys
+ List.map (fun ((p, o), prf) -> (cstr_of_poly (p, o), prf)) sys
let nlia enum prfdepth sys =
let sys = develop_constraints prfdepth z_spec sys in
- let is_linear = List.for_all (fun ((p,_),_) -> LinPoly.is_linear p) sys in
-
+ let is_linear = List.for_all (fun ((p, _), _) -> LinPoly.is_linear p) sys in
if debug then begin
- Printf.fprintf stdout "Input problem\n";
- List.iter (fun s -> Printf.fprintf stdout "%a\n" WithProof.output s) sys;
- end;
-
- if is_linear
- then xlia (List.map fst sys) enum reduction_equations (make_cstr_system sys)
+ Printf.fprintf stdout "Input problem\n";
+ List.iter (fun s -> Printf.fprintf stdout "%a\n" WithProof.output s) sys
+ end;
+ if is_linear then
+ xlia (List.map fst sys) enum reduction_equations (make_cstr_system sys)
else
(*
let sys1 = elim_every_substitution sys in
@@ -1068,23 +1052,15 @@ let nlia enum prfdepth sys =
*)
let sys1 = elim_simple_linear_equality sys in
let sys2 = saturate_by_linear_equalities sys1 in
- let sys3 = nlinear_preprocess (sys1@sys2) in
-
- let sys4 = make_cstr_system ((*sys2@*)sys3) in
+ let sys3 = nlinear_preprocess (sys1 @ sys2) in
+ let sys4 = make_cstr_system (*sys2@*) sys3 in
(* [reduction_equations] is too brutal - there should be some non-linear reasoning *)
- xlia (List.map fst sys) enum reduction_equations sys4
+ xlia (List.map fst sys) enum reduction_equations sys4
(* For regression testing, if bench = true generate a Coq goal *)
-let lia can_enum prfdepth sys =
- gen_bench ("lia",lia) can_enum prfdepth sys
-
-let nlia enum prfdepth sys =
- gen_bench ("nia",nlia) enum prfdepth sys
-
-
-
-
+let lia can_enum prfdepth sys = gen_bench ("lia", lia) can_enum prfdepth sys
+let nlia enum prfdepth sys = gen_bench ("nia", nlia) enum prfdepth sys
(* Local Variables: *)
(* coding: utf-8 *)
diff --git a/plugins/micromega/certificate.mli b/plugins/micromega/certificate.mli
index cd26b72a27..a8cc595ddf 100644
--- a/plugins/micromega/certificate.mli
+++ b/plugins/micromega/certificate.mli
@@ -10,42 +10,36 @@
module Mc = Micromega
-
+val use_simplex : bool ref
(** [use_simplex] is bound to the Coq option Simplex.
If set, use the Simplex method, otherwise use Fourier *)
-val use_simplex : bool ref
-
-type ('prf,'model) res =
- | Prf of 'prf
- | Model of 'model
- | Unknown
-
-type zres = (Mc.zArithProof , (int * Mc.z list)) res
-type qres = (Mc.q Mc.psatz , (int * Mc.q list)) res
+type ('prf, 'model) res = Prf of 'prf | Model of 'model | Unknown
+type zres = (Mc.zArithProof, int * Mc.z list) res
+type qres = (Mc.q Mc.psatz, int * Mc.q list) res
+val dump_file : string option ref
(** [dump_file] is bound to the Coq option Dump Arith.
If set to some [file], arithmetic goals are dumped in filexxx.v *)
-val dump_file : string option ref
-(** [q_cert_of_pos prf] converts a Sos proof into a rational Coq proof *)
val q_cert_of_pos : Sos_types.positivstellensatz -> Mc.q Mc.psatz
+(** [q_cert_of_pos prf] converts a Sos proof into a rational Coq proof *)
-(** [z_cert_of_pos prf] converts a Sos proof into an integer Coq proof *)
val z_cert_of_pos : Sos_types.positivstellensatz -> Mc.z Mc.psatz
+(** [z_cert_of_pos prf] converts a Sos proof into an integer Coq proof *)
+val lia : bool -> int -> (Mc.z Mc.pExpr * Mc.op1) list -> zres
(** [lia enum depth sys] generates an unsat proof for the linear constraints in [sys].
If the Simplex option is set, any failure to find a proof should be considered as a bug. *)
-val lia : bool -> int -> (Mc.z Mc.pExpr * Mc.op1) list -> zres
+val nlia : bool -> int -> (Mc.z Mc.pExpr * Mc.op1) list -> zres
(** [nlia enum depth sys] generates an unsat proof for the non-linear constraints in [sys].
The solver is incomplete -- the problem is undecidable *)
-val nlia : bool -> int -> (Mc.z Mc.pExpr * Mc.op1) list -> zres
+val linear_prover_with_cert : int -> (Mc.q Mc.pExpr * Mc.op1) list -> qres
(** [linear_prover_with_cert depth sys] generates an unsat proof for the linear constraints in [sys].
Over the rationals, the solver is complete. *)
-val linear_prover_with_cert : int -> (Mc.q Mc.pExpr * Mc.op1) list -> qres
+val nlinear_prover : int -> (Mc.q Mc.pExpr * Mc.op1) list -> qres
(** [nlinear depth sys] generates an unsat proof for the non-linear constraints in [sys].
The solver is incompete -- the problem is decidable. *)
-val nlinear_prover : int -> (Mc.q Mc.pExpr * Mc.op1) list -> qres
diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml
index 1772a3c333..f14db40d45 100644
--- a/plugins/micromega/coq_micromega.ml
+++ b/plugins/micromega/coq_micromega.ml
@@ -39,16 +39,11 @@ 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_enum = ref true
let lia_proof_depth = ref max_depth
-
-let get_lia_option () =
- (!Certificate.use_simplex,!lia_enum,!lia_proof_depth)
-
-let get_lra_option () =
- !lra_proof_depth
+let get_lia_option () = (!Certificate.use_simplex, !lia_enum, !lia_proof_depth)
+let get_lra_option () = !lra_proof_depth
(* Enable/disable caches *)
@@ -58,87 +53,72 @@ let use_nra_cache = ref true
let use_csdp_cache = ref true
let () =
-
- let int_opt l vref =
- {
- 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 =
- {
- optdepr = false;
- optname = "Lia Enum";
- optkey = ["Lia";"Enum"];
- optread = (fun () -> !lia_enum);
- optwrite = (fun x -> lia_enum := x)
- } in
-
- let solver_opt =
- {
- optdepr = false;
- optname = "Use the Simplex instead of Fourier elimination";
- optkey = ["Simplex"];
- optread = (fun () -> !Certificate.use_simplex);
- optwrite = (fun x -> Certificate.use_simplex := x)
- } in
-
- let dump_file_opt =
- {
- optdepr = false;
- optname = "Generate Coq goals in file from calls to 'lia' 'nia'";
- optkey = ["Dump"; "Arith"];
- optread = (fun () -> !Certificate.dump_file);
- optwrite = (fun x -> Certificate.dump_file := x)
- } in
-
- let lia_cache_opt =
- {
- optdepr = false;
- optname = "cache of lia (.lia.cache)";
- optkey = ["Lia" ; "Cache"];
- optread = (fun () -> !use_lia_cache);
- optwrite = (fun x -> use_lia_cache := x)
- } in
-
- let nia_cache_opt =
- {
- optdepr = false;
- optname = "cache of nia (.nia.cache)";
- optkey = ["Nia" ; "Cache"];
- optread = (fun () -> !use_nia_cache);
- optwrite = (fun x -> use_nia_cache := x)
- } in
-
- let nra_cache_opt =
- {
- optdepr = false;
- optname = "cache of nra (.nra.cache)";
- optkey = ["Nra" ; "Cache"];
- optread = (fun () -> !use_nra_cache);
- optwrite = (fun x -> use_nra_cache := x)
- } in
-
-
- let () = declare_bool_option solver_opt in
- let () = declare_bool_option lia_cache_opt in
- let () = declare_bool_option nia_cache_opt in
- let () = declare_bool_option nra_cache_opt in
- let () = declare_stringopt_option dump_file_opt 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
- ()
-
+ let int_opt l vref =
+ { 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 =
+ { optdepr = false
+ ; optname = "Lia Enum"
+ ; optkey = ["Lia"; "Enum"]
+ ; optread = (fun () -> !lia_enum)
+ ; optwrite = (fun x -> lia_enum := x) }
+ in
+ let solver_opt =
+ { optdepr = false
+ ; optname = "Use the Simplex instead of Fourier elimination"
+ ; optkey = ["Simplex"]
+ ; optread = (fun () -> !Certificate.use_simplex)
+ ; optwrite = (fun x -> Certificate.use_simplex := x) }
+ in
+ let dump_file_opt =
+ { optdepr = false
+ ; optname = "Generate Coq goals in file from calls to 'lia' 'nia'"
+ ; optkey = ["Dump"; "Arith"]
+ ; optread = (fun () -> !Certificate.dump_file)
+ ; optwrite = (fun x -> Certificate.dump_file := x) }
+ in
+ let lia_cache_opt =
+ { optdepr = false
+ ; optname = "cache of lia (.lia.cache)"
+ ; optkey = ["Lia"; "Cache"]
+ ; optread = (fun () -> !use_lia_cache)
+ ; optwrite = (fun x -> use_lia_cache := x) }
+ in
+ let nia_cache_opt =
+ { optdepr = false
+ ; optname = "cache of nia (.nia.cache)"
+ ; optkey = ["Nia"; "Cache"]
+ ; optread = (fun () -> !use_nia_cache)
+ ; optwrite = (fun x -> use_nia_cache := x) }
+ in
+ let nra_cache_opt =
+ { optdepr = false
+ ; optname = "cache of nra (.nra.cache)"
+ ; optkey = ["Nra"; "Cache"]
+ ; optread = (fun () -> !use_nra_cache)
+ ; optwrite = (fun x -> use_nra_cache := x) }
+ in
+ let () = declare_bool_option solver_opt in
+ let () = declare_bool_option lia_cache_opt in
+ let () = declare_bool_option nia_cache_opt in
+ let () = declare_bool_option nra_cache_opt in
+ let () = declare_stringopt_option dump_file_opt 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).
*)
type tag = Tag.t
+
module Mc = Micromega
(**
@@ -150,29 +130,26 @@ module Mc = Micromega
type 'cst atom = 'cst Mc.formula
-type 'cst formula = ('cst atom, EConstr.constr,tag * EConstr.constr,Names.Id.t) Mc.gFormula
+type 'cst formula =
+ ('cst atom, EConstr.constr, tag * EConstr.constr, Names.Id.t) Mc.gFormula
type 'cst clause = ('cst Mc.nFormula, tag * EConstr.constr) Mc.clause
type 'cst cnf = ('cst Mc.nFormula, tag * EConstr.constr) Mc.cnf
-
-let rec pp_formula o (f:'cst formula) =
+let rec pp_formula o (f : 'cst formula) =
Mc.(
- match f with
- | TT -> output_string o "tt"
- | FF -> output_string o "ff"
+ match f with
+ | TT -> output_string o "tt"
+ | FF -> output_string o "ff"
| X c -> output_string o "X "
- | A(_,(t,_)) -> Printf.fprintf o "A(%a)" Tag.pp t
- | Cj(f1,f2) -> Printf.fprintf o "C(%a,%a)" pp_formula f1 pp_formula f2
- | D(f1,f2) -> Printf.fprintf o "D(%a,%a)" pp_formula f1 pp_formula f2
- | I(f1,n,f2) -> Printf.fprintf o "I(%a,%s,%a)"
- pp_formula f1
- (match n with
- | Some id -> Names.Id.to_string id
- | None -> "") pp_formula f2
- | N(f) -> Printf.fprintf o "N(%a)" pp_formula f
- )
-
+ | A (_, (t, _)) -> Printf.fprintf o "A(%a)" Tag.pp t
+ | Cj (f1, f2) -> Printf.fprintf o "C(%a,%a)" pp_formula f1 pp_formula f2
+ | D (f1, f2) -> Printf.fprintf o "D(%a,%a)" pp_formula f1 pp_formula f2
+ | I (f1, n, f2) ->
+ Printf.fprintf o "I(%a,%s,%a)" pp_formula f1
+ (match n with Some id -> Names.Id.to_string id | None -> "")
+ pp_formula f2
+ | N f -> Printf.fprintf o "N(%a)" pp_formula f)
(**
* Given a set of integers s=\{i0,...,iN\} and a list m, return the list of
@@ -182,9 +159,11 @@ let rec pp_formula o (f:'cst formula) =
let selecti s m =
let rec xselecti i m =
match m with
- | [] -> []
- | e::m -> if ISet.mem i s then e::(xselecti (i+1) m) else xselecti (i+1) m in
- xselecti 0 m
+ | [] -> []
+ | e :: m ->
+ if ISet.mem i s then e :: xselecti (i + 1) m else xselecti (i + 1) m
+ in
+ xselecti 0 m
(**
* MODULE: Mapping of the Coq data-strustures into Caml and Caml extracted
@@ -194,57 +173,62 @@ let selecti s m =
* Opened here and in csdpcert.ml.
*)
-module M =
-struct
-
+(**
+ * MODULE END: M
+ *)
+module M = struct
(**
* Location of the Coq libraries.
*)
- let logic_dir = ["Coq";"Logic";"Decidable"]
+ let logic_dir = ["Coq"; "Logic"; "Decidable"]
let mic_modules =
- [
- ["Coq";"Lists";"List"];
- ["Coq"; "micromega";"ZMicromega"];
- ["Coq"; "micromega";"Tauto"];
- ["Coq"; "micromega"; "DeclConstant"];
- ["Coq"; "micromega";"RingMicromega"];
- ["Coq"; "micromega";"EnvRing"];
- ["Coq"; "micromega"; "ZMicromega"];
- ["Coq"; "micromega"; "RMicromega"];
- ["Coq" ; "micromega" ; "Tauto"];
- ["Coq" ; "micromega" ; "RingMicromega"];
- ["Coq" ; "micromega" ; "EnvRing"];
- ["Coq";"QArith"; "QArith_base"];
- ["Coq";"Reals" ; "Rdefinitions"];
- ["Coq";"Reals" ; "Rpow_def"];
- ["LRing_normalise"]]
-
-[@@@ocaml.warning "-3"]
+ [ ["Coq"; "Lists"; "List"]
+ ; ["Coq"; "micromega"; "ZMicromega"]
+ ; ["Coq"; "micromega"; "Tauto"]
+ ; ["Coq"; "micromega"; "DeclConstant"]
+ ; ["Coq"; "micromega"; "RingMicromega"]
+ ; ["Coq"; "micromega"; "EnvRing"]
+ ; ["Coq"; "micromega"; "ZMicromega"]
+ ; ["Coq"; "micromega"; "RMicromega"]
+ ; ["Coq"; "micromega"; "Tauto"]
+ ; ["Coq"; "micromega"; "RingMicromega"]
+ ; ["Coq"; "micromega"; "EnvRing"]
+ ; ["Coq"; "QArith"; "QArith_base"]
+ ; ["Coq"; "Reals"; "Rdefinitions"]
+ ; ["Coq"; "Reals"; "Rpow_def"]
+ ; ["LRing_normalise"] ]
+
+ [@@@ocaml.warning "-3"]
let coq_modules =
- Coqlib.(init_modules @
- [logic_dir] @ arith_modules @ zarith_base_modules @ mic_modules)
+ Coqlib.(
+ init_modules @ [logic_dir] @ arith_modules @ zarith_base_modules
+ @ mic_modules)
- let bin_module = [["Coq";"Numbers";"BinNums"]]
+ let bin_module = [["Coq"; "Numbers"; "BinNums"]]
let r_modules =
- [["Coq";"Reals" ; "Rdefinitions"];
- ["Coq";"Reals" ; "Rpow_def"] ;
- ["Coq";"Reals" ; "Raxioms"] ;
- ["Coq";"QArith"; "Qreals"] ;
- ]
+ [ ["Coq"; "Reals"; "Rdefinitions"]
+ ; ["Coq"; "Reals"; "Rpow_def"]
+ ; ["Coq"; "Reals"; "Raxioms"]
+ ; ["Coq"; "QArith"; "Qreals"] ]
- let z_modules = [["Coq";"ZArith";"BinInt"]]
+ let z_modules = [["Coq"; "ZArith"; "BinInt"]]
(**
* Initialization : a large amount of Caml symbols are derived from
* ZMicromega.v
*)
- let gen_constant_in_modules s m n = EConstr.of_constr (UnivGen.constr_of_monomorphic_global @@ Coqlib.gen_reference_in_modules s m n)
+ let gen_constant_in_modules s m n =
+ EConstr.of_constr
+ ( UnivGen.constr_of_monomorphic_global
+ @@ Coqlib.gen_reference_in_modules s m n )
+
let init_constant = gen_constant_in_modules "ZMicromega" Coqlib.init_modules
+
[@@@ocaml.warning "+3"]
let constant = gen_constant_in_modules "ZMicromega" coq_modules
@@ -252,98 +236,77 @@ struct
let r_constant = gen_constant_in_modules "ZMicromega" r_modules
let z_constant = gen_constant_in_modules "ZMicromega" z_modules
let m_constant = gen_constant_in_modules "ZMicromega" mic_modules
-
let coq_and = lazy (init_constant "and")
let coq_or = lazy (init_constant "or")
let coq_not = lazy (init_constant "not")
-
let coq_iff = lazy (init_constant "iff")
let coq_True = lazy (init_constant "True")
let coq_False = lazy (init_constant "False")
-
let coq_cons = lazy (constant "cons")
let coq_nil = lazy (constant "nil")
let coq_list = lazy (constant "list")
-
let coq_O = lazy (init_constant "O")
let coq_S = lazy (init_constant "S")
-
let coq_nat = lazy (init_constant "nat")
let coq_unit = lazy (init_constant "unit")
+
(* let coq_option = lazy (init_constant "option")*)
let coq_None = lazy (init_constant "None")
let coq_tt = lazy (init_constant "tt")
let coq_Inl = lazy (init_constant "inl")
let coq_Inr = lazy (init_constant "inr")
-
-
let coq_N0 = lazy (bin_constant "N0")
let coq_Npos = lazy (bin_constant "Npos")
-
let coq_xH = lazy (bin_constant "xH")
let coq_xO = lazy (bin_constant "xO")
let coq_xI = lazy (bin_constant "xI")
-
let coq_Z = lazy (bin_constant "Z")
let coq_ZERO = lazy (bin_constant "Z0")
let coq_POS = lazy (bin_constant "Zpos")
let coq_NEG = lazy (bin_constant "Zneg")
-
let coq_Q = lazy (constant "Q")
let coq_R = lazy (constant "R")
-
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")
- let coq_CZ = lazy (m_constant "CZ")
+ let coq_C0 = lazy (m_constant "C0")
+ let coq_C1 = lazy (m_constant "C1")
+ let coq_CQ = lazy (m_constant "CQ")
+ let coq_CZ = lazy (m_constant "CZ")
let coq_CPlus = lazy (m_constant "CPlus")
let coq_CMinus = lazy (m_constant "CMinus")
- let coq_CMult = lazy (m_constant "CMult")
- let coq_CPow = lazy (m_constant "CPow")
- let coq_CInv = lazy (m_constant "CInv")
- let coq_COpp = lazy (m_constant "COpp")
-
-
- let coq_R0 = lazy (constant "R0")
- let coq_R1 = lazy (constant "R1")
-
+ let coq_CMult = lazy (m_constant "CMult")
+ let coq_CPow = lazy (m_constant "CPow")
+ let coq_CInv = lazy (m_constant "CInv")
+ let coq_COpp = lazy (m_constant "COpp")
+ let coq_R0 = lazy (constant "R0")
+ let coq_R1 = lazy (constant "R1")
let coq_proofTerm = lazy (constant "ZArithProof")
let coq_doneProof = lazy (constant "DoneProof")
let coq_ratProof = lazy (constant "RatProof")
let coq_cutProof = lazy (constant "CutProof")
let coq_enumProof = lazy (constant "EnumProof")
-
let coq_Zgt = lazy (z_constant "Z.gt")
let coq_Zge = lazy (z_constant "Z.ge")
let coq_Zle = lazy (z_constant "Z.le")
let coq_Zlt = lazy (z_constant "Z.lt")
- let coq_Eq = lazy (init_constant "eq")
-
+ let coq_Eq = lazy (init_constant "eq")
let coq_Zplus = lazy (z_constant "Z.add")
let coq_Zminus = lazy (z_constant "Z.sub")
let coq_Zopp = lazy (z_constant "Z.opp")
let coq_Zmult = lazy (z_constant "Z.mul")
let coq_Zpower = lazy (z_constant "Z.pow")
-
let coq_Qle = lazy (constant "Qle")
let coq_Qlt = lazy (constant "Qlt")
let coq_Qeq = lazy (constant "Qeq")
-
let coq_Qplus = lazy (constant "Qplus")
let coq_Qminus = lazy (constant "Qminus")
let coq_Qopp = lazy (constant "Qopp")
let coq_Qmult = lazy (constant "Qmult")
let coq_Qpower = lazy (constant "Qpower")
-
let coq_Rgt = lazy (r_constant "Rgt")
let coq_Rge = lazy (r_constant "Rge")
let coq_Rle = lazy (r_constant "Rle")
let coq_Rlt = lazy (r_constant "Rlt")
-
let coq_Rplus = lazy (r_constant "Rplus")
let coq_Rminus = lazy (r_constant "Rminus")
let coq_Ropp = lazy (r_constant "Ropp")
@@ -351,85 +314,112 @@ struct
let coq_Rinv = lazy (r_constant "Rinv")
let coq_Rpower = lazy (r_constant "pow")
let coq_powerZR = lazy (r_constant "powerRZ")
- let coq_IZR = lazy (r_constant "IZR")
- let coq_IQR = lazy (r_constant "Q2R")
-
-
- let coq_PEX = lazy (constant "PEX" )
- let coq_PEc = lazy (constant"PEc")
+ let coq_IZR = lazy (r_constant "IZR")
+ let coq_IQR = lazy (r_constant "Q2R")
+ let coq_PEX = lazy (constant "PEX")
+ let coq_PEc = lazy (constant "PEc")
let coq_PEadd = lazy (constant "PEadd")
let coq_PEopp = lazy (constant "PEopp")
let coq_PEmul = lazy (constant "PEmul")
let coq_PEsub = lazy (constant "PEsub")
let coq_PEpow = lazy (constant "PEpow")
-
- let coq_PX = lazy (constant "PX" )
- let coq_Pc = lazy (constant"Pc")
+ let coq_PX = lazy (constant "PX")
+ let coq_Pc = lazy (constant "Pc")
let coq_Pinj = lazy (constant "Pinj")
-
let coq_OpEq = lazy (constant "OpEq")
let coq_OpNEq = lazy (constant "OpNEq")
let coq_OpLe = lazy (constant "OpLe")
- let coq_OpLt = lazy (constant "OpLt")
+ let coq_OpLt = lazy (constant "OpLt")
let coq_OpGe = lazy (constant "OpGe")
- let coq_OpGt = lazy (constant "OpGt")
-
+ let coq_OpGt = lazy (constant "OpGt")
let coq_PsatzIn = lazy (constant "PsatzIn")
let coq_PsatzSquare = lazy (constant "PsatzSquare")
let coq_PsatzMulE = lazy (constant "PsatzMulE")
let coq_PsatzMultC = lazy (constant "PsatzMulC")
- let coq_PsatzAdd = lazy (constant "PsatzAdd")
- let coq_PsatzC = lazy (constant "PsatzC")
- let coq_PsatzZ = lazy (constant "PsatzZ")
+ let coq_PsatzAdd = lazy (constant "PsatzAdd")
+ let coq_PsatzC = lazy (constant "PsatzC")
+ let coq_PsatzZ = lazy (constant "PsatzZ")
(* let coq_GT = lazy (m_constant "GT")*)
let coq_DeclaredConstant = lazy (m_constant "DeclaredConstant")
- let coq_TT = lazy
- (gen_constant_in_modules "ZMicromega"
- [["Coq" ; "micromega" ; "Tauto"];["Tauto"]] "TT")
- let coq_FF = lazy
- (gen_constant_in_modules "ZMicromega"
- [["Coq" ; "micromega" ; "Tauto"];["Tauto"]] "FF")
- let coq_And = lazy
- (gen_constant_in_modules "ZMicromega"
- [["Coq" ; "micromega" ; "Tauto"];["Tauto"]] "Cj")
- let coq_Or = lazy
- (gen_constant_in_modules "ZMicromega"
- [["Coq" ; "micromega" ; "Tauto"];["Tauto"]] "D")
- let coq_Neg = lazy
- (gen_constant_in_modules "ZMicromega"
- [["Coq" ; "micromega" ; "Tauto"];["Tauto"]] "N")
- let coq_Atom = lazy
- (gen_constant_in_modules "ZMicromega"
- [["Coq" ; "micromega" ; "Tauto"];["Tauto"]] "A")
- let coq_X = lazy
- (gen_constant_in_modules "ZMicromega"
- [["Coq" ; "micromega" ; "Tauto"];["Tauto"]] "X")
- let coq_Impl = lazy
- (gen_constant_in_modules "ZMicromega"
- [["Coq" ; "micromega" ; "Tauto"];["Tauto"]] "I")
- let coq_Formula = lazy
- (gen_constant_in_modules "ZMicromega"
- [["Coq" ; "micromega" ; "Tauto"];["Tauto"]] "BFormula")
+ let coq_TT =
+ lazy
+ (gen_constant_in_modules "ZMicromega"
+ [["Coq"; "micromega"; "Tauto"]; ["Tauto"]]
+ "TT")
+
+ let coq_FF =
+ lazy
+ (gen_constant_in_modules "ZMicromega"
+ [["Coq"; "micromega"; "Tauto"]; ["Tauto"]]
+ "FF")
+
+ let coq_And =
+ lazy
+ (gen_constant_in_modules "ZMicromega"
+ [["Coq"; "micromega"; "Tauto"]; ["Tauto"]]
+ "Cj")
+
+ let coq_Or =
+ lazy
+ (gen_constant_in_modules "ZMicromega"
+ [["Coq"; "micromega"; "Tauto"]; ["Tauto"]]
+ "D")
+
+ let coq_Neg =
+ lazy
+ (gen_constant_in_modules "ZMicromega"
+ [["Coq"; "micromega"; "Tauto"]; ["Tauto"]]
+ "N")
+
+ let coq_Atom =
+ lazy
+ (gen_constant_in_modules "ZMicromega"
+ [["Coq"; "micromega"; "Tauto"]; ["Tauto"]]
+ "A")
+
+ let coq_X =
+ lazy
+ (gen_constant_in_modules "ZMicromega"
+ [["Coq"; "micromega"; "Tauto"]; ["Tauto"]]
+ "X")
+
+ let coq_Impl =
+ lazy
+ (gen_constant_in_modules "ZMicromega"
+ [["Coq"; "micromega"; "Tauto"]; ["Tauto"]]
+ "I")
+
+ let coq_Formula =
+ lazy
+ (gen_constant_in_modules "ZMicromega"
+ [["Coq"; "micromega"; "Tauto"]; ["Tauto"]]
+ "BFormula")
(**
* Initialization : a few Caml symbols are derived from other libraries;
* QMicromega, ZArithRing, RingMicromega.
*)
- let coq_QWitness = lazy
- (gen_constant_in_modules "QMicromega"
- [["Coq"; "micromega"; "QMicromega"]] "QWitness")
+ let coq_QWitness =
+ lazy
+ (gen_constant_in_modules "QMicromega"
+ [["Coq"; "micromega"; "QMicromega"]]
+ "QWitness")
- let coq_Build = lazy
- (gen_constant_in_modules "RingMicromega"
- [["Coq" ; "micromega" ; "RingMicromega"] ; ["RingMicromega"] ]
- "Build_Formula")
- let coq_Cstr = lazy
- (gen_constant_in_modules "RingMicromega"
- [["Coq" ; "micromega" ; "RingMicromega"] ; ["RingMicromega"] ] "Formula")
+ let coq_Build =
+ lazy
+ (gen_constant_in_modules "RingMicromega"
+ [["Coq"; "micromega"; "RingMicromega"]; ["RingMicromega"]]
+ "Build_Formula")
+
+ let coq_Cstr =
+ lazy
+ (gen_constant_in_modules "RingMicromega"
+ [["Coq"; "micromega"; "RingMicromega"]; ["RingMicromega"]]
+ "Formula")
(**
* Parsing and dumping : transformation functions between Caml and Coq
@@ -445,35 +435,34 @@ struct
(* A simple but useful getter function *)
let get_left_construct sigma term =
- match EConstr.kind sigma term with
- | Construct((_,i),_) -> (i,[| |])
- | App(l,rst) ->
- (match EConstr.kind sigma l with
- | Construct((_,i),_) -> (i,rst)
- | _ -> raise ParseError
- )
- | _ -> raise ParseError
+ match EConstr.kind sigma term with
+ | Construct ((_, i), _) -> (i, [||])
+ | App (l, rst) -> (
+ match EConstr.kind sigma l with
+ | Construct ((_, i), _) -> (i, rst)
+ | _ -> raise ParseError )
+ | _ -> raise ParseError
(* Access the Micromega module *)
(* parse/dump/print from numbers up to expressions and formulas *)
let rec parse_nat sigma term =
- let (i,c) = get_left_construct sigma term in
+ let i, c = get_left_construct sigma term in
match i with
- | 1 -> Mc.O
- | 2 -> Mc.S (parse_nat sigma (c.(0)))
- | i -> raise ParseError
+ | 1 -> Mc.O
+ | 2 -> Mc.S (parse_nat sigma c.(0))
+ | i -> raise ParseError
let pp_nat o n = Printf.fprintf o "%i" (CoqToCaml.nat n)
let rec dump_nat x =
- match x with
+ match x with
| Mc.O -> Lazy.force coq_O
- | Mc.S p -> EConstr.mkApp(Lazy.force coq_S,[| dump_nat p |])
+ | Mc.S p -> EConstr.mkApp (Lazy.force coq_S, [|dump_nat p|])
let rec parse_positive sigma term =
- let (i,c) = get_left_construct sigma term in
+ let i, c = get_left_construct sigma term in
match i with
| 1 -> Mc.XI (parse_positive sigma c.(0))
| 2 -> Mc.XO (parse_positive sigma c.(0))
@@ -483,15 +472,15 @@ struct
let rec dump_positive x =
match x with
| Mc.XH -> Lazy.force coq_xH
- | Mc.XO p -> EConstr.mkApp(Lazy.force coq_xO,[| dump_positive p |])
- | Mc.XI p -> EConstr.mkApp(Lazy.force coq_xI,[| dump_positive p |])
+ | Mc.XO p -> EConstr.mkApp (Lazy.force coq_xO, [|dump_positive p|])
+ | Mc.XI p -> EConstr.mkApp (Lazy.force coq_xI, [|dump_positive p|])
let pp_positive o x = Printf.fprintf o "%i" (CoqToCaml.positive x)
let dump_n x =
match x with
| Mc.N0 -> Lazy.force coq_N0
- | Mc.Npos p -> EConstr.mkApp(Lazy.force coq_Npos,[| dump_positive p|])
+ | Mc.Npos p -> EConstr.mkApp (Lazy.force coq_Npos, [|dump_positive p|])
(** [is_ground_term env sigma term] holds if the term [term]
is an instance of the typeclass [DeclConstant.GT term]
@@ -502,26 +491,26 @@ struct
let is_declared_term env evd t =
match EConstr.kind evd t with
- | Const _ | Construct _ -> (* Restrict typeclass resolution to trivial cases *)
- begin
- let typ = Retyping.get_type_of env evd t in
- try
- ignore (Typeclasses.resolve_one_typeclass env evd (EConstr.mkApp(Lazy.force coq_DeclaredConstant,[| typ;t|]))) ; true
- with Not_found -> false
- end
- | _ -> false
+ | Const _ | Construct _ -> (
+ (* Restrict typeclass resolution to trivial cases *)
+ let typ = Retyping.get_type_of env evd t in
+ try
+ ignore
+ (Typeclasses.resolve_one_typeclass env evd
+ (EConstr.mkApp (Lazy.force coq_DeclaredConstant, [|typ; t|])));
+ true
+ with Not_found -> false )
+ | _ -> false
let rec is_ground_term env evd term =
match EConstr.kind evd term with
- | App(c,args) ->
- is_declared_term env evd c &&
- Array.for_all (is_ground_term env evd) args
+ | App (c, args) ->
+ is_declared_term env evd c && Array.for_all (is_ground_term env evd) args
| Const _ | Construct _ -> is_declared_term env evd term
- | _ -> false
-
+ | _ -> false
let parse_z sigma term =
- let (i,c) = get_left_construct sigma term in
+ let i, c = get_left_construct sigma term in
match i with
| 1 -> Mc.Z0
| 2 -> Mc.Zpos (parse_positive sigma c.(0))
@@ -529,221 +518,246 @@ struct
| i -> raise ParseError
let dump_z x =
- match x with
- | Mc.Z0 ->Lazy.force coq_ZERO
- | Mc.Zpos p -> EConstr.mkApp(Lazy.force coq_POS,[| dump_positive p|])
- | Mc.Zneg p -> EConstr.mkApp(Lazy.force coq_NEG,[| dump_positive p|])
+ match x with
+ | Mc.Z0 -> Lazy.force coq_ZERO
+ | Mc.Zpos p -> EConstr.mkApp (Lazy.force coq_POS, [|dump_positive p|])
+ | Mc.Zneg p -> EConstr.mkApp (Lazy.force coq_NEG, [|dump_positive p|])
- let pp_z o x = Printf.fprintf o "%s" (Big_int.string_of_big_int (CoqToCaml.z_big_int x))
+ let pp_z o x =
+ Printf.fprintf o "%s" (Big_int.string_of_big_int (CoqToCaml.z_big_int x))
let dump_q q =
- EConstr.mkApp(Lazy.force coq_Qmake,
- [| dump_z q.Micromega.qnum ; dump_positive q.Micromega.qden|])
+ EConstr.mkApp
+ ( Lazy.force coq_Qmake
+ , [|dump_z q.Micromega.qnum; dump_positive q.Micromega.qden|] )
let parse_q sigma term =
- match EConstr.kind sigma term with
- | App(c, args) -> if EConstr.eq_constr sigma c (Lazy.force coq_Qmake) then
- {Mc.qnum = parse_z sigma args.(0) ; Mc.qden = parse_positive sigma args.(1) }
- else raise ParseError
- | _ -> raise ParseError
-
+ match EConstr.kind sigma term with
+ | App (c, args) ->
+ if EConstr.eq_constr sigma c (Lazy.force coq_Qmake) then
+ { Mc.qnum = parse_z sigma args.(0)
+ ; Mc.qden = parse_positive sigma args.(1) }
+ else raise ParseError
+ | _ -> raise ParseError
let rec pp_Rcst o cst =
match cst with
- | Mc.C0 -> output_string o "C0"
- | Mc.C1 -> output_string o "C1"
- | Mc.CQ q -> output_string o "CQ _"
- | Mc.CZ z -> pp_z o z
- | Mc.CPlus(x,y) -> Printf.fprintf o "(%a + %a)" pp_Rcst x pp_Rcst y
- | Mc.CMinus(x,y) -> Printf.fprintf o "(%a - %a)" pp_Rcst x pp_Rcst y
- | Mc.CMult(x,y) -> Printf.fprintf o "(%a * %a)" pp_Rcst x pp_Rcst y
- | Mc.CPow(x,y) -> Printf.fprintf o "(%a ^ _)" pp_Rcst x
- | Mc.CInv t -> Printf.fprintf o "(/ %a)" pp_Rcst t
- | Mc.COpp t -> Printf.fprintf o "(- %a)" pp_Rcst t
-
+ | Mc.C0 -> output_string o "C0"
+ | Mc.C1 -> output_string o "C1"
+ | Mc.CQ q -> output_string o "CQ _"
+ | Mc.CZ z -> pp_z o z
+ | Mc.CPlus (x, y) -> Printf.fprintf o "(%a + %a)" pp_Rcst x pp_Rcst y
+ | Mc.CMinus (x, y) -> Printf.fprintf o "(%a - %a)" pp_Rcst x pp_Rcst y
+ | Mc.CMult (x, y) -> Printf.fprintf o "(%a * %a)" pp_Rcst x pp_Rcst y
+ | Mc.CPow (x, y) -> Printf.fprintf o "(%a ^ _)" pp_Rcst x
+ | Mc.CInv t -> Printf.fprintf o "(/ %a)" pp_Rcst t
+ | Mc.COpp t -> Printf.fprintf o "(- %a)" pp_Rcst t
let rec dump_Rcst cst =
match cst with
- | Mc.C0 -> Lazy.force coq_C0
- | Mc.C1 -> Lazy.force coq_C1
- | Mc.CQ q -> EConstr.mkApp(Lazy.force coq_CQ, [| dump_q q |])
- | Mc.CZ z -> EConstr.mkApp(Lazy.force coq_CZ, [| dump_z z |])
- | Mc.CPlus(x,y) -> EConstr.mkApp(Lazy.force coq_CPlus, [| dump_Rcst x ; dump_Rcst y |])
- | Mc.CMinus(x,y) -> EConstr.mkApp(Lazy.force coq_CMinus, [| dump_Rcst x ; dump_Rcst y |])
- | Mc.CMult(x,y) -> EConstr.mkApp(Lazy.force coq_CMult, [| dump_Rcst x ; dump_Rcst y |])
- | Mc.CPow(x,y) -> EConstr.mkApp(Lazy.force coq_CPow, [| dump_Rcst x ;
- match y with
- | Mc.Inl z -> EConstr.mkApp(Lazy.force coq_Inl,[| Lazy.force coq_Z ; Lazy.force coq_nat; dump_z z|])
- | Mc.Inr n -> EConstr.mkApp(Lazy.force coq_Inr,[| Lazy.force coq_Z ; Lazy.force coq_nat; dump_nat n|])
- |])
- | Mc.CInv t -> EConstr.mkApp(Lazy.force coq_CInv, [| dump_Rcst t |])
- | Mc.COpp t -> EConstr.mkApp(Lazy.force coq_COpp, [| dump_Rcst t |])
+ | Mc.C0 -> Lazy.force coq_C0
+ | Mc.C1 -> Lazy.force coq_C1
+ | Mc.CQ q -> EConstr.mkApp (Lazy.force coq_CQ, [|dump_q q|])
+ | Mc.CZ z -> EConstr.mkApp (Lazy.force coq_CZ, [|dump_z z|])
+ | Mc.CPlus (x, y) ->
+ EConstr.mkApp (Lazy.force coq_CPlus, [|dump_Rcst x; dump_Rcst y|])
+ | Mc.CMinus (x, y) ->
+ EConstr.mkApp (Lazy.force coq_CMinus, [|dump_Rcst x; dump_Rcst y|])
+ | Mc.CMult (x, y) ->
+ EConstr.mkApp (Lazy.force coq_CMult, [|dump_Rcst x; dump_Rcst y|])
+ | Mc.CPow (x, y) ->
+ EConstr.mkApp
+ ( Lazy.force coq_CPow
+ , [| dump_Rcst x
+ ; ( match y with
+ | Mc.Inl z ->
+ EConstr.mkApp
+ ( Lazy.force coq_Inl
+ , [|Lazy.force coq_Z; Lazy.force coq_nat; dump_z z|] )
+ | Mc.Inr n ->
+ EConstr.mkApp
+ ( Lazy.force coq_Inr
+ , [|Lazy.force coq_Z; Lazy.force coq_nat; dump_nat n|] ) ) |]
+ )
+ | Mc.CInv t -> EConstr.mkApp (Lazy.force coq_CInv, [|dump_Rcst t|])
+ | Mc.COpp t -> EConstr.mkApp (Lazy.force coq_COpp, [|dump_Rcst t|])
let rec dump_list typ dump_elt l =
- match l with
- | [] -> EConstr.mkApp(Lazy.force coq_nil,[| typ |])
- | e :: l -> EConstr.mkApp(Lazy.force coq_cons,
- [| typ; dump_elt e;dump_list typ dump_elt l|])
+ match l with
+ | [] -> EConstr.mkApp (Lazy.force coq_nil, [|typ|])
+ | e :: l ->
+ EConstr.mkApp
+ (Lazy.force coq_cons, [|typ; dump_elt e; dump_list typ dump_elt l|])
let pp_list op cl elt o l =
- let rec _pp o l =
- match l with
- | [] -> ()
- | [e] -> Printf.fprintf o "%a" elt e
- | e::l -> Printf.fprintf o "%a ,%a" elt e _pp l in
+ let rec _pp o l =
+ match l with
+ | [] -> ()
+ | [e] -> Printf.fprintf o "%a" elt e
+ | e :: l -> Printf.fprintf o "%a ,%a" elt e _pp l
+ in
Printf.fprintf o "%s%a%s" op _pp l cl
let dump_var = dump_positive
let dump_expr typ dump_z e =
- let rec dump_expr e =
- match e with
- | Mc.PEX n -> EConstr.mkApp(Lazy.force coq_PEX,[| typ; dump_var n |])
- | Mc.PEc z -> EConstr.mkApp(Lazy.force coq_PEc,[| typ ; dump_z z |])
- | Mc.PEadd(e1,e2) -> EConstr.mkApp(Lazy.force coq_PEadd,
- [| typ; dump_expr e1;dump_expr e2|])
- | Mc.PEsub(e1,e2) -> EConstr.mkApp(Lazy.force coq_PEsub,
- [| typ; dump_expr e1;dump_expr e2|])
- | Mc.PEopp e -> EConstr.mkApp(Lazy.force coq_PEopp,
- [| typ; dump_expr e|])
- | Mc.PEmul(e1,e2) -> EConstr.mkApp(Lazy.force coq_PEmul,
- [| typ; dump_expr e1;dump_expr e2|])
- | Mc.PEpow(e,n) -> EConstr.mkApp(Lazy.force coq_PEpow,
- [| typ; dump_expr e; dump_n n|])
- in
+ let rec dump_expr e =
+ match e with
+ | Mc.PEX n -> EConstr.mkApp (Lazy.force coq_PEX, [|typ; dump_var n|])
+ | Mc.PEc z -> EConstr.mkApp (Lazy.force coq_PEc, [|typ; dump_z z|])
+ | Mc.PEadd (e1, e2) ->
+ EConstr.mkApp (Lazy.force coq_PEadd, [|typ; dump_expr e1; dump_expr e2|])
+ | Mc.PEsub (e1, e2) ->
+ EConstr.mkApp (Lazy.force coq_PEsub, [|typ; dump_expr e1; dump_expr e2|])
+ | Mc.PEopp e -> EConstr.mkApp (Lazy.force coq_PEopp, [|typ; dump_expr e|])
+ | Mc.PEmul (e1, e2) ->
+ EConstr.mkApp (Lazy.force coq_PEmul, [|typ; dump_expr e1; dump_expr e2|])
+ | Mc.PEpow (e, n) ->
+ EConstr.mkApp (Lazy.force coq_PEpow, [|typ; dump_expr e; dump_n n|])
+ in
dump_expr e
let dump_pol typ dump_c e =
let rec dump_pol e =
match e with
- | Mc.Pc n -> EConstr.mkApp(Lazy.force coq_Pc, [|typ ; dump_c n|])
- | Mc.Pinj(p,pol) -> EConstr.mkApp(Lazy.force coq_Pinj , [| typ ; dump_positive p ; dump_pol pol|])
- | Mc.PX(pol1,p,pol2) -> EConstr.mkApp(Lazy.force coq_PX, [| typ ; dump_pol pol1 ; dump_positive p ; dump_pol pol2|]) in
- dump_pol e
+ | Mc.Pc n -> EConstr.mkApp (Lazy.force coq_Pc, [|typ; dump_c n|])
+ | Mc.Pinj (p, pol) ->
+ EConstr.mkApp
+ (Lazy.force coq_Pinj, [|typ; dump_positive p; dump_pol pol|])
+ | Mc.PX (pol1, p, pol2) ->
+ EConstr.mkApp
+ ( Lazy.force coq_PX
+ , [|typ; dump_pol pol1; dump_positive p; dump_pol pol2|] )
+ in
+ dump_pol e
let pp_pol pp_c o e =
let rec pp_pol o e =
match e with
- | Mc.Pc n -> Printf.fprintf o "Pc %a" pp_c n
- | Mc.Pinj(p,pol) -> Printf.fprintf o "Pinj(%a,%a)" pp_positive p pp_pol pol
- | Mc.PX(pol1,p,pol2) -> Printf.fprintf o "PX(%a,%a,%a)" pp_pol pol1 pp_positive p pp_pol pol2 in
- pp_pol o e
+ | Mc.Pc n -> Printf.fprintf o "Pc %a" pp_c n
+ | Mc.Pinj (p, pol) ->
+ Printf.fprintf o "Pinj(%a,%a)" pp_positive p pp_pol pol
+ | Mc.PX (pol1, p, pol2) ->
+ Printf.fprintf o "PX(%a,%a,%a)" pp_pol pol1 pp_positive p pp_pol pol2
+ in
+ pp_pol o e
-(* let pp_clause pp_c o (f: 'cst clause) =
+ (* let pp_clause pp_c o (f: 'cst clause) =
List.iter (fun ((p,_),(t,_)) -> Printf.fprintf o "(%a @%a)" (pp_pol pp_c) p Tag.pp t) f *)
- let pp_clause_tag o (f: 'cst clause) =
- List.iter (fun ((p,_),(t,_)) -> Printf.fprintf o "(_ @%a)" Tag.pp t) f
+ let pp_clause_tag o (f : 'cst clause) =
+ List.iter (fun ((p, _), (t, _)) -> Printf.fprintf o "(_ @%a)" Tag.pp t) f
-(* let pp_cnf pp_c o (f:'cst cnf) =
+ (* let pp_cnf pp_c o (f:'cst cnf) =
List.iter (fun l -> Printf.fprintf o "[%a]" (pp_clause pp_c) l) f *)
- let pp_cnf_tag o (f:'cst cnf) =
+ let pp_cnf_tag o (f : 'cst cnf) =
List.iter (fun l -> Printf.fprintf o "[%a]" pp_clause_tag l) f
-
let dump_psatz typ dump_z e =
- let z = Lazy.force typ in
- let rec dump_cone e =
- match e with
- | Mc.PsatzIn n -> EConstr.mkApp(Lazy.force coq_PsatzIn,[| z; dump_nat n |])
- | Mc.PsatzMulC(e,c) -> EConstr.mkApp(Lazy.force coq_PsatzMultC,
- [| z; dump_pol z dump_z e ; dump_cone c |])
- | Mc.PsatzSquare e -> EConstr.mkApp(Lazy.force coq_PsatzSquare,
- [| z;dump_pol z dump_z e|])
- | Mc.PsatzAdd(e1,e2) -> EConstr.mkApp(Lazy.force coq_PsatzAdd,
- [| z; dump_cone e1; dump_cone e2|])
- | Mc.PsatzMulE(e1,e2) -> EConstr.mkApp(Lazy.force coq_PsatzMulE,
- [| z; dump_cone e1; dump_cone e2|])
- | Mc.PsatzC p -> EConstr.mkApp(Lazy.force coq_PsatzC,[| z; dump_z p|])
- | Mc.PsatzZ -> EConstr.mkApp(Lazy.force coq_PsatzZ,[| z|]) in
- dump_cone e
-
- let pp_psatz pp_z o e =
- let rec pp_cone o e =
- match e with
- | Mc.PsatzIn n ->
- Printf.fprintf o "(In %a)%%nat" pp_nat n
- | Mc.PsatzMulC(e,c) ->
+ let z = Lazy.force typ in
+ let rec dump_cone e =
+ match e with
+ | Mc.PsatzIn n -> EConstr.mkApp (Lazy.force coq_PsatzIn, [|z; dump_nat n|])
+ | Mc.PsatzMulC (e, c) ->
+ EConstr.mkApp
+ (Lazy.force coq_PsatzMultC, [|z; dump_pol z dump_z e; dump_cone c|])
+ | Mc.PsatzSquare e ->
+ EConstr.mkApp (Lazy.force coq_PsatzSquare, [|z; dump_pol z dump_z e|])
+ | Mc.PsatzAdd (e1, e2) ->
+ EConstr.mkApp
+ (Lazy.force coq_PsatzAdd, [|z; dump_cone e1; dump_cone e2|])
+ | Mc.PsatzMulE (e1, e2) ->
+ EConstr.mkApp
+ (Lazy.force coq_PsatzMulE, [|z; dump_cone e1; dump_cone e2|])
+ | Mc.PsatzC p -> EConstr.mkApp (Lazy.force coq_PsatzC, [|z; dump_z p|])
+ | Mc.PsatzZ -> EConstr.mkApp (Lazy.force coq_PsatzZ, [|z|])
+ in
+ dump_cone e
+
+ let pp_psatz pp_z o e =
+ let rec pp_cone o e =
+ match e with
+ | Mc.PsatzIn n -> Printf.fprintf o "(In %a)%%nat" pp_nat n
+ | Mc.PsatzMulC (e, c) ->
Printf.fprintf o "( %a [*] %a)" (pp_pol pp_z) e pp_cone c
- | Mc.PsatzSquare e ->
- Printf.fprintf o "(%a^2)" (pp_pol pp_z) e
- | Mc.PsatzAdd(e1,e2) ->
+ | Mc.PsatzSquare e -> Printf.fprintf o "(%a^2)" (pp_pol pp_z) e
+ | Mc.PsatzAdd (e1, e2) ->
Printf.fprintf o "(%a [+] %a)" pp_cone e1 pp_cone e2
- | Mc.PsatzMulE(e1,e2) ->
+ | Mc.PsatzMulE (e1, e2) ->
Printf.fprintf o "(%a [*] %a)" pp_cone e1 pp_cone e2
- | Mc.PsatzC p ->
- Printf.fprintf o "(%a)%%positive" pp_z p
- | Mc.PsatzZ ->
- Printf.fprintf o "0" in
+ | Mc.PsatzC p -> Printf.fprintf o "(%a)%%positive" pp_z p
+ | Mc.PsatzZ -> Printf.fprintf o "0"
+ in
pp_cone o e
let dump_op = function
- | Mc.OpEq-> Lazy.force coq_OpEq
- | Mc.OpNEq-> Lazy.force coq_OpNEq
- | Mc.OpLe -> Lazy.force coq_OpLe
- | Mc.OpGe -> Lazy.force coq_OpGe
- | Mc.OpGt-> Lazy.force coq_OpGt
- | Mc.OpLt-> Lazy.force coq_OpLt
-
- let dump_cstr typ dump_constant {Mc.flhs = e1 ; Mc.fop = o ; Mc.frhs = e2} =
- EConstr.mkApp(Lazy.force coq_Build,
- [| typ; dump_expr typ dump_constant e1 ;
- dump_op o ;
- dump_expr typ dump_constant e2|])
+ | Mc.OpEq -> Lazy.force coq_OpEq
+ | Mc.OpNEq -> Lazy.force coq_OpNEq
+ | Mc.OpLe -> Lazy.force coq_OpLe
+ | Mc.OpGe -> Lazy.force coq_OpGe
+ | Mc.OpGt -> Lazy.force coq_OpGt
+ | Mc.OpLt -> Lazy.force coq_OpLt
+
+ let dump_cstr typ dump_constant {Mc.flhs = e1; Mc.fop = o; Mc.frhs = e2} =
+ EConstr.mkApp
+ ( Lazy.force coq_Build
+ , [| typ
+ ; dump_expr typ dump_constant e1
+ ; dump_op o
+ ; dump_expr typ dump_constant e2 |] )
let assoc_const sigma x l =
- try
- snd (List.find (fun (x',y) -> EConstr.eq_constr sigma x (Lazy.force x')) l)
- with
- Not_found -> raise ParseError
-
- let zop_table = [
- coq_Zgt, Mc.OpGt ;
- coq_Zge, Mc.OpGe ;
- coq_Zlt, Mc.OpLt ;
- coq_Zle, Mc.OpLe ]
-
- let rop_table = [
- coq_Rgt, Mc.OpGt ;
- coq_Rge, Mc.OpGe ;
- coq_Rlt, Mc.OpLt ;
- coq_Rle, Mc.OpLe ]
-
- let qop_table = [
- coq_Qlt, Mc.OpLt ;
- coq_Qle, Mc.OpLe ;
- coq_Qeq, Mc.OpEq
- ]
-
- type gl = { env : Environ.env; sigma : Evd.evar_map }
-
- let is_convertible gl t1 t2 =
- Reductionops.is_conv gl.env gl.sigma t1 t2
-
- let parse_zop gl (op,args) =
+ try
+ snd
+ (List.find (fun (x', y) -> EConstr.eq_constr sigma x (Lazy.force x')) l)
+ with Not_found -> raise ParseError
+
+ let zop_table =
+ [ (coq_Zgt, Mc.OpGt)
+ ; (coq_Zge, Mc.OpGe)
+ ; (coq_Zlt, Mc.OpLt)
+ ; (coq_Zle, Mc.OpLe) ]
+
+ let rop_table =
+ [ (coq_Rgt, Mc.OpGt)
+ ; (coq_Rge, Mc.OpGe)
+ ; (coq_Rlt, Mc.OpLt)
+ ; (coq_Rle, Mc.OpLe) ]
+
+ let qop_table = [(coq_Qlt, Mc.OpLt); (coq_Qle, Mc.OpLe); (coq_Qeq, Mc.OpEq)]
+
+ type gl = {env : Environ.env; sigma : Evd.evar_map}
+
+ let is_convertible gl t1 t2 = Reductionops.is_conv gl.env gl.sigma t1 t2
+
+ let parse_zop gl (op, args) =
let sigma = gl.sigma in
match args with
- | [| a1 ; a2|] -> assoc_const sigma op zop_table, a1, a2
- | [| ty ; a1 ; a2|] ->
- if EConstr.eq_constr sigma op (Lazy.force coq_Eq) && is_convertible gl ty (Lazy.force coq_Z)
- then (Mc.OpEq, args.(1), args.(2))
- else raise ParseError
- | _ -> raise ParseError
-
- let parse_rop gl (op,args) =
+ | [|a1; a2|] -> (assoc_const sigma op zop_table, a1, a2)
+ | [|ty; a1; a2|] ->
+ if
+ EConstr.eq_constr sigma op (Lazy.force coq_Eq)
+ && is_convertible gl ty (Lazy.force coq_Z)
+ then (Mc.OpEq, args.(1), args.(2))
+ else raise ParseError
+ | _ -> raise ParseError
+
+ let parse_rop gl (op, args) =
let sigma = gl.sigma in
match args with
- | [| a1 ; a2|] -> assoc_const sigma op rop_table, a1 , a2
- | [| ty ; a1 ; a2|] ->
- if EConstr.eq_constr sigma op (Lazy.force coq_Eq) && is_convertible gl ty (Lazy.force coq_R)
- then (Mc.OpEq, a1, a2)
- else raise ParseError
- | _ -> raise ParseError
-
- let parse_qop gl (op,args) =
- if Array.length args = 2
- then (assoc_const gl.sigma op qop_table, args.(0) , args.(1))
+ | [|a1; a2|] -> (assoc_const sigma op rop_table, a1, a2)
+ | [|ty; a1; a2|] ->
+ if
+ EConstr.eq_constr sigma op (Lazy.force coq_Eq)
+ && is_convertible gl ty (Lazy.force coq_R)
+ then (Mc.OpEq, a1, a2)
+ else raise ParseError
+ | _ -> raise ParseError
+
+ let parse_qop gl (op, args) =
+ if Array.length args = 2 then
+ (assoc_const gl.sigma op qop_table, args.(0), args.(1))
else raise ParseError
type 'a op =
@@ -753,74 +767,65 @@ struct
| Ukn of string
let assoc_ops sigma x l =
- try
- snd (List.find (fun (x',y) -> EConstr.eq_constr sigma x (Lazy.force x')) l)
- with
- Not_found -> Ukn "Oups"
+ try
+ snd
+ (List.find (fun (x', y) -> EConstr.eq_constr sigma x (Lazy.force x')) l)
+ with Not_found -> Ukn "Oups"
(**
* MODULE: Env is for environment.
*)
- module Env =
- struct
-
- type t = {
- vars : EConstr.t list ;
- (* The list represents a mapping from EConstr.t to indexes. *)
- gl : gl;
- (* The evar_map may be updated due to unification of universes *)
- }
-
- let empty gl =
- {
- vars = [];
- gl = gl
- }
+ module Env = struct
+ type t =
+ { vars : EConstr.t list
+ ; (* The list represents a mapping from EConstr.t to indexes. *)
+ gl : gl
+ (* The evar_map may be updated due to unification of universes *) }
+ let empty gl = {vars = []; gl}
(** [eq_constr gl x y] returns an updated [gl] if x and y can be unified *)
let eq_constr gl x y =
let evd = gl.sigma in
match EConstr.eq_constr_universes_proj gl.env evd x y with
- | Some csts ->
- let csts = UnivProblem.to_constraints ~force_weak:false (Evd.universes evd) csts in
- begin
- match Evd.add_constraints evd csts with
- | evd -> Some {gl with sigma = evd}
- | exception Univ.UniverseInconsistency _ -> None
- end
+ | Some csts -> (
+ let csts =
+ UnivProblem.to_constraints ~force_weak:false (Evd.universes evd) csts
+ in
+ match Evd.add_constraints evd csts with
+ | evd -> Some {gl with sigma = evd}
+ | exception Univ.UniverseInconsistency _ -> None )
| None -> None
let compute_rank_add env v =
let rec _add gl vars n v =
match vars with
- | [] -> (gl, [v] ,n)
- | e::l ->
- match eq_constr gl e v with
- | Some gl' -> (gl', vars , n)
- | None ->
- let (gl,l',n) = _add gl l ( n+1) v in
- (gl,e::l',n) in
- let (gl',vars', n) = _add env.gl env.vars 1 v in
- ({vars=vars';gl=gl'}, CamlToCoq.positive n)
-
- let get_rank env v =
- let gl = env.gl in
-
- let rec _get_rank env n =
- match env with
- | [] -> raise (Invalid_argument "get_rank")
- | e::l ->
- match eq_constr gl e v with
- | Some _ -> n
- | None -> _get_rank l (n+1)
- in
- _get_rank env.vars 1
+ | [] -> (gl, [v], n)
+ | e :: l -> (
+ match eq_constr gl e v with
+ | Some gl' -> (gl', vars, n)
+ | None ->
+ let gl, l', n = _add gl l (n + 1) v in
+ (gl, e :: l', n) )
+ in
+ let gl', vars', n = _add env.gl env.vars 1 v in
+ ({vars = vars'; gl = gl'}, CamlToCoq.positive n)
+
+ let get_rank env v =
+ let gl = env.gl in
+ let rec _get_rank env n =
+ match env with
+ | [] -> raise (Invalid_argument "get_rank")
+ | e :: l -> (
+ match eq_constr gl e v with Some _ -> n | None -> _get_rank l (n + 1)
+ )
+ in
+ _get_rank env.vars 1
- let elements env = env.vars
+ let elements env = env.vars
-(* let string_of_env gl env =
+ (* let string_of_env gl env =
let rec string_of_env i env acc =
match env with
| [] -> acc
@@ -830,101 +835,103 @@ struct
(Printer.pr_econstr_env gl.env gl.sigma e)) acc) in
string_of_env 1 env IMap.empty
*)
- let pp gl env =
- let ppl = List.mapi (fun i e -> Pp.str "x" ++ Pp.int (i+1) ++ Pp.str ":" ++ Printer.pr_econstr_env gl.env gl.sigma e)env in
- List.fold_right (fun e p -> e ++ Pp.str " ; " ++ p ) ppl (Pp.str "\n")
+ let pp gl env =
+ let ppl =
+ List.mapi
+ (fun i e ->
+ Pp.str "x"
+ ++ Pp.int (i + 1)
+ ++ Pp.str ":"
+ ++ Printer.pr_econstr_env gl.env gl.sigma e)
+ env
+ in
+ List.fold_right (fun e p -> e ++ Pp.str " ; " ++ p) ppl (Pp.str "\n")
+ end
- end (* MODULE END: Env *)
+ (* MODULE END: Env *)
(**
* This is the big generic function for expression parsers.
*)
let parse_expr gl parse_constant parse_exp ops_spec env term =
- if debug
- then (
- Feedback.msg_debug (Pp.str "parse_expr: " ++ Printer.pr_leconstr_env gl.env gl.sigma term));
-
+ if debug then
+ Feedback.msg_debug
+ (Pp.str "parse_expr: " ++ Printer.pr_leconstr_env gl.env gl.sigma term);
let parse_variable env term =
- let (env,n) = Env.compute_rank_add env term in
- (Mc.PEX n , env) in
-
+ let env, n = Env.compute_rank_add env term in
+ (Mc.PEX n, env)
+ in
let rec parse_expr env term =
- let combine env op (t1,t2) =
- let (expr1,env) = parse_expr env t1 in
- let (expr2,env) = parse_expr env t2 in
- (op expr1 expr2,env) in
-
- try (Mc.PEc (parse_constant gl term) , env)
- with ParseError ->
- match EConstr.kind gl.sigma term with
- | App(t,args) ->
- (
- match EConstr.kind gl.sigma t with
- | Const c ->
- ( match assoc_ops gl.sigma t ops_spec with
- | Binop f -> combine env f (args.(0),args.(1))
- | Opp -> let (expr,env) = parse_expr env args.(0) in
- (Mc.PEopp expr, env)
- | Power ->
- begin
- try
- let (expr,env) = parse_expr env args.(0) in
- let power = (parse_exp expr args.(1)) in
- (power , env)
- with ParseError ->
- (* if the exponent is a variable *)
- let (env,n) = Env.compute_rank_add env term in (Mc.PEX n, env)
- end
- | Ukn s ->
- if debug
- then (Printf.printf "unknown op: %s\n" s; flush stdout;);
- let (env,n) = Env.compute_rank_add env term in (Mc.PEX n, env)
- )
- | _ -> parse_variable env term
- )
- | _ -> parse_variable env term in
- parse_expr env term
+ let combine env op (t1, t2) =
+ let expr1, env = parse_expr env t1 in
+ let expr2, env = parse_expr env t2 in
+ (op expr1 expr2, env)
+ in
+ try (Mc.PEc (parse_constant gl term), env)
+ with ParseError -> (
+ match EConstr.kind gl.sigma term with
+ | App (t, args) -> (
+ match EConstr.kind gl.sigma t with
+ | Const c -> (
+ match assoc_ops gl.sigma t ops_spec with
+ | Binop f -> combine env f (args.(0), args.(1))
+ | Opp ->
+ let expr, env = parse_expr env args.(0) in
+ (Mc.PEopp expr, env)
+ | Power -> (
+ try
+ let expr, env = parse_expr env args.(0) in
+ let power = parse_exp expr args.(1) in
+ (power, env)
+ with ParseError ->
+ (* if the exponent is a variable *)
+ let env, n = Env.compute_rank_add env term in
+ (Mc.PEX n, env) )
+ | Ukn s ->
+ if debug then (
+ Printf.printf "unknown op: %s\n" s;
+ flush stdout );
+ let env, n = Env.compute_rank_add env term in
+ (Mc.PEX n, env) )
+ | _ -> parse_variable env term )
+ | _ -> parse_variable env term )
+ in
+ parse_expr env term
let zop_spec =
- [
- coq_Zplus , Binop (fun x y -> Mc.PEadd(x,y)) ;
- coq_Zminus , Binop (fun x y -> Mc.PEsub(x,y)) ;
- coq_Zmult , Binop (fun x y -> Mc.PEmul (x,y)) ;
- coq_Zopp , Opp ;
- coq_Zpower , Power]
+ [ (coq_Zplus, Binop (fun x y -> Mc.PEadd (x, y)))
+ ; (coq_Zminus, Binop (fun x y -> Mc.PEsub (x, y)))
+ ; (coq_Zmult, Binop (fun x y -> Mc.PEmul (x, y)))
+ ; (coq_Zopp, Opp)
+ ; (coq_Zpower, Power) ]
let qop_spec =
- [
- coq_Qplus , Binop (fun x y -> Mc.PEadd(x,y)) ;
- coq_Qminus , Binop (fun x y -> Mc.PEsub(x,y)) ;
- coq_Qmult , Binop (fun x y -> Mc.PEmul (x,y)) ;
- coq_Qopp , Opp ;
- coq_Qpower , Power]
+ [ (coq_Qplus, Binop (fun x y -> Mc.PEadd (x, y)))
+ ; (coq_Qminus, Binop (fun x y -> Mc.PEsub (x, y)))
+ ; (coq_Qmult, Binop (fun x y -> Mc.PEmul (x, y)))
+ ; (coq_Qopp, Opp)
+ ; (coq_Qpower, Power) ]
let rop_spec =
- [
- coq_Rplus , Binop (fun x y -> Mc.PEadd(x,y)) ;
- coq_Rminus , Binop (fun x y -> Mc.PEsub(x,y)) ;
- coq_Rmult , Binop (fun x y -> Mc.PEmul (x,y)) ;
- coq_Ropp , Opp ;
- coq_Rpower , Power]
+ [ (coq_Rplus, Binop (fun x y -> Mc.PEadd (x, y)))
+ ; (coq_Rminus, Binop (fun x y -> Mc.PEsub (x, y)))
+ ; (coq_Rmult, Binop (fun x y -> Mc.PEmul (x, y)))
+ ; (coq_Ropp, Opp)
+ ; (coq_Rpower, Power) ]
- let parse_constant parse gl t = parse gl.sigma t
+ let parse_constant parse gl t = parse gl.sigma t
(** [parse_more_constant parse gl t] returns the reification of term [t].
If [t] is a ground term, then it is first reduced to normal form
before using a 'syntactic' parser *)
let parse_more_constant parse gl t =
- try
- parse gl t
- with ParseError ->
- begin
- if debug then Feedback.msg_debug Pp.(str "try harder");
- if is_ground_term gl.env gl.sigma t
- then parse gl (Redexpr.cbv_vm gl.env gl.sigma t)
- else raise ParseError
- end
+ try parse gl t
+ with ParseError ->
+ if debug then Feedback.msg_debug Pp.(str "try harder");
+ if is_ground_term gl.env gl.sigma t then
+ parse gl (Redexpr.cbv_vm gl.env gl.sigma t)
+ else raise ParseError
let zconstant = parse_constant parse_z
let qconstant = parse_constant parse_q
@@ -935,22 +942,17 @@ struct
[parse_constant_expr] returns a constant if the argument is an expression without variables. *)
let rec parse_zexpr gl =
- parse_expr gl
- zconstant
- (fun expr (x:EConstr.t) ->
+ parse_expr gl zconstant
+ (fun expr (x : EConstr.t) ->
let z = parse_zconstant gl x in
match z with
| Mc.Zneg _ -> Mc.PEc Mc.Z0
- | _ -> Mc.PEpow(expr, Mc.Z.to_N z)
- )
- zop_spec
- and parse_zconstant gl e =
- let (e,_) = parse_zexpr gl (Env.empty gl) e in
- match Mc.zeval_const e with
- | None -> raise ParseError
- | Some z -> z
-
+ | _ -> Mc.PEpow (expr, Mc.Z.to_N z))
+ zop_spec
+ and parse_zconstant gl e =
+ let e, _ = parse_zexpr gl (Env.empty gl) e in
+ match Mc.zeval_const e with None -> raise ParseError | Some z -> z
(* NB: R is a different story.
Because it is axiomatised, reducing would not be effective.
@@ -958,389 +960,387 @@ struct
*)
let rconst_assoc =
- [
- 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_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)) ;*) ]
let rconstant gl term =
-
let sigma = gl.sigma in
-
let rec rconstant term =
match EConstr.kind sigma term with
| Const x ->
- if EConstr.eq_constr sigma term (Lazy.force coq_R0)
- then Mc.C0
- else if EConstr.eq_constr sigma term (Lazy.force coq_R1)
- then Mc.C1
- else raise ParseError
- | App(op,args) ->
- begin
- try
- (* the evaluation order is important in the following *)
- let f = assoc_const sigma op rconst_assoc in
- let a = rconstant args.(0) in
- let b = rconstant args.(1) in
- f a b
- with
- ParseError ->
- match op with
- | op when EConstr.eq_constr sigma 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 EConstr.eq_constr sigma op (Lazy.force coq_Rpower) ->
- Mc.CPow(rconstant args.(0) , Mc.Inr (parse_more_constant nconstant gl args.(1)))
- | op when EConstr.eq_constr sigma op (Lazy.force coq_IQR) ->
- Mc.CQ (qconstant gl args.(0))
- | op when EConstr.eq_constr sigma op (Lazy.force coq_IZR) ->
- Mc.CZ (parse_more_constant zconstant gl args.(0))
- | _ -> raise ParseError
- end
- | _ -> raise ParseError in
-
+ if EConstr.eq_constr sigma term (Lazy.force coq_R0) then Mc.C0
+ else if EConstr.eq_constr sigma term (Lazy.force coq_R1) then Mc.C1
+ else raise ParseError
+ | App (op, args) -> (
+ try
+ (* the evaluation order is important in the following *)
+ let f = assoc_const sigma op rconst_assoc in
+ let a = rconstant args.(0) in
+ let b = rconstant args.(1) in
+ f a b
+ with ParseError -> (
+ match op with
+ | op when EConstr.eq_constr sigma 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 EConstr.eq_constr sigma op (Lazy.force coq_Rpower) ->
+ Mc.CPow
+ ( rconstant args.(0)
+ , Mc.Inr (parse_more_constant nconstant gl args.(1)) )
+ | op when EConstr.eq_constr sigma op (Lazy.force coq_IQR) ->
+ Mc.CQ (qconstant gl args.(0))
+ | op when EConstr.eq_constr sigma op (Lazy.force coq_IZR) ->
+ Mc.CZ (parse_more_constant zconstant gl args.(0))
+ | _ -> raise ParseError ) )
+ | _ -> raise ParseError
+ in
rconstant term
-
-
let rconstant gl term =
- if debug
- then Feedback.msg_debug (Pp.str "rconstant: " ++ Printer.pr_leconstr_env gl.env gl.sigma term ++ fnl ());
+ if debug then
+ Feedback.msg_debug
+ ( Pp.str "rconstant: "
+ ++ Printer.pr_leconstr_env gl.env gl.sigma term
+ ++ fnl () );
let res = rconstant gl term in
- if debug then
- (Printf.printf "rconstant -> %a\n" pp_Rcst res ; flush stdout) ;
- res
-
-
-
- let parse_qexpr gl = parse_expr gl
- qconstant
- (fun expr x ->
- let exp = zconstant gl x in
+ if debug then (
+ Printf.printf "rconstant -> %a\n" pp_Rcst res;
+ flush stdout );
+ res
+
+ let parse_qexpr gl =
+ parse_expr gl qconstant
+ (fun expr x ->
+ let exp = zconstant gl x in
match exp with
- | Mc.Zneg _ ->
- begin
- match expr with
- | Mc.PEc q -> Mc.PEc (Mc.qpower q exp)
- | _ -> raise ParseError
- end
- | _ -> let exp = Mc.Z.to_N exp in
- Mc.PEpow(expr,exp))
- qop_spec
-
- let parse_rexpr gl = parse_expr gl
- rconstant
- (fun expr x ->
- let exp = Mc.N.of_nat (parse_nat gl.sigma x) in
- Mc.PEpow(expr,exp))
- rop_spec
-
- let parse_arith parse_op parse_expr env cstr gl =
+ | Mc.Zneg _ -> (
+ match expr with
+ | Mc.PEc q -> Mc.PEc (Mc.qpower q exp)
+ | _ -> raise ParseError )
+ | _ ->
+ let exp = Mc.Z.to_N exp in
+ Mc.PEpow (expr, exp))
+ qop_spec
+
+ let parse_rexpr gl =
+ parse_expr gl rconstant
+ (fun expr x ->
+ let exp = Mc.N.of_nat (parse_nat gl.sigma x) in
+ Mc.PEpow (expr, exp))
+ rop_spec
+
+ let parse_arith parse_op parse_expr env cstr gl =
let sigma = gl.sigma in
- if debug
- then Feedback.msg_debug (Pp.str "parse_arith: " ++ Printer.pr_leconstr_env gl.env sigma cstr ++ fnl ());
+ if debug then
+ Feedback.msg_debug
+ ( Pp.str "parse_arith: "
+ ++ Printer.pr_leconstr_env gl.env sigma cstr
+ ++ fnl () );
match EConstr.kind sigma cstr with
- | App(op,args) ->
- let (op,lhs,rhs) = parse_op gl (op,args) in
- let (e1,env) = parse_expr gl env lhs in
- let (e2,env) = parse_expr gl env rhs in
- ({Mc.flhs = e1; Mc.fop = op;Mc.frhs = e2},env)
- | _ -> failwith "error : parse_arith(2)"
+ | App (op, args) ->
+ let op, lhs, rhs = parse_op gl (op, args) in
+ let e1, env = parse_expr gl env lhs in
+ let e2, env = parse_expr gl env rhs in
+ ({Mc.flhs = e1; Mc.fop = op; Mc.frhs = e2}, env)
+ | _ -> failwith "error : parse_arith(2)"
let parse_zarith = parse_arith parse_zop parse_zexpr
-
let parse_qarith = parse_arith parse_qop parse_qexpr
-
let parse_rarith = parse_arith parse_rop parse_rexpr
(* generic parsing of arithmetic expressions *)
- let mkC f1 f2 = Mc.Cj(f1,f2)
- let mkD f1 f2 = Mc.D(f1,f2)
- let mkIff f1 f2 = Mc.Cj(Mc.I(f1,None,f2),Mc.I(f2,None,f1))
- let mkI f1 f2 = Mc.I(f1,None,f2)
+ let mkC f1 f2 = Mc.Cj (f1, f2)
+ let mkD f1 f2 = Mc.D (f1, f2)
+ let mkIff f1 f2 = Mc.Cj (Mc.I (f1, None, f2), Mc.I (f2, None, f1))
+ let mkI f1 f2 = Mc.I (f1, None, f2)
let mkformula_binary g term f1 f2 =
- match f1 , f2 with
- | Mc.X _ , Mc.X _ -> Mc.X(term)
- | _ -> g f1 f2
+ match (f1, f2) with Mc.X _, Mc.X _ -> Mc.X term | _ -> g f1 f2
(**
* This is the big generic function for formula parsers.
*)
let is_prop env sigma term =
- let sort = Retyping.get_sort_of env sigma term in
+ let sort = Retyping.get_sort_of env sigma term in
Sorts.is_prop sort
let parse_formula gl parse_atom env tg term =
let sigma = gl.sigma in
-
let is_prop term = is_prop gl.env gl.sigma term in
-
let parse_atom env tg t =
try
- let (at,env) = parse_atom env t gl in
- (Mc.A(at,(tg,t)), env,Tag.next tg)
+ let at, env = parse_atom env t gl in
+ (Mc.A (at, (tg, t)), env, Tag.next tg)
with ParseError ->
- if is_prop t
- then (Mc.X(t),env,tg)
- else raise ParseError
+ if is_prop t then (Mc.X t, env, tg) else raise ParseError
in
-
let rec xparse_formula env tg term =
match EConstr.kind sigma term with
- | App(l,rst) ->
- (match rst with
- | [|a;b|] when EConstr.eq_constr sigma 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 EConstr.eq_constr sigma 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
+ | App (l, rst) -> (
+ match rst with
+ | [|a; b|] when EConstr.eq_constr sigma 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 EConstr.eq_constr sigma 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 EConstr.eq_constr sigma l (Lazy.force coq_not) ->
- let (f,env,tg) = xparse_formula env tg a in (Mc.N(f), env,tg)
- | [|a;b|] when EConstr.eq_constr sigma 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 typ.binder_name = Anonymous || EConstr.Vars.noccurn sigma 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
- | _ -> if EConstr.eq_constr sigma term (Lazy.force coq_True)
- then (Mc.TT,env,tg)
- else if EConstr.eq_constr sigma term (Lazy.force coq_False)
- then Mc.(FF,env,tg)
- else if is_prop term then Mc.X(term),env,tg
- else raise ParseError
+ let f, env, tg = xparse_formula env tg a in
+ (Mc.N f, env, tg)
+ | [|a; b|] when EConstr.eq_constr sigma 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 typ.binder_name = Anonymous || EConstr.Vars.noccurn sigma 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)
+ | _ ->
+ if EConstr.eq_constr sigma term (Lazy.force coq_True) then
+ (Mc.TT, env, tg)
+ else if EConstr.eq_constr sigma term (Lazy.force coq_False) then
+ Mc.(FF, env, tg)
+ else if is_prop term then (Mc.X term, env, tg)
+ else raise ParseError
in
- xparse_formula env tg ((*Reductionops.whd_zeta*) term)
+ xparse_formula env tg (*Reductionops.whd_zeta*) term
let dump_formula typ dump_atom f =
let app_ctor c args =
- EConstr.mkApp(Lazy.force c, Array.of_list (typ::EConstr.mkProp::Lazy.force coq_unit :: Lazy.force coq_unit :: args)) in
-
+ EConstr.mkApp
+ ( Lazy.force c
+ , Array.of_list
+ ( typ :: EConstr.mkProp :: Lazy.force coq_unit
+ :: Lazy.force coq_unit :: args ) )
+ in
let rec xdump f =
- match f with
- | Mc.TT -> app_ctor coq_TT []
- | Mc.FF -> app_ctor coq_FF []
- | Mc.Cj(x,y) -> app_ctor coq_And [xdump x ; xdump y]
- | Mc.D(x,y) -> app_ctor coq_Or [xdump x ; xdump y]
- | Mc.I(x,_,y) -> app_ctor coq_Impl [xdump x ; EConstr.mkApp(Lazy.force coq_None,[|Lazy.force coq_unit|]); xdump y]
- | Mc.N(x) -> app_ctor coq_Neg [xdump x]
- | Mc.A(x,_) -> app_ctor coq_Atom [dump_atom x;Lazy.force coq_tt]
- | Mc.X(t) -> app_ctor coq_X [t] in
- xdump f
-
+ match f with
+ | Mc.TT -> app_ctor coq_TT []
+ | Mc.FF -> app_ctor coq_FF []
+ | Mc.Cj (x, y) -> app_ctor coq_And [xdump x; xdump y]
+ | Mc.D (x, y) -> app_ctor coq_Or [xdump x; xdump y]
+ | Mc.I (x, _, y) ->
+ app_ctor coq_Impl
+ [ xdump x
+ ; EConstr.mkApp (Lazy.force coq_None, [|Lazy.force coq_unit|])
+ ; xdump y ]
+ | Mc.N x -> app_ctor coq_Neg [xdump x]
+ | Mc.A (x, _) -> app_ctor coq_Atom [dump_atom x; Lazy.force coq_tt]
+ | Mc.X t -> app_ctor coq_X [t]
+ in
+ xdump f
let prop_env_of_formula gl form =
Mc.(
- let rec doit env = function
- | TT | FF | A(_,_) -> env
- | X t -> fst (Env.compute_rank_add env t)
- | Cj(f1,f2) | D(f1,f2) | I(f1,_,f2) ->
- doit (doit env f1) f2
- | N f -> doit env f
- in
-
- doit (Env.empty gl) form)
+ let rec doit env = function
+ | TT | FF | A (_, _) -> env
+ | X t -> fst (Env.compute_rank_add env t)
+ | Cj (f1, f2) | D (f1, f2) | I (f1, _, f2) -> doit (doit env f1) f2
+ | N f -> doit env f
+ in
+ doit (Env.empty gl) form)
let var_env_of_formula form =
-
- let rec vars_of_expr = function
+ 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) ->
+ | 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
+ | 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
+ Mc.(
+ let rec doit = function
+ | TT | FF | X _ -> ISet.empty
+ | A (a, (t, c)) -> vars_of_atom a
+ | Cj (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 : EConstr.constr
+ ; dump_cst : 'cst -> EConstr.constr
+ ; dump_add : EConstr.constr
+ ; dump_sub : EConstr.constr
+ ; dump_opp : EConstr.constr
+ ; dump_mul : EConstr.constr
+ ; dump_pow : EConstr.constr
+ ; dump_pow_arg : Mc.n -> EConstr.constr
+ ; dump_op : (Mc.op2 * EConstr.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 rec dump_Rcst_as_R cst =
+ match cst with
+ | Mc.C0 -> Lazy.force coq_R0
+ | Mc.C1 -> Lazy.force coq_R1
+ | Mc.CQ q -> EConstr.mkApp (Lazy.force coq_IQR, [|dump_q q|])
+ | Mc.CZ z -> EConstr.mkApp (Lazy.force coq_IZR, [|dump_z z|])
+ | Mc.CPlus (x, y) ->
+ EConstr.mkApp
+ (Lazy.force coq_Rplus, [|dump_Rcst_as_R x; dump_Rcst_as_R y|])
+ | Mc.CMinus (x, y) ->
+ EConstr.mkApp
+ (Lazy.force coq_Rminus, [|dump_Rcst_as_R x; dump_Rcst_as_R y|])
+ | Mc.CMult (x, y) ->
+ EConstr.mkApp
+ (Lazy.force coq_Rmult, [|dump_Rcst_as_R x; dump_Rcst_as_R y|])
+ | Mc.CPow (x, y) -> (
+ match y with
+ | Mc.Inl z ->
+ EConstr.mkApp (Lazy.force coq_powerZR, [|dump_Rcst_as_R x; dump_z z|])
+ | Mc.Inr n ->
+ EConstr.mkApp (Lazy.force coq_Rpower, [|dump_Rcst_as_R x; dump_nat n|])
+ )
+ | Mc.CInv t -> EConstr.mkApp (Lazy.force coq_Rinv, [|dump_Rcst_as_R t|])
+ | Mc.COpp t -> EConstr.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 }
+
+ let prodn n env b =
+ let rec prodrec = function
+ | 0, env, b -> b
+ | n, (v, t) :: l, b ->
+ prodrec (n - 1, l, EConstr.mkProd (make_annot v Sorts.Relevant, t, b))
+ | _ -> assert false
+ in
+ prodrec (n, env, b)
- let vars_of_atom {Mc.flhs ; Mc.fop; Mc.frhs} =
- ISet.union (vars_of_expr flhs) (vars_of_expr frhs) in
- Mc.(
- let rec doit = function
- | TT | FF | X _ -> ISet.empty
- | A (a,(t,c)) -> vars_of_atom a
- | Cj(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 : EConstr.constr;
- dump_cst : 'cst -> EConstr.constr;
- dump_add : EConstr.constr;
- dump_sub : EConstr.constr;
- dump_opp : EConstr.constr;
- dump_mul : EConstr.constr;
- dump_pow : EConstr.constr;
- dump_pow_arg : Mc.n -> EConstr.constr;
- dump_op : (Mc.op2 * EConstr.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 rec dump_Rcst_as_R cst =
- match cst with
- | Mc.C0 -> Lazy.force coq_R0
- | Mc.C1 -> Lazy.force coq_R1
- | Mc.CQ q -> EConstr.mkApp(Lazy.force coq_IQR, [| dump_q q |])
- | Mc.CZ z -> EConstr.mkApp(Lazy.force coq_IZR, [| dump_z z |])
- | Mc.CPlus(x,y) -> EConstr.mkApp(Lazy.force coq_Rplus, [| dump_Rcst_as_R x ; dump_Rcst_as_R y |])
- | Mc.CMinus(x,y) -> EConstr.mkApp(Lazy.force coq_Rminus, [| dump_Rcst_as_R x ; dump_Rcst_as_R y |])
- | Mc.CMult(x,y) -> EConstr.mkApp(Lazy.force coq_Rmult, [| dump_Rcst_as_R x ; dump_Rcst_as_R y |])
- | Mc.CPow(x,y) ->
- begin
- match y with
- | Mc.Inl z -> EConstr.mkApp(Lazy.force coq_powerZR,[| dump_Rcst_as_R x ; dump_z z|])
- | Mc.Inr n -> EConstr.mkApp(Lazy.force coq_Rpower,[| dump_Rcst_as_R x ; dump_nat n|])
- end
- | Mc.CInv t -> EConstr.mkApp(Lazy.force coq_Rinv, [| dump_Rcst_as_R t |])
- | Mc.COpp t -> EConstr.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
- }
-
-
-
-
-let prodn n env b =
- let rec prodrec = function
- | (0, env, b) -> b
- | (n, ((v,t)::l), b) -> prodrec (n-1, l, EConstr.mkProd (make_annot v Sorts.Relevant,t,b))
- | _ -> assert false
- in
- prodrec (n,env,b)
-
-(** [make_goal_of_formula depxr vars props form] where
+ (** [make_goal_of_formula depxr vars props form] where
- vars is an environment for the arithmetic variables occurring in form
- props is an environment for the propositions occurring in form
@return a goal where all the variables and propositions of the formula are quantified
*)
-let make_goal_of_formula gl 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 gl form in
-
- let fresh_var str i = Names.Id.of_string (str^(string_of_int i)) in
-
- let fresh_prop str i =
- Names.Id.of_string (str^(string_of_int i)) in
-
- let vars_n = List.map (fun (_,i) -> fresh_var "__x" i, dexpr.interp_typ) vars_idx in
- let props_n = List.mapi (fun i _ -> fresh_prop "__p" (i+1) , EConstr.mkProp) (Env.elements 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 -> EConstr.mkRel (i+(List.assoc (CoqToCaml.positive n) vars_idx))
- | Mc.PEc z -> dexpr.dump_cst z
- | Mc.PEadd(e1,e2) -> EConstr.mkApp(dexpr.dump_add,
- [| dump_expr e1;dump_expr e2|])
- | Mc.PEsub(e1,e2) -> EConstr.mkApp(dexpr.dump_sub,
- [| dump_expr e1;dump_expr e2|])
- | Mc.PEopp e -> EConstr.mkApp(dexpr.dump_opp,
- [| dump_expr e|])
- | Mc.PEmul(e1,e2) -> EConstr.mkApp(dexpr.dump_mul,
- [| dump_expr e1;dump_expr e2|])
- | Mc.PEpow(e,n) -> EConstr.mkApp(dexpr.dump_pow,
- [| dump_expr e; dexpr.dump_pow_arg n|])
- in dump_expr e in
-
- let mkop op e1 e2 =
- try
- EConstr.mkApp(List.assoc op dexpr.dump_op, [| e1; e2|])
+ let make_goal_of_formula gl 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 gl form in
+ let fresh_var str i = Names.Id.of_string (str ^ string_of_int i) in
+ let fresh_prop str i = Names.Id.of_string (str ^ string_of_int i) in
+ let vars_n =
+ List.map (fun (_, i) -> (fresh_var "__x" i, dexpr.interp_typ)) vars_idx
+ in
+ let props_n =
+ List.mapi
+ (fun i _ -> (fresh_prop "__p" (i + 1), EConstr.mkProp))
+ (Env.elements 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 ->
+ EConstr.mkRel (i + List.assoc (CoqToCaml.positive n) vars_idx)
+ | Mc.PEc z -> dexpr.dump_cst z
+ | Mc.PEadd (e1, e2) ->
+ EConstr.mkApp (dexpr.dump_add, [|dump_expr e1; dump_expr e2|])
+ | Mc.PEsub (e1, e2) ->
+ EConstr.mkApp (dexpr.dump_sub, [|dump_expr e1; dump_expr e2|])
+ | Mc.PEopp e -> EConstr.mkApp (dexpr.dump_opp, [|dump_expr e|])
+ | Mc.PEmul (e1, e2) ->
+ EConstr.mkApp (dexpr.dump_mul, [|dump_expr e1; dump_expr e2|])
+ | Mc.PEpow (e, n) ->
+ EConstr.mkApp (dexpr.dump_pow, [|dump_expr e; dexpr.dump_pow_arg n|])
+ in
+ dump_expr e
+ in
+ let mkop op e1 e2 =
+ try EConstr.mkApp (List.assoc op dexpr.dump_op, [|e1; e2|])
with Not_found ->
- EConstr.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
- | Mc.TT -> Lazy.force coq_True
- | Mc.FF -> Lazy.force coq_False
- | Mc.Cj(x,y) -> EConstr.mkApp(Lazy.force coq_and,[|xdump pi xi x ; xdump pi xi y|])
- | Mc.D(x,y) -> EConstr.mkApp(Lazy.force coq_or,[| xdump pi xi x ; xdump pi xi y|])
- | Mc.I(x,_,y) -> EConstr.mkArrow (xdump pi xi x) Sorts.Relevant (xdump (pi+1) (xi+1) y)
- | Mc.N(x) -> EConstr.mkArrow (xdump pi xi x) Sorts.Relevant (Lazy.force coq_False)
- | Mc.A(x,_) -> dump_cstr xi x
- | Mc.X(t) -> let idx = Env.get_rank props t in
- EConstr.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
- EConstr.mkVar (Names.Id.of_string (Printf.sprintf "__p%i" idx)) in
-
- let form' = Mc.mapX subst_prop form in
-
- (prodn nb_props (List.map (fun (x,y) -> Name.Name x,y) props_n)
- (prodn nb_vars (List.map (fun (x,y) -> Name.Name x,y) vars_n)
- (xdump (List.length vars_n) 0 form)),
- List.rev props_n, List.rev var_name_pos,form')
+ EConstr.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
+ | Mc.TT -> Lazy.force coq_True
+ | Mc.FF -> Lazy.force coq_False
+ | Mc.Cj (x, y) ->
+ EConstr.mkApp (Lazy.force coq_and, [|xdump pi xi x; xdump pi xi y|])
+ | Mc.D (x, y) ->
+ EConstr.mkApp (Lazy.force coq_or, [|xdump pi xi x; xdump pi xi y|])
+ | Mc.I (x, _, y) ->
+ EConstr.mkArrow (xdump pi xi x) Sorts.Relevant
+ (xdump (pi + 1) (xi + 1) y)
+ | Mc.N x ->
+ EConstr.mkArrow (xdump pi xi x) Sorts.Relevant (Lazy.force coq_False)
+ | Mc.A (x, _) -> dump_cstr xi x
+ | Mc.X t ->
+ let idx = Env.get_rank props t in
+ EConstr.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
+ EConstr.mkVar (Names.Id.of_string (Printf.sprintf "__p%i" idx))
+ in
+ let form' = Mc.mapX subst_prop form in
+ ( prodn nb_props
+ (List.map (fun (x, y) -> (Name.Name x, y)) props_n)
+ (prodn nb_vars
+ (List.map (fun (x, y) -> (Name.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
@@ -1349,177 +1349,188 @@ let make_goal_of_formula gl dexpr form =
*)
let set l concl =
- let rec xset acc = function
- | [] -> acc
- | (e::l) ->
- let (name,expr,typ) = e in
- xset (EConstr.mkNamedLetIn
- (make_annot (Names.Id.of_string name) Sorts.Relevant)
- expr typ acc) l in
+ let rec xset acc = function
+ | [] -> acc
+ | e :: l ->
+ let name, expr, typ = e in
+ xset
+ (EConstr.mkNamedLetIn
+ (make_annot (Names.Id.of_string name) Sorts.Relevant)
+ expr typ acc)
+ l
+ in
xset concl l
-
-end (**
- * MODULE END: M
- *)
+end
open M
let coq_Branch =
- lazy (gen_constant_in_modules "VarMap"
- [["Coq" ; "micromega" ; "VarMap"];["VarMap"]] "Branch")
+ lazy
+ (gen_constant_in_modules "VarMap"
+ [["Coq"; "micromega"; "VarMap"]; ["VarMap"]]
+ "Branch")
+
let coq_Elt =
- lazy (gen_constant_in_modules "VarMap"
- [["Coq" ; "micromega" ; "VarMap"];["VarMap"]] "Elt")
+ lazy
+ (gen_constant_in_modules "VarMap"
+ [["Coq"; "micromega"; "VarMap"]; ["VarMap"]]
+ "Elt")
+
let coq_Empty =
- lazy (gen_constant_in_modules "VarMap"
- [["Coq" ; "micromega" ;"VarMap"];["VarMap"]] "Empty")
+ lazy
+ (gen_constant_in_modules "VarMap"
+ [["Coq"; "micromega"; "VarMap"]; ["VarMap"]]
+ "Empty")
let coq_VarMap =
- lazy (gen_constant_in_modules "VarMap"
- [["Coq" ; "micromega" ; "VarMap"] ; ["VarMap"]] "t")
-
+ lazy
+ (gen_constant_in_modules "VarMap"
+ [["Coq"; "micromega"; "VarMap"]; ["VarMap"]]
+ "t")
let rec dump_varmap typ m =
match m with
- | Mc.Empty -> EConstr.mkApp(Lazy.force coq_Empty,[| typ |])
- | Mc.Elt v -> EConstr.mkApp(Lazy.force coq_Elt,[| typ; v|])
- | Mc.Branch(l,o,r) ->
- EConstr.mkApp (Lazy.force coq_Branch, [| typ; dump_varmap typ l; o ; dump_varmap typ r |])
-
+ | Mc.Empty -> EConstr.mkApp (Lazy.force coq_Empty, [|typ|])
+ | Mc.Elt v -> EConstr.mkApp (Lazy.force coq_Elt, [|typ; v|])
+ | Mc.Branch (l, o, r) ->
+ EConstr.mkApp
+ (Lazy.force coq_Branch, [|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
+ | (d, _) :: _ ->
+ List.fold_left
+ (fun vm (c, i) -> Mc.vm_add d (CamlToCoq.positive i) c vm)
+ Mc.Empty env
let rec dump_proof_term = function
| Micromega.DoneProof -> Lazy.force coq_doneProof
- | Micromega.RatProof(cone,rst) ->
- EConstr.mkApp(Lazy.force coq_ratProof, [| dump_psatz coq_Z dump_z cone; dump_proof_term rst|])
- | Micromega.CutProof(cone,prf) ->
- EConstr.mkApp(Lazy.force coq_cutProof,
- [| dump_psatz coq_Z dump_z cone ;
- dump_proof_term prf|])
- | Micromega.EnumProof(c1,c2,prfs) ->
- EConstr.mkApp (Lazy.force coq_enumProof,
- [| dump_psatz coq_Z dump_z c1 ; dump_psatz coq_Z dump_z c2 ;
- dump_list (Lazy.force coq_proofTerm) dump_proof_term prfs |])
-
+ | Micromega.RatProof (cone, rst) ->
+ EConstr.mkApp
+ ( Lazy.force coq_ratProof
+ , [|dump_psatz coq_Z dump_z cone; dump_proof_term rst|] )
+ | Micromega.CutProof (cone, prf) ->
+ EConstr.mkApp
+ ( Lazy.force coq_cutProof
+ , [|dump_psatz coq_Z dump_z cone; dump_proof_term prf|] )
+ | Micromega.EnumProof (c1, c2, prfs) ->
+ EConstr.mkApp
+ ( Lazy.force coq_enumProof
+ , [| dump_psatz coq_Z dump_z c1
+ ; dump_psatz coq_Z dump_z c2
+ ; dump_list (Lazy.force coq_proofTerm) dump_proof_term prfs |] )
let rec size_of_psatz = function
| Micromega.PsatzIn _ -> 1
| Micromega.PsatzSquare _ -> 1
- | Micromega.PsatzMulC(_,p) -> 1 + (size_of_psatz p)
- | Micromega.PsatzMulE(p1,p2) | Micromega.PsatzAdd(p1,p2) -> size_of_psatz p1 + size_of_psatz p2
+ | Micromega.PsatzMulC (_, p) -> 1 + size_of_psatz p
+ | Micromega.PsatzMulE (p1, p2) | Micromega.PsatzAdd (p1, p2) ->
+ size_of_psatz p1 + size_of_psatz p2
| Micromega.PsatzC _ -> 1
- | Micromega.PsatzZ -> 1
+ | Micromega.PsatzZ -> 1
let rec size_of_pf = function
| Micromega.DoneProof -> 1
- | Micromega.RatProof(p,a) -> (size_of_pf a) + (size_of_psatz p)
- | Micromega.CutProof(p,a) -> (size_of_pf a) + (size_of_psatz p)
- | Micromega.EnumProof(p1,p2,l) -> (size_of_psatz p1) + (size_of_psatz p2) + (List.fold_left (fun acc p -> size_of_pf p + acc) 0 l)
+ | Micromega.RatProof (p, a) -> size_of_pf a + size_of_psatz p
+ | Micromega.CutProof (p, a) -> size_of_pf a + size_of_psatz p
+ | Micromega.EnumProof (p1, p2, l) ->
+ size_of_psatz p1 + size_of_psatz p2
+ + List.fold_left (fun acc p -> size_of_pf p + acc) 0 l
let dump_proof_term t =
- if debug then Printf.printf "dump_proof_term %i\n" (size_of_pf t) ;
+ if debug then Printf.printf "dump_proof_term %i\n" (size_of_pf t);
dump_proof_term t
-
-
-let pp_q o q = Printf.fprintf o "%a/%a" pp_z q.Micromega.qnum pp_positive q.Micromega.qden
-
+let pp_q o q =
+ Printf.fprintf o "%a/%a" pp_z q.Micromega.qnum pp_positive q.Micromega.qden
let rec pp_proof_term o = function
| Micromega.DoneProof -> Printf.fprintf o "D"
- | Micromega.RatProof(cone,rst) -> Printf.fprintf o "R[%a,%a]" (pp_psatz pp_z) cone pp_proof_term rst
- | Micromega.CutProof(cone,rst) -> Printf.fprintf o "C[%a,%a]" (pp_psatz pp_z) cone pp_proof_term rst
- | Micromega.EnumProof(c1,c2,rst) ->
- Printf.fprintf o "EP[%a,%a,%a]"
- (pp_psatz pp_z) c1 (pp_psatz pp_z) c2
- (pp_list "[" "]" pp_proof_term) rst
+ | Micromega.RatProof (cone, rst) ->
+ Printf.fprintf o "R[%a,%a]" (pp_psatz pp_z) cone pp_proof_term rst
+ | Micromega.CutProof (cone, rst) ->
+ Printf.fprintf o "C[%a,%a]" (pp_psatz pp_z) cone pp_proof_term rst
+ | Micromega.EnumProof (c1, c2, rst) ->
+ Printf.fprintf o "EP[%a,%a,%a]" (pp_psatz pp_z) c1 (pp_psatz pp_z) c2
+ (pp_list "[" "]" pp_proof_term)
+ rst
let rec parse_hyps gl parse_arith env tg hyps =
- match hyps with
- | [] -> ([],env,tg)
- | (i,t)::l ->
- let (lhyps,env,tg) = parse_hyps gl parse_arith env tg l in
- if is_prop gl.env gl.sigma t
- then
- try
- let (c,env,tg) = parse_formula gl parse_arith env tg t in
- ((i,c)::lhyps, env,tg)
- with ParseError -> (lhyps,env,tg)
- else (lhyps,env,tg)
-
-
-let parse_goal gl parse_arith (env:Env.t) hyps term =
- let (f,env,tg) = parse_formula gl parse_arith env (Tag.from 0) term in
- let (lhyps,env,tg) = parse_hyps gl parse_arith env tg hyps in
- (lhyps,f,env)
-
+ match hyps with
+ | [] -> ([], env, tg)
+ | (i, t) :: l ->
+ let lhyps, env, tg = parse_hyps gl parse_arith env tg l in
+ if is_prop gl.env gl.sigma t then
+ try
+ let c, env, tg = parse_formula gl parse_arith env tg t in
+ ((i, c) :: lhyps, env, tg)
+ with ParseError -> (lhyps, env, tg)
+ else (lhyps, env, tg)
+
+let parse_goal gl parse_arith (env : Env.t) hyps term =
+ let f, env, tg = parse_formula gl parse_arith env (Tag.from 0) term in
+ let lhyps, env, tg = parse_hyps gl parse_arith env tg hyps in
+ (lhyps, f, env)
+
+type ('synt_c, 'prf) domain_spec =
+ { typ : EConstr.constr
+ ; (* is the type of the interpretation domain - Z, Q, R*)
+ coeff : EConstr.constr
+ ; (* is the type of the syntactic coeffs - Z , Q , Rcst *)
+ dump_coeff : 'synt_c -> EConstr.constr
+ ; proof_typ : EConstr.constr
+ ; dump_proof : 'prf -> EConstr.constr }
(**
* The datastructures that aggregate theory-dependent proof values.
*)
-type ('synt_c, 'prf) domain_spec = {
- typ : EConstr.constr; (* is the type of the interpretation domain - Z, Q, R*)
- coeff : EConstr.constr ; (* is the type of the syntactic coeffs - Z , Q , Rcst *)
- dump_coeff : 'synt_c -> EConstr.constr ;
- proof_typ : EConstr.constr ;
- dump_proof : 'prf -> EConstr.constr
-}
-
-let zz_domain_spec = lazy {
- typ = Lazy.force coq_Z;
- coeff = Lazy.force coq_Z;
- dump_coeff = dump_z ;
- proof_typ = Lazy.force coq_proofTerm ;
- dump_proof = dump_proof_term
-}
-
-let qq_domain_spec = lazy {
- typ = Lazy.force coq_Q;
- coeff = Lazy.force coq_Q;
- dump_coeff = dump_q ;
- proof_typ = Lazy.force coq_QWitness ;
- dump_proof = dump_psatz coq_Q dump_q
-}
-
-let max_tag f = 1 + (Tag.to_int (Mc.foldA (fun t1 (t2,_) -> Tag.max t1 t2) f (Tag.from 0)))
+let zz_domain_spec =
+ lazy
+ { typ = Lazy.force coq_Z
+ ; coeff = Lazy.force coq_Z
+ ; dump_coeff = dump_z
+ ; proof_typ = Lazy.force coq_proofTerm
+ ; dump_proof = dump_proof_term }
+
+let qq_domain_spec =
+ lazy
+ { typ = Lazy.force coq_Q
+ ; coeff = Lazy.force coq_Q
+ ; dump_coeff = dump_q
+ ; proof_typ = Lazy.force coq_QWitness
+ ; dump_proof = dump_psatz coq_Q dump_q }
+
+let max_tag f =
+ 1 + Tag.to_int (Mc.foldA (fun t1 (t2, _) -> Tag.max t1 t2) f (Tag.from 0))
(** For completeness of the cutting-plane procedure,
each variable 'x' is replaced by 'y' - 'z' where
'y' and 'z' are positive *)
let pre_processZ mt f =
-
let x0 i = 2 * i in
- let x1 i = 2 * i + 1 in
-
+ let x1 i = (2 * i) + 1 in
let tag_of_var fr p b =
-
- let ip = CoqToCaml.positive fr + (CoqToCaml.positive p) in
-
+ let ip = CoqToCaml.positive fr + CoqToCaml.positive p in
match b with
| None ->
- let y = Mc.XO (Mc.Coq_Pos.add fr p) in
- let z = Mc.XI (Mc.Coq_Pos.add fr p) in
- let tag = Tag.from (- x0 (x0 ip)) in
- let constr = Mc.mk_eq_pos p y z in
- (tag, dump_cstr (Lazy.force coq_Z) dump_z constr)
+ let y = Mc.XO (Mc.Coq_Pos.add fr p) in
+ let z = Mc.XI (Mc.Coq_Pos.add fr p) in
+ let tag = Tag.from (-x0 (x0 ip)) in
+ let constr = Mc.mk_eq_pos p y z in
+ (tag, dump_cstr (Lazy.force coq_Z) dump_z constr)
| Some false ->
- let y = Mc.XO (Mc.Coq_Pos.add fr p) in
- let tag = Tag.from (- x0 (x1 ip)) in
- let constr = Mc.bound_var (Mc.XO y) in
- (tag, dump_cstr (Lazy.force coq_Z) dump_z constr)
+ let y = Mc.XO (Mc.Coq_Pos.add fr p) in
+ let tag = Tag.from (-x0 (x1 ip)) in
+ let constr = Mc.bound_var (Mc.XO y) in
+ (tag, dump_cstr (Lazy.force coq_Z) dump_z constr)
| Some true ->
- let z = Mc.XI (Mc.Coq_Pos.add fr p) in
- let tag = Tag.from (- x1 (x1 ip)) in
- let constr = Mc.bound_var (Mc.XI z) in
- (tag, dump_cstr (Lazy.force coq_Z) dump_z constr) in
-
- Mc.bound_problem_fr tag_of_var mt f
+ let z = Mc.XI (Mc.Coq_Pos.add fr p) in
+ let tag = Tag.from (-x1 (x1 ip)) in
+ let constr = Mc.bound_var (Mc.XI z) in
+ (tag, dump_cstr (Lazy.force coq_Z) dump_z constr)
+ in
+ Mc.bound_problem_fr tag_of_var mt f
(** Naive topological sort of constr according to the subterm-ordering *)
(* An element is minimal x is minimal w.r.t y if
@@ -1530,26 +1541,25 @@ let pre_processZ mt f =
* witness.
*)
-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 = (EConstr.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) (vm_of_list env) in
- (* todo : directly generate the proof term - or generalize before conversion? *)
- Proofview.Goal.enter begin fun gl ->
- Tacticals.New.tclTHENLIST
- [
- Tactics.change_concl
- (set
- [
- ("__ff", ff, EConstr.mkApp(Lazy.force coq_Formula, [|formula_typ |]));
- ("__varmap", vm, EConstr.mkApp(Lazy.force coq_VarMap, [|spec.typ|]));
- ("__wit", cert, cert_typ)
- ]
- (Tacmach.New.pf_concl gl))
- ]
- end
-
+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 = EConstr.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 (vm_of_list env) in
+ (* todo : directly generate the proof term - or generalize before conversion? *)
+ Proofview.Goal.enter (fun gl ->
+ Tacticals.New.tclTHENLIST
+ [ Tactics.change_concl
+ (set
+ [ ( "__ff"
+ , ff
+ , EConstr.mkApp (Lazy.force coq_Formula, [|formula_typ|]) )
+ ; ( "__varmap"
+ , vm
+ , EConstr.mkApp (Lazy.force coq_VarMap, [|spec.typ|]) )
+ ; ("__wit", cert, cert_typ) ]
+ (Tacmach.New.pf_concl gl)) ])
(**
* The datastructures that aggregate prover attributes.
@@ -1557,17 +1567,21 @@ let micromega_order_change spec cert cert_typ env ff (*: unit Proofview.tactic*
open Certificate
-type ('option,'a,'prf,'model) prover = {
- name : string ; (* name of the prover *)
- get_option : unit ->'option ; (* find the options of the prover *)
- prover : ('option * 'a list) -> ('prf, 'model) Certificate.res ; (* 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)*)
-}
-
-
+type ('option, 'a, 'prf, 'model) prover =
+ { name : string
+ ; (* name of the prover *)
+ get_option : unit -> 'option
+ ; (* find the options of the prover *)
+ prover : 'option * 'a list -> ('prf, 'model) Certificate.res
+ ; (* 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 prover and a disjunction of atoms, find a proof of any of
@@ -1575,34 +1589,36 @@ type ('option,'a,'prf,'model) prover = {
* datastructure.
*)
-let find_witness p polys1 =
+let find_witness p polys1 =
let polys1 = List.map fst polys1 in
match p.prover (p.get_option (), polys1) with
| Model m -> Model m
| Unknown -> Unknown
- | Prf prf -> Prf(prf,p)
+ | Prf prf -> Prf (prf, p)
(**
* Given a prover and a CNF, find a proof for each of the clauses.
* Return the proofs as a list.
*)
-let witness_list prover l =
- let rec xwitness_list l =
- match l with
- | [] -> Prf []
- | e :: l ->
+let witness_list prover l =
+ let rec xwitness_list l =
+ match l with
+ | [] -> Prf []
+ | e :: l -> (
match xwitness_list l with
- | Model (m,e) -> Model (m,e)
- | Unknown -> Unknown
- | Prf l ->
- match find_witness prover e with
- | Model m -> Model (m,e)
- | Unknown -> Unknown
- | Prf w -> Prf (w::l) in
- xwitness_list l
-
-let witness_list_tags p g = witness_list p g
+ | Model (m, e) -> Model (m, e)
+ | Unknown -> Unknown
+ | Prf l -> (
+ match find_witness prover e with
+ | Model m -> Model (m, e)
+ | Unknown -> Unknown
+ | Prf w -> Prf (w :: l) ) )
+ in
+ xwitness_list l
+
+let witness_list_tags p g = witness_list p g
+
(* let t1 = System.get_time () in
let res = witness_list p g in
let t2 = System.get_time () in
@@ -1614,15 +1630,17 @@ let witness_list_tags p g = witness_list p g
* Prune the proof object, according to the 'diff' between two cnf formulas.
*)
-
-let compact_proofs (cnf_ff: 'cst cnf) res (cnf_ff': 'cst cnf) =
-
- let compact_proof (old_cl:'cst clause) (prf,prover) (new_cl:'cst clause) =
- let new_cl = List.mapi (fun i (f,_) -> (f,i)) new_cl in
+let compact_proofs (cnf_ff : 'cst cnf) res (cnf_ff' : 'cst cnf) =
+ let compact_proof (old_cl : 'cst clause) (prf, prover) (new_cl : 'cst clause)
+ =
+ let new_cl = List.mapi (fun i (f, _) -> (f, i)) new_cl in
let remap i =
- let formula = try fst (List.nth old_cl i) with Failure _ -> failwith "bad old index" in
- List.assoc formula new_cl in
-(* if debug then
+ let formula =
+ try fst (List.nth old_cl i) with Failure _ -> failwith "bad old index"
+ in
+ List.assoc formula new_cl
+ in
+ (* if debug then
begin
Printf.printf "\ncompact_proof : %a %a %a"
(pp_ml_list prover.pp_f) (List.map fst old_cl)
@@ -1630,91 +1648,96 @@ 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 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 (prover.get_option (), List.map fst new_cl) with
+ 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 (prover.get_option (), List.map fst new_cl) with
| Unknown | Model _ -> failwith "proof compaction error"
- | Prf p -> p
+ | Prf p -> p )
in
- if debug then
- begin
- Printf.printf " -> %a\n"
- prover.pp_prf res ;
- flush stdout
- end ;
- res in
-
- let is_proof_compatible (old_cl:'cst clause) (prf,prover) (new_cl:'cst clause) =
+ if debug then begin
+ Printf.printf " -> %a\n" prover.pp_prf res;
+ flush stdout
+ end;
+ res
+ in
+ let is_proof_compatible (old_cl : 'cst clause) (prf, prover)
+ (new_cl : 'cst clause) =
let hyps_idx = prover.hyps prf in
let hyps = selecti hyps_idx old_cl in
- is_sublist (=) hyps new_cl in
-
-
-
- let cnf_res = List.combine cnf_ff res in (* we get pairs clause * proof *)
- if debug then
- begin
- Printf.printf "CNFRES\n"; flush stdout;
- Printf.printf "CNFOLD %a\n" pp_cnf_tag cnf_ff;
- List.iter (fun (cl,(prf,prover)) ->
- let hyps_idx = prover.hyps prf in
- let hyps = selecti hyps_idx cl in
- Printf.printf "\nProver %a -> %a\n"
- pp_clause_tag cl pp_clause_tag hyps;flush stdout) cnf_res;
- Printf.printf "CNFNEW %a\n" pp_cnf_tag cnf_ff';
-
- end;
-
- List.map (fun x ->
- let (o,p) =
- try
- List.find (fun (l,p) -> is_proof_compatible l p x) cnf_res
+ is_sublist ( = ) hyps new_cl
+ in
+ let cnf_res = List.combine cnf_ff res in
+ (* we get pairs clause * proof *)
+ if debug then begin
+ Printf.printf "CNFRES\n";
+ flush stdout;
+ Printf.printf "CNFOLD %a\n" pp_cnf_tag cnf_ff;
+ List.iter
+ (fun (cl, (prf, prover)) ->
+ let hyps_idx = prover.hyps prf in
+ let hyps = selecti hyps_idx cl in
+ Printf.printf "\nProver %a -> %a\n" pp_clause_tag cl pp_clause_tag hyps;
+ flush stdout)
+ cnf_res;
+ Printf.printf "CNFNEW %a\n" pp_cnf_tag cnf_ff'
+ end;
+ List.map
+ (fun x ->
+ let o, p =
+ try List.find (fun (l, p) -> is_proof_compatible l p x) cnf_res
with Not_found ->
- begin
- Printf.printf "ERROR: no compatible proof" ; flush stdout;
- failwith "Cannot find compatible proof" end
- in
- compact_proof o p x) cnf_ff'
-
+ Printf.printf "ERROR: no compatible proof";
+ flush stdout;
+ failwith "Cannot find compatible proof"
+ in
+ compact_proof o p x)
+ cnf_ff'
(**
* "Hide out" tagged atoms of a formula by transforming them into generic
* variables. See the Tag module in mutils.ml for more.
*)
-
-
let abstract_formula : TagSet.t -> 'a formula -> 'a formula =
- fun hyps f ->
- let to_constr = Mc.({
- mkTT = Lazy.force coq_True;
- mkFF = Lazy.force coq_False;
- mkA = (fun a (tg, t) -> t);
- mkCj = (let coq_and = Lazy.force coq_and in
- fun x y -> EConstr.mkApp(coq_and,[|x;y|]));
- mkD = (let coq_or = Lazy.force coq_or in
- fun x y -> EConstr.mkApp(coq_or,[|x;y|]));
- mkI = (fun x y -> EConstr.mkArrow x Sorts.Relevant y);
- mkN = (let coq_not = Lazy.force coq_not in
- (fun x -> EConstr.mkApp(coq_not,[|x|])))
- }) in
- Mc.abst_form to_constr (fun (t,_) -> TagSet.mem t hyps) true f
-
+ fun hyps f ->
+ let to_constr =
+ Mc.
+ { mkTT = Lazy.force coq_True
+ ; mkFF = Lazy.force coq_False
+ ; mkA = (fun a (tg, t) -> t)
+ ; mkCj =
+ (let coq_and = Lazy.force coq_and in
+ fun x y -> EConstr.mkApp (coq_and, [|x; y|]))
+ ; mkD =
+ (let coq_or = Lazy.force coq_or in
+ fun x y -> EConstr.mkApp (coq_or, [|x; y|]))
+ ; mkI = (fun x y -> EConstr.mkArrow x Sorts.Relevant y)
+ ; mkN =
+ (let coq_not = Lazy.force coq_not in
+ fun x -> EConstr.mkApp (coq_not, [|x|])) }
+ in
+ Mc.abst_form to_constr (fun (t, _) -> TagSet.mem t hyps) true f
(* [abstract_wrt_formula] is used in contexts whre f1 is already an abstraction of f2 *)
let rec abstract_wrt_formula f1 f2 =
Mc.(
- match f1 , f2 with
- | X c , _ -> X c
- | A _ , A _ -> f2
- | Cj(a,b) , Cj(a',b') -> Cj(abstract_wrt_formula a a', abstract_wrt_formula b b')
- | D(a,b) , D(a',b') -> D(abstract_wrt_formula a a', abstract_wrt_formula b b')
- | I(a,_,b) , I(a',x,b') -> I(abstract_wrt_formula a a',x, abstract_wrt_formula b b')
- | FF , FF -> FF
- | TT , TT -> TT
- | N x , N y -> N(abstract_wrt_formula x y)
- | _ -> failwith "abstract_wrt_formula")
+ match (f1, f2) with
+ | X c, _ -> X c
+ | A _, A _ -> f2
+ | Cj (a, b), Cj (a', b') ->
+ Cj (abstract_wrt_formula a a', abstract_wrt_formula b b')
+ | D (a, b), D (a', b') ->
+ D (abstract_wrt_formula a a', abstract_wrt_formula b b')
+ | I (a, _, b), I (a', x, b') ->
+ I (abstract_wrt_formula a a', x, abstract_wrt_formula b b')
+ | FF, FF -> FF
+ | TT, TT -> TT
+ | N x, N y -> N (abstract_wrt_formula x y)
+ | _ -> failwith "abstract_wrt_formula")
(**
* This exception is raised by really_call_csdpcert if Coq's configure didn't
@@ -1723,7 +1746,6 @@ 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.
@@ -1731,12 +1753,11 @@ exception CsdpNotFound
let formula_hyps_concl hyps concl =
List.fold_right
- (fun (id,f) (cc,ids) ->
- match f with
- Mc.X _ -> (cc,ids)
- | _ -> (Mc.I(f,Some id,cc), id::ids))
- hyps (concl,[])
-
+ (fun (id, f) (cc, ids) ->
+ match f with
+ | Mc.X _ -> (cc, ids)
+ | _ -> (Mc.I (f, Some id, cc), id :: ids))
+ hyps (concl, [])
(* let time str f x =
let t1 = System.get_time () in
@@ -1746,70 +1767,76 @@ let formula_hyps_concl hyps concl =
res
*)
-let micromega_tauto pre_process cnf spec prover env (polys1: (Names.Id.t * 'cst formula) list) (polys2: 'cst formula) gl =
-
- (* Express the goal as one big implication *)
- let (ff,ids) = formula_hyps_concl polys1 polys2 in
- let mt = CamlToCoq.positive (max_tag ff) in
-
- (* Construction of cnf *)
- let pre_ff = pre_process mt (ff:'a formula) in
- let (cnf_ff,cnf_ff_tags) = cnf pre_ff in
-
- match witness_list_tags prover cnf_ff with
- | Model m -> Model m
- | Unknown -> Unknown
- | Prf res -> (*Printf.printf "\nList %i" (List.length `res); *)
- let deps = List.fold_left
- (fun s (cl,(prf,p)) ->
- let tags = ISet.fold (fun i s ->
- let t = fst (snd (List.nth cl i)) in
- if debug then (Printf.fprintf stdout "T : %i -> %a" i Tag.pp t) ;
- (*try*) TagSet.add t s (* with Invalid_argument _ -> s*)) (p.hyps prf) TagSet.empty in
- TagSet.union s tags) (List.fold_left (fun s (i,_) -> TagSet.add i s) TagSet.empty cnf_ff_tags) (List.combine cnf_ff res) in
-
- let ff' = abstract_formula deps ff in
-
- let pre_ff' = pre_process mt ff' in
-
- let (cnf_ff',_) = cnf pre_ff' in
-
- if debug then
- begin
+let micromega_tauto pre_process cnf spec prover env
+ (polys1 : (Names.Id.t * 'cst formula) list) (polys2 : 'cst formula) gl =
+ (* Express the goal as one big implication *)
+ let ff, ids = formula_hyps_concl polys1 polys2 in
+ let mt = CamlToCoq.positive (max_tag ff) in
+ (* Construction of cnf *)
+ let pre_ff = pre_process mt (ff : 'a formula) in
+ let cnf_ff, cnf_ff_tags = cnf pre_ff in
+ match witness_list_tags prover cnf_ff with
+ | Model m -> Model m
+ | Unknown -> Unknown
+ | Prf res ->
+ (*Printf.printf "\nList %i" (List.length `res); *)
+ let deps =
+ List.fold_left
+ (fun s (cl, (prf, p)) ->
+ let tags =
+ ISet.fold
+ (fun i s ->
+ let t = fst (snd (List.nth cl i)) in
+ if debug then Printf.fprintf stdout "T : %i -> %a" i Tag.pp t;
+ (*try*) TagSet.add t s
+ (* with Invalid_argument _ -> s*))
+ (p.hyps prf) TagSet.empty
+ in
+ TagSet.union s tags)
+ (List.fold_left
+ (fun s (i, _) -> TagSet.add i s)
+ TagSet.empty cnf_ff_tags)
+ (List.combine cnf_ff res)
+ in
+ let ff' = abstract_formula deps ff in
+ let pre_ff' = pre_process mt ff' in
+ let cnf_ff', _ = cnf pre_ff' in
+ if debug then begin
output_string stdout "\n";
- Printf.printf "TForm : %a\n" pp_formula ff ; flush stdout;
- Printf.printf "CNF : %a\n" pp_cnf_tag cnf_ff ; flush stdout;
- Printf.printf "TFormAbs : %a\n" pp_formula ff' ; flush stdout;
- Printf.printf "TFormPre : %a\n" pp_formula pre_ff ; flush stdout;
- Printf.printf "TFormPreAbs : %a\n" pp_formula pre_ff' ; flush stdout;
- Printf.printf "CNF : %a\n" pp_cnf_tag cnf_ff' ; flush stdout;
+ Printf.printf "TForm : %a\n" pp_formula ff;
+ flush stdout;
+ Printf.printf "CNF : %a\n" pp_cnf_tag cnf_ff;
+ flush stdout;
+ Printf.printf "TFormAbs : %a\n" pp_formula ff';
+ flush stdout;
+ Printf.printf "TFormPre : %a\n" pp_formula pre_ff;
+ flush stdout;
+ Printf.printf "TFormPreAbs : %a\n" pp_formula pre_ff';
+ flush stdout;
+ Printf.printf "CNF : %a\n" pp_cnf_tag cnf_ff';
+ flush stdout
end;
-
- (* Even if it does not work, this does not mean it is not provable
+ (* Even if it does not work, this does not mean it is not provable
-- the prover is REALLY incomplete *)
- (* if debug then
+ (* if debug then
begin
(* recompute the proofs *)
match witness_list_tags prover cnf_ff' with
| None -> failwith "abstraction is wrong"
| Some res -> ()
end ; *)
+ let res' = compact_proofs cnf_ff res cnf_ff' in
+ let ff', res', ids = (ff', res', Mc.ids_of_formula ff') in
+ let res' = dump_list spec.proof_typ spec.dump_proof res' in
+ Prf (ids, ff', res')
- let res' = compact_proofs cnf_ff res cnf_ff' in
-
- let (ff',res',ids) = (ff',res', Mc.ids_of_formula ff') in
-
- let res' = dump_list (spec.proof_typ) spec.dump_proof res' in
- Prf (ids,ff',res')
-
-let micromega_tauto pre_process cnf spec prover env (polys1: (Names.Id.t * 'cst formula) list) (polys2: 'cst formula) gl =
+let micromega_tauto pre_process cnf spec prover env
+ (polys1 : (Names.Id.t * 'cst formula) list) (polys2 : 'cst formula) gl =
try micromega_tauto pre_process cnf spec prover env polys1 polys2 gl
with Not_found ->
- begin
- Printexc.print_backtrace stdout; flush stdout;
- Unknown
- end
-
+ Printexc.print_backtrace stdout;
+ flush stdout;
+ Unknown
(**
* Parse the proof environment, and call micromega_tauto
@@ -1818,194 +1845,234 @@ let fresh_id avoid id gl =
Tactics.fresh_id_in_env avoid id (Proofview.Goal.env gl)
let clear_all_no_check =
- Proofview.Goal.enter begin fun gl ->
- let concl = Tacmach.New.pf_concl gl in
- let env = Environ.reset_with_named_context Environ.empty_named_context_val (Tacmach.New.pf_env gl) in
- (Refine.refine ~typecheck:false begin fun sigma ->
- Evarutil.new_evar env sigma ~principal:true concl
- end)
- end
-
-
-
-let micromega_gen
- parse_arith
- pre_process
- cnf
- spec dumpexpr prover tac =
- Proofview.Goal.enter begin fun gl ->
- let sigma = Tacmach.New.project gl in
- let concl = Tacmach.New.pf_concl gl in
- let hyps = Tacmach.New.pf_hyps_types gl in
- try
- let gl0 = { env = Tacmach.New.pf_env gl; sigma } in
- let (hyps,concl,env) = parse_goal gl0 parse_arith (Env.empty gl0) hyps concl in
- let env = Env.elements env in
- let spec = Lazy.force spec in
- let dumpexpr = Lazy.force dumpexpr in
-
-
- if debug then Feedback.msg_debug (Pp.str "Env " ++ (Env.pp gl0 env)) ;
-
- match micromega_tauto pre_process cnf spec prover env hyps concl gl0 with
- | Unknown -> flush stdout ; Tacticals.New.tclFAIL 0 (Pp.str " Cannot find witness")
- | Model(m,e) -> Tacticals.New.tclFAIL 0 (Pp.str " Cannot find witness")
- | Prf (ids,ff',res') ->
- let (arith_goal,props,vars,ff_arith) = make_goal_of_formula gl0 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 (CAst.make @@ IntroNaming (Namegen.IntroIdentifier id)) in*)
- let goal_name = fresh_id Id.Set.empty (Names.Id.of_string "__arith") gl in
- let env' = List.map (fun (id,i) -> EConstr.mkVar id,i) vars in
-
- let tac_arith = Tacticals.New.tclTHENLIST [ clear_all_no_check ;intro_props ; intro_vars ;
- micromega_order_change spec res'
- (EConstr.mkApp(Lazy.force coq_list, [|spec.proof_typ|])) env' ff_arith ] in
-
- let goal_props = List.rev (Env.elements (prop_env_of_formula gl0 ff')) in
-
- let goal_vars = List.map (fun (_,i) -> List.nth env (i-1)) vars in
-
- let arith_args = goal_props @ goal_vars in
+ Proofview.Goal.enter (fun gl ->
+ let concl = Tacmach.New.pf_concl gl in
+ let env =
+ Environ.reset_with_named_context Environ.empty_named_context_val
+ (Tacmach.New.pf_env gl)
+ in
+ Refine.refine ~typecheck:false (fun sigma ->
+ Evarutil.new_evar env sigma ~principal:true concl))
- let kill_arith = Tacticals.New.tclTHEN tac_arith tac in
-(*
+let micromega_gen parse_arith pre_process cnf spec dumpexpr prover tac =
+ Proofview.Goal.enter (fun gl ->
+ let sigma = Tacmach.New.project gl in
+ let concl = Tacmach.New.pf_concl gl in
+ let hyps = Tacmach.New.pf_hyps_types gl in
+ try
+ let gl0 = {env = Tacmach.New.pf_env gl; sigma} in
+ let hyps, concl, env =
+ parse_goal gl0 parse_arith (Env.empty gl0) hyps concl
+ in
+ let env = Env.elements env in
+ let spec = Lazy.force spec in
+ let dumpexpr = Lazy.force dumpexpr in
+ if debug then Feedback.msg_debug (Pp.str "Env " ++ Env.pp gl0 env);
+ match
+ micromega_tauto pre_process cnf spec prover env hyps concl gl0
+ with
+ | Unknown ->
+ flush stdout;
+ Tacticals.New.tclFAIL 0 (Pp.str " Cannot find witness")
+ | Model (m, e) ->
+ Tacticals.New.tclFAIL 0 (Pp.str " Cannot find witness")
+ | Prf (ids, ff', res') ->
+ let arith_goal, props, vars, ff_arith =
+ make_goal_of_formula gl0 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 (CAst.make @@ IntroNaming (Namegen.IntroIdentifier id)) in*)
+ let goal_name =
+ fresh_id Id.Set.empty (Names.Id.of_string "__arith") gl
+ in
+ let env' = List.map (fun (id, i) -> (EConstr.mkVar id, i)) vars in
+ let tac_arith =
+ Tacticals.New.tclTHENLIST
+ [ clear_all_no_check
+ ; intro_props
+ ; intro_vars
+ ; micromega_order_change spec res'
+ (EConstr.mkApp (Lazy.force coq_list, [|spec.proof_typ|]))
+ env' ff_arith ]
+ in
+ let goal_props =
+ List.rev (Env.elements (prop_env_of_formula gl0 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 tac_arith tac in
+ (*
(*tclABSTRACT fails in certain corner cases.*)
Tacticals.New.tclTHEN
clear_all_no_check
(Abstract.tclABSTRACT ~opaque:false None (Tacticals.New.tclTHEN tac_arith tac)) in *)
-
- Tacticals.New.tclTHEN
- (Tactics.assert_by (Names.Name goal_name) arith_goal
- ((*Proofview.tclTIME (Some "kill_arith")*) kill_arith))
- ((*Proofview.tclTIME (Some "apply_arith") *)
- (Tactics.exact_check (EConstr.applist (EConstr.mkVar goal_name, arith_args@(List.map EConstr.mkVar ids)))))
- with
- | 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"))
- | x -> begin if debug then Tacticals.New.tclFAIL 0 (Pp.str (Printexc.get_backtrace ()))
- else raise x
- end
- end
-
-let micromega_order_changer cert env ff =
+ Tacticals.New.tclTHEN
+ (Tactics.assert_by (Names.Name goal_name) arith_goal
+ (*Proofview.tclTIME (Some "kill_arith")*) kill_arith)
+ ((*Proofview.tclTIME (Some "apply_arith") *)
+ Tactics.exact_check
+ (EConstr.applist
+ ( EConstr.mkVar goal_name
+ , arith_args @ List.map EConstr.mkVar ids )))
+ with
+ | 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" ))
+ | x ->
+ if debug then
+ Tacticals.New.tclFAIL 0 (Pp.str (Printexc.get_backtrace ()))
+ else raise x)
+
+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 = (EConstr.mkApp(Lazy.force coq_list, [|Lazy.force coq_QWitness |])) in
-
- let formula_typ = (EConstr.mkApp (Lazy.force coq_Cstr,[| coeff|])) in
+ let typ = Lazy.force coq_R in
+ let cert_typ =
+ EConstr.mkApp (Lazy.force coq_list, [|Lazy.force coq_QWitness|])
+ in
+ let formula_typ = EConstr.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.enter begin fun gl ->
- Tacticals.New.tclTHENLIST
- [
- (Tactics.change_concl
- (set
- [
- ("__ff", ff, EConstr.mkApp(Lazy.force coq_Formula, [|formula_typ |]));
- ("__varmap", vm, EConstr.mkApp
- (gen_constant_in_modules "VarMap"
- [["Coq" ; "micromega" ; "VarMap"] ; ["VarMap"]] "t", [|typ|]));
- ("__wit", cert, cert_typ)
- ]
- (Tacmach.New.pf_concl gl)));
- (* Tacticals.New.tclTHENLIST (List.map (fun id -> (Tactics.introduction id)) ids)*)
- ]
- end
+ let vm = dump_varmap typ (vm_of_list env) in
+ Proofview.Goal.enter (fun gl ->
+ Tacticals.New.tclTHENLIST
+ [ Tactics.change_concl
+ (set
+ [ ( "__ff"
+ , ff
+ , EConstr.mkApp (Lazy.force coq_Formula, [|formula_typ|]) )
+ ; ( "__varmap"
+ , vm
+ , EConstr.mkApp
+ ( gen_constant_in_modules "VarMap"
+ [["Coq"; "micromega"; "VarMap"]; ["VarMap"]]
+ "t"
+ , [|typ|] ) )
+ ; ("__wit", cert, cert_typ) ]
+ (Tacmach.New.pf_concl gl))
+ (* Tacticals.New.tclTHENLIST (List.map (fun id -> (Tactics.introduction id)) ids)*)
+ ])
let micromega_genr prover tac =
let parse_arith = parse_rarith in
- let spec = lazy {
- typ = Lazy.force coq_R;
- coeff = Lazy.force coq_Rcst;
- dump_coeff = dump_q;
- proof_typ = Lazy.force coq_QWitness ;
- dump_proof = dump_psatz coq_Q dump_q
- } in
- Proofview.Goal.enter begin fun gl ->
- let sigma = Tacmach.New.project gl in
- let concl = Tacmach.New.pf_concl gl in
- let hyps = Tacmach.New.pf_hyps_types gl in
-
- try
- let gl0 = { env = Tacmach.New.pf_env gl; sigma } in
- let (hyps,concl,env) = parse_goal gl0 parse_arith (Env.empty gl0) hyps concl in
- let env = Env.elements env in
- let spec = Lazy.force spec in
-
- let hyps' = List.map (fun (n,f) -> (n, Mc.map_bformula (Micromega.map_Formula Micromega.q_of_Rcst) f)) hyps in
- let concl' = Mc.map_bformula (Micromega.map_Formula Micromega.q_of_Rcst) concl in
-
- match micromega_tauto (fun _ x -> x) Mc.cnfQ spec prover env hyps' concl' gl0 with
- | Unknown | Model _ -> flush stdout ; Tacticals.New.tclFAIL 0 (Pp.str " Cannot find witness")
- | Prf (ids,ff',res') ->
- 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 gl0 (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 (CAst.make @@ IntroNaming (Namegen.IntroIdentifier id)) in
- let goal_name = fresh_id Id.Set.empty (Names.Id.of_string "__arith") gl in
- let env' = List.map (fun (id,i) -> EConstr.mkVar id,i) vars in
-
- let tac_arith = Tacticals.New.tclTHENLIST [ clear_all_no_check ; intro_props ; intro_vars ;
- micromega_order_changer res' env' ff_arith ] in
-
- let goal_props = List.rev (Env.elements (prop_env_of_formula gl0 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 tac_arith tac in
- (* Tacticals.New.tclTHEN
+ let spec =
+ lazy
+ { typ = Lazy.force coq_R
+ ; coeff = Lazy.force coq_Rcst
+ ; dump_coeff = dump_q
+ ; proof_typ = Lazy.force coq_QWitness
+ ; dump_proof = dump_psatz coq_Q dump_q }
+ in
+ Proofview.Goal.enter (fun gl ->
+ let sigma = Tacmach.New.project gl in
+ let concl = Tacmach.New.pf_concl gl in
+ let hyps = Tacmach.New.pf_hyps_types gl in
+ try
+ let gl0 = {env = Tacmach.New.pf_env gl; sigma} in
+ let hyps, concl, env =
+ parse_goal gl0 parse_arith (Env.empty gl0) hyps concl
+ in
+ let env = Env.elements env in
+ let spec = Lazy.force spec in
+ let hyps' =
+ List.map
+ (fun (n, f) ->
+ (n, Mc.map_bformula (Micromega.map_Formula Micromega.q_of_Rcst) f))
+ hyps
+ in
+ let concl' =
+ Mc.map_bformula (Micromega.map_Formula Micromega.q_of_Rcst) concl
+ in
+ match
+ micromega_tauto
+ (fun _ x -> x)
+ Mc.cnfQ spec prover env hyps' concl' gl0
+ with
+ | Unknown | Model _ ->
+ flush stdout;
+ Tacticals.New.tclFAIL 0 (Pp.str " Cannot find witness")
+ | Prf (ids, ff', res') ->
+ 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 gl0 (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 (CAst.make @@ IntroNaming (Namegen.IntroIdentifier id))
+ in
+ let goal_name =
+ fresh_id Id.Set.empty (Names.Id.of_string "__arith") gl
+ in
+ let env' = List.map (fun (id, i) -> (EConstr.mkVar id, i)) vars in
+ let tac_arith =
+ Tacticals.New.tclTHENLIST
+ [ clear_all_no_check
+ ; intro_props
+ ; intro_vars
+ ; micromega_order_changer res' env' ff_arith ]
+ in
+ let goal_props =
+ List.rev (Env.elements (prop_env_of_formula gl0 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 tac_arith tac in
+ (* Tacticals.New.tclTHEN
(Tactics.keep [])
(Tactics.tclABSTRACT None*)
-
- Tacticals.New.tclTHENS
- (Tactics.forward true (Some None) (ipat_of_name goal_name) arith_goal)
- [
- kill_arith;
- (Tacticals.New.tclTHENLIST
- [(Tactics.generalize (List.map EConstr.mkVar ids));
- (Tactics.exact_check (EConstr.applist (EConstr.mkVar goal_name, arith_args)))
- ] )
- ]
-
- with
- | 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 lift_ratproof prover l =
- match prover l with
+ Tacticals.New.tclTHENS
+ (Tactics.forward true (Some None) (ipat_of_name goal_name)
+ arith_goal)
+ [ kill_arith
+ ; Tacticals.New.tclTHENLIST
+ [ Tactics.generalize (List.map EConstr.mkVar ids)
+ ; Tactics.exact_check
+ (EConstr.applist (EConstr.mkVar goal_name, arith_args)) ] ]
+ with
+ | 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" )))
+
+let lift_ratproof prover l =
+ match prover l with
| Unknown | Model _ -> Unknown
- | Prf c -> Prf (Mc.RatProof( c,Mc.DoneProof))
+ | Prf c -> Prf (Mc.RatProof (c, Mc.DoneProof))
type micromega_polys = (Micromega.q Mc.pol * Mc.op1) list
[@@@ocaml.warning "-37"]
+
type csdp_certificate = S of Sos_types.positivstellensatz option | F of string
+
(* Used to read the result of the execution of csdpcert *)
type provername = string * int option
@@ -2016,47 +2083,47 @@ type provername = string * int option
open Persistent_cache
+module MakeCache (T : sig
+ type prover_option
+ type coeff
-module MakeCache(T : sig type prover_option
- type coeff
- val hash_prover_option : int -> prover_option -> int
- val hash_coeff : int -> coeff -> int
- val eq_prover_option : prover_option -> prover_option -> bool
- val eq_coeff : coeff -> coeff -> bool
-
- end) =
- struct
- module E =
- struct
- type t = T.prover_option * (T.coeff Mc.pol * Mc.op1) list
-
- let equal = Hash.(eq_pair T.eq_prover_option (CList.equal (eq_pair (eq_pol T.eq_coeff) Hash.eq_op1)))
+ val hash_prover_option : int -> prover_option -> int
+ val hash_coeff : int -> coeff -> int
+ val eq_prover_option : prover_option -> prover_option -> bool
+ val eq_coeff : coeff -> coeff -> bool
+end) =
+struct
+ module E = struct
+ type t = T.prover_option * (T.coeff Mc.pol * Mc.op1) list
- let hash =
- let hash_cstr = Hash.(hash_pair (hash_pol T.hash_coeff) hash_op1) in
- Hash.( (hash_pair T.hash_prover_option (List.fold_left hash_cstr)) 0)
- end
+ let equal =
+ Hash.(
+ eq_pair T.eq_prover_option
+ (CList.equal (eq_pair (eq_pol T.eq_coeff) Hash.eq_op1)))
- include PHashtable(E)
+ let hash =
+ let hash_cstr = Hash.(hash_pair (hash_pol T.hash_coeff) hash_op1) in
+ Hash.((hash_pair T.hash_prover_option (List.fold_left hash_cstr)) 0)
+ end
- let memo_opt use_cache cache_file f =
- let memof = memo cache_file f in
- fun x -> if !use_cache then memof x else f x
+ include PHashtable (E)
- end
+ let memo_opt use_cache cache_file f =
+ let memof = memo cache_file f in
+ fun x -> if !use_cache then memof x else f x
+end
+module CacheCsdp = MakeCache (struct
+ type prover_option = provername
+ type coeff = Mc.q
+ let hash_prover_option =
+ Hash.(hash_pair hash_string (hash_elt (Option.hash (fun x -> x))))
-module CacheCsdp = MakeCache(struct
- type prover_option = provername
- type coeff = Mc.q
- let hash_prover_option = Hash.(hash_pair hash_string
- (hash_elt (Option.hash (fun x -> x))))
- let eq_prover_option = Hash.(eq_pair String.equal
- (Option.equal Int.equal))
- let hash_coeff = Hash.hash_q
- let eq_coeff = Hash.eq_q
- end)
+ let eq_prover_option = Hash.(eq_pair String.equal (Option.equal Int.equal))
+ let hash_coeff = Hash.hash_q
+ let eq_coeff = Hash.eq_q
+end)
(**
* Build the command to call csdpcert, and launch it. This in turn will call
@@ -2065,233 +2132,235 @@ module CacheCsdp = MakeCache(struct
*)
let require_csdp =
- if System.is_in_system_path "csdp"
- then lazy ()
- else lazy (raise CsdpNotFound)
-
-let really_call_csdpcert : provername -> micromega_polys -> Sos_types.positivstellensatz option =
- fun provername poly ->
+ if System.is_in_system_path "csdp" then lazy () else lazy (raise CsdpNotFound)
+let really_call_csdpcert :
+ provername -> micromega_polys -> Sos_types.positivstellensatz option =
+ fun provername poly ->
Lazy.force require_csdp;
-
let cmdname =
List.fold_left Filename.concat (Envars.coqlib ())
- ["plugins"; "micromega"; "csdpcert" ^ Coq_config.exec_extension] in
-
- match ((command cmdname [|cmdname|] (provername,poly)) : csdp_certificate) with
- | F str ->
- if debug then Printf.fprintf stdout "really_call_csdpcert : %s\n" str;
- raise (failwith str)
- | S res -> res
+ ["plugins"; "micromega"; "csdpcert" ^ Coq_config.exec_extension]
+ in
+ match (command cmdname [|cmdname|] (provername, poly) : csdp_certificate) with
+ | F str ->
+ if debug then Printf.fprintf stdout "really_call_csdpcert : %s\n" str;
+ raise (failwith str)
+ | S res -> res
(**
* Check the cache before calling the prover.
*)
let xcall_csdpcert =
- CacheCsdp.memo_opt use_csdp_cache ".csdp.cache" (fun (prover,pb) -> really_call_csdpcert prover pb)
+ CacheCsdp.memo_opt use_csdp_cache ".csdp.cache" (fun (prover, pb) ->
+ really_call_csdpcert prover pb)
(**
* Prover callback functions.
*)
-let call_csdpcert prover pb = xcall_csdpcert (prover,pb)
+let call_csdpcert prover pb = xcall_csdpcert (prover, pb)
let rec z_to_q_pol e =
- match e with
- | Mc.Pc z -> Mc.Pc {Mc.qnum = z ; Mc.qden = Mc.XH}
- | Mc.Pinj(p,pol) -> Mc.Pinj(p,z_to_q_pol pol)
- | Mc.PX(pol1,p,pol2) -> Mc.PX(z_to_q_pol pol1, p, z_to_q_pol pol2)
+ match e with
+ | Mc.Pc z -> Mc.Pc {Mc.qnum = z; Mc.qden = Mc.XH}
+ | Mc.Pinj (p, pol) -> Mc.Pinj (p, z_to_q_pol pol)
+ | Mc.PX (pol1, p, pol2) -> Mc.PX (z_to_q_pol pol1, p, z_to_q_pol pol2)
let call_csdpcert_q provername poly =
- match call_csdpcert provername poly with
+ match call_csdpcert provername poly with
| None -> Unknown
| Some cert ->
- let cert = Certificate.q_cert_of_pos cert in
- if Mc.qWeakChecker poly cert
- then Prf cert
- else ((print_string "buggy certificate") ;Unknown)
+ let cert = Certificate.q_cert_of_pos cert in
+ if Mc.qWeakChecker poly cert then Prf cert
+ else (
+ print_string "buggy certificate";
+ Unknown )
let call_csdpcert_z provername poly =
- let l = List.map (fun (e,o) -> (z_to_q_pol e,o)) poly in
+ let l = List.map (fun (e, o) -> (z_to_q_pol e, o)) poly in
match call_csdpcert provername l with
- | None -> Unknown
- | Some cert ->
- let cert = Certificate.z_cert_of_pos cert in
- if Mc.zWeakChecker poly cert
- then Prf cert
- else ((print_string "buggy certificate" ; flush stdout) ;Unknown)
+ | None -> Unknown
+ | Some cert ->
+ let cert = Certificate.z_cert_of_pos cert in
+ if Mc.zWeakChecker poly cert then Prf cert
+ else (
+ print_string "buggy certificate";
+ flush stdout;
+ Unknown )
let xhyps_of_cone base acc prf =
let rec xtract e acc =
match e with
| Mc.PsatzC _ | Mc.PsatzZ | Mc.PsatzSquare _ -> acc
- | Mc.PsatzIn n -> let n = (CoqToCaml.nat n) in
- if n >= base
- then ISet.add (n-base) acc
- else acc
- | Mc.PsatzMulC(_,c) -> xtract c acc
- | Mc.PsatzAdd(e1,e2) | Mc.PsatzMulE(e1,e2) -> xtract e1 (xtract e2 acc) in
-
- xtract prf acc
+ | Mc.PsatzIn n ->
+ let n = CoqToCaml.nat n in
+ if n >= base then ISet.add (n - base) acc else acc
+ | Mc.PsatzMulC (_, c) -> xtract c acc
+ | Mc.PsatzAdd (e1, e2) | Mc.PsatzMulE (e1, e2) -> xtract e1 (xtract e2 acc)
+ in
+ xtract prf acc
let hyps_of_cone prf = xhyps_of_cone 0 ISet.empty prf
-let compact_cone prf f =
+let compact_cone prf f =
let np n = CamlToCoq.nat (f (CoqToCaml.nat n)) in
-
let rec xinterp prf =
match prf with
| Mc.PsatzC _ | Mc.PsatzZ | Mc.PsatzSquare _ -> prf
| Mc.PsatzIn n -> Mc.PsatzIn (np n)
- | Mc.PsatzMulC(e,c) -> Mc.PsatzMulC(e,xinterp c)
- | Mc.PsatzAdd(e1,e2) -> Mc.PsatzAdd(xinterp e1,xinterp e2)
- | Mc.PsatzMulE(e1,e2) -> Mc.PsatzMulE(xinterp e1,xinterp e2) in
-
- xinterp prf
+ | Mc.PsatzMulC (e, c) -> Mc.PsatzMulC (e, xinterp c)
+ | Mc.PsatzAdd (e1, e2) -> Mc.PsatzAdd (xinterp e1, xinterp e2)
+ | Mc.PsatzMulE (e1, e2) -> Mc.PsatzMulE (xinterp e1, xinterp e2)
+ in
+ xinterp prf
let hyps_of_pt pt =
-
let rec xhyps base pt acc =
match pt with
- | Mc.DoneProof -> acc
- | Mc.RatProof(c,pt) -> xhyps (base+1) pt (xhyps_of_cone base acc c)
- | Mc.CutProof(c,pt) -> xhyps (base+1) pt (xhyps_of_cone base acc c)
- | Mc.EnumProof(c1,c2,l) ->
- let s = xhyps_of_cone base (xhyps_of_cone base acc c2) c1 in
- List.fold_left (fun s x -> xhyps (base + 1) x s) s l in
-
- xhyps 0 pt ISet.empty
+ | Mc.DoneProof -> acc
+ | Mc.RatProof (c, pt) -> xhyps (base + 1) pt (xhyps_of_cone base acc c)
+ | Mc.CutProof (c, pt) -> xhyps (base + 1) pt (xhyps_of_cone base acc c)
+ | Mc.EnumProof (c1, c2, l) ->
+ let s = xhyps_of_cone base (xhyps_of_cone base acc c2) c1 in
+ List.fold_left (fun s x -> xhyps (base + 1) x s) s l
+ in
+ xhyps 0 pt ISet.empty
let compact_pt pt f =
- let translate ofset x =
- if x < ofset then x
- else (f (x-ofset) + ofset) in
-
+ let translate ofset x = if x < ofset then x else f (x - ofset) + ofset in
let rec compact_pt ofset pt =
match pt with
- | Mc.DoneProof -> Mc.DoneProof
- | Mc.RatProof(c,pt) -> Mc.RatProof(compact_cone c (translate (ofset)), compact_pt (ofset+1) pt )
- | Mc.CutProof(c,pt) -> Mc.CutProof(compact_cone c (translate (ofset)), compact_pt (ofset+1) pt )
- | Mc.EnumProof(c1,c2,l) -> Mc.EnumProof(compact_cone c1 (translate (ofset)), compact_cone c2 (translate (ofset)),
- Mc.map (fun x -> compact_pt (ofset+1) x) l) in
- compact_pt 0 pt
+ | Mc.DoneProof -> Mc.DoneProof
+ | Mc.RatProof (c, pt) ->
+ Mc.RatProof (compact_cone c (translate ofset), compact_pt (ofset + 1) pt)
+ | Mc.CutProof (c, pt) ->
+ Mc.CutProof (compact_cone c (translate ofset), compact_pt (ofset + 1) pt)
+ | Mc.EnumProof (c1, c2, l) ->
+ Mc.EnumProof
+ ( compact_cone c1 (translate ofset)
+ , compact_cone c2 (translate ofset)
+ , Mc.map (fun x -> compact_pt (ofset + 1) x) l )
+ in
+ compact_pt 0 pt
(**
* Definition of provers.
* Instantiates the type ('a,'prf) prover defined above.
*)
-let lift_pexpr_prover p l = p (List.map (fun (e,o) -> Mc.denorm e , o) l)
-
-
-module CacheZ = MakeCache(struct
- type prover_option = bool * bool* int
- type coeff = Mc.z
- let hash_prover_option : int -> prover_option -> int = Hash.hash_elt Hashtbl.hash
- let eq_prover_option : prover_option -> prover_option -> bool = (=)
- let eq_coeff = Hash.eq_z
- let hash_coeff = Hash.hash_z
- end)
-
-module CacheQ = MakeCache(struct
- type prover_option = int
- type coeff = Mc.q
- let hash_prover_option : int -> int -> int = Hash.hash_elt Hashtbl.hash
- let eq_prover_option = Int.equal
- let eq_coeff = Hash.eq_q
- let hash_coeff = Hash.hash_q
- end)
-
-let memo_lia = CacheZ.memo_opt use_lia_cache ".lia.cache"
- (fun ((_,ce,b),s) -> lift_pexpr_prover (Certificate.lia ce b) s)
-let memo_nlia = CacheZ.memo_opt use_nia_cache ".nia.cache"
- (fun ((_,ce,b),s) -> lift_pexpr_prover (Certificate.nlia ce b) s)
-let memo_nra = CacheQ.memo_opt use_nra_cache ".nra.cache"
- (fun (o,s) -> lift_pexpr_prover (Certificate.nlinear_prover o) s)
-
-
-
-let linear_prover_Q = {
- name = "linear prover";
- get_option = get_lra_option ;
- prover = (fun (o,l) -> lift_pexpr_prover (Certificate.linear_prover_with_cert o ) 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";
- get_option = get_lra_option ;
- prover = (fun (o,l) -> lift_pexpr_prover (Certificate.linear_prover_with_cert o ) 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";
- 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 ;
- pp_f = fun o x -> pp_pol pp_q o (fst x)
-}
-
-let non_linear_prover_R str o = {
- name = "real nonlinear prover";
- 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;
- pp_f = fun o x -> pp_pol pp_q o (fst x)
-}
-
-let non_linear_prover_Z str o = {
- name = "real nonlinear prover";
- 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)
-}
-
-let linear_Z = {
- name = "lia";
- get_option = get_lia_option;
- prover = memo_lia ;
- hyps = hyps_of_pt;
- compact = compact_pt;
- pp_prf = pp_proof_term;
- pp_f = fun o x -> pp_pol pp_z o (fst x)
-}
-
-let nlinear_Z = {
- name = "nlia";
- get_option = get_lia_option;
- prover = memo_nlia ;
- hyps = hyps_of_pt;
- compact = compact_pt;
- pp_prf = pp_proof_term;
- pp_f = fun o x -> pp_pol pp_z o (fst x)
-}
+let lift_pexpr_prover p l = p (List.map (fun (e, o) -> (Mc.denorm e, o)) l)
+
+module CacheZ = MakeCache (struct
+ type prover_option = bool * bool * int
+ type coeff = Mc.z
+
+ let hash_prover_option : int -> prover_option -> int =
+ Hash.hash_elt Hashtbl.hash
+
+ let eq_prover_option : prover_option -> prover_option -> bool = ( = )
+ let eq_coeff = Hash.eq_z
+ let hash_coeff = Hash.hash_z
+end)
+
+module CacheQ = MakeCache (struct
+ type prover_option = int
+ type coeff = Mc.q
+
+ let hash_prover_option : int -> int -> int = Hash.hash_elt Hashtbl.hash
+ let eq_prover_option = Int.equal
+ let eq_coeff = Hash.eq_q
+ let hash_coeff = Hash.hash_q
+end)
+
+let memo_lia =
+ CacheZ.memo_opt use_lia_cache ".lia.cache" (fun ((_, ce, b), s) ->
+ lift_pexpr_prover (Certificate.lia ce b) s)
+
+let memo_nlia =
+ CacheZ.memo_opt use_nia_cache ".nia.cache" (fun ((_, ce, b), s) ->
+ lift_pexpr_prover (Certificate.nlia ce b) s)
+
+let memo_nra =
+ CacheQ.memo_opt use_nra_cache ".nra.cache" (fun (o, s) ->
+ lift_pexpr_prover (Certificate.nlinear_prover o) s)
+
+let linear_prover_Q =
+ { name = "linear prover"
+ ; get_option = get_lra_option
+ ; prover =
+ (fun (o, l) ->
+ lift_pexpr_prover (Certificate.linear_prover_with_cert o) 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"
+ ; get_option = get_lra_option
+ ; prover =
+ (fun (o, l) ->
+ lift_pexpr_prover (Certificate.linear_prover_with_cert o) 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"
+ ; 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
+ ; pp_f = (fun o x -> pp_pol pp_q o (fst x)) }
+
+let non_linear_prover_R str o =
+ { name = "real nonlinear prover"
+ ; 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
+ ; pp_f = (fun o x -> pp_pol pp_q o (fst x)) }
+
+let non_linear_prover_Z str o =
+ { name = "real nonlinear prover"
+ ; 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)) }
+
+let linear_Z =
+ { name = "lia"
+ ; get_option = get_lia_option
+ ; prover = memo_lia
+ ; hyps = hyps_of_pt
+ ; compact = compact_pt
+ ; pp_prf = pp_proof_term
+ ; pp_f = (fun o x -> pp_pol pp_z o (fst x)) }
+
+let nlinear_Z =
+ { name = "nlia"
+ ; get_option = get_lia_option
+ ; prover = memo_nlia
+ ; hyps = hyps_of_pt
+ ; compact = compact_pt
+ ; pp_prf = pp_proof_term
+ ; pp_f = (fun o x -> pp_pol pp_z o (fst x)) }
(**
* Functions instantiating micromega_gen with the appropriate theories and
@@ -2299,67 +2368,70 @@ let nlinear_Z = {
*)
let exfalso_if_concl_not_Prop =
- Proofview.Goal.enter begin fun gl ->
- Tacmach.New.(
- if is_prop (pf_env gl) (project gl) (pf_concl gl)
- then Tacticals.New.tclIDTAC
- else Tactics.elim_type (Lazy.force coq_False)
- )
- end
+ Proofview.Goal.enter (fun gl ->
+ Tacmach.New.(
+ if is_prop (pf_env gl) (project gl) (pf_concl gl) then
+ Tacticals.New.tclIDTAC
+ else Tactics.elim_type (Lazy.force coq_False)))
let micromega_gen parse_arith pre_process cnf spec dumpexpr prover tac =
- Tacticals.New.tclTHEN exfalso_if_concl_not_Prop (micromega_gen parse_arith pre_process cnf spec dumpexpr prover tac)
+ Tacticals.New.tclTHEN exfalso_if_concl_not_Prop
+ (micromega_gen parse_arith pre_process cnf spec dumpexpr prover tac)
let micromega_genr prover tac =
Tacticals.New.tclTHEN exfalso_if_concl_not_Prop (micromega_genr prover tac)
let lra_Q =
- micromega_gen parse_qarith (fun _ x -> x) Mc.cnfQ qq_domain_spec dump_qexpr
- linear_prover_Q
+ micromega_gen parse_qarith
+ (fun _ x -> x)
+ Mc.cnfQ qq_domain_spec dump_qexpr linear_prover_Q
-let psatz_Q i =
- micromega_gen parse_qarith (fun _ x -> x) Mc.cnfQ qq_domain_spec dump_qexpr
- (non_linear_prover_Q "real_nonlinear_prover" (Some i) )
+let psatz_Q i =
+ micromega_gen parse_qarith
+ (fun _ x -> x)
+ Mc.cnfQ qq_domain_spec dump_qexpr
+ (non_linear_prover_Q "real_nonlinear_prover" (Some i))
-let lra_R =
- micromega_genr linear_prover_R
+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 =
+ micromega_genr (non_linear_prover_R "real_nonlinear_prover" (Some i))
+let psatz_Z i =
+ micromega_gen parse_zarith
+ (fun _ x -> x)
+ Mc.cnfZ zz_domain_spec dump_zexpr
+ (non_linear_prover_Z "real_nonlinear_prover" (Some i))
-let psatz_Z i =
- micromega_gen parse_zarith (fun _ x -> x) Mc.cnfZ zz_domain_spec dump_zexpr
- (non_linear_prover_Z "real_nonlinear_prover" (Some i) )
+let sos_Z =
+ micromega_gen parse_zarith
+ (fun _ x -> x)
+ Mc.cnfZ zz_domain_spec dump_zexpr
+ (non_linear_prover_Z "pure_sos" None)
-let sos_Z =
- micromega_gen parse_zarith (fun _ x -> x) Mc.cnfZ zz_domain_spec dump_zexpr
- (non_linear_prover_Z "pure_sos" None)
-
-let sos_Q =
- micromega_gen parse_qarith (fun _ x -> x) Mc.cnfQ qq_domain_spec dump_qexpr
- (non_linear_prover_Q "pure_sos" None)
-
-
-let sos_R =
- micromega_genr (non_linear_prover_R "pure_sos" None)
+let sos_Q =
+ micromega_gen parse_qarith
+ (fun _ x -> x)
+ Mc.cnfQ qq_domain_spec dump_qexpr
+ (non_linear_prover_Q "pure_sos" None)
+let sos_R = micromega_genr (non_linear_prover_R "pure_sos" None)
let xlia =
- micromega_gen parse_zarith pre_processZ Mc.cnfZ zz_domain_spec dump_zexpr
- linear_Z
-
+ micromega_gen parse_zarith pre_processZ Mc.cnfZ zz_domain_spec dump_zexpr
+ linear_Z
-let xnlia =
- micromega_gen parse_zarith (fun _ x -> x) Mc.cnfZ zz_domain_spec dump_zexpr
- nlinear_Z
+let xnlia =
+ micromega_gen parse_zarith
+ (fun _ x -> x)
+ Mc.cnfZ zz_domain_spec dump_zexpr nlinear_Z
-let nra =
- micromega_genr nlinear_prover_R
+let nra = micromega_genr nlinear_prover_R
-let nqa =
- micromega_gen parse_qarith (fun _ x -> x) Mc.cnfQ qq_domain_spec dump_qexpr
- nlinear_prover_R
+let nqa =
+ micromega_gen parse_qarith
+ (fun _ x -> x)
+ Mc.cnfQ qq_domain_spec dump_qexpr nlinear_prover_R
(* Local Variables: *)
(* coding: utf-8 *)
diff --git a/plugins/micromega/coq_micromega.mli b/plugins/micromega/coq_micromega.mli
index 844ff5b1a6..37ea560241 100644
--- a/plugins/micromega/coq_micromega.mli
+++ b/plugins/micromega/coq_micromega.mli
@@ -22,8 +22,7 @@ val sos_R : unit Proofview.tactic -> unit Proofview.tactic
val lra_Q : unit Proofview.tactic -> unit Proofview.tactic
val lra_R : unit Proofview.tactic -> unit Proofview.tactic
-
(** {5 Use Micromega independently from tactics. } *)
-(** [dump_proof_term] generates the Coq representation of a Micromega proof witness *)
val dump_proof_term : Micromega.zArithProof -> EConstr.t
+(** [dump_proof_term] generates the Coq representation of a Micromega proof witness *)
diff --git a/plugins/micromega/csdpcert.ml b/plugins/micromega/csdpcert.ml
index 09e354957a..90dd81adf4 100644
--- a/plugins/micromega/csdpcert.ml
+++ b/plugins/micromega/csdpcert.ml
@@ -18,7 +18,6 @@ open Num
open Sos
open Sos_types
open Sos_lib
-
module Mc = Micromega
module C2Ml = Mutils.CoqToCaml
@@ -26,157 +25,179 @@ type micromega_polys = (Micromega.q Mc.pol * Mc.op1) list
type csdp_certificate = S of Sos_types.positivstellensatz option | F of string
type provername = string * int option
-
-let flags = [Open_append;Open_binary;Open_creat]
-
+let flags = [Open_append; Open_binary; Open_creat]
let chan = open_out_gen flags 0o666 "trace"
+module M = struct
+ open Mc
-module M =
-struct
- open Mc
-
- let rec expr_to_term = function
- | PEc z -> Const (C2Ml.q_to_num z)
- | PEX v -> Var ("x"^(string_of_int (C2Ml.index v)))
- | PEmul(p1,p2) ->
+ let rec expr_to_term = function
+ | PEc z -> Const (C2Ml.q_to_num z)
+ | PEX v -> Var ("x" ^ string_of_int (C2Ml.index v))
+ | PEmul (p1, p2) ->
let p1 = expr_to_term p1 in
let p2 = expr_to_term p2 in
- let res = Mul(p1,p2) in res
-
- | PEadd(p1,p2) -> Add(expr_to_term p1, expr_to_term p2)
- | PEsub(p1,p2) -> Sub(expr_to_term p1, expr_to_term p2)
- | PEpow(p,n) -> Pow(expr_to_term p , C2Ml.n n)
- | PEopp p -> Opp (expr_to_term p)
-
-
+ let res = Mul (p1, p2) in
+ res
+ | PEadd (p1, p2) -> Add (expr_to_term p1, expr_to_term p2)
+ | PEsub (p1, p2) -> Sub (expr_to_term p1, expr_to_term p2)
+ | PEpow (p, n) -> Pow (expr_to_term p, C2Ml.n n)
+ | PEopp p -> Opp (expr_to_term p)
end
+
open M
let partition_expr l =
- let rec f i = function
- | [] -> ([],[],[])
- | (e,k)::l ->
- let (eq,ge,neq) = f (i+1) l in
+ let rec f i = function
+ | [] -> ([], [], [])
+ | (e, k) :: l -> (
+ let eq, ge, neq = f (i + 1) l in
match k with
- | Mc.Equal -> ((e,i)::eq,ge,neq)
- | Mc.NonStrict -> (eq,(e,Axiom_le i)::ge,neq)
- | Mc.Strict -> (* e > 0 == e >= 0 /\ e <> 0 *)
- (eq, (e,Axiom_lt i)::ge,(e,Axiom_lt i)::neq)
- | Mc.NonEqual -> (eq,ge,(e,Axiom_eq i)::neq)
- (* Not quite sure -- Coq interface has changed *)
- in f 0 l
-
+ | Mc.Equal -> ((e, i) :: eq, ge, neq)
+ | Mc.NonStrict -> (eq, (e, Axiom_le i) :: ge, neq)
+ | Mc.Strict ->
+ (* e > 0 == e >= 0 /\ e <> 0 *)
+ (eq, (e, Axiom_lt i) :: ge, (e, Axiom_lt i) :: neq)
+ | Mc.NonEqual -> (eq, ge, (e, Axiom_eq i) :: neq) )
+ (* Not quite sure -- Coq interface has changed *)
+ in
+ f 0 l
let rec sets_of_list l =
- match l with
+ match l with
| [] -> [[]]
- | e::l -> let s = sets_of_list l in
- s@(List.map (fun s0 -> e::s0) s)
+ | e :: l ->
+ let s = sets_of_list l in
+ s @ List.map (fun s0 -> e :: s0) s
(* The exploration is probably not complete - for simple cases, it works... *)
let real_nonlinear_prover d l =
- let l = List.map (fun (e,op) -> (Mc.denorm e,op)) l in
- try
- let (eq,ge,neq) = partition_expr l in
-
- let rec elim_const = function
- [] -> []
- | (x,y)::l -> let p = poly_of_term (expr_to_term x) in
- if poly_isconst p
- then elim_const l
- else (p,y)::(elim_const l) in
-
- let eq = elim_const eq in
- let peq = List.map fst eq in
-
- let pge = List.map
- (fun (e,psatz) -> poly_of_term (expr_to_term e),psatz) ge in
-
- let monoids = List.map (fun m -> (List.fold_right (fun (p,kd) y ->
- let p = poly_of_term (expr_to_term p) in
- match kd with
- | Axiom_lt i -> poly_mul p y
- | Axiom_eq i -> poly_mul (poly_pow p 2) y
- | _ -> failwith "monoids") m (poly_const (Int 1)) , List.map snd m))
- (sets_of_list neq) in
-
- let (cert_ideal, cert_cone,monoid) = deepen_until d (fun d ->
- tryfind (fun m -> let (ci,cc) =
- real_positivnullstellensatz_general false d peq pge (poly_neg (fst m) ) in
- (ci,cc,snd m)) monoids) 0 in
-
- let proofs_ideal = List.map2 (fun q i -> Eqmul(term_of_poly q,Axiom_eq i))
- cert_ideal (List.map snd eq) in
-
- let proofs_cone = List.map term_of_sos cert_cone in
-
- let proof_ne =
- let (neq , lt) = List.partition
- (function Axiom_eq _ -> true | _ -> false ) monoid in
- let sq = match
- (List.map (function Axiom_eq i -> i | _ -> failwith "error") neq)
- with
- | [] -> Rational_lt (Int 1)
- | l -> Monoid l in
- List.fold_right (fun x y -> Product(x,y)) lt sq in
-
- let proof = end_itlist
- (fun s t -> Sum(s,t)) (proof_ne :: proofs_ideal @ proofs_cone) in
- S (Some proof)
- with
+ let l = List.map (fun (e, op) -> (Mc.denorm e, op)) l in
+ try
+ let eq, ge, neq = partition_expr l in
+ let rec elim_const = function
+ | [] -> []
+ | (x, y) :: l ->
+ let p = poly_of_term (expr_to_term x) in
+ if poly_isconst p then elim_const l else (p, y) :: elim_const l
+ in
+ let eq = elim_const eq in
+ let peq = List.map fst eq in
+ let pge =
+ List.map (fun (e, psatz) -> (poly_of_term (expr_to_term e), psatz)) ge
+ in
+ let monoids =
+ List.map
+ (fun m ->
+ ( List.fold_right
+ (fun (p, kd) y ->
+ let p = poly_of_term (expr_to_term p) in
+ match kd with
+ | Axiom_lt i -> poly_mul p y
+ | Axiom_eq i -> poly_mul (poly_pow p 2) y
+ | _ -> failwith "monoids")
+ m (poly_const (Int 1))
+ , List.map snd m ))
+ (sets_of_list neq)
+ in
+ let cert_ideal, cert_cone, monoid =
+ deepen_until d
+ (fun d ->
+ tryfind
+ (fun m ->
+ let ci, cc =
+ real_positivnullstellensatz_general false d peq pge
+ (poly_neg (fst m))
+ in
+ (ci, cc, snd m))
+ monoids)
+ 0
+ in
+ let proofs_ideal =
+ List.map2
+ (fun q i -> Eqmul (term_of_poly q, Axiom_eq i))
+ cert_ideal (List.map snd eq)
+ in
+ let proofs_cone = List.map term_of_sos cert_cone in
+ let proof_ne =
+ let neq, lt =
+ List.partition (function Axiom_eq _ -> true | _ -> false) monoid
+ in
+ let sq =
+ match
+ List.map (function Axiom_eq i -> i | _ -> failwith "error") neq
+ with
+ | [] -> Rational_lt (Int 1)
+ | l -> Monoid l
+ in
+ List.fold_right (fun x y -> Product (x, y)) lt sq
+ in
+ let proof =
+ end_itlist
+ (fun s t -> Sum (s, t))
+ ((proof_ne :: proofs_ideal) @ proofs_cone)
+ in
+ S (Some proof)
+ with
| Sos_lib.TooDeep -> S None
| any -> F (Printexc.to_string any)
(* This is somewhat buggy, over Z, strict inequality vanish... *)
-let pure_sos l =
- let l = List.map (fun (e,o) -> Mc.denorm e, o) l in
-
- (* If there is no strict inequality,
+let pure_sos l =
+ let l = List.map (fun (e, o) -> (Mc.denorm e, o)) l in
+ (* If there is no strict inequality,
I should nonetheless be able to try something - over Z > is equivalent to -1 >= *)
- try
- let l = List.combine l (CList.interval 0 (List.length l -1)) in
- let (lt,i) = try (List.find (fun (x,_) -> (=) (snd x) Mc.Strict) l)
- with Not_found -> List.hd l in
- let plt = poly_neg (poly_of_term (expr_to_term (fst lt))) in
- let (n,polys) = sumofsquares plt in (* n * (ci * pi^2) *)
- let pos = Product (Rational_lt n,
- List.fold_right (fun (c,p) rst -> Sum (Product (Rational_lt c, Square
- (term_of_poly p)), rst))
- polys (Rational_lt (Int 0))) in
- let proof = Sum(Axiom_lt i, pos) in
-(* let s,proof' = scale_certificate proof in
+ try
+ let l = List.combine l (CList.interval 0 (List.length l - 1)) in
+ let lt, i =
+ try List.find (fun (x, _) -> snd x = Mc.Strict) l
+ with Not_found -> List.hd l
+ in
+ let plt = poly_neg (poly_of_term (expr_to_term (fst lt))) in
+ let n, polys = sumofsquares plt in
+ (* n * (ci * pi^2) *)
+ let pos =
+ Product
+ ( Rational_lt n
+ , List.fold_right
+ (fun (c, p) rst ->
+ Sum (Product (Rational_lt c, Square (term_of_poly p)), rst))
+ polys (Rational_lt (Int 0)) )
+ in
+ let proof = Sum (Axiom_lt i, pos) in
+ (* let s,proof' = scale_certificate proof in
let cert = snd (cert_of_pos proof') in *)
S (Some proof)
- with
-(* | Sos.CsdpNotFound -> F "Sos.CsdpNotFound" *)
- | any -> (* May be that could be refined *) S None
-
-
+ with (* | Sos.CsdpNotFound -> F "Sos.CsdpNotFound" *)
+ | any ->
+ (* May be that could be refined *) S None
let run_prover prover pb =
- match prover with
- | "real_nonlinear_prover", Some d -> real_nonlinear_prover d pb
- | "pure_sos", None -> pure_sos pb
- | prover, _ -> (Printf.printf "unknown prover: %s\n" prover; exit 1)
+ match prover with
+ | "real_nonlinear_prover", Some d -> real_nonlinear_prover d pb
+ | "pure_sos", None -> pure_sos pb
+ | prover, _ ->
+ Printf.printf "unknown prover: %s\n" prover;
+ exit 1
let main () =
try
- let (prover,poly) = (input_value stdin : provername * micromega_polys) in
+ let (prover, poly) = (input_value stdin : provername * micromega_polys) in
let cert = run_prover prover poly in
-(* Printf.fprintf chan "%a -> %a" print_list_term poly output_csdp_certificate cert ;
+ (* Printf.fprintf chan "%a -> %a" print_list_term poly output_csdp_certificate cert ;
close_out chan ; *)
-
- output_value stdout (cert:csdp_certificate);
- flush stdout ;
- Marshal.to_channel chan (cert:csdp_certificate) [] ;
- flush chan ;
- exit 0
- with any -> (Printf.fprintf chan "error %s" (Printexc.to_string any) ; exit 1)
+ output_value stdout (cert : csdp_certificate);
+ flush stdout;
+ Marshal.to_channel chan (cert : csdp_certificate) [];
+ flush chan;
+ exit 0
+ with any ->
+ Printf.fprintf chan "error %s" (Printexc.to_string any);
+ exit 1
;;
-
-let _ = main () in ()
+let _ = main () in
+()
(* Local Variables: *)
(* coding: utf-8 *)
diff --git a/plugins/micromega/itv.ml b/plugins/micromega/itv.ml
index 533b060dd3..214edb46ba 100644
--- a/plugins/micromega/itv.ml
+++ b/plugins/micromega/itv.ml
@@ -12,9 +12,9 @@
open Num
- (** The type of intervals is *)
- type interval = num option * num option
- (** None models the absence of bound i.e. infinity
+(** The type of intervals is *)
+type interval = num option * num option
+(** None models the absence of bound i.e. infinity
As a result,
- None , None -> \]-oo,+oo\[
- None , Some v -> \]-oo,v\]
@@ -23,59 +23,51 @@ open Num
Intervals needs to be explicitly normalised.
*)
- let pp o (n1,n2) =
- (match n1 with
- | None -> output_string o "]-oo"
- | Some n -> Printf.fprintf o "[%s" (string_of_num n)
- );
- output_string o ",";
- (match n2 with
- | None -> output_string o "+oo["
- | Some n -> Printf.fprintf o "%s]" (string_of_num n)
- )
+let pp o (n1, n2) =
+ ( match n1 with
+ | None -> output_string o "]-oo"
+ | Some n -> Printf.fprintf o "[%s" (string_of_num n) );
+ output_string o ",";
+ match n2 with
+ | None -> output_string o "+oo["
+ | Some n -> Printf.fprintf o "%s]" (string_of_num n)
-
-
- (** if then interval [itv] is empty, [norm_itv itv] returns [None]
+(** if then interval [itv] is empty, [norm_itv itv] returns [None]
otherwise, it returns [Some itv] *)
- let norm_itv itv =
- match itv with
- | Some a , Some b -> if a <=/ b then Some itv else None
- | _ -> Some itv
+let norm_itv itv =
+ match itv with
+ | Some a, Some b -> if a <=/ b then Some itv else None
+ | _ -> Some itv
(** [inter i1 i2 = None] if the intersection of intervals is empty
[inter i1 i2 = Some i] if [i] is the intersection of the intervals [i1] and [i2] *)
- let inter i1 i2 =
- let (l1,r1) = i1
- and (l2,r2) = i2 in
-
- let inter f o1 o2 =
- match o1 , o2 with
- | None , None -> None
- | Some _ , None -> o1
- | None , Some _ -> o2
- | Some n1 , Some n2 -> Some (f n1 n2) in
-
- norm_itv (inter max_num l1 l2 , inter min_num r1 r2)
-
- let range = function
- | None,_ | _,None -> None
- | Some i,Some j -> Some (floor_num j -/ceiling_num i +/ (Int 1))
-
-
- let smaller_itv i1 i2 =
- match range i1 , range i2 with
- | None , _ -> false
- | _ , None -> true
- | Some i , Some j -> i <=/ j
-
+let inter i1 i2 =
+ let l1, r1 = i1 and l2, r2 = i2 in
+ let inter f o1 o2 =
+ match (o1, o2) with
+ | None, None -> None
+ | Some _, None -> o1
+ | None, Some _ -> o2
+ | Some n1, Some n2 -> Some (f n1 n2)
+ in
+ norm_itv (inter max_num l1 l2, inter min_num r1 r2)
+
+let range = function
+ | None, _ | _, None -> None
+ | Some i, Some j -> Some (floor_num j -/ ceiling_num i +/ Int 1)
+
+let smaller_itv i1 i2 =
+ match (range i1, range i2) with
+ | None, _ -> false
+ | _, None -> true
+ | Some i, Some j -> i <=/ j
(** [in_bound bnd v] checks whether [v] is within the bounds [bnd] *)
let in_bound bnd v =
- let (l,r) = bnd in
- match l , r with
- | None , None -> true
- | None , Some a -> v <=/ a
- | Some a , None -> a <=/ v
- | Some a , Some b -> a <=/ v && v <=/ b
+ let l, r = bnd in
+ match (l, r) with
+ | None, None -> true
+ | None, Some a -> v <=/ a
+ | Some a, None -> a <=/ v
+ | Some a, Some b -> a <=/ v && v <=/ b
diff --git a/plugins/micromega/itv.mli b/plugins/micromega/itv.mli
index 7b7edc64ea..c7164f2c98 100644
--- a/plugins/micromega/itv.mli
+++ b/plugins/micromega/itv.mli
@@ -10,6 +10,7 @@
open Num
type interval = num option * num option
+
val pp : out_channel -> interval -> unit
val inter : interval -> interval -> interval option
val range : interval -> num option
diff --git a/plugins/micromega/mfourier.ml b/plugins/micromega/mfourier.ml
index 75cdfa24f1..da75137185 100644
--- a/plugins/micromega/mfourier.ml
+++ b/plugins/micromega/mfourier.ml
@@ -14,37 +14,25 @@ open Polynomial
open Vect
let debug = false
-
let compare_float (p : float) q = pervasives_compare p q
-(** Implementation of intervals *)
open Itv
+(** Implementation of intervals *)
+
type vector = Vect.t
(** 'cstr' is the type of constraints.
{coeffs = v ; bound = (l,r) } models the constraints l <= v <= r
**)
-module ISet = Set.Make(Int)
+module ISet = Set.Make (Int)
+module System = Hashtbl.Make (Vect)
-module System = Hashtbl.Make(Vect)
+type proof = Assum of int | Elim of var * proof * proof | And of proof * proof
-type proof =
-| Assum of int
-| Elim of var * proof * proof
-| And of proof * proof
-
-type system = {
- sys : cstr_info ref System.t ;
- vars : ISet.t
-}
-and cstr_info = {
- bound : interval ;
- prf : proof ;
- pos : int ;
- neg : int ;
-}
+type system = {sys : cstr_info ref System.t; vars : ISet.t}
+and cstr_info = {bound : interval; prf : proof; pos : int; neg : int}
(** 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
@@ -58,31 +46,29 @@ and cstr_info = {
[v] is an upper-bound of the set of variables which appear in [s].
*)
-(** To be thrown when a system has no solution *)
exception SystemContradiction of proof
+(** To be thrown when a system has no solution *)
(** Pretty printing *)
- let rec pp_proof o prf =
- match prf with
- | Assum i -> Printf.fprintf o "H%i" i
- | Elim(v, prf1,prf2) -> Printf.fprintf o "E(%i,%a,%a)" v pp_proof prf1 pp_proof prf2
- | And(prf1,prf2) -> Printf.fprintf o "A(%a,%a)" pp_proof prf1 pp_proof prf2
-
-let pp_cstr o (vect,bnd) =
- let (l,r) = bnd in
- (match l with
- | None -> ()
- | Some n -> Printf.fprintf o "%s <= " (string_of_num n))
- ;
- Vect.pp o vect ;
- (match r with
- | None -> output_string o"\n"
- | Some n -> Printf.fprintf o "<=%s\n" (string_of_num n))
-
-
-let pp_system o sys=
- System.iter (fun vect ibnd ->
- pp_cstr o (vect,(!ibnd).bound)) sys
+let rec pp_proof o prf =
+ match prf with
+ | Assum i -> Printf.fprintf o "H%i" i
+ | Elim (v, prf1, prf2) ->
+ Printf.fprintf o "E(%i,%a,%a)" v pp_proof prf1 pp_proof prf2
+ | And (prf1, prf2) -> Printf.fprintf o "A(%a,%a)" pp_proof prf1 pp_proof prf2
+
+let pp_cstr o (vect, bnd) =
+ let l, r = bnd in
+ ( match l with
+ | None -> ()
+ | Some n -> Printf.fprintf o "%s <= " (string_of_num n) );
+ Vect.pp o vect;
+ match r with
+ | None -> output_string o "\n"
+ | Some n -> Printf.fprintf o "<=%s\n" (string_of_num n)
+
+let pp_system o sys =
+ System.iter (fun vect ibnd -> pp_cstr o (vect, !ibnd.bound)) sys
(** [merge_cstr_info] takes:
- the intersection of bounds and
@@ -90,13 +76,12 @@ let pp_system o sys=
- [pos] and [neg] fields should be identical *)
let merge_cstr_info i1 i2 =
- let { pos = p1 ; neg = n1 ; bound = i1 ; prf = prf1 } = i1
- and { pos = p2 ; neg = n2 ; bound = i2 ; prf = prf2 } = i2 in
- assert (Int.equal p1 p2 && Int.equal n1 n2) ;
- match inter i1 i2 with
- | None -> None (* Could directly raise a system contradiction exception *)
- | Some bnd ->
- Some { pos = p1 ; neg = n1 ; bound = bnd ; prf = And(prf1,prf2) }
+ let {pos = p1; neg = n1; bound = i1; prf = prf1} = i1
+ and {pos = p2; neg = n2; bound = i2; prf = prf2} = i2 in
+ assert (Int.equal p1 p2 && Int.equal n1 n2);
+ match inter i1 i2 with
+ | None -> None (* Could directly raise a system contradiction exception *)
+ | Some bnd -> Some {pos = p1; neg = n1; bound = bnd; prf = And (prf1, prf2)}
(** [xadd_cstr vect cstr_info] loads an constraint into the system.
The constraint is neither redundant nor contradictory.
@@ -104,90 +89,96 @@ let merge_cstr_info i1 i2 =
*)
let xadd_cstr vect cstr_info sys =
- 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)))
- | Some info' -> info := info'
- with
- | Not_found -> System.replace sys vect (ref cstr_info)
+ match merge_cstr_info cstr_info !info with
+ | None -> raise (SystemContradiction (And (cstr_info.prf, !info.prf)))
+ | Some info' -> info := info'
+ 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
+ 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.
+ | Contradiction
+ (** The constraint is contradictory.
Typically, a [SystemContradiction] exception will be raised. *)
- | Redundant (** The constrain is redundant.
+ | Redundant
+ (** The constrain is redundant.
Typically, the constraint will be dropped *)
- | Cstr of vector * cstr_info (** Taken alone, the constraint is neither contradictory nor redundant.
+ | Cstr of vector * cstr_info
+ (** Taken alone, the constraint is neither contradictory nor redundant.
Typically, it will be added to the constraint system. *)
(** [normalise_cstr] : vector -> cstr_info -> cstr_ext *)
let normalise_cstr vect cinfo =
match norm_itv cinfo.bound with
- | None -> Contradiction
- | Some (l,r) ->
- match Vect.choose vect with
- | None -> if Itv.in_bound (l,r) (Int 0) then Redundant else Contradiction
- | Some (_,n,_) -> Cstr(Vect.div n vect,
- let divn x = x // n in
- if Int.equal (sign_num n) 1
- then{cinfo with bound = (Option.map divn l , Option.map divn r) }
- else {cinfo with pos = cinfo.neg ; neg = cinfo.pos ; bound = (Option.map divn r , Option.map divn l)})
-
+ | None -> Contradiction
+ | Some (l, r) -> (
+ match Vect.choose vect with
+ | None -> if Itv.in_bound (l, r) (Int 0) then Redundant else Contradiction
+ | Some (_, n, _) ->
+ Cstr
+ ( Vect.div n vect
+ , let divn x = x // n in
+ if Int.equal (sign_num n) 1 then
+ {cinfo with bound = (Option.map divn l, Option.map divn r)}
+ else
+ { cinfo with
+ pos = cinfo.neg
+ ; neg = cinfo.pos
+ ; bound = (Option.map divn r, Option.map divn l) } ) )
(** For compatibility, there is an external representation of constraints *)
-
let count v =
- Vect.fold (fun (n,p) _ vl ->
+ Vect.fold
+ (fun (n, p) _ vl ->
let sg = sign_num vl in
- assert (sg <> 0) ;
- if Int.equal sg 1 then (n,p+1)else (n+1, p)) (0,0) v
-
-
-let norm_cstr {coeffs = v ; op = o ; cst = c} idx =
- let (n,p) = count v in
-
- normalise_cstr v {pos = p ; neg = n ; bound =
- (match o with
- | Eq -> Some c , Some c
- | Ge -> Some c , None
- | Gt -> raise Polynomial.Strict
- ) ;
- prf = Assum idx }
-
+ assert (sg <> 0);
+ if Int.equal sg 1 then (n, p + 1) else (n + 1, p))
+ (0, 0) v
+
+let norm_cstr {coeffs = v; op = o; cst = c} idx =
+ let n, p = count v in
+ normalise_cstr v
+ { pos = p
+ ; neg = n
+ ; bound =
+ ( match o with
+ | Eq -> (Some c, Some c)
+ | Ge -> (Some c, None)
+ | Gt -> raise Polynomial.Strict )
+ ; prf = Assum idx }
(** [load_system l] takes a list of constraints of type [cstr_compat]
@return a system of constraints
@raise SystemContradiction if a contradiction is found
*)
let load_system l =
-
let sys = System.create 1000 in
-
- let li = List.mapi (fun i e -> (e,i)) l in
-
- let vars = List.fold_left (fun vrs (cstr,i) ->
- match norm_cstr cstr i with
- | Contradiction -> raise (SystemContradiction (Assum i))
- | Redundant -> vrs
- | Cstr(vect,info) ->
- xadd_cstr vect info sys ;
- Vect.fold (fun s v _ -> ISet.add v s) vrs cstr.coeffs) ISet.empty li in
-
- {sys = sys ;vars = vars}
+ let li = List.mapi (fun i e -> (e, i)) l in
+ let vars =
+ List.fold_left
+ (fun vrs (cstr, i) ->
+ match norm_cstr cstr i with
+ | Contradiction -> raise (SystemContradiction (Assum i))
+ | Redundant -> vrs
+ | Cstr (vect, info) ->
+ xadd_cstr vect info sys;
+ Vect.fold (fun s v _ -> ISet.add v s) vrs cstr.coeffs)
+ ISet.empty li
+ in
+ {sys; vars}
let system_list sys =
- let { sys = s ; vars = v } = sys in
- System.fold (fun k bi l -> (k, !bi)::l) s []
-
+ let {sys = s; vars = v} = sys in
+ System.fold (fun k bi l -> (k, !bi) :: l) s []
(** [add (v1,c1) (v2,c2) ]
precondition: (c1 <>/ Int 0 && c2 <>/ Int 0)
@@ -196,15 +187,15 @@ let system_list sys =
Note that the resulting vector is not normalised.
*)
-let add (v1,c1) (v2,c2) =
- assert (c1 <>/ Int 0 && c2 <>/ Int 0) ;
- let res = mul_add (Int 1 // c1) v1 (Int 1 // c2) v2 in
- (res, count res)
+let add (v1, c1) (v2, c2) =
+ assert (c1 <>/ Int 0 && c2 <>/ Int 0);
+ let res = mul_add (Int 1 // c1) v1 (Int 1 // c2) v2 in
+ (res, count res)
-let add (v1,c1) (v2,c2) =
- let res = add (v1,c1) (v2,c2) in
- (* Printf.printf "add(%a,%s,%a,%s) -> %a\n" pp_vect v1 (string_of_num c1) pp_vect v2 (string_of_num c2) pp_vect (fst res) ;*)
- res
+let add (v1, c1) (v2, c2) =
+ let res = add (v1, c1) (v2, c2) in
+ (* Printf.printf "add(%a,%s,%a,%s) -> %a\n" pp_vect v1 (string_of_num c1) pp_vect v2 (string_of_num c2) pp_vect (fst res) ;*)
+ res
(** To perform Fourier elimination, constraints are categorised depending on the sign of the variable to eliminate. *)
@@ -215,54 +206,59 @@ let add (v1,c1) (v2,c2) =
@param m contains constraints which do not mention [x]
*)
-let split x (vect: vector) info (l,m,r) =
- match get x vect with
- | Int 0 -> (* The constraint does not mention [x], store it in m *)
- (l,(vect,info)::m,r)
- | vl -> (* otherwise *)
-
- let cons_bound lst bd =
- match bd with
- | None -> lst
- | Some bnd -> (vl,vect,{info with bound = Some bnd,None})::lst in
-
- let lb,rb = info.bound in
- if Int.equal (sign_num vl) 1
- then (cons_bound l lb,m,cons_bound r rb)
- else (* sign_num vl = -1 *)
- (cons_bound l rb,m,cons_bound r lb)
-
+let split x (vect : vector) info (l, m, r) =
+ match get x vect with
+ | Int 0 ->
+ (* The constraint does not mention [x], store it in m *)
+ (l, (vect, info) :: m, r)
+ | vl ->
+ (* otherwise *)
+ let cons_bound lst bd =
+ match bd with
+ | None -> lst
+ | Some bnd -> (vl, vect, {info with bound = (Some bnd, None)}) :: lst
+ in
+ let lb, rb = info.bound in
+ if Int.equal (sign_num vl) 1 then (cons_bound l lb, m, cons_bound r rb)
+ else (* sign_num vl = -1 *)
+ (cons_bound l rb, m, cons_bound r lb)
(** [project vr sys] projects system [sys] over the set of variables [ISet.remove vr sys.vars ].
This is a one step Fourier elimination.
*)
let project vr sys =
-
- let (l,m,r) = System.fold (fun vect rf l_m_r -> split vr vect !rf l_m_r) sys.sys ([],[],[]) in
-
+ let l, m, r =
+ System.fold
+ (fun vect rf l_m_r -> split vr vect !rf l_m_r)
+ sys.sys ([], [], [])
+ in
let new_sys = System.create (System.length sys.sys) in
-
- (* Constraints in [m] belong to the projection - for those [vr] is already projected out *)
- List.iter (fun (vect,info) -> System.replace new_sys vect (ref info) ) m ;
-
- let elim (v1,vect1,info1) (v2,vect2,info2) =
- let {neg = n1 ; pos = p1 ; bound = bound1 ; prf = prf1} = info1
- and {neg = n2 ; pos = p2 ; bound = bound2 ; prf = prf2} = info2 in
-
- let bnd1 = Option.get (fst bound1)
- and bnd2 = Option.get (fst bound2) in
- let bound = (bnd1 // v1) +/ (bnd2 // minus_num v2) in
- let vres,(n,p) = add (vect1,v1) (vect2,minus_num v2) in
- (vres,{neg = n ; pos = p ; bound = (Some bound, None); prf = Elim(vr,info1.prf,info2.prf)}) in
-
- List.iter(fun l_elem -> List.iter (fun r_elem ->
- let (vect,info) = elim l_elem r_elem in
+ (* Constraints in [m] belong to the projection - for those [vr] is already projected out *)
+ List.iter (fun (vect, info) -> System.replace new_sys vect (ref info)) m;
+ let elim (v1, vect1, info1) (v2, vect2, info2) =
+ let {neg = n1; pos = p1; bound = bound1; prf = prf1} = info1
+ and {neg = n2; pos = p2; bound = bound2; prf = prf2} = info2 in
+ let bnd1 = Option.get (fst bound1) and bnd2 = Option.get (fst bound2) in
+ let bound = (bnd1 // v1) +/ (bnd2 // minus_num v2) in
+ let vres, (n, p) = add (vect1, v1) (vect2, minus_num v2) in
+ ( vres
+ , { neg = n
+ ; pos = p
+ ; bound = (Some bound, None)
+ ; prf = Elim (vr, info1.prf, info2.prf) } )
+ in
+ List.iter
+ (fun l_elem ->
+ List.iter
+ (fun r_elem ->
+ let vect, info = elim l_elem r_elem in
match normalise_cstr vect info with
- | Redundant -> ()
- | Contradiction -> raise (SystemContradiction info.prf)
- | Cstr(vect,info) -> xadd_cstr vect info new_sys) r ) l;
- {sys = new_sys ; vars = ISet.remove vr sys.vars}
-
+ | Redundant -> ()
+ | Contradiction -> raise (SystemContradiction info.prf)
+ | Cstr (vect, info) -> xadd_cstr vect info new_sys)
+ r)
+ l;
+ {sys = new_sys; vars = ISet.remove vr sys.vars}
(** [project_using_eq] performs elimination by pivoting using an equation.
This is the counter_part of the [elim] sub-function of [!project].
@@ -273,103 +269,92 @@ let project vr sys =
@param prf is the proof of the equation
*)
-let project_using_eq vr c vect bound prf (vect',info') =
- match get vr vect' with
- | Int 0 -> (vect',info')
- | c2 ->
- let c1 = if c2 >=/ Int 0 then minus_num c else c in
-
- let c2 = abs_num c2 in
-
- let (vres,(n,p)) = add (vect,c1) (vect', c2) in
-
- let cst = bound // c1 in
-
- let bndres =
- let f x = cst +/ x // c2 in
- let (l,r) = info'.bound in
- (Option.map f l , Option.map f r) in
-
- (vres,{neg = n ; pos = p ; bound = bndres ; prf = Elim(vr,prf,info'.prf)})
-
+let project_using_eq vr c vect bound prf (vect', info') =
+ match get vr vect' with
+ | Int 0 -> (vect', info')
+ | c2 ->
+ let c1 = if c2 >=/ Int 0 then minus_num c else c in
+ let c2 = abs_num c2 in
+ let vres, (n, p) = add (vect, c1) (vect', c2) in
+ let cst = bound // c1 in
+ let bndres =
+ let f x = cst +/ (x // c2) in
+ let l, r = info'.bound in
+ (Option.map f l, Option.map f r)
+ in
+ (vres, {neg = n; pos = p; bound = bndres; prf = Elim (vr, prf, info'.prf)})
-let elim_var_using_eq vr vect cst prf sys =
+let elim_var_using_eq vr vect cst prf sys =
let c = get vr vect in
-
- let elim_var = project_using_eq vr c vect cst prf in
-
- let new_sys = System.create (System.length sys.sys) in
-
- System.iter(fun vect iref ->
- let (vect',info') = elim_var (vect,!iref) in
- match normalise_cstr vect' info' with
- | Redundant -> ()
- | Contradiction -> raise (SystemContradiction info'.prf)
- | Cstr(vect,info') -> xadd_cstr vect info' new_sys) sys.sys ;
-
- {sys = new_sys ; vars = ISet.remove vr sys.vars}
-
+ let elim_var = project_using_eq vr c vect cst prf in
+ let new_sys = System.create (System.length sys.sys) in
+ System.iter
+ (fun vect iref ->
+ let vect', info' = elim_var (vect, !iref) in
+ match normalise_cstr vect' info' with
+ | Redundant -> ()
+ | Contradiction -> raise (SystemContradiction info'.prf)
+ | Cstr (vect, info') -> xadd_cstr vect info' new_sys)
+ sys.sys;
+ {sys = new_sys; vars = ISet.remove vr sys.vars}
(** [size sys] computes the number of entries in the system of constraints *)
-let size sys = System.fold (fun v iref s -> s + (!iref).neg + (!iref).pos) sys 0
+let size sys = System.fold (fun v iref s -> s + !iref.neg + !iref.pos) sys 0
-module IMap = CMap.Make(Int)
+module IMap = CMap.Make (Int)
(** [eval_vect map vect] evaluates vector [vect] using the values of [map].
If [map] binds all the variables of [vect], we get
[eval_vect map [(x1,v1);...;(xn,vn)] = (IMap.find x1 map * v1) + ... + (IMap.find xn map) * vn , []]
The function returns as second argument, a sub-vector consisting in the variables that are not in [map]. *)
-let eval_vect map vect =
- Vect.fold (fun (sum,rst) v vl ->
+let eval_vect map vect =
+ Vect.fold
+ (fun (sum, rst) v vl ->
try
let val_v = IMap.find v map in
(sum +/ (val_v */ vl), rst)
- with
- Not_found -> (sum, Vect.set v vl rst)) (Int 0,Vect.null) vect
-
-
+ with Not_found -> (sum, Vect.set v vl rst))
+ (Int 0, Vect.null) vect
(** [restrict_bound n sum itv] returns the interval of [x]
given that (fst itv) <= x * n + sum <= (snd itv) *)
-let restrict_bound n sum (itv:interval) =
- let f x = (x -/ sum) // n in
- let l,r = itv in
- match sign_num n with
- | 0 -> if in_bound itv sum
- then (None,None) (* redundant *)
- else failwith "SystemContradiction"
- | 1 -> Option.map f l , Option.map f r
- | _ -> Option.map f r , Option.map f l
-
+let restrict_bound n sum (itv : interval) =
+ let f x = (x -/ sum) // n in
+ let l, r = itv in
+ match sign_num n with
+ | 0 ->
+ if in_bound itv sum then (None, None) (* redundant *)
+ else failwith "SystemContradiction"
+ | 1 -> (Option.map f l, Option.map f r)
+ | _ -> (Option.map f r, Option.map f l)
(** [bound_of_variable map v sys] computes the interval of [v] in
[sys] given a mapping [map] binding all the other variables *)
let bound_of_variable map v sys =
- System.fold (fun vect iref bnd ->
- let sum,rst = eval_vect map vect in
- let vl = Vect.get v rst in
- match inter bnd (restrict_bound vl sum (!iref).bound) with
+ System.fold
+ (fun vect iref bnd ->
+ let sum, rst = eval_vect map vect in
+ let vl = Vect.get v rst in
+ match inter bnd (restrict_bound vl sum !iref.bound) with
| None ->
- Printf.fprintf stdout "bound_of_variable: eval_vecr %a = %s,%a\n"
- Vect.pp vect (Num.string_of_num sum) Vect.pp rst ;
- Printf.fprintf stdout "current interval: %a\n" Itv.pp (!iref).bound;
- failwith "bound_of_variable: impossible"
- | Some itv -> itv) sys (None,None)
-
+ Printf.fprintf stdout "bound_of_variable: eval_vecr %a = %s,%a\n"
+ Vect.pp vect (Num.string_of_num sum) Vect.pp rst;
+ Printf.fprintf stdout "current interval: %a\n" Itv.pp !iref.bound;
+ failwith "bound_of_variable: impossible"
+ | Some itv -> itv)
+ sys (None, None)
(** [pick_small_value bnd] picks a value being closed to zero within the interval *)
let pick_small_value bnd =
match bnd with
- | None , None -> Int 0
- | None , Some i -> if (Int 0) <=/ (floor_num i) then Int 0 else floor_num i
- | Some i,None -> if i <=/ (Int 0) then Int 0 else ceiling_num i
- | Some i,Some j ->
- if i <=/ Int 0 && Int 0 <=/ j
- then Int 0
- else if ceiling_num i <=/ floor_num j
- then ceiling_num i (* why not *) else i
-
+ | None, None -> Int 0
+ | None, Some i -> if Int 0 <=/ floor_num i then Int 0 else floor_num i
+ | Some i, None -> if i <=/ Int 0 then Int 0 else ceiling_num i
+ | Some i, Some j ->
+ if i <=/ Int 0 && Int 0 <=/ j then Int 0
+ else if ceiling_num i <=/ floor_num j then ceiling_num i (* why not *)
+ else i
(** [solution s1 sys_l = Some(sn,\[(vn-1,sn-1);...; (v1,s1)\]\@sys_l)]
then [sn] is a system which contains only [black_v] -- if it existed in [s1]
@@ -378,262 +363,242 @@ let pick_small_value bnd =
*)
let solve_sys black_v choose_eq choose_variable sys sys_l =
-
let rec solve_sys sys sys_l =
- if debug then Printf.printf "S #%i size %i\n" (System.length sys.sys) (size sys.sys);
- if debug then Printf.printf "solve_sys :\n %a" pp_system sys.sys ;
-
+ if debug then
+ Printf.printf "S #%i size %i\n" (System.length sys.sys) (size sys.sys);
+ if debug then Printf.printf "solve_sys :\n %a" pp_system sys.sys;
let eqs = choose_eq sys in
+ try
+ let v, vect, cst, ln =
+ fst (List.find (fun ((v, _, _, _), _) -> v <> black_v) eqs)
+ in
+ if debug then (
+ Printf.printf "\nE %a = %s variable %i\n" Vect.pp vect
+ (string_of_num cst) v;
+ flush stdout );
+ let sys' = elim_var_using_eq v vect cst ln sys in
+ solve_sys sys' ((v, sys) :: sys_l)
+ with Not_found -> (
+ let vars = choose_variable sys in
try
- let (v,vect,cst,ln) = fst (List.find (fun ((v,_,_,_),_) -> v <> black_v) eqs) in
- if debug then
- (Printf.printf "\nE %a = %s variable %i\n" Vect.pp vect (string_of_num cst) v ;
- flush stdout);
- let sys' = elim_var_using_eq v vect cst ln sys in
- solve_sys sys' ((v,sys)::sys_l)
- with Not_found ->
- let vars = choose_variable sys in
- try
- let (v,est) = (List.find (fun (v,_) -> v <> black_v) vars) in
- if debug then (Printf.printf "\nV : %i estimate %f\n" v est ; flush stdout) ;
- let sys' = project v sys in
- solve_sys sys' ((v,sys)::sys_l)
- with Not_found -> (* we are done *) Inl (sys,sys_l) in
- solve_sys sys sys_l
-
-
-
-
-let solve black_v choose_eq choose_variable cstrs =
-
+ let v, est = List.find (fun (v, _) -> v <> black_v) vars in
+ if debug then (
+ Printf.printf "\nV : %i estimate %f\n" v est;
+ flush stdout );
+ let sys' = project v sys in
+ solve_sys sys' ((v, sys) :: sys_l)
+ with Not_found -> (* we are done *) Inl (sys, sys_l) )
+ in
+ solve_sys sys sys_l
+
+let solve black_v choose_eq choose_variable cstrs =
try
let sys = load_system cstrs in
- if debug then Printf.printf "solve :\n %a" pp_system sys.sys ;
- solve_sys black_v choose_eq choose_variable sys []
+ if debug then Printf.printf "solve :\n %a" pp_system sys.sys;
+ solve_sys black_v choose_eq choose_variable sys []
with SystemContradiction prf -> Inr prf
-
(** The purpose of module [EstimateElimVar] is to try to estimate the cost of eliminating a variable.
The output is an ordered list of (variable,cost).
*)
-module EstimateElimVar =
-struct
+module EstimateElimVar = struct
type sys_list = (vector * cstr_info) list
- let abstract_partition (v:int) (l: sys_list) =
-
- let rec xpart (l:sys_list) (ltl:sys_list) (n:int list) (z:int) (p:int list) =
+ let abstract_partition (v : int) (l : sys_list) =
+ let rec xpart (l : sys_list) (ltl : sys_list) (n : int list) (z : int)
+ (p : int list) =
match l with
- | [] -> (ltl, n,z,p)
- | (l1,info) ::rl ->
- match Vect.choose l1 with
- | None -> xpart rl ((Vect.null,info)::ltl) n (info.neg+info.pos+z) p
- | Some(vr, vl, rl1) ->
- if Int.equal v vr
- then
- let cons_bound lst bd =
- match bd with
- | None -> lst
- | Some bnd -> info.neg+info.pos::lst in
-
- let lb,rb = info.bound in
- if Int.equal (sign_num vl) 1
- then xpart rl ((rl1,info)::ltl) (cons_bound n lb) z (cons_bound p rb)
- else xpart rl ((rl1,info)::ltl) (cons_bound n rb) z (cons_bound p lb)
- else
- (* the variable is greater *)
- xpart rl ((l1,info)::ltl) n (info.neg+info.pos+z) p
-
+ | [] -> (ltl, n, z, p)
+ | (l1, info) :: rl -> (
+ match Vect.choose l1 with
+ | None ->
+ xpart rl ((Vect.null, info) :: ltl) n (info.neg + info.pos + z) p
+ | Some (vr, vl, rl1) ->
+ if Int.equal v vr then
+ let cons_bound lst bd =
+ match bd with
+ | None -> lst
+ | Some bnd -> (info.neg + info.pos) :: lst
+ in
+ let lb, rb = info.bound in
+ if Int.equal (sign_num vl) 1 then
+ xpart rl ((rl1, info) :: ltl) (cons_bound n lb) z
+ (cons_bound p rb)
+ else
+ xpart rl ((rl1, info) :: ltl) (cons_bound n rb) z
+ (cons_bound p lb)
+ else
+ (* the variable is greater *)
+ xpart rl ((l1, info) :: ltl) n (info.neg + info.pos + z) p )
in
- let (sys',n,z,p) = xpart l [] [] 0 [] in
-
+ let sys', n, z, p = xpart l [] [] 0 [] in
let ln = float_of_int (List.length n) in
- let sn = float_of_int (List.fold_left (+) 0 n) in
+ let sn = float_of_int (List.fold_left ( + ) 0 n) in
let lp = float_of_int (List.length p) in
- let sp = float_of_int (List.fold_left (+) 0 p) in
- (sys', float_of_int z +. lp *. sn +. ln *. sp -. lp*.ln)
-
-
- let choose_variable sys =
- let {sys = s ; vars = v} = sys in
+ let sp = float_of_int (List.fold_left ( + ) 0 p) in
+ (sys', float_of_int z +. (lp *. sn) +. (ln *. sp) -. (lp *. ln))
+ let choose_variable sys =
+ let {sys = s; vars = v} = sys in
let sl = system_list sys in
-
- let evals = fst
- (ISet.fold (fun v (eval,s) -> let ts,vl = abstract_partition v s in
- ((v,vl)::eval, ts)) v ([],sl)) in
-
- List.sort (fun x y -> compare_float (snd x) (snd y) ) evals
-
-
+ let evals =
+ fst
+ (ISet.fold
+ (fun v (eval, s) ->
+ let ts, vl = abstract_partition v s in
+ ((v, vl) :: eval, ts))
+ v ([], sl))
+ in
+ List.sort (fun x y -> compare_float (snd x) (snd y)) evals
end
+
open EstimateElimVar
(** The module [EstimateElimEq] is similar to [EstimateElimVar] but it orders equations.
*)
-module EstimateElimEq =
-struct
-
- let itv_point bnd =
- match bnd with
- |(Some a, Some b) -> a =/ b
- | _ -> false
+module EstimateElimEq = struct
+ let itv_point bnd = match bnd with Some a, Some b -> a =/ b | _ -> false
let rec unroll_until v l =
match Vect.choose l with
- | None -> (false,Vect.null)
- | Some(i,_,rl) -> if Int.equal i v
- then (true,rl)
- else if i < v then unroll_until v rl else (false,l)
-
-
+ | None -> (false, Vect.null)
+ | Some (i, _, rl) ->
+ if Int.equal i v then (true, rl)
+ else if i < v then unroll_until v rl
+ else (false, l)
let rec choose_simple_equation eqs =
match eqs with
- | [] -> None
- | (vect,a,prf,ln)::eqs ->
- match Vect.choose vect with
- | Some(i,v,rst) -> if Vect.is_null rst
- then Some (i,vect,a,prf,ln)
- else choose_simple_equation eqs
- | _ -> choose_simple_equation eqs
-
-
- let choose_primal_equation eqs (sys_l: (Vect.t *cstr_info) list) =
+ | [] -> None
+ | (vect, a, prf, ln) :: eqs -> (
+ match Vect.choose vect with
+ | Some (i, v, rst) ->
+ if Vect.is_null rst then Some (i, vect, a, prf, ln)
+ else choose_simple_equation eqs
+ | _ -> choose_simple_equation eqs )
+ let choose_primal_equation eqs (sys_l : (Vect.t * cstr_info) list) =
(* Counts the number of equations referring to variable [v] --
It looks like nb_cst is dead...
*)
let is_primal_equation_var v =
- List.fold_left (fun nb_eq (vect,info) ->
- if fst (unroll_until v vect)
- then if itv_point info.bound then nb_eq + 1 else nb_eq
- else nb_eq) 0 sys_l in
-
+ List.fold_left
+ (fun nb_eq (vect, info) ->
+ if fst (unroll_until v vect) then
+ if itv_point info.bound then nb_eq + 1 else nb_eq
+ else nb_eq)
+ 0 sys_l
+ in
let rec find_var vect =
match Vect.choose vect with
- | None -> None
- | Some(i,_,vect) ->
- let nb_eq = is_primal_equation_var i in
- if Int.equal nb_eq 2
- then Some i else find_var vect in
-
+ | None -> None
+ | Some (i, _, vect) ->
+ let nb_eq = is_primal_equation_var i in
+ if Int.equal nb_eq 2 then Some i else find_var vect
+ in
let rec find_eq_var eqs =
match eqs with
- | [] -> None
- | (vect,a,prf,ln)::l ->
- match find_var vect with
- | None -> find_eq_var l
- | Some r -> Some (r,vect,a,prf,ln)
+ | [] -> None
+ | (vect, a, prf, ln) :: l -> (
+ match find_var vect with
+ | None -> find_eq_var l
+ | Some r -> Some (r, vect, a, prf, ln) )
in
- match choose_simple_equation eqs with
- | None -> find_eq_var eqs
- | Some res -> Some res
-
-
-
- let choose_equality_var sys =
+ match choose_simple_equation eqs with
+ | None -> find_eq_var eqs
+ | Some res -> Some res
+ let choose_equality_var sys =
let sys_l = system_list sys in
-
- let equalities = List.fold_left
- (fun l (vect,info) ->
- match info.bound with
- | Some a , Some b ->
- if a =/ b then (* This an equation *)
- (vect,a,info.prf,info.neg+info.pos)::l else l
- | _ -> l
- ) [] sys_l in
-
+ let equalities =
+ List.fold_left
+ (fun l (vect, info) ->
+ match info.bound with
+ | Some a, Some b ->
+ if a =/ b then
+ (* This an equation *)
+ (vect, a, info.prf, info.neg + info.pos) :: l
+ else l
+ | _ -> l)
+ [] sys_l
+ in
let rec estimate_cost v ct sysl acc tlsys =
match sysl with
- | [] -> (acc,tlsys)
- | (l,info)::rsys ->
- let ln = info.pos + info.neg in
- let (b,l) = unroll_until v l in
- match b with
- | true ->
- if itv_point info.bound
- then estimate_cost v ct rsys (acc+ln) ((l,info)::tlsys) (* this is free *)
- else estimate_cost v ct rsys (acc+ln+ct) ((l,info)::tlsys) (* should be more ? *)
- | false -> estimate_cost v ct rsys (acc+ln) ((l,info)::tlsys) in
-
- match choose_primal_equation equalities sys_l with
- | None ->
- let cost_eq eq const prf ln acc_costs =
-
- let rec cost_eq eqr sysl costs =
- match Vect.choose eqr with
- | None -> costs
- | Some(v,_,eqr) -> let (cst,tlsys) = estimate_cost v (ln-1) sysl 0 [] in
- cost_eq eqr tlsys (((v,eq,const,prf),cst)::costs) in
- cost_eq eq sys_l acc_costs in
-
- let all_costs = List.fold_left (fun all_costs (vect,const,prf,ln) -> cost_eq vect const prf ln all_costs) [] equalities in
-
- (* pp_list (fun o ((v,eq,_,_),cst) -> Printf.fprintf o "((%i,%a),%i)\n" v pp_vect eq cst) stdout all_costs ; *)
-
- List.sort (fun x y -> Int.compare (snd x) (snd y) ) all_costs
- | Some (v,vect, const,prf,_) -> [(v,vect,const,prf),0]
-
-
+ | [] -> (acc, tlsys)
+ | (l, info) :: rsys -> (
+ let ln = info.pos + info.neg in
+ let b, l = unroll_until v l in
+ match b with
+ | true ->
+ if itv_point info.bound then
+ estimate_cost v ct rsys (acc + ln) ((l, info) :: tlsys)
+ (* this is free *)
+ else estimate_cost v ct rsys (acc + ln + ct) ((l, info) :: tlsys)
+ (* should be more ? *)
+ | false -> estimate_cost v ct rsys (acc + ln) ((l, info) :: tlsys) )
+ in
+ match choose_primal_equation equalities sys_l with
+ | None ->
+ let cost_eq eq const prf ln acc_costs =
+ let rec cost_eq eqr sysl costs =
+ match Vect.choose eqr with
+ | None -> costs
+ | Some (v, _, eqr) ->
+ let cst, tlsys = estimate_cost v (ln - 1) sysl 0 [] in
+ cost_eq eqr tlsys (((v, eq, const, prf), cst) :: costs)
+ in
+ cost_eq eq sys_l acc_costs
+ in
+ let all_costs =
+ List.fold_left
+ (fun all_costs (vect, const, prf, ln) ->
+ cost_eq vect const prf ln all_costs)
+ [] equalities
+ in
+ (* pp_list (fun o ((v,eq,_,_),cst) -> Printf.fprintf o "((%i,%a),%i)\n" v pp_vect eq cst) stdout all_costs ; *)
+ List.sort (fun x y -> Int.compare (snd x) (snd y)) all_costs
+ | Some (v, vect, const, prf, _) -> [((v, vect, const, prf), 0)]
end
-open EstimateElimEq
-module Fourier =
-struct
+open EstimateElimEq
+module Fourier = struct
let optimise vect l =
(* We add a dummy (fresh) variable for vector *)
- let fresh =
- List.fold_left (fun fr c -> max fr (Vect.fresh c.coeffs)) 0 l in
- let cstr = {
- coeffs = Vect.set fresh (Int (-1)) vect ;
- op = Eq ;
- cst = (Int 0)} in
- match solve fresh choose_equality_var choose_variable (cstr::l) with
- | Inr prf -> None (* This is an unsatisfiability proof *)
- | Inl (s,_) ->
- try
- Some (bound_of_variable IMap.empty fresh s.sys)
- with x when CErrors.noncritical x ->
- Printf.printf "optimise Exception : %s" (Printexc.to_string x);
- None
-
+ let fresh = List.fold_left (fun fr c -> max fr (Vect.fresh c.coeffs)) 0 l in
+ let cstr =
+ {coeffs = Vect.set fresh (Int (-1)) vect; op = Eq; cst = Int 0}
+ in
+ match solve fresh choose_equality_var choose_variable (cstr :: l) with
+ | Inr prf -> None (* This is an unsatisfiability proof *)
+ | Inl (s, _) -> (
+ try Some (bound_of_variable IMap.empty fresh s.sys)
+ with x when CErrors.noncritical x ->
+ Printf.printf "optimise Exception : %s" (Printexc.to_string x);
+ None )
let find_point cstrs =
-
match solve max_int choose_equality_var choose_variable cstrs with
- | Inr prf -> Inr prf
- | Inl (_,l) ->
-
- let rec rebuild_solution l map =
- match l with
- | [] -> map
- | (v,e)::l ->
- let itv = bound_of_variable map v e.sys in
- let map = IMap.add v (pick_small_value itv) map in
- rebuild_solution l map
- in
-
- let map = rebuild_solution l IMap.empty in
- let vect = IMap.fold (fun v i vect -> Vect.set v i vect) map Vect.null in
- if debug then Printf.printf "SOLUTION %a" Vect.pp vect ;
- let res = Inl vect in
- res
-
-
+ | Inr prf -> Inr prf
+ | Inl (_, l) ->
+ let rec rebuild_solution l map =
+ match l with
+ | [] -> map
+ | (v, e) :: l ->
+ let itv = bound_of_variable map v e.sys in
+ let map = IMap.add v (pick_small_value itv) map in
+ rebuild_solution l map
+ in
+ let map = rebuild_solution l IMap.empty in
+ let vect = IMap.fold (fun v i vect -> Vect.set v i vect) map Vect.null in
+ if debug then Printf.printf "SOLUTION %a" Vect.pp vect;
+ let res = Inl vect in
+ res
end
-
-module Proof =
-struct
-
-
-
-
-(** A proof term in the sense of a ZMicromega.RatProof is a positive combination of the hypotheses which leads to a contradiction.
+module Proof = struct
+ (** A proof term in the sense of a ZMicromega.RatProof is a positive combination of the hypotheses which leads to a contradiction.
The proofs constructed by Fourier elimination are more like execution traces:
- certain facts are recorded but are useless
- certain inferences are implicit.
@@ -641,124 +606,123 @@ struct
*)
let add x y = fst (add x y)
-
let forall_pairs f l1 l2 =
- List.fold_left (fun acc e1 ->
- List.fold_left (fun acc e2 ->
- match f e1 e2 with
- | None -> acc
- | Some v -> v::acc) acc l2) [] l1
-
-
- let add_op x y =
- match x , y with
- | Eq , Eq -> Eq
- | _ -> Ge
-
-
- let pivot v (p1,c1) (p2,c2) =
- let {coeffs = v1 ; op = op1 ; cst = n1} = c1
- and {coeffs = v2 ; op = op2 ; cst = n2} = c2 in
-
- match Vect.get v v1 , Vect.get v v2 with
- | Int 0 , _ | _ , Int 0 -> None
- | a , b ->
- if Int.equal ((sign_num a) * (sign_num b)) (-1)
- then
- Some (add (p1,abs_num a) (p2,abs_num b) ,
- {coeffs = add (v1,abs_num a) (v2,abs_num b) ;
- op = add_op op1 op2 ;
- cst = n1 // (abs_num a) +/ n2 // (abs_num b) })
- else if op1 == Eq
- then Some (add (p1,minus_num (a // b)) (p2,Int 1),
- {coeffs = add (v1,minus_num (a// b)) (v2 ,Int 1) ;
- op = add_op op1 op2;
- cst = n1 // (minus_num (a// b)) +/ n2 // (Int 1)})
- else if op2 == Eq
- then
- Some (add (p2,minus_num (b // a)) (p1,Int 1),
- {coeffs = add (v2,minus_num (b// a)) (v1 ,Int 1) ;
- op = add_op op1 op2;
- cst = n2 // (minus_num (b// a)) +/ n1 // (Int 1)})
- else None (* op2 could be Eq ... this might happen *)
-
+ List.fold_left
+ (fun acc e1 ->
+ List.fold_left
+ (fun acc e2 -> match f e1 e2 with None -> acc | Some v -> v :: acc)
+ acc l2)
+ [] l1
+
+ let add_op x y = match (x, y) with Eq, Eq -> Eq | _ -> Ge
+
+ let pivot v (p1, c1) (p2, c2) =
+ let {coeffs = v1; op = op1; cst = n1} = c1
+ and {coeffs = v2; op = op2; cst = n2} = c2 in
+ match (Vect.get v v1, Vect.get v v2) with
+ | Int 0, _ | _, Int 0 -> None
+ | a, b ->
+ if Int.equal (sign_num a * sign_num b) (-1) then
+ Some
+ ( add (p1, abs_num a) (p2, abs_num b)
+ , { coeffs = add (v1, abs_num a) (v2, abs_num b)
+ ; op = add_op op1 op2
+ ; cst = (n1 // abs_num a) +/ (n2 // abs_num b) } )
+ else if op1 == Eq then
+ Some
+ ( add (p1, minus_num (a // b)) (p2, Int 1)
+ , { coeffs = add (v1, minus_num (a // b)) (v2, Int 1)
+ ; op = add_op op1 op2
+ ; cst = (n1 // minus_num (a // b)) +/ (n2 // Int 1) } )
+ else if op2 == Eq then
+ Some
+ ( add (p2, minus_num (b // a)) (p1, Int 1)
+ , { coeffs = add (v2, minus_num (b // a)) (v1, Int 1)
+ ; op = add_op op1 op2
+ ; cst = (n2 // minus_num (b // a)) +/ (n1 // Int 1) } )
+ else None
+
+ (* op2 could be Eq ... this might happen *)
let normalise_proofs l =
- List.fold_left (fun acc (prf,cstr) ->
- match acc with
+ List.fold_left
+ (fun acc (prf, cstr) ->
+ match acc with
| Inr _ -> acc (* I already found a contradiction *)
- | Inl acc ->
- match norm_cstr cstr 0 with
- | Redundant -> Inl acc
- | Contradiction -> Inr (prf,cstr)
- | Cstr(v,info) -> Inl ((prf,cstr,v,info)::acc)) (Inl []) l
-
+ | Inl acc -> (
+ match norm_cstr cstr 0 with
+ | Redundant -> Inl acc
+ | Contradiction -> Inr (prf, cstr)
+ | Cstr (v, info) -> Inl ((prf, cstr, v, info) :: acc) ))
+ (Inl []) l
type oproof = (vector * cstr * num) option
- let merge_proof (oleft:oproof) (prf,cstr,v,info) (oright:oproof) =
- let (l,r) = info.bound in
-
+ let merge_proof (oleft : oproof) (prf, cstr, v, info) (oright : oproof) =
+ let l, r = info.bound in
let keep p ob bd =
- match ob , bd with
- | None , None -> None
- | None , Some b -> Some(prf,cstr,b)
- | Some _ , None -> ob
- | Some(prfl,cstrl,bl) , Some b -> if p bl b then Some(prf,cstr, b) else ob in
-
- let oleft = keep (<=/) oleft l in
- let oright = keep (>=/) oright r in
- (* Now, there might be a contradiction *)
- match oleft , oright with
- | None , _ | _ , None -> Inl (oleft,oright)
- | Some(prfl,cstrl,l) , Some(prfr,cstrr,r) ->
- if l <=/ r
- then Inl (oleft,oright)
- else (* There is a contradiction - it should show up by scaling up the vectors - any pivot should do*)
- match Vect.choose cstrr.coeffs with
- | None -> Inr (add (prfl,Int 1) (prfr,Int 1), cstrr) (* this is wrong *)
- | Some(v,_,_) ->
- match pivot v (prfl,cstrl) (prfr,cstrr) with
- | None -> failwith "merge_proof : pivot is not possible"
- | Some x -> Inr x
-
-let mk_proof hyps prf =
- (* I am keeping list - I might have a proof for the left bound and a proof for the right bound.
+ match (ob, bd) with
+ | None, None -> None
+ | None, Some b -> Some (prf, cstr, b)
+ | Some _, None -> ob
+ | Some (prfl, cstrl, bl), Some b ->
+ if p bl b then Some (prf, cstr, b) else ob
+ in
+ let oleft = keep ( <=/ ) oleft l in
+ let oright = keep ( >=/ ) oright r in
+ (* Now, there might be a contradiction *)
+ match (oleft, oright) with
+ | None, _ | _, None -> Inl (oleft, oright)
+ | Some (prfl, cstrl, l), Some (prfr, cstrr, r) -> (
+ if l <=/ r then Inl (oleft, oright)
+ else
+ (* There is a contradiction - it should show up by scaling up the vectors - any pivot should do*)
+ match Vect.choose cstrr.coeffs with
+ | None ->
+ Inr (add (prfl, Int 1) (prfr, Int 1), cstrr) (* this is wrong *)
+ | Some (v, _, _) -> (
+ match pivot v (prfl, cstrl) (prfr, cstrr) with
+ | None -> failwith "merge_proof : pivot is not possible"
+ | Some x -> Inr x ) )
+
+ let mk_proof hyps prf =
+ (* I am keeping list - I might have a proof for the left bound and a proof for the right bound.
If I perform aggressive elimination of redundancies, I expect the list to be of length at most 2.
For each proof list, all the vectors should be of the form a.v for different constants a.
*)
-
- let rec mk_proof prf =
- match prf with
- | Assum i -> [ (Vect.set i (Int 1) Vect.null , List.nth hyps i) ]
-
- | 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 meaningful *)
- forall_pairs (pivot v) prfsl prfsr
- | And(prf1,prf2) ->
- let prfsl1 = mk_proof prf1
- and prfsl2 = mk_proof prf2 in
- (* detect trivial redundancies and contradictions *)
- match normalise_proofs (prfsl1@prfsl2) with
- | Inr x -> [x] (* This is a contradiction - this should be the end of the proof *)
- | Inl l -> (* All the vectors are the same *)
- let prfs =
- List.fold_left (fun acc e ->
- match acc with
- | Inr _ -> acc (* I have a contradiction *)
- | Inl (oleft,oright) -> merge_proof oleft e oright) (Inl(None,None)) l in
- match prfs with
- | Inr x -> [x]
- | Inl (oleft,oright) ->
- match oleft , oright with
- | None , None -> []
- | None , Some(prf,cstr,_) | Some(prf,cstr,_) , None -> [prf,cstr]
- | Some(prf1,cstr1,_) , Some(prf2,cstr2,_) -> [prf1,cstr1;prf2,cstr2] in
-
+ let rec mk_proof prf =
+ match prf with
+ | Assum i -> [(Vect.set i (Int 1) Vect.null, List.nth hyps i)]
+ | 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 meaningful *)
+ forall_pairs (pivot v) prfsl prfsr
+ | And (prf1, prf2) -> (
+ let prfsl1 = mk_proof prf1 and prfsl2 = mk_proof prf2 in
+ (* detect trivial redundancies and contradictions *)
+ match normalise_proofs (prfsl1 @ prfsl2) with
+ | Inr x -> [x]
+ (* This is a contradiction - this should be the end of the proof *)
+ | Inl l -> (
+ (* All the vectors are the same *)
+ let prfs =
+ List.fold_left
+ (fun acc e ->
+ match acc with
+ | Inr _ -> acc (* I have a contradiction *)
+ | Inl (oleft, oright) -> merge_proof oleft e oright)
+ (Inl (None, None))
+ l
+ in
+ match prfs with
+ | Inr x -> [x]
+ | Inl (oleft, oright) -> (
+ match (oleft, oright) with
+ | None, None -> []
+ | None, Some (prf, cstr, _) | Some (prf, cstr, _), None ->
+ [(prf, cstr)]
+ | Some (prf1, cstr1, _), Some (prf2, cstr2, _) ->
+ [(prf1, cstr1); (prf2, cstr2)] ) ) )
+ in
mk_proof prf
-
-
end
-
diff --git a/plugins/micromega/mfourier.mli b/plugins/micromega/mfourier.mli
index 16cb49c85e..8743f0ccc4 100644
--- a/plugins/micromega/mfourier.mli
+++ b/plugins/micromega/mfourier.mli
@@ -13,26 +13,17 @@ module IMap : CSig.MapS with type key = int
type proof
module Fourier : sig
-
-
- val find_point : Polynomial.cstr list ->
- (Vect.t, proof) Util.union
-
- val optimise : Vect.t ->
- Polynomial.cstr list ->
- Itv.interval option
-
+ val find_point : Polynomial.cstr list -> (Vect.t, proof) Util.union
+ val optimise : Vect.t -> Polynomial.cstr list -> Itv.interval option
end
val pp_proof : out_channel -> proof -> unit
module Proof : sig
-
- val mk_proof : Polynomial.cstr list ->
- proof -> (Vect.t * Polynomial.cstr) list
+ val mk_proof :
+ Polynomial.cstr list -> proof -> (Vect.t * Polynomial.cstr) list
val add_op : Polynomial.op -> Polynomial.op -> Polynomial.op
-
end
exception TimeOut
diff --git a/plugins/micromega/micromega.mli b/plugins/micromega/micromega.mli
index 822fde9ab0..033f5d8732 100644
--- a/plugins/micromega/micromega.mli
+++ b/plugins/micromega/micromega.mli
@@ -1,275 +1,274 @@
-
type __ = Obj.t
-
-type unit0 =
-| Tt
+type unit0 = Tt
val negb : bool -> bool
-type nat =
-| O
-| S of nat
-
-type ('a, 'b) sum =
-| Inl of 'a
-| Inr of 'b
-
-val fst : ('a1 * 'a2) -> 'a1
-
-val snd : ('a1 * 'a2) -> 'a2
+type nat = O | S of nat
+type ('a, 'b) sum = Inl of 'a | Inr of 'b
+val fst : 'a1 * 'a2 -> 'a1
+val snd : 'a1 * 'a2 -> 'a2
val app : 'a1 list -> 'a1 list -> 'a1 list
-type comparison =
-| Eq
-| Lt
-| Gt
+type comparison = Eq | Lt | Gt
val compOpp : comparison -> comparison
-
val add : nat -> nat -> nat
-
val nth : nat -> 'a1 list -> 'a1 -> 'a1
-
val rev_append : 'a1 list -> 'a1 list -> 'a1 list
-
val map : ('a1 -> 'a2) -> 'a1 list -> 'a2 list
-
val fold_left : ('a1 -> 'a2 -> 'a1) -> 'a2 list -> 'a1 -> 'a1
-
val fold_right : ('a2 -> 'a1 -> 'a1) -> 'a1 -> 'a2 list -> 'a1
-type positive =
-| XI of positive
-| XO of positive
-| XH
-
-type n =
-| N0
-| Npos of positive
-
-type z =
-| Z0
-| Zpos of positive
-| Zneg of positive
-
-module Pos :
- sig
- type mask =
- | IsNul
- | IsPos of positive
- | IsNeg
- end
-
-module Coq_Pos :
- sig
- val succ : positive -> positive
+type positive = XI of positive | XO of positive | XH
+type n = N0 | Npos of positive
+type z = Z0 | Zpos of positive | Zneg of positive
- val add : positive -> positive -> positive
+module Pos : sig
+ type mask = IsNul | IsPos of positive | IsNeg
+end
+module Coq_Pos : sig
+ val succ : positive -> positive
+ val add : positive -> positive -> positive
val add_carry : positive -> positive -> positive
-
val pred_double : positive -> positive
- type mask = Pos.mask =
- | IsNul
- | IsPos of positive
- | IsNeg
+ type mask = Pos.mask = IsNul | IsPos of positive | IsNeg
val succ_double_mask : mask -> mask
-
val double_mask : mask -> mask
-
val double_pred_mask : positive -> 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 : ('a1 -> 'a1) -> 'a1 -> positive -> 'a1
-
val size_nat : positive -> nat
-
val compare_cont : comparison -> positive -> positive -> comparison
-
val compare : positive -> positive -> comparison
-
val gcdn : nat -> positive -> positive -> positive
-
val gcd : positive -> positive -> positive
-
val of_succ_nat : nat -> positive
- end
+end
-module N :
- sig
+module N : sig
val of_nat : nat -> n
- end
+end
val pow_pos : ('a1 -> 'a1 -> 'a1) -> 'a1 -> positive -> 'a1
-module Z :
- sig
+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 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 leb : z -> z -> bool
-
val ltb : z -> z -> bool
-
val gtb : z -> z -> bool
-
val max : z -> z -> z
-
val abs : z -> z
-
val to_N : z -> n
-
val of_nat : nat -> z
-
val of_N : n -> z
-
val pos_div_eucl : positive -> z -> z * z
-
val div_eucl : z -> z -> z * z
-
val div : z -> z -> z
-
val gcd : z -> z -> z
- end
+end
val zeq_bool : z -> z -> bool
type 'c pol =
-| Pc of 'c
-| Pinj of positive * 'c pol
-| PX of 'c pol * positive * 'c pol
+ | Pc of 'c
+ | Pinj of positive * 'c pol
+ | PX of 'c pol * positive * 'c pol
val p0 : 'a1 -> 'a1 pol
-
val p1 : 'a1 -> 'a1 pol
-
val peq : ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> bool
-
val mkPinj : positive -> 'a1 pol -> 'a1 pol
-
val mkPinj_pred : positive -> 'a1 pol -> 'a1 pol
-val mkPX : 'a1 -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol
+val mkPX :
+ 'a1 -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol
val mkXi : 'a1 -> 'a1 -> positive -> 'a1 pol
-
val mkX : 'a1 -> 'a1 -> 'a1 pol
-
val popp : ('a1 -> 'a1) -> 'a1 pol -> 'a1 pol
-
val paddC : ('a1 -> 'a1 -> 'a1) -> 'a1 pol -> 'a1 -> 'a1 pol
-
val psubC : ('a1 -> 'a1 -> 'a1) -> 'a1 pol -> 'a1 -> 'a1 pol
val paddI :
- ('a1 -> 'a1 -> 'a1) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1 pol -> positive -> 'a1 pol ->
- 'a1 pol
+ ('a1 -> 'a1 -> 'a1)
+ -> ('a1 pol -> 'a1 pol -> 'a1 pol)
+ -> 'a1 pol
+ -> positive
+ -> 'a1 pol
+ -> 'a1 pol
val psubI :
- ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1 pol -> positive
- -> 'a1 pol -> 'a1 pol
+ ('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
+val padd :
+ 'a1
+ -> ('a1 -> 'a1 -> 'a1)
+ -> ('a1 -> 'a1 -> bool)
+ -> 'a1 pol
+ -> 'a1 pol
+ -> 'a1 pol
val psub :
- 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) ->
- 'a1 pol -> 'a1 pol -> 'a1 pol
+ 'a1
+ -> ('a1 -> 'a1 -> 'a1)
+ -> ('a1 -> 'a1 -> 'a1)
+ -> ('a1 -> 'a1)
+ -> ('a1 -> 'a1 -> bool)
+ -> 'a1 pol
+ -> 'a1 pol
+ -> 'a1 pol
-val pmulC_aux : 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 -> 'a1 pol
+val pmulC_aux :
+ 'a1
+ -> ('a1 -> 'a1 -> 'a1)
+ -> ('a1 -> 'a1 -> bool)
+ -> 'a1 pol
+ -> 'a1
+ -> 'a1 pol
val pmulC :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 -> 'a1 pol
+ '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
+ '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
+ -> 'a1
+ -> ('a1 -> 'a1 -> 'a1)
+ -> ('a1 -> 'a1 -> 'a1)
+ -> ('a1 -> 'a1 -> bool)
+ -> 'a1 pol
-> 'a1 pol
type 'c pExpr =
-| PEc of 'c
-| PEX of positive
-| PEadd of 'c pExpr * 'c pExpr
-| PEsub of 'c pExpr * 'c pExpr
-| PEmul of 'c pExpr * 'c pExpr
-| PEopp of 'c pExpr
-| PEpow of 'c pExpr * n
+ | PEc of 'c
+ | PEX of positive
+ | PEadd of 'c pExpr * 'c pExpr
+ | PEsub of 'c pExpr * 'c pExpr
+ | PEmul of 'c pExpr * 'c pExpr
+ | PEopp of 'c pExpr
+ | PEpow of 'c pExpr * n
val mk_X : 'a1 -> 'a1 -> positive -> 'a1 pol
val ppow_pos :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol
- -> 'a1 pol) -> 'a1 pol -> 'a1 pol -> positive -> 'a1 pol
+ '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 ('tA, 'tX, 'aA, 'aF) gFormula =
-| TT
-| FF
-| X of 'tX
-| A of 'tA * 'aA
-| Cj of ('tA, 'tX, 'aA, 'aF) gFormula * ('tA, 'tX, 'aA, 'aF) gFormula
-| D of ('tA, 'tX, 'aA, 'aF) gFormula * ('tA, 'tX, 'aA, 'aF) gFormula
-| N of ('tA, 'tX, 'aA, 'aF) gFormula
-| I of ('tA, 'tX, 'aA, 'aF) gFormula * 'aF option * ('tA, 'tX, 'aA, 'aF) gFormula
-
-val mapX : ('a2 -> 'a2) -> ('a1, 'a2, 'a3, 'a4) gFormula -> ('a1, 'a2, 'a3, 'a4) gFormula
+ | TT
+ | FF
+ | X of 'tX
+ | A of 'tA * 'aA
+ | Cj of ('tA, 'tX, 'aA, 'aF) gFormula * ('tA, 'tX, 'aA, 'aF) gFormula
+ | D of ('tA, 'tX, 'aA, 'aF) gFormula * ('tA, 'tX, 'aA, 'aF) gFormula
+ | N of ('tA, 'tX, 'aA, 'aF) gFormula
+ | I of
+ ('tA, 'tX, 'aA, 'aF) gFormula * 'aF option * ('tA, 'tX, 'aA, 'aF) gFormula
+
+val mapX :
+ ('a2 -> 'a2) -> ('a1, 'a2, 'a3, 'a4) gFormula -> ('a1, 'a2, 'a3, 'a4) gFormula
val foldA : ('a5 -> 'a3 -> 'a5) -> ('a1, 'a2, 'a3, 'a4) gFormula -> 'a5 -> 'a5
-
val cons_id : 'a1 option -> 'a1 list -> 'a1 list
-
val ids_of_formula : ('a1, 'a2, 'a3, 'a4) gFormula -> 'a4 list
-
val collect_annot : ('a1, 'a2, 'a3, 'a4) gFormula -> 'a3 list
type 'a bFormula = ('a, __, unit0, unit0) gFormula
@@ -278,411 +277,482 @@ val map_bformula :
('a1 -> 'a2) -> ('a1, 'a3, 'a4, 'a5) gFormula -> ('a2, 'a3, 'a4, 'a5) gFormula
type ('x, 'annot) clause = ('x * 'annot) list
-
type ('x, 'annot) cnf = ('x, 'annot) clause list
val cnf_tt : ('a1, 'a2) cnf
-
val cnf_ff : ('a1, 'a2) cnf
val add_term :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1 * 'a2) -> ('a1, 'a2) clause -> ('a1, 'a2)
- clause option
+ ('a1 -> bool)
+ -> ('a1 -> 'a1 -> 'a1 option)
+ -> 'a1 * 'a2
+ -> ('a1, 'a2) clause
+ -> ('a1, 'a2) clause option
val or_clause :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) clause -> ('a1, 'a2) clause ->
- ('a1, 'a2) clause option
+ ('a1 -> bool)
+ -> ('a1 -> 'a1 -> 'a1 option)
+ -> ('a1, 'a2) clause
+ -> ('a1, 'a2) clause
+ -> ('a1, 'a2) clause option
val xor_clause_cnf :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) clause -> ('a1, 'a2) cnf -> ('a1,
- 'a2) cnf
+ ('a1 -> bool)
+ -> ('a1 -> 'a1 -> 'a1 option)
+ -> ('a1, 'a2) clause
+ -> ('a1, 'a2) cnf
+ -> ('a1, 'a2) cnf
val or_clause_cnf :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) clause -> ('a1, 'a2) cnf -> ('a1,
- 'a2) cnf
+ ('a1 -> bool)
+ -> ('a1 -> 'a1 -> 'a1 option)
+ -> ('a1, 'a2) clause
+ -> ('a1, 'a2) cnf
+ -> ('a1, 'a2) cnf
val or_cnf :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf -> ('a1, 'a2)
- cnf
+ ('a1 -> bool)
+ -> ('a1 -> 'a1 -> 'a1 option)
+ -> ('a1, 'a2) cnf
+ -> ('a1, 'a2) cnf
+ -> ('a1, 'a2) cnf
val and_cnf : ('a1, 'a2) cnf -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf
type ('term, 'annot, 'tX, 'aF) tFormula = ('term, 'tX, 'annot, 'aF) gFormula
val is_cnf_tt : ('a1, 'a2) cnf -> bool
-
val is_cnf_ff : ('a1, 'a2) cnf -> bool
-
val and_cnf_opt : ('a1, 'a2) cnf -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf
val or_cnf_opt :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf -> ('a1, 'a2)
- cnf
+ ('a1 -> bool)
+ -> ('a1 -> 'a1 -> 'a1 option)
+ -> ('a1, 'a2) cnf
+ -> ('a1, 'a2) cnf
+ -> ('a1, 'a2) cnf
val xcnf :
- ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a3 -> ('a2, 'a3) cnf) -> ('a1 -> 'a3
- -> ('a2, 'a3) cnf) -> bool -> ('a1, 'a3, 'a4, 'a5) tFormula -> ('a2, 'a3) cnf
+ ('a2 -> bool)
+ -> ('a2 -> 'a2 -> 'a2 option)
+ -> ('a1 -> 'a3 -> ('a2, 'a3) cnf)
+ -> ('a1 -> 'a3 -> ('a2, 'a3) cnf)
+ -> bool
+ -> ('a1, 'a3, 'a4, 'a5) tFormula
+ -> ('a2, 'a3) cnf
val radd_term :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1 * 'a2) -> ('a1, 'a2) clause -> (('a1,
- 'a2) clause, 'a2 list) sum
+ ('a1 -> bool)
+ -> ('a1 -> 'a1 -> 'a1 option)
+ -> 'a1 * 'a2
+ -> ('a1, 'a2) clause
+ -> (('a1, 'a2) clause, 'a2 list) sum
val ror_clause :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1 * 'a2) list -> ('a1, 'a2) clause ->
- (('a1, 'a2) clause, 'a2 list) sum
+ ('a1 -> bool)
+ -> ('a1 -> 'a1 -> 'a1 option)
+ -> ('a1 * 'a2) list
+ -> ('a1, 'a2) clause
+ -> (('a1, 'a2) clause, 'a2 list) sum
val xror_clause_cnf :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1 * 'a2) list -> ('a1, 'a2) clause list ->
- ('a1, 'a2) clause list * 'a2 list
+ ('a1 -> bool)
+ -> ('a1 -> 'a1 -> 'a1 option)
+ -> ('a1 * 'a2) list
+ -> ('a1, 'a2) clause list
+ -> ('a1, 'a2) clause list * 'a2 list
val ror_clause_cnf :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1 * 'a2) list -> ('a1, 'a2) clause list ->
- ('a1, 'a2) clause list * 'a2 list
+ ('a1 -> bool)
+ -> ('a1 -> 'a1 -> 'a1 option)
+ -> ('a1 * 'a2) list
+ -> ('a1, 'a2) clause list
+ -> ('a1, 'a2) clause list * 'a2 list
val ror_cnf :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) clause list -> ('a1, 'a2) clause
- list -> ('a1, 'a2) cnf * 'a2 list
+ ('a1 -> bool)
+ -> ('a1 -> 'a1 -> 'a1 option)
+ -> ('a1, 'a2) clause list
+ -> ('a1, 'a2) clause list
+ -> ('a1, 'a2) cnf * 'a2 list
val ror_cnf_opt :
- ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> ('a1, 'a2) cnf -> ('a1, 'a2) cnf -> ('a1, 'a2)
- cnf * 'a2 list
+ ('a1 -> bool)
+ -> ('a1 -> 'a1 -> 'a1 option)
+ -> ('a1, 'a2) cnf
+ -> ('a1, 'a2) cnf
+ -> ('a1, 'a2) cnf * 'a2 list
val ratom : ('a1, 'a2) cnf -> 'a2 -> ('a1, 'a2) cnf * 'a2 list
val rxcnf :
- ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a3 -> ('a2, 'a3) cnf) -> ('a1 -> 'a3
- -> ('a2, 'a3) cnf) -> bool -> ('a1, 'a3, 'a4, 'a5) tFormula -> ('a2, 'a3) cnf * 'a3 list
-
-type ('term, 'annot, 'tX) to_constrT = { mkTT : 'tX; mkFF : 'tX;
- mkA : ('term -> 'annot -> 'tX);
- mkCj : ('tX -> 'tX -> 'tX); mkD : ('tX -> 'tX -> 'tX);
- mkI : ('tX -> 'tX -> 'tX); mkN : ('tX -> 'tX) }
-
-val aformula : ('a1, 'a2, 'a3) to_constrT -> ('a1, 'a2, 'a3, 'a4) tFormula -> 'a3
+ ('a2 -> bool)
+ -> ('a2 -> 'a2 -> 'a2 option)
+ -> ('a1 -> 'a3 -> ('a2, 'a3) cnf)
+ -> ('a1 -> 'a3 -> ('a2, 'a3) cnf)
+ -> bool
+ -> ('a1, 'a3, 'a4, 'a5) tFormula
+ -> ('a2, 'a3) cnf * 'a3 list
+
+type ('term, 'annot, 'tX) to_constrT =
+ { mkTT : 'tX
+ ; mkFF : 'tX
+ ; mkA : 'term -> 'annot -> 'tX
+ ; mkCj : 'tX -> 'tX -> 'tX
+ ; mkD : 'tX -> 'tX -> 'tX
+ ; mkI : 'tX -> 'tX -> 'tX
+ ; mkN : 'tX -> 'tX }
+
+val aformula :
+ ('a1, 'a2, 'a3) to_constrT -> ('a1, 'a2, 'a3, 'a4) tFormula -> 'a3
val is_X : ('a1, 'a2, 'a3, 'a4) tFormula -> 'a3 option
val abs_and :
- ('a1, 'a2, 'a3) to_constrT -> ('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4) tFormula
- -> (('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4)
- tFormula) -> ('a1, 'a3, 'a2, 'a4) gFormula
+ ('a1, 'a2, 'a3) to_constrT
+ -> ('a1, 'a2, 'a3, 'a4) tFormula
+ -> ('a1, 'a2, 'a3, 'a4) tFormula
+ -> ( ('a1, 'a2, 'a3, 'a4) tFormula
+ -> ('a1, 'a2, 'a3, 'a4) tFormula
+ -> ('a1, 'a2, 'a3, 'a4) tFormula)
+ -> ('a1, 'a3, 'a2, 'a4) gFormula
val abs_or :
- ('a1, 'a2, 'a3) to_constrT -> ('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4) tFormula
- -> (('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4)
- tFormula) -> ('a1, 'a3, 'a2, 'a4) gFormula
+ ('a1, 'a2, 'a3) to_constrT
+ -> ('a1, 'a2, 'a3, 'a4) tFormula
+ -> ('a1, 'a2, 'a3, 'a4) tFormula
+ -> ( ('a1, 'a2, 'a3, 'a4) tFormula
+ -> ('a1, 'a2, 'a3, 'a4) tFormula
+ -> ('a1, 'a2, 'a3, 'a4) tFormula)
+ -> ('a1, 'a3, 'a2, 'a4) gFormula
val mk_arrow :
- 'a4 option -> ('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2, 'a3, 'a4) tFormula -> ('a1, 'a2,
- 'a3, 'a4) tFormula
+ 'a4 option
+ -> ('a1, 'a2, 'a3, 'a4) tFormula
+ -> ('a1, 'a2, 'a3, 'a4) tFormula
+ -> ('a1, 'a2, 'a3, 'a4) tFormula
val abst_form :
- ('a1, 'a2, 'a3) to_constrT -> ('a2 -> bool) -> bool -> ('a1, 'a2, 'a3, 'a4) tFormula -> ('a1,
- 'a3, 'a2, 'a4) gFormula
+ ('a1, 'a2, 'a3) to_constrT
+ -> ('a2 -> bool)
+ -> bool
+ -> ('a1, 'a2, 'a3, 'a4) tFormula
+ -> ('a1, 'a3, 'a2, 'a4) gFormula
-val cnf_checker : (('a1 * 'a2) list -> 'a3 -> bool) -> ('a1, 'a2) cnf -> 'a3 list -> bool
+val cnf_checker :
+ (('a1 * 'a2) list -> 'a3 -> bool) -> ('a1, 'a2) cnf -> 'a3 list -> bool
val tauto_checker :
- ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a3 -> ('a2, 'a3) cnf) -> ('a1 -> 'a3
- -> ('a2, 'a3) cnf) -> (('a2 * 'a3) list -> 'a4 -> bool) -> ('a1, __, 'a3, unit0) gFormula ->
- 'a4 list -> bool
+ ('a2 -> bool)
+ -> ('a2 -> 'a2 -> 'a2 option)
+ -> ('a1 -> 'a3 -> ('a2, 'a3) cnf)
+ -> ('a1 -> 'a3 -> ('a2, 'a3) cnf)
+ -> (('a2 * 'a3) list -> 'a4 -> bool)
+ -> ('a1, __, 'a3, unit0) gFormula
+ -> 'a4 list
+ -> bool
val cneqb : ('a1 -> 'a1 -> bool) -> 'a1 -> 'a1 -> bool
-
val cltb : ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 -> 'a1 -> bool
type 'c polC = 'c pol
-
-type op1 =
-| Equal
-| NonEqual
-| Strict
-| NonStrict
-
+type op1 = Equal | NonEqual | Strict | NonStrict
type 'c nFormula = 'c polC * op1
val opMult : op1 -> op1 -> op1 option
-
val opAdd : op1 -> op1 -> op1 option
type 'c psatz =
-| PsatzIn of nat
-| PsatzSquare of 'c polC
-| PsatzMulC of 'c polC * 'c psatz
-| PsatzMulE of 'c psatz * 'c psatz
-| PsatzAdd of 'c psatz * 'c psatz
-| PsatzC of 'c
-| PsatzZ
+ | PsatzIn of nat
+ | PsatzSquare of 'c polC
+ | PsatzMulC of 'c polC * 'c psatz
+ | PsatzMulE of 'c psatz * 'c psatz
+ | PsatzAdd of 'c psatz * 'c psatz
+ | PsatzC of 'c
+ | PsatzZ
val map_option : ('a1 -> 'a2 option) -> 'a1 option -> 'a2 option
-val map_option2 : ('a1 -> 'a2 -> 'a3 option) -> 'a1 option -> 'a2 option -> 'a3 option
+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
val check_normalised_formulas :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 ->
- 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz -> bool
-
-type op2 =
-| OpEq
-| OpNEq
-| OpLe
-| OpGe
-| OpLt
-| OpGt
-
-type 't formula = { flhs : 't pExpr; fop : op2; frhs : 't pExpr }
+ 'a1
+ -> 'a1
+ -> ('a1 -> 'a1 -> 'a1)
+ -> ('a1 -> 'a1 -> 'a1)
+ -> ('a1 -> 'a1 -> bool)
+ -> ('a1 -> 'a1 -> bool)
+ -> 'a1 nFormula list
+ -> 'a1 psatz
+ -> bool
+
+type op2 = OpEq | OpNEq | OpLe | OpGe | OpLt | OpGt
+type 't formula = {flhs : 't pExpr; fop : op2; frhs : 't 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
+val padd0 :
+ 'a1
+ -> ('a1 -> 'a1 -> 'a1)
+ -> ('a1 -> 'a1 -> bool)
+ -> 'a1 pol
+ -> 'a1 pol
+ -> 'a1 pol
val popp0 : ('a1 -> 'a1) -> 'a1 pol -> 'a1 pol
val normalise :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
- 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1 nFormula
+ 'a1
+ -> 'a1
+ -> ('a1 -> 'a1 -> 'a1)
+ -> ('a1 -> 'a1 -> 'a1)
+ -> ('a1 -> 'a1 -> 'a1)
+ -> ('a1 -> 'a1)
+ -> ('a1 -> 'a1 -> bool)
+ -> 'a1 formula
+ -> 'a1 nFormula
val xnormalise : ('a1 -> 'a1) -> 'a1 nFormula -> 'a1 nFormula list
-
val xnegate : ('a1 -> 'a1) -> 'a1 nFormula -> 'a1 nFormula list
val cnf_of_list :
- 'a1 -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a2 -> ('a1
- nFormula, 'a2) cnf
+ 'a1
+ -> ('a1 -> 'a1 -> bool)
+ -> ('a1 -> 'a1 -> bool)
+ -> 'a1 nFormula list
+ -> 'a2
+ -> ('a1 nFormula, 'a2) cnf
val cnf_normalise :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
- 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a2 -> ('a1 nFormula,
- 'a2) cnf
+ 'a1
+ -> 'a1
+ -> ('a1 -> 'a1 -> 'a1)
+ -> ('a1 -> 'a1 -> 'a1)
+ -> ('a1 -> 'a1 -> 'a1)
+ -> ('a1 -> 'a1)
+ -> ('a1 -> 'a1 -> bool)
+ -> ('a1 -> 'a1 -> bool)
+ -> 'a1 formula
+ -> 'a2
+ -> ('a1 nFormula, 'a2) cnf
val cnf_negate :
- 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 ->
- 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a2 -> ('a1 nFormula,
- 'a2) cnf
+ 'a1
+ -> 'a1
+ -> ('a1 -> 'a1 -> 'a1)
+ -> ('a1 -> 'a1 -> 'a1)
+ -> ('a1 -> 'a1 -> 'a1)
+ -> ('a1 -> 'a1)
+ -> ('a1 -> 'a1 -> bool)
+ -> ('a1 -> 'a1 -> bool)
+ -> 'a1 formula
+ -> 'a2
+ -> ('a1 nFormula, 'a2) cnf
val xdenorm : positive -> 'a1 pol -> 'a1 pExpr
-
val denorm : 'a1 pol -> 'a1 pExpr
-
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
-module PositiveSet :
- sig
- type tree =
- | Leaf
- | Node of tree * bool * tree
- end
+module PositiveSet : sig
+ type tree = Leaf | Node of tree * bool * tree
+end
-type q = { qnum : z; qden : positive }
+type q = {qnum : z; qden : positive}
val qeq_bool : q -> q -> bool
-
val qle_bool : q -> q -> bool
-
val qplus : q -> q -> q
-
val qmult : q -> q -> q
-
val qopp : q -> q
-
val qminus : q -> q -> q
-
val qinv : q -> q
-
val qpower_positive : q -> positive -> q
-
val qpower : q -> z -> q
-type 'a t =
-| Empty
-| Elt of 'a
-| Branch of 'a t * 'a * 'a t
+type 'a t = Empty | Elt of 'a | Branch of 'a t * 'a * 'a t
val find : 'a1 -> 'a1 t -> positive -> 'a1
-
val singleton : 'a1 -> positive -> 'a1 -> 'a1 t
-
val vm_add : 'a1 -> positive -> 'a1 -> 'a1 t -> 'a1 t
-
val zeval_const : z pExpr -> z option
type zWitness = z psatz
val zWeakChecker : z nFormula list -> z psatz -> bool
-
val psub1 : z pol -> z pol -> z pol
-
val padd1 : z pol -> z pol -> z pol
-
val normZ : z pExpr -> z pol
-
val zunsat : z nFormula -> bool
-
val zdeduce : z nFormula -> z nFormula -> z nFormula option
-
val xnnormalise : z formula -> z nFormula
-
val xnormalise0 : z nFormula -> z nFormula list
-
val cnf_of_list0 : 'a1 -> z nFormula list -> (z nFormula * 'a1) list list
-
val normalise0 : z formula -> 'a1 -> (z nFormula, 'a1) cnf
-
val xnegate0 : z nFormula -> z nFormula list
-
val negate : z formula -> 'a1 -> (z nFormula, 'a1) cnf
-val cnfZ : (z formula, 'a1, 'a2, 'a3) tFormula -> (z nFormula, 'a1) cnf * 'a1 list
+val cnfZ :
+ (z formula, 'a1, 'a2, 'a3) tFormula -> (z nFormula, 'a1) cnf * 'a1 list
val ceiling : z -> z -> z
type zArithProof =
-| DoneProof
-| RatProof of zWitness * zArithProof
-| CutProof of zWitness * zArithProof
-| EnumProof of zWitness * zWitness * zArithProof list
+ | DoneProof
+ | RatProof of zWitness * zArithProof
+ | CutProof of zWitness * zArithProof
+ | EnumProof of zWitness * zWitness * zArithProof list
val zgcdM : z -> z -> z
-
val zgcd_pol : z polC -> z * z
-
val zdiv_pol : z polC -> z -> z polC
-
val makeCuttingPlane : z polC -> z polC * z
-
val genCuttingPlane : z nFormula -> ((z polC * z) * op1) option
-
-val nformula_of_cutting_plane : ((z polC * z) * op1) -> z nFormula
-
+val nformula_of_cutting_plane : (z polC * z) * op1 -> z nFormula
val is_pol_Z0 : z polC -> bool
-
val eval_Psatz0 : z nFormula list -> zWitness -> z nFormula option
-
val valid_cut_sign : op1 -> bool
-module Vars :
- sig
+module Vars : sig
type elt = positive
-
- type tree = PositiveSet.tree =
- | Leaf
- | Node of tree * bool * tree
-
+ type tree = PositiveSet.tree = Leaf | Node of tree * bool * tree
type t = tree
val empty : t
-
val add : elt -> t -> t
-
val singleton : elt -> t
-
val union : t -> t -> t
-
val rev_append : elt -> elt -> elt
-
val rev : elt -> elt
-
val xfold : (elt -> 'a1 -> 'a1) -> t -> 'a1 -> elt -> 'a1
-
val fold : (elt -> 'a1 -> 'a1) -> t -> 'a1 -> 'a1
- end
+end
val vars_of_pexpr : z pExpr -> Vars.t
-
val vars_of_formula : z formula -> Vars.t
-
val vars_of_bformula : (z formula, 'a1, 'a2, 'a3) gFormula -> Vars.t
-
val bound_var : positive -> z formula
-
val mk_eq_pos : positive -> positive -> positive -> z formula
val bound_vars :
- (positive -> positive -> bool option -> 'a2) -> positive -> Vars.t -> (z formula, 'a1, 'a2,
- 'a3) gFormula
+ (positive -> positive -> bool option -> 'a2)
+ -> positive
+ -> Vars.t
+ -> (z formula, 'a1, 'a2, 'a3) gFormula
val bound_problem_fr :
- (positive -> positive -> bool option -> 'a2) -> positive -> (z formula, 'a1, 'a2, 'a3)
- gFormula -> (z formula, 'a1, 'a2, 'a3) gFormula
+ (positive -> positive -> bool option -> 'a2)
+ -> positive
+ -> (z formula, 'a1, 'a2, 'a3) gFormula
+ -> (z formula, 'a1, 'a2, 'a3) gFormula
val zChecker : z nFormula list -> zArithProof -> bool
-
val zTautoChecker : z formula bFormula -> zArithProof list -> bool
type qWitness = q psatz
val qWeakChecker : q nFormula list -> q psatz -> bool
-
val qnormalise : q formula -> 'a1 -> (q nFormula, 'a1) cnf
-
val qnegate : q formula -> 'a1 -> (q nFormula, 'a1) cnf
-
val qunsat : q nFormula -> bool
-
val qdeduce : q nFormula -> q nFormula -> q nFormula option
-
val normQ : q pExpr -> q pol
-val cnfQ : (q formula, 'a1, 'a2, 'a3) tFormula -> (q nFormula, 'a1) cnf * 'a1 list
+val cnfQ :
+ (q formula, 'a1, 'a2, 'a3) tFormula -> (q nFormula, 'a1) cnf * 'a1 list
val qTautoChecker : q formula bFormula -> qWitness list -> bool
type rcst =
-| C0
-| C1
-| CQ of q
-| CZ of z
-| CPlus of rcst * rcst
-| CMinus of rcst * rcst
-| CMult of rcst * rcst
-| CPow of rcst * (z, nat) sum
-| CInv of rcst
-| COpp of rcst
+ | C0
+ | C1
+ | CQ of q
+ | CZ of z
+ | CPlus of rcst * rcst
+ | CMinus of rcst * rcst
+ | CMult of rcst * rcst
+ | CPow of rcst * (z, nat) sum
+ | CInv of rcst
+ | COpp of rcst
val z_of_exp : (z, nat) sum -> z
-
val q_of_Rcst : rcst -> q
type rWitness = q psatz
val rWeakChecker : q nFormula list -> q psatz -> bool
-
val rnormalise : q formula -> 'a1 -> (q nFormula, 'a1) cnf
-
val rnegate : q formula -> 'a1 -> (q nFormula, 'a1) cnf
-
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/mutils.ml b/plugins/micromega/mutils.ml
index a30e963f2a..d92b409ba1 100644
--- a/plugins/micromega/mutils.ml
+++ b/plugins/micromega/mutils.ml
@@ -21,294 +21,246 @@
module Int = struct
type t = int
+
let compare : int -> int -> int = compare
- let equal : int -> int -> bool = (=)
+ let equal : int -> int -> bool = ( = )
end
-module ISet =
- struct
- include Set.Make(Int)
+module ISet = struct
+ include Set.Make (Int)
- let pp o s = iter (fun i -> Printf.fprintf o "%i " i) s
- end
+ let pp o s = iter (fun i -> Printf.fprintf o "%i " i) s
+end
-module IMap =
- struct
- include Map.Make(Int)
+module IMap = struct
+ include Map.Make (Int)
- let from k m =
- let (_,_,r) = split (k-1) m in
- r
- end
+ let from k m =
+ let _, _, r = split (k - 1) m in
+ r
+end
let rec pp_list s f o l =
match l with
- | [] -> ()
- | [e] -> f o e
- | e::l -> f o e ; output_string o s ; pp_list s f o l
+ | [] -> ()
+ | [e] -> f o e
+ | e :: l -> f o e; output_string o s; pp_list s f o l
let finally f rst =
try
let res = f () in
- rst () ; res
+ rst (); res
with reraise ->
- (try rst ()
- with any -> raise reraise
- ); raise reraise
+ (try rst () with any -> raise reraise);
+ raise reraise
let rec try_any l x =
- match l with
+ match l with
| [] -> None
- | (f,s)::l -> match f x with
- | None -> try_any l x
- | x -> x
+ | (f, s) :: l -> ( match f x with None -> try_any l x | x -> x )
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 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
+ match l with [] -> acc | e :: lx -> xpairs (pair_with acc e l) lx
+ in
+ xpairs [] l
let rec is_sublist f l1 l2 =
- match l1 ,l2 with
- | [] ,_ -> true
- | e::l1', [] -> false
- | e::l1' , e'::l2' ->
- if f e e' then is_sublist f l1' l2'
- else is_sublist f l1 l2'
+ match (l1, l2) with
+ | [], _ -> true
+ | e :: l1', [] -> false
+ | e :: l1', e' :: l2' ->
+ if f e e' then is_sublist f l1' l2' else is_sublist f l1 l2'
let extract pred l =
- List.fold_left (fun (fd,sys) e ->
- match fd with
- | None ->
- begin
- match pred e with
- | None -> fd, e::sys
- | Some v -> Some(v,e) , sys
- end
- | _ -> (fd, e::sys)
- ) (None,[]) l
+ List.fold_left
+ (fun (fd, sys) e ->
+ match fd with
+ | None -> (
+ match pred e with None -> (fd, e :: sys) | Some v -> (Some (v, e), sys)
+ )
+ | _ -> (fd, e :: sys))
+ (None, []) l
let extract_best red lt l =
let rec extractb c e rst l =
match l with
- [] -> Some (c,e) , rst
- | e'::l' -> match red e' with
- | None -> extractb c e (e'::rst) l'
- | Some c' -> if lt c' c
- then extractb c' e' (e::rst) l'
- else extractb c e (e'::rst) l' in
+ | [] -> (Some (c, e), rst)
+ | e' :: l' -> (
+ match red e' with
+ | None -> extractb c e (e' :: rst) l'
+ | Some c' ->
+ if lt c' c then extractb c' e' (e :: rst) l'
+ else extractb c e (e' :: rst) l' )
+ in
match extract red l with
- | None , _ -> None,l
- | Some(c,e), rst -> extractb c e [] rst
-
+ | None, _ -> (None, l)
+ | Some (c, e), rst -> extractb c e [] rst
let rec find_option pred l =
match l with
| [] -> raise Not_found
- | e::l -> match pred e with
- | Some r -> r
- | None -> find_option pred l
+ | e :: l -> ( match pred e with Some r -> r | None -> find_option pred l )
-let find_some pred l =
- try Some (find_option pred l) with Not_found -> None
+let find_some pred l = try Some (find_option pred l) with Not_found -> None
-
-let extract_all pred l =
- List.fold_left (fun (s1,s2) e ->
- match pred e with
- | None -> s1,e::s2
- | Some v -> (v,e)::s1 , s2) ([],[]) l
+let extract_all pred l =
+ List.fold_left
+ (fun (s1, s2) e ->
+ match pred e with None -> (s1, e :: s2) | Some v -> ((v, e) :: s1, s2))
+ ([], []) l
let simplify f sys =
- let (sys',b) =
- List.fold_left (fun (sys',b) c ->
- match f c with
- | None -> (c::sys',b)
- | Some c' ->
- (c'::sys',true)
- ) ([],false) sys in
+ let sys', b =
+ List.fold_left
+ (fun (sys', b) c ->
+ match f c with None -> (c :: sys', b) | Some c' -> (c' :: sys', true))
+ ([], false) sys
+ in
if b then Some sys' else None
let generate_acc f acc sys =
- List.fold_left (fun sys' c -> match f c with
- | None -> sys'
- | Some c' -> c'::sys'
- ) acc sys
-
+ List.fold_left
+ (fun sys' c -> match f c with None -> sys' | Some c' -> c' :: sys')
+ acc sys
let generate f sys = generate_acc f [] sys
-
let saturate p f sys =
- let rec sat acc l =
+ let rec sat acc l =
match extract p l with
- | None,_ -> acc
- | Some r,l' ->
- let n = generate (f r) (l'@acc) in
- sat (n@acc) l' in
- try sat [] sys with
- x ->
- begin
- Printexc.print_backtrace stdout ;
- raise x
- end
-
+ | None, _ -> acc
+ | Some r, l' ->
+ let n = generate (f r) (l' @ acc) in
+ sat (n @ acc) l'
+ in
+ try sat [] sys
+ with x ->
+ Printexc.print_backtrace stdout;
+ raise x
open Num
open Big_int
let ppcm x y =
- let g = gcd_big_int x y in
- let x' = div_big_int x g in
- let y' = div_big_int y g in
+ let g = gcd_big_int x y in
+ let x' = div_big_int x g in
+ let y' = div_big_int y g in
mult_big_int g (mult_big_int x' y')
let denominator = function
- | Int _ | Big_int _ -> unit_big_int
- | Ratio r -> Ratio.denominator_ratio r
+ | Int _ | Big_int _ -> unit_big_int
+ | Ratio r -> Ratio.denominator_ratio r
let numerator = function
- | Ratio r -> Ratio.numerator_ratio r
- | Int i -> Big_int.big_int_of_int i
- | Big_int i -> i
+ | Ratio r -> Ratio.numerator_ratio r
+ | Int i -> Big_int.big_int_of_int i
+ | Big_int i -> i
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 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'
-
+ match l with
+ | [] -> None
+ | f :: fl -> ( match f x with None -> app_funs fl x | Some x' -> Some x' )
(**
* MODULE: Coq to Caml data-structure mappings
*)
-module CoqToCaml =
-struct
- open Micromega
+module CoqToCaml = struct
+ open Micromega
- let rec nat = function
- | O -> 0
- | S n -> (nat n) + 1
+ let rec nat = function O -> 0 | S n -> nat n + 1
+ let rec positive p =
+ match p with
+ | XH -> 1
+ | XI p -> 1 + (2 * positive p)
+ | XO p -> 2 * positive p
- let rec positive p =
- match p with
- | XH -> 1
- | XI p -> 1+ 2*(positive p)
- | XO p -> 2*(positive p)
+ let n nt = match nt with N0 -> 0 | Npos p -> positive p
- let n nt =
- match nt with
- | N0 -> 0
- | Npos p -> positive p
+ let rec index i =
+ (* Swap left-right ? *)
+ match i with XH -> 1 | XI i -> 1 + (2 * index i) | XO i -> 2 * index i
- let rec index i = (* Swap left-right ? *)
- match i with
- | XH -> 1
- | XI i -> 1+(2*(index i))
- | XO i -> 2*(index i)
+ open Big_int
- open Big_int
+ let rec positive_big_int p =
+ match p with
+ | XH -> unit_big_int
+ | XI p -> add_int_big_int 1 (mult_int_big_int 2 (positive_big_int p))
+ | XO p -> mult_int_big_int 2 (positive_big_int p)
- let rec positive_big_int p =
- match p with
- | XH -> unit_big_int
- | XI p -> add_int_big_int 1 (mult_int_big_int 2 (positive_big_int p))
- | XO p -> (mult_int_big_int 2 (positive_big_int p))
+ let z_big_int x =
+ match x with
+ | Z0 -> zero_big_int
+ | Zpos p -> positive_big_int p
+ | Zneg p -> minus_big_int (positive_big_int p)
- let z_big_int x =
- match x with
- | Z0 -> zero_big_int
- | Zpos p -> (positive_big_int p)
- | Zneg p -> minus_big_int (positive_big_int p)
-
- let z x =
- match x with
- | Z0 -> 0
- | Zpos p -> index p
- | Zneg p -> - (index p)
-
-
- let q_to_num {qnum = x ; qden = y} =
- Big_int (z_big_int x) // (Big_int (z_big_int (Zpos y)))
+ let z x = match x with Z0 -> 0 | Zpos p -> index p | Zneg p -> -index p
+ let q_to_num {qnum = x; qden = y} =
+ Big_int (z_big_int x) // Big_int (z_big_int (Zpos y))
end
-
(**
* MODULE: Caml to Coq data-structure mappings
*)
-module CamlToCoq =
-struct
- open Micromega
+module CamlToCoq = struct
+ open Micromega
- let rec nat = function
- | 0 -> O
- | n -> S (nat (n-1))
+ let rec nat = function 0 -> O | n -> S (nat (n - 1))
+ let rec positive n =
+ if Int.equal n 1 then XH
+ else if Int.equal (n land 1) 1 then XI (positive (n lsr 1))
+ else XO (positive (n lsr 1))
- let rec positive n =
- if Int.equal n 1 then XH
- else if Int.equal (n land 1) 1 then XI (positive (n lsr 1))
- else XO (positive (n lsr 1))
+ let n nt =
+ if nt < 0 then assert false
+ else if Int.equal nt 0 then N0
+ else Npos (positive nt)
- let n nt =
- if nt < 0
- then assert false
- else if Int.equal nt 0 then N0
- else Npos (positive nt)
+ let rec index n =
+ if Int.equal n 1 then XH
+ else if Int.equal (n land 1) 1 then XI (index (n lsr 1))
+ else XO (index (n lsr 1))
- let rec index n =
- if Int.equal n 1 then XH
- else if Int.equal (n land 1) 1 then XI (index (n lsr 1))
- else XO (index (n lsr 1))
-
-
- let z x =
- match compare x 0 with
- | 0 -> Z0
- | 1 -> Zpos (positive x)
- | _ -> (* this should be -1 *)
+ let z x =
+ match compare x 0 with
+ | 0 -> Z0
+ | 1 -> Zpos (positive x)
+ | _ ->
+ (* this should be -1 *)
Zneg (positive (-x))
- open Big_int
-
- let positive_big_int n =
- let two = big_int_of_int 2 in
- let rec _pos n =
- if eq_big_int n unit_big_int then XH
- else
- let (q,m) = quomod_big_int n two in
- if eq_big_int unit_big_int m
- then XI (_pos q)
- else XO (_pos q) in
- _pos n
-
- let bigint x =
- match sign_big_int x with
- | 0 -> Z0
- | 1 -> Zpos (positive_big_int x)
- | _ -> Zneg (positive_big_int (minus_big_int x))
-
- let q n =
- {Micromega.qnum = bigint (numerator n) ;
- Micromega.qden = positive_big_int (denominator n)}
-
+ open Big_int
+
+ let positive_big_int n =
+ let two = big_int_of_int 2 in
+ let rec _pos n =
+ if eq_big_int n unit_big_int then XH
+ else
+ let q, m = quomod_big_int n two in
+ if eq_big_int unit_big_int m then XI (_pos q) else XO (_pos q)
+ in
+ _pos n
+
+ let bigint x =
+ match sign_big_int x with
+ | 0 -> Z0
+ | 1 -> Zpos (positive_big_int x)
+ | _ -> Zneg (positive_big_int (minus_big_int x))
+
+ let q n =
+ { Micromega.qnum = bigint (numerator n)
+ ; Micromega.qden = positive_big_int (denominator n) }
end
(**
@@ -316,25 +268,22 @@ end
* between two lists given an ordering, and using a hash computation
*)
-module Cmp =
-struct
-
- let rec compare_lexical l =
- match l with
- | [] -> 0 (* Equal *)
- | f::l ->
+module Cmp = struct
+ let rec compare_lexical l =
+ match l with
+ | [] -> 0 (* Equal *)
+ | f :: l ->
let cmp = f () in
- if Int.equal cmp 0 then compare_lexical l else cmp
-
- let rec compare_list cmp l1 l2 =
- match l1 , l2 with
- | [] , [] -> 0
- | [] , _ -> -1
- | _ , [] -> 1
- | e1::l1 , e2::l2 ->
+ if Int.equal cmp 0 then compare_lexical l else cmp
+
+ let rec compare_list cmp l1 l2 =
+ match (l1, l2) with
+ | [], [] -> 0
+ | [], _ -> -1
+ | _, [] -> 1
+ | e1 :: l1, e2 :: l2 ->
let c = cmp e1 e2 in
- if Int.equal c 0 then compare_list cmp l1 l2 else c
-
+ if Int.equal c 0 then compare_list cmp l1 l2 else c
end
(**
@@ -344,9 +293,7 @@ end
* superfluous items, which speeds the translation up a bit.
*)
-module type Tag =
-sig
-
+module type Tag = sig
type t
val from : int -> t
@@ -354,12 +301,10 @@ sig
val pp : out_channel -> t -> unit
val compare : t -> t -> int
val max : t -> t -> t
- val to_int : t -> int
+ val to_int : t -> int
end
-module Tag : Tag =
-struct
-
+module Tag : Tag = struct
type t = int
let from i = i
@@ -368,14 +313,13 @@ struct
let pp o i = output_string o (string_of_int i)
let compare : int -> int -> int = Int.compare
let to_int x = x
-
end
(**
* MODULE: Ordered sets of tags.
*)
-module TagSet = Set.Make(Tag)
+module TagSet = Set.Make (Tag)
(** As for Unix.close_process, our Unix.waipid will ignore all EINTR *)
@@ -389,120 +333,100 @@ let rec waitpid_non_intr pid =
let command exe_path args vl =
(* creating pipes for stdin, stdout, stderr *)
- let (stdin_read,stdin_write) = Unix.pipe ()
- and (stdout_read,stdout_write) = Unix.pipe ()
- and (stderr_read,stderr_write) = Unix.pipe () in
-
+ let stdin_read, stdin_write = Unix.pipe ()
+ and stdout_read, stdout_write = Unix.pipe ()
+ and stderr_read, stderr_write = Unix.pipe () in
(* Create the process *)
- let pid = Unix.create_process exe_path args stdin_read stdout_write stderr_write in
-
+ let pid =
+ Unix.create_process exe_path args stdin_read stdout_write stderr_write
+ in
(* Write the data on the stdin of the created process *)
let outch = Unix.out_channel_of_descr stdin_write in
- output_value outch vl ;
- flush outch ;
-
+ output_value outch vl;
+ flush outch;
(* Wait for its completion *)
- let status = waitpid_non_intr pid in
-
- finally
- (* Recover the result *)
- (fun () ->
- match status with
- | Unix.WEXITED 0 ->
- let inch = Unix.in_channel_of_descr stdout_read in
- begin
- try Marshal.from_channel inch
- with any ->
- failwith
- (Printf.sprintf "command \"%s\" exited %s" exe_path
- (Printexc.to_string any))
- end
- | Unix.WEXITED i ->
- failwith (Printf.sprintf "command \"%s\" exited %i" exe_path i)
- | Unix.WSIGNALED i ->
- failwith (Printf.sprintf "command \"%s\" killed %i" exe_path i)
- | Unix.WSTOPPED i ->
- failwith (Printf.sprintf "command \"%s\" stopped %i" exe_path i))
- (* Cleanup *)
- (fun () ->
- List.iter (fun x -> try Unix.close x with any -> ())
- [stdin_read; stdin_write;
- stdout_read; stdout_write;
- stderr_read; stderr_write])
+ let status = waitpid_non_intr pid in
+ finally
+ (* Recover the result *)
+ (fun () ->
+ match status with
+ | Unix.WEXITED 0 -> (
+ let inch = Unix.in_channel_of_descr stdout_read in
+ try Marshal.from_channel inch
+ with any ->
+ failwith
+ (Printf.sprintf "command \"%s\" exited %s" exe_path
+ (Printexc.to_string any)) )
+ | Unix.WEXITED i ->
+ failwith (Printf.sprintf "command \"%s\" exited %i" exe_path i)
+ | Unix.WSIGNALED i ->
+ failwith (Printf.sprintf "command \"%s\" killed %i" exe_path i)
+ | Unix.WSTOPPED i ->
+ failwith (Printf.sprintf "command \"%s\" stopped %i" exe_path i))
+ (* Cleanup *)
+ (fun () ->
+ List.iter
+ (fun x -> try Unix.close x with any -> ())
+ [ stdin_read
+ ; stdin_write
+ ; stdout_read
+ ; stdout_write
+ ; stderr_read
+ ; stderr_write ])
(** Hashing utilities *)
-module Hash =
- struct
-
- module Mc = Micromega
-
- open Hashset.Combine
-
- let int_of_eq_op1 = Mc.(function
- | Equal -> 0
- | NonEqual -> 1
- | Strict -> 2
- | NonStrict -> 3)
-
- let eq_op1 o1 o2 = int_of_eq_op1 o1 = int_of_eq_op1 o2
-
- let hash_op1 h o = combine h (int_of_eq_op1 o)
-
-
- let rec eq_positive p1 p2 =
- match p1 , p2 with
- | Mc.XH , Mc.XH -> true
- | Mc.XI p1 , Mc.XI p2 -> eq_positive p1 p2
- | Mc.XO p1 , Mc.XO p2 -> eq_positive p1 p2
- | _ , _ -> false
-
- let eq_z z1 z2 =
- match z1 , z2 with
- | Mc.Z0 , Mc.Z0 -> true
- | Mc.Zpos p1, Mc.Zpos p2
- | Mc.Zneg p1, Mc.Zneg p2 -> eq_positive p1 p2
- | _ , _ -> false
-
- let eq_q {Mc.qnum = qn1 ; Mc.qden = qd1} {Mc.qnum = qn2 ; Mc.qden = qd2} =
- eq_z qn1 qn2 && eq_positive qd1 qd2
-
- let rec eq_pol eq p1 p2 =
- match p1 , p2 with
- | Mc.Pc c1 , Mc.Pc c2 -> eq c1 c2
- | Mc.Pinj(i1,p1) , Mc.Pinj(i2,p2) -> eq_positive i1 i2 && eq_pol eq p1 p2
- | Mc.PX(p1,i1,p1') , Mc.PX(p2,i2,p2') ->
- eq_pol eq p1 p2 && eq_positive i1 i2 && eq_pol eq p1' p2'
- | _ , _ -> false
-
-
- let eq_pair eq1 eq2 (x1,y1) (x2,y2) =
- eq1 x1 x2 && eq2 y1 y2
-
-
- let hash_pol helt =
- let rec hash acc = function
- | Mc.Pc c -> helt (combine acc 1) c
- | Mc.Pinj(p,c) -> hash (combine (combine acc 1) (CoqToCaml.index p)) c
- | Mc.PX(p1,i,p2) -> hash (hash (combine (combine acc 2) (CoqToCaml.index i)) p1) p2 in
- hash
-
-
- let hash_pair h1 h2 h (e1,e2) =
- h2 (h1 h e1) e2
-
- let hash_elt f h e = combine h (f e)
-
- let hash_string h (e:string) = hash_elt Hashtbl.hash h e
-
- let hash_z = hash_elt CoqToCaml.z
-
- let hash_q = hash_elt (fun q -> Hashtbl.hash (CoqToCaml.q_to_num q))
-
- end
-
-
-
+module Hash = struct
+ module Mc = Micromega
+ open Hashset.Combine
+
+ let int_of_eq_op1 =
+ Mc.(function Equal -> 0 | NonEqual -> 1 | Strict -> 2 | NonStrict -> 3)
+
+ let eq_op1 o1 o2 = int_of_eq_op1 o1 = int_of_eq_op1 o2
+ let hash_op1 h o = combine h (int_of_eq_op1 o)
+
+ let rec eq_positive p1 p2 =
+ match (p1, p2) with
+ | Mc.XH, Mc.XH -> true
+ | Mc.XI p1, Mc.XI p2 -> eq_positive p1 p2
+ | Mc.XO p1, Mc.XO p2 -> eq_positive p1 p2
+ | _, _ -> false
+
+ let eq_z z1 z2 =
+ match (z1, z2) with
+ | Mc.Z0, Mc.Z0 -> true
+ | Mc.Zpos p1, Mc.Zpos p2 | Mc.Zneg p1, Mc.Zneg p2 -> eq_positive p1 p2
+ | _, _ -> false
+
+ let eq_q {Mc.qnum = qn1; Mc.qden = qd1} {Mc.qnum = qn2; Mc.qden = qd2} =
+ eq_z qn1 qn2 && eq_positive qd1 qd2
+
+ let rec eq_pol eq p1 p2 =
+ match (p1, p2) with
+ | Mc.Pc c1, Mc.Pc c2 -> eq c1 c2
+ | Mc.Pinj (i1, p1), Mc.Pinj (i2, p2) -> eq_positive i1 i2 && eq_pol eq p1 p2
+ | Mc.PX (p1, i1, p1'), Mc.PX (p2, i2, p2') ->
+ eq_pol eq p1 p2 && eq_positive i1 i2 && eq_pol eq p1' p2'
+ | _, _ -> false
+
+ let eq_pair eq1 eq2 (x1, y1) (x2, y2) = eq1 x1 x2 && eq2 y1 y2
+
+ let hash_pol helt =
+ let rec hash acc = function
+ | Mc.Pc c -> helt (combine acc 1) c
+ | Mc.Pinj (p, c) -> hash (combine (combine acc 1) (CoqToCaml.index p)) c
+ | Mc.PX (p1, i, p2) ->
+ hash (hash (combine (combine acc 2) (CoqToCaml.index i)) p1) p2
+ in
+ hash
+
+ let hash_pair h1 h2 h (e1, e2) = h2 (h1 h e1) e2
+ let hash_elt f h e = combine h (f e)
+ let hash_string h (e : string) = hash_elt Hashtbl.hash h e
+ let hash_z = hash_elt CoqToCaml.z
+ let hash_q = hash_elt (fun q -> Hashtbl.hash (CoqToCaml.q_to_num q))
+end
(* Local Variables: *)
(* coding: utf-8 *)
diff --git a/plugins/micromega/mutils.mli b/plugins/micromega/mutils.mli
index 9692bc631b..ef8d154b13 100644
--- a/plugins/micromega/mutils.mli
+++ b/plugins/micromega/mutils.mli
@@ -8,51 +8,50 @@
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
-module Int : sig type t = int val compare : int -> int -> int val equal : int -> int -> bool end
+module Int : sig
+ type t = int
+ val compare : int -> int -> int
+ val equal : int -> int -> bool
+end
module ISet : sig
include Set.S with type elt = int
+
val pp : out_channel -> t -> unit
end
-module IMap :
-sig
+module IMap : sig
include Map.S with type key = int
- (** [from k m] returns the submap of [m] with keys greater or equal k *)
val from : key -> 'elt t -> 'elt t
-
+ (** [from k m] returns the submap of [m] with keys greater or equal k *)
end
val numerator : Num.num -> Big_int.big_int
val denominator : Num.num -> Big_int.big_int
module Cmp : sig
-
val compare_list : ('a -> 'b -> int) -> 'a list -> 'b list -> int
val compare_lexical : (unit -> int) list -> int
-
end
module Tag : sig
-
type t
val pp : out_channel -> t -> unit
val next : t -> t
- val max : t -> t -> t
+ val max : t -> t -> t
val from : int -> t
val to_int : t -> int
-
end
module TagSet : CSig.SetS with type elt = Tag.t
-val pp_list : string -> (out_channel -> 'a -> unit) -> out_channel -> 'a list -> unit
+val pp_list :
+ string -> (out_channel -> 'a -> unit) -> out_channel -> 'a list -> unit
module CamlToCoq : sig
-
val positive : int -> Micromega.positive
val bigint : Big_int.big_int -> Micromega.z
val n : int -> Micromega.n
@@ -61,74 +60,62 @@ module CamlToCoq : sig
val index : int -> Micromega.positive
val z : int -> Micromega.z
val positive_big_int : Big_int.big_int -> Micromega.positive
-
end
module CoqToCaml : sig
-
val z_big_int : Micromega.z -> Big_int.big_int
- val z : Micromega.z -> int
- val q_to_num : Micromega.q -> Num.num
- val positive : Micromega.positive -> int
- val n : Micromega.n -> int
- val nat : Micromega.nat -> int
- val index : Micromega.positive -> int
-
+ val z : Micromega.z -> int
+ val q_to_num : Micromega.q -> Num.num
+ val positive : Micromega.positive -> int
+ val n : Micromega.n -> int
+ val nat : Micromega.nat -> int
+ val index : Micromega.positive -> int
end
module Hash : sig
-
- val eq_op1 : Micromega.op1 -> Micromega.op1 -> bool
-
+ val eq_op1 : Micromega.op1 -> Micromega.op1 -> bool
val eq_positive : Micromega.positive -> Micromega.positive -> bool
-
val eq_z : Micromega.z -> Micromega.z -> bool
-
val eq_q : Micromega.q -> Micromega.q -> bool
- val eq_pol : ('a -> 'a -> bool) -> 'a Micromega.pol -> 'a Micromega.pol -> bool
+ val eq_pol :
+ ('a -> 'a -> bool) -> 'a Micromega.pol -> 'a Micromega.pol -> bool
- val eq_pair : ('a -> 'a -> bool) -> ('b -> 'b -> bool) -> 'a * 'b -> 'a * 'b -> bool
+ val eq_pair :
+ ('a -> 'a -> bool) -> ('b -> 'b -> bool) -> 'a * 'b -> 'a * 'b -> bool
val hash_op1 : int -> Micromega.op1 -> int
+ val hash_pol : (int -> 'a -> int) -> int -> 'a Micromega.pol -> int
- val hash_pol : (int -> 'a -> int) -> int -> 'a Micromega.pol -> int
-
- val hash_pair : (int -> 'a -> int) -> (int -> 'b -> int) -> int -> 'a * 'b -> int
-
- val hash_z : int -> Micromega.z -> int
-
- val hash_q : int -> Micromega.q -> int
+ val hash_pair :
+ (int -> 'a -> int) -> (int -> 'b -> int) -> int -> 'a * 'b -> int
+ val hash_z : int -> Micromega.z -> int
+ val hash_q : int -> Micromega.q -> int
val hash_string : int -> string -> int
-
val hash_elt : ('a -> int) -> int -> 'a -> int
-
end
-
val ppcm : Big_int.big_int -> Big_int.big_int -> Big_int.big_int
-
val all_pairs : ('a -> 'a -> 'b) -> 'a list -> 'b list
val try_any : (('a -> 'b option) * 'c) list -> 'a -> 'b option
val is_sublist : ('a -> 'b -> bool) -> 'a list -> 'b list -> bool
-
val extract : ('a -> 'b option) -> 'a list -> ('b * 'a) option * 'a list
-
val extract_all : ('a -> 'b option) -> 'a list -> ('b * 'a) list * 'a list
-val extract_best : ('a -> 'b option) -> ('b -> 'b -> bool) -> 'a list -> ('b *'a) option * 'a list
-
-val find_some : ('a -> 'b option) -> 'a list -> 'b option
+val extract_best :
+ ('a -> 'b option)
+ -> ('b -> 'b -> bool)
+ -> 'a list
+ -> ('b * 'a) option * 'a list
+val find_some : ('a -> 'b option) -> 'a list -> 'b option
val iterate_until_stable : ('a -> 'a option) -> 'a -> 'a
-
val simplify : ('a -> 'a option) -> 'a list -> 'a list option
-val saturate : ('a -> 'b option) -> ('b * 'a -> 'a -> 'a option) -> 'a list -> 'a list
+val saturate :
+ ('a -> 'b option) -> ('b * 'a -> 'a -> 'a option) -> 'a list -> 'a list
val generate : ('a -> 'b option) -> 'a list -> 'b list
-
val app_funs : ('a -> 'b option) list -> 'a -> 'b option
-
val command : string -> string array -> 'a -> 'b
diff --git a/plugins/micromega/persistent_cache.ml b/plugins/micromega/persistent_cache.ml
index 28d8d5a020..d5b28cb03e 100644
--- a/plugins/micromega/persistent_cache.ml
+++ b/plugins/micromega/persistent_cache.ml
@@ -14,207 +14,158 @@
(* *)
(************************************************************************)
-module type PHashtable =
- sig
- (* see documentation in [persistent_cache.mli] *)
- type 'a t
- type key
-
- val open_in : string -> 'a t
-
- val find : 'a t -> key -> 'a
-
- val add : 'a t -> key -> 'a -> unit
-
- val memo : string -> (key -> 'a) -> (key -> 'a)
-
- val memo_cond : string -> (key -> bool) -> (key -> 'a) -> (key -> 'a)
-
- end
+module type PHashtable = sig
+ (* see documentation in [persistent_cache.mli] *)
+ type 'a t
+ type key
+
+ val open_in : string -> 'a t
+ val find : 'a t -> key -> 'a
+ val add : 'a t -> key -> 'a -> unit
+ val memo : string -> (key -> 'a) -> key -> 'a
+ val memo_cond : string -> (key -> bool) -> (key -> 'a) -> key -> 'a
+end
open Hashtbl
-module PHashtable(Key:HashedType) : PHashtable with type key = Key.t =
-struct
+module PHashtable (Key : HashedType) : PHashtable with type key = Key.t = struct
open Unix
type key = Key.t
- module Table = Hashtbl.Make(Key)
+ module Table = Hashtbl.Make (Key)
exception InvalidTableFormat
exception UnboundTable
type mode = Closed | Open
+ type 'a t = {outch : out_channel; mutable status : mode; htbl : 'a Table.t}
- type 'a t =
- {
- outch : out_channel ;
- mutable status : mode ;
- htbl : 'a Table.t
- }
-
-
-let finally f rst =
- try
- let res = f () in
- rst () ; res
- with reraise ->
- (try rst ()
- with any -> raise reraise
- ); raise reraise
-
-
-let read_key_elem inch =
- try
- Some (Marshal.from_channel inch)
- with
+ let finally f rst =
+ try
+ let res = f () in
+ rst (); res
+ with reraise ->
+ (try rst () with any -> raise reraise);
+ raise reraise
+
+ let read_key_elem inch =
+ try Some (Marshal.from_channel inch) with
| End_of_file -> None
| e when CErrors.noncritical e -> raise InvalidTableFormat
-(**
+ (**
We used to only lock/unlock regions.
Is-it more robust/portable to lock/unlock a fixed region e.g. [0;1]?
In case of locking failure, the cache is not used.
**)
-type lock_kind = Read | Write
-
-let lock kd fd =
- let pos = lseek fd 0 SEEK_CUR in
- let success =
- try
- ignore (lseek fd 0 SEEK_SET);
- let lk = match kd with
- | Read -> F_RLOCK
- | Write -> F_LOCK in
- lockf fd lk 1; true
- with Unix.Unix_error(_,_,_) -> false in
- ignore (lseek fd pos SEEK_SET) ;
- success
-
-let unlock fd =
- let pos = lseek fd 0 SEEK_CUR in
- try
- ignore (lseek fd 0 SEEK_SET) ;
- lockf fd F_ULOCK 1
- with
- Unix.Unix_error(_,_,_) -> ()
- (* Here, this is really bad news --
+ type lock_kind = Read | Write
+
+ let lock kd fd =
+ let pos = lseek fd 0 SEEK_CUR in
+ let success =
+ try
+ ignore (lseek fd 0 SEEK_SET);
+ let lk = match kd with Read -> F_RLOCK | Write -> F_LOCK in
+ lockf fd lk 1; true
+ with Unix.Unix_error (_, _, _) -> false
+ in
+ ignore (lseek fd pos SEEK_SET);
+ success
+
+ let unlock fd =
+ let pos = lseek fd 0 SEEK_CUR in
+ try
+ ignore (lseek fd 0 SEEK_SET);
+ lockf fd F_ULOCK 1
+ with Unix.Unix_error (_, _, _) ->
+ ()
+ (* Here, this is really bad news --
there is a pending lock which could cause a deadlock.
Should it be an anomaly or produce a warning ?
*);
- ignore (lseek fd pos SEEK_SET)
-
-
-(* We make the assumption that an acquired lock can always be released *)
+ ignore (lseek fd pos SEEK_SET)
-let do_under_lock kd fd f =
- if lock kd fd
- then
- finally f (fun () -> unlock fd)
- else f ()
+ (* We make the assumption that an acquired lock can always be released *)
+ let do_under_lock kd fd f =
+ if lock kd fd then finally f (fun () -> unlock fd) else f ()
-
-let open_in f =
- let flags = [O_RDONLY ; O_CREAT] in
- let finch = openfile f flags 0o666 in
- let inch = in_channel_of_descr finch in
- let htbl = Table.create 100 in
-
- let rec xload () =
- match read_key_elem inch with
+ let open_in f =
+ let flags = [O_RDONLY; O_CREAT] in
+ let finch = openfile f flags 0o666 in
+ let inch = in_channel_of_descr finch in
+ let htbl = Table.create 100 in
+ let rec xload () =
+ match read_key_elem inch with
| None -> ()
- | Some (key,elem) ->
- Table.add htbl key elem ;
- xload () in
+ | Some (key, elem) -> Table.add htbl key elem; xload ()
+ in
try
(* Locking of the (whole) file while reading *)
- do_under_lock Read finch xload ;
- close_in_noerr inch ;
- {
- outch = out_channel_of_descr (openfile f [O_WRONLY;O_APPEND;O_CREAT] 0o666) ;
- status = Open ;
- htbl = htbl
- }
+ do_under_lock Read finch xload;
+ close_in_noerr inch;
+ { outch =
+ out_channel_of_descr (openfile f [O_WRONLY; O_APPEND; O_CREAT] 0o666)
+ ; status = Open
+ ; htbl }
with InvalidTableFormat ->
- (* The file is corrupted *)
- begin
- close_in_noerr inch ;
- let flags = [O_WRONLY; O_TRUNC;O_CREAT] in
- let out = (openfile f flags 0o666) in
+ (* The file is corrupted *)
+ close_in_noerr inch;
+ let flags = [O_WRONLY; O_TRUNC; O_CREAT] in
+ let out = openfile f flags 0o666 in
let outch = out_channel_of_descr out in
- do_under_lock Write out
- (fun () ->
- Table.iter
- (fun k e -> Marshal.to_channel outch (k,e) [Marshal.No_sharing]) htbl;
- flush outch) ;
- { outch = outch ;
- status = Open ;
- htbl = htbl
- }
- end
-
-
-let add t k e =
- let {outch = outch ; status = status ; htbl = tbl} = t in
- if status == Closed
- then raise UnboundTable
+ do_under_lock Write out (fun () ->
+ Table.iter
+ (fun k e -> Marshal.to_channel outch (k, e) [Marshal.No_sharing])
+ htbl;
+ flush outch);
+ {outch; status = Open; htbl}
+
+ let add t k e =
+ let {outch; status; htbl = tbl} = t in
+ if status == Closed then raise UnboundTable
else
let fd = descr_of_out_channel outch in
- begin
- Table.add tbl k e ;
- do_under_lock Write fd
- (fun _ ->
- Marshal.to_channel outch (k,e) [Marshal.No_sharing] ;
- flush outch
- )
- end
-
-let find t k =
- let {outch = outch ; status = status ; htbl = tbl} = t in
- if status == Closed
- then raise UnboundTable
+ Table.add tbl k e;
+ do_under_lock Write fd (fun _ ->
+ Marshal.to_channel outch (k, e) [Marshal.No_sharing];
+ flush outch)
+
+ let find t k =
+ let {outch; status; htbl = tbl} = t in
+ if status == Closed then raise UnboundTable
else
let res = Table.find tbl k in
- res
-
-let memo cache f =
- 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
- Not_found ->
- let res = f x in
- add tbl x res ;
- res
-
-let memo_cond cache cond f =
- let tbl = lazy (try Some (open_in cache) with _ -> None) in
- fun x ->
- match Lazy.force tbl with
- | None -> f x
- | Some tbl ->
- if cond x
- then
- begin
- try find tbl x
- with Not_found ->
- let res = f x in
- add tbl x res ;
- res
- end
- else f x
-
-
+ res
+
+ let memo cache f =
+ 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 Not_found ->
+ let res = f x in
+ add tbl x res; res )
+
+ let memo_cond cache cond f =
+ let tbl = lazy (try Some (open_in cache) with _ -> None) in
+ fun x ->
+ match Lazy.force tbl with
+ | None -> f x
+ | Some tbl ->
+ if cond x then begin
+ try find tbl x
+ with Not_found ->
+ let res = f x in
+ add tbl x res; res
+ end
+ else f x
end
-
(* Local Variables: *)
(* coding: utf-8 *)
(* End: *)
diff --git a/plugins/micromega/persistent_cache.mli b/plugins/micromega/persistent_cache.mli
index cb14d73972..7d459a66e7 100644
--- a/plugins/micromega/persistent_cache.mli
+++ b/plugins/micromega/persistent_cache.mli
@@ -10,32 +10,29 @@
open Hashtbl
-module type PHashtable =
- sig
- type 'a t
- type key
+module type PHashtable = sig
+ type 'a t
+ type key
- val open_in : string -> 'a t
- (** [open_in f] rebuilds a table from the records stored in file [f].
+ val open_in : string -> 'a t
+ (** [open_in f] rebuilds a table from the records stored in file [f].
As marshaling is not type-safe, it might segfault.
*)
- val find : 'a t -> key -> 'a
- (** find has the specification of Hashtable.find *)
+ val find : 'a t -> key -> 'a
+ (** find has the specification of Hashtable.find *)
- val add : 'a t -> key -> 'a -> unit
- (** [add tbl key elem] adds the binding [key] [elem] to the table [tbl].
+ val add : 'a t -> key -> 'a -> unit
+ (** [add tbl key elem] adds the binding [key] [elem] to the table [tbl].
(and writes the binding to the file associated with [tbl].)
If [key] is already bound, raises KeyAlreadyBound *)
- val memo : string -> (key -> 'a) -> (key -> 'a)
- (** [memo cache f] returns a memo function for [f] using file [cache] as persistent table.
+ val memo : string -> (key -> 'a) -> key -> 'a
+ (** [memo cache f] returns a memo function for [f] using file [cache] as persistent table.
Note that the cache will only be loaded when the function is used for the first time *)
- val memo_cond : string -> (key -> bool) -> (key -> 'a) -> (key -> 'a)
- (** [memo cache cond f] only use the cache if [cond k] holds for the key [k]. *)
+ val memo_cond : string -> (key -> bool) -> (key -> 'a) -> key -> 'a
+ (** [memo cache cond f] only use the cache if [cond k] holds for the key [k]. *)
+end
-
- end
-
-module PHashtable(Key:HashedType) : PHashtable with type key = Key.t
+module PHashtable (Key : HashedType) : PHashtable with type key = Key.t
diff --git a/plugins/micromega/polynomial.ml b/plugins/micromega/polynomial.ml
index 1a31a36732..adf54713b9 100644
--- a/plugins/micromega/polynomial.ml
+++ b/plugins/micromega/polynomial.ml
@@ -17,7 +17,6 @@
open Num
module Utils = Mutils
open Utils
-
module Mc = Micromega
let max_nb_cstr = ref max_int
@@ -25,165 +24,150 @@ let max_nb_cstr = ref max_int
type var = int
let debug = false
+let ( <+> ) = add_num
+let ( <*> ) = mult_num
-let (<+>) = add_num
-let (<*>) = mult_num
-
-module Monomial :
-sig
+module Monomial : sig
type t
+
val const : t
val is_const : t -> bool
val var : var -> t
val is_var : t -> bool
val get_var : t -> var option
val prod : t -> t -> t
- val exp : t -> int -> t
- val div : t -> t -> t * int
+ val exp : t -> int -> t
+ val div : t -> t -> t * int
val compare : t -> t -> int
val pp : out_channel -> t -> unit
val fold : (var -> int -> 'a -> 'a) -> t -> 'a -> 'a
val sqrt : t -> t option
val variables : t -> ISet.t
-end
- = struct
+end = struct
(* A monomial is represented by a multiset of variables *)
- module Map = Map.Make(Int)
+ module Map = Map.Make (Int)
open Map
type t = int Map.t
let is_singleton m =
try
- let (k,v) = choose m in
- let (l,e,r) = split k m in
- if is_empty l && is_empty r
- then Some(k,v) else None
+ let k, v = choose m in
+ let l, e, r = split k m in
+ if is_empty l && is_empty r then Some (k, v) else None
with Not_found -> None
let pp o m =
- let pp_elt o (k,v)=
- if v = 1 then Printf.fprintf o "x%i" k
- else Printf.fprintf o "x%i^%i" k v in
-
+ let pp_elt o (k, v) =
+ if v = 1 then Printf.fprintf o "x%i" k else Printf.fprintf o "x%i^%i" k v
+ in
let rec pp_list o l =
match l with
- [] -> ()
+ | [] -> ()
| [e] -> pp_elt o e
- | e::l -> Printf.fprintf o "%a*%a" pp_elt e pp_list l in
-
- pp_list o (Map.bindings m)
-
-
+ | e :: l -> Printf.fprintf o "%a*%a" pp_elt e pp_list l
+ in
+ pp_list o (Map.bindings m)
(* The monomial that corresponds to a constant *)
let const = Map.empty
-
let sum_degree m = Map.fold (fun _ n s -> s + n) m 0
(* Total ordering of monomials *)
- let compare: t -> t -> int =
- fun m1 m2 ->
- let s1 = sum_degree m1
- and s2 = sum_degree m2 in
- if Int.equal s1 s2 then Map.compare Int.compare m1 m2
- else Int.compare s1 s2
+ let compare : t -> t -> int =
+ fun m1 m2 ->
+ let s1 = sum_degree m1 and s2 = sum_degree m2 in
+ if Int.equal s1 s2 then Map.compare Int.compare m1 m2 else Int.compare s1 s2
- let is_const m = (m = Map.empty)
+ let is_const m = m = Map.empty
(* The monomial 'x' *)
let var x = Map.add x 1 Map.empty
let is_var m =
- match is_singleton m with
- | None -> false
- | Some (_,i) -> i = 1
+ match is_singleton m with None -> false | Some (_, i) -> i = 1
let get_var m =
match is_singleton m with
| None -> None
- | Some (k,i) -> if i = 1 then Some k else None
-
+ | Some (k, i) -> if i = 1 then Some k else None
let sqrt m =
if is_const m then None
else
try
- Some (Map.fold (fun v i acc ->
- let i' = i / 2 in
- if i mod 2 = 0
- then add v i' acc
- else raise Not_found) m const)
+ Some
+ (Map.fold
+ (fun v i acc ->
+ let i' = i / 2 in
+ if i mod 2 = 0 then add v i' acc else raise Not_found)
+ m const)
with Not_found -> None
-
(* Get the degre of a variable in a monomial *)
let find x m = try find x m with Not_found -> 0
(* Product of monomials *)
- let prod m1 m2 = Map.fold (fun k d m -> add k ((find k m) + d) m) m1 m2
+ let prod m1 m2 = Map.fold (fun k d m -> add k (find k m + d) m) m1 m2
let exp m n =
- let rec exp acc n =
- if n = 0 then acc
- else exp (prod acc m) (n - 1) in
-
+ let rec exp acc n = if n = 0 then acc else exp (prod acc m) (n - 1) in
exp const n
(* [div m1 m2 = mr,n] such that mr * (m2)^n = m1 *)
let div m1 m2 =
- let n = fold (fun x i n -> let i' = find x m1 in
- let nx = i' / i in
- min n nx) m2 max_int in
-
- let mr = fold (fun x i' m ->
- let i = find x m2 in
- let ir = i' - i * n in
- if ir = 0 then m
- else add x ir m) m1 empty in
- (mr,n)
-
+ let n =
+ fold
+ (fun x i n ->
+ let i' = find x m1 in
+ let nx = i' / i in
+ min n nx)
+ m2 max_int
+ in
+ let mr =
+ fold
+ (fun x i' m ->
+ let i = find x m2 in
+ let ir = i' - (i * n) in
+ if ir = 0 then m else add x ir m)
+ m1 empty
+ in
+ (mr, n)
let variables m = fold (fun v i acc -> ISet.add v acc) m ISet.empty
-
let fold = fold
-
end
-module MonMap =
- struct
- include Map.Make(Monomial)
+module MonMap = struct
+ include Map.Make (Monomial)
- let union f = merge
- (fun x v1 v2 ->
- match v1 , v2 with
- | None , None -> None
- | Some v , None | None , Some v -> Some v
- | Some v1 , Some v2 -> f x v1 v2)
- end
+ let union f =
+ merge (fun x v1 v2 ->
+ match (v1, v2) with
+ | None, None -> None
+ | Some v, None | None, Some v -> Some v
+ | Some v1, Some v2 -> f x v1 v2)
+end
let pp_mon o (m, i) =
- if Monomial.is_const m
- then if eq_num (Int 0) i then ()
- else Printf.fprintf o "%s" (string_of_num i)
+ if Monomial.is_const m then
+ if eq_num (Int 0) i then () else Printf.fprintf o "%s" (string_of_num i)
else
match i with
- | Int 1 -> Monomial.pp o m
+ | Int 1 -> Monomial.pp o m
| Int -1 -> Printf.fprintf o "-%a" Monomial.pp m
- | Int 0 -> ()
- | _ -> Printf.fprintf o "%s*%a" (string_of_num i) Monomial.pp m
-
+ | Int 0 -> ()
+ | _ -> Printf.fprintf o "%s*%a" (string_of_num i) Monomial.pp m
-
-module Poly :
-(* A polynomial is a map of monomials *)
-(*
+module Poly : (* A polynomial is a map of monomials *)
+ (*
This is probably a naive implementation
(expected to be fast enough - Coq is probably the bottleneck)
*The new ring contribution is using a sparse Horner representation.
*)
sig
type t
+
val pp : out_channel -> t -> unit
val get : Monomial.t -> t -> num
val variable : var -> t
@@ -193,42 +177,34 @@ sig
val addition : t -> t -> t
val uminus : t -> t
val fold : (Monomial.t -> num -> 'a -> 'a) -> t -> 'a -> 'a
- val factorise : var -> t -> t * t
-end = struct
+ val factorise : var -> t -> t * t
+end = struct
(*normalisation bug : 0*x ... *)
- module P = Map.Make(Monomial)
+ module P = Map.Make (Monomial)
open P
type t = num P.t
-
- let pp o p = P.iter (fun mn i -> Printf.fprintf o "%a + " pp_mon (mn, i)) p
-
+ let pp o p = P.iter (fun mn i -> Printf.fprintf o "%a + " pp_mon (mn, i)) p
(* Get the coefficient of monomial mn *)
let get : Monomial.t -> t -> num =
- fun mn p -> try find mn p with Not_found -> (Int 0)
-
+ fun mn p -> try find mn p with Not_found -> Int 0
(* The polynomial 1.x *)
- let variable : var -> t =
- fun x -> add (Monomial.var x) (Int 1) empty
+ let variable : var -> t = fun x -> add (Monomial.var x) (Int 1) empty
(*The constant polynomial *)
- let constant : num -> t =
- fun c -> add (Monomial.const) c empty
+ let constant : num -> t = fun c -> add Monomial.const c empty
(* The addition of a monomial *)
let add : Monomial.t -> num -> t -> t =
- fun mn v p ->
+ fun mn v p ->
if sign_num v = 0 then p
else
- let vl = (get mn p) <+> v in
- if sign_num vl = 0 then
- remove mn p
- else add mn vl p
-
+ let vl = get mn p <+> v in
+ if sign_num vl = 0 then remove mn p else add mn vl p
(** Design choice: empty is not a polynomial
I do not remember why ....
@@ -236,76 +212,56 @@ end = struct
(* The product by a monomial *)
let mult : Monomial.t -> num -> t -> t =
- fun mn v p ->
- if sign_num v = 0
- then constant (Int 0)
+ fun mn v p ->
+ if sign_num v = 0 then constant (Int 0)
else
- fold (fun mn' v' res -> P.add (Monomial.prod mn mn') (v<*>v') res) p empty
-
-
- let addition : t -> t -> t =
- fun p1 p2 -> fold (fun mn v p -> add mn v p) p1 p2
+ fold
+ (fun mn' v' res -> P.add (Monomial.prod mn mn') (v <*> v') res)
+ p empty
+ let addition : t -> t -> t =
+ fun p1 p2 -> fold (fun mn v p -> add mn v p) p1 p2
let product : t -> t -> t =
- fun p1 p2 ->
- fold (fun mn v res -> addition (mult mn v p2) res ) p1 empty
-
-
- let uminus : t -> t =
- fun p -> map (fun v -> minus_num v) p
+ fun p1 p2 -> fold (fun mn v res -> addition (mult mn v p2) res) p1 empty
+ let uminus : t -> t = fun p -> map (fun v -> minus_num v) p
let fold = P.fold
let factorise x p =
let x = Monomial.var x in
- P.fold (fun m v (px,cx) ->
- let (m1,i) = Monomial.div m x in
- if i = 0
- then (px, add m v cx)
+ P.fold
+ (fun m v (px, cx) ->
+ let m1, i = Monomial.div m x in
+ if i = 0 then (px, add m v cx)
else
- let mx = Monomial.prod m1 (Monomial.exp x (i-1)) in
- (add mx v px,cx) ) p (constant (Int 0) , constant (Int 0))
-
+ let mx = Monomial.prod m1 (Monomial.exp x (i - 1)) in
+ (add mx v px, cx))
+ p
+ (constant (Int 0), constant (Int 0))
end
-
-
type vector = Vect.t
-type cstr = {coeffs : vector ; op : op ; cst : num}
-and op = |Eq | Ge | Gt
-
-exception Strict
+type cstr = {coeffs : vector; op : op; cst : num}
-let is_strict c = (=) c.op Gt
-
-let eval_op = function
- | Eq -> (=/)
- | Ge -> (>=/)
- | Gt -> (>/)
+and op = Eq | Ge | Gt
+exception Strict
+let is_strict c = c.op = Gt
+let eval_op = function Eq -> ( =/ ) | Ge -> ( >=/ ) | Gt -> ( >/ )
let string_of_op = function Eq -> "=" | Ge -> ">=" | Gt -> ">"
-let output_cstr o { coeffs ; op ; cst } =
- Printf.fprintf o "%a %s %s" Vect.pp coeffs (string_of_op op) (string_of_num cst)
-
+let output_cstr o {coeffs; op; cst} =
+ Printf.fprintf o "%a %s %s" Vect.pp coeffs (string_of_op op)
+ (string_of_num cst)
let opMult o1 o2 =
- match o1, o2 with
- | Eq , _ | _ , Eq -> Eq
- | Ge , _ | _ , Ge -> Ge
- | Gt , Gt -> Gt
+ match (o1, o2) with Eq, _ | _, Eq -> Eq | Ge, _ | _, Ge -> Ge | Gt, Gt -> Gt
let opAdd o1 o2 =
- match o1, o2 with
- | Eq , x | x , Eq -> x
- | Gt , x | x , Gt -> Gt
- | Ge , Ge -> Ge
-
-
-
+ match (o1, o2) with Eq, x | x, Eq -> x | Gt, x | x, Gt -> Gt | Ge, Ge -> Ge
module LinPoly = struct
(** A linear polynomial a0 + a1.x1 + ... + an.xn
@@ -314,36 +270,32 @@ module LinPoly = struct
type t = Vect.t
- module MonT = struct
- module MonoMap = Map.Make(Monomial)
- module IntMap = Map.Make(Int)
+ module MonT = struct
+ module MonoMap = Map.Make (Monomial)
+ module IntMap = Map.Make (Int)
(** A hash table might be preferable but requires a hash function. *)
- let (index_of_monomial : int MonoMap.t ref) = ref (MonoMap.empty)
- let (monomial_of_index : Monomial.t IntMap.t ref) = ref (IntMap.empty)
+ let (index_of_monomial : int MonoMap.t ref) = ref MonoMap.empty
+
+ let (monomial_of_index : Monomial.t IntMap.t ref) = ref IntMap.empty
let fresh = ref 0
let clear () =
index_of_monomial := MonoMap.empty;
- monomial_of_index := IntMap.empty ;
+ monomial_of_index := IntMap.empty;
fresh := 0
-
let register m =
- try
- MonoMap.find m !index_of_monomial
+ try MonoMap.find m !index_of_monomial
with Not_found ->
- begin
- let res = !fresh in
- index_of_monomial := MonoMap.add m res !index_of_monomial ;
- monomial_of_index := IntMap.add res m !monomial_of_index ;
- incr fresh ; res
- end
+ let res = !fresh in
+ index_of_monomial := MonoMap.add m res !index_of_monomial;
+ monomial_of_index := IntMap.add res m !monomial_of_index;
+ incr fresh;
+ res
let retrieve i = IntMap.find i !monomial_of_index
-
let _ = register Monomial.const
-
end
let var v = Vect.set (MonT.register (Monomial.var v)) (Int 1) Vect.null
@@ -353,126 +305,122 @@ module LinPoly = struct
Vect.set v (Int 1) Vect.null
let linpol_of_pol p =
- Poly.fold
- (fun mon num vct ->
- let vr = MonT.register mon in
- Vect.set vr num vct) p Vect.null
+ Poly.fold
+ (fun mon num vct ->
+ let vr = MonT.register mon in
+ Vect.set vr num vct)
+ p Vect.null
let pol_of_linpol v =
- Vect.fold (fun p vr n -> Poly.add (MonT.retrieve vr) n p) (Poly.constant (Int 0)) v
-
- let coq_poly_of_linpol cst p =
+ Vect.fold
+ (fun p vr n -> Poly.add (MonT.retrieve vr) n p)
+ (Poly.constant (Int 0)) v
+ let coq_poly_of_linpol cst p =
let pol_of_mon m =
- Monomial.fold (fun x v p -> Mc.PEmul(Mc.PEpow(Mc.PEX(CamlToCoq.positive x),CamlToCoq.n v),p)) m (Mc.PEc (cst (Int 1))) in
-
- Vect.fold (fun acc x v ->
+ Monomial.fold
+ (fun x v p ->
+ Mc.PEmul (Mc.PEpow (Mc.PEX (CamlToCoq.positive x), CamlToCoq.n v), p))
+ m
+ (Mc.PEc (cst (Int 1)))
+ in
+ Vect.fold
+ (fun acc x v ->
let mn = MonT.retrieve x in
- Mc.PEadd(Mc.PEmul(Mc.PEc (cst v), pol_of_mon mn),acc)) (Mc.PEc (cst (Int 0))) p
+ Mc.PEadd (Mc.PEmul (Mc.PEc (cst v), pol_of_mon mn), acc))
+ (Mc.PEc (cst (Int 0)))
+ p
let pp_var o vr =
- try
- Monomial.pp o (MonT.retrieve vr) (* this is a non-linear variable *)
- with Not_found -> Printf.fprintf o "v%i" vr
-
-
- let pp o p = Vect.pp_gen pp_var o p
-
-
- let constant c =
- if sign_num c = 0
- then Vect.null
- else Vect.set 0 c Vect.null
+ try Monomial.pp o (MonT.retrieve vr) (* this is a non-linear variable *)
+ with Not_found -> Printf.fprintf o "v%i" vr
+ let pp o p = Vect.pp_gen pp_var o p
+ let constant c = if sign_num c = 0 then Vect.null else Vect.set 0 c Vect.null
let is_linear p =
- Vect.for_all (fun v _ ->
- let mn = (MonT.retrieve v) in
- Monomial.is_var mn || Monomial.is_const mn) p
+ Vect.for_all
+ (fun v _ ->
+ let mn = MonT.retrieve v in
+ Monomial.is_var mn || Monomial.is_const mn)
+ p
let is_variable p =
- let ((x,v),r) = Vect.decomp_fst p in
- if Vect.is_null r && v >/ Int 0
- then Monomial.get_var (MonT.retrieve x)
+ let (x, v), r = Vect.decomp_fst p in
+ if Vect.is_null r && v >/ Int 0 then Monomial.get_var (MonT.retrieve x)
else None
-
let factorise x p =
- let (px,cx) = Poly.factorise x (pol_of_linpol p) in
+ let px, cx = Poly.factorise x (pol_of_linpol p) in
(linpol_of_pol px, linpol_of_pol cx)
-
let is_linear_for x p =
- let (a,b) = factorise x p in
+ let a, b = factorise x p in
Vect.is_constant a
let search_all_linear p l =
- Vect.fold (fun acc x v ->
- if p v
- then
+ Vect.fold
+ (fun acc x v ->
+ if p v then
let x' = MonT.retrieve x in
match Monomial.get_var x' with
| None -> acc
- | Some x ->
- if is_linear_for x l
- then x::acc
- else acc
- else acc) [] l
+ | Some x -> if is_linear_for x l then x :: acc else acc
+ else acc)
+ [] l
- let min_list (l:int list) =
- match l with
- | [] -> None
- | e::l -> Some (List.fold_left min e l)
-
- let search_linear p l =
- min_list (search_all_linear p l)
+ let min_list (l : int list) =
+ match l with [] -> None | e :: l -> Some (List.fold_left min e l)
+ let search_linear p l = min_list (search_all_linear p l)
let product p1 p2 =
linpol_of_pol (Poly.product (pol_of_linpol p1) (pol_of_linpol p2))
let addition p1 p2 = Vect.add p1 p2
-
let of_vect v =
- Vect.fold (fun acc v vl -> addition (product (var v) (constant vl)) acc) Vect.null v
-
- let variables p = Vect.fold
- (fun acc v _ ->
- ISet.union (Monomial.variables (MonT.retrieve v)) acc) ISet.empty p
-
-
- let pp_goal typ o l =
- let vars = List.fold_left (fun acc p -> ISet.union acc (variables (fst p))) ISet.empty l in
- let pp_vars o i = ISet.iter (fun v -> Printf.fprintf o "(x%i : %s) " v typ) vars in
-
- Printf.fprintf o "forall %a\n" pp_vars vars ;
- List.iteri (fun i (p,op) -> Printf.fprintf o "(H%i : %a %s 0)\n" i pp p (string_of_op op)) l;
+ Vect.fold
+ (fun acc v vl -> addition (product (var v) (constant vl)) acc)
+ Vect.null v
+
+ let variables p =
+ Vect.fold
+ (fun acc v _ -> ISet.union (Monomial.variables (MonT.retrieve v)) acc)
+ ISet.empty p
+
+ let pp_goal typ o l =
+ let vars =
+ List.fold_left
+ (fun acc p -> ISet.union acc (variables (fst p)))
+ ISet.empty l
+ in
+ let pp_vars o i =
+ ISet.iter (fun v -> Printf.fprintf o "(x%i : %s) " v typ) vars
+ in
+ Printf.fprintf o "forall %a\n" pp_vars vars;
+ List.iteri
+ (fun i (p, op) ->
+ Printf.fprintf o "(H%i : %a %s 0)\n" i pp p (string_of_op op))
+ l;
Printf.fprintf o ", False\n"
-
-
-
-
- let collect_square p =
- Vect.fold (fun acc v _ ->
- let m = (MonT.retrieve v) in
- match Monomial.sqrt m with
- | None -> acc
- | Some s -> MonMap.add s m acc
- ) MonMap.empty p
-
-
+ let collect_square p =
+ Vect.fold
+ (fun acc v _ ->
+ let m = MonT.retrieve v in
+ match Monomial.sqrt m with None -> acc | Some s -> MonMap.add s m acc)
+ MonMap.empty p
end
-module ProofFormat = struct
+module ProofFormat = struct
open Big_int
type prf_rule =
| Annot of string * prf_rule
| Hyp of int
| Def of int
- | Cst of Num.num
+ | Cst of Num.num
| Zero
| Square of Vect.t
| MulC of Vect.t * prf_rule
@@ -486,264 +434,257 @@ module ProofFormat = struct
| Step of int * prf_rule * proof
| Enum of int * prf_rule * Vect.t * prf_rule * proof list
-
let rec output_prf_rule o = function
- | Annot(s,p) -> Printf.fprintf o "(%a)@%s" output_prf_rule p s
+ | Annot (s, p) -> Printf.fprintf o "(%a)@%s" output_prf_rule p s
| Hyp i -> Printf.fprintf o "Hyp %i" i
| Def i -> Printf.fprintf o "Def %i" i
| Cst c -> Printf.fprintf o "Cst %s" (string_of_num c)
- | Zero -> Printf.fprintf o "Zero"
+ | Zero -> Printf.fprintf o "Zero"
| Square s -> Printf.fprintf o "(%a)^2" Poly.pp (LinPoly.pol_of_linpol s)
- | MulC(p,pr) -> Printf.fprintf o "(%a) * (%a)" Poly.pp (LinPoly.pol_of_linpol p) output_prf_rule pr
- | MulPrf(p1,p2) -> Printf.fprintf o "(%a) * (%a)" output_prf_rule p1 output_prf_rule p2
- | AddPrf(p1,p2) -> Printf.fprintf o "%a + %a" output_prf_rule p1 output_prf_rule p2
- | CutPrf(p) -> Printf.fprintf o "[%a]" output_prf_rule p
- | Gcd(c,p) -> Printf.fprintf o "(%a)/%s" output_prf_rule p (string_of_big_int c)
+ | MulC (p, pr) ->
+ Printf.fprintf o "(%a) * (%a)" Poly.pp (LinPoly.pol_of_linpol p)
+ output_prf_rule pr
+ | MulPrf (p1, p2) ->
+ Printf.fprintf o "(%a) * (%a)" output_prf_rule p1 output_prf_rule p2
+ | AddPrf (p1, p2) ->
+ Printf.fprintf o "%a + %a" output_prf_rule p1 output_prf_rule p2
+ | CutPrf p -> Printf.fprintf o "[%a]" output_prf_rule p
+ | Gcd (c, p) ->
+ Printf.fprintf o "(%a)/%s" output_prf_rule p (string_of_big_int c)
let rec output_proof o = function
| Done -> Printf.fprintf o "."
- | Step(i,p,pf) -> Printf.fprintf o "%i:= %a ; %a" i output_prf_rule p output_proof pf
- | Enum(i,p1,v,p2,pl) -> Printf.fprintf o "%i{%a<=%a<=%a}%a" i
- output_prf_rule p1 Vect.pp v output_prf_rule p2
- (pp_list ";" output_proof) pl
+ | Step (i, p, pf) ->
+ Printf.fprintf o "%i:= %a ; %a" i output_prf_rule p output_proof pf
+ | Enum (i, p1, v, p2, pl) ->
+ Printf.fprintf o "%i{%a<=%a<=%a}%a" i output_prf_rule p1 Vect.pp v
+ output_prf_rule p2 (pp_list ";" output_proof) pl
let rec pr_size = function
- | Annot(_,p) -> pr_size p
- | Zero| Square _ -> Int 0
- | Hyp _ -> Int 1
- | Def _ -> Int 1
- | Cst n -> n
- | Gcd(i, p) -> pr_size p // (Big_int i)
- | MulPrf(p1,p2) | AddPrf(p1,p2) -> pr_size p1 +/ pr_size p2
- | CutPrf p -> pr_size p
- | MulC(v, p) -> pr_size p
-
+ | Annot (_, p) -> pr_size p
+ | Zero | Square _ -> Int 0
+ | Hyp _ -> Int 1
+ | Def _ -> Int 1
+ | Cst n -> n
+ | Gcd (i, p) -> pr_size p // Big_int i
+ | MulPrf (p1, p2) | AddPrf (p1, p2) -> pr_size p1 +/ pr_size p2
+ | CutPrf p -> pr_size p
+ | MulC (v, p) -> pr_size p
let rec pr_rule_max_id = function
- | Annot(_,p) -> pr_rule_max_id p
+ | Annot (_, p) -> pr_rule_max_id p
| Hyp i | Def i -> i
| Cst _ | Zero | Square _ -> -1
- | MulC(_,p) | CutPrf p | Gcd(_,p) -> pr_rule_max_id p
- | MulPrf(p1,p2)| AddPrf(p1,p2) -> max (pr_rule_max_id p1) (pr_rule_max_id p2)
+ | MulC (_, p) | CutPrf p | Gcd (_, p) -> pr_rule_max_id p
+ | MulPrf (p1, p2) | AddPrf (p1, p2) ->
+ max (pr_rule_max_id p1) (pr_rule_max_id p2)
let rec proof_max_id = function
| Done -> -1
- | Step(i,pr,prf) -> max i (max (pr_rule_max_id pr) (proof_max_id prf))
- | Enum(i,p1,_,p2,l) ->
- let m = max (pr_rule_max_id p1) (pr_rule_max_id p2) in
- List.fold_left (fun i prf -> max i (proof_max_id prf)) (max i m) l
-
+ | Step (i, pr, prf) -> max i (max (pr_rule_max_id pr) (proof_max_id prf))
+ | Enum (i, p1, _, p2, l) ->
+ let m = max (pr_rule_max_id p1) (pr_rule_max_id p2) in
+ List.fold_left (fun i prf -> max i (proof_max_id prf)) (max i m) l
let rec pr_rule_def_cut id = function
- | Annot(_,p) -> pr_rule_def_cut id p
- | MulC(p,prf) ->
- let (bds,id',prf') = pr_rule_def_cut id prf in
- (bds, id', MulC(p,prf'))
- | MulPrf(p1,p2) ->
- let (bds1,id,p1) = pr_rule_def_cut id p1 in
- let (bds2,id,p2) = pr_rule_def_cut id p2 in
- (bds2@bds1,id,MulPrf(p1,p2))
- | AddPrf(p1,p2) ->
- let (bds1,id,p1) = pr_rule_def_cut id p1 in
- let (bds2,id,p2) = pr_rule_def_cut id p2 in
- (bds2@bds1,id,AddPrf(p1,p2))
+ | Annot (_, p) -> pr_rule_def_cut id p
+ | MulC (p, prf) ->
+ let bds, id', prf' = pr_rule_def_cut id prf in
+ (bds, id', MulC (p, prf'))
+ | MulPrf (p1, p2) ->
+ let bds1, id, p1 = pr_rule_def_cut id p1 in
+ let bds2, id, p2 = pr_rule_def_cut id p2 in
+ (bds2 @ bds1, id, MulPrf (p1, p2))
+ | AddPrf (p1, p2) ->
+ let bds1, id, p1 = pr_rule_def_cut id p1 in
+ let bds2, id, p2 = pr_rule_def_cut id p2 in
+ (bds2 @ bds1, id, AddPrf (p1, p2))
| CutPrf p ->
- let (bds,id,p) = pr_rule_def_cut id p in
- ((id,p)::bds,id+1,Def id)
- | Gcd(c,p) ->
- let (bds,id,p) = pr_rule_def_cut id p in
- ((id,p)::bds,id+1,Def id)
- | Square _|Cst _|Def _|Hyp _|Zero as x -> ([],id,x)
-
+ let bds, id, p = pr_rule_def_cut id p in
+ ((id, p) :: bds, id + 1, Def id)
+ | Gcd (c, p) ->
+ let bds, id, p = pr_rule_def_cut id p in
+ ((id, p) :: bds, id + 1, Def id)
+ | (Square _ | Cst _ | Def _ | Hyp _ | Zero) as x -> ([], id, x)
(* Do not define top-level cuts *)
let pr_rule_def_cut id = function
| CutPrf p ->
- let (bds,ids,p') = pr_rule_def_cut id p in
- bds,ids, CutPrf p'
- | p -> pr_rule_def_cut id p
-
-
- let rec implicit_cut p =
- match p with
- | CutPrf p -> implicit_cut p
- | _ -> p
+ let bds, ids, p' = pr_rule_def_cut id p in
+ (bds, ids, CutPrf p')
+ | p -> pr_rule_def_cut id p
+ let rec implicit_cut p = match p with CutPrf p -> implicit_cut p | _ -> p
let rec pr_rule_collect_hyps pr =
match pr with
- | Annot(_,pr) -> pr_rule_collect_hyps pr
+ | Annot (_, pr) -> pr_rule_collect_hyps pr
| Hyp i | Def i -> ISet.add i ISet.empty
| Cst _ | Zero | Square _ -> ISet.empty
- | MulC(_,pr) | Gcd(_,pr)| CutPrf pr -> pr_rule_collect_hyps pr
- | MulPrf(p1,p2) | AddPrf(p1,p2) -> ISet.union (pr_rule_collect_hyps p1) (pr_rule_collect_hyps p2)
+ | MulC (_, pr) | Gcd (_, pr) | CutPrf pr -> pr_rule_collect_hyps pr
+ | MulPrf (p1, p2) | AddPrf (p1, p2) ->
+ ISet.union (pr_rule_collect_hyps p1) (pr_rule_collect_hyps p2)
- let simplify_proof p =
+ let simplify_proof p =
let rec simplify_proof p =
match p with
| Done -> (Done, ISet.empty)
- | Step(i,pr,Done) -> (p, ISet.add i (pr_rule_collect_hyps pr))
- | Step(i,pr,prf) ->
- let (prf',hyps) = simplify_proof prf in
- if not (ISet.mem i hyps)
- then (prf',hyps)
- else
- (Step(i,pr,prf'), ISet.add i (ISet.union (pr_rule_collect_hyps pr) hyps))
- | Enum(i,p1,v,p2,pl) ->
- let (pl,hl) = List.split (List.map simplify_proof pl) in
- let hyps = List.fold_left ISet.union ISet.empty hl in
- (Enum(i,p1,v,p2,pl),ISet.add i (ISet.union (ISet.union (pr_rule_collect_hyps p1) (pr_rule_collect_hyps p2)) hyps)) in
+ | Step (i, pr, Done) -> (p, ISet.add i (pr_rule_collect_hyps pr))
+ | Step (i, pr, prf) ->
+ let prf', hyps = simplify_proof prf in
+ if not (ISet.mem i hyps) then (prf', hyps)
+ else
+ ( Step (i, pr, prf')
+ , ISet.add i (ISet.union (pr_rule_collect_hyps pr) hyps) )
+ | Enum (i, p1, v, p2, pl) ->
+ let pl, hl = List.split (List.map simplify_proof pl) in
+ let hyps = List.fold_left ISet.union ISet.empty hl in
+ ( Enum (i, p1, v, p2, pl)
+ , ISet.add i
+ (ISet.union
+ (ISet.union (pr_rule_collect_hyps p1) (pr_rule_collect_hyps p2))
+ hyps) )
+ in
fst (simplify_proof p)
-
let rec normalise_proof id prf =
match prf with
- | Done -> (id,Done)
- | Step(i,Gcd(c,p),Done) -> normalise_proof id (Step(i,p,Done))
- | Step(i,p,prf) ->
- let bds,id,p' = pr_rule_def_cut id p in
- let (id,prf) = normalise_proof id prf in
- let prf = List.fold_left (fun acc (i,p) -> Step(i, CutPrf p,acc))
- (Step(i,p',prf)) bds in
-
- (id,prf)
- | Enum(i,p1,v,p2,pl) ->
- (* Why do I have top-level cuts ? *)
- (* let p1 = implicit_cut p1 in
+ | Done -> (id, Done)
+ | Step (i, Gcd (c, p), Done) -> normalise_proof id (Step (i, p, Done))
+ | Step (i, p, prf) ->
+ let bds, id, p' = pr_rule_def_cut id p in
+ let id, prf = normalise_proof id prf in
+ let prf =
+ List.fold_left
+ (fun acc (i, p) -> Step (i, CutPrf p, acc))
+ (Step (i, p', prf))
+ bds
+ in
+ (id, prf)
+ | Enum (i, p1, v, p2, pl) ->
+ (* Why do I have top-level cuts ? *)
+ (* let p1 = implicit_cut p1 in
let p2 = implicit_cut p2 in
let (ids,prfs) = List.split (List.map (normalise_proof id) pl) in
(List.fold_left max 0 ids ,
Enum(i,p1,v,p2,prfs))
*)
-
- let bds1,id,p1' = pr_rule_def_cut id (implicit_cut p1) in
- let bds2,id,p2' = pr_rule_def_cut id (implicit_cut p2) in
- let (ids,prfs) = List.split (List.map (normalise_proof id) pl) in
- (List.fold_left max 0 ids ,
- List.fold_left (fun acc (i,p) -> Step(i, CutPrf p,acc))
- (Enum(i,p1',v,p2',prfs)) (bds2@bds1))
-
+ let bds1, id, p1' = pr_rule_def_cut id (implicit_cut p1) in
+ let bds2, id, p2' = pr_rule_def_cut id (implicit_cut p2) in
+ let ids, prfs = List.split (List.map (normalise_proof id) pl) in
+ ( List.fold_left max 0 ids
+ , List.fold_left
+ (fun acc (i, p) -> Step (i, CutPrf p, acc))
+ (Enum (i, p1', v, p2', prfs))
+ (bds2 @ bds1) )
let normalise_proof id prf =
let prf = simplify_proof prf in
let res = normalise_proof id prf in
- if debug then Printf.printf "normalise_proof %a -> %a" output_proof prf output_proof (snd res) ;
+ if debug then
+ Printf.printf "normalise_proof %a -> %a" output_proof prf output_proof
+ (snd res);
res
- module OrdPrfRule =
- struct
- type t = prf_rule
-
- let id_of_constr = function
- | Annot _ -> 0
- | Hyp _ -> 1
- | Def _ -> 2
- | Cst _ -> 3
- | Zero -> 4
- | Square _ -> 5
- | MulC _ -> 6
- | Gcd _ -> 7
- | MulPrf _ -> 8
- | AddPrf _ -> 9
- | CutPrf _ -> 10
-
- let cmp_pair c1 c2 (x1,x2) (y1,y2) =
- match c1 x1 y1 with
- | 0 -> c2 x2 y2
- | i -> i
-
-
- let rec compare p1 p2 =
- match p1, p2 with
- | Annot(s1,p1) , Annot(s2,p2) -> if s1 = s2 then compare p1 p2
- else Util.pervasives_compare s1 s2
- | Hyp i , Hyp j -> Util.pervasives_compare i j
- | Def i , Def j -> Util.pervasives_compare i j
- | Cst n , Cst m -> Num.compare_num n m
- | Zero , Zero -> 0
- | Square v1 , Square v2 -> Vect.compare v1 v2
- | MulC(v1,p1) , MulC(v2,p2) -> cmp_pair Vect.compare compare (v1,p1) (v2,p2)
- | Gcd(b1,p1) , Gcd(b2,p2) -> cmp_pair Big_int.compare_big_int compare (b1,p1) (b2,p2)
- | MulPrf(p1,q1) , MulPrf(p2,q2) -> cmp_pair compare compare (p1,q1) (p2,q2)
- | AddPrf(p1,q1) , MulPrf(p2,q2) -> cmp_pair compare compare (p1,q1) (p2,q2)
- | CutPrf p , CutPrf p' -> compare p p'
- | _ , _ -> Util.pervasives_compare (id_of_constr p1) (id_of_constr p2)
-
- end
-
-
-
+ module OrdPrfRule = struct
+ type t = prf_rule
+
+ let id_of_constr = function
+ | Annot _ -> 0
+ | Hyp _ -> 1
+ | Def _ -> 2
+ | Cst _ -> 3
+ | Zero -> 4
+ | Square _ -> 5
+ | MulC _ -> 6
+ | Gcd _ -> 7
+ | MulPrf _ -> 8
+ | AddPrf _ -> 9
+ | CutPrf _ -> 10
+
+ let cmp_pair c1 c2 (x1, x2) (y1, y2) =
+ match c1 x1 y1 with 0 -> c2 x2 y2 | i -> i
+
+ let rec compare p1 p2 =
+ match (p1, p2) with
+ | Annot (s1, p1), Annot (s2, p2) ->
+ if s1 = s2 then compare p1 p2 else Util.pervasives_compare s1 s2
+ | Hyp i, Hyp j -> Util.pervasives_compare i j
+ | Def i, Def j -> Util.pervasives_compare i j
+ | Cst n, Cst m -> Num.compare_num n m
+ | Zero, Zero -> 0
+ | Square v1, Square v2 -> Vect.compare v1 v2
+ | MulC (v1, p1), MulC (v2, p2) ->
+ cmp_pair Vect.compare compare (v1, p1) (v2, p2)
+ | Gcd (b1, p1), Gcd (b2, p2) ->
+ cmp_pair Big_int.compare_big_int compare (b1, p1) (b2, p2)
+ | MulPrf (p1, q1), MulPrf (p2, q2) ->
+ cmp_pair compare compare (p1, q1) (p2, q2)
+ | AddPrf (p1, q1), MulPrf (p2, q2) ->
+ cmp_pair compare compare (p1, q1) (p2, q2)
+ | CutPrf p, CutPrf p' -> compare p p'
+ | _, _ -> Util.pervasives_compare (id_of_constr p1) (id_of_constr p2)
+ end
let add_proof x y =
- match x, y with
- | Zero , p | p , Zero -> p
- | _ -> AddPrf(x,y)
-
+ match (x, y) with Zero, p | p, Zero -> p | _ -> AddPrf (x, y)
let rec mul_cst_proof c p =
match p with
- | Annot(s,p) -> Annot(s,mul_cst_proof c p)
- | MulC(v,p') -> MulC(Vect.mul c v,p')
- | _ ->
- match sign_num c with
- | 0 -> Zero (* This is likely to be a bug *)
- | -1 -> MulC(LinPoly.constant c, p) (* [p] should represent an equality *)
- | 1 ->
- if eq_num (Int 1) c
- then p
- else MulPrf(Cst c,p)
- | _ -> assert false
-
+ | Annot (s, p) -> Annot (s, mul_cst_proof c p)
+ | MulC (v, p') -> MulC (Vect.mul c v, p')
+ | _ -> (
+ match sign_num c with
+ | 0 -> Zero (* This is likely to be a bug *)
+ | -1 ->
+ MulC (LinPoly.constant c, p) (* [p] should represent an equality *)
+ | 1 -> if eq_num (Int 1) c then p else MulPrf (Cst c, p)
+ | _ -> assert false )
let sMulC v p =
- let (c,v') = Vect.decomp_cst v in
- if Vect.is_null v' then mul_cst_proof c p
- else MulC(v,p)
-
+ let c, v' = Vect.decomp_cst v in
+ if Vect.is_null v' then mul_cst_proof c p else MulC (v, p)
let mul_proof p1 p2 =
- match p1 , p2 with
- | Zero , _ | _ , Zero -> Zero
- | Cst c , p | p , Cst c -> mul_cst_proof c p
- | _ , _ ->
- MulPrf(p1,p2)
-
+ match (p1, p2) with
+ | Zero, _ | _, Zero -> Zero
+ | Cst c, p | p, Cst c -> mul_cst_proof c p
+ | _, _ -> MulPrf (p1, p2)
- module PrfRuleMap = Map.Make(OrdPrfRule)
+ module PrfRuleMap = Map.Make (OrdPrfRule)
- let prf_rule_of_map m =
- PrfRuleMap.fold (fun k v acc -> add_proof (sMulC v k) acc) m Zero
+ let prf_rule_of_map m =
+ PrfRuleMap.fold (fun k v acc -> add_proof (sMulC v k) acc) m Zero
-
- let rec dev_prf_rule p =
- match p with
- | Annot(s,p) -> dev_prf_rule p
- | Hyp _ | Def _ | Cst _ | Zero | Square _ -> PrfRuleMap.singleton p (LinPoly.constant (Int 1))
- | MulC(v,p) -> PrfRuleMap.map (fun v1 -> LinPoly.product v v1) (dev_prf_rule p)
- | AddPrf(p1,p2) -> PrfRuleMap.merge (fun k o1 o2 ->
- match o1 , o2 with
- | None , None -> None
- | None , Some v | Some v, None -> Some v
- | Some v1 , Some v2 -> Some (LinPoly.addition v1 v2)) (dev_prf_rule p1) (dev_prf_rule p2)
- | MulPrf(p1, p2) ->
- begin
- let p1' = dev_prf_rule p1 in
- let p2' = dev_prf_rule p2 in
-
- let p1'' = prf_rule_of_map p1' in
- let p2'' = prf_rule_of_map p2' in
-
- match p1'' with
- | Cst c -> PrfRuleMap.map (fun v1 -> Vect.mul c v1) p2'
- | _ -> PrfRuleMap.singleton (MulPrf(p1'',p2'')) (LinPoly.constant (Int 1))
- end
- | _ -> PrfRuleMap.singleton p (LinPoly.constant (Int 1))
-
- let simplify_prf_rule p =
- prf_rule_of_map (dev_prf_rule p)
-
-
- (*
+ let rec dev_prf_rule p =
+ match p with
+ | Annot (s, p) -> dev_prf_rule p
+ | Hyp _ | Def _ | Cst _ | Zero | Square _ ->
+ PrfRuleMap.singleton p (LinPoly.constant (Int 1))
+ | MulC (v, p) ->
+ PrfRuleMap.map (fun v1 -> LinPoly.product v v1) (dev_prf_rule p)
+ | AddPrf (p1, p2) ->
+ PrfRuleMap.merge
+ (fun k o1 o2 ->
+ match (o1, o2) with
+ | None, None -> None
+ | None, Some v | Some v, None -> Some v
+ | Some v1, Some v2 -> Some (LinPoly.addition v1 v2))
+ (dev_prf_rule p1) (dev_prf_rule p2)
+ | MulPrf (p1, p2) -> (
+ let p1' = dev_prf_rule p1 in
+ let p2' = dev_prf_rule p2 in
+ let p1'' = prf_rule_of_map p1' in
+ let p2'' = prf_rule_of_map p2' in
+ match p1'' with
+ | Cst c -> PrfRuleMap.map (fun v1 -> Vect.mul c v1) p2'
+ | _ ->
+ PrfRuleMap.singleton (MulPrf (p1'', p2'')) (LinPoly.constant (Int 1)) )
+ | _ -> PrfRuleMap.singleton p (LinPoly.constant (Int 1))
+
+ let simplify_prf_rule p = prf_rule_of_map (dev_prf_rule p)
+
+ (*
let mul_proof p1 p2 =
let res = mul_proof p1 p2 in
Printf.printf "mul_proof %a %a = %a\n"
@@ -767,309 +708,285 @@ module ProofFormat = struct
*)
let proof_of_farkas env vect =
- Vect.fold (fun prf x n ->
- add_proof (mul_cst_proof n (IMap.find x env)) prf) Zero vect
-
-
-
-
- module Env = struct
+ Vect.fold
+ (fun prf x n -> add_proof (mul_cst_proof n (IMap.find x env)) prf)
+ Zero vect
+ module Env = struct
let rec string_of_int_list l =
match l with
| [] -> ""
- | i::l -> Printf.sprintf "%i,%s" i (string_of_int_list l)
-
+ | i :: l -> Printf.sprintf "%i,%s" i (string_of_int_list l)
let id_of_hyp hyp l =
let rec xid_of_hyp i l' =
match l' with
- | [] -> failwith (Printf.sprintf "id_of_hyp %i %s" hyp (string_of_int_list l))
- | hyp'::l' -> if (=) hyp hyp' then i else xid_of_hyp (i+1) l' in
+ | [] ->
+ failwith (Printf.sprintf "id_of_hyp %i %s" hyp (string_of_int_list l))
+ | hyp' :: l' -> if hyp = hyp' then i else xid_of_hyp (i + 1) l'
+ in
xid_of_hyp 0 l
-
end
- let cmpl_prf_rule norm (cst:num-> 'a) env prf =
- let rec cmpl =
- function
- | Annot(s,p) -> cmpl p
+ let cmpl_prf_rule norm (cst : num -> 'a) env prf =
+ let rec cmpl = function
+ | Annot (s, p) -> cmpl p
| Hyp i | Def i -> Mc.PsatzIn (CamlToCoq.nat (Env.id_of_hyp i env))
- | Cst i -> Mc.PsatzC (cst i)
- | Zero -> Mc.PsatzZ
- | MulPrf(p1,p2) -> Mc.PsatzMulE(cmpl p1, cmpl p2)
- | AddPrf(p1,p2) -> Mc.PsatzAdd(cmpl p1 , cmpl p2)
- | MulC(lp,p) -> let lp = norm (LinPoly.coq_poly_of_linpol cst lp) in
- Mc.PsatzMulC(lp,cmpl p)
- | Square lp -> Mc.PsatzSquare (norm (LinPoly.coq_poly_of_linpol cst lp))
- | _ -> failwith "Cuts should already be compiled" in
+ | Cst i -> Mc.PsatzC (cst i)
+ | Zero -> Mc.PsatzZ
+ | MulPrf (p1, p2) -> Mc.PsatzMulE (cmpl p1, cmpl p2)
+ | AddPrf (p1, p2) -> Mc.PsatzAdd (cmpl p1, cmpl p2)
+ | MulC (lp, p) ->
+ let lp = norm (LinPoly.coq_poly_of_linpol cst lp) in
+ Mc.PsatzMulC (lp, cmpl p)
+ | Square lp -> Mc.PsatzSquare (norm (LinPoly.coq_poly_of_linpol cst lp))
+ | _ -> failwith "Cuts should already be compiled"
+ in
cmpl prf
+ let cmpl_prf_rule_z env r =
+ cmpl_prf_rule Mc.normZ (fun x -> CamlToCoq.bigint (numerator x)) env r
-
-
- let cmpl_prf_rule_z env r = cmpl_prf_rule Mc.normZ (fun x -> CamlToCoq.bigint (numerator x)) env r
-
- let rec cmpl_proof env = function
- | Done -> Mc.DoneProof
- | Step(i,p,prf) ->
- begin
- match p with
- | CutPrf p' ->
- Mc.CutProof(cmpl_prf_rule_z env p', cmpl_proof (i::env) prf)
- | _ -> Mc.RatProof(cmpl_prf_rule_z env p,cmpl_proof (i::env) prf)
- end
- | Enum(i,p1,_,p2,l) ->
- Mc.EnumProof(cmpl_prf_rule_z env p1,cmpl_prf_rule_z env p2,List.map (cmpl_proof (i::env)) l)
-
+ let rec cmpl_proof env = function
+ | Done -> Mc.DoneProof
+ | Step (i, p, prf) -> (
+ match p with
+ | CutPrf p' ->
+ Mc.CutProof (cmpl_prf_rule_z env p', cmpl_proof (i :: env) prf)
+ | _ -> Mc.RatProof (cmpl_prf_rule_z env p, cmpl_proof (i :: env) prf) )
+ | Enum (i, p1, _, p2, l) ->
+ Mc.EnumProof
+ ( cmpl_prf_rule_z env p1
+ , cmpl_prf_rule_z 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
+ let _, prf = normalise_proof id prf in
cmpl_proof env prf
let rec eval_prf_rule env = function
- | Annot(s,p) -> eval_prf_rule env p
+ | Annot (s, p) -> eval_prf_rule env p
| Hyp i | Def i -> env i
- | Cst n -> (Vect.set 0 n Vect.null,
- match Num.compare_num n (Int 0) with
- | 0 -> Ge
- | 1 -> Gt
- | _ -> failwith "eval_prf_rule : negative constant"
- )
- | Zero -> (Vect.null, Ge)
- | Square v -> (LinPoly.product v v,Ge)
- | MulC(v, p) ->
- let (p1,o) = eval_prf_rule env p in
- begin match o with
- | Eq -> (LinPoly.product v p1,Eq)
- | _ ->
- Printf.fprintf stdout "MulC(%a,%a) invalid 2d arg %a %s" Vect.pp v output_prf_rule p Vect.pp p1 (string_of_op o);
- failwith "eval_prf_rule : not an equality"
- end
- | Gcd(g,p) -> let (v,op) = eval_prf_rule env p in
- (Vect.div (Big_int g) v, op)
- | MulPrf(p1,p2) ->
- let (v1,o1) = eval_prf_rule env p1 in
- let (v2,o2) = eval_prf_rule env p2 in
- (LinPoly.product v1 v2, opMult o1 o2)
- | AddPrf(p1,p2) ->
- let (v1,o1) = eval_prf_rule env p1 in
- let (v2,o2) = eval_prf_rule env p2 in
- (LinPoly.addition v1 v2, opAdd o1 o2)
- | CutPrf p -> eval_prf_rule env p
-
-
- let is_unsat (p,o) =
- let (c,r) = Vect.decomp_cst p in
- if Vect.is_null r
- then not (eval_op o c (Int 0))
- else false
+ | Cst n -> (
+ ( Vect.set 0 n Vect.null
+ , match Num.compare_num n (Int 0) with
+ | 0 -> Ge
+ | 1 -> Gt
+ | _ -> failwith "eval_prf_rule : negative constant" ) )
+ | Zero -> (Vect.null, Ge)
+ | Square v -> (LinPoly.product v v, Ge)
+ | MulC (v, p) -> (
+ let p1, o = eval_prf_rule env p in
+ match o with
+ | Eq -> (LinPoly.product v p1, Eq)
+ | _ ->
+ Printf.fprintf stdout "MulC(%a,%a) invalid 2d arg %a %s" Vect.pp v
+ output_prf_rule p Vect.pp p1 (string_of_op o);
+ failwith "eval_prf_rule : not an equality" )
+ | Gcd (g, p) ->
+ let v, op = eval_prf_rule env p in
+ (Vect.div (Big_int g) v, op)
+ | MulPrf (p1, p2) ->
+ let v1, o1 = eval_prf_rule env p1 in
+ let v2, o2 = eval_prf_rule env p2 in
+ (LinPoly.product v1 v2, opMult o1 o2)
+ | AddPrf (p1, p2) ->
+ let v1, o1 = eval_prf_rule env p1 in
+ let v2, o2 = eval_prf_rule env p2 in
+ (LinPoly.addition v1 v2, opAdd o1 o2)
+ | CutPrf p -> eval_prf_rule env p
+
+ let is_unsat (p, o) =
+ let c, r = Vect.decomp_cst p in
+ if Vect.is_null r then not (eval_op o c (Int 0)) else false
let rec eval_proof env p =
match p with
| Done -> failwith "Proof is not finished"
- | Step(i, prf, rst) ->
- let (p,o) = eval_prf_rule (fun i -> IMap.find i env) prf in
- if is_unsat (p,o) then true
- else
- if (=) rst Done
- then
- begin
- Printf.fprintf stdout "Last inference %a %s\n" LinPoly.pp p (string_of_op o);
- false
- end
- else eval_proof (IMap.add i (p,o) env) rst
- | Enum(i,r1,v,r2,l) -> let _ = eval_prf_rule (fun i -> IMap.find i env) r1 in
- let _ = eval_prf_rule (fun i -> IMap.find i env) r2 in
- (* Should check bounds *)
- failwith "Not implemented"
-
+ | Step (i, prf, rst) ->
+ let p, o = eval_prf_rule (fun i -> IMap.find i env) prf in
+ if is_unsat (p, o) then true
+ else if rst = Done then begin
+ Printf.fprintf stdout "Last inference %a %s\n" LinPoly.pp p
+ (string_of_op o);
+ false
+ end
+ else eval_proof (IMap.add i (p, o) env) rst
+ | Enum (i, r1, v, r2, l) ->
+ let _ = eval_prf_rule (fun i -> IMap.find i env) r1 in
+ let _ = eval_prf_rule (fun i -> IMap.find i env) r2 in
+ (* Should check bounds *)
+ failwith "Not implemented"
end
-module WithProof = struct
-
- type t = ((LinPoly.t * op) * ProofFormat.prf_rule)
+module WithProof = struct
+ type t = (LinPoly.t * op) * ProofFormat.prf_rule
- let annot s (p,prf) = (p, ProofFormat.Annot(s,prf))
+ let annot s (p, prf) = (p, ProofFormat.Annot (s, prf))
- let output o ((lp,op),prf) =
- Printf.fprintf o "%a %s 0 by %a\n" LinPoly.pp lp (string_of_op op) ProofFormat.output_prf_rule prf
+ let output o ((lp, op), prf) =
+ Printf.fprintf o "%a %s 0 by %a\n" LinPoly.pp lp (string_of_op op)
+ ProofFormat.output_prf_rule prf
- let output_sys o l =
- List.iter (Printf.fprintf o "%a\n" output) l
+ let output_sys o l = List.iter (Printf.fprintf o "%a\n" output) l
exception InvalidProof
- let zero = ((Vect.null,Eq), ProofFormat.Zero)
+ let zero = ((Vect.null, Eq), ProofFormat.Zero)
+ let const n = ((LinPoly.constant n, Ge), ProofFormat.Cst n)
+ let of_cstr (c, prf) = ((Vect.set 0 (Num.minus_num c.cst) c.coeffs, c.op), prf)
- let const n = ((LinPoly.constant n,Ge), ProofFormat.Cst n)
-
- let of_cstr (c,prf) =
- (Vect.set 0 (Num.minus_num (c.cst)) c.coeffs,c.op), prf
-
- let product : t -> t -> t = fun ((p1,o1),prf1) ((p2,o2),prf2) ->
- ((LinPoly.product p1 p2 , opMult o1 o2), ProofFormat.mul_proof prf1 prf2)
+ let product : t -> t -> t =
+ fun ((p1, o1), prf1) ((p2, o2), prf2) ->
+ ((LinPoly.product p1 p2, opMult o1 o2), ProofFormat.mul_proof prf1 prf2)
- let addition : t -> t -> t = fun ((p1,o1),prf1) ((p2,o2),prf2) ->
+ let addition : t -> t -> t =
+ fun ((p1, o1), prf1) ((p2, o2), prf2) ->
((Vect.add p1 p2, opAdd o1 o2), ProofFormat.add_proof prf1 prf2)
- let mult p ((p1,o1),prf1) =
+ let mult p ((p1, o1), prf1) =
match o1 with
- | Eq -> ((LinPoly.product p p1,o1), ProofFormat.sMulC p prf1)
- | Gt| Ge -> let (n,r) = Vect.decomp_cst p in
- if Vect.is_null r && n >/ Int 0
- then ((LinPoly.product p p1, o1), ProofFormat.mul_cst_proof n prf1)
- else raise InvalidProof
-
-
- let cutting_plane ((p,o),prf) =
- let (c,p') = Vect.decomp_cst p in
- let g = (Vect.gcd p') in
- if (Big_int.eq_big_int Big_int.unit_big_int g) || c =/ Int 0 ||
- not (Big_int.eq_big_int (denominator c) Big_int.unit_big_int)
+ | Eq -> ((LinPoly.product p p1, o1), ProofFormat.sMulC p prf1)
+ | Gt | Ge ->
+ let n, r = Vect.decomp_cst p in
+ if Vect.is_null r && n >/ Int 0 then
+ ((LinPoly.product p p1, o1), ProofFormat.mul_cst_proof n prf1)
+ else raise InvalidProof
+
+ let cutting_plane ((p, o), prf) =
+ let c, p' = Vect.decomp_cst p in
+ let g = Vect.gcd p' in
+ if
+ Big_int.eq_big_int Big_int.unit_big_int g
+ || c =/ Int 0
+ || not (Big_int.eq_big_int (denominator c) Big_int.unit_big_int)
then None (* Nothing to do *)
else
- let c1 = c // (Big_int g) in
+ let c1 = c // Big_int g in
let c1' = Num.floor_num c1 in
- if c1 =/ c1'
- then None
+ if c1 =/ c1' then None
else
match o with
- | Eq -> Some ((Vect.set 0 (Int (-1)) Vect.null,Eq), ProofFormat.Gcd(g,prf))
+ | Eq ->
+ Some ((Vect.set 0 (Int (-1)) Vect.null, Eq), ProofFormat.Gcd (g, prf))
| Gt -> failwith "cutting_plane ignore strict constraints"
| Ge ->
- (* This is a non-trivial common divisor *)
- Some ((Vect.set 0 c1' (Vect.div (Big_int g) p),o),ProofFormat.Gcd(g, prf))
-
+ (* This is a non-trivial common divisor *)
+ Some
+ ( (Vect.set 0 c1' (Vect.div (Big_int g) p), o)
+ , ProofFormat.Gcd (g, prf) )
let construct_sign p =
- let (c,p') = Vect.decomp_cst p in
- if Vect.is_null p'
- then
- Some (begin match sign_num c with
- | 0 -> (true, Eq, ProofFormat.Zero)
- | 1 -> (true,Gt, ProofFormat.Cst c)
- | _ (*-1*) -> (false,Gt, ProofFormat.Cst (minus_num c))
- end)
+ let c, p' = Vect.decomp_cst p in
+ if Vect.is_null p' then
+ Some
+ ( match sign_num c with
+ | 0 -> (true, Eq, ProofFormat.Zero)
+ | 1 -> (true, Gt, ProofFormat.Cst c)
+ | _ (*-1*) -> (false, Gt, ProofFormat.Cst (minus_num c)) )
else None
-
let get_sign l p =
match construct_sign p with
- | None -> begin
+ | None -> (
+ try
+ let (p', o), prf =
+ List.find (fun ((p', o), prf) -> Vect.equal p p') l
+ in
+ Some (true, o, prf)
+ with Not_found -> (
+ let p = Vect.uminus p in
try
- let ((p',o),prf) =
- List.find (fun ((p',o),prf) -> Vect.equal p p') l in
- Some (true,o,prf)
- with Not_found ->
- let p = Vect.uminus p in
- try
- let ((p',o),prf) = List.find (fun ((p',o),prf) -> Vect.equal p p') l in
- Some (false,o,prf)
- with Not_found -> None
- end
+ let (p', o), prf =
+ List.find (fun ((p', o), prf) -> Vect.equal p p') l
+ in
+ Some (false, o, prf)
+ with Not_found -> None ) )
| Some s -> Some s
+ let mult_sign : bool -> t -> t =
+ fun b ((p, o), prf) -> if b then ((p, o), prf) else ((Vect.uminus p, o), prf)
- let mult_sign : bool -> t -> t = fun b ((p,o),prf) ->
- if b then ((p,o),prf)
- else ((Vect.uminus p,o),prf)
-
-
- let rec linear_pivot sys ((lp1,op1), prf1) x ((lp2,op2), prf2) =
-
+ let rec linear_pivot sys ((lp1, op1), prf1) x ((lp2, op2), prf2) =
(* lp1 = a1.x + b1 *)
- let (a1,b1) = LinPoly.factorise x lp1 in
-
+ let a1, b1 = LinPoly.factorise x lp1 in
(* lp2 = a2.x + b2 *)
- let (a2,b2) = LinPoly.factorise x lp2 in
-
- if Vect.is_null a2
- then (* We are done *)
- Some ((lp2,op2),prf2)
+ let a2, b2 = LinPoly.factorise x lp2 in
+ if Vect.is_null a2 then (* We are done *)
+ Some ((lp2, op2), prf2)
else
- match op1,op2 with
- | Eq , (Ge|Gt) -> begin
- match get_sign sys a1 with
- | None -> None (* Impossible to pivot without sign information *)
- | Some(b,o,prf) ->
- let sa1 = mult_sign b ((a1,o),prf) in
- let sa2 = if b then (Vect.uminus a2) else a2 in
-
- let ((lp2,op2),prf2) =
- addition (product sa1 ((lp2,op2),prf2))
- (mult sa2 ((lp1,op1),prf1)) in
- linear_pivot sys ((lp1,op1),prf1) x ((lp2,op2),prf2)
-
- end
- | Eq , Eq ->
- let ((lp2,op2),prf2) = addition (mult a1 ((lp2,op2),prf2))
- (mult (Vect.uminus a2) ((lp1,op1),prf1)) in
- linear_pivot sys ((lp1,op1),prf1) x ((lp2,op2),prf2)
-
- | (Ge | Gt) , (Ge| Gt) -> begin
- match get_sign sys a1 , get_sign sys a2 with
- | Some(b1,o1,p1) , Some(b2,o2,p2) ->
- if b1 <> b2
- then
- let ((lp2,op2),prf2) =
- addition (product (mult_sign b1 ((a1,o1), p1)) ((lp2,op2),prf2))
- (product (mult_sign b2 ((a2,o2), p2)) ((lp1,op1),prf1)) in
- linear_pivot sys ((lp1,op1),prf1) x ((lp2,op2),prf2)
- else None
- | _ -> None
- end
- | (Ge|Gt) , Eq -> failwith "pivot: equality as second argument"
-
- let linear_pivot sys ((lp1,op1), prf1) x ((lp2,op2), prf2) =
- match linear_pivot sys ((lp1,op1), prf1) x ((lp2,op2), prf2) with
+ match (op1, op2) with
+ | Eq, (Ge | Gt) -> (
+ match get_sign sys a1 with
+ | None -> None (* Impossible to pivot without sign information *)
+ | Some (b, o, prf) ->
+ let sa1 = mult_sign b ((a1, o), prf) in
+ let sa2 = if b then Vect.uminus a2 else a2 in
+ let (lp2, op2), prf2 =
+ addition
+ (product sa1 ((lp2, op2), prf2))
+ (mult sa2 ((lp1, op1), prf1))
+ in
+ linear_pivot sys ((lp1, op1), prf1) x ((lp2, op2), prf2) )
+ | Eq, Eq ->
+ let (lp2, op2), prf2 =
+ addition
+ (mult a1 ((lp2, op2), prf2))
+ (mult (Vect.uminus a2) ((lp1, op1), prf1))
+ in
+ linear_pivot sys ((lp1, op1), prf1) x ((lp2, op2), prf2)
+ | (Ge | Gt), (Ge | Gt) -> (
+ match (get_sign sys a1, get_sign sys a2) with
+ | Some (b1, o1, p1), Some (b2, o2, p2) ->
+ if b1 <> b2 then
+ let (lp2, op2), prf2 =
+ addition
+ (product (mult_sign b1 ((a1, o1), p1)) ((lp2, op2), prf2))
+ (product (mult_sign b2 ((a2, o2), p2)) ((lp1, op1), prf1))
+ in
+ linear_pivot sys ((lp1, op1), prf1) x ((lp2, op2), prf2)
+ else None
+ | _ -> None )
+ | (Ge | Gt), Eq -> failwith "pivot: equality as second argument"
+
+ let linear_pivot sys ((lp1, op1), prf1) x ((lp2, op2), prf2) =
+ match linear_pivot sys ((lp1, op1), prf1) x ((lp2, op2), prf2) with
| None -> None
- | Some (c,p) -> Some(c, ProofFormat.simplify_prf_rule p)
+ | Some (c, p) -> Some (c, ProofFormat.simplify_prf_rule p)
+ let is_substitution strict ((p, o), prf) =
+ let pred v = if strict then v =/ Int 1 || v =/ Int (-1) else true in
+ match o with Eq -> LinPoly.search_linear pred p | _ -> None
-let is_substitution strict ((p,o),prf) =
- let pred v = if strict then v =/ Int 1 || v =/ Int (-1) else true in
-
- match o with
- | Eq -> LinPoly.search_linear pred p
- | _ -> None
-
-
-let subst1 sys0 =
- let (oeq,sys') = extract (is_substitution true) sys0 in
- match oeq with
- | None -> sys0
- | Some(v,pc) ->
- match simplify (linear_pivot sys0 pc v) sys' with
- | None -> sys0
- | Some sys' -> sys'
-
-
-
-let subst sys0 =
- let elim sys =
- let (oeq,sys') = extract (is_substitution true) sys in
+ let subst1 sys0 =
+ let oeq, sys' = extract (is_substitution true) sys0 in
match oeq with
- | None -> None
- | Some(v,pc) -> simplify (linear_pivot sys0 pc v) sys' in
-
- iterate_until_stable elim sys0
-
-
-let saturate_subst b sys0 =
- let select = is_substitution b in
- let gen (v,pc) ((c,op),prf) =
- if ISet.mem v (LinPoly.variables c)
- then linear_pivot sys0 pc v ((c,op),prf)
- else None
- in
- saturate select gen sys0
-
-
+ | None -> sys0
+ | Some (v, pc) -> (
+ match simplify (linear_pivot sys0 pc v) sys' with
+ | None -> sys0
+ | Some sys' -> sys' )
+
+ let subst sys0 =
+ let elim sys =
+ let oeq, sys' = extract (is_substitution true) sys in
+ match oeq with
+ | None -> None
+ | Some (v, pc) -> simplify (linear_pivot sys0 pc v) sys'
+ in
+ iterate_until_stable elim sys0
+
+ let saturate_subst b sys0 =
+ let select = is_substitution b in
+ let gen (v, pc) ((c, op), prf) =
+ if ISet.mem v (LinPoly.variables c) then
+ linear_pivot sys0 pc v ((c, op), prf)
+ else None
+ in
+ saturate select gen sys0
end
-
(* Local Variables: *)
(* coding: utf-8 *)
(* End: *)
diff --git a/plugins/micromega/polynomial.mli b/plugins/micromega/polynomial.mli
index cfb1bb914c..e7c619affc 100644
--- a/plugins/micromega/polynomial.mli
+++ b/plugins/micromega/polynomial.mli
@@ -9,7 +9,6 @@
(************************************************************************)
open Mutils
-
module Mc = Micromega
val max_nb_cstr : int ref
@@ -17,46 +16,45 @@ val max_nb_cstr : int ref
type var = int
module Monomial : sig
- (** A monomial is represented by a multiset of variables *)
type t
+ (** A monomial is represented by a multiset of variables *)
- (** [fold f m acc]
- folds over the variables with multiplicities *)
val fold : (var -> int -> 'a -> 'a) -> t -> 'a -> 'a
+ (** [fold f m acc]
+ folds over the variables with multiplicities *)
+ val const : t
(** [const]
@return the empty monomial i.e. without any variable *)
- val const : t
val is_const : t -> bool
+ val var : var -> t
(** [var x]
@return the monomial x^1 *)
- val var : var -> t
+ val sqrt : t -> t option
(** [sqrt m]
@return [Some r] iff r^2 = m *)
- val sqrt : t -> t option
+ val is_var : t -> bool
(** [is_var m]
@return [true] iff m = x^1 for some variable x *)
- val is_var : t -> bool
+ val get_var : t -> var option
(** [get_var m]
@return [x] iff m = x^1 for variable x *)
- val get_var : t -> var option
-
+ val div : t -> t -> t * int
(** [div m1 m2]
@return a pair [mr,n] such that mr * (m2)^n = m1 where n is maximum *)
- val div : t -> t -> t * int
- (** [compare m1 m2] provides a total order over monomials*)
val compare : t -> t -> int
+ (** [compare m1 m2] provides a total order over monomials*)
+ val variables : t -> ISet.t
(** [variables m]
@return the set of variables with (strictly) positive multiplicities *)
- val variables : t -> ISet.t
end
module MonMap : sig
@@ -76,52 +74,52 @@ module Poly : sig
type t
+ val constant : Num.num -> t
(** [constant c]
@return the constant polynomial c *)
- val constant : Num.num -> t
+ val variable : var -> t
(** [variable x]
@return the polynomial 1.x^1 *)
- val variable : var -> t
+ val addition : t -> t -> t
(** [addition p1 p2]
@return the polynomial p1+p2 *)
- val addition : t -> t -> t
+ val product : t -> t -> t
(** [product p1 p2]
@return the polynomial p1*p2 *)
- val product : t -> t -> t
+ val uminus : t -> t
(** [uminus p]
@return the polynomial -p i.e product by -1 *)
- val uminus : t -> t
+ val get : Monomial.t -> t -> Num.num
(** [get mi p]
@return the coefficient ai of the monomial mi. *)
- val get : Monomial.t -> t -> Num.num
-
- (** [fold f p a] folds f over the monomials of p with non-zero coefficient *)
val fold : (Monomial.t -> Num.num -> 'a -> 'a) -> t -> 'a -> 'a
+ (** [fold f p a] folds f over the monomials of p with non-zero coefficient *)
+ val add : Monomial.t -> Num.num -> t -> t
(** [add m n p]
@return the polynomial n*m + p *)
- val add : Monomial.t -> Num.num -> t -> t
-
end
-type cstr = {coeffs : Vect.t ; op : op ; cst : Num.num} (* Representation of linear constraints *)
+type cstr = {coeffs : Vect.t; op : op; cst : Num.num}
+
+(* Representation of linear constraints *)
and op = Eq | Ge | Gt
val eval_op : op -> Num.num -> Num.num -> bool
(*val opMult : op -> op -> op*)
-val opAdd : op -> op -> op
+val opAdd : op -> op -> op
+val is_strict : cstr -> bool
(** [is_strict c]
@return whether the constraint is strict i.e. c.op = Gt *)
-val is_strict : cstr -> bool
exception Strict
@@ -141,65 +139,64 @@ module LinPoly : sig
This is done using the monomial tables of the module MonT. *)
module MonT : sig
- (** [clear ()] clears the mapping. *)
val clear : unit -> unit
+ (** [clear ()] clears the mapping. *)
+ val retrieve : int -> Monomial.t
(** [retrieve x]
@return the monomial corresponding to the variable [x] *)
- val retrieve : int -> Monomial.t
+ val register : Monomial.t -> int
(** [register m]
@return the variable index for the monomial m *)
- val register : Monomial.t -> int
-
end
- (** [linpol_of_pol p] linearise the polynomial p *)
val linpol_of_pol : Poly.t -> t
+ (** [linpol_of_pol p] linearise the polynomial p *)
+ val var : var -> t
(** [var x]
@return 1.y where y is the variable index of the monomial x^1.
*)
- val var : var -> t
+ val coq_poly_of_linpol : (Num.num -> 'a) -> t -> 'a Mc.pExpr
(** [coq_poly_of_linpol c p]
@param p is a multi-variate polynomial.
@param c maps a rational to a Coq polynomial coefficient.
@return the coq expression corresponding to polynomial [p].*)
- val coq_poly_of_linpol : (Num.num -> 'a) -> t -> 'a Mc.pExpr
+ val of_monomial : Monomial.t -> t
(** [of_monomial m]
@returns 1.x where x is the variable (index) for monomial m *)
- val of_monomial : Monomial.t -> t
- (** [of_vect v]
+ val of_vect : Vect.t -> t
+ (** [of_vect v]
@returns a1.x1 + ... + an.xn
This is not the identity because xi is the variable index of xi^1
*)
- val of_vect : Vect.t -> t
+ val variables : t -> ISet.t
(** [variables p]
@return the set of variables of the polynomial p
interpreted as a multi-variate polynomial *)
- val variables : t -> ISet.t
+ val is_variable : t -> var option
(** [is_variable p]
@return Some x if p = a.x for a >= 0 *)
- val is_variable : t -> var option
+ val is_linear : t -> bool
(** [is_linear p]
@return whether the multi-variate polynomial is linear. *)
- val is_linear : t -> bool
+ val is_linear_for : var -> t -> bool
(** [is_linear_for x p]
@return true if the polynomial is linear in x
i.e can be written c*x+r where c is a constant and r is independent from x *)
- val is_linear_for : var -> t -> bool
+ val constant : Num.num -> t
(** [constant c]
@return the constant polynomial c
*)
- val constant : Num.num -> t
(** [search_linear pred p]
@return a variable x such p = a.x + b such that
@@ -208,36 +205,34 @@ module LinPoly : sig
val search_linear : (Num.num -> bool) -> t -> var option
+ val search_all_linear : (Num.num -> bool) -> t -> var list
(** [search_all_linear pred p]
@return all the variables x such p = a.x + b such that
p is linear in x i.e x does not occur in b and
a is a constant such that [pred a] *)
- val search_all_linear : (Num.num -> bool) -> t -> var list
- (** [product p q]
- @return the product of the polynomial [p*q] *)
val product : t -> t -> t
+ (** [product p q]
+ @return the product of the polynomial [p*q] *)
+ val factorise : var -> t -> t * t
(** [factorise x p]
@return [a,b] such that [p = a.x + b]
and [x] does not occur in [b] *)
- val factorise : var -> t -> t * t
+ val collect_square : t -> Monomial.t MonMap.t
(** [collect_square p]
@return a mapping m such that m[s] = s^2
for every s^2 that is a monomial of [p] *)
- val collect_square : t -> Monomial.t MonMap.t
-
- (** [pp_var o v] pretty-prints a monomial indexed by v. *)
val pp_var : out_channel -> var -> unit
+ (** [pp_var o v] pretty-prints a monomial indexed by v. *)
- (** [pp o p] pretty-prints a polynomial. *)
val pp : out_channel -> t -> unit
+ (** [pp o p] pretty-prints a polynomial. *)
- (** [pp_goal typ o l] pretty-prints the list of constraints as a Coq goal. *)
val pp_goal : string -> out_channel -> (t * op) list -> unit
-
+ (** [pp_goal typ o l] pretty-prints the list of constraints as a Coq goal. *)
end
module ProofFormat : sig
@@ -252,7 +247,7 @@ module ProofFormat : sig
| Annot of string * prf_rule
| Hyp of int
| Def of int
- | Cst of Num.num
+ | Cst of Num.num
| Zero
| Square of Vect.t
| MulC of Vect.t * prf_rule
@@ -267,90 +262,77 @@ module ProofFormat : sig
| Enum of int * prf_rule * Vect.t * prf_rule * proof list
val pr_size : prf_rule -> Num.num
-
val pr_rule_max_id : prf_rule -> int
-
val proof_max_id : proof -> int
-
val normalise_proof : int -> proof -> int * proof
-
val output_prf_rule : out_channel -> prf_rule -> unit
-
val output_proof : out_channel -> proof -> unit
-
val add_proof : prf_rule -> prf_rule -> prf_rule
-
val mul_cst_proof : Num.num -> prf_rule -> prf_rule
-
val mul_proof : prf_rule -> prf_rule -> prf_rule
-
val compile_proof : int list -> proof -> Micromega.zArithProof
- val cmpl_prf_rule : ('a Micromega.pExpr -> 'a Micromega.pol) ->
- (Num.num -> 'a) -> (int list) -> prf_rule -> 'a Micromega.psatz
+ val cmpl_prf_rule :
+ ('a Micromega.pExpr -> 'a Micromega.pol)
+ -> (Num.num -> 'a)
+ -> int list
+ -> prf_rule
+ -> 'a Micromega.psatz
val proof_of_farkas : prf_rule IMap.t -> Vect.t -> prf_rule
-
val eval_prf_rule : (int -> LinPoly.t * op) -> prf_rule -> LinPoly.t * op
-
val eval_proof : (LinPoly.t * op) IMap.t -> proof -> bool
-
end
val output_cstr : out_channel -> cstr -> unit
-
val opMult : op -> op -> op
(** [module WithProof] constructs polynomials packed with the proof that their sign is correct. *)
-module WithProof :
-sig
-
+module WithProof : sig
type t = (LinPoly.t * op) * ProofFormat.prf_rule
- (** [InvalidProof] is raised if the operation is invalid. *)
exception InvalidProof
+ (** [InvalidProof] is raised if the operation is invalid. *)
val annot : string -> t -> t
-
val of_cstr : cstr * ProofFormat.prf_rule -> t
- (** [out_channel chan c] pretty-prints the constraint [c] over the channel [chan] *)
val output : out_channel -> t -> unit
+ (** [out_channel chan c] pretty-prints the constraint [c] over the channel [chan] *)
val output_sys : out_channel -> t list -> unit
- (** [zero] represents the tautology (0=0) *)
val zero : t
+ (** [zero] represents the tautology (0=0) *)
- (** [const n] represents the tautology (n>=0) *)
val const : Num.num -> t
+ (** [const n] represents the tautology (n>=0) *)
+ val product : t -> t -> t
(** [product p q]
@return the polynomial p*q with its sign and proof *)
- val product : t -> t -> t
+ val addition : t -> t -> t
(** [addition p q]
@return the polynomial p+q with its sign and proof *)
- val addition : t -> t -> t
+ val mult : LinPoly.t -> t -> t
(** [mult p q]
@return the polynomial p*q with its sign and proof.
@raise InvalidProof if p is not a constant and p is not an equality *)
- val mult : LinPoly.t -> t -> t
- (** [cutting_plane p] does integer reasoning and adjust the constant to be integral *)
val cutting_plane : t -> t option
+ (** [cutting_plane p] does integer reasoning and adjust the constant to be integral *)
+ val linear_pivot : t list -> t -> Vect.var -> t -> t option
(** [linear_pivot sys p x q]
@return the polynomial [q] where [x] is eliminated using the polynomial [p]
The pivoting operation is only defined if
- p is linear in x i.e p = a.x+b and x neither occurs in a and b
- The pivoting also requires some sign conditions for [a]
*)
- val linear_pivot : t list -> t -> Vect.var -> t -> t option
-
-(** [subst sys] performs the equivalent of the 'subst' tactic of Coq.
+ (** [subst sys] performs the equivalent of the 'subst' tactic of Coq.
For every p=0 \in sys such that p is linear in x with coefficient +/- 1
i.e. p = 0 <-> x = e and x \notin e.
Replace x by e in sys
@@ -361,12 +343,9 @@ sig
val subst : t list -> t list
- (** [subst1 sys] performs a single substitution *)
val subst1 : t list -> t list
+ (** [subst1 sys] performs a single substitution *)
val saturate_subst : bool -> t list -> t list
-
-
val is_substitution : bool -> t -> var option
-
end
diff --git a/plugins/micromega/simplex.ml b/plugins/micromega/simplex.ml
index 4c95e6da75..5bda268035 100644
--- a/plugins/micromega/simplex.ml
+++ b/plugins/micromega/simplex.ml
@@ -8,73 +8,62 @@
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
-(** A naive simplex *)
open Polynomial
+(** A naive simplex *)
+
open Num
+
(*open Util*)
open Mutils
-type ('a,'b) sum = Inl of 'a | Inr of 'b
+type ('a, 'b) sum = Inl of 'a | Inr of 'b
let debug = false
type iset = unit IMap.t
-type tableau = Vect.t IMap.t (** Mapping basic variables to their equation.
+type tableau = Vect.t IMap.t
+(** Mapping basic variables to their equation.
All variables >= than a threshold rst are restricted.*)
-module Restricted =
- struct
- type t =
- {
- base : int; (** All variables above [base] are restricted *)
- exc : int option (** Except [exc] which is currently optimised *)
- }
-
- let pp o {base;exc} =
- Printf.fprintf o ">= %a " LinPoly.pp_var base;
- match exc with
- | None ->Printf.fprintf o "-"
- | Some x ->Printf.fprintf o "-%a" LinPoly.pp_var base
-
- let is_exception (x:var) (r:t) =
- match r.exc with
- | None -> false
- | Some x' -> x = x'
-
- let restrict x rst =
- if is_exception x rst
- then
- {base = rst.base;exc= None}
- else failwith (Printf.sprintf "Cannot restrict %i" x)
+module Restricted = struct
+ type t =
+ { base : int (** All variables above [base] are restricted *)
+ ; exc : int option (** Except [exc] which is currently optimised *) }
+ let pp o {base; exc} =
+ Printf.fprintf o ">= %a " LinPoly.pp_var base;
+ match exc with
+ | None -> Printf.fprintf o "-"
+ | Some x -> Printf.fprintf o "-%a" LinPoly.pp_var base
- let is_restricted x r0 =
- x >= r0.base && not (is_exception x r0)
+ let is_exception (x : var) (r : t) =
+ match r.exc with None -> false | Some x' -> x = x'
- let make x = {base = x ; exc = None}
-
- let set_exc x rst = {base = rst.base ; exc = Some x}
-
- let fold rst f m acc =
- IMap.fold (fun k v acc ->
- if is_exception k rst then acc
- else f k v acc) (IMap.from rst.base m) acc
-
- end
+ let restrict x rst =
+ if is_exception x rst then {base = rst.base; exc = None}
+ else failwith (Printf.sprintf "Cannot restrict %i" x)
+ let is_restricted x r0 = x >= r0.base && not (is_exception x r0)
+ let make x = {base = x; exc = None}
+ let set_exc x rst = {base = rst.base; exc = Some x}
+ let fold rst f m acc =
+ IMap.fold
+ (fun k v acc -> if is_exception k rst then acc else f k v acc)
+ (IMap.from rst.base m) acc
+end
let pp_row o v = LinPoly.pp o v
-let output_tableau o t =
- IMap.iter (fun k v ->
- Printf.fprintf o "%a = %a\n" LinPoly.pp_var k pp_row v) t
+let output_tableau o t =
+ IMap.iter
+ (fun k v -> Printf.fprintf o "%a = %a\n" LinPoly.pp_var k pp_row v)
+ t
let output_vars o m =
IMap.iter (fun k _ -> Printf.fprintf o "%a " LinPoly.pp_var k) m
-
(** A tableau is feasible iff for every basic restricted variable xi,
we have ci>=0.
@@ -83,12 +72,10 @@ let output_vars o m =
if ci>=0.
*)
-
-let unfeasible (rst:Restricted.t) tbl =
- Restricted.fold rst (fun k v m ->
- if Vect.get_cst v >=/ Int 0 then m
- else IMap.add k () m) tbl IMap.empty
-
+let unfeasible (rst : Restricted.t) tbl =
+ Restricted.fold rst
+ (fun k v m -> if Vect.get_cst v >=/ Int 0 then m else IMap.add k () m)
+ tbl IMap.empty
let is_feasible rst tb = IMap.is_empty (unfeasible rst tb)
@@ -105,11 +92,10 @@ let is_feasible rst tb = IMap.is_empty (unfeasible rst tb)
*)
let is_maximised_vect rst v =
- Vect.for_all (fun xi ai ->
- if ai >/ Int 0
- then false
- else Restricted.is_restricted xi rst) v
-
+ Vect.for_all
+ (fun xi ai ->
+ if ai >/ Int 0 then false else Restricted.is_restricted xi rst)
+ v
(** [is_maximised rst v]
@return None if the variable is not maximised
@@ -117,10 +103,8 @@ let is_maximised_vect rst v =
*)
let is_maximised rst v =
try
- let (vl,v) = Vect.decomp_cst v in
- if is_maximised_vect rst v
- then Some vl
- else None
+ let vl, v = Vect.decomp_cst v in
+ if is_maximised_vect rst v then Some vl else None
with Not_found -> None
(** A variable xi is unbounded if for every
@@ -132,21 +116,13 @@ let is_maximised rst v =
violating a restriction.
*)
-
type result =
- | Max of num (** Maximum is reached *)
+ | Max of num (** Maximum is reached *)
| Ubnd of var (** Problem is unbounded *)
- | Feas (** Problem is feasible *)
-
-type pivot =
- | Done of result
- | Pivot of int * int * num
-
+ | Feas (** Problem is feasible *)
-
-
-type simplex =
- | Opt of tableau * result
+type pivot = Done of result | Pivot of int * int * num
+type simplex = Opt of tableau * result
(** For a row, x = ao.xo+...+ai.xi
a valid pivot variable is such that it can improve the value of xi.
@@ -156,15 +132,16 @@ type simplex =
This is the entering variable.
*)
-let rec find_pivot_column (rst:Restricted.t) (r:Vect.t) =
+let rec find_pivot_column (rst : Restricted.t) (r : Vect.t) =
match Vect.choose r with
| None -> failwith "find_pivot_column"
- | Some(xi,ai,r') -> if ai </ Int 0
- then if Restricted.is_restricted xi rst
- then find_pivot_column rst r' (* ai.xi cannot be improved *)
- else (xi, -1) (* r is not restricted, sign of ai does not matter *)
- else (* ai is positive, xi can be increased *)
- (xi,1)
+ | Some (xi, ai, r') ->
+ if ai </ Int 0 then
+ if Restricted.is_restricted xi rst then find_pivot_column rst r'
+ (* ai.xi cannot be improved *)
+ else (xi, -1) (* r is not restricted, sign of ai does not matter *)
+ else (* ai is positive, xi can be increased *)
+ (xi, 1)
(** Finding the variable leaving the basis is more subtle because we need to:
- increase the objective function
@@ -173,46 +150,46 @@ let rec find_pivot_column (rst:Restricted.t) (r:Vect.t) =
This explains why we choose the pivot with the smallest score
*)
-let min_score s (i1,sc1) =
+let min_score s (i1, sc1) =
match s with
- | None -> Some (i1,sc1)
- | Some(i0,sc0) ->
- if sc0 </ sc1 then s
- else if sc1 </ sc0 then Some (i1,sc1)
- else if i0 < i1 then s else Some(i1,sc1)
+ | None -> Some (i1, sc1)
+ | Some (i0, sc0) ->
+ if sc0 </ sc1 then s
+ else if sc1 </ sc0 then Some (i1, sc1)
+ else if i0 < i1 then s
+ else Some (i1, sc1)
let find_pivot_row rst tbl j sgn =
Restricted.fold rst
(fun i' v res ->
let aij = Vect.get j v in
- if (Int sgn) */ aij </ Int 0
- then (* This would improve *)
- let score' = Num.abs_num ((Vect.get_cst v) // aij) in
- min_score res (i',score')
- else res) tbl None
+ if Int sgn */ aij </ Int 0 then
+ (* This would improve *)
+ let score' = Num.abs_num (Vect.get_cst v // aij) in
+ min_score res (i', score')
+ else res)
+ tbl None
let safe_find err x t =
- try
- IMap.find x t
+ try IMap.find x t
with Not_found ->
- if debug
- then Printf.fprintf stdout "safe_find %s x%i %a\n" err x output_tableau t;
- failwith err
-
+ if debug then
+ Printf.fprintf stdout "safe_find %s x%i %a\n" err x output_tableau t;
+ failwith err
(** [find_pivot vr t] aims at improving the objective function of the basic variable vr *)
-let find_pivot vr (rst:Restricted.t) tbl =
+let find_pivot vr (rst : Restricted.t) tbl =
(* Get the objective of the basic variable vr *)
- let v = safe_find "find_pivot" vr tbl in
+ let v = safe_find "find_pivot" vr tbl in
match is_maximised rst v with
| Some mx -> Done (Max mx) (* Maximum is reached; we are done *)
- | None ->
- (* Extract the vector *)
- let (_,v) = Vect.decomp_cst v in
- let (j',sgn) = find_pivot_column rst v in
- match find_pivot_row rst (IMap.remove vr tbl) j' sgn with
- | None -> Done (Ubnd j')
- | Some (i',sc) -> Pivot(i', j', sc)
+ | None -> (
+ (* Extract the vector *)
+ let _, v = Vect.decomp_cst v in
+ let j', sgn = find_pivot_column rst v in
+ match find_pivot_row rst (IMap.remove vr tbl) j' sgn with
+ | None -> Done (Ubnd j')
+ | Some (i', sc) -> Pivot (i', j', sc) )
(** [solve_column c r e]
@param c is a non-basic variable
@@ -223,12 +200,11 @@ let find_pivot vr (rst:Restricted.t) tbl =
c = (r - e')/ai
*)
-let solve_column (c : var) (r : var) (e : Vect.t) : Vect.t =
+let solve_column (c : var) (r : var) (e : Vect.t) : Vect.t =
let a = Vect.get c e in
- if a =/ Int 0
- then failwith "Cannot solve column"
+ if a =/ Int 0 then failwith "Cannot solve column"
else
- let a' = (Int (-1) // a) in
+ let a' = Int (-1) // a in
Vect.mul a' (Vect.set r (Int (-1)) (Vect.set c (Int 0) e))
(** [pivot_row r c e]
@@ -236,439 +212,424 @@ let solve_column (c : var) (r : var) (e : Vect.t) : Vect.t =
@param r is a vector r = g.c + r'
@return g.e+r' *)
-let pivot_row (row: Vect.t) (c : var) (e : Vect.t) : Vect.t =
+let pivot_row (row : Vect.t) (c : var) (e : Vect.t) : Vect.t =
let g = Vect.get c row in
- if g =/ Int 0
- then row
- else Vect.mul_add g e (Int 1) (Vect.set c (Int 0) row)
+ if g =/ Int 0 then row else Vect.mul_add g e (Int 1) (Vect.set c (Int 0) row)
-let pivot_with (m : tableau) (v: var) (p : Vect.t) =
- IMap.map (fun (r:Vect.t) -> pivot_row r v p) m
+let pivot_with (m : tableau) (v : var) (p : Vect.t) =
+ IMap.map (fun (r : Vect.t) -> pivot_row r v p) m
let pivot (m : tableau) (r : var) (c : var) =
- let row = safe_find "pivot" r m in
+ let row = safe_find "pivot" r m in
let piv = solve_column c r row in
IMap.add c piv (pivot_with (IMap.remove r m) c piv)
-
let adapt_unbounded vr x rst tbl =
- if Vect.get_cst (IMap.find vr tbl) >=/ Int 0
- then tbl
- else pivot tbl vr x
+ if Vect.get_cst (IMap.find vr tbl) >=/ Int 0 then tbl else pivot tbl vr x
-module BaseSet = Set.Make(struct type t = iset let compare = IMap.compare (fun x y -> 0) end)
+module BaseSet = Set.Make (struct
+ type t = iset
+
+ let compare = IMap.compare (fun x y -> 0)
+end)
let get_base tbl = IMap.mapi (fun k _ -> ()) tbl
let simplex opt vr rst tbl =
let b = ref BaseSet.empty in
-
-let rec simplex opt vr rst tbl =
-
- if debug then begin
+ let rec simplex opt vr rst tbl =
+ ( if debug then
let base = get_base tbl in
- if BaseSet.mem base !b
- then Printf.fprintf stdout "Cycling detected\n"
- else b := BaseSet.add base !b
- end;
-
- if debug && not (is_feasible rst tbl)
- then
- begin
+ if BaseSet.mem base !b then Printf.fprintf stdout "Cycling detected\n"
+ else b := BaseSet.add base !b );
+ if debug && not (is_feasible rst tbl) then begin
let m = unfeasible rst tbl in
Printf.fprintf stdout "Simplex error\n";
Printf.fprintf stdout "The current tableau is not feasible\n";
- Printf.fprintf stdout "Restricted >= %a\n" Restricted.pp rst ;
+ Printf.fprintf stdout "Restricted >= %a\n" Restricted.pp rst;
output_tableau stdout tbl;
Printf.fprintf stdout "Error for variables %a\n" output_vars m
end;
-
- if not opt && (Vect.get_cst (IMap.find vr tbl) >=/ Int 0)
- then Opt(tbl,Feas)
- else
- match find_pivot vr rst tbl with
- | Done r ->
- begin match r with
- | Max _ -> Opt(tbl, r)
- | Ubnd x ->
+ if (not opt) && Vect.get_cst (IMap.find vr tbl) >=/ Int 0 then
+ Opt (tbl, Feas)
+ else
+ match find_pivot vr rst tbl with
+ | Done r -> (
+ match r with
+ | Max _ -> Opt (tbl, r)
+ | Ubnd x ->
let t' = adapt_unbounded vr x rst tbl in
- Opt(t',r)
- | Feas -> raise (Invalid_argument "find_pivot")
- end
- | Pivot(i,j,s) ->
- if debug then begin
- Printf.fprintf stdout "Find pivot for x%i(%s)\n" vr (string_of_num s);
- Printf.fprintf stdout "Leaving variable x%i\n" i;
- Printf.fprintf stdout "Entering variable x%i\n" j;
- end;
- let m' = pivot tbl i j in
- simplex opt vr rst m' in
-
-simplex opt vr rst tbl
-
-
+ Opt (t', r)
+ | Feas -> raise (Invalid_argument "find_pivot") )
+ | Pivot (i, j, s) ->
+ if debug then begin
+ Printf.fprintf stdout "Find pivot for x%i(%s)\n" vr (string_of_num s);
+ Printf.fprintf stdout "Leaving variable x%i\n" i;
+ Printf.fprintf stdout "Entering variable x%i\n" j
+ end;
+ let m' = pivot tbl i j in
+ simplex opt vr rst m'
+ in
+ simplex opt vr rst tbl
-type certificate =
- | Unsat of Vect.t
- | Sat of tableau * var option
+type certificate = Unsat of Vect.t | Sat of tableau * var option
(** [normalise_row t v]
@return a row obtained by pivoting the basic variables of the vector v
*)
-let normalise_row (t : tableau) (v: Vect.t) =
- Vect.fold (fun acc vr ai -> try
+let normalise_row (t : tableau) (v : Vect.t) =
+ Vect.fold
+ (fun acc vr ai ->
+ try
let e = IMap.find vr t in
Vect.add (Vect.mul ai e) acc
with Not_found -> Vect.add (Vect.set vr ai Vect.null) acc)
Vect.null v
-let normalise_row (t : tableau) (v: Vect.t) =
+let normalise_row (t : tableau) (v : Vect.t) =
let v' = normalise_row t v in
if debug then Printf.fprintf stdout "Normalised Optimising %a\n" LinPoly.pp v';
v'
-let add_row (nw :var) (t : tableau) (v : Vect.t) : tableau =
+let add_row (nw : var) (t : tableau) (v : Vect.t) : tableau =
IMap.add nw (normalise_row t v) t
-
-
(** [push_real] performs reasoning over the rationals *)
-let push_real (opt : bool) (nw : var) (v : Vect.t) (rst: Restricted.t) (t : tableau) : certificate =
- if debug
- then begin Printf.fprintf stdout "Current Tableau\n%a\n" output_tableau t;
- Printf.fprintf stdout "Optimising %a=%a\n" LinPoly.pp_var nw LinPoly.pp v
- end;
+let push_real (opt : bool) (nw : var) (v : Vect.t) (rst : Restricted.t)
+ (t : tableau) : certificate =
+ if debug then begin
+ Printf.fprintf stdout "Current Tableau\n%a\n" output_tableau t;
+ Printf.fprintf stdout "Optimising %a=%a\n" LinPoly.pp_var nw LinPoly.pp v
+ end;
match simplex opt nw rst (add_row nw t v) with
- | Opt(t',r) -> (* Look at the optimal *)
- match r with
- | Ubnd x->
- if debug then Printf.printf "The objective is unbounded (variable %a)\n" LinPoly.pp_var x;
- Sat (t',Some x) (* This is sat and we can extract a value *)
- | Feas -> Sat (t',None)
- | Max n ->
- if debug then begin
- Printf.printf "The objective is maximised %s\n" (string_of_num n);
- Printf.printf "%a = %a\n" LinPoly.pp_var nw pp_row (IMap.find nw t')
- end;
-
- if n >=/ Int 0
- then Sat (t',None)
- else
- let v' = safe_find "push_real" nw t' in
- Unsat (Vect.set nw (Int 1) (Vect.set 0 (Int 0) (Vect.mul (Int (-1)) v')))
-
+ | Opt (t', r) -> (
+ (* Look at the optimal *)
+ match r with
+ | Ubnd x ->
+ if debug then
+ Printf.printf "The objective is unbounded (variable %a)\n"
+ LinPoly.pp_var x;
+ Sat (t', Some x) (* This is sat and we can extract a value *)
+ | Feas -> Sat (t', None)
+ | Max n ->
+ if debug then begin
+ Printf.printf "The objective is maximised %s\n" (string_of_num n);
+ Printf.printf "%a = %a\n" LinPoly.pp_var nw pp_row (IMap.find nw t')
+ end;
+ if n >=/ Int 0 then Sat (t', None)
+ else
+ let v' = safe_find "push_real" nw t' in
+ Unsat
+ (Vect.set nw (Int 1) (Vect.set 0 (Int 0) (Vect.mul (Int (-1)) v'))) )
+open Mutils
(** One complication is that equalities needs some pre-processing.
*)
-open Mutils
-open Polynomial
-let fresh_var l =
- 1 +
- try
- (ISet.max_elt (List.fold_left (fun acc c -> ISet.union acc (Vect.variables c.coeffs)) ISet.empty l))
- with Not_found -> 0
+open Polynomial
+let fresh_var l =
+ 1
+ +
+ try
+ ISet.max_elt
+ (List.fold_left
+ (fun acc c -> ISet.union acc (Vect.variables c.coeffs))
+ ISet.empty l)
+ with Not_found -> 0
(*type varmap = (int * bool) IMap.t*)
-
let make_certificate vm l =
- Vect.normalise (Vect.fold (fun acc x n ->
- let (x',b) = IMap.find x vm in
- Vect.set x' (if b then n else Num.minus_num n) acc) Vect.null l)
-
-
-
-
-
-let eliminate_equalities (vr0:var) (l:Polynomial.cstr list) =
+ Vect.normalise
+ (Vect.fold
+ (fun acc x n ->
+ let x', b = IMap.find x vm in
+ Vect.set x' (if b then n else Num.minus_num n) acc)
+ Vect.null l)
+
+let eliminate_equalities (vr0 : var) (l : Polynomial.cstr list) =
let rec elim idx vr vm l acc =
match l with
- | [] -> (vr,vm,acc)
- | c::l -> match c.op with
- | Ge -> let v = Vect.set 0 (minus_num c.cst) c.coeffs in
- elim (idx+1) (vr+1) (IMap.add vr (idx,true) vm) l ((vr,v)::acc)
- | Eq -> let v1 = Vect.set 0 (minus_num c.cst) c.coeffs in
- let v2 = Vect.mul (Int (-1)) v1 in
- let vm = IMap.add vr (idx,true) (IMap.add (vr+1) (idx,false) vm) in
- elim (idx+1) (vr+2) vm l ((vr,v1)::(vr+1,v2)::acc)
- | Gt -> raise Strict in
+ | [] -> (vr, vm, acc)
+ | c :: l -> (
+ match c.op with
+ | Ge ->
+ let v = Vect.set 0 (minus_num c.cst) c.coeffs in
+ elim (idx + 1) (vr + 1) (IMap.add vr (idx, true) vm) l ((vr, v) :: acc)
+ | Eq ->
+ let v1 = Vect.set 0 (minus_num c.cst) c.coeffs in
+ let v2 = Vect.mul (Int (-1)) v1 in
+ let vm = IMap.add vr (idx, true) (IMap.add (vr + 1) (idx, false) vm) in
+ elim (idx + 1) (vr + 2) vm l ((vr, v1) :: (vr + 1, v2) :: acc)
+ | Gt -> raise Strict )
+ in
elim 0 vr0 IMap.empty l []
let find_solution rst tbl =
- IMap.fold (fun vr v res -> if Restricted.is_restricted vr rst
- then res
- else Vect.set vr (Vect.get_cst v) res) tbl Vect.null
+ IMap.fold
+ (fun vr v res ->
+ if Restricted.is_restricted vr rst then res
+ else Vect.set vr (Vect.get_cst v) res)
+ tbl Vect.null
-let choose_conflict (sol:Vect.t) (l: (var * Vect.t) list) =
+let choose_conflict (sol : Vect.t) (l : (var * Vect.t) list) =
let esol = Vect.set 0 (Int 1) sol in
-
- let rec most_violating l e (x,v) rst =
+ let rec most_violating l e (x, v) rst =
match l with
- | [] -> Some((x,v),rst)
- | (x',v')::l ->
- let e' = Vect.dotproduct esol v' in
- if e' <=/ e
- then most_violating l e' (x',v') ((x,v)::rst)
- else most_violating l e (x,v) ((x',v')::rst) in
-
+ | [] -> Some ((x, v), rst)
+ | (x', v') :: l ->
+ let e' = Vect.dotproduct esol v' in
+ if e' <=/ e then most_violating l e' (x', v') ((x, v) :: rst)
+ else most_violating l e (x, v) ((x', v') :: rst)
+ in
match l with
| [] -> None
- | (x,v)::l -> let e = Vect.dotproduct esol v in
- most_violating l e (x,v) []
-
+ | (x, v) :: l ->
+ let e = Vect.dotproduct esol v in
+ most_violating l e (x, v) []
-
-let rec solve opt l (rst:Restricted.t) (t:tableau) =
+let rec solve opt l (rst : Restricted.t) (t : tableau) =
let sol = find_solution rst t in
match choose_conflict sol l with
- | None -> Inl (rst,t,None)
- | Some((vr,v),l) ->
- match push_real opt vr v (Restricted.set_exc vr rst) t with
- | Sat (t',x) ->
- (* let t' = remove_redundant rst t' in*)
- begin
- match l with
- | [] -> Inl(rst,t', x)
- | _ -> solve opt l rst t'
- end
- | Unsat c -> Inr c
-
-let find_unsat_certificate (l : Polynomial.cstr list ) =
+ | None -> Inl (rst, t, None)
+ | Some ((vr, v), l) -> (
+ match push_real opt vr v (Restricted.set_exc vr rst) t with
+ | Sat (t', x) -> (
+ (* let t' = remove_redundant rst t' in*)
+ match l with
+ | [] -> Inl (rst, t', x)
+ | _ -> solve opt l rst t' )
+ | Unsat c -> Inr c )
+
+let find_unsat_certificate (l : Polynomial.cstr list) =
let vr = fresh_var l in
- let (_,vm,l') = eliminate_equalities vr l in
-
- match solve false l' (Restricted.make vr) IMap.empty with
- | Inr c -> Some (make_certificate vm c)
+ let _, vm, l' = eliminate_equalities vr l in
+ match solve false l' (Restricted.make vr) IMap.empty with
+ | Inr c -> Some (make_certificate vm c)
| Inl _ -> None
-
-
let find_point (l : Polynomial.cstr list) =
let vr = fresh_var l in
- let (_,vm,l') = eliminate_equalities vr l in
-
+ let _, vm, l' = eliminate_equalities vr l in
match solve false l' (Restricted.make vr) IMap.empty with
- | Inl (rst,t,_) -> Some (find_solution rst t)
- | _ -> None
-
-
+ | Inl (rst, t, _) -> Some (find_solution rst t)
+ | _ -> None
let optimise obj l =
let vr0 = fresh_var l in
- let (_,vm,l') = eliminate_equalities (vr0+1) l in
-
+ let _, vm, l' = eliminate_equalities (vr0 + 1) l in
let bound pos res =
match res with
- | Opt(_,Max n) -> Some (if pos then n else minus_num n)
- | Opt(_,Ubnd _) -> None
- | Opt(_,Feas) -> None
+ | Opt (_, Max n) -> Some (if pos then n else minus_num n)
+ | Opt (_, Ubnd _) -> None
+ | Opt (_, Feas) -> None
in
-
match solve false l' (Restricted.make vr0) IMap.empty with
- | Inl (rst,t,_) ->
- Some (bound false
- (simplex true vr0 rst (add_row vr0 t (Vect.uminus obj))),
- bound true
- (simplex true vr0 rst (add_row vr0 t obj)))
- | _ -> None
-
-
+ | Inl (rst, t, _) ->
+ Some
+ ( bound false (simplex true vr0 rst (add_row vr0 t (Vect.uminus obj)))
+ , bound true (simplex true vr0 rst (add_row vr0 t obj)) )
+ | _ -> None
open Polynomial
let env_of_list l =
- List.fold_left (fun (i,m) l -> (i+1, IMap.add i l m)) (0,IMap.empty) l
-
+ List.fold_left (fun (i, m) l -> (i + 1, IMap.add i l m)) (0, IMap.empty) l
open ProofFormat
-let make_farkas_certificate (env: WithProof.t IMap.t) vm v =
- Vect.fold (fun acc x n ->
+let make_farkas_certificate (env : WithProof.t IMap.t) vm v =
+ Vect.fold
+ (fun acc x n ->
add_proof acc
begin
try
- let (x',b) = IMap.find x vm in
- (mul_cst_proof
- (if b then n else (Num.minus_num n))
- (snd (IMap.find x' env)))
- with Not_found -> (* This is an introduced hypothesis *)
- (mul_cst_proof n (snd (IMap.find x env)))
- end) Zero v
-
-let make_farkas_proof (env: WithProof.t IMap.t) vm v =
- Vect.fold (fun wp x n ->
- WithProof.addition wp begin
+ let x', b = IMap.find x vm in
+ mul_cst_proof
+ (if b then n else Num.minus_num n)
+ (snd (IMap.find x' env))
+ with Not_found ->
+ (* This is an introduced hypothesis *)
+ mul_cst_proof n (snd (IMap.find x env))
+ end)
+ Zero v
+
+let make_farkas_proof (env : WithProof.t IMap.t) vm v =
+ Vect.fold
+ (fun wp x n ->
+ WithProof.addition wp
+ begin
try
- let (x', b) = IMap.find x vm in
- let n = if b then n else Num.minus_num n in
+ let x', b = IMap.find x vm in
+ let n = if b then n else Num.minus_num n in
WithProof.mult (Vect.cst n) (IMap.find x' env)
- with Not_found ->
- WithProof.mult (Vect.cst n) (IMap.find x env)
- end) WithProof.zero v
-
+ with Not_found -> WithProof.mult (Vect.cst n) (IMap.find x env)
+ end)
+ WithProof.zero v
let frac_num n = n -/ Num.floor_num n
-
(* [resolv_var v rst tbl] returns (if it exists) a restricted variable vr such that v = vr *)
exception FoundVar of int
let resolve_var v rst tbl =
- let v = Vect.set v (Int 1) Vect.null in
+ let v = Vect.set v (Int 1) Vect.null in
try
- IMap.iter (fun k vect ->
- if Restricted.is_restricted k rst
- then if Vect.equal v vect then raise (FoundVar k)
- else ()) tbl ; None
+ IMap.iter
+ (fun k vect ->
+ if Restricted.is_restricted k rst then
+ if Vect.equal v vect then raise (FoundVar k) else ())
+ tbl;
+ None
with FoundVar k -> Some k
let prepare_cut env rst tbl x v =
(* extract the unrestricted part *)
- let (unrst,rstv) = Vect.partition (fun x vl -> not (Restricted.is_restricted x rst) && frac_num vl <>/ Int 0) (Vect.set 0 (Int 0) v) in
- if Vect.is_null unrst
- then Some rstv
- else Some (Vect.fold (fun acc k i ->
- match resolve_var k rst tbl with
- | None -> acc (* Should not happen *)
- | Some v' -> Vect.set v' i acc)
- rstv unrst)
-
-let cut env rmin sol vm (rst:Restricted.t) tbl (x,v) =
- begin
- (* Printf.printf "Trying to cut %i\n" x;*)
- let (n,r) = Vect.decomp_cst v in
-
-
+ let unrst, rstv =
+ Vect.partition
+ (fun x vl ->
+ (not (Restricted.is_restricted x rst)) && frac_num vl <>/ Int 0)
+ (Vect.set 0 (Int 0) v)
+ in
+ if Vect.is_null unrst then Some rstv
+ else
+ Some
+ (Vect.fold
+ (fun acc k i ->
+ match resolve_var k rst tbl with
+ | None -> acc (* Should not happen *)
+ | Some v' -> Vect.set v' i acc)
+ rstv unrst)
+
+let cut env rmin sol vm (rst : Restricted.t) tbl (x, v) =
+ (* Printf.printf "Trying to cut %i\n" x;*)
+ let n, r = Vect.decomp_cst v in
let f = frac_num n in
-
- if f =/ Int 0
- then None (* The solution is integral *)
+ if f =/ Int 0 then None (* The solution is integral *)
else
(* This is potentially a cut *)
- let t =
- if f </ (Int 1) // (Int 2)
- then
- let t' = ((Int 1) // f) in
- if Num.is_integer_num t'
- then t' -/ Int 1
- else Num.floor_num t'
- else Int 1 in
-
+ let t =
+ if f </ Int 1 // Int 2 then
+ let t' = Int 1 // f in
+ if Num.is_integer_num t' then t' -/ Int 1 else Num.floor_num t'
+ else Int 1
+ in
let cut_coeff1 v =
let fv = frac_num v in
- if fv <=/ (Int 1 -/ f)
- then fv // (Int 1 -/ f)
- else (Int 1 -/ fv) // f in
-
+ if fv <=/ Int 1 -/ f then fv // (Int 1 -/ f) else (Int 1 -/ fv) // f
+ in
let cut_coeff2 v = frac_num (t */ v) in
-
let cut_vector ccoeff =
match prepare_cut env rst tbl x v with
| None -> Vect.null
| Some r ->
- (*Printf.printf "Potential cut %a\n" LinPoly.pp r;*)
- Vect.fold (fun acc x n -> Vect.set x (ccoeff n) acc) Vect.null r
+ (*Printf.printf "Potential cut %a\n" LinPoly.pp r;*)
+ Vect.fold (fun acc x n -> Vect.set x (ccoeff n) acc) Vect.null r
in
-
- let lcut = List.map (fun cv -> Vect.normalise (cut_vector cv)) [cut_coeff1 ; cut_coeff2] in
-
- let lcut = List.map (make_farkas_proof env vm) lcut in
-
+ let lcut =
+ List.map
+ (fun cv -> Vect.normalise (cut_vector cv))
+ [cut_coeff1; cut_coeff2]
+ in
+ let lcut = List.map (make_farkas_proof env vm) lcut in
let check_cutting_plane c =
match WithProof.cutting_plane c with
| None ->
- if debug then Printf.printf "This is not cutting plane for %a\n%a:" LinPoly.pp_var x WithProof.output c;
- None
- | Some(v,prf) ->
- if debug then begin
- Printf.printf "This is a cutting plane for %a:" LinPoly.pp_var x;
- Printf.printf " %a\n" WithProof.output (v,prf);
- end;
- if (=) (snd v) Eq
- then (* Unsat *) Some (x,(v,prf))
- else
- let vl = (Vect.dotproduct (fst v) (Vect.set 0 (Int 1) sol)) in
- if eval_op Ge vl (Int 0)
- then begin
- (* Can this happen? *)
- if debug then Printf.printf "The cut is feasible %s >= 0 ??\n" (Num.string_of_num vl);
- None
- end
- else Some(x,(v,prf)) in
-
+ if debug then
+ Printf.printf "This is not cutting plane for %a\n%a:" LinPoly.pp_var x
+ WithProof.output c;
+ None
+ | Some (v, prf) ->
+ if debug then begin
+ Printf.printf "This is a cutting plane for %a:" LinPoly.pp_var x;
+ Printf.printf " %a\n" WithProof.output (v, prf)
+ end;
+ if snd v = Eq then (* Unsat *) Some (x, (v, prf))
+ else
+ let vl = Vect.dotproduct (fst v) (Vect.set 0 (Int 1) sol) in
+ if eval_op Ge vl (Int 0) then begin
+ (* Can this happen? *)
+ if debug then
+ Printf.printf "The cut is feasible %s >= 0 ??\n"
+ (Num.string_of_num vl);
+ None
+ end
+ else Some (x, (v, prf))
+ in
find_some check_cutting_plane lcut
- end
let find_cut nb env u sol vm rst tbl =
- if nb = 0
- then
- IMap.fold (fun x v acc ->
- match acc with
- | None -> cut env u sol vm rst tbl (x,v)
- | Some c -> Some c) tbl None
+ if nb = 0 then
+ IMap.fold
+ (fun x v acc ->
+ match acc with
+ | None -> cut env u sol vm rst tbl (x, v)
+ | Some c -> Some c)
+ tbl None
else
- IMap.fold (fun x v acc ->
- match cut env u sol vm rst tbl (x,v) , acc with
- | None , Some r | Some r , None -> Some r
- | None , None -> None
- | Some (v,((lp,o),p1)) , Some (v',((lp',o'),p2)) ->
- Some (if ProofFormat.pr_size p1 </ ProofFormat.pr_size p2
- then (v,((lp,o),p1)) else (v',((lp',o'),p2)))
- ) tbl None
-
-
+ IMap.fold
+ (fun x v acc ->
+ match (cut env u sol vm rst tbl (x, v), acc) with
+ | None, Some r | Some r, None -> Some r
+ | None, None -> None
+ | Some (v, ((lp, o), p1)), Some (v', ((lp', o'), p2)) ->
+ Some
+ ( if ProofFormat.pr_size p1 </ ProofFormat.pr_size p2 then
+ (v, ((lp, o), p1))
+ else (v', ((lp', o'), p2)) ))
+ tbl None
let integer_solver lp =
- let (l,_) = List.split lp in
+ let l, _ = List.split lp in
let vr0 = fresh_var l in
- let (vr,vm,l') = eliminate_equalities vr0 l in
-
- let _,env = env_of_list (List.map WithProof.of_cstr lp) in
-
+ let vr, vm, l' = eliminate_equalities vr0 l in
+ let _, env = env_of_list (List.map WithProof.of_cstr lp) in
let insert_row vr v rst tbl =
match push_real true vr v rst tbl with
- | Sat (t',x) -> Inl (Restricted.restrict vr rst,t',x)
- | Unsat c -> Inr c in
-
+ | Sat (t', x) -> Inl (Restricted.restrict vr rst, t', x)
+ | Unsat c -> Inr c
+ in
let nb = ref 0 in
-
let rec isolve env cr vr res =
incr nb;
match res with
- | Inr c -> Some (Step (vr, make_farkas_certificate env vm (Vect.normalise c),Done))
- | Inl (rst,tbl,x) ->
- if debug then begin
- Printf.fprintf stdout "Looking for a cut\n";
- Printf.fprintf stdout "Restricted %a ...\n" Restricted.pp rst;
- Printf.fprintf stdout "Current Tableau\n%a\n" output_tableau tbl;
- (* Printf.fprintf stdout "Bounding box\n%a\n" output_box (bounding_box (vr+1) rst tbl l')*)
- end;
- let sol = find_solution rst tbl in
-
- match find_cut (!nb mod 2) env cr (*x*) sol vm rst tbl with
- | None -> None
- | Some(cr,((v,op),cut)) ->
- if (=) op Eq
- then (* This is a contradiction *)
- Some(Step(vr,CutPrf cut, Done))
- else
- let res = insert_row vr v (Restricted.set_exc vr rst) tbl in
- let prf = isolve (IMap.add vr ((v,op),Def vr) env) (Some cr) (vr+1) res in
- match prf with
- | None -> None
- | Some p -> Some (Step(vr,CutPrf cut,p)) in
-
+ | Inr c ->
+ Some (Step (vr, make_farkas_certificate env vm (Vect.normalise c), Done))
+ | Inl (rst, tbl, x) -> (
+ if debug then begin
+ Printf.fprintf stdout "Looking for a cut\n";
+ Printf.fprintf stdout "Restricted %a ...\n" Restricted.pp rst;
+ Printf.fprintf stdout "Current Tableau\n%a\n" output_tableau tbl
+ (* Printf.fprintf stdout "Bounding box\n%a\n" output_box (bounding_box (vr+1) rst tbl l')*)
+ end;
+ let sol = find_solution rst tbl in
+ match find_cut (!nb mod 2) env cr (*x*) sol vm rst tbl with
+ | None -> None
+ | Some (cr, ((v, op), cut)) -> (
+ if op = Eq then
+ (* This is a contradiction *)
+ Some (Step (vr, CutPrf cut, Done))
+ else
+ let res = insert_row vr v (Restricted.set_exc vr rst) tbl in
+ let prf =
+ isolve (IMap.add vr ((v, op), Def vr) env) (Some cr) (vr + 1) res
+ in
+ match prf with
+ | None -> None
+ | Some p -> Some (Step (vr, CutPrf cut, p)) ) )
+ in
let res = solve true l' (Restricted.make vr0) IMap.empty in
isolve env None vr res
let integer_solver lp =
- if debug then Printf.printf "Input integer solver\n%a\n" WithProof.output_sys (List.map WithProof.of_cstr lp);
-
+ if debug then
+ Printf.printf "Input integer solver\n%a\n" WithProof.output_sys
+ (List.map WithProof.of_cstr lp);
match integer_solver lp with
| None -> None
- | Some prf -> if debug
- then Printf.fprintf stdout "Proof %a\n" ProofFormat.output_proof prf ;
- Some prf
+ | Some prf ->
+ if debug then
+ Printf.fprintf stdout "Proof %a\n" ProofFormat.output_proof prf;
+ Some prf
diff --git a/plugins/micromega/simplex.mli b/plugins/micromega/simplex.mli
index cba8e94ea7..19bcce3590 100644
--- a/plugins/micromega/simplex.mli
+++ b/plugins/micromega/simplex.mli
@@ -10,9 +10,8 @@
open Polynomial
val optimise : Vect.t -> cstr list -> (Num.num option * Num.num option) option
-
val find_point : cstr list -> Vect.t option
-
val find_unsat_certificate : cstr list -> Vect.t option
-val integer_solver : (cstr * ProofFormat.prf_rule) list -> ProofFormat.proof option
+val integer_solver :
+ (cstr * ProofFormat.prf_rule) list -> ProofFormat.proof option
diff --git a/plugins/micromega/sos.ml b/plugins/micromega/sos.ml
index f2dfaa42a5..772ed7a8c5 100644
--- a/plugins/micromega/sos.ml
+++ b/plugins/micromega/sos.ml
@@ -9,17 +9,17 @@
(* ========================================================================= *)
(* Nonlinear universal reals procedure using SOS decomposition. *)
(* ========================================================================= *)
-open Num;;
-open Sos_types;;
-open Sos_lib;;
+open Num
+open Sos_types
+open Sos_lib
(*
prioritize_real();;
*)
-let debugging = ref false;;
+let debugging = ref false
-exception Sanity;;
+exception Sanity
(* ------------------------------------------------------------------------- *)
(* Turn a rational into a decimal string with d sig digits. *)
@@ -29,228 +29,224 @@ let decimalize =
let rec normalize y =
if abs_num y </ Int 1 // Int 10 then normalize (Int 10 */ y) - 1
else if abs_num y >=/ Int 1 then normalize (y // Int 10) + 1
- else 0 in
+ else 0
+ in
fun d x ->
- if x =/ Int 0 then "0.0" else
- let y = abs_num x in
- let e = normalize y in
- let z = pow10(-e) */ y +/ Int 1 in
- let k = round_num(pow10 d */ z) in
- (if x </ Int 0 then "-0." else "0.") ^
- implode(List.tl(explode(string_of_num k))) ^
- (if e = 0 then "" else "e"^string_of_int e);;
+ if x =/ Int 0 then "0.0"
+ else
+ let y = abs_num x in
+ let e = normalize y in
+ let z = (pow10 (-e) */ y) +/ Int 1 in
+ let k = round_num (pow10 d */ z) in
+ (if x </ Int 0 then "-0." else "0.")
+ ^ implode (List.tl (explode (string_of_num k)))
+ ^ if e = 0 then "" else "e" ^ string_of_int e
(* ------------------------------------------------------------------------- *)
(* Iterations over numbers, and lists indexed by numbers. *)
(* ------------------------------------------------------------------------- *)
let rec itern k l f a =
- match l with
- [] -> a
- | h::t -> itern (k + 1) t f (f h k a);;
+ match l with [] -> a | h :: t -> itern (k + 1) t f (f h k a)
-let rec iter (m,n) f a =
- if n < m then a
- else iter (m+1,n) f (f m a);;
+let rec iter (m, n) f a = if n < m then a else iter (m + 1, n) f (f m a)
(* ------------------------------------------------------------------------- *)
(* The main types. *)
(* ------------------------------------------------------------------------- *)
-type vector = int*(int,num)func;;
-
-type matrix = (int*int)*(int*int,num)func;;
-
-type monomial = (vname,int)func;;
-
-type poly = (monomial,num)func;;
+type vector = int * (int, num) func
+type matrix = (int * int) * (int * int, num) func
+type monomial = (vname, int) func
+type poly = (monomial, num) func
(* ------------------------------------------------------------------------- *)
(* Assignment avoiding zeros. *)
(* ------------------------------------------------------------------------- *)
-let (|-->) x y a = if y =/ Int 0 then a else (x |-> y) a;;
+let ( |--> ) x y a = if y =/ Int 0 then a else (x |-> y) a
(* ------------------------------------------------------------------------- *)
(* This can be generic. *)
(* ------------------------------------------------------------------------- *)
-let element (d,v) i = tryapplyd v i (Int 0);;
-
-let mapa f (d,v) =
- d,foldl (fun a i c -> (i |--> f(c)) a) undefined v;;
-
-let is_zero (d,v) =
- match v with
- Empty -> true
- | _ -> false;;
+let element (d, v) i = tryapplyd v i (Int 0)
+let mapa f (d, v) = (d, foldl (fun a i c -> (i |--> f c) a) undefined v)
+let is_zero (d, v) = match v with Empty -> true | _ -> false
(* ------------------------------------------------------------------------- *)
(* Vectors. Conventionally indexed 1..n. *)
(* ------------------------------------------------------------------------- *)
-let vector_0 n = (n,undefined:vector);;
-
-let dim (v:vector) = fst v;;
+let vector_0 n = ((n, undefined) : vector)
+let dim (v : vector) = fst v
let vector_const c n =
if c =/ Int 0 then vector_0 n
- else (n,List.fold_right (fun k -> k |-> c) (1--n) undefined :vector);;
+ else ((n, List.fold_right (fun k -> k |-> c) (1 -- n) undefined) : vector)
-let vector_cmul c (v:vector) =
+let vector_cmul c (v : vector) =
let n = dim v in
- if c =/ Int 0 then vector_0 n
- else n,mapf (fun x -> c */ x) (snd v)
+ if c =/ Int 0 then vector_0 n else (n, mapf (fun x -> c */ x) (snd v))
let vector_of_list l =
let n = List.length l in
- (n,List.fold_right2 (|->) (1--n) l undefined :vector);;
+ ((n, List.fold_right2 ( |-> ) (1 -- n) l undefined) : vector)
(* ------------------------------------------------------------------------- *)
(* Matrices; again rows and columns indexed from 1. *)
(* ------------------------------------------------------------------------- *)
-let matrix_0 (m,n) = ((m,n),undefined:matrix);;
-
-let dimensions (m:matrix) = fst m;;
+let matrix_0 (m, n) = (((m, n), undefined) : matrix)
+let dimensions (m : matrix) = fst m
-let matrix_cmul c (m:matrix) =
- let (i,j) = dimensions m in
- if c =/ Int 0 then matrix_0 (i,j)
- else (i,j),mapf (fun x -> c */ x) (snd m);;
+let matrix_cmul c (m : matrix) =
+ let i, j = dimensions m in
+ if c =/ Int 0 then matrix_0 (i, j)
+ else ((i, j), mapf (fun x -> c */ x) (snd m))
-let matrix_neg (m:matrix) = (dimensions m,mapf minus_num (snd m) :matrix);;
+let matrix_neg (m : matrix) = ((dimensions m, mapf minus_num (snd m)) : matrix)
-let matrix_add (m1:matrix) (m2:matrix) =
+let matrix_add (m1 : matrix) (m2 : matrix) =
let d1 = dimensions m1 and d2 = dimensions m2 in
if d1 <> d2 then failwith "matrix_add: incompatible dimensions"
- else (d1,combine (+/) (fun x -> x =/ Int 0) (snd m1) (snd m2) :matrix);;
-
-let row k (m:matrix) =
- let i,j = dimensions m in
- (j,
- foldl (fun a (i,j) c -> if i = k then (j |-> c) a else a) undefined (snd m)
- : vector);;
-
-let column k (m:matrix) =
- let i,j = dimensions m in
- (i,
- foldl (fun a (i,j) c -> if j = k then (i |-> c) a else a) undefined (snd m)
- : vector);;
-
-let diagonal (v:vector) =
+ else ((d1, combine ( +/ ) (fun x -> x =/ Int 0) (snd m1) (snd m2)) : matrix)
+
+let row k (m : matrix) =
+ let i, j = dimensions m in
+ ( ( j
+ , foldl
+ (fun a (i, j) c -> if i = k then (j |-> c) a else a)
+ undefined (snd m) )
+ : vector )
+
+let column k (m : matrix) =
+ let i, j = dimensions m in
+ ( ( i
+ , foldl
+ (fun a (i, j) c -> if j = k then (i |-> c) a else a)
+ undefined (snd m) )
+ : vector )
+
+let diagonal (v : vector) =
let n = dim v in
- ((n,n),foldl (fun a i c -> ((i,i) |-> c) a) undefined (snd v) : matrix);;
+ (((n, n), foldl (fun a i c -> ((i, i) |-> c) a) undefined (snd v)) : matrix)
(* ------------------------------------------------------------------------- *)
(* Monomials. *)
(* ------------------------------------------------------------------------- *)
-let monomial_1 = (undefined:monomial);;
-
-let monomial_var x = (x |=> 1 :monomial);;
+let monomial_1 = (undefined : monomial)
+let monomial_var x = (x |=> 1 : monomial)
-let (monomial_mul:monomial->monomial->monomial) =
- combine (+) (fun x -> false);;
+let (monomial_mul : monomial -> monomial -> monomial) =
+ combine ( + ) (fun x -> false)
-let monomial_degree x (m:monomial) = tryapplyd m x 0;;
-
-let monomial_multidegree (m:monomial) = foldl (fun a x k -> k + a) 0 m;;
-
-let monomial_variables m = dom m;;
+let monomial_degree x (m : monomial) = tryapplyd m x 0
+let monomial_multidegree (m : monomial) = foldl (fun a x k -> k + a) 0 m
+let monomial_variables m = dom m
(* ------------------------------------------------------------------------- *)
(* Polynomials. *)
(* ------------------------------------------------------------------------- *)
-let poly_0 = (undefined:poly);;
-
-let poly_isconst (p:poly) = foldl (fun a m c -> m = monomial_1 && a) true p;;
-
-let poly_var x = ((monomial_var x) |=> Int 1 :poly);;
+let poly_0 = (undefined : poly)
+let poly_isconst (p : poly) = foldl (fun a m c -> m = monomial_1 && a) true p
+let poly_var x = (monomial_var x |=> Int 1 : poly)
+let poly_const c = if c =/ Int 0 then poly_0 else monomial_1 |=> c
-let poly_const c =
- if c =/ Int 0 then poly_0 else (monomial_1 |=> c);;
+let poly_cmul c (p : poly) =
+ if c =/ Int 0 then poly_0 else mapf (fun x -> c */ x) p
-let poly_cmul c (p:poly) =
- if c =/ Int 0 then poly_0
- else mapf (fun x -> c */ x) p;;
-
-let poly_neg (p:poly) = (mapf minus_num p :poly);;
+let poly_neg (p : poly) = (mapf minus_num p : poly)
-let poly_add (p1:poly) (p2:poly) =
- (combine (+/) (fun x -> x =/ Int 0) p1 p2 :poly);;
+let poly_add (p1 : poly) (p2 : poly) =
+ (combine ( +/ ) (fun x -> x =/ Int 0) p1 p2 : poly)
-let poly_sub p1 p2 = poly_add p1 (poly_neg p2);;
+let poly_sub p1 p2 = poly_add p1 (poly_neg p2)
-let poly_cmmul (c,m) (p:poly) =
+let poly_cmmul (c, m) (p : poly) =
if c =/ Int 0 then poly_0
else if m = monomial_1 then mapf (fun d -> c */ d) p
- else foldl (fun a m' d -> (monomial_mul m m' |-> c */ d) a) poly_0 p;;
+ else foldl (fun a m' d -> (monomial_mul m m' |-> c */ d) a) poly_0 p
-let poly_mul (p1:poly) (p2:poly) =
- foldl (fun a m c -> poly_add (poly_cmmul (c,m) p2) a) poly_0 p1;;
+let poly_mul (p1 : poly) (p2 : poly) =
+ foldl (fun a m c -> poly_add (poly_cmmul (c, m) p2) a) poly_0 p1
-let poly_square p = poly_mul p p;;
+let poly_square p = poly_mul p p
let rec poly_pow p k =
if k = 0 then poly_const (Int 1)
else if k = 1 then p
- else let q = poly_square(poly_pow p (k / 2)) in
- if k mod 2 = 1 then poly_mul p q else q;;
+ else
+ let q = poly_square (poly_pow p (k / 2)) in
+ if k mod 2 = 1 then poly_mul p q else q
-let degree x (p:poly) = foldl (fun a m c -> max (monomial_degree x m) a) 0 p;;
+let degree x (p : poly) = foldl (fun a m c -> max (monomial_degree x m) a) 0 p
-let multidegree (p:poly) =
- foldl (fun a m c -> max (monomial_multidegree m) a) 0 p;;
+let multidegree (p : poly) =
+ foldl (fun a m c -> max (monomial_multidegree m) a) 0 p
-let poly_variables (p:poly) =
- foldr (fun m c -> union (monomial_variables m)) p [];;
+let poly_variables (p : poly) =
+ foldr (fun m c -> union (monomial_variables m)) p []
(* ------------------------------------------------------------------------- *)
(* Order monomials for human presentation. *)
(* ------------------------------------------------------------------------- *)
-let humanorder_varpow (x1,k1) (x2,k2) = x1 < x2 || x1 = x2 && k1 > k2;;
+let humanorder_varpow (x1, k1) (x2, k2) = x1 < x2 || (x1 = x2 && k1 > k2)
let humanorder_monomial =
- let rec ord l1 l2 = match (l1,l2) with
- _,[] -> true
- | [],_ -> false
- | h1::t1,h2::t2 -> humanorder_varpow h1 h2 || h1 = h2 && ord t1 t2 in
- fun m1 m2 -> m1 = m2 ||
- ord (sort humanorder_varpow (graph m1))
- (sort humanorder_varpow (graph m2));;
+ let rec ord l1 l2 =
+ match (l1, l2) with
+ | _, [] -> true
+ | [], _ -> false
+ | h1 :: t1, h2 :: t2 -> humanorder_varpow h1 h2 || (h1 = h2 && ord t1 t2)
+ in
+ fun m1 m2 ->
+ m1 = m2
+ || ord
+ (sort humanorder_varpow (graph m1))
+ (sort humanorder_varpow (graph m2))
(* ------------------------------------------------------------------------- *)
(* Conversions to strings. *)
(* ------------------------------------------------------------------------- *)
-let string_of_vname (v:vname): string = (v: string);;
+let string_of_vname (v : vname) : string = (v : string)
let string_of_varpow x k =
- if k = 1 then string_of_vname x else string_of_vname x^"^"^string_of_int k;;
+ if k = 1 then string_of_vname x else string_of_vname x ^ "^" ^ string_of_int k
let string_of_monomial m =
- if m = monomial_1 then "1" else
- let vps = List.fold_right (fun (x,k) a -> string_of_varpow x k :: a)
- (sort humanorder_varpow (graph m)) [] in
- String.concat "*" vps;;
-
-let string_of_cmonomial (c,m) =
+ if m = monomial_1 then "1"
+ else
+ let vps =
+ List.fold_right
+ (fun (x, k) a -> string_of_varpow x k :: a)
+ (sort humanorder_varpow (graph m))
+ []
+ in
+ String.concat "*" vps
+
+let string_of_cmonomial (c, m) =
if m = monomial_1 then string_of_num c
else if c =/ Int 1 then string_of_monomial m
- else string_of_num c ^ "*" ^ string_of_monomial m;;
-
-let string_of_poly (p:poly) =
- if p = poly_0 then "<<0>>" else
- let cms = sort (fun (m1,_) (m2,_) -> humanorder_monomial m1 m2) (graph p) in
- let s =
- List.fold_left (fun a (m,c) ->
- if c </ Int 0 then a ^ " - " ^ string_of_cmonomial(minus_num c,m)
- else a ^ " + " ^ string_of_cmonomial(c,m))
- "" cms in
- let s1 = String.sub s 0 3
- and s2 = String.sub s 3 (String.length s - 3) in
- "<<" ^(if s1 = " + " then s2 else "-"^s2)^">>";;
+ else string_of_num c ^ "*" ^ string_of_monomial m
+
+let string_of_poly (p : poly) =
+ if p = poly_0 then "<<0>>"
+ else
+ let cms =
+ sort (fun (m1, _) (m2, _) -> humanorder_monomial m1 m2) (graph p)
+ in
+ let s =
+ List.fold_left
+ (fun a (m, c) ->
+ if c </ Int 0 then a ^ " - " ^ string_of_cmonomial (minus_num c, m)
+ else a ^ " + " ^ string_of_cmonomial (c, m))
+ "" cms
+ in
+ let s1 = String.sub s 0 3 and s2 = String.sub s 3 (String.length s - 3) in
+ "<<" ^ (if s1 = " + " then s2 else "-" ^ s2) ^ ">>"
(* ------------------------------------------------------------------------- *)
(* Printers. *)
@@ -275,38 +271,41 @@ let print_poly m = Format.print_string(string_of_poly m);;
(* Conversion from term. *)
(* ------------------------------------------------------------------------- *)
-let rec poly_of_term t = match t with
- Zero -> poly_0
-| Const n -> poly_const n
-| Var x -> poly_var x
-| Opp t1 -> poly_neg (poly_of_term t1)
-| Add (l, r) -> poly_add (poly_of_term l) (poly_of_term r)
-| Sub (l, r) -> poly_sub (poly_of_term l) (poly_of_term r)
-| Mul (l, r) -> poly_mul (poly_of_term l) (poly_of_term r)
-| Pow (t, n) ->
- poly_pow (poly_of_term t) n;;
+let rec poly_of_term t =
+ match t with
+ | Zero -> poly_0
+ | Const n -> poly_const n
+ | Var x -> poly_var x
+ | Opp t1 -> poly_neg (poly_of_term t1)
+ | Add (l, r) -> poly_add (poly_of_term l) (poly_of_term r)
+ | Sub (l, r) -> poly_sub (poly_of_term l) (poly_of_term r)
+ | Mul (l, r) -> poly_mul (poly_of_term l) (poly_of_term r)
+ | Pow (t, n) -> poly_pow (poly_of_term t) n
(* ------------------------------------------------------------------------- *)
(* String of vector (just a list of space-separated numbers). *)
(* ------------------------------------------------------------------------- *)
-let sdpa_of_vector (v:vector) =
+let sdpa_of_vector (v : vector) =
let n = dim v in
- let strs = List.map (o (decimalize 20) (element v)) (1--n) in
- String.concat " " strs ^ "\n";;
+ let strs = List.map (o (decimalize 20) (element v)) (1 -- n) in
+ String.concat " " strs ^ "\n"
(* ------------------------------------------------------------------------- *)
(* String for a matrix numbered k, in SDPA sparse format. *)
(* ------------------------------------------------------------------------- *)
-let sdpa_of_matrix k (m:matrix) =
+let sdpa_of_matrix k (m : matrix) =
let pfx = string_of_int k ^ " 1 " in
- let ms = foldr (fun (i,j) c a -> if i > j then a else ((i,j),c)::a)
- (snd m) [] in
+ let ms =
+ foldr (fun (i, j) c a -> if i > j then a else ((i, j), c) :: a) (snd m) []
+ in
let mss = sort (increasing fst) ms in
- List.fold_right (fun ((i,j),c) a ->
- pfx ^ string_of_int i ^ " " ^ string_of_int j ^
- " " ^ decimalize 20 c ^ "\n" ^ a) mss "";;
+ List.fold_right
+ (fun ((i, j), c) a ->
+ pfx ^ string_of_int i ^ " " ^ string_of_int j ^ " " ^ decimalize 20 c
+ ^ "\n" ^ a)
+ mss ""
(* ------------------------------------------------------------------------- *)
(* String in SDPA sparse format for standard SDP problem: *)
@@ -316,85 +315,88 @@ let sdpa_of_matrix k (m:matrix) =
(* ------------------------------------------------------------------------- *)
let sdpa_of_problem comment obj mats =
- let m = List.length mats - 1
- and n,_ = dimensions (List.hd mats) in
- "\"" ^ comment ^ "\"\n" ^
- string_of_int m ^ "\n" ^
- "1\n" ^
- string_of_int n ^ "\n" ^
- sdpa_of_vector obj ^
- List.fold_right2 (fun k m a -> sdpa_of_matrix (k - 1) m ^ a)
- (1--List.length mats) mats "";;
+ let m = List.length mats - 1 and n, _ = dimensions (List.hd mats) in
+ "\"" ^ comment ^ "\"\n" ^ string_of_int m ^ "\n" ^ "1\n" ^ string_of_int n
+ ^ "\n" ^ sdpa_of_vector obj
+ ^ List.fold_right2
+ (fun k m a -> sdpa_of_matrix (k - 1) m ^ a)
+ (1 -- List.length mats)
+ mats ""
(* ------------------------------------------------------------------------- *)
(* More parser basics. *)
(* ------------------------------------------------------------------------- *)
let word s =
- end_itlist (fun p1 p2 -> (p1 ++ p2) >> (fun (s,t) -> s^t))
- (List.map a (explode s));;
+ end_itlist
+ (fun p1 p2 -> p1 ++ p2 >> fun (s, t) -> s ^ t)
+ (List.map a (explode s))
+
let token s =
- many (some isspace) ++ word s ++ many (some isspace)
- >> (fun ((_,t),_) -> t);;
+ many (some isspace) ++ word s ++ many (some isspace) >> fun ((_, t), _) -> t
let decimal =
- let (||) = parser_or in
+ let ( || ) = parser_or in
let numeral = some isnum in
- let decimalint = atleast 1 numeral >> ((o) Num.num_of_string implode) in
- let decimalfrac = atleast 1 numeral
- >> (fun s -> Num.num_of_string(implode s) // pow10 (List.length s)) in
+ let decimalint = atleast 1 numeral >> o Num.num_of_string implode in
+ let decimalfrac =
+ atleast 1 numeral
+ >> fun s -> Num.num_of_string (implode s) // pow10 (List.length s)
+ in
let decimalsig =
decimalint ++ possibly (a "." ++ decimalfrac >> snd)
- >> (function (h,[x]) -> h +/ x | (h,_) -> h) in
+ >> function h, [x] -> h +/ x | h, _ -> h
+ in
let signed prs =
- a "-" ++ prs >> ((o) minus_num snd)
- || a "+" ++ prs >> snd
- || prs in
+ a "-" ++ prs >> o minus_num snd || a "+" ++ prs >> snd || prs
+ in
let exponent = (a "e" || a "E") ++ signed decimalint >> snd in
- signed decimalsig ++ possibly exponent
- >> (function (h,[x]) -> h */ power_num (Int 10) x | (h,_) -> h);;
+ signed decimalsig ++ possibly exponent
+ >> function h, [x] -> h */ power_num (Int 10) x | h, _ -> h
let mkparser p s =
- let x,rst = p(explode s) in
- if rst = [] then x else failwith "mkparser: unparsed input";;
+ let x, rst = p (explode s) in
+ if rst = [] then x else failwith "mkparser: unparsed input"
(* ------------------------------------------------------------------------- *)
(* Parse back a vector. *)
(* ------------------------------------------------------------------------- *)
let _parse_sdpaoutput, parse_csdpoutput =
- let (||) = parser_or in
+ let ( || ) = parser_or in
let vector =
token "{" ++ listof decimal (token ",") "decimal" ++ token "}"
- >> (fun ((_,v),_) -> vector_of_list v) in
+ >> fun ((_, v), _) -> vector_of_list v
+ in
let rec skipupto dscr prs inp =
- (dscr ++ prs >> snd
- || some (fun c -> true) ++ skipupto dscr prs >> snd) inp in
- let ignore inp = (),[] in
+ (dscr ++ prs >> snd || some (fun c -> true) ++ skipupto dscr prs >> snd) inp
+ in
+ let ignore inp = ((), []) in
let sdpaoutput =
- skipupto (word "xVec" ++ token "=")
- (vector ++ ignore >> fst) in
+ skipupto (word "xVec" ++ token "=") (vector ++ ignore >> fst)
+ in
let csdpoutput =
- (decimal ++ many(a " " ++ decimal >> snd) >> (fun (h,t) -> h::t)) ++
- (a " " ++ a "\n" ++ ignore) >> ((o) vector_of_list fst) in
- mkparser sdpaoutput,mkparser csdpoutput;;
+ (decimal ++ many (a " " ++ decimal >> snd) >> fun (h, t) -> h :: t)
+ ++ (a " " ++ a "\n" ++ ignore)
+ >> o vector_of_list fst
+ in
+ (mkparser sdpaoutput, mkparser csdpoutput)
(* ------------------------------------------------------------------------- *)
(* The default parameters. Unfortunately this goes to a fixed file. *)
(* ------------------------------------------------------------------------- *)
let _sdpa_default_parameters =
-"100 unsigned int maxIteration;\
-\n1.0E-7 double 0.0 < epsilonStar;\
-\n1.0E2 double 0.0 < lambdaStar;\
-\n2.0 double 1.0 < omegaStar;\
-\n-1.0E5 double lowerBound;\
-\n1.0E5 double upperBound;\
-\n0.1 double 0.0 <= betaStar < 1.0;\
-\n0.2 double 0.0 <= betaBar < 1.0, betaStar <= betaBar;\
-\n0.9 double 0.0 < gammaStar < 1.0;\
-\n1.0E-7 double 0.0 < epsilonDash;\
-\n";;
+ "100 unsigned int maxIteration;\n\
+ 1.0E-7 double 0.0 < epsilonStar;\n\
+ 1.0E2 double 0.0 < lambdaStar;\n\
+ 2.0 double 1.0 < omegaStar;\n\
+ -1.0E5 double lowerBound;\n\
+ 1.0E5 double upperBound;\n\
+ 0.1 double 0.0 <= betaStar < 1.0;\n\
+ 0.2 double 0.0 <= betaBar < 1.0, betaStar <= betaBar;\n\
+ 0.9 double 0.0 < gammaStar < 1.0;\n\
+ 1.0E-7 double 0.0 < epsilonDash;\n"
(* ------------------------------------------------------------------------- *)
(* These were suggested by Makoto Yamashita for problems where we are *)
@@ -402,42 +404,40 @@ let _sdpa_default_parameters =
(* ------------------------------------------------------------------------- *)
let sdpa_alt_parameters =
-"1000 unsigned int maxIteration;\
-\n1.0E-7 double 0.0 < epsilonStar;\
-\n1.0E4 double 0.0 < lambdaStar;\
-\n2.0 double 1.0 < omegaStar;\
-\n-1.0E5 double lowerBound;\
-\n1.0E5 double upperBound;\
-\n0.1 double 0.0 <= betaStar < 1.0;\
-\n0.2 double 0.0 <= betaBar < 1.0, betaStar <= betaBar;\
-\n0.9 double 0.0 < gammaStar < 1.0;\
-\n1.0E-7 double 0.0 < epsilonDash;\
-\n";;
+ "1000 unsigned int maxIteration;\n\
+ 1.0E-7 double 0.0 < epsilonStar;\n\
+ 1.0E4 double 0.0 < lambdaStar;\n\
+ 2.0 double 1.0 < omegaStar;\n\
+ -1.0E5 double lowerBound;\n\
+ 1.0E5 double upperBound;\n\
+ 0.1 double 0.0 <= betaStar < 1.0;\n\
+ 0.2 double 0.0 <= betaBar < 1.0, betaStar <= betaBar;\n\
+ 0.9 double 0.0 < gammaStar < 1.0;\n\
+ 1.0E-7 double 0.0 < epsilonDash;\n"
-let _sdpa_params = sdpa_alt_parameters;;
+let _sdpa_params = sdpa_alt_parameters
(* ------------------------------------------------------------------------- *)
(* CSDP parameters; so far I'm sticking with the defaults. *)
(* ------------------------------------------------------------------------- *)
let csdp_default_parameters =
-"axtol=1.0e-8\
-\natytol=1.0e-8\
-\nobjtol=1.0e-8\
-\npinftol=1.0e8\
-\ndinftol=1.0e8\
-\nmaxiter=100\
-\nminstepfrac=0.9\
-\nmaxstepfrac=0.97\
-\nminstepp=1.0e-8\
-\nminstepd=1.0e-8\
-\nusexzgap=1\
-\ntweakgap=0\
-\naffine=0\
-\nprintlevel=1\
-\n";;
-
-let csdp_params = csdp_default_parameters;;
+ "axtol=1.0e-8\n\
+ atytol=1.0e-8\n\
+ objtol=1.0e-8\n\
+ pinftol=1.0e8\n\
+ dinftol=1.0e8\n\
+ maxiter=100\n\
+ minstepfrac=0.9\n\
+ maxstepfrac=0.97\n\
+ minstepp=1.0e-8\n\
+ minstepd=1.0e-8\n\
+ usexzgap=1\n\
+ tweakgap=0\n\
+ affine=0\n\
+ printlevel=1\n"
+
+let csdp_params = csdp_default_parameters
(* ------------------------------------------------------------------------- *)
(* Now call CSDP on a problem and parse back the output. *)
@@ -450,14 +450,15 @@ let run_csdp dbg obj mats =
and params_file = Filename.concat temp_path "param.csdp" in
file_of_string input_file (sdpa_of_problem "" obj mats);
file_of_string params_file csdp_params;
- let rv = Sys.command("cd "^temp_path^"; csdp "^input_file ^
- " " ^ output_file ^
- (if dbg then "" else "> /dev/null")) in
+ let rv =
+ Sys.command
+ ( "cd " ^ temp_path ^ "; csdp " ^ input_file ^ " " ^ output_file
+ ^ if dbg then "" else "> /dev/null" )
+ in
let op = string_of_file output_file in
let res = parse_csdpoutput op in
- ((if dbg then ()
- else (Sys.remove input_file; Sys.remove output_file));
- rv,res);;
+ if dbg then () else (Sys.remove input_file; Sys.remove output_file);
+ (rv, res)
(* ------------------------------------------------------------------------- *)
(* Try some apparently sensible scaling first. Note that this is purely to *)
@@ -470,27 +471,27 @@ let scale_then =
let common_denominator amat acc =
foldl (fun a m c -> lcm_num (denominator c) a) acc amat
and maximal_element amat acc =
- foldl (fun maxa m c -> max_num maxa (abs_num c)) acc amat in
+ foldl (fun maxa m c -> max_num maxa (abs_num c)) acc amat
+ in
fun solver obj mats ->
let cd1 = List.fold_right common_denominator mats (Int 1)
- and cd2 = common_denominator (snd obj) (Int 1) in
+ and cd2 = common_denominator (snd obj) (Int 1) in
let mats' = List.map (mapf (fun x -> cd1 */ x)) mats
and obj' = vector_cmul cd2 obj in
let max1 = List.fold_right maximal_element mats' (Int 0)
and max2 = maximal_element (snd obj') (Int 0) in
- let scal1 = pow2 (20-int_of_float(log(float_of_num max1) /. log 2.0))
- and scal2 = pow2 (20-int_of_float(log(float_of_num max2) /. log 2.0)) in
+ let scal1 = pow2 (20 - int_of_float (log (float_of_num max1) /. log 2.0))
+ and scal2 = pow2 (20 - int_of_float (log (float_of_num max2) /. log 2.0)) in
let mats'' = List.map (mapf (fun x -> x */ scal1)) mats'
and obj'' = vector_cmul scal2 obj' in
- solver obj'' mats'';;
+ solver obj'' mats''
(* ------------------------------------------------------------------------- *)
(* Round a vector to "nice" rationals. *)
(* ------------------------------------------------------------------------- *)
-let nice_rational n x = round_num (n */ x) // n;;
-
-let nice_vector n = mapa (nice_rational n);;
+let nice_rational n x = round_num (n */ x) // n
+let nice_vector n = mapa (nice_rational n)
(* ------------------------------------------------------------------------- *)
(* Reduce linear program to SDP (diagonal matrices) and test with CSDP. This *)
@@ -498,13 +499,13 @@ let nice_vector n = mapa (nice_rational n);;
(* ------------------------------------------------------------------------- *)
let linear_program_basic a =
- let m,n = dimensions a in
- let mats = List.map (fun j -> diagonal (column j a)) (1--n)
+ let m, n = dimensions a in
+ let mats = List.map (fun j -> diagonal (column j a)) (1 -- n)
and obj = vector_const (Int 1) m in
- let rv,res = run_csdp false obj mats in
+ let rv, res = run_csdp false obj mats in
if rv = 1 || rv = 2 then false
else if rv = 0 then true
- else failwith "linear_program: An error occurred in the SDP solver";;
+ else failwith "linear_program: An error occurred in the SDP solver"
(* ------------------------------------------------------------------------- *)
(* Test whether a point is in the convex hull of others. Rather than use *)
@@ -513,16 +514,17 @@ let linear_program_basic a =
(* ------------------------------------------------------------------------- *)
let in_convex_hull pts pt =
- let pts1 = (1::pt) :: List.map (fun x -> 1::x) pts in
+ let pts1 = (1 :: pt) :: List.map (fun x -> 1 :: x) pts in
let pts2 = List.map (fun p -> List.map (fun x -> -x) p @ p) pts1 in
- let n = List.length pts + 1
- and v = 2 * (List.length pt + 1) in
+ let n = List.length pts + 1 and v = 2 * (List.length pt + 1) in
let m = v + n - 1 in
let mat =
- (m,n),
- itern 1 pts2 (fun pts j -> itern 1 pts (fun x i -> (i,j) |-> Int x))
- (iter (1,n) (fun i -> (v + i,i+1) |-> Int 1) undefined) in
- linear_program_basic mat;;
+ ( (m, n)
+ , itern 1 pts2
+ (fun pts j -> itern 1 pts (fun x i -> (i, j) |-> Int x))
+ (iter (1, n) (fun i -> (v + i, i + 1) |-> Int 1) undefined) )
+ in
+ linear_program_basic mat
(* ------------------------------------------------------------------------- *)
(* Filter down a set of points to a minimal set with the same convex hull. *)
@@ -531,24 +533,23 @@ let in_convex_hull pts pt =
let minimal_convex_hull =
let augment1 = function
| [] -> assert false
- | (m::ms) -> if in_convex_hull ms m then ms else ms@[m] in
- let augment m ms = funpow 3 augment1 (m::ms) in
+ | m :: ms -> if in_convex_hull ms m then ms else ms @ [m]
+ in
+ let augment m ms = funpow 3 augment1 (m :: ms) in
fun mons ->
let mons' = List.fold_right augment (List.tl mons) [List.hd mons] in
- funpow (List.length mons') augment1 mons';;
+ funpow (List.length mons') augment1 mons'
(* ------------------------------------------------------------------------- *)
(* Stuff for "equations" (generic A->num functions). *)
(* ------------------------------------------------------------------------- *)
-let equation_cmul c eq =
- if c =/ Int 0 then Empty else mapf (fun d -> c */ d) eq;;
-
-let equation_add eq1 eq2 = combine (+/) (fun x -> x =/ Int 0) eq1 eq2;;
+let equation_cmul c eq = if c =/ Int 0 then Empty else mapf (fun d -> c */ d) eq
+let equation_add eq1 eq2 = combine ( +/ ) (fun x -> x =/ Int 0) eq1 eq2
let equation_eval assig eq =
let value v = apply assig v in
- foldl (fun a v c -> a +/ value(v) */ c) (Int 0) eq;;
+ foldl (fun a v c -> a +/ (value v */ c)) (Int 0) eq
(* ------------------------------------------------------------------------- *)
(* Eliminate all variables, in an essentially arbitrary order. *)
@@ -556,29 +557,35 @@ let equation_eval assig eq =
let eliminate_all_equations one =
let choose_variable eq =
- let (v,_) = choose eq in
+ let v, _ = choose eq in
if v = one then
let eq' = undefine v eq in
- if is_undefined eq' then failwith "choose_variable" else
- let (w,_) = choose eq' in w
- else v in
+ if is_undefined eq' then failwith "choose_variable"
+ else
+ let w, _ = choose eq' in
+ w
+ else v
+ in
let rec eliminate dun eqs =
match eqs with
- [] -> dun
- | eq::oeqs ->
- if is_undefined eq then eliminate dun oeqs else
+ | [] -> dun
+ | eq :: oeqs ->
+ if is_undefined eq then eliminate dun oeqs
+ else
let v = choose_variable eq in
let a = apply eq v in
- let eq' = equation_cmul (Int(-1) // a) (undefine v eq) in
+ let eq' = equation_cmul (Int (-1) // a) (undefine v eq) in
let elim e =
let b = tryapplyd e v (Int 0) in
- if b =/ Int 0 then e else
- equation_add e (equation_cmul (minus_num b // a) eq) in
- eliminate ((v |-> eq') (mapf elim dun)) (List.map elim oeqs) in
+ if b =/ Int 0 then e
+ else equation_add e (equation_cmul (minus_num b // a) eq)
+ in
+ eliminate ((v |-> eq') (mapf elim dun)) (List.map elim oeqs)
+ in
fun eqs ->
let assig = eliminate undefined eqs in
let vs = foldl (fun a x f -> subtract (dom f) [one] @ a) [] assig in
- setify vs,assig;;
+ (setify vs, assig)
(* ------------------------------------------------------------------------- *)
(* Hence produce the "relevant" monomials: those whose squares lie in the *)
@@ -593,14 +600,23 @@ let eliminate_all_equations one =
let newton_polytope pol =
let vars = poly_variables pol in
- let mons = List.map (fun m -> List.map (fun x -> monomial_degree x m) vars) (dom pol)
+ let mons =
+ List.map (fun m -> List.map (fun x -> monomial_degree x m) vars) (dom pol)
and ds = List.map (fun x -> (degree x pol + 1) / 2) vars in
- let all = List.fold_right (fun n -> allpairs (fun h t -> h::t) (0--n)) ds [[]]
+ let all =
+ List.fold_right (fun n -> allpairs (fun h t -> h :: t) (0 -- n)) ds [[]]
and mons' = minimal_convex_hull mons in
let all' =
- List.filter (fun m -> in_convex_hull mons' (List.map (fun x -> 2 * x) m)) all in
- List.map (fun m -> List.fold_right2 (fun v i a -> if i = 0 then a else (v |-> i) a)
- vars m monomial_1) (List.rev all');;
+ List.filter
+ (fun m -> in_convex_hull mons' (List.map (fun x -> 2 * x) m))
+ all
+ in
+ List.map
+ (fun m ->
+ List.fold_right2
+ (fun v i a -> if i = 0 then a else (v |-> i) a)
+ vars m monomial_1)
+ (List.rev all')
(* ------------------------------------------------------------------------- *)
(* Diagonalize (Cholesky/LDU) the matrix corresponding to a quadratic form. *)
@@ -609,40 +625,55 @@ let newton_polytope pol =
let diag m =
let nn = dimensions m in
let n = fst nn in
- if snd nn <> n then failwith "diagonalize: non-square matrix" else
- let rec diagonalize i m =
- if is_zero m then [] else
- let a11 = element m (i,i) in
- if a11 </ Int 0 then failwith "diagonalize: not PSD"
- else if a11 =/ Int 0 then
- if is_zero(row i m) then diagonalize (i + 1) m
- else failwith "diagonalize: not PSD"
- else
- let v = row i m in
- let v' = mapa (fun a1k -> a1k // a11) v in
- let m' =
- (n,n),
- iter (i+1,n) (fun j ->
- iter (i+1,n) (fun k ->
- ((j,k) |--> (element m (j,k) -/ element v j */ element v' k))))
- undefined in
- (a11,v')::diagonalize (i + 1) m' in
- diagonalize 1 m;;
+ if snd nn <> n then failwith "diagonalize: non-square matrix"
+ else
+ let rec diagonalize i m =
+ if is_zero m then []
+ else
+ let a11 = element m (i, i) in
+ if a11 </ Int 0 then failwith "diagonalize: not PSD"
+ else if a11 =/ Int 0 then
+ if is_zero (row i m) then diagonalize (i + 1) m
+ else failwith "diagonalize: not PSD"
+ else
+ let v = row i m in
+ let v' = mapa (fun a1k -> a1k // a11) v in
+ let m' =
+ ( (n, n)
+ , iter
+ (i + 1, n)
+ (fun j ->
+ iter
+ (i + 1, n)
+ (fun k ->
+ (j, k)
+ |--> element m (j, k) -/ (element v j */ element v' k)))
+ undefined )
+ in
+ (a11, v') :: diagonalize (i + 1) m'
+ in
+ diagonalize 1 m
(* ------------------------------------------------------------------------- *)
(* Adjust a diagonalization to collect rationals at the start. *)
(* ------------------------------------------------------------------------- *)
let deration d =
- if d = [] then Int 0,d else
- let adj(c,l) =
- let a = foldl (fun a i c -> lcm_num a (denominator c)) (Int 1) (snd l) //
- foldl (fun a i c -> gcd_num a (numerator c)) (Int 0) (snd l) in
- (c // (a */ a)),mapa (fun x -> a */ x) l in
- let d' = List.map adj d in
- let a = List.fold_right ((o) lcm_num ( (o) denominator fst)) d' (Int 1) //
- List.fold_right ((o) gcd_num ( (o) numerator fst)) d' (Int 0) in
- (Int 1 // a),List.map (fun (c,l) -> (a */ c,l)) d';;
+ if d = [] then (Int 0, d)
+ else
+ let adj (c, l) =
+ let a =
+ foldl (fun a i c -> lcm_num a (denominator c)) (Int 1) (snd l)
+ // foldl (fun a i c -> gcd_num a (numerator c)) (Int 0) (snd l)
+ in
+ (c // (a */ a), mapa (fun x -> a */ x) l)
+ in
+ let d' = List.map adj d in
+ let a =
+ List.fold_right (o lcm_num (o denominator fst)) d' (Int 1)
+ // List.fold_right (o gcd_num (o numerator fst)) d' (Int 0)
+ in
+ (Int 1 // a, List.map (fun (c, l) -> (a */ c, l)) d')
(* ------------------------------------------------------------------------- *)
(* Enumeration of monomials with given multidegree bound. *)
@@ -651,12 +682,18 @@ let deration d =
let rec enumerate_monomials d vars =
if d < 0 then []
else if d = 0 then [undefined]
- else if vars = [] then [monomial_1] else
- let alts =
- List.map (fun k -> let oths = enumerate_monomials (d - k) (List.tl vars) in
- List.map (fun ks -> if k = 0 then ks else (List.hd vars |-> k) ks) oths)
- (0--d) in
- end_itlist (@) alts;;
+ else if vars = [] then [monomial_1]
+ else
+ let alts =
+ List.map
+ (fun k ->
+ let oths = enumerate_monomials (d - k) (List.tl vars) in
+ List.map
+ (fun ks -> if k = 0 then ks else (List.hd vars |-> k) ks)
+ oths)
+ (0 -- d)
+ in
+ end_itlist ( @ ) alts
(* ------------------------------------------------------------------------- *)
(* Enumerate products of distinct input polys with degree <= d. *)
@@ -665,46 +702,57 @@ let rec enumerate_monomials d vars =
(* ------------------------------------------------------------------------- *)
let rec enumerate_products d pols =
- if d = 0 then [poly_const num_1,Rational_lt num_1] else if d < 0 then [] else
- match pols with
- [] -> [poly_const num_1,Rational_lt num_1]
- | (p,b)::ps -> let e = multidegree p in
- if e = 0 then enumerate_products d ps else
- enumerate_products d ps @
- List.map (fun (q,c) -> poly_mul p q,Product(b,c))
- (enumerate_products (d - e) ps);;
+ if d = 0 then [(poly_const num_1, Rational_lt num_1)]
+ else if d < 0 then []
+ else
+ match pols with
+ | [] -> [(poly_const num_1, Rational_lt num_1)]
+ | (p, b) :: ps ->
+ let e = multidegree p in
+ if e = 0 then enumerate_products d ps
+ else
+ enumerate_products d ps
+ @ List.map
+ (fun (q, c) -> (poly_mul p q, Product (b, c)))
+ (enumerate_products (d - e) ps)
(* ------------------------------------------------------------------------- *)
(* Multiply equation-parametrized poly by regular poly and add accumulator. *)
(* ------------------------------------------------------------------------- *)
let epoly_pmul p q acc =
- foldl (fun a m1 c ->
- foldl (fun b m2 e ->
- let m = monomial_mul m1 m2 in
- let es = tryapplyd b m undefined in
- (m |-> equation_add (equation_cmul c e) es) b)
- a q) acc p;;
+ foldl
+ (fun a m1 c ->
+ foldl
+ (fun b m2 e ->
+ let m = monomial_mul m1 m2 in
+ let es = tryapplyd b m undefined in
+ (m |-> equation_add (equation_cmul c e) es) b)
+ a q)
+ acc p
(* ------------------------------------------------------------------------- *)
(* Convert regular polynomial. Note that we treat (0,0,0) as -1. *)
(* ------------------------------------------------------------------------- *)
let epoly_of_poly p =
- foldl (fun a m c -> (m |-> ((0,0,0) |=> minus_num c)) a) undefined p;;
+ foldl (fun a m c -> (m |-> ((0, 0, 0) |=> minus_num c)) a) undefined p
(* ------------------------------------------------------------------------- *)
(* String for block diagonal matrix numbered k. *)
(* ------------------------------------------------------------------------- *)
let sdpa_of_blockdiagonal k m =
- let pfx = string_of_int k ^" " in
+ let pfx = string_of_int k ^ " " in
let ents =
- foldl (fun a (b,i,j) c -> if i > j then a else ((b,i,j),c)::a) [] m in
+ foldl (fun a (b, i, j) c -> if i > j then a else ((b, i, j), c) :: a) [] m
+ in
let entss = sort (increasing fst) ents in
- List.fold_right (fun ((b,i,j),c) a ->
- pfx ^ string_of_int b ^ " " ^ string_of_int i ^ " " ^ string_of_int j ^
- " " ^ decimalize 20 c ^ "\n" ^ a) entss "";;
+ List.fold_right
+ (fun ((b, i, j), c) a ->
+ pfx ^ string_of_int b ^ " " ^ string_of_int i ^ " " ^ string_of_int j
+ ^ " " ^ decimalize 20 c ^ "\n" ^ a)
+ entss ""
(* ------------------------------------------------------------------------- *)
(* SDPA for problem using block diagonal (i.e. multiple SDPs) *)
@@ -712,14 +760,14 @@ let sdpa_of_blockdiagonal k m =
let sdpa_of_blockproblem comment nblocks blocksizes obj mats =
let m = List.length mats - 1 in
- "\"" ^ comment ^ "\"\n" ^
- string_of_int m ^ "\n" ^
- string_of_int nblocks ^ "\n" ^
- (String.concat " " (List.map string_of_int blocksizes)) ^
- "\n" ^
- sdpa_of_vector obj ^
- List.fold_right2 (fun k m a -> sdpa_of_blockdiagonal (k - 1) m ^ a)
- (1--List.length mats) mats "";;
+ "\"" ^ comment ^ "\"\n" ^ string_of_int m ^ "\n" ^ string_of_int nblocks
+ ^ "\n"
+ ^ String.concat " " (List.map string_of_int blocksizes)
+ ^ "\n" ^ sdpa_of_vector obj
+ ^ List.fold_right2
+ (fun k m a -> sdpa_of_blockdiagonal (k - 1) m ^ a)
+ (1 -- List.length mats)
+ mats ""
(* ------------------------------------------------------------------------- *)
(* Hence run CSDP on a problem in block diagonal form. *)
@@ -731,254 +779,319 @@ let run_csdp dbg nblocks blocksizes obj mats =
String.sub input_file 0 (String.length input_file - 6) ^ ".out"
and params_file = Filename.concat temp_path "param.csdp" in
file_of_string input_file
- (sdpa_of_blockproblem "" nblocks blocksizes obj mats);
+ (sdpa_of_blockproblem "" nblocks blocksizes obj mats);
file_of_string params_file csdp_params;
- let rv = Sys.command("cd "^temp_path^"; csdp "^input_file ^
- " " ^ output_file ^
- (if dbg then "" else "> /dev/null")) in
+ let rv =
+ Sys.command
+ ( "cd " ^ temp_path ^ "; csdp " ^ input_file ^ " " ^ output_file
+ ^ if dbg then "" else "> /dev/null" )
+ in
let op = string_of_file output_file in
let res = parse_csdpoutput op in
- ((if dbg then ()
- else (Sys.remove input_file; Sys.remove output_file));
- rv,res);;
+ if dbg then () else (Sys.remove input_file; Sys.remove output_file);
+ (rv, res)
let csdp nblocks blocksizes obj mats =
- let rv,res = run_csdp (!debugging) nblocks blocksizes obj mats in
- (if rv = 1 || rv = 2 then failwith "csdp: Problem is infeasible"
- else if rv = 3 then ()
+ let rv, res = run_csdp !debugging nblocks blocksizes obj mats in
+ if rv = 1 || rv = 2 then failwith "csdp: Problem is infeasible"
+ else if rv = 3 then ()
(*Format.print_string "csdp warning: Reduced accuracy";
Format.print_newline() *)
- else if rv <> 0 then failwith("csdp: error "^string_of_int rv)
- else ());
- res;;
+ else if rv <> 0 then failwith ("csdp: error " ^ string_of_int rv)
+ else ();
+ res
(* ------------------------------------------------------------------------- *)
(* 3D versions of matrix operations to consider blocks separately. *)
(* ------------------------------------------------------------------------- *)
-let bmatrix_add = combine (+/) (fun x -> x =/ Int 0);;
+let bmatrix_add = combine ( +/ ) (fun x -> x =/ Int 0)
let bmatrix_cmul c bm =
- if c =/ Int 0 then undefined
- else mapf (fun x -> c */ x) bm;;
+ if c =/ Int 0 then undefined else mapf (fun x -> c */ x) bm
-let bmatrix_neg = bmatrix_cmul (Int(-1));;
+let bmatrix_neg = bmatrix_cmul (Int (-1))
(* ------------------------------------------------------------------------- *)
(* Smash a block matrix into components. *)
(* ------------------------------------------------------------------------- *)
let blocks blocksizes bm =
- List.map (fun (bs,b0) ->
- let m = foldl
- (fun a (b,i,j) c -> if b = b0 then ((i,j) |-> c) a else a)
- undefined bm in
- (((bs,bs),m):matrix))
- (List.combine blocksizes (1--List.length blocksizes));;
+ List.map
+ (fun (bs, b0) ->
+ let m =
+ foldl
+ (fun a (b, i, j) c -> if b = b0 then ((i, j) |-> c) a else a)
+ undefined bm
+ in
+ (((bs, bs), m) : matrix))
+ (List.combine blocksizes (1 -- List.length blocksizes))
(* ------------------------------------------------------------------------- *)
(* Positiv- and Nullstellensatz. Flag "linf" forces a linear representation. *)
(* ------------------------------------------------------------------------- *)
let real_positivnullstellensatz_general linf d eqs leqs pol =
- let vars = List.fold_right ((o) union poly_variables) (pol::eqs @ List.map fst leqs) [] in
+ let vars =
+ List.fold_right (o union poly_variables)
+ ((pol :: eqs) @ List.map fst leqs)
+ []
+ in
let monoid =
if linf then
- (poly_const num_1,Rational_lt num_1)::
- (List.filter (fun (p,c) -> multidegree p <= d) leqs)
- else enumerate_products d leqs in
+ (poly_const num_1, Rational_lt num_1)
+ :: List.filter (fun (p, c) -> multidegree p <= d) leqs
+ else enumerate_products d leqs
+ in
let nblocks = List.length monoid in
let mk_idmultiplier k p =
let e = d - multidegree p in
let mons = enumerate_monomials e vars in
- let nons = List.combine mons (1--List.length mons) in
- mons,
- List.fold_right (fun (m,n) -> (m |-> ((-k,-n,n) |=> Int 1))) nons undefined in
- let mk_sqmultiplier k (p,c) =
+ let nons = List.combine mons (1 -- List.length mons) in
+ ( mons
+ , List.fold_right
+ (fun (m, n) -> m |-> ((-k, -n, n) |=> Int 1))
+ nons undefined )
+ in
+ let mk_sqmultiplier k (p, c) =
let e = (d - multidegree p) / 2 in
let mons = enumerate_monomials e vars in
- let nons = List.combine mons (1--List.length mons) in
- mons,
- List.fold_right (fun (m1,n1) ->
- List.fold_right (fun (m2,n2) a ->
- let m = monomial_mul m1 m2 in
- if n1 > n2 then a else
- let c = if n1 = n2 then Int 1 else Int 2 in
- let e = tryapplyd a m undefined in
- (m |-> equation_add ((k,n1,n2) |=> c) e) a)
- nons)
- nons undefined in
- let sqmonlist,sqs = List.split(List.map2 mk_sqmultiplier (1--List.length monoid) monoid)
- and idmonlist,ids = List.split(List.map2 mk_idmultiplier (1--List.length eqs) eqs) in
+ let nons = List.combine mons (1 -- List.length mons) in
+ ( mons
+ , List.fold_right
+ (fun (m1, n1) ->
+ List.fold_right
+ (fun (m2, n2) a ->
+ let m = monomial_mul m1 m2 in
+ if n1 > n2 then a
+ else
+ let c = if n1 = n2 then Int 1 else Int 2 in
+ let e = tryapplyd a m undefined in
+ (m |-> equation_add ((k, n1, n2) |=> c) e) a)
+ nons)
+ nons undefined )
+ in
+ let sqmonlist, sqs =
+ List.split (List.map2 mk_sqmultiplier (1 -- List.length monoid) monoid)
+ and idmonlist, ids =
+ List.split (List.map2 mk_idmultiplier (1 -- List.length eqs) eqs)
+ in
let blocksizes = List.map List.length sqmonlist in
let bigsum =
- List.fold_right2 (fun p q a -> epoly_pmul p q a) eqs ids
- (List.fold_right2 (fun (p,c) s a -> epoly_pmul p s a) monoid sqs
- (epoly_of_poly(poly_neg pol))) in
- let eqns = foldl (fun a m e -> e::a) [] bigsum in
- let pvs,assig = eliminate_all_equations (0,0,0) eqns in
- let qvars = (0,0,0)::pvs in
- let allassig = List.fold_right (fun v -> (v |-> (v |=> Int 1))) pvs assig in
+ List.fold_right2
+ (fun p q a -> epoly_pmul p q a)
+ eqs ids
+ (List.fold_right2
+ (fun (p, c) s a -> epoly_pmul p s a)
+ monoid sqs
+ (epoly_of_poly (poly_neg pol)))
+ in
+ let eqns = foldl (fun a m e -> e :: a) [] bigsum in
+ let pvs, assig = eliminate_all_equations (0, 0, 0) eqns in
+ let qvars = (0, 0, 0) :: pvs in
+ let allassig = List.fold_right (fun v -> v |-> (v |=> Int 1)) pvs assig in
let mk_matrix v =
- foldl (fun m (b,i,j) ass -> if b < 0 then m else
- let c = tryapplyd ass v (Int 0) in
- if c =/ Int 0 then m else
- ((b,j,i) |-> c) (((b,i,j) |-> c) m))
- undefined allassig in
- let diagents = foldl
- (fun a (b,i,j) e -> if b > 0 && i = j then equation_add e a else a)
- undefined allassig in
+ foldl
+ (fun m (b, i, j) ass ->
+ if b < 0 then m
+ else
+ let c = tryapplyd ass v (Int 0) in
+ if c =/ Int 0 then m else ((b, j, i) |-> c) (((b, i, j) |-> c) m))
+ undefined allassig
+ in
+ let diagents =
+ foldl
+ (fun a (b, i, j) e -> if b > 0 && i = j then equation_add e a else a)
+ undefined allassig
+ in
let mats = List.map mk_matrix qvars
- and obj = List.length pvs,
- itern 1 pvs (fun v i -> (i |--> tryapplyd diagents v (Int 0)))
- undefined in
- let raw_vec = if pvs = [] then vector_0 0
- else scale_then (csdp nblocks blocksizes) obj mats in
+ and obj =
+ ( List.length pvs
+ , itern 1 pvs (fun v i -> i |--> tryapplyd diagents v (Int 0)) undefined )
+ in
+ let raw_vec =
+ if pvs = [] then vector_0 0
+ else scale_then (csdp nblocks blocksizes) obj mats
+ in
let find_rounding d =
- (if !debugging then
- (Format.print_string("Trying rounding with limit "^string_of_num d);
- Format.print_newline())
- else ());
+ if !debugging then (
+ Format.print_string ("Trying rounding with limit " ^ string_of_num d);
+ Format.print_newline () )
+ else ();
let vec = nice_vector d raw_vec in
- let blockmat = iter (1,dim vec)
- (fun i a -> bmatrix_add (bmatrix_cmul (element vec i) (List.nth mats i)) a)
- (bmatrix_neg (List.nth mats 0)) in
+ let blockmat =
+ iter
+ (1, dim vec)
+ (fun i a ->
+ bmatrix_add (bmatrix_cmul (element vec i) (List.nth mats i)) a)
+ (bmatrix_neg (List.nth mats 0))
+ in
let allmats = blocks blocksizes blockmat in
- vec,List.map diag allmats in
- let vec,ratdias =
+ (vec, List.map diag allmats)
+ in
+ let vec, ratdias =
if pvs = [] then find_rounding num_1
- else tryfind find_rounding (List.map Num.num_of_int (1--31) @
- List.map pow2 (5--66)) in
+ else
+ tryfind find_rounding
+ (List.map Num.num_of_int (1 -- 31) @ List.map pow2 (5 -- 66))
+ in
let newassigs =
- List.fold_right (fun k -> List.nth pvs (k - 1) |-> element vec k)
- (1--dim vec) ((0,0,0) |=> Int(-1)) in
+ List.fold_right
+ (fun k -> List.nth pvs (k - 1) |-> element vec k)
+ (1 -- dim vec)
+ ((0, 0, 0) |=> Int (-1))
+ in
let finalassigs =
- foldl (fun a v e -> (v |-> equation_eval newassigs e) a) newassigs
- allassig in
+ foldl (fun a v e -> (v |-> equation_eval newassigs e) a) newassigs allassig
+ in
let poly_of_epoly p =
- foldl (fun a v e -> (v |--> equation_eval finalassigs e) a)
- undefined p in
+ foldl (fun a v e -> (v |--> equation_eval finalassigs e) a) undefined p
+ in
let mk_sos mons =
- let mk_sq (c,m) =
- c,List.fold_right (fun k a -> (List.nth mons (k - 1) |--> element m k) a)
- (1--List.length mons) undefined in
- List.map mk_sq in
+ let mk_sq (c, m) =
+ ( c
+ , List.fold_right
+ (fun k a -> (List.nth mons (k - 1) |--> element m k) a)
+ (1 -- List.length mons)
+ undefined )
+ in
+ List.map mk_sq
+ in
let sqs = List.map2 mk_sos sqmonlist ratdias
and cfs = List.map poly_of_epoly ids in
- let msq = List.filter (fun (a,b) -> b <> []) (List.map2 (fun a b -> a,b) monoid sqs) in
- let eval_sq sqs = List.fold_right
- (fun (c,q) -> poly_add (poly_cmul c (poly_mul q q))) sqs poly_0 in
+ let msq =
+ List.filter
+ (fun (a, b) -> b <> [])
+ (List.map2 (fun a b -> (a, b)) monoid sqs)
+ in
+ let eval_sq sqs =
+ List.fold_right
+ (fun (c, q) -> poly_add (poly_cmul c (poly_mul q q)))
+ sqs poly_0
+ in
let sanity =
- List.fold_right (fun ((p,c),s) -> poly_add (poly_mul p (eval_sq s))) msq
- (List.fold_right2 (fun p q -> poly_add (poly_mul p q)) cfs eqs
- (poly_neg pol)) in
- if not(is_undefined sanity) then raise Sanity else
- cfs,List.map (fun (a,b) -> snd a,b) msq;;
+ List.fold_right
+ (fun ((p, c), s) -> poly_add (poly_mul p (eval_sq s)))
+ msq
+ (List.fold_right2
+ (fun p q -> poly_add (poly_mul p q))
+ cfs eqs (poly_neg pol))
+ in
+ if not (is_undefined sanity) then raise Sanity
+ else (cfs, List.map (fun (a, b) -> (snd a, b)) msq)
(* ------------------------------------------------------------------------- *)
(* The ordering so we can create canonical HOL polynomials. *)
(* ------------------------------------------------------------------------- *)
-let dest_monomial mon = sort (increasing fst) (graph mon);;
+let dest_monomial mon = sort (increasing fst) (graph mon)
let monomial_order =
let rec lexorder l1 l2 =
- match (l1,l2) with
- [],[] -> true
- | vps,[] -> false
- | [],vps -> true
- | ((x1,n1)::vs1),((x2,n2)::vs2) ->
- if x1 < x2 then true
- else if x2 < x1 then false
- else if n1 < n2 then false
- else if n2 < n1 then true
- else lexorder vs1 vs2 in
+ match (l1, l2) with
+ | [], [] -> true
+ | vps, [] -> false
+ | [], vps -> true
+ | (x1, n1) :: vs1, (x2, n2) :: vs2 ->
+ if x1 < x2 then true
+ else if x2 < x1 then false
+ else if n1 < n2 then false
+ else if n2 < n1 then true
+ else lexorder vs1 vs2
+ in
fun m1 m2 ->
- if m2 = monomial_1 then true else if m1 = monomial_1 then false else
- let mon1 = dest_monomial m1 and mon2 = dest_monomial m2 in
- let deg1 = List.fold_right ((o) (+) snd) mon1 0
- and deg2 = List.fold_right ((o) (+) snd) mon2 0 in
- if deg1 < deg2 then false else if deg1 > deg2 then true
- else lexorder mon1 mon2;;
+ if m2 = monomial_1 then true
+ else if m1 = monomial_1 then false
+ else
+ let mon1 = dest_monomial m1 and mon2 = dest_monomial m2 in
+ let deg1 = List.fold_right (o ( + ) snd) mon1 0
+ and deg2 = List.fold_right (o ( + ) snd) mon2 0 in
+ if deg1 < deg2 then false
+ else if deg1 > deg2 then true
+ else lexorder mon1 mon2
(* ------------------------------------------------------------------------- *)
(* Map back polynomials and their composites to HOL. *)
(* ------------------------------------------------------------------------- *)
-let term_of_varpow =
- fun x k ->
- if k = 1 then Var x else Pow (Var x, k);;
+let term_of_varpow x k = if k = 1 then Var x else Pow (Var x, k)
-let term_of_monomial =
- fun m -> if m = monomial_1 then Const num_1 else
- let m' = dest_monomial m in
- let vps = List.fold_right (fun (x,k) a -> term_of_varpow x k :: a) m' [] in
- end_itlist (fun s t -> Mul (s,t)) vps;;
+let term_of_monomial m =
+ if m = monomial_1 then Const num_1
+ else
+ let m' = dest_monomial m in
+ let vps = List.fold_right (fun (x, k) a -> term_of_varpow x k :: a) m' [] in
+ end_itlist (fun s t -> Mul (s, t)) vps
-let term_of_cmonomial =
- fun (m,c) ->
- if m = monomial_1 then Const c
- else if c =/ num_1 then term_of_monomial m
- else Mul (Const c,term_of_monomial m);;
+let term_of_cmonomial (m, c) =
+ if m = monomial_1 then Const c
+ else if c =/ num_1 then term_of_monomial m
+ else Mul (Const c, term_of_monomial m)
-let term_of_poly =
- fun p ->
- if p = poly_0 then Zero else
- let cms = List.map term_of_cmonomial
- (sort (fun (m1,_) (m2,_) -> monomial_order m1 m2) (graph p)) in
- end_itlist (fun t1 t2 -> Add (t1,t2)) cms;;
+let term_of_poly p =
+ if p = poly_0 then Zero
+ else
+ let cms =
+ List.map term_of_cmonomial
+ (sort (fun (m1, _) (m2, _) -> monomial_order m1 m2) (graph p))
+ in
+ end_itlist (fun t1 t2 -> Add (t1, t2)) cms
-let term_of_sqterm (c,p) =
- Product(Rational_lt c,Square(term_of_poly p));;
+let term_of_sqterm (c, p) = Product (Rational_lt c, Square (term_of_poly p))
-let term_of_sos (pr,sqs) =
+let term_of_sos (pr, sqs) =
if sqs = [] then pr
- else Product(pr,end_itlist (fun a b -> Sum(a,b)) (List.map term_of_sqterm sqs));;
+ else
+ Product
+ (pr, end_itlist (fun a b -> Sum (a, b)) (List.map term_of_sqterm sqs))
(* ------------------------------------------------------------------------- *)
(* Some combinatorial helper functions. *)
(* ------------------------------------------------------------------------- *)
let rec allpermutations l =
- if l = [] then [[]] else
- List.fold_right (fun h acc -> List.map (fun t -> h::t)
- (allpermutations (subtract l [h])) @ acc) l [];;
+ if l = [] then [[]]
+ else
+ List.fold_right
+ (fun h acc ->
+ List.map (fun t -> h :: t) (allpermutations (subtract l [h])) @ acc)
+ l []
-let changevariables_monomial zoln (m:monomial) =
- foldl (fun a x k -> (List.assoc x zoln |-> k) a) monomial_1 m;;
+let changevariables_monomial zoln (m : monomial) =
+ foldl (fun a x k -> (List.assoc x zoln |-> k) a) monomial_1 m
let changevariables zoln pol =
- foldl (fun a m c -> (changevariables_monomial zoln m |-> c) a)
- poly_0 pol;;
+ foldl (fun a m c -> (changevariables_monomial zoln m |-> c) a) poly_0 pol
(* ------------------------------------------------------------------------- *)
(* Return to original non-block matrices. *)
(* ------------------------------------------------------------------------- *)
-let sdpa_of_vector (v:vector) =
+let sdpa_of_vector (v : vector) =
let n = dim v in
- let strs = List.map (o (decimalize 20) (element v)) (1--n) in
- String.concat " " strs ^ "\n";;
+ let strs = List.map (o (decimalize 20) (element v)) (1 -- n) in
+ String.concat " " strs ^ "\n"
-let sdpa_of_matrix k (m:matrix) =
+let sdpa_of_matrix k (m : matrix) =
let pfx = string_of_int k ^ " 1 " in
- let ms = foldr (fun (i,j) c a -> if i > j then a else ((i,j),c)::a)
- (snd m) [] in
+ let ms =
+ foldr (fun (i, j) c a -> if i > j then a else ((i, j), c) :: a) (snd m) []
+ in
let mss = sort (increasing fst) ms in
- List.fold_right (fun ((i,j),c) a ->
- pfx ^ string_of_int i ^ " " ^ string_of_int j ^
- " " ^ decimalize 20 c ^ "\n" ^ a) mss "";;
+ List.fold_right
+ (fun ((i, j), c) a ->
+ pfx ^ string_of_int i ^ " " ^ string_of_int j ^ " " ^ decimalize 20 c
+ ^ "\n" ^ a)
+ mss ""
let sdpa_of_problem comment obj mats =
- let m = List.length mats - 1
- and n,_ = dimensions (List.hd mats) in
- "\"" ^ comment ^ "\"\n" ^
- string_of_int m ^ "\n" ^
- "1\n" ^
- string_of_int n ^ "\n" ^
- sdpa_of_vector obj ^
- List.fold_right2 (fun k m a -> sdpa_of_matrix (k - 1) m ^ a)
- (1--List.length mats) mats "";;
+ let m = List.length mats - 1 and n, _ = dimensions (List.hd mats) in
+ "\"" ^ comment ^ "\"\n" ^ string_of_int m ^ "\n" ^ "1\n" ^ string_of_int n
+ ^ "\n" ^ sdpa_of_vector obj
+ ^ List.fold_right2
+ (fun k m a -> sdpa_of_matrix (k - 1) m ^ a)
+ (1 -- List.length mats)
+ mats ""
let run_csdp dbg obj mats =
let input_file = Filename.temp_file "sos" ".dat-s" in
@@ -987,109 +1100,139 @@ let run_csdp dbg obj mats =
and params_file = Filename.concat temp_path "param.csdp" in
file_of_string input_file (sdpa_of_problem "" obj mats);
file_of_string params_file csdp_params;
- let rv = Sys.command("cd "^temp_path^"; csdp "^input_file ^
- " " ^ output_file ^
- (if dbg then "" else "> /dev/null")) in
+ let rv =
+ Sys.command
+ ( "cd " ^ temp_path ^ "; csdp " ^ input_file ^ " " ^ output_file
+ ^ if dbg then "" else "> /dev/null" )
+ in
let op = string_of_file output_file in
let res = parse_csdpoutput op in
- ((if dbg then ()
- else (Sys.remove input_file; Sys.remove output_file));
- rv,res);;
+ if dbg then () else (Sys.remove input_file; Sys.remove output_file);
+ (rv, res)
let csdp obj mats =
- let rv,res = run_csdp (!debugging) obj mats in
- (if rv = 1 || rv = 2 then failwith "csdp: Problem is infeasible"
- else if rv = 3 then ()
-(* (Format.print_string "csdp warning: Reduced accuracy";
+ let rv, res = run_csdp !debugging obj mats in
+ if rv = 1 || rv = 2 then failwith "csdp: Problem is infeasible"
+ else if rv = 3 then ()
+ (* (Format.print_string "csdp warning: Reduced accuracy";
Format.print_newline()) *)
- else if rv <> 0 then failwith("csdp: error "^string_of_int rv)
- else ());
- res;;
+ else if rv <> 0 then failwith ("csdp: error " ^ string_of_int rv)
+ else ();
+ res
(* ------------------------------------------------------------------------- *)
(* Sum-of-squares function with some lowbrow symmetry reductions. *)
(* ------------------------------------------------------------------------- *)
let sumofsquares_general_symmetry tool pol =
- let vars = poly_variables pol
- and lpps = newton_polytope pol in
+ let vars = poly_variables pol and lpps = newton_polytope pol in
let n = List.length lpps in
let sym_eqs =
- let invariants = List.filter
- (fun vars' ->
- is_undefined(poly_sub pol (changevariables (List.combine vars vars') pol)))
- (allpermutations vars) in
- let lpns = List.combine lpps (1--List.length lpps) in
+ let invariants =
+ List.filter
+ (fun vars' ->
+ is_undefined
+ (poly_sub pol (changevariables (List.combine vars vars') pol)))
+ (allpermutations vars)
+ in
+ let lpns = List.combine lpps (1 -- List.length lpps) in
let lppcs =
- List.filter (fun (m,(n1,n2)) -> n1 <= n2)
- (allpairs
- (fun (m1,n1) (m2,n2) -> (m1,m2),(n1,n2)) lpns lpns) in
- let clppcs = end_itlist (@)
- (List.map (fun ((m1,m2),(n1,n2)) ->
- List.map (fun vars' ->
- (changevariables_monomial (List.combine vars vars') m1,
- changevariables_monomial (List.combine vars vars') m2),(n1,n2))
- invariants)
- lppcs) in
- let clppcs_dom = setify(List.map fst clppcs) in
- let clppcs_cls = List.map (fun d -> List.filter (fun (e,_) -> e = d) clppcs)
- clppcs_dom in
+ List.filter
+ (fun (m, (n1, n2)) -> n1 <= n2)
+ (allpairs (fun (m1, n1) (m2, n2) -> ((m1, m2), (n1, n2))) lpns lpns)
+ in
+ let clppcs =
+ end_itlist ( @ )
+ (List.map
+ (fun ((m1, m2), (n1, n2)) ->
+ List.map
+ (fun vars' ->
+ ( ( changevariables_monomial (List.combine vars vars') m1
+ , changevariables_monomial (List.combine vars vars') m2 )
+ , (n1, n2) ))
+ invariants)
+ lppcs)
+ in
+ let clppcs_dom = setify (List.map fst clppcs) in
+ let clppcs_cls =
+ List.map (fun d -> List.filter (fun (e, _) -> e = d) clppcs) clppcs_dom
+ in
let eqvcls = List.map (o setify (List.map snd)) clppcs_cls in
let mk_eq cls acc =
match cls with
- [] -> raise Sanity
+ | [] -> raise Sanity
| [h] -> acc
- | h::t -> List.map (fun k -> (k |-> Int(-1)) (h |=> Int 1)) t @ acc in
- List.fold_right mk_eq eqvcls [] in
- let eqs = foldl (fun a x y -> y::a) []
- (itern 1 lpps (fun m1 n1 ->
- itern 1 lpps (fun m2 n2 f ->
- let m = monomial_mul m1 m2 in
- if n1 > n2 then f else
- let c = if n1 = n2 then Int 1 else Int 2 in
- (m |-> ((n1,n2) |-> c) (tryapplyd f m undefined)) f))
- (foldl (fun a m c -> (m |-> ((0,0)|=>c)) a)
- undefined pol)) @
- sym_eqs in
- let pvs,assig = eliminate_all_equations (0,0) eqs in
- let allassig = List.fold_right (fun v -> (v |-> (v |=> Int 1))) pvs assig in
- let qvars = (0,0)::pvs in
+ | h :: t -> List.map (fun k -> (k |-> Int (-1)) (h |=> Int 1)) t @ acc
+ in
+ List.fold_right mk_eq eqvcls []
+ in
+ let eqs =
+ foldl
+ (fun a x y -> y :: a)
+ []
+ (itern 1 lpps
+ (fun m1 n1 ->
+ itern 1 lpps (fun m2 n2 f ->
+ let m = monomial_mul m1 m2 in
+ if n1 > n2 then f
+ else
+ let c = if n1 = n2 then Int 1 else Int 2 in
+ (m |-> ((n1, n2) |-> c) (tryapplyd f m undefined)) f))
+ (foldl (fun a m c -> (m |-> ((0, 0) |=> c)) a) undefined pol))
+ @ sym_eqs
+ in
+ let pvs, assig = eliminate_all_equations (0, 0) eqs in
+ let allassig = List.fold_right (fun v -> v |-> (v |=> Int 1)) pvs assig in
+ let qvars = (0, 0) :: pvs in
let diagents =
- end_itlist equation_add (List.map (fun i -> apply allassig (i,i)) (1--n)) in
+ end_itlist equation_add (List.map (fun i -> apply allassig (i, i)) (1 -- n))
+ in
let mk_matrix v =
- ((n,n),
- foldl (fun m (i,j) ass -> let c = tryapplyd ass v (Int 0) in
- if c =/ Int 0 then m else
- ((j,i) |-> c) (((i,j) |-> c) m))
- undefined allassig :matrix) in
+ ( ( (n, n)
+ , foldl
+ (fun m (i, j) ass ->
+ let c = tryapplyd ass v (Int 0) in
+ if c =/ Int 0 then m else ((j, i) |-> c) (((i, j) |-> c) m))
+ undefined allassig )
+ : matrix )
+ in
let mats = List.map mk_matrix qvars
- and obj = List.length pvs,
- itern 1 pvs (fun v i -> (i |--> tryapplyd diagents v (Int 0)))
- undefined in
+ and obj =
+ ( List.length pvs
+ , itern 1 pvs (fun v i -> i |--> tryapplyd diagents v (Int 0)) undefined )
+ in
let raw_vec = if pvs = [] then vector_0 0 else tool obj mats in
let find_rounding d =
- (if !debugging then
- (Format.print_string("Trying rounding with limit "^string_of_num d);
- Format.print_newline())
- else ());
+ if !debugging then (
+ Format.print_string ("Trying rounding with limit " ^ string_of_num d);
+ Format.print_newline () )
+ else ();
let vec = nice_vector d raw_vec in
- let mat = iter (1,dim vec)
- (fun i a -> matrix_add (matrix_cmul (element vec i) (List.nth mats i)) a)
- (matrix_neg (List.nth mats 0)) in
- deration(diag mat) in
- let rat,dia =
+ let mat =
+ iter
+ (1, dim vec)
+ (fun i a ->
+ matrix_add (matrix_cmul (element vec i) (List.nth mats i)) a)
+ (matrix_neg (List.nth mats 0))
+ in
+ deration (diag mat)
+ in
+ let rat, dia =
if pvs = [] then
- let mat = matrix_neg (List.nth mats 0) in
- deration(diag mat)
+ let mat = matrix_neg (List.nth mats 0) in
+ deration (diag mat)
else
- tryfind find_rounding (List.map Num.num_of_int (1--31) @
- List.map pow2 (5--66)) in
- let poly_of_lin(d,v) =
- d,foldl(fun a i c -> (List.nth lpps (i - 1) |-> c) a) undefined (snd v) in
+ tryfind find_rounding
+ (List.map Num.num_of_int (1 -- 31) @ List.map pow2 (5 -- 66))
+ in
+ let poly_of_lin (d, v) =
+ (d, foldl (fun a i c -> (List.nth lpps (i - 1) |-> c) a) undefined (snd v))
+ in
let lins = List.map poly_of_lin dia in
- let sqs = List.map (fun (d,l) -> poly_mul (poly_const d) (poly_pow l 2)) lins in
+ let sqs =
+ List.map (fun (d, l) -> poly_mul (poly_const d) (poly_pow l 2)) lins
+ in
let sos = poly_cmul rat (end_itlist poly_add sqs) in
- if is_undefined(poly_sub sos pol) then rat,lins else raise Sanity;;
-
-let sumofsquares = sumofsquares_general_symmetry csdp;;
+ if is_undefined (poly_sub sos pol) then (rat, lins) else raise Sanity
+let sumofsquares = sumofsquares_general_symmetry csdp
diff --git a/plugins/micromega/sos.mli b/plugins/micromega/sos.mli
index c9181953c8..ac75bd37f0 100644
--- a/plugins/micromega/sos.mli
+++ b/plugins/micromega/sos.mli
@@ -13,26 +13,24 @@ open Sos_types
type poly
val poly_isconst : poly -> bool
-
val poly_neg : poly -> poly
-
val poly_mul : poly -> poly -> poly
-
val poly_pow : poly -> int -> poly
-
val poly_const : Num.num -> poly
-
val poly_of_term : term -> poly
-
val term_of_poly : poly -> term
-val term_of_sos : positivstellensatz * (Num.num * poly) list ->
- positivstellensatz
+val term_of_sos :
+ positivstellensatz * (Num.num * poly) list -> positivstellensatz
val string_of_poly : poly -> string
-val real_positivnullstellensatz_general : bool -> int -> poly list ->
- (poly * positivstellensatz) list ->
- poly -> poly list * (positivstellensatz * (Num.num * poly) list) list
+val real_positivnullstellensatz_general :
+ bool
+ -> int
+ -> poly list
+ -> (poly * positivstellensatz) list
+ -> poly
+ -> poly list * (positivstellensatz * (Num.num * poly) list) list
-val sumofsquares : poly -> Num.num * ( Num.num * poly) list
+val sumofsquares : poly -> Num.num * (Num.num * poly) list
diff --git a/plugins/micromega/sos_lib.ml b/plugins/micromega/sos_lib.ml
index 0a0ffc7947..51221aa6b9 100644
--- a/plugins/micromega/sos_lib.ml
+++ b/plugins/micromega/sos_lib.ml
@@ -13,47 +13,45 @@ open Num
(* Comparisons that are reflexive on NaN and also short-circuiting. *)
(* ------------------------------------------------------------------------- *)
-let cmp = compare (** FIXME *)
+(** FIXME *)
+let cmp = compare
-let (=?) = fun x y -> cmp x y = 0;;
-let (<?) = fun x y -> cmp x y < 0;;
-let (<=?) = fun x y -> cmp x y <= 0;;
-let (>?) = fun x y -> cmp x y > 0;;
+let ( =? ) x y = cmp x y = 0
+let ( <? ) x y = cmp x y < 0
+let ( <=? ) x y = cmp x y <= 0
+let ( >? ) x y = cmp x y > 0
(* ------------------------------------------------------------------------- *)
(* Combinators. *)
(* ------------------------------------------------------------------------- *)
-let (o) = fun f g x -> f(g x);;
+let o f g x = f (g x)
(* ------------------------------------------------------------------------- *)
(* Some useful functions on "num" type. *)
(* ------------------------------------------------------------------------- *)
-
let num_0 = Int 0
and num_1 = Int 1
and num_2 = Int 2
-and num_10 = Int 10;;
+and num_10 = Int 10
-let pow2 n = power_num num_2 (Int n);;
-let pow10 n = power_num num_10 (Int n);;
+let pow2 n = power_num num_2 (Int n)
+let pow10 n = power_num num_10 (Int n)
let numdom r =
let r' = Ratio.normalize_ratio (ratio_of_num r) in
- num_of_big_int(Ratio.numerator_ratio r'),
- num_of_big_int(Ratio.denominator_ratio r');;
+ ( num_of_big_int (Ratio.numerator_ratio r')
+ , num_of_big_int (Ratio.denominator_ratio r') )
-let numerator = (o) fst numdom
-and denominator = (o) snd numdom;;
+let numerator = o fst numdom
+and denominator = o snd numdom
let gcd_num n1 n2 =
- num_of_big_int(Big_int.gcd_big_int (big_int_of_num n1) (big_int_of_num n2));;
+ num_of_big_int (Big_int.gcd_big_int (big_int_of_num n1) (big_int_of_num n2))
let lcm_num x y =
- if x =/ num_0 && y =/ num_0 then num_0
- else abs_num((x */ y) // gcd_num x y);;
-
+ if x =/ num_0 && y =/ num_0 then num_0 else abs_num (x */ y // gcd_num x y)
(* ------------------------------------------------------------------------- *)
(* Various versions of list iteration. *)
@@ -61,9 +59,9 @@ let lcm_num x y =
let rec end_itlist f l =
match l with
- [] -> failwith "end_itlist"
- | [x] -> x
- | (h::t) -> f h (end_itlist f t);;
+ | [] -> failwith "end_itlist"
+ | [x] -> x
+ | h :: t -> f h (end_itlist f t)
(* ------------------------------------------------------------------------- *)
(* All pairs arising from applying a function over two lists. *)
@@ -71,36 +69,32 @@ let rec end_itlist f l =
let rec allpairs f l1 l2 =
match l1 with
- h1::t1 -> List.fold_right (fun x a -> f h1 x :: a) l2 (allpairs f t1 l2)
- | [] -> [];;
+ | h1 :: t1 -> List.fold_right (fun x a -> f h1 x :: a) l2 (allpairs f t1 l2)
+ | [] -> []
(* ------------------------------------------------------------------------- *)
(* String operations (surely there is a better way...) *)
(* ------------------------------------------------------------------------- *)
-let implode l = List.fold_right (^) l "";;
+let implode l = List.fold_right ( ^ ) l ""
let explode s =
let rec exap n l =
- if n < 0 then l else
- exap (n - 1) ((String.sub s n 1)::l) in
- exap (String.length s - 1) [];;
-
+ if n < 0 then l else exap (n - 1) (String.sub s n 1 :: l)
+ in
+ exap (String.length s - 1) []
(* ------------------------------------------------------------------------- *)
(* Repetition of a function. *)
(* ------------------------------------------------------------------------- *)
-let rec funpow n f x =
- if n < 1 then x else funpow (n-1) f (f x);;
-
-
+let rec funpow n f x = if n < 1 then x else funpow (n - 1) f (f x)
(* ------------------------------------------------------------------------- *)
(* Sequences. *)
(* ------------------------------------------------------------------------- *)
-let rec (--) = fun m n -> if m > n then [] else m::((m + 1) -- n);;
+let rec ( -- ) m n = if m > n then [] else m :: (m + 1 -- n)
(* ------------------------------------------------------------------------- *)
(* Various useful list operations. *)
@@ -108,39 +102,29 @@ let rec (--) = fun m n -> if m > n then [] else m::((m + 1) -- n);;
let rec tryfind f l =
match l with
- [] -> failwith "tryfind"
- | (h::t) -> try f h with Failure _ -> tryfind f t;;
+ | [] -> failwith "tryfind"
+ | h :: t -> ( try f h with Failure _ -> tryfind f t )
(* ------------------------------------------------------------------------- *)
(* "Set" operations on lists. *)
(* ------------------------------------------------------------------------- *)
-let rec mem x lis =
- match lis with
- [] -> false
- | (h::t) -> x =? h || mem x t;;
-
-let insert x l =
- if mem x l then l else x::l;;
-
-let union l1 l2 = List.fold_right insert l1 l2;;
-
-let subtract l1 l2 = List.filter (fun x -> not (mem x l2)) l1;;
+let rec mem x lis = match lis with [] -> false | h :: t -> x =? h || mem x t
+let insert x l = if mem x l then l else x :: l
+let union l1 l2 = List.fold_right insert l1 l2
+let subtract l1 l2 = List.filter (fun x -> not (mem x l2)) l1
(* ------------------------------------------------------------------------- *)
(* Common measure predicates to use with "sort". *)
(* ------------------------------------------------------------------------- *)
-let increasing f x y = f x <? f y;;
+let increasing f x y = f x <? f y
(* ------------------------------------------------------------------------- *)
(* Iterating functions over lists. *)
(* ------------------------------------------------------------------------- *)
-let rec do_list f l =
- match l with
- [] -> ()
- | (h::t) -> (f h; do_list f t);;
+let rec do_list f l = match l with [] -> () | h :: t -> f h; do_list f t
(* ------------------------------------------------------------------------- *)
(* Sorting. *)
@@ -148,10 +132,10 @@ let rec do_list f l =
let rec sort cmp lis =
match lis with
- [] -> []
- | piv::rest ->
- let r,l = List.partition (cmp piv) rest in
- (sort cmp l) @ (piv::(sort cmp r));;
+ | [] -> []
+ | piv :: rest ->
+ let r, l = List.partition (cmp piv) rest in
+ sort cmp l @ (piv :: sort cmp r)
(* ------------------------------------------------------------------------- *)
(* Removing adjacent (NB!) equal elements from list. *)
@@ -159,16 +143,16 @@ let rec sort cmp lis =
let rec uniq l =
match l with
- x::(y::_ as t) -> let t' = uniq t in
- if x =? y then t' else
- if t'==t then l else x::t'
- | _ -> l;;
+ | x :: (y :: _ as t) ->
+ let t' = uniq t in
+ if x =? y then t' else if t' == t then l else x :: t'
+ | _ -> l
(* ------------------------------------------------------------------------- *)
(* Convert list into set by eliminating duplicates. *)
(* ------------------------------------------------------------------------- *)
-let setify s = uniq (sort (<=?) s);;
+let setify s = uniq (sort ( <=? ) s)
(* ------------------------------------------------------------------------- *)
(* Polymorphic finite partial functions via Patricia trees. *)
@@ -179,25 +163,22 @@ let setify s = uniq (sort (<=?) s);;
(* Idea due to Diego Olivier Fernandez Pons (OCaml list, 2003/11/10). *)
(* ------------------------------------------------------------------------- *)
-type ('a,'b)func =
- Empty
- | Leaf of int * ('a*'b)list
- | Branch of int * int * ('a,'b)func * ('a,'b)func;;
+type ('a, 'b) func =
+ | Empty
+ | Leaf of int * ('a * 'b) list
+ | Branch of int * int * ('a, 'b) func * ('a, 'b) func
(* ------------------------------------------------------------------------- *)
(* Undefined function. *)
(* ------------------------------------------------------------------------- *)
-let undefined = Empty;;
+let undefined = Empty
(* ------------------------------------------------------------------------- *)
(* In case of equality comparison worries, better use this. *)
(* ------------------------------------------------------------------------- *)
-let is_undefined f =
- match f with
- Empty -> true
- | _ -> false;;
+let is_undefined f = match f with Empty -> true | _ -> false
(* ------------------------------------------------------------------------- *)
(* Operation analogous to "map" for lists. *)
@@ -205,15 +186,15 @@ let is_undefined f =
let mapf =
let rec map_list f l =
- match l with
- [] -> []
- | (x,y)::t -> (x,f(y))::(map_list f t) in
+ match l with [] -> [] | (x, y) :: t -> (x, f y) :: map_list f t
+ in
let rec mapf f t =
match t with
- Empty -> Empty
- | Leaf(h,l) -> Leaf(h,map_list f l)
- | Branch(p,b,l,r) -> Branch(p,b,mapf f l,mapf f r) in
- mapf;;
+ | Empty -> Empty
+ | Leaf (h, l) -> Leaf (h, map_list f l)
+ | Branch (p, b, l, r) -> Branch (p, b, mapf f l, mapf f r)
+ in
+ mapf
(* ------------------------------------------------------------------------- *)
(* Operations analogous to "fold" for lists. *)
@@ -221,119 +202,125 @@ let mapf =
let foldl =
let rec foldl_list f a l =
- match l with
- [] -> a
- | (x,y)::t -> foldl_list f (f a x y) t in
+ match l with [] -> a | (x, y) :: t -> foldl_list f (f a x y) t
+ in
let rec foldl f a t =
match t with
- Empty -> a
- | Leaf(h,l) -> foldl_list f a l
- | Branch(p,b,l,r) -> foldl f (foldl f a l) r in
- foldl;;
+ | Empty -> a
+ | Leaf (h, l) -> foldl_list f a l
+ | Branch (p, b, l, r) -> foldl f (foldl f a l) r
+ in
+ foldl
let foldr =
let rec foldr_list f l a =
- match l with
- [] -> a
- | (x,y)::t -> f x y (foldr_list f t a) in
+ match l with [] -> a | (x, y) :: t -> f x y (foldr_list f t a)
+ in
let rec foldr f t a =
match t with
- Empty -> a
- | Leaf(h,l) -> foldr_list f l a
- | Branch(p,b,l,r) -> foldr f l (foldr f r a) in
- foldr;;
+ | Empty -> a
+ | Leaf (h, l) -> foldr_list f l a
+ | Branch (p, b, l, r) -> foldr f l (foldr f r a)
+ in
+ foldr
(* ------------------------------------------------------------------------- *)
(* Redefinition and combination. *)
(* ------------------------------------------------------------------------- *)
-let (|->),combine =
- let ldb x y = let z = x lxor y in z land (-z) in
+let ( |-> ), combine =
+ let ldb x y =
+ let z = x lxor y in
+ z land -z
+ in
let newbranch p1 t1 p2 t2 =
let b = ldb p1 p2 in
let p = p1 land (b - 1) in
- if p1 land b = 0 then Branch(p,b,t1,t2)
- else Branch(p,b,t2,t1) in
- let rec define_list (x,y as xy) l =
+ if p1 land b = 0 then Branch (p, b, t1, t2) else Branch (p, b, t2, t1)
+ in
+ let rec define_list ((x, y) as xy) l =
match l with
- (a,b as ab)::t ->
- if x =? a then xy::t
- else if x <? a then xy::l
- else ab::(define_list xy t)
+ | ((a, b) as ab) :: t ->
+ if x =? a then xy :: t
+ else if x <? a then xy :: l
+ else ab :: define_list xy t
| [] -> [xy]
and combine_list op z l1 l2 =
- match (l1,l2) with
- [],_ -> l2
- | _,[] -> l1
- | ((x1,y1 as xy1)::t1,(x2,y2 as xy2)::t2) ->
- if x1 <? x2 then xy1::(combine_list op z t1 l2)
- else if x2 <? x1 then xy2::(combine_list op z l1 t2) else
- let y = op y1 y2 and l = combine_list op z t1 t2 in
- if z(y) then l else (x1,y)::l in
- let (|->) x y =
+ match (l1, l2) with
+ | [], _ -> l2
+ | _, [] -> l1
+ | ((x1, y1) as xy1) :: t1, ((x2, y2) as xy2) :: t2 ->
+ if x1 <? x2 then xy1 :: combine_list op z t1 l2
+ else if x2 <? x1 then xy2 :: combine_list op z l1 t2
+ else
+ let y = op y1 y2 and l = combine_list op z t1 t2 in
+ if z y then l else (x1, y) :: l
+ in
+ let ( |-> ) x y =
let k = Hashtbl.hash x in
let rec upd t =
match t with
- Empty -> Leaf (k,[x,y])
- | Leaf(h,l) ->
- if h = k then Leaf(h,define_list (x,y) l)
- else newbranch h t k (Leaf(k,[x,y]))
- | Branch(p,b,l,r) ->
- if k land (b - 1) <> p then newbranch p t k (Leaf(k,[x,y]))
- else if k land b = 0 then Branch(p,b,upd l,r)
- else Branch(p,b,l,upd r) in
- upd in
+ | Empty -> Leaf (k, [(x, y)])
+ | Leaf (h, l) ->
+ if h = k then Leaf (h, define_list (x, y) l)
+ else newbranch h t k (Leaf (k, [(x, y)]))
+ | Branch (p, b, l, r) ->
+ if k land (b - 1) <> p then newbranch p t k (Leaf (k, [(x, y)]))
+ else if k land b = 0 then Branch (p, b, upd l, r)
+ else Branch (p, b, l, upd r)
+ in
+ upd
+ in
let rec combine op z t1 t2 =
- match (t1,t2) with
- Empty,_ -> t2
- | _,Empty -> t1
- | Leaf(h1,l1),Leaf(h2,l2) ->
- if h1 = h2 then
- let l = combine_list op z l1 l2 in
- if l = [] then Empty else Leaf(h1,l)
- else newbranch h1 t1 h2 t2
- | (Leaf(k,lis) as lf),(Branch(p,b,l,r) as br) |
- (Branch(p,b,l,r) as br),(Leaf(k,lis) as lf) ->
- if k land (b - 1) = p then
- if k land b = 0 then
- let l' = combine op z lf l in
- if is_undefined l' then r else Branch(p,b,l',r)
- else
- let r' = combine op z lf r in
- if is_undefined r' then l else Branch(p,b,l,r')
- else
- newbranch k lf p br
- | Branch(p1,b1,l1,r1),Branch(p2,b2,l2,r2) ->
- if b1 < b2 then
- if p2 land (b1 - 1) <> p1 then newbranch p1 t1 p2 t2
- else if p2 land b1 = 0 then
- let l = combine op z l1 t2 in
- if is_undefined l then r1 else Branch(p1,b1,l,r1)
- else
- let r = combine op z r1 t2 in
- if is_undefined r then l1 else Branch(p1,b1,l1,r)
- else if b2 < b1 then
- if p1 land (b2 - 1) <> p2 then newbranch p1 t1 p2 t2
- else if p1 land b2 = 0 then
- let l = combine op z t1 l2 in
- if is_undefined l then r2 else Branch(p2,b2,l,r2)
- else
- let r = combine op z t1 r2 in
- if is_undefined r then l2 else Branch(p2,b2,l2,r)
- else if p1 = p2 then
- let l = combine op z l1 l2 and r = combine op z r1 r2 in
- if is_undefined l then r
- else if is_undefined r then l else Branch(p1,b1,l,r)
- else
- newbranch p1 t1 p2 t2 in
- (|->),combine;;
+ match (t1, t2) with
+ | Empty, _ -> t2
+ | _, Empty -> t1
+ | Leaf (h1, l1), Leaf (h2, l2) ->
+ if h1 = h2 then
+ let l = combine_list op z l1 l2 in
+ if l = [] then Empty else Leaf (h1, l)
+ else newbranch h1 t1 h2 t2
+ | (Leaf (k, lis) as lf), (Branch (p, b, l, r) as br)
+ |(Branch (p, b, l, r) as br), (Leaf (k, lis) as lf) ->
+ if k land (b - 1) = p then
+ if k land b = 0 then
+ let l' = combine op z lf l in
+ if is_undefined l' then r else Branch (p, b, l', r)
+ else
+ let r' = combine op z lf r in
+ if is_undefined r' then l else Branch (p, b, l, r')
+ else newbranch k lf p br
+ | Branch (p1, b1, l1, r1), Branch (p2, b2, l2, r2) ->
+ if b1 < b2 then
+ if p2 land (b1 - 1) <> p1 then newbranch p1 t1 p2 t2
+ else if p2 land b1 = 0 then
+ let l = combine op z l1 t2 in
+ if is_undefined l then r1 else Branch (p1, b1, l, r1)
+ else
+ let r = combine op z r1 t2 in
+ if is_undefined r then l1 else Branch (p1, b1, l1, r)
+ else if b2 < b1 then
+ if p1 land (b2 - 1) <> p2 then newbranch p1 t1 p2 t2
+ else if p1 land b2 = 0 then
+ let l = combine op z t1 l2 in
+ if is_undefined l then r2 else Branch (p2, b2, l, r2)
+ else
+ let r = combine op z t1 r2 in
+ if is_undefined r then l2 else Branch (p2, b2, l2, r)
+ else if p1 = p2 then
+ let l = combine op z l1 l2 and r = combine op z r1 r2 in
+ if is_undefined l then r
+ else if is_undefined r then l
+ else Branch (p1, b1, l, r)
+ else newbranch p1 t1 p2 t2
+ in
+ (( |-> ), combine)
(* ------------------------------------------------------------------------- *)
(* Special case of point function. *)
(* ------------------------------------------------------------------------- *)
-let (|=>) = fun x y -> (x |-> y) undefined;;
-
+let ( |=> ) x y = (x |-> y) undefined
(* ------------------------------------------------------------------------- *)
(* Grab an arbitrary element. *)
@@ -341,9 +328,9 @@ let (|=>) = fun x y -> (x |-> y) undefined;;
let rec choose t =
match t with
- Empty -> failwith "choose: completely undefined function"
- | Leaf(h,l) -> List.hd l
- | Branch(b,p,t1,t2) -> choose t1;;
+ | Empty -> failwith "choose: completely undefined function"
+ | Leaf (h, l) -> List.hd l
+ | Branch (b, p, t1, t2) -> choose t1
(* ------------------------------------------------------------------------- *)
(* Application. *)
@@ -352,21 +339,22 @@ let rec choose t =
let applyd =
let rec apply_listd l d x =
match l with
- (a,b)::t -> if x =? a then b
- else if x >? a then apply_listd t d x else d x
- | [] -> d x in
+ | (a, b) :: t ->
+ if x =? a then b else if x >? a then apply_listd t d x else d x
+ | [] -> d x
+ in
fun f d x ->
let k = Hashtbl.hash x in
let rec look t =
match t with
- Leaf(h,l) when h = k -> apply_listd l d x
- | Branch(p,b,l,r) -> look (if k land b = 0 then l else r)
- | _ -> d x in
- look f;;
-
-let apply f = applyd f (fun x -> failwith "apply");;
+ | Leaf (h, l) when h = k -> apply_listd l d x
+ | Branch (p, b, l, r) -> look (if k land b = 0 then l else r)
+ | _ -> d x
+ in
+ look f
-let tryapplyd f a d = applyd f (fun x -> d) a;;
+let apply f = applyd f (fun x -> failwith "apply")
+let tryapplyd f a d = applyd f (fun x -> d) a
(* ------------------------------------------------------------------------- *)
(* Undefinition. *)
@@ -375,161 +363,166 @@ let tryapplyd f a d = applyd f (fun x -> d) a;;
let undefine =
let rec undefine_list x l =
match l with
- (a,b as ab)::t ->
- if x =? a then t
- else if x <? a then l else
- let t' = undefine_list x t in
- if t' == t then l else ab::t'
- | [] -> [] in
+ | ((a, b) as ab) :: t ->
+ if x =? a then t
+ else if x <? a then l
+ else
+ let t' = undefine_list x t in
+ if t' == t then l else ab :: t'
+ | [] -> []
+ in
fun x ->
let k = Hashtbl.hash x in
let rec und t =
match t with
- Leaf(h,l) when h = k ->
- let l' = undefine_list x l in
+ | Leaf (h, l) when h = k ->
+ let l' = undefine_list x l in
+ if l' == l then t else if l' = [] then Empty else Leaf (h, l')
+ | Branch (p, b, l, r) when k land (b - 1) = p ->
+ if k land b = 0 then
+ let l' = und l in
if l' == l then t
- else if l' = [] then Empty
- else Leaf(h,l')
- | Branch(p,b,l,r) when k land (b - 1) = p ->
- if k land b = 0 then
- let l' = und l in
- if l' == l then t
- else if is_undefined l' then r
- else Branch(p,b,l',r)
- else
- let r' = und r in
- if r' == r then t
- else if is_undefined r' then l
- else Branch(p,b,l,r')
- | _ -> t in
- und;;
-
+ else if is_undefined l' then r
+ else Branch (p, b, l', r)
+ else
+ let r' = und r in
+ if r' == r then t
+ else if is_undefined r' then l
+ else Branch (p, b, l, r')
+ | _ -> t
+ in
+ und
(* ------------------------------------------------------------------------- *)
(* Mapping to sorted-list representation of the graph, domain and range. *)
(* ------------------------------------------------------------------------- *)
-let graph f = setify (foldl (fun a x y -> (x,y)::a) [] f);;
-
-let dom f = setify(foldl (fun a x y -> x::a) [] f);;
+let graph f = setify (foldl (fun a x y -> (x, y) :: a) [] f)
+let dom f = setify (foldl (fun a x y -> x :: a) [] f)
(* ------------------------------------------------------------------------- *)
(* More parser basics. *)
(* ------------------------------------------------------------------------- *)
-exception Noparse;;
+exception Noparse
-
-let isspace,isnum =
- let charcode s = Char.code(String.get s 0) in
+let isspace, isnum =
+ let charcode s = Char.code s.[0] in
let spaces = " \t\n\r"
and separators = ",;"
and brackets = "()[]{}"
and symbs = "\\!@#$%^&*-+|\\<=>/?~.:"
and alphas = "'abcdefghijklmnopqrstuvwxyz_ABCDEFGHIJKLMNOPQRSTUVWXYZ"
and nums = "0123456789" in
- let allchars = spaces^separators^brackets^symbs^alphas^nums in
- let csetsize = List.fold_right ((o) max charcode) (explode allchars) 256 in
+ let allchars = spaces ^ separators ^ brackets ^ symbs ^ alphas ^ nums in
+ let csetsize = List.fold_right (o max charcode) (explode allchars) 256 in
let ctable = Array.make csetsize 0 in
- do_list (fun c -> Array.set ctable (charcode c) 1) (explode spaces);
- do_list (fun c -> Array.set ctable (charcode c) 2) (explode separators);
- do_list (fun c -> Array.set ctable (charcode c) 4) (explode brackets);
- do_list (fun c -> Array.set ctable (charcode c) 8) (explode symbs);
- do_list (fun c -> Array.set ctable (charcode c) 16) (explode alphas);
- do_list (fun c -> Array.set ctable (charcode c) 32) (explode nums);
- let isspace c = Array.get ctable (charcode c) = 1
- and isnum c = Array.get ctable (charcode c) = 32 in
- isspace,isnum;;
+ do_list (fun c -> ctable.(charcode c) <- 1) (explode spaces);
+ do_list (fun c -> ctable.(charcode c) <- 2) (explode separators);
+ do_list (fun c -> ctable.(charcode c) <- 4) (explode brackets);
+ do_list (fun c -> ctable.(charcode c) <- 8) (explode symbs);
+ do_list (fun c -> ctable.(charcode c) <- 16) (explode alphas);
+ do_list (fun c -> ctable.(charcode c) <- 32) (explode nums);
+ let isspace c = ctable.(charcode c) = 1
+ and isnum c = ctable.(charcode c) = 32 in
+ (isspace, isnum)
let parser_or parser1 parser2 input =
- try parser1 input
- with Noparse -> parser2 input;;
+ try parser1 input with Noparse -> parser2 input
-let (++) parser1 parser2 input =
- let result1,rest1 = parser1 input in
- let result2,rest2 = parser2 rest1 in
- (result1,result2),rest2;;
+let ( ++ ) parser1 parser2 input =
+ let result1, rest1 = parser1 input in
+ let result2, rest2 = parser2 rest1 in
+ ((result1, result2), rest2)
let rec many prs input =
- try let result,next = prs input in
- let results,rest = many prs next in
- (result::results),rest
- with Noparse -> [],input;;
+ try
+ let result, next = prs input in
+ let results, rest = many prs next in
+ (result :: results, rest)
+ with Noparse -> ([], input)
-let (>>) prs treatment input =
- let result,rest = prs input in
- treatment(result),rest;;
+let ( >> ) prs treatment input =
+ let result, rest = prs input in
+ (treatment result, rest)
let fix err prs input =
- try prs input
- with Noparse -> failwith (err ^ " expected");;
+ try prs input with Noparse -> failwith (err ^ " expected")
let listof prs sep err =
- prs ++ many (sep ++ fix err prs >> snd) >> (fun (h,t) -> h::t);;
+ prs ++ many (sep ++ fix err prs >> snd) >> fun (h, t) -> h :: t
let possibly prs input =
- try let x,rest = prs input in [x],rest
- with Noparse -> [],input;;
+ try
+ let x, rest = prs input in
+ ([x], rest)
+ with Noparse -> ([], input)
-let some p =
- function
- [] -> raise Noparse
- | (h::t) -> if p h then (h,t) else raise Noparse;;
+let some p = function
+ | [] -> raise Noparse
+ | h :: t -> if p h then (h, t) else raise Noparse
-let a tok = some (fun item -> item = tok);;
+let a tok = some (fun item -> item = tok)
let rec atleast n prs i =
- (if n <= 0 then many prs
- else prs ++ atleast (n - 1) prs >> (fun (h,t) -> h::t)) i;;
+ ( if n <= 0 then many prs
+ else prs ++ atleast (n - 1) prs >> fun (h, t) -> h :: t )
+ i
(* ------------------------------------------------------------------------- *)
-let temp_path = Filename.get_temp_dir_name ();;
+let temp_path = Filename.get_temp_dir_name ()
(* ------------------------------------------------------------------------- *)
(* Convenient conversion between files and (lists of) strings. *)
(* ------------------------------------------------------------------------- *)
let strings_of_file filename =
- let fd = try open_in filename
- with Sys_error _ ->
- failwith("strings_of_file: can't open "^filename) in
+ let fd =
+ try open_in filename
+ with Sys_error _ -> failwith ("strings_of_file: can't open " ^ filename)
+ in
let rec suck_lines acc =
- try let l = input_line fd in
- suck_lines (l::acc)
- with End_of_file -> List.rev acc in
+ try
+ let l = input_line fd in
+ suck_lines (l :: acc)
+ with End_of_file -> List.rev acc
+ in
let data = suck_lines [] in
- (close_in fd; data);;
+ close_in fd; data
-let string_of_file filename =
- String.concat "\n" (strings_of_file filename);;
+let string_of_file filename = String.concat "\n" (strings_of_file filename)
let file_of_string filename s =
let fd = open_out filename in
- output_string fd s; close_out fd;;
-
+ output_string fd s; close_out fd
(* ------------------------------------------------------------------------- *)
(* Iterative deepening. *)
(* ------------------------------------------------------------------------- *)
let rec deepen f n =
- try (*print_string "Searching with depth limit ";
- print_int n; print_newline();*) f n
- with Failure _ -> deepen f (n + 1);;
+ try
+ (*print_string "Searching with depth limit ";
+ print_int n; print_newline();*)
+ f n
+ with Failure _ -> deepen f (n + 1)
exception TooDeep
let deepen_until limit f n =
match compare limit 0 with
- | 0 -> raise TooDeep
- | -1 -> deepen f n
- | _ ->
- let rec d_until f n =
- try(* if !debugging
+ | 0 -> raise TooDeep
+ | -1 -> deepen f n
+ | _ ->
+ let rec d_until f n =
+ try
+ (* if !debugging
then (print_string "Searching with depth limit ";
- print_int n; print_newline()) ;*) f n
- with Failure x ->
- (*if !debugging then (Printf.printf "solver error : %s\n" x) ; *)
- if n = limit then raise TooDeep else d_until f (n + 1) in
- d_until f n
+ print_int n; print_newline()) ;*)
+ f n
+ with Failure x ->
+ (*if !debugging then (Printf.printf "solver error : %s\n" x) ; *)
+ if n = limit then raise TooDeep else d_until f (n + 1)
+ in
+ d_until f n
diff --git a/plugins/micromega/sos_lib.mli b/plugins/micromega/sos_lib.mli
index f01b632c67..2bbcbf336b 100644
--- a/plugins/micromega/sos_lib.mli
+++ b/plugins/micromega/sos_lib.mli
@@ -9,58 +9,54 @@
(************************************************************************)
val o : ('a -> 'b) -> ('c -> 'a) -> 'c -> 'b
-
val num_1 : Num.num
val pow10 : int -> Num.num
val pow2 : int -> Num.num
-
val implode : string list -> string
val explode : string -> string list
-
val funpow : int -> ('a -> 'a) -> 'a -> 'a
val tryfind : ('a -> 'b) -> 'a list -> 'b
-type ('a,'b) func =
- | Empty
- | Leaf of int * ('a*'b) list
- | Branch of int * int * ('a,'b) func * ('a,'b) func
+type ('a, 'b) func =
+ | Empty
+ | Leaf of int * ('a * 'b) list
+ | Branch of int * int * ('a, 'b) func * ('a, 'b) func
val undefined : ('a, 'b) func
val is_undefined : ('a, 'b) func -> bool
-val (|->) : 'a -> 'b -> ('a, 'b) func -> ('a, 'b) func
-val (|=>) : 'a -> 'b -> ('a, 'b) func
+val ( |-> ) : 'a -> 'b -> ('a, 'b) func -> ('a, 'b) func
+val ( |=> ) : 'a -> 'b -> ('a, 'b) func
val choose : ('a, 'b) func -> 'a * 'b
-val combine : ('a -> 'a -> 'a) -> ('a -> bool) -> ('b, 'a) func -> ('b, 'a) func -> ('b, 'a) func
-val (--) : int -> int -> int list
+val combine :
+ ('a -> 'a -> 'a)
+ -> ('a -> bool)
+ -> ('b, 'a) func
+ -> ('b, 'a) func
+ -> ('b, 'a) func
+
+val ( -- ) : int -> int -> int list
val tryapplyd : ('a, 'b) func -> 'a -> 'b -> 'b
val apply : ('a, 'b) func -> 'a -> 'b
-
val foldl : ('a -> 'b -> 'c -> 'a) -> 'a -> ('b, 'c) func -> 'a
val foldr : ('a -> 'b -> 'c -> 'c) -> ('a, 'b) func -> 'c -> 'c
val mapf : ('a -> 'b) -> ('c, 'a) func -> ('c, 'b) func
-
val undefine : 'a -> ('a, 'b) func -> ('a, 'b) func
-
val dom : ('a, 'b) func -> 'a list
val graph : ('a, 'b) func -> ('a * 'b) list
-
val union : 'a list -> 'a list -> 'a list
val subtract : 'a list -> 'a list -> 'a list
val sort : ('a -> 'a -> bool) -> 'a list -> 'a list
val setify : 'a list -> 'a list
val increasing : ('a -> 'b) -> 'a -> 'a -> bool
val allpairs : ('a -> 'b -> 'c) -> 'a list -> 'b list -> 'c list
-
val gcd_num : Num.num -> Num.num -> Num.num
val lcm_num : Num.num -> Num.num -> Num.num
val numerator : Num.num -> Num.num
val denominator : Num.num -> Num.num
val end_itlist : ('a -> 'a -> 'a) -> 'a list -> 'a
-
-val (>>) : ('a -> 'b * 'c) -> ('b -> 'd) -> 'a -> 'd * 'c
-val (++) : ('a -> 'b * 'c) -> ('c -> 'd * 'e) -> 'a -> ('b * 'd) * 'e
-
+val ( >> ) : ('a -> 'b * 'c) -> ('b -> 'd) -> 'a -> 'd * 'c
+val ( ++ ) : ('a -> 'b * 'c) -> ('c -> 'd * 'e) -> 'a -> ('b * 'd) * 'e
val a : 'a -> 'a list -> 'a * 'a list
val many : ('a -> 'b * 'a) -> 'a -> 'b list * 'a
val some : ('a -> bool) -> 'a list -> 'a * 'a list
@@ -70,10 +66,9 @@ val parser_or : ('a -> 'b) -> ('a -> 'b) -> 'a -> 'b
val isnum : string -> bool
val atleast : int -> ('a -> 'b * 'a) -> 'a -> 'b list * 'a
val listof : ('a -> 'b * 'c) -> ('c -> 'd * 'a) -> string -> 'a -> 'b list * 'c
-
val temp_path : string
val string_of_file : string -> string
val file_of_string : string -> string -> unit
-
val deepen_until : int -> (int -> 'a) -> int -> 'a
+
exception TooDeep
diff --git a/plugins/micromega/sos_types.ml b/plugins/micromega/sos_types.ml
index 0ba76fc0ea..988024968b 100644
--- a/plugins/micromega/sos_types.ml
+++ b/plugins/micromega/sos_types.ml
@@ -14,53 +14,53 @@ open Num
type vname = string
type term =
-| Zero
-| Const of Num.num
-| Var of vname
-| Opp of term
-| Add of (term * term)
-| Sub of (term * term)
-| Mul of (term * term)
-| Pow of (term * int)
-
+ | Zero
+ | Const of Num.num
+ | Var of vname
+ | Opp of term
+ | Add of (term * term)
+ | Sub of (term * term)
+ | Mul of (term * term)
+ | Pow of (term * int)
let rec output_term o t =
match t with
- | Zero -> output_string o "0"
- | Const n -> output_string o (string_of_num n)
- | Var n -> Printf.fprintf o "v%s" n
- | Opp t -> Printf.fprintf o "- (%a)" output_term t
- | Add(t1,t2) -> Printf.fprintf o "(%a)+(%a)" output_term t1 output_term t2
- | Sub(t1,t2) -> Printf.fprintf o "(%a)-(%a)" output_term t1 output_term t2
- | Mul(t1,t2) -> Printf.fprintf o "(%a)*(%a)" output_term t1 output_term t2
- | Pow(t1,i) -> Printf.fprintf o "(%a)^(%i)" output_term t1 i
+ | Zero -> output_string o "0"
+ | Const n -> output_string o (string_of_num n)
+ | Var n -> Printf.fprintf o "v%s" n
+ | Opp t -> Printf.fprintf o "- (%a)" output_term t
+ | Add (t1, t2) -> Printf.fprintf o "(%a)+(%a)" output_term t1 output_term t2
+ | Sub (t1, t2) -> Printf.fprintf o "(%a)-(%a)" output_term t1 output_term t2
+ | Mul (t1, t2) -> Printf.fprintf o "(%a)*(%a)" output_term t1 output_term t2
+ | Pow (t1, i) -> Printf.fprintf o "(%a)^(%i)" output_term t1 i
+
(* ------------------------------------------------------------------------- *)
(* Data structure for Positivstellensatz refutations. *)
(* ------------------------------------------------------------------------- *)
type positivstellensatz =
- Axiom_eq of int
- | Axiom_le of int
- | Axiom_lt of int
- | Rational_eq of num
- | Rational_le of num
- | Rational_lt of num
- | Square of term
- | Monoid of int list
- | Eqmul of term * positivstellensatz
- | Sum of positivstellensatz * positivstellensatz
- | Product of positivstellensatz * positivstellensatz;;
-
+ | Axiom_eq of int
+ | Axiom_le of int
+ | Axiom_lt of int
+ | Rational_eq of num
+ | Rational_le of num
+ | Rational_lt of num
+ | Square of term
+ | Monoid of int list
+ | Eqmul of term * positivstellensatz
+ | Sum of positivstellensatz * positivstellensatz
+ | Product of positivstellensatz * positivstellensatz
let rec output_psatz o = function
| Axiom_eq i -> Printf.fprintf o "Aeq(%i)" i
| Axiom_le i -> Printf.fprintf o "Ale(%i)" i
| Axiom_lt i -> Printf.fprintf o "Alt(%i)" i
- | Rational_eq n -> Printf.fprintf o "eq(%s)" (string_of_num n)
- | Rational_le n -> Printf.fprintf o "le(%s)" (string_of_num n)
- | Rational_lt n -> Printf.fprintf o "lt(%s)" (string_of_num n)
- | Square t -> Printf.fprintf o "(%a)^2" output_term t
- | Monoid l -> Printf.fprintf o "monoid"
- | Eqmul (t,ps) -> Printf.fprintf o "%a * %a" output_term t output_psatz ps
- | Sum (t1,t2) -> Printf.fprintf o "%a + %a" output_psatz t1 output_psatz t2
- | Product (t1,t2) -> Printf.fprintf o "%a * %a" output_psatz t1 output_psatz t2
+ | Rational_eq n -> Printf.fprintf o "eq(%s)" (string_of_num n)
+ | Rational_le n -> Printf.fprintf o "le(%s)" (string_of_num n)
+ | Rational_lt n -> Printf.fprintf o "lt(%s)" (string_of_num n)
+ | Square t -> Printf.fprintf o "(%a)^2" output_term t
+ | Monoid l -> Printf.fprintf o "monoid"
+ | Eqmul (t, ps) -> Printf.fprintf o "%a * %a" output_term t output_psatz ps
+ | Sum (t1, t2) -> Printf.fprintf o "%a + %a" output_psatz t1 output_psatz t2
+ | Product (t1, t2) ->
+ Printf.fprintf o "%a * %a" output_psatz t1 output_psatz t2
diff --git a/plugins/micromega/sos_types.mli b/plugins/micromega/sos_types.mli
index c55bb69e8a..ca9a43b1d0 100644
--- a/plugins/micromega/sos_types.mli
+++ b/plugins/micromega/sos_types.mli
@@ -13,28 +13,28 @@
type vname = string
type term =
-| Zero
-| Const of Num.num
-| Var of vname
-| Opp of term
-| Add of (term * term)
-| Sub of (term * term)
-| Mul of (term * term)
-| Pow of (term * int)
+ | Zero
+ | Const of Num.num
+ | Var of vname
+ | Opp of term
+ | Add of (term * term)
+ | Sub of (term * term)
+ | Mul of (term * term)
+ | Pow of (term * int)
val output_term : out_channel -> term -> unit
type positivstellensatz =
- Axiom_eq of int
- | Axiom_le of int
- | Axiom_lt of int
- | Rational_eq of Num.num
- | Rational_le of Num.num
- | Rational_lt of Num.num
- | Square of term
- | Monoid of int list
- | Eqmul of term * positivstellensatz
- | Sum of positivstellensatz * positivstellensatz
- | Product of positivstellensatz * positivstellensatz
+ | Axiom_eq of int
+ | Axiom_le of int
+ | Axiom_lt of int
+ | Rational_eq of Num.num
+ | Rational_le of Num.num
+ | Rational_lt of Num.num
+ | Square of term
+ | Monoid of int list
+ | Eqmul of term * positivstellensatz
+ | Sum of positivstellensatz * positivstellensatz
+ | Product of positivstellensatz * positivstellensatz
val output_psatz : out_channel -> positivstellensatz -> unit
diff --git a/plugins/micromega/vect.ml b/plugins/micromega/vect.ml
index a5f3b83c48..581dc79cef 100644
--- a/plugins/micromega/vect.ml
+++ b/plugins/micromega/vect.ml
@@ -11,177 +11,157 @@
open Num
open Mutils
+type var = int
(** [t] is the type of vectors.
A vector [(x1,v1) ; ... ; (xn,vn)] is such that:
- variables indexes are ordered (x1 < ... < xn
- values are all non-zero
*)
-type var = int
+
type t = (var * num) list
(** [equal v1 v2 = true] if the vectors are syntactically equal. *)
let rec equal v1 v2 =
- match v1 , v2 with
- | [] , [] -> true
- | [] , _ -> false
- | _::_ , [] -> false
- | (i1,n1)::v1 , (i2,n2)::v2 ->
- (Int.equal i1 i2) && n1 =/ n2 && equal v1 v2
+ match (v1, v2) with
+ | [], [] -> true
+ | [], _ -> false
+ | _ :: _, [] -> false
+ | (i1, n1) :: v1, (i2, n2) :: v2 -> Int.equal i1 i2 && n1 =/ n2 && equal v1 v2
let hash v =
let rec hash i = function
| [] -> i
- | (vr,vl)::l -> hash (i + (Hashtbl.hash (vr, float_of_num vl))) l in
- Hashtbl.hash (hash 0 v )
-
+ | (vr, vl) :: l -> hash (i + Hashtbl.hash (vr, float_of_num vl)) l
+ in
+ Hashtbl.hash (hash 0 v)
let null = []
+let is_null v = match v with [] | [(0, Int 0)] -> true | _ -> false
-let is_null v =
- match v with
- | [] | [0,Int 0] -> true
- | _ -> false
-
-let pp_var_num pp_var o (v,n) =
- if Int.equal v 0
- then if eq_num (Int 0) n then ()
- else Printf.fprintf o "%s" (string_of_num n)
+let pp_var_num pp_var o (v, n) =
+ if Int.equal v 0 then
+ if eq_num (Int 0) n then () else Printf.fprintf o "%s" (string_of_num n)
else
match n with
- | Int 1 -> pp_var o v
+ | Int 1 -> pp_var o v
| Int -1 -> Printf.fprintf o "-%a" pp_var v
- | Int 0 -> ()
- | _ -> Printf.fprintf o "%s*%a" (string_of_num n) pp_var v
+ | Int 0 -> ()
+ | _ -> Printf.fprintf o "%s*%a" (string_of_num n) pp_var v
-let pp_var_num_smt pp_var o (v,n) =
- if Int.equal v 0
- then if eq_num (Int 0) n then ()
- else Printf.fprintf o "%s" (string_of_num n)
+let pp_var_num_smt pp_var o (v, n) =
+ if Int.equal v 0 then
+ if eq_num (Int 0) n then () else Printf.fprintf o "%s" (string_of_num n)
else
match n with
- | Int 1 -> pp_var o v
+ | Int 1 -> pp_var o v
| Int -1 -> Printf.fprintf o "(- %a)" pp_var v
- | Int 0 -> ()
- | _ -> Printf.fprintf o "(* %s %a)" (string_of_num n) pp_var v
-
+ | Int 0 -> ()
+ | _ -> Printf.fprintf o "(* %s %a)" (string_of_num n) pp_var v
let rec pp_gen pp_var o v =
match v with
| [] -> output_string o "0"
| [e] -> pp_var_num pp_var o e
- | e::l -> Printf.fprintf o "%a + %a" (pp_var_num pp_var) e (pp_gen pp_var) l
-
+ | e :: l -> Printf.fprintf o "%a + %a" (pp_var_num pp_var) e (pp_gen pp_var) l
let pp_var o v = Printf.fprintf o "x%i" v
-
let pp o v = pp_gen pp_var o v
-let pp_smt o v =
- let list o v = List.iter (fun e -> Printf.fprintf o "%a " (pp_var_num_smt pp_var) e) v in
+let pp_smt o v =
+ let list o v =
+ List.iter (fun e -> Printf.fprintf o "%a " (pp_var_num_smt pp_var) e) v
+ in
Printf.fprintf o "(+ %a)" list v
-let from_list (l: num list) =
+let from_list (l : num list) =
let rec xfrom_list i l =
match l with
| [] -> []
- | e::l ->
- if e <>/ Int 0
- then (i,e)::(xfrom_list (i+1) l)
- else xfrom_list (i+1) l in
-
+ | e :: l ->
+ if e <>/ Int 0 then (i, e) :: xfrom_list (i + 1) l
+ else xfrom_list (i + 1) l
+ in
xfrom_list 0 l
let zero_num = Int 0
-
let to_list m =
let rec xto_list i l =
match l with
| [] -> []
- | (x,v)::l' ->
- if i = x then v::(xto_list (i+1) l') else zero_num ::(xto_list (i+1) l) in
+ | (x, v) :: l' ->
+ if i = x then v :: xto_list (i + 1) l' else zero_num :: xto_list (i + 1) l
+ in
xto_list 0 m
-
-let cons i v rst = if v =/ Int 0 then rst else (i,v)::rst
+let cons i v rst = if v =/ Int 0 then rst else (i, v) :: rst
let rec update i f t =
match t with
| [] -> cons i (f zero_num) []
- | (k,v)::l ->
- match Int.compare i k with
- | 0 -> cons k (f v) l
- | -1 -> cons i (f zero_num) t
- | 1 -> (k,v) ::(update i f l)
- | _ -> failwith "compare_num"
+ | (k, v) :: l -> (
+ match Int.compare i k with
+ | 0 -> cons k (f v) l
+ | -1 -> cons i (f zero_num) t
+ | 1 -> (k, v) :: update i f l
+ | _ -> failwith "compare_num" )
let rec set i n t =
match t with
| [] -> cons i n []
- | (k,v)::l ->
- match Int.compare i k with
- | 0 -> cons k n l
- | -1 -> cons i n t
- | 1 -> (k,v) :: (set i n l)
- | _ -> failwith "compare_num"
-
-let cst n = if n =/ Int 0 then [] else [0,n]
+ | (k, v) :: l -> (
+ match Int.compare i k with
+ | 0 -> cons k n l
+ | -1 -> cons i n t
+ | 1 -> (k, v) :: set i n l
+ | _ -> failwith "compare_num" )
+let cst n = if n =/ Int 0 then [] else [(0, n)]
let mul z t =
match z with
| Int 0 -> []
| Int 1 -> t
- | _ -> List.map (fun (i,n) -> (i, mult_num z n)) t
+ | _ -> List.map (fun (i, n) -> (i, mult_num z n)) t
let div z t =
- if z <>/ Int 1
- then List.map (fun (x,nx) -> (x,nx // z)) t
- else t
-
-
-let uminus t = List.map (fun (i,n) -> i, minus_num n) t
-
-
-let rec add (ve1:t) (ve2:t) =
- match ve1 , ve2 with
- | [] , v | v , [] -> v
- | (v1,c1)::l1 , (v2,c2)::l2 ->
- let cmp = Util.pervasives_compare v1 v2 in
- if cmp == 0 then
- let s = add_num c1 c2 in
- if eq_num (Int 0) s
- then add l1 l2
- else (v1,s)::(add l1 l2)
- else if cmp < 0 then (v1,c1) :: (add l1 ve2)
- else (v2,c2) :: (add l2 ve1)
-
-
-let rec xmul_add (n1:num) (ve1:t) (n2:num) (ve2:t) =
- match ve1 , ve2 with
- | [] , _ -> mul n2 ve2
- | _ , [] -> mul n1 ve1
- | (v1,c1)::l1 , (v2,c2)::l2 ->
- let cmp = Util.pervasives_compare v1 v2 in
- if cmp == 0 then
- let s = ( n1 */ c1) +/ (n2 */ c2) in
- if eq_num (Int 0) s
- then xmul_add n1 l1 n2 l2
- else (v1,s)::(xmul_add n1 l1 n2 l2)
- else if cmp < 0 then (v1,n1 */ c1) :: (xmul_add n1 l1 n2 ve2)
- else (v2,n2 */c2) :: (xmul_add n1 ve1 n2 l2)
+ if z <>/ Int 1 then List.map (fun (x, nx) -> (x, nx // z)) t else t
+
+let uminus t = List.map (fun (i, n) -> (i, minus_num n)) t
+
+let rec add (ve1 : t) (ve2 : t) =
+ match (ve1, ve2) with
+ | [], v | v, [] -> v
+ | (v1, c1) :: l1, (v2, c2) :: l2 ->
+ let cmp = Util.pervasives_compare v1 v2 in
+ if cmp == 0 then
+ let s = add_num c1 c2 in
+ if eq_num (Int 0) s then add l1 l2 else (v1, s) :: add l1 l2
+ else if cmp < 0 then (v1, c1) :: add l1 ve2
+ else (v2, c2) :: add l2 ve1
+
+let rec xmul_add (n1 : num) (ve1 : t) (n2 : num) (ve2 : t) =
+ match (ve1, ve2) with
+ | [], _ -> mul n2 ve2
+ | _, [] -> mul n1 ve1
+ | (v1, c1) :: l1, (v2, c2) :: l2 ->
+ let cmp = Util.pervasives_compare v1 v2 in
+ if cmp == 0 then
+ let s = (n1 */ c1) +/ (n2 */ c2) in
+ if eq_num (Int 0) s then xmul_add n1 l1 n2 l2
+ else (v1, s) :: xmul_add n1 l1 n2 l2
+ else if cmp < 0 then (v1, n1 */ c1) :: xmul_add n1 l1 n2 ve2
+ else (v2, n2 */ c2) :: xmul_add n1 ve1 n2 l2
let mul_add n1 ve1 n2 ve2 =
- if n1 =/ Int 1 && n2 =/ Int 1
- then add ve1 ve2
- else xmul_add n1 ve1 n2 ve2
+ if n1 =/ Int 1 && n2 =/ Int 1 then add ve1 ve2 else xmul_add n1 ve1 n2 ve2
-
-let compare : t -> t -> int = Mutils.Cmp.compare_list (fun x y -> Mutils.Cmp.compare_lexical
- [
- (fun () -> Int.compare (fst x) (fst y));
- (fun () -> compare_num (snd x) (snd y))])
+let compare : t -> t -> int =
+ Mutils.Cmp.compare_list (fun x y ->
+ Mutils.Cmp.compare_lexical
+ [ (fun () -> Int.compare (fst x) (fst y))
+ ; (fun () -> compare_num (snd x) (snd y)) ])
(** [tail v vect] returns
- [None] if [v] is not a variable of the vector [vect]
@@ -189,150 +169,103 @@ let compare : t -> t -> int = Mutils.Cmp.compare_list (fun x y -> Mutils.Cmp.com
and [rst] is the remaining of the vector
We exploit that vectors are ordered lists
*)
-let rec tail (v:var) (vect:t) =
+let rec tail (v : var) (vect : t) =
match vect with
| [] -> None
- | (v',vl)::vect' ->
- match Int.compare v' v with
- | 0 -> Some (vl,vect) (* Ok, found *)
- | -1 -> tail v vect' (* Might be in the tail *)
- | _ -> None (* Hopeless *)
-
-let get v vect =
- match tail v vect with
- | None -> Int 0
- | Some(vl,_) -> vl
-
-let is_constant v =
- match v with
- | [] | [0,_] -> true
- | _ -> false
-
-
-
-let get_cst vect =
- match vect with
- | (0,v)::_ -> v
- | _ -> Int 0
-
-let choose v =
- match v with
- | [] -> None
- | (vr,vl)::rst -> Some (vr,vl,rst)
-
-
-let rec fresh v =
- match v with
- | [] -> 1
- | [v,_] -> v + 1
- | _::v -> fresh v
-
-
-let variables v =
- List.fold_left (fun acc (x,_) -> ISet.add x acc) ISet.empty v
-
-let decomp_cst v =
- match v with
- | (0,vl)::v -> vl,v
- | _ -> Int 0,v
+ | (v', vl) :: vect' -> (
+ match Int.compare v' v with
+ | 0 -> Some (vl, vect) (* Ok, found *)
+ | -1 -> tail v vect' (* Might be in the tail *)
+ | _ -> None )
+
+(* Hopeless *)
+
+let get v vect = match tail v vect with None -> Int 0 | Some (vl, _) -> vl
+let is_constant v = match v with [] | [(0, _)] -> true | _ -> false
+let get_cst vect = match vect with (0, v) :: _ -> v | _ -> Int 0
+let choose v = match v with [] -> None | (vr, vl) :: rst -> Some (vr, vl, rst)
+let rec fresh v = match v with [] -> 1 | [(v, _)] -> v + 1 | _ :: v -> fresh v
+let variables v = List.fold_left (fun acc (x, _) -> ISet.add x acc) ISet.empty v
+let decomp_cst v = match v with (0, vl) :: v -> (vl, v) | _ -> (Int 0, v)
let rec decomp_at i v =
match v with
- | [] -> (Int 0 , null)
- | (vr,vl)::r -> if i = vr then (vl,r)
- else if i < vr then (Int 0,v)
- else decomp_at i r
-
-let decomp_fst v =
- match v with
- | [] -> ((0,Int 0),[])
- | x::v -> (x,v)
+ | [] -> (Int 0, null)
+ | (vr, vl) :: r ->
+ if i = vr then (vl, r) else if i < vr then (Int 0, v) else decomp_at i r
-
-let fold f acc v =
- List.fold_left (fun acc (v,i) -> f acc v i) acc v
+let decomp_fst v = match v with [] -> ((0, Int 0), []) | x :: v -> (x, v)
+let fold f acc v = List.fold_left (fun acc (v, i) -> f acc v i) acc v
let fold_error f acc v =
let rec fold acc v =
match v with
| [] -> Some acc
- | (x,i)::v' -> match f acc x i with
- | None -> None
- | Some acc' -> fold acc' v' in
+ | (x, i) :: v' -> (
+ match f acc x i with None -> None | Some acc' -> fold acc' v' )
+ in
fold acc v
-
-
let rec find p v =
match v with
| [] -> None
- | (v,n)::v' -> match p v n with
- | None -> find p v'
- | Some r -> Some r
-
-
-let for_all p l =
- List.for_all (fun (v,n) -> p v n) l
+ | (v, n) :: v' -> ( match p v n with None -> find p v' | Some r -> Some r )
-
-let decr_var i v = List.map (fun (v,n) -> (v-i,n)) v
-let incr_var i v = List.map (fun (v,n) -> (v+i,n)) v
+let for_all p l = List.for_all (fun (v, n) -> p v n) l
+let decr_var i v = List.map (fun (v, n) -> (v - i, n)) v
+let incr_var i v = List.map (fun (v, n) -> (v + i, n)) v
open Big_int
let gcd v =
- let res = fold (fun c _ n ->
- assert (Int.equal (compare_big_int (denominator n) unit_big_int) 0);
- gcd_big_int c (numerator n)) zero_big_int v in
- if Int.equal (compare_big_int res zero_big_int) 0
- then unit_big_int else res
+ let res =
+ fold
+ (fun c _ n ->
+ assert (Int.equal (compare_big_int (denominator n) unit_big_int) 0);
+ gcd_big_int c (numerator n))
+ zero_big_int v
+ in
+ if Int.equal (compare_big_int res zero_big_int) 0 then unit_big_int else res
let normalise v =
let ppcm = fold (fun c _ n -> ppcm c (denominator n)) unit_big_int v in
- let gcd =
+ let gcd =
let gcd = fold (fun c _ n -> gcd_big_int c (numerator n)) zero_big_int v in
- if Int.equal (compare_big_int gcd zero_big_int) 0 then unit_big_int else gcd in
- List.map (fun (x,v) -> (x, v */ (Big_int ppcm) // (Big_int gcd))) v
+ if Int.equal (compare_big_int gcd zero_big_int) 0 then unit_big_int else gcd
+ in
+ List.map (fun (x, v) -> (x, v */ Big_int ppcm // Big_int gcd)) v
let rec exists2 p vect1 vect2 =
- match vect1 , vect2 with
- | _ , [] | [], _ -> None
- | (v1,n1)::vect1' , (v2, n2) :: vect2' ->
- if Int.equal v1 v2
- then
- if p n1 n2
- then Some (v1,n1,n2)
- else
- exists2 p vect1' vect2'
- else
- if v1 < v2
- then exists2 p vect1' vect2
- else exists2 p vect1 vect2'
+ match (vect1, vect2) with
+ | _, [] | [], _ -> None
+ | (v1, n1) :: vect1', (v2, n2) :: vect2' ->
+ if Int.equal v1 v2 then
+ if p n1 n2 then Some (v1, n1, n2) else exists2 p vect1' vect2'
+ else if v1 < v2 then exists2 p vect1' vect2
+ else exists2 p vect1 vect2'
let dotproduct v1 v2 =
let rec dot acc v1 v2 =
- match v1, v2 with
- | [] , _ | _ , [] -> acc
- | (x1,n1)::v1', (x2,n2)::v2' ->
- if x1 == x2
- then dot (acc +/ n1 */ n2) v1' v2'
- else if x1 < x2
- then dot acc v1' v2
- else dot acc v1 v2' in
+ match (v1, v2) with
+ | [], _ | _, [] -> acc
+ | (x1, n1) :: v1', (x2, n2) :: v2' ->
+ if x1 == x2 then dot (acc +/ (n1 */ n2)) v1' v2'
+ else if x1 < x2 then dot acc v1' v2
+ else dot acc v1 v2'
+ in
dot (Int 0) v1 v2
-
-let map f v = List.map (fun (x,v) -> f x v) v
+let map f v = List.map (fun (x, v) -> f x v) v
let abs_min_elt v =
match v with
| [] -> None
- | (v,vl)::r ->
- Some (List.fold_left (fun (v1,vl1) (v2,vl2) ->
- if abs_num vl1 </ abs_num vl2
- then (v1,vl1) else (v2,vl2) ) (v,vl) r)
-
-
-let partition p = List.partition (fun (vr,vl) -> p vr vl)
-
+ | (v, vl) :: r ->
+ Some
+ (List.fold_left
+ (fun (v1, vl1) (v2, vl2) ->
+ if abs_num vl1 </ abs_num vl2 then (v1, vl1) else (v2, vl2))
+ (v, vl) r)
+
+let partition p = List.partition (fun (vr, vl) -> p vr vl)
let mkvar x = set x (Int 1) null
diff --git a/plugins/micromega/vect.mli b/plugins/micromega/vect.mli
index 40ef8078e4..83d9ef32b0 100644
--- a/plugins/micromega/vect.mli
+++ b/plugins/micromega/vect.mli
@@ -11,9 +11,11 @@
open Num
open Mutils
-type var = int (** Variables are simply (positive) integers. *)
+type var = int
+(** Variables are simply (positive) integers. *)
-type t (** The type of vectors or equivalently linear expressions.
+type t
+(** The type of vectors or equivalently linear expressions.
The current implementation is using association lists.
A list [(0,c),(x1,ai),...,(xn,an)] represents the linear expression
c + a1.xn + ... an.xn where ai are rational constants and xi are variables.
@@ -34,140 +36,137 @@ val compare : t -> t -> int
(** {1 Basic accessors and utility functions} *)
+val pp_gen : (out_channel -> var -> unit) -> out_channel -> t -> unit
(** [pp_gen pp_var o v] prints the representation of the vector [v] over the channel [o] *)
-val pp_gen : (out_channel -> var -> unit) -> out_channel -> t -> unit
-(** [pp o v] prints the representation of the vector [v] over the channel [o] *)
val pp : out_channel -> t -> unit
+(** [pp o v] prints the representation of the vector [v] over the channel [o] *)
-(** [pp_smt o v] prints the representation of the vector [v] over the channel [o] using SMTLIB conventions *)
val pp_smt : out_channel -> t -> unit
+(** [pp_smt o v] prints the representation of the vector [v] over the channel [o] using SMTLIB conventions *)
-(** [variables v] returns the set of variables with non-zero coefficients *)
val variables : t -> ISet.t
+(** [variables v] returns the set of variables with non-zero coefficients *)
-(** [get_cst v] returns c i.e. the coefficient of the variable zero *)
val get_cst : t -> num
+(** [get_cst v] returns c i.e. the coefficient of the variable zero *)
-(** [decomp_cst v] returns the pair (c,a1.x1+...+an.xn) *)
val decomp_cst : t -> num * t
+(** [decomp_cst v] returns the pair (c,a1.x1+...+an.xn) *)
-(** [decomp_cst v] returns the pair (ai, ai+1.xi+...+an.xn) *)
val decomp_at : int -> t -> num * t
+(** [decomp_cst v] returns the pair (ai, ai+1.xi+...+an.xn) *)
val decomp_fst : t -> (var * num) * t
-(** [cst c] returns the vector v=c+0.x1+...+0.xn *)
val cst : num -> t
+(** [cst c] returns the vector v=c+0.x1+...+0.xn *)
+val is_constant : t -> bool
(** [is_constant v] holds if [v] is a constant vector i.e. v=c+0.x1+...+0.xn
*)
-val is_constant : t -> bool
-(** [null] is the empty vector i.e. 0+0.x1+...+0.xn *)
val null : t
+(** [null] is the empty vector i.e. 0+0.x1+...+0.xn *)
-(** [is_null v] returns whether [v] is the [null] vector i.e [equal v null] *)
val is_null : t -> bool
+(** [is_null v] returns whether [v] is the [null] vector i.e [equal v null] *)
+val get : var -> t -> num
(** [get xi v] returns the coefficient ai of the variable [xi].
[get] is also defined for the variable 0 *)
-val get : var -> t -> num
+val set : var -> num -> t -> t
(** [set xi ai' v] returns the vector c+a1.x1+...ai'.xi+...+an.xn
i.e. the coefficient of the variable xi is set to ai' *)
-val set : var -> num -> t -> t
-(** [mkvar xi] returns 1.xi *)
val mkvar : var -> t
+(** [mkvar xi] returns 1.xi *)
+val update : var -> (num -> num) -> t -> t
(** [update xi f v] returns c+a1.x1+...+f(ai).xi+...+an.xn *)
-val update : var -> (num -> num) -> t -> t
-(** [fresh v] return the fresh variable with index 1+ max (variables v) *)
val fresh : t -> int
+(** [fresh v] return the fresh variable with index 1+ max (variables v) *)
+val choose : t -> (var * num * t) option
(** [choose v] decomposes a vector [v] depending on whether it is [null] or not.
@return None if v is [null]
@return Some(x,n,r) where v = r + n.x x is the smallest variable with non-zero coefficient n <> 0.
*)
-val choose : t -> (var * num * t) option
-(** [from_list l] returns the vector c+a1.x1...an.xn from the list of coefficient [l=c;a1;...;an] *)
val from_list : num list -> t
+(** [from_list l] returns the vector c+a1.x1...an.xn from the list of coefficient [l=c;a1;...;an] *)
+val to_list : t -> num list
(** [to_list v] returns the list of all coefficient of the vector v i.e. [c;a1;...;an]
The list representation is (obviously) not sparsed
and therefore certain ai may be 0 *)
-val to_list : t -> num list
+val decr_var : int -> t -> t
(** [decr_var i v] decrements the variables of the vector [v] by the amount [i].
Beware, it is only defined if all the variables of v are greater than i
*)
-val decr_var : int -> t -> t
+val incr_var : int -> t -> t
(** [incr_var i v] increments the variables of the vector [v] by the amount [i].
*)
-val incr_var : int -> t -> t
+val gcd : t -> Big_int.big_int
(** [gcd v] returns gcd(num(c),num(a1),...,num(an)) where num extracts
the numerator of a rational value. *)
-val gcd : t -> Big_int.big_int
-(** [normalise v] returns a vector with only integer coefficients *)
val normalise : t -> t
-
+(** [normalise v] returns a vector with only integer coefficients *)
(** {1 Linear arithmetics} *)
+val add : t -> t -> t
(** [add v1 v2] is vector addition.
@param v1 is of the form c +a1.x1 +...+an.xn
@param v2 is of the form c'+a1'.x1 +...+an'.xn
@return c1+c1'+ (a1+a1').x1 + ... + (an+an').xn
*)
-val add : t -> t -> t
+val mul : num -> t -> t
(** [mul a v] is vector multiplication of vector [v] by a scalar [a].
@return a.v = a.c+a.a1.x1+...+a.an.xn *)
-val mul : num -> t -> t
-(** [mul_add c1 v1 c2 v2] returns the linear combination c1.v1+c2.v2 *)
val mul_add : num -> t -> num -> t -> t
+(** [mul_add c1 v1 c2 v2] returns the linear combination c1.v1+c2.v2 *)
-(** [div c1 v1] returns the mutiplication by the inverse of c1 i.e (1/c1).v1 *)
val div : num -> t -> t
+(** [div c1 v1] returns the mutiplication by the inverse of c1 i.e (1/c1).v1 *)
-(** [uminus v] @return -v the opposite vector of v i.e. (-1).v *)
val uminus : t -> t
+(** [uminus v] @return -v the opposite vector of v i.e. (-1).v *)
(** {1 Iterators} *)
-(** [fold f acc v] returns f (f (f acc 0 c ) x1 a1 ) ... xn an *)
val fold : ('acc -> var -> num -> 'acc) -> 'acc -> t -> 'acc
+(** [fold f acc v] returns f (f (f acc 0 c ) x1 a1 ) ... xn an *)
+val fold_error : ('acc -> var -> num -> 'acc option) -> 'acc -> t -> 'acc option
(** [fold_error f acc v] is the same as
[fold (fun acc x i -> match acc with None -> None | Some acc' -> f acc' x i) (Some acc) v]
but with early exit...
*)
-val fold_error : ('acc -> var -> num -> 'acc option) -> 'acc -> t -> 'acc option
+val find : (var -> num -> 'c option) -> t -> 'c option
(** [find f v] returns the first [f xi ai] such that [f xi ai <> None].
If no such xi ai exists, it returns None *)
-val find : (var -> num -> 'c option) -> t -> 'c option
-(** [for_all p v] returns /\_{i>=0} (f xi ai) *)
val for_all : (var -> num -> bool) -> t -> bool
+(** [for_all p v] returns /\_{i>=0} (f xi ai) *)
+val exists2 : (num -> num -> bool) -> t -> t -> (var * num * num) option
(** [exists2 p v v'] returns Some(xi,ai,ai')
if p(xi,ai,ai') holds and ai,ai' <> 0.
It returns None if no such pair of coefficient exists. *)
-val exists2 : (num -> num -> bool) -> t -> t -> (var * num * num) option
-(** [dotproduct v1 v2] is the dot product of v1 and v2. *)
val dotproduct : t -> t -> num
+(** [dotproduct v1 v2] is the dot product of v1 and v2. *)
val map : (var -> num -> 'a) -> t -> 'a list
-
val abs_min_elt : t -> (var * num) option
-
val partition : (var -> num -> bool) -> t -> t * t
diff --git a/plugins/micromega/zify.ml b/plugins/micromega/zify.ml
index 0a57677220..5d8ae83853 100644
--- a/plugins/micromega/zify.ml
+++ b/plugins/micromega/zify.ml
@@ -27,10 +27,7 @@ let pr_constr env evd e = Printer.pr_econstr_env env evd e
let rec find_option pred l =
match l with
| [] -> raise Not_found
- | e::l -> match pred e with
- | Some r -> r
- | None -> find_option pred l
-
+ | e :: l -> ( match pred e with Some r -> r | None -> find_option pred l )
(** [HConstr] is a map indexed by EConstr.t.
It should only be used using closed terms.
@@ -39,8 +36,7 @@ module HConstr = struct
module M = Map.Make (struct
type t = EConstr.t
- let compare c c' =
- Constr.compare (unsafe_to_constr c) (unsafe_to_constr c')
+ let compare c c' = Constr.compare (unsafe_to_constr c) (unsafe_to_constr c')
end)
type 'a t = 'a list M.t
@@ -52,91 +48,89 @@ module HConstr = struct
M.add h (e :: l) m
let empty = M.empty
-
let find h m = match lfind h m with e :: _ -> e | [] -> raise Not_found
-
let find_all = lfind
let fold f m acc =
M.fold (fun k l acc -> List.fold_left (fun acc e -> f k e acc) acc l) m acc
-
end
-
(** [get_projections_from_constant (evd,c) ]
returns an array of constr [| a1,.. an|] such that [c] is defined as
Definition c := mk a1 .. an with mk a constructor.
ai is therefore either a type parameter or a projection.
*)
-
let get_projections_from_constant (evd, i) =
- match EConstr.kind evd (Reductionops.clos_whd_flags CClosure.all (Global.env ()) evd i) with
+ match
+ EConstr.kind evd
+ (Reductionops.clos_whd_flags CClosure.all (Global.env ()) evd i)
+ with
| App (c, a) -> Some a
| _ ->
- raise (CErrors.user_err Pp.(str "The hnf of term " ++ pr_constr (Global.env ()) evd i
- ++ str " should be an application i.e. (c a1 ... an)"))
+ raise
+ (CErrors.user_err
+ Pp.(
+ str "The hnf of term "
+ ++ pr_constr (Global.env ()) evd i
+ ++ str " should be an application i.e. (c a1 ... an)"))
(** An instance of type, say T, is registered into a hashtable, say TableT. *)
type 'a decl =
- { decl: EConstr.t
+ { decl : EConstr.t
; (* Registered type instance *)
- deriv: 'a
- (* Projections of insterest *) }
-
+ deriv : 'a (* Projections of insterest *) }
module EInjT = struct
type t =
- { isid: bool
+ { isid : bool
; (* S = T -> inj = fun x -> x*)
- source: EConstr.t
+ source : EConstr.t
; (* S *)
- target: EConstr.t
+ target : EConstr.t
; (* T *)
(* projections *)
- inj: EConstr.t
+ inj : EConstr.t
; (* S -> T *)
- pred: EConstr.t
+ pred : EConstr.t
; (* T -> Prop *)
- cstr: EConstr.t option
- (* forall x, pred (inj x) *) }
+ cstr : EConstr.t option (* forall x, pred (inj x) *) }
end
module EBinOpT = struct
type t =
{ (* Op : source1 -> source2 -> source3 *)
- source1: EConstr.t
- ; source2: EConstr.t
- ; source3: EConstr.t
- ; target: EConstr.t
- ; inj1: EConstr.t
+ source1 : EConstr.t
+ ; source2 : EConstr.t
+ ; source3 : EConstr.t
+ ; target : EConstr.t
+ ; inj1 : EConstr.t
; (* InjTyp source1 target *)
- inj2: EConstr.t
+ inj2 : EConstr.t
; (* InjTyp source2 target *)
- inj3: EConstr.t
+ inj3 : EConstr.t
; (* InjTyp source3 target *)
- tbop: EConstr.t
- (* TBOpInj *) }
+ tbop : EConstr.t (* TBOpInj *) }
end
module ECstOpT = struct
- type t = {source: EConstr.t; target: EConstr.t; inj: EConstr.t}
+ type t = {source : EConstr.t; target : EConstr.t; inj : EConstr.t}
end
module EUnOpT = struct
type t =
- { source1: EConstr.t
- ; source2: EConstr.t
- ; target: EConstr.t
- ; inj1_t: EConstr.t
- ; inj2_t: EConstr.t
- ; unop: EConstr.t }
+ { source1 : EConstr.t
+ ; source2 : EConstr.t
+ ; target : EConstr.t
+ ; inj1_t : EConstr.t
+ ; inj2_t : EConstr.t
+ ; unop : EConstr.t }
end
module EBinRelT = struct
type t =
- {source: EConstr.t; target: EConstr.t; inj: EConstr.t; brel: EConstr.t}
+ {source : EConstr.t; target : EConstr.t; inj : EConstr.t; brel : EConstr.t}
end
module EPropBinOpT = struct
@@ -147,37 +141,32 @@ module EPropUnOpT = struct
type t = EConstr.t
end
-
module ESatT = struct
- type t = {parg1: EConstr.t; parg2: EConstr.t; satOK: EConstr.t}
+ type t = {parg1 : EConstr.t; parg2 : EConstr.t; satOK : EConstr.t}
end
(* Different type of declarations *)
type decl_kind =
| PropOp of EPropBinOpT.t decl
- | PropUnOp of EPropUnOpT.t decl
- | InjTyp of EInjT.t decl
- | BinRel of EBinRelT.t decl
- | BinOp of EBinOpT.t decl
- | UnOp of EUnOpT.t decl
- | CstOp of ECstOpT.t decl
- | Saturate of ESatT.t decl
-
-
-let get_decl = function
+ | PropUnOp of EPropUnOpT.t decl
+ | InjTyp of EInjT.t decl
+ | BinRel of EBinRelT.t decl
+ | BinOp of EBinOpT.t decl
+ | UnOp of EUnOpT.t decl
+ | CstOp of ECstOpT.t decl
+ | Saturate of ESatT.t decl
+
+let get_decl = function
| PropOp d -> d.decl
- | PropUnOp d -> d.decl
- | InjTyp d -> d.decl
- | BinRel d -> d.decl
- | BinOp d -> d.decl
- | UnOp d -> d.decl
- | CstOp d -> d.decl
- | Saturate d -> d.decl
-
-type term_kind =
- | Application of EConstr.constr
- | OtherTerm of EConstr.constr
+ | PropUnOp d -> d.decl
+ | InjTyp d -> d.decl
+ | BinRel d -> d.decl
+ | BinOp d -> d.decl
+ | UnOp d -> d.decl
+ | CstOp d -> d.decl
+ | Saturate d -> d.decl
+type term_kind = Application of EConstr.constr | OtherTerm of EConstr.constr
module type Elt = sig
type elt
@@ -185,11 +174,9 @@ module type Elt = sig
val name : string
(** name *)
- val table : (term_kind * decl_kind) HConstr.t ref
-
+ val table : (term_kind * decl_kind) HConstr.t ref
val cast : elt decl -> decl_kind
-
- val dest : decl_kind -> (elt decl) option
+ val dest : decl_kind -> elt decl option
val get_key : int
(** [get_key] is the type-index used as key for the instance *)
@@ -199,19 +186,14 @@ module type Elt = sig
built from the type-instance i and the arguments (type indexes and projections)
of the type-class constructor. *)
- (* val arity : int*)
-
+ (* val arity : int*)
end
-
-let table = Summary.ref ~name:("zify_table") HConstr.empty
-
-let saturate = Summary.ref ~name:("zify_saturate") HConstr.empty
-
+let table = Summary.ref ~name:"zify_table" HConstr.empty
+let saturate = Summary.ref ~name:"zify_saturate" HConstr.empty
let table_cache = ref HConstr.empty
let saturate_cache = ref HConstr.empty
-
(** Each type-class gives rise to a different table.
They only differ on how projections are extracted. *)
module EInj = struct
@@ -220,186 +202,129 @@ module EInj = struct
type elt = EInjT.t
let name = "EInj"
-
let table = table
-
let cast x = InjTyp x
-
- let dest = function
- | InjTyp x -> Some x
- | _ -> None
-
+ let dest = function InjTyp x -> Some x | _ -> None
let mk_elt evd i (a : EConstr.t array) =
let isid = EConstr.eq_constr evd a.(0) a.(1) in
{ isid
- ; source= a.(0)
- ; target= a.(1)
- ; inj= a.(2)
- ; pred= a.(3)
- ; cstr= (if isid then None else Some a.(4)) }
+ ; source = a.(0)
+ ; target = a.(1)
+ ; inj = a.(2)
+ ; pred = a.(3)
+ ; cstr = (if isid then None else Some a.(4)) }
let get_key = 0
-
end
module EBinOp = struct
type elt = EBinOpT.t
+
open EBinOpT
let name = "BinOp"
-
let table = table
let mk_elt evd i a =
- { source1= a.(0)
- ; source2= a.(1)
- ; source3= a.(2)
- ; target= a.(3)
- ; inj1= a.(5)
- ; inj2= a.(6)
- ; inj3= a.(7)
- ; tbop= a.(9) }
+ { source1 = a.(0)
+ ; source2 = a.(1)
+ ; source3 = a.(2)
+ ; target = a.(3)
+ ; inj1 = a.(5)
+ ; inj2 = a.(6)
+ ; inj3 = a.(7)
+ ; tbop = a.(9) }
let get_key = 4
-
-
let cast x = BinOp x
-
- let dest = function
- | BinOp x -> Some x
- | _ -> None
-
+ let dest = function BinOp x -> Some x | _ -> None
end
module ECstOp = struct
type elt = ECstOpT.t
+
open ECstOpT
let name = "CstOp"
-
let table = table
-
let cast x = CstOp x
-
- let dest = function
- | CstOp x -> Some x
- | _ -> None
-
-
- let mk_elt evd i a = {source= a.(0); target= a.(1); inj= a.(3)}
-
+ let dest = function CstOp x -> Some x | _ -> None
+ let mk_elt evd i a = {source = a.(0); target = a.(1); inj = a.(3)}
let get_key = 2
-
end
module EUnOp = struct
type elt = EUnOpT.t
+
open EUnOpT
let name = "UnOp"
-
let table = table
-
let cast x = UnOp x
-
- let dest = function
- | UnOp x -> Some x
- | _ -> None
-
+ let dest = function UnOp x -> Some x | _ -> None
let mk_elt evd i a =
- { source1= a.(0)
- ; source2= a.(1)
- ; target= a.(2)
- ; inj1_t= a.(4)
- ; inj2_t= a.(5)
- ; unop= a.(6) }
+ { source1 = a.(0)
+ ; source2 = a.(1)
+ ; target = a.(2)
+ ; inj1_t = a.(4)
+ ; inj2_t = a.(5)
+ ; unop = a.(6) }
let get_key = 3
-
end
module EBinRel = struct
type elt = EBinRelT.t
+
open EBinRelT
let name = "BinRel"
-
let table = table
-
let cast x = BinRel x
+ let dest = function BinRel x -> Some x | _ -> None
- let dest = function
- | BinRel x -> Some x
- | _ -> None
-
- let mk_elt evd i a = {source= a.(0); target= a.(1); inj= a.(3); brel= a.(4)}
+ let mk_elt evd i a =
+ {source = a.(0); target = a.(1); inj = a.(3); brel = a.(4)}
let get_key = 2
-
end
module EPropOp = struct
type elt = EConstr.t
let name = "PropBinOp"
-
let table = table
-
let cast x = PropOp x
-
- let dest = function
- | PropOp x -> Some x
- | _ -> None
-
+ let dest = function PropOp x -> Some x | _ -> None
let mk_elt evd i a = i
-
let get_key = 0
-
end
module EPropUnOp = struct
type elt = EConstr.t
let name = "PropUnOp"
-
let table = table
-
let cast x = PropUnOp x
-
- let dest = function
- | PropUnOp x -> Some x
- | _ -> None
-
+ let dest = function PropUnOp x -> Some x | _ -> None
let mk_elt evd i a = i
-
let get_key = 0
-
end
-
-
-let constr_of_term_kind = function
- | Application c -> c
- | OtherTerm c -> c
-
-
+let constr_of_term_kind = function Application c -> c | OtherTerm c -> c
let fold_declared_const f evd acc =
HConstr.fold
- (fun _ (_,e) acc -> f (fst (EConstr.destConst evd (get_decl e))) acc)
- (!table_cache) acc
-
-
+ (fun _ (_, e) acc -> f (fst (EConstr.destConst evd (get_decl e))) acc)
+ !table_cache acc
module type S = sig
val register : Constrexpr.constr_expr -> unit
-
val print : unit -> unit
end
-
module MakeTable (E : Elt) = struct
(** Given a term [c] and its arguments ai,
we construct a HConstr.t table that is
@@ -410,33 +335,34 @@ module MakeTable (E : Elt) = struct
let make_elt (evd, i) =
match get_projections_from_constant (evd, i) with
| None ->
- let env = Global.env () in
- let t = string_of_ppcmds (pr_constr env evd i) in
- failwith ("Cannot register term " ^ t)
+ let env = Global.env () in
+ let t = string_of_ppcmds (pr_constr env evd i) in
+ failwith ("Cannot register term " ^ t)
| Some a -> E.mk_elt evd i a
- let register_hint evd t elt =
+ let register_hint evd t elt =
match EConstr.kind evd t with
- | App(c,_) ->
- E.table := HConstr.add c (Application t, E.cast elt) !E.table
- | _ -> E.table := HConstr.add t (OtherTerm t, E.cast elt) !E.table
-
-
-
+ | App (c, _) ->
+ E.table := HConstr.add c (Application t, E.cast elt) !E.table
+ | _ -> E.table := HConstr.add t (OtherTerm t, E.cast elt) !E.table
let register_constr env evd c =
let c = EConstr.of_constr c in
let t = get_type_of env evd c in
match EConstr.kind evd t with
| App (intyp, args) ->
- let styp = args.(E.get_key) in
- let elt = {decl= c; deriv= (make_elt (evd, c))} in
- register_hint evd styp elt
+ let styp = args.(E.get_key) in
+ let elt = {decl = c; deriv = make_elt (evd, c)} in
+ register_hint evd styp elt
| _ ->
- let env = Global.env () in
- raise (CErrors.user_err Pp.
- (str ": Cannot register term "++pr_constr env evd c++
- str ". It has type "++pr_constr env evd t++str " which should be of the form [F X1 .. Xn]"))
+ let env = Global.env () in
+ raise
+ (CErrors.user_err
+ Pp.(
+ str ": Cannot register term "
+ ++ pr_constr env evd c ++ str ". It has type "
+ ++ pr_constr env evd t
+ ++ str " which should be of the form [F X1 .. Xn]"))
let register_obj : Constr.constr -> Libobject.obj =
let cache_constr (_, c) =
@@ -447,7 +373,7 @@ module MakeTable (E : Elt) = struct
let subst_constr (subst, c) = Mod_subst.subst_mps subst c in
Libobject.declare_object
@@ Libobject.superglobal_object_nodischarge
- ("register-zify-" ^ E.name)
+ ("register-zify-" ^ E.name)
~cache:cache_constr ~subst:(Some subst_constr)
(** [register c] is called from the VERNACULAR ADD [name] constr(t).
@@ -455,52 +381,40 @@ module MakeTable (E : Elt) = struct
registered as a [superglobal_object_nodischarge].
TODO: pre-compute [get_type_of] - [cache_constr] is using another environment.
*)
- let register = fun c ->
+ let register c =
let env = Global.env () in
let evd = Evd.from_env env in
let evd, c = Constrintern.interp_open_constr env evd c in
let _ = Lib.add_anonymous_leaf (register_obj (EConstr.to_constr evd c)) in
()
-
let pp_keys () =
let env = Global.env () in
let evd = Evd.from_env env in
HConstr.fold
- (fun _ (k,d) acc ->
+ (fun _ (k, d) acc ->
match E.dest d with
| None -> acc
| Some _ ->
- Pp.(pr_constr env evd (constr_of_term_kind k) ++ str " " ++ acc))
- (!E.table) (Pp.str "")
-
-
- let print () = Feedback.msg_info (pp_keys ())
+ Pp.(pr_constr env evd (constr_of_term_kind k) ++ str " " ++ acc))
+ !E.table (Pp.str "")
+ let print () = Feedback.msg_info (pp_keys ())
end
-
module InjTable = MakeTable (EInj)
-
module ESat = struct
type elt = ESatT.t
+
open ESatT
let name = "Saturate"
-
let table = saturate
-
let cast x = Saturate x
-
- let dest = function
- | Saturate x -> Some x
- | _ -> None
-
- let mk_elt evd i a = {parg1= a.(2); parg2= a.(3); satOK= a.(5)}
-
+ let dest = function Saturate x -> Some x | _ -> None
+ let mk_elt evd i a = {parg1 = a.(2); parg2 = a.(3); satOK = a.(5)}
let get_key = 1
-
end
module BinOp = MakeTable (EBinOp)
@@ -512,10 +426,9 @@ module PropUnOp = MakeTable (EPropUnOp)
module Saturate = MakeTable (ESat)
let init_cache () =
- table_cache := !table;
+ table_cache := !table;
saturate_cache := !saturate
-
(** The module [Spec] is used to register
the instances of [BinOpSpec], [UnOpSpec].
They are not indexed and stored in a list. *)
@@ -556,7 +469,6 @@ module Spec = struct
Feedback.msg_notice l
end
-
let unfold_decl evd =
let f cst acc = cst :: acc in
fold_declared_const f evd []
@@ -578,33 +490,19 @@ let locate_const str =
(* The following [constr] are necessary for constructing the proof terms *)
let mkapp2 = lazy (zify "mkapp2")
-
let mkapp = lazy (zify "mkapp")
-
let mkapp0 = lazy (zify "mkapp0")
-
let mkdp = lazy (zify "mkinjterm")
-
let eq_refl = lazy (zify "eq_refl")
-
let mkrel = lazy (zify "mkrel")
-
let mkprop_op = lazy (zify "mkprop_op")
-
let mkuprop_op = lazy (zify "mkuprop_op")
-
let mkdpP = lazy (zify "mkinjprop")
-
let iff_refl = lazy (zify "iff_refl")
-
let q = lazy (zify "target_prop")
-
let ieq = lazy (zify "injprop_ok")
-
let iff = lazy (zify "iff")
-
-
(* A super-set of the previous are needed to unfold the generated proof terms. *)
let to_unfold =
@@ -631,7 +529,6 @@ let to_unfold =
; "mkapp0"
; "mkprop_op" ])
-
(** Module [CstrTable] records terms [x] injected into [inj x]
together with the corresponding type constraint.
The terms are stored by side-effect during the traversal
@@ -644,17 +541,15 @@ module CstrTable = struct
type t = EConstr.t
let hash c = Constr.hash (unsafe_to_constr c)
-
let equal c c' = Constr.equal (unsafe_to_constr c) (unsafe_to_constr c')
end)
let table : EConstr.t HConstr.t = HConstr.create 10
-
let register evd t (i : EConstr.t) = HConstr.add table t i
let get () =
let l = HConstr.fold (fun k i acc -> (k, i) :: acc) table [] in
- HConstr.clear table ; l
+ HConstr.clear table; l
(** [gen_cstr table] asserts (cstr k) for each element of the table (k,cstr).
NB: the constraint is only asserted if it does not already exist in the context.
@@ -667,7 +562,7 @@ module CstrTable = struct
let hyps_table = HConstr.create 20 in
List.iter
(fun (_, (t : EConstr.types)) -> HConstr.add hyps_table t ())
- (Tacmach.New.pf_hyps_types gl) ;
+ (Tacmach.New.pf_hyps_types gl);
fun c -> HConstr.mem hyps_table c
in
(* Add the constraint (cstr k) if it is not already present *)
@@ -683,17 +578,16 @@ module CstrTable = struct
(Names.Id.of_string "cstr")
env
in
- Tactics.pose_proof (Names.Name n) term )
+ Tactics.pose_proof (Names.Name n) term)
in
List.fold_left
(fun acc (k, i) -> Tacticals.New.tclTHEN (gen k i) acc)
- Tacticals.New.tclIDTAC table )
+ Tacticals.New.tclIDTAC table)
end
let mkvar red evd inj v =
( if not red then
- match inj.cstr with None -> () | Some ctr -> CstrTable.register evd v ctr
- ) ;
+ match inj.cstr with None -> () | Some ctr -> CstrTable.register evd v ctr );
let iv = EConstr.mkApp (inj.inj, [|v|]) in
let iv = if red then Tacred.compute (Global.env ()) evd iv else iv in
EConstr.mkApp
@@ -724,11 +618,8 @@ let inj_term_of_texpr evd = function
| Var (inj, e) -> mkvar false evd inj e
| Constant (inj, e) -> mkvar true evd inj e
-let mkapp2_id evd i (* InjTyp S3 T *)
- inj (* deriv i *)
- t (* S1 -> S2 -> S3 *)
- b (* Binop S1 S2 S3 t ... *)
- dbop (* deriv b *) e1 e2 =
+let mkapp2_id evd i (* InjTyp S3 T *) inj (* deriv i *) t (* S1 -> S2 -> S3 *) b
+ (* Binop S1 S2 S3 t ... *) dbop (* deriv b *) e1 e2 =
let default () =
let e1' = inj_term_of_texpr evd e1 in
let e2' = inj_term_of_texpr evd e2 in
@@ -755,15 +646,16 @@ let mkapp2_id evd i (* InjTyp S3 T *)
|Var (_, e1), Var (_, e2)
|Constant (_, e1), Var (_, e2)
|Var (_, e1), Constant (_, e2) ->
- Var (inj, EConstr.mkApp (t, [|e1; e2|]))
+ Var (inj, EConstr.mkApp (t, [|e1; e2|]))
| _, _ -> default ()
let mkapp_id evd i inj (unop, u) f e1 =
- EUnOpT.(if EConstr.eq_constr evd u.unop f then
- (* Injection does nothing *)
- match e1 with
- | Constant (_, e) | Var (_, e) -> Var (inj, EConstr.mkApp (f, [|e|]))
- | Injterm e1 ->
+ EUnOpT.(
+ if EConstr.eq_constr evd u.unop f then
+ (* Injection does nothing *)
+ match e1 with
+ | Constant (_, e) | Var (_, e) -> Var (inj, EConstr.mkApp (f, [|e|]))
+ | Injterm e1 ->
Injterm
(EConstr.mkApp
( force mkapp
@@ -775,124 +667,128 @@ let mkapp_id evd i inj (unop, u) f e1 =
; u.inj2_t
; unop
; e1 |] ))
- else
- let e1 = inj_term_of_texpr evd e1 in
- Injterm
- (EConstr.mkApp
- ( force mkapp
- , [|u.source1; u.source2; u.target; f; u.inj1_t; u.inj2_t; unop; e1|]
- )))
-
-type typed_constr = {constr: EConstr.t; typ: EConstr.t}
-
+ else
+ let e1 = inj_term_of_texpr evd e1 in
+ Injterm
+ (EConstr.mkApp
+ ( force mkapp
+ , [|u.source1; u.source2; u.target; f; u.inj1_t; u.inj2_t; unop; e1|]
+ )))
+type typed_constr = {constr : EConstr.t; typ : EConstr.t}
let get_injection env evd t =
match snd (HConstr.find t !table_cache) with
| InjTyp i -> i
- | _ -> raise Not_found
-
-
- (* [arrow] is the term (fun (x:Prop) (y : Prop) => x -> y) *)
- let arrow =
- let name x =
- Context.make_annot (Name.mk_name (Names.Id.of_string x)) Sorts.Relevant in
- EConstr.mkLambda
- ( name "x"
- , EConstr.mkProp
- , EConstr.mkLambda
- ( name "y"
- , EConstr.mkProp
- , EConstr.mkProd
- ( Context.make_annot Names.Anonymous Sorts.Relevant
- , EConstr.mkRel 2
- , EConstr.mkRel 2 ) ) )
-
-
- let is_prop env sigma term =
- let sort = Retyping.get_sort_of env sigma term in
+ | _ -> raise Not_found
+
+(* [arrow] is the term (fun (x:Prop) (y : Prop) => x -> y) *)
+let arrow =
+ let name x =
+ Context.make_annot (Name.mk_name (Names.Id.of_string x)) Sorts.Relevant
+ in
+ EConstr.mkLambda
+ ( name "x"
+ , EConstr.mkProp
+ , EConstr.mkLambda
+ ( name "y"
+ , EConstr.mkProp
+ , EConstr.mkProd
+ ( Context.make_annot Names.Anonymous Sorts.Relevant
+ , EConstr.mkRel 2
+ , EConstr.mkRel 2 ) ) )
+
+let is_prop env sigma term =
+ let sort = Retyping.get_sort_of env sigma term in
Sorts.is_prop sort
- (** [get_application env evd e] expresses [e] as an application (c a)
+(** [get_application env evd e] expresses [e] as an application (c a)
where c is the head symbol and [a] is the array of arguments.
The function also transforms (x -> y) as (arrow x y) *)
- let get_operator env evd e =
- let is_arrow a p1 p2 =
- is_prop env evd p1 && is_prop (EConstr.push_rel (Context.Rel.Declaration.LocalAssum(a,p1)) env) evd p2
- && (a.Context.binder_name = Names.Anonymous || EConstr.Vars.noccurn evd 1 p2) in
- match EConstr.kind evd e with
- | Prod (a, p1, p2) when is_arrow a p1 p2 ->
- (arrow,[|p1 ;p2|])
- | App(c,a) -> (c,a)
- | _ -> (e,[||])
-
+let get_operator env evd e =
+ let is_arrow a p1 p2 =
+ is_prop env evd p1
+ && is_prop
+ (EConstr.push_rel (Context.Rel.Declaration.LocalAssum (a, p1)) env)
+ evd p2
+ && (a.Context.binder_name = Names.Anonymous || EConstr.Vars.noccurn evd 1 p2)
+ in
+ match EConstr.kind evd e with
+ | Prod (a, p1, p2) when is_arrow a p1 p2 -> (arrow, [|p1; p2|])
+ | App (c, a) -> (c, a)
+ | _ -> (e, [||])
- let is_convertible env evd k t =
- Reductionops.check_conv env evd k t
+let is_convertible env evd k t = Reductionops.check_conv env evd k t
- (** [match_operator env evd hd arg (t,d)]
+(** [match_operator env evd hd arg (t,d)]
- hd is head operator of t
- If t = OtherTerm _, then t = hd
- If t = Application _, then
we extract the relevant number of arguments from arg
and check for convertibility *)
- let match_operator env evd hd args (t, d) =
- let decomp t i =
- let n = Array.length args in
- let t' = EConstr.mkApp(hd,Array.sub args 0 (n-i)) in
- if is_convertible env evd t' t
- then Some (d,t)
- else None in
-
- match t with
- | OtherTerm t -> Some(d,t)
- | Application t ->
- match d with
- | CstOp _ -> decomp t 0
- | UnOp _ -> decomp t 1
- | BinOp _ -> decomp t 2
- | BinRel _ -> decomp t 2
- | PropOp _ -> decomp t 2
- | PropUnOp _ -> decomp t 1
- | _ -> None
-
-
- let rec trans_expr env evd e =
+let match_operator env evd hd args (t, d) =
+ let decomp t i =
+ let n = Array.length args in
+ let t' = EConstr.mkApp (hd, Array.sub args 0 (n - i)) in
+ if is_convertible env evd t' t then Some (d, t) else None
+ in
+ match t with
+ | OtherTerm t -> Some (d, t)
+ | Application t -> (
+ match d with
+ | CstOp _ -> decomp t 0
+ | UnOp _ -> decomp t 1
+ | BinOp _ -> decomp t 2
+ | BinRel _ -> decomp t 2
+ | PropOp _ -> decomp t 2
+ | PropUnOp _ -> decomp t 1
+ | _ -> None )
+
+let rec trans_expr env evd e =
(* Get the injection *)
- let {decl= i; deriv= inj} = get_injection env evd e.typ in
+ let {decl = i; deriv = inj} = get_injection env evd e.typ in
let e = e.constr in
if EConstr.isConstruct evd e then Constant (inj, e) (* Evaluate later *)
else
- let (c,a) = get_operator env evd e in
+ let c, a = get_operator env evd e in
try
- let (k,t) = find_option (match_operator env evd c a) (HConstr.find_all c !table_cache) in
+ let k, t =
+ find_option
+ (match_operator env evd c a)
+ (HConstr.find_all c !table_cache)
+ in
let n = Array.length a in
- match k with
- | CstOp {decl = c'} ->
- Injterm (EConstr.mkApp (force mkapp0, [|inj.source; inj.target; e; i; c'|]))
- | UnOp {decl = unop ; deriv = u} ->
- let a' = trans_expr env evd {constr= a.(n-1); typ= u.EUnOpT.source1} in
- if is_constant a' && EConstr.isConstruct evd t then
- Constant (inj, e)
- else mkapp_id evd i inj (unop, u) t a'
- | BinOp {decl = binop ; deriv = b} ->
- let a0 = trans_expr env evd {constr= a.(n-2); typ= b.EBinOpT.source1} in
- let a1 = trans_expr env evd {constr= a.(n-1); typ= b.EBinOpT.source2} in
- if is_constant a0 && is_constant a1 && EConstr.isConstruct evd t
- then Constant (inj, e)
- else mkapp2_id evd i inj t binop b a0 a1
- | d ->
- Var (inj,e)
- with Not_found -> Var (inj,e)
+ match k with
+ | CstOp {decl = c'} ->
+ Injterm
+ (EConstr.mkApp (force mkapp0, [|inj.source; inj.target; e; i; c'|]))
+ | UnOp {decl = unop; deriv = u} ->
+ let a' =
+ trans_expr env evd {constr = a.(n - 1); typ = u.EUnOpT.source1}
+ in
+ if is_constant a' && EConstr.isConstruct evd t then Constant (inj, e)
+ else mkapp_id evd i inj (unop, u) t a'
+ | BinOp {decl = binop; deriv = b} ->
+ let a0 =
+ trans_expr env evd {constr = a.(n - 2); typ = b.EBinOpT.source1}
+ in
+ let a1 =
+ trans_expr env evd {constr = a.(n - 1); typ = b.EBinOpT.source2}
+ in
+ if is_constant a0 && is_constant a1 && EConstr.isConstruct evd t then
+ Constant (inj, e)
+ else mkapp2_id evd i inj t binop b a0 a1
+ | d -> Var (inj, e)
+ with Not_found -> Var (inj, e)
let trans_expr env evd e =
- try trans_expr env evd e with Not_found ->
+ try trans_expr env evd e
+ with Not_found ->
raise
(CErrors.user_err
( Pp.str "Missing injection for type "
++ Printer.pr_leconstr_env env evd e.typ ))
-
type tprop =
| TProp of EConstr.t (** Transformed proposition *)
| IProp of EConstr.t (** Identical proposition *)
@@ -903,72 +799,72 @@ let mk_iprop e =
let inj_prop_of_tprop = function TProp p -> p | IProp e -> mk_iprop e
let rec trans_prop env evd e =
- let (c,a) = get_operator env evd e in
+ let c, a = get_operator env evd e in
try
- let (k,t) = find_option (match_operator env evd c a) (HConstr.find_all c !table_cache) in
+ let k, t =
+ find_option (match_operator env evd c a) (HConstr.find_all c !table_cache)
+ in
let n = Array.length a in
match k with
- | PropOp {decl= rop} ->
- begin
- try
- let t1 = trans_prop env evd a.(n-2) in
- let t2 = trans_prop env evd a.(n-1) in
- match (t1, t2) with
- | IProp _, IProp _ -> IProp e
- | _, _ ->
- let t1 = inj_prop_of_tprop t1 in
- let t2 = inj_prop_of_tprop t2 in
- TProp (EConstr.mkApp (force mkprop_op, [|t; rop; t1; t2|]))
- with Not_found -> IProp e
- end
- | BinRel {decl = br ; deriv = rop} ->
- begin
- try
- let a1 = trans_expr env evd {constr = a.(n-2) ; typ = rop.EBinRelT.source} in
- let a2 = trans_expr env evd {constr = a.(n-1) ; typ = rop.EBinRelT.source} in
- if EConstr.eq_constr evd t rop.EBinRelT.brel then
- match (constr_of_texpr a1, constr_of_texpr a2) with
- | Some e1, Some e2 -> IProp (EConstr.mkApp (t, [|e1; e2|]))
- | _, _ ->
- let a1 = inj_term_of_texpr evd a1 in
- let a2 = inj_term_of_texpr evd a2 in
- TProp
- (EConstr.mkApp
- ( force mkrel
- , [| rop.EBinRelT.source
- ; rop.EBinRelT.target
- ; t
- ; rop.EBinRelT.inj
- ; br
- ; a1
- ; a2 |] ))
- else
- let a1 = inj_term_of_texpr evd a1 in
- let a2 = inj_term_of_texpr evd a2 in
- TProp
- (EConstr.mkApp
- ( force mkrel
- , [| rop.EBinRelT.source
- ; rop.EBinRelT.target
- ; t
- ; rop.EBinRelT.inj
- ; br
- ; a1
- ; a2 |] ))
- with Not_found -> IProp e
- end
- | PropUnOp {decl = rop} ->
- begin
- try
- let t1 = trans_prop env evd a.(n-1) in
- match t1 with
- | IProp _ -> IProp e
- | _ ->
- let t1 = inj_prop_of_tprop t1 in
- TProp (EConstr.mkApp (force mkuprop_op, [|t; rop; t1|]))
- with Not_found -> IProp e
- end
- | _ -> IProp e
+ | PropOp {decl = rop} -> (
+ try
+ let t1 = trans_prop env evd a.(n - 2) in
+ let t2 = trans_prop env evd a.(n - 1) in
+ match (t1, t2) with
+ | IProp _, IProp _ -> IProp e
+ | _, _ ->
+ let t1 = inj_prop_of_tprop t1 in
+ let t2 = inj_prop_of_tprop t2 in
+ TProp (EConstr.mkApp (force mkprop_op, [|t; rop; t1; t2|]))
+ with Not_found -> IProp e )
+ | BinRel {decl = br; deriv = rop} -> (
+ try
+ let a1 =
+ trans_expr env evd {constr = a.(n - 2); typ = rop.EBinRelT.source}
+ in
+ let a2 =
+ trans_expr env evd {constr = a.(n - 1); typ = rop.EBinRelT.source}
+ in
+ if EConstr.eq_constr evd t rop.EBinRelT.brel then
+ match (constr_of_texpr a1, constr_of_texpr a2) with
+ | Some e1, Some e2 -> IProp (EConstr.mkApp (t, [|e1; e2|]))
+ | _, _ ->
+ let a1 = inj_term_of_texpr evd a1 in
+ let a2 = inj_term_of_texpr evd a2 in
+ TProp
+ (EConstr.mkApp
+ ( force mkrel
+ , [| rop.EBinRelT.source
+ ; rop.EBinRelT.target
+ ; t
+ ; rop.EBinRelT.inj
+ ; br
+ ; a1
+ ; a2 |] ))
+ else
+ let a1 = inj_term_of_texpr evd a1 in
+ let a2 = inj_term_of_texpr evd a2 in
+ TProp
+ (EConstr.mkApp
+ ( force mkrel
+ , [| rop.EBinRelT.source
+ ; rop.EBinRelT.target
+ ; t
+ ; rop.EBinRelT.inj
+ ; br
+ ; a1
+ ; a2 |] ))
+ with Not_found -> IProp e )
+ | PropUnOp {decl = rop} -> (
+ try
+ let t1 = trans_prop env evd a.(n - 1) in
+ match t1 with
+ | IProp _ -> IProp e
+ | _ ->
+ let t1 = inj_prop_of_tprop t1 in
+ TProp (EConstr.mkApp (force mkuprop_op, [|t; rop; t1|]))
+ with Not_found -> IProp e )
+ | _ -> IProp e
with Not_found -> IProp e
let unfold n env evd c =
@@ -984,14 +880,14 @@ let unfold n env evd c =
match n with
| None -> c
| Some n ->
- Tacred.unfoldn [(Locus.AllOccurrences, Names.EvalVarRef n)] env evd c
+ Tacred.unfoldn [(Locus.AllOccurrences, Names.EvalVarRef n)] env evd c
in
(* Reduce the term *)
- let c = cbv (List.rev_append (force to_unfold) unfold_decl) env evd c in
+ let c = cbv (List.rev_append (force to_unfold) unfold_decl) env evd c in
c
let trans_check_prop env evd t =
- if is_prop env evd t then
+ if is_prop env evd t then
(*let t = Tacred.unfoldn [Locus.AllOccurrences, Names.EvalConstRef coq_not] env evd t in*)
match trans_prop env evd t with IProp e -> None | TProp e -> Some e
else None
@@ -1001,7 +897,7 @@ let trans_hyps env evd l =
(fun acc (h, p) ->
match trans_check_prop env evd p with
| None -> acc
- | Some p' -> (h, p, p') :: acc )
+ | Some p' -> (h, p, p') :: acc)
[] (List.rev l)
(* Only used if a direct rewrite fails *)
@@ -1016,7 +912,7 @@ let trans_hyp h t =
let h' = fresh_id_in_env Id.Set.empty h env in
tclTHENLIST
[ letin_tac None (Names.Name n) t None
- Locus.{onhyps= None; concl_occs= NoOccurrences}
+ Locus.{onhyps = None; concl_occs = NoOccurrences}
; assert_by (Name.Name h')
(EConstr.mkApp (force q, [|EConstr.mkVar n|]))
(tclTHEN
@@ -1027,19 +923,19 @@ let trans_hyp h t =
(h', Locus.InHyp)
; clear [n]
; (* [clear H] may fail if [h] has dependencies *)
- tclTRY (clear [h]) ] )))
+ tclTRY (clear [h]) ])))
let is_progress_rewrite evd t rew =
match EConstr.kind evd rew with
| App (c, [|lhs; rhs|]) ->
- if EConstr.eq_constr evd (force iff) c then
- (* This is a successful rewriting *)
- not (EConstr.eq_constr evd lhs rhs)
- else
- CErrors.anomaly
- Pp.(
- str "is_progress_rewrite: not a rewrite"
- ++ pr_constr (Global.env ()) evd rew)
+ if EConstr.eq_constr evd (force iff) c then
+ (* This is a successful rewriting *)
+ not (EConstr.eq_constr evd lhs rhs)
+ else
+ CErrors.anomaly
+ Pp.(
+ str "is_progress_rewrite: not a rewrite"
+ ++ pr_constr (Global.env ()) evd rew)
| _ -> failwith "is_progress_rewrite: not even an application"
let trans_hyp h t0 t =
@@ -1050,10 +946,10 @@ let trans_hyp h t0 t =
let t' = unfold None env evd (EConstr.mkApp (force ieq, [|t|])) in
if is_progress_rewrite evd t0 (get_type_of env evd t') then
tclFIRST
- [ Equality.general_rewrite_in true Locus.AllOccurrences true false
- h t' false
+ [ Equality.general_rewrite_in true Locus.AllOccurrences true false h
+ t' false
; trans_hyp h t ]
- else tclIDTAC ))
+ else tclIDTAC))
let trans_concl t =
Tacticals.New.(
@@ -1064,15 +960,15 @@ let trans_concl t =
let t' = unfold None env evd (EConstr.mkApp (force ieq, [|t|])) in
if is_progress_rewrite evd concl (get_type_of env evd t') then
Equality.general_rewrite true Locus.AllOccurrences true false t'
- else tclIDTAC ))
+ else tclIDTAC))
let tclTHENOpt e tac tac' =
match e with None -> tac' | Some e' -> Tacticals.New.tclTHEN (tac e') tac'
let zify_tac =
Proofview.Goal.enter (fun gl ->
- Coqlib.check_required_library ["Coq"; "micromega"; "ZifyClasses"] ;
- Coqlib.check_required_library ["Coq"; "micromega"; "ZifyInst"] ;
+ Coqlib.check_required_library ["Coq"; "micromega"; "ZifyClasses"];
+ Coqlib.check_required_library ["Coq"; "micromega"; "ZifyInst"];
init_cache ();
let evd = Tacmach.New.project gl in
let env = Tacmach.New.pf_env gl in
@@ -1083,15 +979,16 @@ let zify_tac =
(Tacticals.New.tclTHEN
(Tacticals.New.tclTHENLIST
(List.rev_map (fun (h, p, t) -> trans_hyp h p t) hyps))
- (CstrTable.gen_cstr l)) )
+ (CstrTable.gen_cstr l)))
let iter_specs tac =
Tacticals.New.tclTHENLIST
- (List.fold_left (fun acc d -> tac d :: acc) [] (Spec.get ()))
+ (List.fold_left (fun acc d -> tac d :: acc) [] (Spec.get ()))
-
-let iter_specs (tac: Ltac_plugin.Tacinterp.Value.t) =
- iter_specs (fun c -> Ltac_plugin.Tacinterp.Value.apply tac [Ltac_plugin.Tacinterp.Value.of_constr c])
+let iter_specs (tac : Ltac_plugin.Tacinterp.Value.t) =
+ iter_specs (fun c ->
+ Ltac_plugin.Tacinterp.Value.apply tac
+ [Ltac_plugin.Tacinterp.Value.of_constr c])
let find_hyp evd t l =
try Some (fst (List.find (fun (h, t') -> EConstr.eq_constr evd t t') l))
@@ -1104,39 +1001,37 @@ let sat_constr c d =
let hyps = Tacmach.New.pf_hyps_types gl in
match EConstr.kind evd c with
| App (c, args) ->
- if Array.length args = 2 then (
- let h1 =
- Tacred.cbv_beta env evd
- (EConstr.mkApp (d.ESatT.parg1, [|args.(0)|]))
+ if Array.length args = 2 then
+ let h1 =
+ Tacred.cbv_beta env evd
+ (EConstr.mkApp (d.ESatT.parg1, [|args.(0)|]))
+ in
+ let h2 =
+ Tacred.cbv_beta env evd
+ (EConstr.mkApp (d.ESatT.parg2, [|args.(1)|]))
+ in
+ match (find_hyp evd h1 hyps, find_hyp evd h2 hyps) with
+ | Some h1, Some h2 ->
+ let n =
+ Tactics.fresh_id_in_env Id.Set.empty
+ (Names.Id.of_string "__sat")
+ env
in
- let h2 =
- Tacred.cbv_beta env evd
- (EConstr.mkApp (d.ESatT.parg2, [|args.(1)|]))
+ let trm =
+ EConstr.mkApp
+ ( d.ESatT.satOK
+ , [|args.(0); args.(1); EConstr.mkVar h1; EConstr.mkVar h2|] )
in
- match (find_hyp evd h1 hyps, find_hyp evd h2 hyps) with
- | Some h1, Some h2 ->
- let n =
- Tactics.fresh_id_in_env Id.Set.empty
- (Names.Id.of_string "__sat")
- env
- in
- let trm =
- EConstr.mkApp
- ( d.ESatT.satOK
- , [|args.(0); args.(1); EConstr.mkVar h1; EConstr.mkVar h2|]
- )
- in
- Tactics.pose_proof (Names.Name n) trm
- | _, _ -> Tacticals.New.tclIDTAC )
- else Tacticals.New.tclIDTAC
- | _ -> Tacticals.New.tclIDTAC )
-
+ Tactics.pose_proof (Names.Name n) trm
+ | _, _ -> Tacticals.New.tclIDTAC
+ else Tacticals.New.tclIDTAC
+ | _ -> Tacticals.New.tclIDTAC)
let get_all_sat env evd c =
- List.fold_left (fun acc e ->
- match e with
- | (_,Saturate s) -> s::acc
- | _ -> acc) [] (HConstr.find_all c !saturate_cache )
+ List.fold_left
+ (fun acc e -> match e with _, Saturate s -> s :: acc | _ -> acc)
+ []
+ (HConstr.find_all c !saturate_cache)
let saturate =
Proofview.Goal.enter (fun gl ->
@@ -1149,21 +1044,19 @@ let saturate =
let rec sat t =
match EConstr.kind evd t with
| App (c, args) ->
- sat c ;
- Array.iter sat args ;
- if Array.length args = 2 then
- let ds = get_all_sat env evd c in
- if ds = [] then ()
- else (
- List.iter (fun x -> CstrTable.HConstr.add table t x.deriv) ds )
- else ()
+ sat c;
+ Array.iter sat args;
+ if Array.length args = 2 then
+ let ds = get_all_sat env evd c in
+ if ds = [] then ()
+ else List.iter (fun x -> CstrTable.HConstr.add table t x.deriv) ds
+ else ()
| Prod (a, t1, t2) when a.Context.binder_name = Names.Anonymous ->
- sat t1 ; sat t2
+ sat t1; sat t2
| _ -> ()
in
(* Collect all the potential saturation lemma *)
- sat concl ;
- List.iter (fun (_, t) -> sat t) hyps ;
+ sat concl;
+ List.iter (fun (_, t) -> sat t) hyps;
Tacticals.New.tclTHENLIST
- (CstrTable.HConstr.fold (fun c d acc -> sat_constr c d :: acc) table [])
- )
+ (CstrTable.HConstr.fold (fun c d acc -> sat_constr c d :: acc) table []))
diff --git a/plugins/micromega/zify.mli b/plugins/micromega/zify.mli
index 54e8f07ddc..9e3cf5d24c 100644
--- a/plugins/micromega/zify.mli
+++ b/plugins/micromega/zify.mli
@@ -9,16 +9,19 @@
(************************************************************************)
open Constrexpr
-module type S = sig val register : constr_expr -> unit val print : unit -> unit end
+module type S = sig
+ val register : constr_expr -> unit
+ val print : unit -> unit
+end
module InjTable : S
-module UnOp : S
-module BinOp : S
-module CstOp : S
-module BinRel : S
-module PropOp : S
+module UnOp : S
+module BinOp : S
+module CstOp : S
+module BinRel : S
+module PropOp : S
module PropUnOp : S
-module Spec : S
+module Spec : S
module Saturate : S
val zify_tac : unit Proofview.tactic