aboutsummaryrefslogtreecommitdiffhomepage
path: root/plugins
diff options
context:
space:
mode:
Diffstat (limited to 'plugins')
-rw-r--r--plugins/btauto/refl_btauto.ml13
-rw-r--r--plugins/cc/ccalgo.ml2
-rw-r--r--plugins/cc/cctac.ml75
-rw-r--r--plugins/cc/g_congruence.ml44
-rw-r--r--plugins/decl_mode/decl_expr.mli2
-rw-r--r--plugins/decl_mode/decl_interp.ml6
-rw-r--r--plugins/decl_mode/decl_mode.ml2
-rw-r--r--plugins/decl_mode/decl_proof_instr.ml42
-rw-r--r--plugins/decl_mode/g_decl_mode.ml48
-rw-r--r--plugins/derive/derive.ml5
-rw-r--r--plugins/derive/g_derive.ml44
-rw-r--r--plugins/extraction/extraction.ml10
-rw-r--r--plugins/extraction/g_extraction.ml45
-rw-r--r--plugins/extraction/table.ml2
-rw-r--r--plugins/firstorder/formula.ml11
-rw-r--r--plugins/firstorder/formula.mli3
-rw-r--r--plugins/firstorder/g_ground.ml421
-rw-r--r--plugins/firstorder/instances.ml8
-rw-r--r--plugins/firstorder/rules.ml10
-rw-r--r--plugins/fourier/fourierR.ml20
-rw-r--r--plugins/funind/functional_principles_proofs.ml90
-rw-r--r--plugins/funind/functional_principles_types.ml52
-rw-r--r--plugins/funind/g_indfun.ml414
-rw-r--r--plugins/funind/glob_term_to_relation.ml54
-rw-r--r--plugins/funind/indfun.ml23
-rw-r--r--plugins/funind/invfun.ml92
-rw-r--r--plugins/funind/merge.ml43
-rw-r--r--plugins/funind/recdef.ml60
-rw-r--r--plugins/micromega/Psatz.v10
-rw-r--r--plugins/micromega/certificate.ml1563
-rw-r--r--plugins/micromega/coq_micromega.ml419
-rw-r--r--plugins/micromega/g_micromega.ml444
-rw-r--r--plugins/micromega/mfourier.ml24
-rw-r--r--plugins/micromega/mutils.ml9
-rw-r--r--plugins/nsatz/g_nsatz.ml417
-rw-r--r--plugins/nsatz/nsatz.ml (renamed from plugins/nsatz/nsatz.ml4)8
-rw-r--r--plugins/nsatz/nsatz_plugin.mllib1
-rw-r--r--plugins/omega/coq_omega.ml112
-rw-r--r--plugins/omega/g_omega.ml417
-rw-r--r--plugins/quote/g_quote.ml417
-rw-r--r--plugins/quote/quote.ml11
-rw-r--r--plugins/romega/ReflOmegaCore.v2
-rw-r--r--plugins/romega/g_romega.ml417
-rw-r--r--plugins/romega/refl_omega.ml2
-rw-r--r--plugins/rtauto/refl_tauto.ml5
-rw-r--r--plugins/rtauto/refl_tauto.mli2
-rw-r--r--plugins/setoid_ring/ArithRing.v4
-rw-r--r--plugins/setoid_ring/InitialRing.v48
-rw-r--r--plugins/setoid_ring/NArithRing.v2
-rw-r--r--plugins/setoid_ring/Ring.v6
-rw-r--r--plugins/setoid_ring/ZArithRing.v8
-rw-r--r--plugins/setoid_ring/g_newring.ml4105
-rw-r--r--plugins/setoid_ring/newring.ml (renamed from plugins/setoid_ring/newring.ml4)255
-rw-r--r--plugins/setoid_ring/newring.mli78
-rw-r--r--plugins/setoid_ring/newring_ast.mli63
-rw-r--r--plugins/setoid_ring/newring_plugin.mllib1
56 files changed, 1981 insertions, 1550 deletions
diff --git a/plugins/btauto/refl_btauto.ml b/plugins/btauto/refl_btauto.ml
index 57268a9cf..aee0bd856 100644
--- a/plugins/btauto/refl_btauto.ml
+++ b/plugins/btauto/refl_btauto.ml
@@ -1,3 +1,4 @@
+open Proofview.Notations
let contrib_name = "btauto"
@@ -11,7 +12,7 @@ let get_constant dir s = lazy (Coqlib.gen_constant contrib_name dir s)
let get_inductive dir s =
let glob_ref () = Coqlib.find_reference contrib_name ("Coq" :: dir) s in
- Lazy.lazy_from_fun (fun () -> Globnames.destIndRef (glob_ref ()))
+ Lazy.from_fun (fun () -> Globnames.destIndRef (glob_ref ()))
let decomp_term (c : Term.constr) =
Term.kind_of_term (Term.strip_outer_cast c)
@@ -216,7 +217,7 @@ module Btauto = struct
Tacticals.tclFAIL 0 msg gl
let try_unification env =
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
let concl = Proofview.Goal.concl gl in
let eq = Lazy.force eq in
let t = decomp_term concl in
@@ -228,10 +229,10 @@ module Btauto = struct
| _ ->
let msg = str "Btauto: Internal error" in
Tacticals.New.tclFAIL 0 msg
- end
+ end }
let tac =
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
let concl = Proofview.Goal.concl gl in
let eq = Lazy.force eq in
let bool = Lazy.force Bool.typ in
@@ -249,12 +250,12 @@ module Btauto = struct
Tacticals.New.tclTHENLIST [
Tactics.change_concl changed_gl;
Tactics.apply (Lazy.force soundness);
- Proofview.V82.tactic (Tactics.normalise_vm_in_concl);
+ Tactics.normalise_vm_in_concl;
try_unification env
]
| _ ->
let msg = str "Cannot recognize a boolean equality" in
Tacticals.New.tclFAIL 0 msg
- end
+ end }
end
diff --git a/plugins/cc/ccalgo.ml b/plugins/cc/ccalgo.ml
index 5d16edfc6..c01214377 100644
--- a/plugins/cc/ccalgo.ml
+++ b/plugins/cc/ccalgo.ml
@@ -824,7 +824,7 @@ let __eps__ = Id.of_string "_eps_"
let new_state_var typ state =
let id = pf_get_new_id __eps__ state.gls in
let {it=gl ; sigma=sigma} = state.gls in
- let gls = Goal.V82.new_goal_with sigma gl [id,None,typ] in
+ let gls = Goal.V82.new_goal_with sigma gl [Context.Named.Declaration.LocalAssum (id,typ)] in
state.gls<- gls;
id
diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml
index df4a7319a..c8924073c 100644
--- a/plugins/cc/cctac.ml
+++ b/plugins/cc/cctac.ml
@@ -22,6 +22,8 @@ open Ccproof
open Pp
open Errors
open Util
+open Proofview.Notations
+open Context.Rel.Declaration
let reference dir s = lazy (Coqlib.gen_reference "CC" dir s)
@@ -46,7 +48,7 @@ let whd_delta env=
(* decompose member of equality in an applicative format *)
(** FIXME: evar leak *)
-let sf_of env sigma c = sort_of env (ref sigma) c
+let sf_of env sigma c = e_sort_of env (ref sigma) c
let rec decompose_term env sigma t=
match kind_of_term (whd env t) with
@@ -151,7 +153,7 @@ let rec quantified_atom_of_constr env sigma nrels term =
let patts=patterns_of_constr env sigma nrels atom in
`Nrule patts
else
- quantified_atom_of_constr (Environ.push_rel (id,None,atom) env) sigma (succ nrels) ff
+ quantified_atom_of_constr (Environ.push_rel (LocalAssum (id,atom)) env) sigma (succ nrels) ff
| _ ->
let patts=patterns_of_constr env sigma nrels term in
`Rule patts
@@ -166,7 +168,7 @@ let litteral_of_constr env sigma term=
else
begin
try
- quantified_atom_of_constr (Environ.push_rel (id,None,atom) env) sigma 1 ff
+ quantified_atom_of_constr (Environ.push_rel (LocalAssum (id,atom)) env) sigma 1 ff
with Not_found ->
`Other (decompose_term env sigma term)
end
@@ -187,7 +189,8 @@ let make_prb gls depth additionnal_terms =
let t = decompose_term env sigma c in
ignore (add_term state t)) additionnal_terms;
List.iter
- (fun (id,_,e) ->
+ (fun decl ->
+ let (id,_,e) = Context.Named.Declaration.to_tuple decl in
begin
let cid=mkVar id in
match litteral_of_constr env sigma e with
@@ -220,24 +223,9 @@ let make_prb gls depth additionnal_terms =
(* indhyps builds the array of arrays of constructor hyps for (ind largs) *)
-let build_projection intype outtype (cstr:pconstructor) special default gls=
- let env=pf_env gls in
- let (h,argv) = try destApp intype with DestKO -> (intype,[||]) in
- let ind,u=destInd h in
- let types=Inductiveops.arities_of_constructors env (ind,u) in
- let lp=Array.length types in
- let ci=pred (snd(fst cstr)) in
- let branch i=
- let ti= prod_appvect types.(i) argv in
- let rc=fst (decompose_prod_assum ti) in
- let head=
- if Int.equal i ci then special else default in
- it_mkLambda_or_LetIn head rc in
- let branches=Array.init lp branch in
- let casee=mkRel 1 in
- let pred=mkLambda(Anonymous,intype,outtype) in
- let case_info=make_case_info (pf_env gls) ind RegularStyle in
- let body= mkCase(case_info, pred, casee, branches) in
+let build_projection intype (cstr:pconstructor) special default gls=
+ let ci= (snd(fst cstr)) in
+ let body=Equality.build_selector (pf_env gls) (project gls) ci (mkRel 1) intype special default in
let id=pf_get_new_id (Id.of_string "t") gls in
mkLambda(Name id,intype,body)
@@ -254,13 +242,13 @@ let new_app_global f args k =
let new_refine c = Proofview.V82.tactic (refine c)
let assert_before n c =
- Proofview.Goal.enter begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let evm, _ = Tacmach.New.pf_apply type_of gl c in
Tacticals.New.tclTHEN (Proofview.V82.tactic (Refiner.tclEVARS evm)) (assert_before n c)
- end
+ end }
let rec proof_tac p : unit Proofview.tactic =
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
let type_of t = Tacmach.New.pf_unsafe_type_of gl t in
try (* type_of can raise exceptions *)
match p.p_rule with
@@ -319,16 +307,16 @@ let rec proof_tac p : unit Proofview.tactic =
let outtype = (* Termops.refresh_universes *) (type_of default) in
let special=mkRel (1+nargs-argind) in
let proj =
- Tacmach.New.of_old (build_projection intype outtype cstr special default) gl
+ Tacmach.New.of_old (build_projection intype cstr special default) gl
in
let injt=
app_global _f_equal [|intype;outtype;proj;ti;tj;_M 1|] in
Tacticals.New.tclTHEN (Proofview.V82.tactic (injt refine)) (proof_tac prf)
with e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e
- end
+ end }
let refute_tac c t1 t2 p =
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
let tt1=constr_of_term t1 and tt2=constr_of_term t2 in
let intype =
Tacmach.New.of_old (fun gls -> (* Termops.refresh_universes *) (pf_unsafe_type_of gls tt1)) gl
@@ -338,14 +326,14 @@ let refute_tac c t1 t2 p =
let false_t=mkApp (c,[|mkVar hid|]) in
Tacticals.New.tclTHENS (neweq (assert_before (Name hid)))
[proof_tac p; simplest_elim false_t]
- end
+ end }
let refine_exact_check c gl =
let evm, _ = pf_apply type_of gl c in
Tacticals.tclTHEN (Refiner.tclEVARS evm) (Proofview.V82.of_tactic (exact_check c)) gl
let convert_to_goal_tac c t1 t2 p =
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
let tt1=constr_of_term t1 and tt2=constr_of_term t2 in
let sort =
Tacmach.New.of_old (fun gls -> (* Termops.refresh_universes *) (pf_unsafe_type_of gls tt2)) gl
@@ -357,20 +345,20 @@ let convert_to_goal_tac c t1 t2 p =
let endt=app_global _eq_rect [|sort;tt1;identity;c;tt2;mkVar e|] in
Tacticals.New.tclTHENS (neweq (assert_before (Name e)))
[proof_tac p; Proofview.V82.tactic (endt refine_exact_check)]
- end
+ end }
let convert_to_hyp_tac c1 t1 c2 t2 p =
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
let tt2=constr_of_term t2 in
let h = Tacmach.New.of_old (pf_get_new_id (Id.of_string "H")) gl in
let false_t=mkApp (c2,[|mkVar h|]) in
Tacticals.New.tclTHENS (assert_before (Name h) tt2)
[convert_to_goal_tac c1 t1 t2 p;
simplest_elim false_t]
- end
+ end }
let discriminate_tac (cstr,u as cstru) p =
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
let t1=constr_of_term p.p_lhs and t2=constr_of_term p.p_rhs in
let intype =
Tacmach.New.of_old (fun gls -> (* Termops.refresh_universes *) (pf_unsafe_type_of gls t1)) gl
@@ -384,11 +372,11 @@ let discriminate_tac (cstr,u as cstru) p =
let identity = Universes.constr_of_global (Lazy.force _I) in
(* let trivial=pf_unsafe_type_of gls identity in *)
let trivial = Universes.constr_of_global (Lazy.force _True) in
- let evm, outtype = Evd.new_sort_variable Evd.univ_flexible (Proofview.Goal.sigma gl) in
+ let evm, outtype = Evd.new_sort_variable Evd.univ_flexible (Tacmach.New.project gl) in
let outtype = mkSort outtype in
let pred=mkLambda(Name xid,outtype,mkRel 1) in
let hid = Tacmach.New.of_old (pf_get_new_id (Id.of_string "Heq")) gl in
- let proj = Tacmach.New.of_old (build_projection intype outtype cstru trivial concl) gl in
+ let proj = Tacmach.New.of_old (build_projection intype cstru trivial concl) gl in
let injt=app_global _f_equal
[|intype;outtype;proj;t1;t2;mkVar hid|] in
let endt k =
@@ -399,7 +387,7 @@ let discriminate_tac (cstr,u as cstru) p =
Tacticals.New.tclTHEN (Proofview.Unsafe.tclEVARS evm)
(Tacticals.New.tclTHENS (neweq (assert_before (Name hid)))
[proof_tac p; Proofview.V82.tactic (endt refine_exact_check)])
- end
+ end }
(* wrap everything *)
@@ -411,7 +399,7 @@ let build_term_to_complete uf meta pac =
applistc (mkConstructU cinfo.ci_constr) all_args
let cc_tactic depth additionnal_terms =
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
Coqlib.check_required_library Coqlib.logic_module_name;
let _ = debug (fun () -> Pp.str "Reading subgoal ...") in
let state = Tacmach.New.of_old (fun gls -> make_prb gls depth additionnal_terms) gl in
@@ -462,7 +450,7 @@ let cc_tactic depth additionnal_terms =
convert_to_goal_tac id ta tb p
| HeqnH (ida,idb) ->
convert_to_hyp_tac ida ta idb tb p
- end
+ end }
let cc_fail gls =
errorlabstrm "Congruence" (Pp.str "congruence failed.")
@@ -485,8 +473,7 @@ let congruence_tac depth l =
let mk_eq f c1 c2 k =
Tacticals.New.pf_constr_of_global (Lazy.force f) (fun fc ->
- Proofview.Goal.enter begin
- fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let open Tacmach.New in
let evm, ty = pf_apply type_of gl c1 in
let evm, ty = Evarsolve.refresh_universes (Some false) (pf_env gl) evm ty in
@@ -494,10 +481,10 @@ let mk_eq f c1 c2 k =
let evm, _ = type_of (pf_env gl) evm term in
Tacticals.New.tclTHEN (Proofview.V82.tactic (Refiner.tclEVARS evm))
(k term)
- end)
+ end })
let f_equal =
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
let concl = Proofview.Goal.concl gl in
let cut_eq c1 c2 =
try (* type_of can raise an exception *)
@@ -523,4 +510,4 @@ let f_equal =
| Type_errors.TypeError _ -> Proofview.tclUNIT ()
| e -> Proofview.tclZERO ~info e
end
- end
+ end }
diff --git a/plugins/cc/g_congruence.ml4 b/plugins/cc/g_congruence.ml4
index 5dbc340ca..9a53e2e16 100644
--- a/plugins/cc/g_congruence.ml4
+++ b/plugins/cc/g_congruence.ml4
@@ -9,6 +9,10 @@
(*i camlp4deps: "grammar/grammar.cma" i*)
open Cctac
+open Stdarg
+open Constrarg
+open Pcoq.Prim
+open Pcoq.Constr
DECLARE PLUGIN "cc_plugin"
diff --git a/plugins/decl_mode/decl_expr.mli b/plugins/decl_mode/decl_expr.mli
index 79ef3d186..9d78a51ef 100644
--- a/plugins/decl_mode/decl_expr.mli
+++ b/plugins/decl_mode/decl_expr.mli
@@ -99,4 +99,4 @@ type proof_instr =
(Term.constr statement,
Term.constr,
proof_pattern,
- Tacexpr.glob_tactic_expr) gen_proof_instr
+ Genarg.Val.t) gen_proof_instr
diff --git a/plugins/decl_mode/decl_interp.ml b/plugins/decl_mode/decl_interp.ml
index 2a44dca21..34307a358 100644
--- a/plugins/decl_mode/decl_interp.ml
+++ b/plugins/decl_mode/decl_interp.ml
@@ -96,7 +96,7 @@ let rec add_vars_of_simple_pattern globs = function
add_vars_of_simple_pattern globs p
| CPatCstr (_,_,pl1,pl2) ->
List.fold_left add_vars_of_simple_pattern
- (List.fold_left add_vars_of_simple_pattern globs pl1) pl2
+ (Option.fold_left (List.fold_left add_vars_of_simple_pattern) globs pl1) pl2
| CPatNotation(_,_,(pl,pll),pl') ->
List.fold_left add_vars_of_simple_pattern globs (List.flatten (pl::pl'::pll))
| CPatAtom (_,Some (Libnames.Ident (_,id))) -> add_var id globs
@@ -384,7 +384,7 @@ let interp_cases info env sigma params (pat:cases_pattern_expr) hyps =
let interp_cut interp_it env sigma cut=
let nenv,nstat = interp_it env sigma cut.cut_stat in
- {cut with
+ { cut_using=Option.map (Tacinterp.Value.of_closure (Tacinterp.default_ist ())) cut.cut_using;
cut_stat=nstat;
cut_by=interp_justification_items nenv sigma cut.cut_by}
@@ -403,7 +403,7 @@ let interp_suffices_clause env sigma (hyps,cot)=
match hyp with
(Hprop st | Hvar st) ->
match st.st_label with
- Name id -> Environ.push_named (id,None,st.st_it) env0
+ Name id -> Environ.push_named (Context.Named.Declaration.LocalAssum (id,st.st_it)) env0
| _ -> env in
let nenv = List.fold_right push_one locvars env in
nenv,res
diff --git a/plugins/decl_mode/decl_mode.ml b/plugins/decl_mode/decl_mode.ml
index acee3d6c2..f9399d682 100644
--- a/plugins/decl_mode/decl_mode.ml
+++ b/plugins/decl_mode/decl_mode.ml
@@ -116,7 +116,7 @@ let get_top_stack pts =
let get_stack pts = Proof.get_at_focus proof_focus pts
let get_last env = match Environ.named_context env with
- | (id,_,_)::_ -> id
+ | decl :: _ -> Context.Named.Declaration.get_id decl
| [] -> error "no previous statement to use"
diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml
index ba9fb728c..090b293f5 100644
--- a/plugins/decl_mode/decl_proof_instr.ml
+++ b/plugins/decl_mode/decl_proof_instr.ml
@@ -29,6 +29,8 @@ open Termops
open Namegen
open Goptions
open Misctypes
+open Sigma.Notations
+open Context.Named.Declaration
(* Strictness option *)
@@ -86,7 +88,7 @@ Please \"suppose\" something or \"end\" it now."
| _ -> ()
let mk_evd metalist gls =
- let evd0= create_goal_evar_defs (sig_sig gls) in
+ let evd0= clear_metas (sig_sig gls) in
let add_one (meta,typ) evd =
meta_declare meta typ evd in
List.fold_right add_one metalist evd0
@@ -228,7 +230,8 @@ let close_previous_case pts =
(* automation *)
let filter_hyps f gls =
- let filter_aux (id,_,_) =
+ let filter_aux id =
+ let id = get_id id in
if f id then
tclIDTAC
else
@@ -330,11 +333,12 @@ let enstack_subsubgoals env se stack gls=
let rc,_ = Reduction.dest_prod env apptype in
let rec meta_aux last lenv = function
[] -> (last,lenv,[])
- | (nam,_,typ)::q ->
+ | decl::q ->
let nlast=succ last in
let (llast,holes,metas) =
meta_aux nlast (mkMeta nlast :: lenv) q in
- (llast,holes,(nlast,special_nf gls (substl lenv typ))::metas) in
+ let open Context.Rel.Declaration in
+ (llast,holes,(nlast,special_nf gls (substl lenv (get_type decl)))::metas) in
let (nlast,holes,nmetas) =
meta_aux se.se_last_meta [] (List.rev rc) in
let refiner = applist (appterm,List.rev holes) in
@@ -403,15 +407,15 @@ let concl_refiner metas body gls =
let concl = pf_concl gls in
let evd = sig_sig gls in
let env = pf_env gls in
- let sort = family_of_sort (Typing.sort_of env (ref evd) concl) in
+ let sort = family_of_sort (Typing.e_sort_of env (ref evd) concl) in
let rec aux env avoid subst = function
[] -> anomaly ~label:"concl_refiner" (Pp.str "cannot happen")
| (n,typ)::rest ->
let _A = subst_meta subst typ in
let x = id_of_name_using_hdchar env _A Anonymous in
let _x = fresh_id avoid x gls in
- let nenv = Environ.push_named (_x,None,_A) env in
- let asort = family_of_sort (Typing.sort_of nenv (ref evd) _A) in
+ let nenv = Environ.push_named (LocalAssum (_x,_A)) env in
+ let asort = family_of_sort (Typing.e_sort_of nenv (ref evd) _A) in
let nsubst = (n,mkVar _x)::subst in
if List.is_empty rest then
asort,_A,mkNamedLambda _x _A (subst_meta nsubst body)
@@ -492,7 +496,7 @@ let just_tac _then cut info gls0 =
None ->
Proofview.V82.of_tactic automation_tac gls
| Some tac ->
- Proofview.V82.of_tactic (Tacinterp.eval_tactic tac) gls in
+ Proofview.V82.of_tactic (Tacinterp.tactic_of_value (Tacinterp.default_ist ()) tac) gls in
justification (tclTHEN items_tac method_tac) gls0
let instr_cut mkstat _thus _then cut gls0 =
@@ -542,7 +546,7 @@ let instr_rew _thus rew_side cut gls0 =
None ->
Proofview.V82.of_tactic automation_tac gls
| Some tac ->
- Proofview.V82.of_tactic (Tacinterp.eval_tactic tac) gls in
+ Proofview.V82.of_tactic (Tacinterp.tactic_of_value (Tacinterp.default_ist ()) tac) gls in
let just_tac gls =
justification (tclTHEN items_tac method_tac) gls in
let (c_id,_) = match cut.cut_stat.st_label with
@@ -605,7 +609,7 @@ let assume_tac hyps gls =
tclTHEN
(push_intro_tac
(fun id ->
- Proofview.V82.of_tactic (convert_hyp (id,None,st.st_it))) st.st_label))
+ Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,st.st_it)))) st.st_label))
hyps tclIDTAC gls
let assume_hyps_or_theses hyps gls =
@@ -615,7 +619,7 @@ let assume_hyps_or_theses hyps gls =
tclTHEN
(push_intro_tac
(fun id ->
- Proofview.V82.of_tactic (convert_hyp (id,None,c))) nam)
+ Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,c)))) nam)
| Hprop {st_label=nam;st_it=Thesis (tk)} ->
tclTHEN
(push_intro_tac
@@ -627,7 +631,7 @@ let assume_st hyps gls =
(fun st ->
tclTHEN
(push_intro_tac
- (fun id -> Proofview.V82.of_tactic (convert_hyp (id,None,st.st_it))) st.st_label))
+ (fun id -> Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,st.st_it)))) st.st_label))
hyps tclIDTAC gls
let assume_st_letin hyps gls =
@@ -636,7 +640,7 @@ let assume_st_letin hyps gls =
tclTHEN
(push_intro_tac
(fun id ->
- Proofview.V82.of_tactic (convert_hyp (id,Some (fst st.st_it),snd st.st_it))) st.st_label))
+ Proofview.V82.of_tactic (convert_hyp (LocalDef (id, fst st.st_it, snd st.st_it)))) st.st_label))
hyps tclIDTAC gls
(* suffices *)
@@ -730,7 +734,7 @@ let rec consider_match may_intro introduced available expected gls =
error "Not enough sub-hypotheses to match statements."
(* should tell which ones *)
| id::rest_ids,(Hvar st | Hprop st)::rest ->
- tclIFTHENELSE (Proofview.V82.of_tactic (convert_hyp (id,None,st.st_it)))
+ tclIFTHENELSE (Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,st.st_it))))
begin
match st.st_label with
Anonymous ->
@@ -798,8 +802,8 @@ let define_tac id args body gls =
let cast_tac id_or_thesis typ gls =
match id_or_thesis with
This id ->
- let (_,body,_) = pf_get_hyp gls id in
- Proofview.V82.of_tactic (convert_hyp (id,body,typ)) gls
+ let body = pf_get_hyp gls id |> get_value in
+ Proofview.V82.of_tactic (convert_hyp (of_tuple (id,body,typ))) gls
| Thesis (For _ ) ->
error "\"thesis for ...\" is not applicable here."
| Thesis Plain ->
@@ -1305,7 +1309,11 @@ let understand_my_constr env sigma c concl =
Pretyping.understand_tcc env sigma ~expected_type:(Pretyping.OfType concl) (frob rawc)
let my_refine c gls =
- let oc sigma = understand_my_constr (pf_env gls) sigma c (pf_concl gls) in
+ let oc = { run = begin fun sigma ->
+ let sigma = Sigma.to_evar_map sigma in
+ let (sigma, c) = understand_my_constr (pf_env gls) sigma c (pf_concl gls) in
+ Sigma.Unsafe.of_pair (c, sigma)
+ end } in
Proofview.V82.of_tactic (Tactics.New.refine oc) gls
(* end focus/claim *)
diff --git a/plugins/decl_mode/g_decl_mode.ml4 b/plugins/decl_mode/g_decl_mode.ml4
index 4c5c65669..802c36109 100644
--- a/plugins/decl_mode/g_decl_mode.ml4
+++ b/plugins/decl_mode/g_decl_mode.ml4
@@ -84,7 +84,7 @@ let vernac_proof_instr instr =
(* Only declared at raw level, because only used in vernac commands. *)
let wit_proof_instr : (raw_proof_instr, glob_proof_instr, proof_instr) Genarg.genarg_type =
- Genarg.make0 None "proof_instr"
+ Genarg.make0 "proof_instr"
(* We create a new parser entry [proof_mode]. The Declarative proof mode
will replace the normal parser entry for tactics with this one. *)
@@ -92,7 +92,7 @@ let proof_mode : vernac_expr Gram.entry =
Gram.entry_create "vernac:proof_command"
(* Auxiliary grammar entry. *)
let proof_instr : raw_proof_instr Gram.entry =
- Pcoq.create_generic_entry "proof_instr" (Genarg.rawwit wit_proof_instr)
+ Pcoq.create_generic_entry Pcoq.utactic "proof_instr" (Genarg.rawwit wit_proof_instr)
let _ = Pptactic.declare_extra_genarg_pprule wit_proof_instr
pr_raw_proof_instr pr_glob_proof_instr pr_proof_instr
@@ -132,7 +132,7 @@ let _ =
set = begin fun () ->
(* We set the command non terminal to
[proof_mode] (which we just defined). *)
- G_vernac.set_command_entry proof_mode ;
+ Pcoq.set_command_entry proof_mode ;
(* We substitute the goal printer, by the one we built
for the proof mode. *)
Printer.set_printer_pr { Printer.default_printer_pr with
@@ -144,7 +144,7 @@ let _ =
reset = begin fun () ->
(* We restore the command non terminal to
[noedit_mode]. *)
- G_vernac.set_command_entry G_vernac.noedit_mode ;
+ Pcoq.set_command_entry Pcoq.Vernac_.noedit_mode ;
(* We restore the goal printer to default *)
Printer.set_printer_pr Printer.default_printer_pr
end
diff --git a/plugins/derive/derive.ml b/plugins/derive/derive.ml
index ce93c5a3f..5d1551106 100644
--- a/plugins/derive/derive.ml
+++ b/plugins/derive/derive.ml
@@ -6,6 +6,8 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Context.Named.Declaration
+
let map_const_entry_body (f:Term.constr->Term.constr) (x:Safe_typing.private_constants Entries.const_entry_body)
: Safe_typing.private_constants Entries.const_entry_body =
Future.chain ~pure:true x begin fun ((b,ctx),fx) ->
@@ -32,7 +34,7 @@ let start_deriving f suchthat lemma =
let open Proofview in
TCons ( env , sigma , f_type_type , (fun sigma f_type ->
TCons ( env , sigma , f_type , (fun sigma ef ->
- let env' = Environ.push_named (f , (Some ef) , f_type) env in
+ let env' = Environ.push_named (LocalDef (f, ef, f_type)) env in
let evdref = ref sigma in
let suchthat = Constrintern.interp_type_evars env' evdref suchthat in
TCons ( env' , !evdref , suchthat , (fun sigma _ ->
@@ -93,6 +95,7 @@ let start_deriving f suchthat lemma =
ignore (Declare.declare_constant lemma lemma_def)
in
+ let terminator = Proof_global.make_terminator terminator in
let () = Proof_global.start_dependent_proof lemma kind goals terminator in
let _ = Proof_global.with_current_proof begin fun _ p ->
Proof.run_tactic env Proofview.(tclFOCUS 1 2 shelve) p
diff --git a/plugins/derive/g_derive.ml4 b/plugins/derive/g_derive.ml4
index 18570a684..35a5a7616 100644
--- a/plugins/derive/g_derive.ml4
+++ b/plugins/derive/g_derive.ml4
@@ -6,6 +6,10 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Constrarg
+open Pcoq.Prim
+open Pcoq.Constr
+
(*i camlp4deps: "grammar/grammar.cma" i*)
let classify_derive_command _ = Vernacexpr.(VtStartProof ("Classic",Doesn'tGuaranteeOpacity,[]),VtLater)
diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml
index 667721e67..098f76bbf 100644
--- a/plugins/extraction/extraction.ml
+++ b/plugins/extraction/extraction.ml
@@ -11,7 +11,6 @@ open Util
open Names
open Term
open Vars
-open Context
open Declarations
open Declareops
open Environ
@@ -26,6 +25,7 @@ open Globnames
open Miniml
open Table
open Mlutil
+open Context.Rel.Declaration
(*i*)
exception I of inductive_kind
@@ -75,7 +75,7 @@ type flag = info * scheme
let rec flag_of_type env t : flag =
let t = whd_betadeltaiota env none t in
match kind_of_term t with
- | Prod (x,t,c) -> flag_of_type (push_rel (x,None,t) env) c
+ | Prod (x,t,c) -> flag_of_type (push_rel (LocalAssum (x,t)) env) c
| Sort s when Sorts.is_prop s -> (Logic,TypeScheme)
| Sort _ -> (Info,TypeScheme)
| _ -> if (sort_of env t) == InProp then (Logic,Default) else (Info,Default)
@@ -248,7 +248,7 @@ let rec extract_type env db j c args =
| _ when sort_of env (applist (c, args)) == InProp -> Tdummy Kprop
| Rel n ->
(match lookup_rel n env with
- | (_,Some t,_) -> extract_type env db j (lift n t) args
+ | LocalDef (_,t,_) -> extract_type env db j (lift n t) args
| _ ->
(* Asks [db] a translation for [n]. *)
if n > List.length db then Tunknown
@@ -561,7 +561,7 @@ let rec extract_term env mle mlt c args =
put_magic_if magic (MLlam (id, d')))
| LetIn (n, c1, t1, c2) ->
let id = id_of_name n in
- let env' = push_rel (Name id, Some c1, t1) env in
+ let env' = push_rel (LocalDef (Name id, c1, t1)) env in
(* We directly push the args inside the [LetIn].
TODO: the opt_let_app flag is supposed to prevent that *)
let args' = List.map (lift 1) args in
@@ -836,7 +836,7 @@ and extract_fix env mle i (fi,ti,ci as recd) mlt =
let decomp_lams_eta_n n m env c t =
let rels = fst (splay_prod_n env none n t) in
- let rels = List.map (fun (id,_,c) -> (id,c)) rels in
+ let rels = List.map (fun (LocalAssum (id,c) | LocalDef (id,_,c)) -> (id,c)) rels in
let rels',c = decompose_lam c in
let d = n - m in
(* we'd better keep rels' as long as possible. *)
diff --git a/plugins/extraction/g_extraction.ml4 b/plugins/extraction/g_extraction.ml4
index aec958689..ca4e13e12 100644
--- a/plugins/extraction/g_extraction.ml4
+++ b/plugins/extraction/g_extraction.ml4
@@ -11,6 +11,10 @@
(* ML names *)
open Genarg
+open Stdarg
+open Constrarg
+open Pcoq.Prim
+open Pcoq.Constr
open Pp
open Names
open Nameops
@@ -31,7 +35,6 @@ let pr_int_or_id _ _ _ = function
| ArgId id -> pr_id id
ARGUMENT EXTEND int_or_id
- TYPED AS int_or_id
PRINTED BY pr_int_or_id
| [ preident(id) ] -> [ ArgId (Id.of_string id) ]
| [ integer(i) ] -> [ ArgInt i ]
diff --git a/plugins/extraction/table.ml b/plugins/extraction/table.ml
index d7842e127..466c8054b 100644
--- a/plugins/extraction/table.ml
+++ b/plugins/extraction/table.ml
@@ -453,7 +453,7 @@ let check_loaded_modfile mp = match base_mp mp with
if not (Library.library_is_loaded dp) then begin
match base_mp (Lib.current_mp ()) with
| MPfile dp' when not (DirPath.equal dp dp') ->
- err (str ("Please load library "^(DirPath.to_string dp^" first.")))
+ err (str "Please load library " ++ pr_dirpath dp ++ str " first.")
| _ -> ()
end
| _ -> ()
diff --git a/plugins/firstorder/formula.ml b/plugins/firstorder/formula.ml
index ae2d059fa..2ed436c6b 100644
--- a/plugins/firstorder/formula.ml
+++ b/plugins/firstorder/formula.ml
@@ -15,6 +15,7 @@ open Tacmach
open Util
open Declarations
open Globnames
+open Context.Rel.Declaration
let qflag=ref true
@@ -139,8 +140,8 @@ let build_atoms gl metagen side cciterm =
negative:= unsigned :: !negative
end;
let v = ind_hyps 0 i l gl in
- let g i _ (_,_,t) =
- build_rec env polarity (lift i t) in
+ let g i _ decl =
+ build_rec env polarity (lift i (get_type decl)) in
let f l =
List.fold_left_i g (1-(List.length l)) () l in
if polarity && (* we have a constant constructor *)
@@ -150,8 +151,8 @@ let build_atoms gl metagen side cciterm =
| Exists(i,l)->
let var=mkMeta (metagen true) in
let v =(ind_hyps 1 i l gl).(0) in
- let g i _ (_,_,t) =
- build_rec (var::env) polarity (lift i t) in
+ let g i _ decl =
+ build_rec (var::env) polarity (lift i (get_type decl)) in
List.fold_left_i g (2-(List.length l)) () v
| Forall(_,b)->
let var=mkMeta (metagen true) in
@@ -224,7 +225,7 @@ let build_formula side nam typ gl metagen=
| And(_,_,_) -> Rand
| Or(_,_,_) -> Ror
| Exists (i,l) ->
- let (_,_,d)=List.last (ind_hyps 0 i l gl).(0) in
+ let d = get_type (List.last (ind_hyps 0 i l gl).(0)) in
Rexists(m,d,trivial)
| Forall (_,a) -> Rforall
| Arrow (a,b) -> Rarrow in
diff --git a/plugins/firstorder/formula.mli b/plugins/firstorder/formula.mli
index 39d99d2e0..0f70d3ea0 100644
--- a/plugins/firstorder/formula.mli
+++ b/plugins/firstorder/formula.mli
@@ -7,7 +7,6 @@
(************************************************************************)
open Term
-open Context
open Globnames
val qflag : bool ref
@@ -27,7 +26,7 @@ type counter = bool -> metavariable
val construct_nhyps : pinductive -> Proof_type.goal Tacmach.sigma -> int array
val ind_hyps : int -> pinductive -> constr list ->
- Proof_type.goal Tacmach.sigma -> rel_context array
+ Proof_type.goal Tacmach.sigma -> Context.Rel.t array
type atoms = {positive:constr list;negative:constr list}
diff --git a/plugins/firstorder/g_ground.ml4 b/plugins/firstorder/g_ground.ml4
index 041526881..587d10d1c 100644
--- a/plugins/firstorder/g_ground.ml4
+++ b/plugins/firstorder/g_ground.ml4
@@ -15,6 +15,10 @@ open Goptions
open Tacticals
open Tacinterp
open Libnames
+open Constrarg
+open Stdarg
+open Pcoq.Prim
+open Pcoq.Tactic
DECLARE PLUGIN "ground_plugin"
@@ -52,8 +56,15 @@ let _=
in
declare_int_option gdopt
+let default_intuition_tac =
+ let tac _ _ = Auto.h_auto None [] None in
+ let name = { Tacexpr.mltac_plugin = "ground_plugin"; mltac_tactic = "auto_with"; } in
+ let entry = { Tacexpr.mltac_name = name; mltac_index = 0 } in
+ Tacenv.register_ml_tactic name [| tac |];
+ Tacexpr.TacML (Loc.ghost, entry, [])
+
let (set_default_solver, default_solver, print_default_solver) =
- Tactic_option.declare_tactic_option ~default:(<:tactic<auto with *>>) "Firstorder default solver"
+ Tactic_option.declare_tactic_option ~default:default_intuition_tac "Firstorder default solver"
VERNAC COMMAND EXTEND Firstorder_Set_Solver CLASSIFIED AS SIDEFF
| [ "Set" "Firstorder" "Solver" tactic(t) ] -> [
@@ -128,17 +139,17 @@ END
TACTIC EXTEND firstorder
[ "firstorder" tactic_opt(t) firstorder_using(l) ] ->
- [ Proofview.V82.tactic (gen_ground_tac true (Option.map eval_tactic t) l []) ]
+ [ Proofview.V82.tactic (gen_ground_tac true (Option.map (tactic_of_value ist) t) l []) ]
| [ "firstorder" tactic_opt(t) "with" ne_preident_list(l) ] ->
- [ Proofview.V82.tactic (gen_ground_tac true (Option.map eval_tactic t) [] l) ]
+ [ Proofview.V82.tactic (gen_ground_tac true (Option.map (tactic_of_value ist) t) [] l) ]
| [ "firstorder" tactic_opt(t) firstorder_using(l)
"with" ne_preident_list(l') ] ->
- [ Proofview.V82.tactic (gen_ground_tac true (Option.map eval_tactic t) l l') ]
+ [ Proofview.V82.tactic (gen_ground_tac true (Option.map (tactic_of_value ist) t) l l') ]
END
TACTIC EXTEND gintuition
[ "gintuition" tactic_opt(t) ] ->
- [ Proofview.V82.tactic (gen_ground_tac false (Option.map eval_tactic t) [] []) ]
+ [ Proofview.V82.tactic (gen_ground_tac false (Option.map (tactic_of_value ist) t) [] []) ]
END
open Proofview.Notations
diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml
index a717cc91e..0bc40136c 100644
--- a/plugins/firstorder/instances.ml
+++ b/plugins/firstorder/instances.ml
@@ -22,6 +22,8 @@ open Formula
open Sequent
open Names
open Misctypes
+open Sigma.Notations
+open Context.Rel.Declaration
let compare_instance inst1 inst2=
match inst1,inst2 with
@@ -116,8 +118,10 @@ let mk_open_instance id idc gl m t=
let rec aux n avoid env evmap decls =
if Int.equal n 0 then evmap, decls else
let nid=(fresh_id avoid var_id gl) in
- let evmap, (c, _) = Evarutil.new_type_evar env evmap Evd.univ_flexible in
- let decl = (Name nid,None,c) in
+ let evmap = Sigma.Unsafe.of_evar_map evmap in
+ let Sigma ((c, _), evmap, _) = Evarutil.new_type_evar env evmap Evd.univ_flexible in
+ let evmap = Sigma.to_evar_map evmap in
+ let decl = LocalAssum (Name nid, c) in
aux (n-1) (nid::avoid) (Environ.push_rel decl env) evmap (decl::decls) in
let evmap, decls = aux m [] env evmap [] in
evmap, decls, revt
diff --git a/plugins/firstorder/rules.ml b/plugins/firstorder/rules.ml
index e676a8a93..c05015c53 100644
--- a/plugins/firstorder/rules.ml
+++ b/plugins/firstorder/rules.ml
@@ -19,6 +19,7 @@ open Formula
open Sequent
open Globnames
open Locus
+open Context.Named.Declaration
type seqtac= (Sequent.t -> tactic) -> Sequent.t -> tactic
@@ -34,12 +35,13 @@ let wrap n b continue seq gls=
if i<=0 then seq else
match nc with
[]->anomaly (Pp.str "Not the expected number of hyps")
- | ((id,_,typ) as nd)::q->
+ | nd::q->
+ let id = get_id nd in
if occur_var env id (pf_concl gls) ||
List.exists (occur_var_in_decl env id) ctx then
(aux (i-1) q (nd::ctx))
else
- add_formula Hyp (VarRef id) typ (aux (i-1) q (nd::ctx)) gls in
+ add_formula Hyp (VarRef id) (get_type nd) (aux (i-1) q (nd::ctx)) gls in
let seq1=aux n nc [] in
let seq2=if b then
add_formula Concl dummy_id (pf_concl gls) seq1 gls else seq1 in
@@ -210,6 +212,6 @@ let defined_connectives=lazy
let normalize_evaluables=
onAllHypsAndConcl
(function
- None->unfold_in_concl (Lazy.force defined_connectives)
+ None-> Proofview.V82.of_tactic (unfold_in_concl (Lazy.force defined_connectives))
| Some id ->
- unfold_in_hyp (Lazy.force defined_connectives) (id,InHypTypeOnly))
+ Proofview.V82.of_tactic (unfold_in_hyp (Lazy.force defined_connectives) (id,InHypTypeOnly)))
diff --git a/plugins/fourier/fourierR.ml b/plugins/fourier/fourierR.ml
index 72e9371be..8bc84608e 100644
--- a/plugins/fourier/fourierR.ml
+++ b/plugins/fourier/fourierR.ml
@@ -19,6 +19,7 @@ open Globnames
open Tacmach
open Fourier
open Contradiction
+open Proofview.Notations
(******************************************************************************
Opérations sur les combinaisons linéaires affines.
@@ -412,13 +413,6 @@ let tac_zero_infeq_false gl (n,d) =
(tac_zero_inf_pos gl (-n,d)))
;;
-let create_meta () = mkMeta(Evarutil.new_meta());;
-
-let my_cut c gl=
- let concl = pf_concl gl in
- apply_type (mkProd(Anonymous,c,concl)) [create_meta()] gl
-;;
-
let exact = exact_check;;
let tac_use h =
@@ -451,7 +445,11 @@ let is_ineq (h,t) =
;;
*)
-let list_of_sign s = List.map (fun (x,_,z)->(x,z)) s;;
+let list_of_sign s =
+ let open Context.Named.Declaration in
+ List.map (function LocalAssum (name, typ) -> name, typ
+ | LocalDef (name, _, typ) -> name, typ)
+ s;;
let mkAppL a =
let l = Array.to_list a in
@@ -462,7 +460,7 @@ exception GoalDone
(* Résolution d'inéquations linéaires dans R *)
let rec fourier () =
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
let concl = Proofview.Goal.concl gl in
Coqlib.check_required_library ["Coq";"fourier";"Fourier"];
let goal = strip_outer_cast concl in
@@ -586,7 +584,7 @@ let rec fourier () =
then tac_zero_inf_false gl (rational_to_fraction cres)
else tac_zero_infeq_false gl (rational_to_fraction cres)
in
- tac:=(Tacticals.New.tclTHENS (Proofview.V82.tactic (my_cut ineq))
+ tac:=(Tacticals.New.tclTHENS (cut ineq)
[Tacticals.New.tclTHEN (change_concl
(mkAppL [| get coq_not; ineq|]
))
@@ -622,7 +620,7 @@ let rec fourier () =
(* ((tclTHEN !tac (tclFAIL 1 (* 1 au hasard... *))) gl) *)
!tac
(* ((tclABSTRACT None !tac) gl) *)
- end
+ end }
;;
(*
diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml
index 169a70600..02cd819f4 100644
--- a/plugins/funind/functional_principles_proofs.ml
+++ b/plugins/funind/functional_principles_proofs.ml
@@ -3,18 +3,19 @@ open Errors
open Util
open Term
open Vars
-open Context
open Namegen
open Names
open Declarations
open Pp
open Tacmach
+open Termops
open Proof_type
open Tacticals
open Tactics
open Indfun_common
open Libnames
open Globnames
+open Context.Rel.Declaration
(* let msgnl = Pp.msgnl *)
@@ -52,10 +53,10 @@ let rec print_debug_queue e =
let _ =
match e with
| Some e ->
- Pp.msg_debug (lmsg ++ (str " raised exception " ++ Errors.print e) ++ str " on goal " ++ goal)
+ Pp.msg_debug (hov 0 (lmsg ++ (str " raised exception " ++ Errors.print e) ++ str " on goal" ++ fnl() ++ goal))
| None ->
begin
- Pp.msg_debug (str " from " ++ lmsg ++ str " on goal " ++ goal);
+ Pp.msg_debug (str " from " ++ lmsg ++ str " on goal" ++ fnl() ++ goal);
end in
print_debug_queue None ;
end
@@ -229,7 +230,7 @@ let nf_betaiotazeta = (* Reductionops.local_strong Reductionops.whd_betaiotazeta
-let change_eq env sigma hyp_id (context:rel_context) x t end_of_type =
+let change_eq env sigma hyp_id (context:Context.Rel.t) x t end_of_type =
let nochange ?t' msg =
begin
observe (str ("Not treating ( "^msg^" )") ++ pr_lconstr t ++ str " " ++ match t' with None -> str "" | Some t -> Printer.pr_lconstr t );
@@ -304,11 +305,11 @@ let change_eq env sigma hyp_id (context:rel_context) x t end_of_type =
in
let new_type_of_hyp,ctxt_size,witness_fun =
List.fold_left_i
- (fun i (end_of_type,ctxt_size,witness_fun) ((x',b',t') as decl) ->
+ (fun i (end_of_type,ctxt_size,witness_fun) decl ->
try
let witness = Int.Map.find i sub in
- if not (Option.is_empty b') then anomaly (Pp.str "can not redefine a rel!");
- (Termops.pop end_of_type,ctxt_size,mkLetIn(x',witness,t',witness_fun))
+ if is_local_def decl then anomaly (Pp.str "can not redefine a rel!");
+ (Termops.pop end_of_type,ctxt_size,mkLetIn (get_name decl, witness, get_type decl, witness_fun))
with Not_found ->
(mkProd_or_LetIn decl end_of_type, ctxt_size + 1, mkLambda_or_LetIn decl witness_fun)
)
@@ -371,12 +372,12 @@ let isLetIn t =
| _ -> false
-let h_reduce_with_zeta =
- reduce
+let h_reduce_with_zeta cl =
+ Proofview.V82.of_tactic (reduce
(Genredexpr.Cbv
{Redops.all_flags
with Genredexpr.rDelta = false;
- })
+ }) cl)
@@ -536,7 +537,7 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma =
(scan_type new_context new_t')
with Failure "NoChange" ->
(* Last thing todo : push the rel in the context and continue *)
- scan_type ((x,None,t_x)::context) t'
+ scan_type (LocalAssum (x,t_x) :: context) t'
end
end
else
@@ -705,9 +706,9 @@ let build_proof
in
tclTHENSEQ
[
- Simple.generalize (term_eq::(List.map mkVar dyn_infos.rec_hyps));
+ generalize (term_eq::(List.map mkVar dyn_infos.rec_hyps));
thin dyn_infos.rec_hyps;
- pattern_option [Locus.AllOccurrencesBut [1],t] None;
+ Proofview.V82.of_tactic (pattern_option [Locus.AllOccurrencesBut [1],t] None);
(fun g -> observe_tac "toto" (
tclTHENSEQ [Proofview.V82.of_tactic (Simple.case t);
(fun g' ->
@@ -736,7 +737,8 @@ let build_proof
tclTHEN
(Proofview.V82.of_tactic intro)
(fun g' ->
- let (id,_,_) = pf_last_hyp g' in
+ let open Context.Named.Declaration in
+ let id = pf_last_hyp g' |> get_id in
let new_term =
pf_nf_betaiota g'
(mkApp(dyn_infos.info,[|mkVar id|]))
@@ -921,7 +923,9 @@ let generalize_non_dep hyp g =
let env = Global.env () in
let hyp_typ = pf_unsafe_type_of g (mkVar hyp) in
let to_revert,_ =
- Environ.fold_named_context_reverse (fun (clear,keep) (hyp,_,_ as decl) ->
+ let open Context.Named.Declaration in
+ Environ.fold_named_context_reverse (fun (clear,keep) decl ->
+ let hyp = get_id decl in
if Id.List.mem hyp hyps
|| List.exists (Termops.occur_var_in_decl env hyp) keep
|| Termops.occur_var env hyp hyp_typ
@@ -932,11 +936,11 @@ let generalize_non_dep hyp g =
in
(* observe (str "to_revert := " ++ prlist_with_sep spc Ppconstr.pr_id to_revert); *)
tclTHEN
- ((* observe_tac "h_generalize" *) (Simple.generalize (List.map mkVar to_revert) ))
+ ((* observe_tac "h_generalize" *) (generalize (List.map mkVar to_revert) ))
((* observe_tac "thin" *) (thin to_revert))
g
-let id_of_decl (na,_,_) = (Nameops.out_name na)
+let id_of_decl decl = Nameops.out_name (get_name decl)
let var_of_decl decl = mkVar (id_of_decl decl)
let revert idl =
tclTHEN
@@ -1044,7 +1048,8 @@ let do_replace (evd:Evd.evar_map ref) params rec_arg_num rev_args_id f fun_num a
(
fun g' ->
let just_introduced = nLastDecls nb_intro_to_do g' in
- let just_introduced_id = List.map (fun (id,_,_) -> id) just_introduced in
+ let open Context.Named.Declaration in
+ let just_introduced_id = List.map get_id just_introduced in
tclTHEN (Proofview.V82.of_tactic (Equality.rewriteLR equation_lemma))
(revert just_introduced_id) g'
)
@@ -1069,11 +1074,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
(Name new_id)
)
in
- let fresh_decl =
- (fun (na,b,t) ->
- (fresh_id na,b,t)
- )
- in
+ let fresh_decl = map_name fresh_id in
let princ_info : elim_scheme =
{ princ_info with
params = List.map fresh_decl princ_info.params;
@@ -1120,11 +1121,11 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
)
in
observe (str "full_params := " ++
- prlist_with_sep spc (fun (na,_,_) -> Ppconstr.pr_id (Nameops.out_name na))
+ prlist_with_sep spc (fun decl -> Ppconstr.pr_id (Nameops.out_name (get_name decl)))
full_params
);
observe (str "princ_params := " ++
- prlist_with_sep spc (fun (na,_,_) -> Ppconstr.pr_id (Nameops.out_name na))
+ prlist_with_sep spc (fun decl -> Ppconstr.pr_id (Nameops.out_name (get_name decl)))
princ_params
);
observe (str "fbody_with_full_params := " ++
@@ -1165,7 +1166,8 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
in
let pte_to_fix,rev_info =
List.fold_left_i
- (fun i (acc_map,acc_info) (pte,_,_) ->
+ (fun i (acc_map,acc_info) decl ->
+ let pte = get_name decl in
let infos = info_array.(i) in
let type_args,_ = decompose_prod infos.types in
let nargs = List.length type_args in
@@ -1259,7 +1261,8 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
let args = nLastDecls nb_args g in
let fix_body = fix_info.body_with_param in
(* observe (str "fix_body := "++ pr_lconstr_env (pf_env gl) fix_body); *)
- let args_id = List.map (fun (id,_,_) -> id) args in
+ let open Context.Named.Declaration in
+ let args_id = List.map get_id args in
let dyn_infos =
{
nb_rec_hyps = -100;
@@ -1276,7 +1279,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
(do_replace evd
full_params
(fix_info.idx + List.length princ_params)
- (args_id@(List.map (fun (id,_,_) -> Nameops.out_name id ) princ_params))
+ (args_id@(List.map (fun decl -> Nameops.out_name (get_name decl)) princ_params))
(all_funs.(fix_info.num_in_block))
fix_info.num_in_block
all_funs
@@ -1317,8 +1320,9 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
[
tclDO nb_args (Proofview.V82.of_tactic intro);
(fun g -> (* replacement of the function by its body *)
- let args = nLastDecls nb_args g in
- let args_id = List.map (fun (id,_,_) -> id) args in
+ let args = nLastDecls nb_args g in
+ let open Context.Named.Declaration in
+ let args_id = List.map get_id args in
let dyn_infos =
{
nb_rec_hyps = -100;
@@ -1334,7 +1338,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
in
let fname = destConst (fst (decompose_app (List.hd (List.rev pte_args)))) in
tclTHENSEQ
- [unfold_in_concl [(Locus.AllOccurrences, Names.EvalConstRef (fst fname))];
+ [Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalConstRef (fst fname))]);
let do_prove =
build_proof
interactive_proof
@@ -1403,7 +1407,7 @@ let prove_with_tcc tcc_lemma_constr eqs : tactic =
(* let ids = List.filter (fun id -> not (List.mem id ids)) ids' in *)
(* rewrite *)
(* ) *)
- Eauto.gen_eauto (false,5) [] (Some [])
+ Proofview.V82.of_tactic (Eauto.gen_eauto (false,5) [] (Some []))
]
gls
@@ -1460,7 +1464,7 @@ let new_prove_with_tcc is_mes acc_inv hrec tcc_hyps eqs : tactic =
(fun g ->
if is_mes
then
- unfold_in_concl [(Locus.AllOccurrences, evaluable_of_global_reference (delayed_force ltof_ref))] g
+ Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, evaluable_of_global_reference (delayed_force ltof_ref))]) g
else tclIDTAC g
);
observe_tac "rew_and_finish"
@@ -1472,7 +1476,7 @@ let new_prove_with_tcc is_mes acc_inv hrec tcc_hyps eqs : tactic =
tclCOMPLETE(
Eauto.eauto_with_bases
(true,5)
- [Evd.empty,Lazy.force refl_equal]
+ [{ Tacexpr.delayed = fun _ sigma -> Sigma.here (Lazy.force refl_equal) sigma}]
[Hints.Hint_db.empty empty_transparent_state false]
)
)
@@ -1520,7 +1524,7 @@ let prove_principle_for_gen
avoid := new_id :: !avoid;
Name new_id
in
- let fresh_decl (na,b,t) = (fresh_id na,b,t) in
+ let fresh_decl = map_name fresh_id in
let princ_info : elim_scheme =
{ princ_info with
params = List.map fresh_decl princ_info.params;
@@ -1550,11 +1554,11 @@ let prove_principle_for_gen
in
let rec_arg_id =
match List.rev post_rec_arg with
- | (Name id,_,_)::_ -> id
+ | (LocalAssum (Name id,_) | LocalDef (Name id,_,_)) :: _ -> id
| _ -> assert false
in
(* observe (str "rec_arg_id := " ++ pr_lconstr (mkVar rec_arg_id)); *)
- let subst_constrs = List.map (fun (na,_,_) -> mkVar (Nameops.out_name na)) (pre_rec_arg@princ_info.params) in
+ let subst_constrs = List.map (fun decl -> mkVar (Nameops.out_name (get_name decl))) (pre_rec_arg@princ_info.params) in
let relation = substl subst_constrs relation in
let input_type = substl subst_constrs rec_arg_type in
let wf_thm_id = Nameops.out_name (fresh_id (Name (Id.of_string "wf_R"))) in
@@ -1562,7 +1566,7 @@ let prove_principle_for_gen
Nameops.out_name (fresh_id (Name (Id.of_string ("Acc_"^(Id.to_string rec_arg_id)))))
in
let revert l =
- tclTHEN (Tactics.Simple.generalize (List.map mkVar l)) (clear l)
+ tclTHEN (Tactics.generalize (List.map mkVar l)) (clear l)
in
let fix_id = Nameops.out_name (fresh_id (Name hrec_id)) in
let prove_rec_arg_acc g =
@@ -1582,7 +1586,7 @@ let prove_principle_for_gen
)
g
in
- let args_ids = List.map (fun (na,_,_) -> Nameops.out_name na) princ_info.args in
+ let args_ids = List.map (fun decl -> Nameops.out_name (get_name decl)) princ_info.args in
let lemma =
match !tcc_lemma_ref with
| None -> error "No tcc proof !!"
@@ -1629,7 +1633,7 @@ let prove_principle_for_gen
[
observe_tac "start_tac" start_tac;
h_intros
- (List.rev_map (fun (na,_,_) -> Nameops.out_name na)
+ (List.rev_map (fun decl -> Nameops.out_name (get_name decl))
(princ_info.args@princ_info.branches@princ_info.predicates@princ_info.params)
);
(* observe_tac "" *) Proofview.V82.of_tactic (assert_by
@@ -1667,7 +1671,7 @@ let prove_principle_for_gen
in
let acc_inv = lazy (mkApp(Lazy.force acc_inv, [|mkVar acc_rec_arg_id|])) in
let predicates_names =
- List.map (fun (na,_,_) -> Nameops.out_name na) princ_info.predicates
+ List.map (fun decl -> Nameops.out_name (get_name decl)) princ_info.predicates
in
let pte_info =
{ proving_tac =
@@ -1683,7 +1687,7 @@ let prove_principle_for_gen
is_mes acc_inv fix_id
(!tcc_list@(List.map
- (fun (na,_,_) -> (Nameops.out_name na))
+ (fun decl -> (Nameops.out_name (get_name decl)))
(princ_info.args@princ_info.params)
)@ ([acc_rec_arg_id])) eqs
)
@@ -1712,7 +1716,7 @@ let prove_principle_for_gen
(* observe_tac "instanciate_hyps_with_args" *)
(instanciate_hyps_with_args
make_proof
- (List.map (fun (na,_,_) -> Nameops.out_name na) princ_info.branches)
+ (List.map (fun decl -> Nameops.out_name (get_name decl)) princ_info.branches)
(List.rev args_ids)
)
gl'
diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml
index 18200307a..91a826c73 100644
--- a/plugins/funind/functional_principles_types.ml
+++ b/plugins/funind/functional_principles_types.ml
@@ -3,15 +3,16 @@ open Errors
open Util
open Term
open Vars
-open Context
open Namegen
open Names
open Pp
open Entries
open Tactics
+open Context.Rel.Declaration
open Indfun_common
open Functional_principles_proofs
open Misctypes
+open Sigma.Notations
exception Toberemoved_with_rel of int*constr
exception Toberemoved
@@ -29,14 +30,16 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
let env = Global.env () in
let env_with_params = Environ.push_rel_context princ_type_info.params env in
let tbl = Hashtbl.create 792 in
- let rec change_predicates_names (avoid:Id.t list) (predicates:rel_context) : rel_context =
+ let rec change_predicates_names (avoid:Id.t list) (predicates:Context.Rel.t) : Context.Rel.t =
match predicates with
| [] -> []
- |(Name x,v,t)::predicates ->
- let id = Namegen.next_ident_away x avoid in
- Hashtbl.add tbl id x;
- (Name id,v,t)::(change_predicates_names (id::avoid) predicates)
- | (Anonymous,_,_)::_ -> anomaly (Pp.str "Anonymous property binder ")
+ | decl :: predicates ->
+ (match Context.Rel.Declaration.get_name decl with
+ | Name x ->
+ let id = Namegen.next_ident_away x avoid in
+ Hashtbl.add tbl id x;
+ set_name (Name id) decl :: change_predicates_names (id::avoid) predicates
+ | Anonymous -> anomaly (Pp.str "Anonymous property binder "))
in
let avoid = (Termops.ids_of_context env_with_params ) in
let princ_type_info =
@@ -46,15 +49,16 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
in
(* observe (str "starting princ_type := " ++ pr_lconstr_env env princ_type); *)
(* observe (str "princ_infos : " ++ pr_elim_scheme princ_type_info); *)
- let change_predicate_sort i (x,_,t) =
+ let change_predicate_sort i decl =
let new_sort = sorts.(i) in
- let args,_ = decompose_prod t in
+ let args,_ = decompose_prod (get_type decl) in
let real_args =
if princ_type_info.indarg_in_concl
then List.tl args
else args
in
- Nameops.out_name x,None,compose_prod real_args (mkSort new_sort)
+ Context.Named.Declaration.LocalAssum (Nameops.out_name (Context.Rel.Declaration.get_name decl),
+ compose_prod real_args (mkSort new_sort))
in
let new_predicates =
List.map_i
@@ -69,7 +73,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
| _ -> error "Not a valid predicate"
)
in
- let ptes_vars = List.map (fun (id,_,_) -> id) new_predicates in
+ let ptes_vars = List.map Context.Named.Declaration.get_id new_predicates in
let is_pte =
let set = List.fold_right Id.Set.add ptes_vars Id.Set.empty in
fun t ->
@@ -114,7 +118,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
| Rel n ->
begin
try match Environ.lookup_rel n env with
- | _,_,t when is_dom t -> raise Toberemoved
+ | LocalAssum (_,t) | LocalDef (_,_,t) when is_dom t -> raise Toberemoved
| _ -> pre_princ,[]
with Not_found -> assert false
end
@@ -159,7 +163,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
try
let new_t,binders_to_remove_from_t = compute_new_princ_type remove env t in
let new_x : Name.t = get_name (Termops.ids_of_context env) x in
- let new_env = Environ.push_rel (x,None,t) env in
+ let new_env = Environ.push_rel (LocalAssum (x,t)) env in
let new_b,binders_to_remove_from_b = compute_new_princ_type remove new_env b in
if List.exists (eq_constr (mkRel 1)) binders_to_remove_from_b
then (Termops.pop new_b), filter_map (eq_constr (mkRel 1)) Termops.pop binders_to_remove_from_b
@@ -188,7 +192,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
let new_t,binders_to_remove_from_t = compute_new_princ_type remove env t in
let new_v,binders_to_remove_from_v = compute_new_princ_type remove env v in
let new_x : Name.t = get_name (Termops.ids_of_context env) x in
- let new_env = Environ.push_rel (x,Some v,t) env in
+ let new_env = Environ.push_rel (LocalDef (x,v,t)) env in
let new_b,binders_to_remove_from_b = compute_new_princ_type remove new_env b in
if List.exists (eq_constr (mkRel 1)) binders_to_remove_from_b
then (Termops.pop new_b),filter_map (eq_constr (mkRel 1)) Termops.pop binders_to_remove_from_b
@@ -227,7 +231,8 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
in
it_mkProd_or_LetIn
(it_mkProd_or_LetIn
- pre_res (List.map (fun (id,t,b) -> Name(Hashtbl.find tbl id), t,b)
+ pre_res (List.map (function Context.Named.Declaration.LocalAssum (id,b) -> LocalAssum (Name (Hashtbl.find tbl id), b)
+ | Context.Named.Declaration.LocalDef (id,t,b) -> LocalDef (Name (Hashtbl.find tbl id), t, b))
new_predicates)
)
princ_type_info.params
@@ -235,10 +240,12 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
let change_property_sort evd toSort princ princName =
+ let open Context.Rel.Declaration in
let princ_info = compute_elim_sig princ in
- let change_sort_in_predicate (x,v,t) =
- (x,None,
- let args,ty = decompose_prod t in
+ let change_sort_in_predicate decl =
+ LocalAssum
+ (get_name decl,
+ let args,ty = decompose_prod (get_type decl) in
let s = destSort ty in
Global.add_constraints (Univ.enforce_leq (univ_of_sort toSort) (univ_of_sort s) Univ.Constraint.empty);
compose_prod args (mkSort toSort)
@@ -648,12 +655,15 @@ let build_case_scheme fa =
let this_block_funs_indexes = Array.to_list this_block_funs_indexes in
List.assoc_f Constant.equal (fst (destConst funs)) this_block_funs_indexes
in
- let ind_fun =
+ let (ind, sf) =
let ind = first_fun_kn,funs_indexes in
(ind,Univ.Instance.empty)(*FIXME*),prop_sort
in
- let sigma, scheme =
- (fun (ind,sf) -> Indrec.build_case_analysis_scheme_default env sigma ind sf) ind_fun in
+ let sigma = Sigma.Unsafe.of_evar_map sigma in
+ let Sigma (scheme, sigma, _) =
+ Indrec.build_case_analysis_scheme_default env sigma ind sf
+ in
+ let sigma = Sigma.to_evar_map sigma in
let scheme_type = (Typing.unsafe_type_of env sigma ) scheme in
let sorts =
(fun (_,_,x) ->
diff --git a/plugins/funind/g_indfun.ml4 b/plugins/funind/g_indfun.ml4
index a15e46bfe..e93c395e3 100644
--- a/plugins/funind/g_indfun.ml4
+++ b/plugins/funind/g_indfun.ml4
@@ -16,8 +16,12 @@ open Constrexpr
open Indfun_common
open Indfun
open Genarg
+open Constrarg
open Tacticals
open Misctypes
+open Pcoq.Prim
+open Pcoq.Constr
+open Pcoq.Tactic
DECLARE PLUGIN "recdef_plugin"
@@ -55,7 +59,9 @@ let pr_with_bindings_typed prc prlc (c,bl) =
let pr_fun_ind_using_typed prc prlc _ opt_c =
match opt_c with
| None -> mt ()
- | Some b -> spc () ++ hov 2 (str "using" ++ spc () ++ pr_with_bindings_typed prc prlc b.Evd.it)
+ | Some b ->
+ let (b, _) = Tactics.run_delayed (Global.env ()) Evd.empty b in
+ spc () ++ hov 2 (str "using" ++ spc () ++ pr_with_bindings_typed prc prlc b)
ARGUMENT EXTEND fun_ind_using
@@ -88,7 +94,7 @@ let out_disjunctive = function
| loc, IntroAction (IntroOrAndPattern l) -> (loc,l)
| _ -> Errors.error "Disjunctive or conjunctive intro pattern expected."
-ARGUMENT EXTEND with_names TYPED AS simple_intropattern_opt PRINTED BY pr_intro_as_pat
+ARGUMENT EXTEND with_names TYPED AS intropattern_opt PRINTED BY pr_intro_as_pat
| [ "as" simple_intropattern(ipat) ] -> [ Some ipat ]
| [] ->[ None ]
END
@@ -144,10 +150,10 @@ module Tactic = Pcoq.Tactic
type function_rec_definition_loc_argtype = (Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) Loc.located
let (wit_function_rec_definition_loc : function_rec_definition_loc_argtype Genarg.uniform_genarg_type) =
- Genarg.create_arg None "function_rec_definition_loc"
+ Genarg.create_arg "function_rec_definition_loc"
let function_rec_definition_loc =
- Pcoq.create_generic_entry "function_rec_definition_loc" (Genarg.rawwit wit_function_rec_definition_loc)
+ Pcoq.create_generic_entry Pcoq.utactic "function_rec_definition_loc" (Genarg.rawwit wit_function_rec_definition_loc)
GEXTEND Gram
GLOBAL: function_rec_definition_loc ;
diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml
index 5d92fca5e..8a0a1a064 100644
--- a/plugins/funind/glob_term_to_relation.ml
+++ b/plugins/funind/glob_term_to_relation.ml
@@ -335,15 +335,17 @@ let raw_push_named (na,raw_value,raw_typ) env =
| Name id ->
let value = Option.map (fun x-> fst (Pretyping.understand env (Evd.from_env env) x)) raw_value in
let typ,ctx = Pretyping.understand env (Evd.from_env env) ~expected_type:Pretyping.IsType raw_typ in
- Environ.push_named (id,value,typ) env
+ let open Context.Named.Declaration in
+ Environ.push_named (of_tuple (id,value,typ)) env
let add_pat_variables pat typ env : Environ.env =
let rec add_pat_variables env pat typ : Environ.env =
+ let open Context.Rel.Declaration in
observe (str "new rel env := " ++ Printer.pr_rel_context_of env (Evd.from_env env));
match pat with
- | PatVar(_,na) -> Environ.push_rel (na,None,typ) env
+ | PatVar(_,na) -> Environ.push_rel (LocalAssum (na,typ)) env
| PatCstr(_,c,patl,na) ->
let Inductiveops.IndType(indf,indargs) =
try Inductiveops.find_rectype env (Evd.from_env env) typ
@@ -351,15 +353,16 @@ let add_pat_variables pat typ env : Environ.env =
in
let constructors = Inductiveops.get_constructors env indf in
let constructor : Inductiveops.constructor_summary = List.find (fun cs -> eq_constructor c (fst cs.Inductiveops.cs_cstr)) (Array.to_list constructors) in
- let cs_args_types :types list = List.map (fun (_,_,t) -> t) constructor.Inductiveops.cs_args in
+ let cs_args_types :types list = List.map get_type constructor.Inductiveops.cs_args in
List.fold_left2 add_pat_variables env patl (List.rev cs_args_types)
in
let new_env = add_pat_variables env pat typ in
let res =
fst (
- Context.fold_rel_context
- (fun (na,v,t) (env,ctxt) ->
- match na with
+ Context.Rel.fold_outside
+ (fun decl (env,ctxt) ->
+ let _,v,t = Context.Rel.Declaration.to_tuple decl in
+ match Context.Rel.Declaration.get_name decl with
| Anonymous -> assert false
| Name id ->
let new_t = substl ctxt t in
@@ -370,7 +373,8 @@ let add_pat_variables pat typ env : Environ.env =
Option.fold_right (fun v _ -> str "old value := " ++ Printer.pr_lconstr v ++ fnl ()) v (mt ()) ++
Option.fold_right (fun v _ -> str "new value := " ++ Printer.pr_lconstr v ++ fnl ()) new_v (mt ())
);
- (Environ.push_named (id,new_v,new_t) env,mkVar id::ctxt)
+ let open Context.Named.Declaration in
+ (Environ.push_named (of_tuple (id,new_v,new_t)) env,mkVar id::ctxt)
)
(Environ.rel_context new_env)
~init:(env,[])
@@ -398,7 +402,8 @@ let rec pattern_to_term_and_type env typ = function
in
let constructors = Inductiveops.get_constructors env indf in
let constructor = List.find (fun cs -> eq_constructor (fst cs.Inductiveops.cs_cstr) constr) (Array.to_list constructors) in
- let cs_args_types :types list = List.map (fun (_,_,t) -> t) constructor.Inductiveops.cs_args in
+ let open Context.Rel.Declaration in
+ let cs_args_types :types list = List.map get_type constructor.Inductiveops.cs_args in
let _,cstl = Inductiveops.dest_ind_family indf in
let csta = Array.of_list cstl in
let implicit_args =
@@ -597,9 +602,10 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
let v_as_constr,ctx = Pretyping.understand env (Evd.from_env env) v in
let v_type = Typing.unsafe_type_of env (Evd.from_env env) v_as_constr in
let new_env =
+ let open Context.Named.Declaration in
match n with
Anonymous -> env
- | Name id -> Environ.push_named (id,Some v_as_constr,v_type) env
+ | Name id -> Environ.push_named (of_tuple (id,Some v_as_constr,v_type)) env
in
let b_res = build_entry_lc new_env funnames avoid b in
combine_results (combine_letin n) v_res b_res
@@ -875,7 +881,7 @@ exception Continue
*)
let rec rebuild_cons env nb_args relname args crossed_types depth rt =
observe (str "rebuilding : " ++ pr_glob_constr rt);
-
+ let open Context.Rel.Declaration in
match rt with
| GProd(_,n,k,t,b) ->
let not_free_in_t id = not (is_free_in id t) in
@@ -895,7 +901,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
mkGApp(mkGVar(mk_rel_id this_relname),args'@[res_rt])
in
let t',ctx = Pretyping.understand env (Evd.from_env env) new_t in
- let new_env = Environ.push_rel (n,None,t') env in
+ let new_env = Environ.push_rel (LocalAssum (n,t')) env in
let new_b,id_to_exclude =
rebuild_cons new_env
nb_args relname
@@ -926,7 +932,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
let subst_b =
if is_in_b then b else replace_var_by_term id rt b
in
- let new_env = Environ.push_rel (n,None,t') env in
+ let new_env = Environ.push_rel (LocalAssum (n,t')) env in
let new_b,id_to_exclude =
rebuild_cons
new_env
@@ -970,9 +976,8 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
(fun acc var_as_constr arg ->
if isRel var_as_constr
then
- let (na,_,_) =
- Environ.lookup_rel (destRel var_as_constr) env
- in
+ let open Context.Rel.Declaration in
+ let na = get_name (Environ.lookup_rel (destRel var_as_constr) env) in
match na with
| Anonymous -> acc
| Name id' ->
@@ -1010,7 +1015,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
in
let new_env =
let t',ctx = Pretyping.understand env (Evd.from_env env) eq' in
- Environ.push_rel (n,None,t') env
+ Environ.push_rel (LocalAssum (n,t')) env
in
let new_b,id_to_exclude =
rebuild_cons
@@ -1048,7 +1053,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
with Continue ->
observe (str "computing new type for prod : " ++ pr_glob_constr rt);
let t',ctx = Pretyping.understand env (Evd.from_env env) t in
- let new_env = Environ.push_rel (n,None,t') env in
+ let new_env = Environ.push_rel (LocalAssum (n,t')) env in
let new_b,id_to_exclude =
rebuild_cons new_env
nb_args relname
@@ -1064,7 +1069,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
| _ ->
observe (str "computing new type for prod : " ++ pr_glob_constr rt);
let t',ctx = Pretyping.understand env (Evd.from_env env) t in
- let new_env = Environ.push_rel (n,None,t') env in
+ let new_env = Environ.push_rel (LocalAssum (n,t')) env in
let new_b,id_to_exclude =
rebuild_cons new_env
nb_args relname
@@ -1085,7 +1090,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
let t',ctx = Pretyping.understand env (Evd.from_env env) t in
match n with
| Name id ->
- let new_env = Environ.push_rel (n,None,t') env in
+ let new_env = Environ.push_rel (LocalAssum (n,t')) env in
let new_b,id_to_exclude =
rebuild_cons new_env
nb_args relname
@@ -1108,7 +1113,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
let t',ctx = Pretyping.understand env evd t in
let evd = Evd.from_ctx ctx in
let type_t' = Typing.unsafe_type_of env evd t' in
- let new_env = Environ.push_rel (n,Some t',type_t') env in
+ let new_env = Environ.push_rel (LocalDef (n,t',type_t')) env in
let new_b,id_to_exclude =
rebuild_cons new_env
nb_args relname
@@ -1132,7 +1137,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
depth t
in
let t',ctx = Pretyping.understand env (Evd.from_env env) new_t in
- let new_env = Environ.push_rel (na,None,t') env in
+ let new_env = Environ.push_rel (LocalAssum (na,t')) env in
let new_b,id_to_exclude =
rebuild_cons new_env
nb_args relname
@@ -1254,12 +1259,13 @@ let do_build_inductive
let relnames = Array.map mk_rel_id funnames in
let relnames_as_set = Array.fold_right Id.Set.add relnames Id.Set.empty in
(* Construction of the pseudo constructors *)
+ let open Context.Named.Declaration in
let evd,env =
Array.fold_right2
(fun id c (evd,env) ->
let evd,t = Typing.type_of env evd (mkConstU c) in
evd,
- Environ.push_named (id,None,t)
+ Environ.push_named (LocalAssum (id,t))
(* try *)
(* Typing.e_type_of env evd (mkConstU c) *)
(* with Not_found -> *)
@@ -1298,8 +1304,8 @@ let do_build_inductive
*)
let rel_arities = Array.mapi rel_arity funsargs in
Util.Array.fold_left2 (fun env rel_name rel_ar ->
- Environ.push_named (rel_name,None,
- fst (with_full_print (Constrintern.interp_constr env evd) rel_ar)) env) env relnames rel_arities
+ Environ.push_named (LocalAssum (rel_name,
+ fst (with_full_print (Constrintern.interp_constr env evd) rel_ar))) env) env relnames rel_arities
in
(* and of the real constructors*)
let constr i res =
diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml
index 3dbd43806..84a4d910e 100644
--- a/plugins/funind/indfun.ml
+++ b/plugins/funind/indfun.ml
@@ -1,3 +1,4 @@
+open Context.Rel.Declaration
open Errors
open Util
open Names
@@ -10,12 +11,13 @@ open Glob_term
open Declarations
open Misctypes
open Decl_kinds
+open Sigma.Notations
let is_rec_info scheme_info =
- let test_branche min acc (_,_,br) =
+ let test_branche min acc decl =
acc || (
let new_branche =
- it_mkProd_or_LetIn mkProp (fst (decompose_prod_assum br)) in
+ it_mkProd_or_LetIn mkProp (fst (decompose_prod_assum (get_type decl))) in
let free_rels_in_br = Termops.free_rels new_branche in
let max = min + scheme_info.Tactics.npredicates in
Int.Set.exists (fun i -> i >= min && i< max) free_rels_in_br
@@ -85,7 +87,7 @@ let functional_induction with_clean c princl pat =
in
let encoded_pat_as_patlist =
List.make (List.length args + List.length c_list - 1) None @ [pat] in
- List.map2 (fun c pat -> ((None,Tacexpr.ElimOnConstr (fun env sigma -> sigma,(c,NoBindings))),(None,pat),None))
+ List.map2 (fun c pat -> ((None,Tacexpr.ElimOnConstr ({ Tacexpr.delayed = fun env sigma -> Sigma ((c,NoBindings), sigma, Sigma.refl) })),(None,pat),None))
(args@c_list) encoded_pat_as_patlist
in
let princ' = Some (princ,bindings) in
@@ -112,7 +114,7 @@ let functional_induction with_clean c princl pat =
in
Tacticals.tclTHEN
(Tacticals.tclMAP (fun id -> Tacticals.tclTRY (Proofview.V82.of_tactic (Equality.subst_gen (do_rewrite_dependent ()) [id]))) idl )
- (Tactics.reduce flag Locusops.allHypsAndConcl)
+ (Proofview.V82.of_tactic (Tactics.reduce flag Locusops.allHypsAndConcl))
g
else Tacticals.tclIDTAC g
in
@@ -152,7 +154,8 @@ let build_newrecursive
let evdref = ref (Evd.from_env env0) in
let _, (_, impls') = Constrintern.interp_context_evars env evdref bl in
let impl = Constrintern.compute_internalization_data env0 Constrintern.Recursive arity impls' in
- (Environ.push_named (recname,None,arity) env, Id.Map.add recname impl impls))
+ let open Context.Named.Declaration in
+ (Environ.push_named (LocalAssum (recname,arity)) env, Id.Map.add recname impl impls))
(env0,Constrintern.empty_internalization_env) lnameargsardef in
let recdef =
(* Declare local notations *)
@@ -727,9 +730,9 @@ let rec add_args id new_args b =
List.map (fun (e,o) -> add_args id new_args e,o) bl)
| CCases(loc,sty,b_option,cel,cal) ->
CCases(loc,sty,Option.map (add_args id new_args) b_option,
- List.map (fun (b,(na,b_option)) ->
+ List.map (fun (b,na,b_option) ->
add_args id new_args b,
- (na, b_option)) cel,
+ na, b_option) cel,
List.map (fun (loc,cpl,e) -> (loc,cpl,add_args id new_args e)) cal
)
| CLetTuple(loc,nal,(na,b_option),b1,b2) ->
@@ -751,10 +754,8 @@ let rec add_args id new_args b =
| CCast(loc,b1,b2) ->
CCast(loc,add_args id new_args b1,
Miscops.map_cast_type (add_args id new_args) b2)
- | CRecord (loc, w, pars) ->
- CRecord (loc,
- (match w with Some w -> Some (add_args id new_args w) | _ -> None),
- List.map (fun (e,o) -> e, add_args id new_args o) pars)
+ | CRecord (loc, pars) ->
+ CRecord (loc, List.map (fun (e,o) -> e, add_args id new_args o) pars)
| CNotation _ -> anomaly ~label:"add_args " (Pp.str "CNotation")
| CGeneralization _ -> anomaly ~label:"add_args " (Pp.str "CGeneralization")
| CPrim _ -> b
diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml
index a800c186a..6a5a5ad53 100644
--- a/plugins/funind/invfun.ml
+++ b/plugins/funind/invfun.ml
@@ -5,6 +5,7 @@
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+
open Tacexpr
open Declarations
open Errors
@@ -19,6 +20,8 @@ open Tactics
open Indfun_common
open Tacmach
open Misctypes
+open Termops
+open Context.Rel.Declaration
(* Some pretty printing function for debugging purpose *)
@@ -70,8 +73,8 @@ let do_observe_tac s tac g =
with reraise ->
let reraise = Errors.push reraise in
let e = Cerrors.process_vernac_interp_error reraise in
- observe (str "observation "++ s++str " raised exception " ++
- Errors.iprint e ++ str " on goal " ++ goal );
+ observe (hov 0 (str "observation "++ s++str " raised exception " ++
+ Errors.iprint e ++ str " on goal" ++ fnl() ++ goal ));
iraise reraise;;
@@ -133,18 +136,21 @@ let generate_type evd g_to_f f graph i =
let fun_ctxt,res_type =
match ctxt with
| [] | [_] -> anomaly (Pp.str "Not a valid context")
- | (_,_,res_type)::fun_ctxt -> fun_ctxt,res_type
+ | decl :: fun_ctxt -> fun_ctxt, get_type decl
in
let rec args_from_decl i accu = function
| [] -> accu
- | (_, Some _, _) :: l ->
+ | LocalDef _ :: l ->
args_from_decl (succ i) accu l
| _ :: l ->
let t = mkRel i in
args_from_decl (succ i) (t :: accu) l
in
(*i We need to name the vars [res] and [fv] i*)
- let filter = function (Name id,_,_) -> Some id | (Anonymous,_,_) -> None in
+ let filter = fun decl -> match get_name decl with
+ | Name id -> Some id
+ | Anonymous -> None
+ in
let named_ctxt = List.map_filter filter fun_ctxt in
let res_id = Namegen.next_ident_away_in_goal (Id.of_string "_res") named_ctxt in
let fv_id = Namegen.next_ident_away_in_goal (Id.of_string "fv") (res_id :: named_ctxt) in
@@ -170,12 +176,12 @@ let generate_type evd g_to_f f graph i =
\[\forall (x_1:t_1)\ldots(x_n:t_n), let fv := f x_1\ldots x_n in, forall res, \]
i*)
let pre_ctxt =
- (Name res_id,None,lift 1 res_type)::(Name fv_id,Some (mkApp(f,args_as_rels)),res_type)::fun_ctxt
+ LocalAssum (Name res_id, lift 1 res_type) :: LocalDef (Name fv_id, mkApp (f,args_as_rels), res_type) :: fun_ctxt
in
(*i and we can return the solution depending on which lemma type we are defining i*)
if g_to_f
- then (Anonymous,None,graph_applied)::pre_ctxt,(lift 1 res_eq_f_of_args),graph
- else (Anonymous,None,res_eq_f_of_args)::pre_ctxt,(lift 1 graph_applied),graph
+ then LocalAssum (Anonymous,graph_applied)::pre_ctxt,(lift 1 res_eq_f_of_args),graph
+ else LocalAssum (Anonymous,res_eq_f_of_args)::pre_ctxt,(lift 1 graph_applied),graph
(*
@@ -259,10 +265,10 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
(* and built the intro pattern for each of them *)
let intro_pats =
List.map
- (fun (_,_,br_type) ->
+ (fun decl ->
List.map
(fun id -> Loc.ghost, IntroNaming (IntroIdentifier id))
- (generate_fresh_id (Id.of_string "y") ids (List.length (fst (decompose_prod_assum br_type))))
+ (generate_fresh_id (Id.of_string "y") ids (List.length (fst (decompose_prod_assum (get_type decl)))))
)
branches
in
@@ -362,14 +368,14 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
(* unfolding of all the defined variables introduced by this branch *)
(* observe_tac "unfolding" pre_tac; *)
(* $zeta$ normalizing of the conclusion *)
- reduce
+ Proofview.V82.of_tactic (reduce
(Genredexpr.Cbv
{ Redops.all_flags with
Genredexpr.rDelta = false ;
Genredexpr.rConst = []
}
)
- Locusops.onConcl;
+ Locusops.onConcl);
observe_tac ("toto ") tclIDTAC;
(* introducing the the result of the graph and the equality hypothesis *)
@@ -389,10 +395,10 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
(fun ((_,(ctxt,concl))) ->
match ctxt with
| [] | [_] | [_;_] -> anomaly (Pp.str "bad context")
- | hres::res::(x,_,t)::ctxt ->
+ | hres::res::decl::ctxt ->
let res = Termops.it_mkLambda_or_LetIn
(Termops.it_mkProd_or_LetIn concl [hres;res])
- ((x,None,t)::ctxt)
+ (LocalAssum (get_name decl, get_type decl) :: ctxt)
in
res
)
@@ -407,8 +413,8 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
let bindings =
let params_bindings,avoid =
List.fold_left2
- (fun (bindings,avoid) (x,_,_) p ->
- let id = Namegen.next_ident_away (Nameops.out_name x) avoid in
+ (fun (bindings,avoid) decl p ->
+ let id = Namegen.next_ident_away (Nameops.out_name (get_name decl)) avoid in
p::bindings,id::avoid
)
([],pf_ids_of_hyps g)
@@ -417,8 +423,8 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
in
let lemmas_bindings =
List.rev (fst (List.fold_left2
- (fun (bindings,avoid) (x,_,_) p ->
- let id = Namegen.next_ident_away (Nameops.out_name x) avoid in
+ (fun (bindings,avoid) decl p ->
+ let id = Namegen.next_ident_away (Nameops.out_name (get_name decl)) avoid in
(nf_zeta p)::bindings,id::avoid)
([],avoid)
princ_infos.predicates
@@ -454,10 +460,11 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
generalize every hypothesis which depends of [x] but [hyp]
*)
let generalize_dependent_of x hyp g =
+ let open Context.Named.Declaration in
tclMAP
(function
- | (id,None,t) when not (Id.equal id hyp) &&
- (Termops.occur_var (pf_env g) x t) -> tclTHEN (Tactics.Simple.generalize [mkVar id]) (thin [id])
+ | LocalAssum (id,t) when not (Id.equal id hyp) &&
+ (Termops.occur_var (pf_env g) x t) -> tclTHEN (Tactics.generalize [mkVar id]) (thin [id])
| _ -> tclIDTAC
)
(pf_hyps g)
@@ -467,6 +474,15 @@ let generalize_dependent_of x hyp g =
(* [intros_with_rewrite] do the intros in each branch and treat each new hypothesis
(unfolding, substituting, destructing cases \ldots)
*)
+let tauto =
+ let dp = List.map Id.of_string ["Tauto" ; "Init"; "Coq"] in
+ let mp = ModPath.MPfile (DirPath.make dp) in
+ let kn = KerName.make2 mp (Label.make "tauto") in
+ Proofview.tclBIND (Proofview.tclUNIT ()) begin fun () ->
+ let body = Tacenv.interp_ltac kn in
+ Tacinterp.eval_tactic body
+ end
+
let rec intros_with_rewrite g =
observe_tac "intros_with_rewrite" intros_with_rewrite_aux g
and intros_with_rewrite_aux : tactic =
@@ -483,15 +499,15 @@ and intros_with_rewrite_aux : tactic =
tclTHENSEQ [ Proofview.V82.of_tactic (Simple.intro id); thin [id]; intros_with_rewrite ] g
else if isVar args.(1) && (Environ.evaluable_named (destVar args.(1)) (pf_env g))
then tclTHENSEQ[
- unfold_in_concl [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(1)))];
- tclMAP (fun id -> tclTRY(unfold_in_hyp [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(1)))] ((destVar args.(1)),Locus.InHyp) ))
+ Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(1)))]);
+ tclMAP (fun id -> tclTRY(Proofview.V82.of_tactic (unfold_in_hyp [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(1)))] ((destVar args.(1)),Locus.InHyp) )))
(pf_ids_of_hyps g);
intros_with_rewrite
] g
else if isVar args.(2) && (Environ.evaluable_named (destVar args.(2)) (pf_env g))
then tclTHENSEQ[
- unfold_in_concl [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(2)))];
- tclMAP (fun id -> tclTRY(unfold_in_hyp [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(2)))] ((destVar args.(2)),Locus.InHyp) ))
+ Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(2)))]);
+ tclMAP (fun id -> tclTRY(Proofview.V82.of_tactic (unfold_in_hyp [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(2)))] ((destVar args.(2)),Locus.InHyp) )))
(pf_ids_of_hyps g);
intros_with_rewrite
] g
@@ -523,7 +539,7 @@ and intros_with_rewrite_aux : tactic =
] g
end
| Ind _ when eq_constr t (Coqlib.build_coq_False ()) ->
- Proofview.V82.of_tactic Tauto.tauto g
+ Proofview.V82.of_tactic tauto g
| Case(_,_,v,_) ->
tclTHENSEQ[
Proofview.V82.of_tactic (simplest_case v);
@@ -531,12 +547,12 @@ and intros_with_rewrite_aux : tactic =
] g
| LetIn _ ->
tclTHENSEQ[
- reduce
+ Proofview.V82.of_tactic (reduce
(Genredexpr.Cbv
{Redops.all_flags
with Genredexpr.rDelta = false;
})
- Locusops.onConcl
+ Locusops.onConcl)
;
intros_with_rewrite
] g
@@ -546,12 +562,12 @@ and intros_with_rewrite_aux : tactic =
end
| LetIn _ ->
tclTHENSEQ[
- reduce
+ Proofview.V82.of_tactic (reduce
(Genredexpr.Cbv
{Redops.all_flags
with Genredexpr.rDelta = false;
})
- Locusops.onConcl
+ Locusops.onConcl)
;
intros_with_rewrite
] g
@@ -662,10 +678,10 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic =
let branches = List.rev princ_infos.branches in
let intro_pats =
List.map
- (fun (_,_,br_type) ->
+ (fun decl ->
List.map
(fun id -> id)
- (generate_fresh_id (Id.of_string "y") ids (nb_prod br_type))
+ (generate_fresh_id (Id.of_string "y") ids (nb_prod (get_type decl)))
)
branches
in
@@ -691,18 +707,18 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic =
Proofview.V82.of_tactic (Equality.rewriteLR (mkConst eq_lemma));
(* Don't forget to $\zeta$ normlize the term since the principles
have been $\zeta$-normalized *)
- reduce
+ Proofview.V82.of_tactic (reduce
(Genredexpr.Cbv
{Redops.all_flags
with Genredexpr.rDelta = false;
})
- Locusops.onConcl
+ Locusops.onConcl)
;
- Simple.generalize (List.map mkVar ids);
+ generalize (List.map mkVar ids);
thin ids
]
else
- unfold_in_concl [(Locus.AllOccurrences, Names.EvalConstRef (fst (destConst f)))]
+ Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalConstRef (fst (destConst f)))])
in
(* The proof of each branche itself *)
let ind_number = ref 0 in
@@ -737,7 +753,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic =
tclTHENSEQ
[ tclMAP (fun id -> Proofview.V82.of_tactic (Simple.intro id)) (args_names@[res;hres]);
observe_tac "h_generalize"
- (Simple.generalize [mkApp(applist(graph_principle,params),Array.map (fun c -> applist(c,params)) lemmas)]);
+ (generalize [mkApp(applist(graph_principle,params),Array.map (fun c -> applist(c,params)) lemmas)]);
Proofview.V82.of_tactic (Simple.intro graph_principle_id);
observe_tac "" (tclTHEN_i
(observe_tac "elim" (Proofview.V82.of_tactic (elim false None (mkVar hres,NoBindings) (Some (mkVar graph_principle_id,NoBindings)))))
@@ -920,7 +936,7 @@ let revert_graph kn post_tac hid g =
let f_args,res = Array.chop (Array.length args - 1) args in
tclTHENSEQ
[
- Simple.generalize [applist(mkConst f_complete,(Array.to_list f_args)@[res.(0);mkVar hid])];
+ generalize [applist(mkConst f_complete,(Array.to_list f_args)@[res.(0);mkVar hid])];
thin [hid];
Proofview.V82.of_tactic (Simple.intro hid);
post_tac hid
@@ -964,7 +980,7 @@ let functional_inversion kn hid fconst f_correct : tactic =
in
tclTHENSEQ[
pre_tac hid;
- Simple.generalize [applist(f_correct,(Array.to_list f_args)@[res;mkVar hid])];
+ generalize [applist(f_correct,(Array.to_list f_args)@[res;mkVar hid])];
thin [hid];
Proofview.V82.of_tactic (Simple.intro hid);
Proofview.V82.of_tactic (Inv.inv FullInversion None (NamedHyp hid));
diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml
index 87d7ca76d..c71d9a9ca 100644
--- a/plugins/funind/merge.ml
+++ b/plugins/funind/merge.ml
@@ -19,12 +19,12 @@ open Pp
open Names
open Term
open Vars
-open Context
open Termops
open Declarations
open Glob_term
open Glob_termops
open Decl_kinds
+open Context.Rel.Declaration
(** {1 Utilities} *)
@@ -135,9 +135,9 @@ let showind (id:Id.t) =
let cstrid = Constrintern.global_reference id in
let ind1,cstrlist = Inductiveops.find_inductive (Global.env()) Evd.empty cstrid in
let mib1,ib1 = Inductive.lookup_mind_specif (Global.env()) (fst ind1) in
- List.iter (fun (nm, optcstr, tp) ->
- print_string (string_of_name nm^":");
- prconstr tp; print_string "\n")
+ List.iter (fun decl ->
+ print_string (string_of_name (Context.Rel.Declaration.get_name decl) ^ ":");
+ prconstr (get_type decl); print_string "\n")
ib1.mind_arity_ctxt;
Printf.printf "arity :"; prconstr (Inductiveops.type_of_inductive (Global.env ()) ind1);
Array.iteri
@@ -258,27 +258,27 @@ type merge_infos =
lnk2: int merged_arg array;
(** rec params which remain rec param (ie not linked) *)
- recprms1: rel_declaration list;
- recprms2: rel_declaration list;
+ recprms1: Context.Rel.Declaration.t list;
+ recprms2: Context.Rel.Declaration.t list;
nrecprms1: int;
nrecprms2: int;
(** rec parms which became non parm (either linked to something
or because after a rec parm that became non parm) *)
- otherprms1: rel_declaration list;
- otherprms2: rel_declaration list;
+ otherprms1: Context.Rel.Declaration.t list;
+ otherprms2: Context.Rel.Declaration.t list;
notherprms1:int;
notherprms2:int;
(** args which remain args in merge *)
- args1:rel_declaration list;
- args2:rel_declaration list;
+ args1:Context.Rel.Declaration.t list;
+ args2:Context.Rel.Declaration.t list;
nargs1:int;
nargs2:int;
(** functional result args *)
- funresprms1: rel_declaration list;
- funresprms2: rel_declaration list;
+ funresprms1: Context.Rel.Declaration.t list;
+ funresprms2: Context.Rel.Declaration.t list;
nfunresprms1:int;
nfunresprms2:int;
}
@@ -460,11 +460,12 @@ let shift_linked_params mib1 mib2 (lnk1:linked_var array) (lnk2:linked_var array
let recprms2,otherprms2,args2,funresprms2 = bldprms (List.rev oib2.mind_arity_ctxt) mlnk2 in
let _ = prstr "\notherprms1:\n" in
let _ =
- List.iter (fun (x,_,y) -> prstr (string_of_name x^" : ");prconstr y;prstr "\n")
+ List.iter (fun decl -> prstr (string_of_name (get_name decl) ^ " : ");
+ prconstr (get_type decl); prstr "\n")
otherprms1 in
let _ = prstr "\notherprms2:\n" in
let _ =
- List.iter (fun (x,_,y) -> prstr (string_of_name x^" : ");prconstr y;prstr "\n")
+ List.iter (fun decl -> prstr (string_of_name (get_name decl) ^ " : "); prconstr (get_type decl); prstr "\n")
otherprms2 in
{
ident=id;
@@ -824,9 +825,11 @@ let merge_rec_params_and_arity prms1 prms2 shift (concl:constr) =
let concl = Constrextern.extern_constr false (Global.env()) Evd.empty concl in
let arity,_ =
List.fold_left
- (fun (acc,env) (nm,_,c) ->
+ (fun (acc,env) decl ->
+ let nm = Context.Rel.Declaration.get_name decl in
+ let c = get_type decl in
let typ = Constrextern.extern_constr false env Evd.empty c in
- let newenv = Environ.push_rel (nm,None,c) env in
+ let newenv = Environ.push_rel (LocalAssum (nm,c)) env in
CProdN (Loc.ghost, [[(Loc.ghost,nm)],Constrexpr_ops.default_binder_kind,typ] , acc) , newenv)
(concl,Global.env())
(shift.funresprms2 @ shift.funresprms1
@@ -851,12 +854,12 @@ let glob_constr_list_to_inductive_expr prms1 prms2 mib1 mib2 shift
lident , bindlist , Some cstr_expr , lcstor_expr
-let mkProd_reldecl (rdecl:rel_declaration) (t2:glob_constr) =
+let mkProd_reldecl (rdecl:Context.Rel.Declaration.t) (t2:glob_constr) =
match rdecl with
- | (nme,None,t) ->
+ | LocalAssum (nme,t) ->
let traw = Detyping.detype false [] (Global.env()) Evd.empty t in
GProd (Loc.ghost,nme,Explicit,traw,t2)
- | (_,Some _,_) -> assert false
+ | LocalDef _ -> assert false
(** [merge_inductive ind1 ind2 lnk] merges two graphs, linking
@@ -970,7 +973,7 @@ let funify_branches relinfo nfuns branch =
| Rel i -> let reali = i-shift in (reali>=0 && reali<relinfo.nbranches)
| _ -> false in
(* FIXME: *)
- (Anonymous,Some mkProp,mkProp)
+ LocalDef (Anonymous,mkProp,mkProp)
let relprinctype_to_funprinctype relprinctype nfuns =
diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml
index 065d0fe53..046c7aa43 100644
--- a/plugins/funind/recdef.ml
+++ b/plugins/funind/recdef.ml
@@ -29,6 +29,7 @@ open Proof_type
open Pfedit
open Glob_term
open Pretyping
+open Termops
open Constrintern
open Misctypes
open Genredexpr
@@ -38,7 +39,8 @@ open Auto
open Eauto
open Indfun_common
-
+open Sigma.Notations
+open Context.Rel.Declaration
(* Ugly things which should not be here *)
@@ -179,7 +181,7 @@ let (value_f:constr list -> global_reference -> constr) =
)
in
let context = List.map
- (fun (x, c) -> Name x, None, c) (List.combine rev_x_id_l (List.rev al))
+ (fun (x, c) -> LocalAssum (Name x, c)) (List.combine rev_x_id_l (List.rev al))
in
let env = Environ.push_rel_context context (Global.env ()) in
let glob_body =
@@ -212,10 +214,10 @@ let rec print_debug_queue b e =
begin
let lmsg,goal = Stack.pop debug_queue in
if b then
- Pp.msg_debug (lmsg ++ (str " raised exception " ++ Errors.print e) ++ str " on goal " ++ goal)
+ Pp.msg_debug (hov 1 (lmsg ++ (str " raised exception " ++ Errors.print e) ++ str " on goal" ++ fnl() ++ goal))
else
begin
- Pp.msg_debug (str " from " ++ lmsg ++ str " on goal " ++ goal);
+ Pp.msg_debug (hov 1 (str " from " ++ lmsg ++ str " on goal"++fnl() ++ goal));
end;
(* print_debug_queue false e; *)
end
@@ -274,8 +276,8 @@ let tclUSER tac is_mes l g =
if is_mes
then observe_tclTHENLIST (str "tclUSER2")
[
- unfold_in_concl [(Locus.AllOccurrences, evaluable_of_global_reference
- (delayed_force Indfun_common.ltof_ref))];
+ Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, evaluable_of_global_reference
+ (delayed_force Indfun_common.ltof_ref))]);
tac
]
else tac
@@ -560,10 +562,10 @@ let rec destruct_bounds_aux infos (bound,hyple,rechyps) lbounds g =
observe_tclTHENLIST (str "destruct_bounds_aux2")[
observe_tac (str "clearing k ") (clear [id]);
h_intros [k;h';def];
- observe_tac (str "simple_iter") (simpl_iter Locusops.onConcl);
+ observe_tac (str "simple_iter") (Proofview.V82.of_tactic (simpl_iter Locusops.onConcl));
observe_tac (str "unfold functional")
- (unfold_in_concl[(Locus.OnlyOccurrences [1],
- evaluable_of_global_reference infos.func)]);
+ (Proofview.V82.of_tactic (unfold_in_concl[(Locus.OnlyOccurrences [1],
+ evaluable_of_global_reference infos.func)]));
(
observe_tclTHENLIST (str "test")[
list_rewrite true
@@ -676,8 +678,10 @@ let mkDestructEq :
let hyps = pf_hyps g in
let to_revert =
Util.List.map_filter
- (fun (id, _, t) ->
- if Id.List.mem id not_on_hyp || not (Termops.occur_term expr t)
+ (fun decl ->
+ let open Context.Named.Declaration in
+ let id = get_id decl in
+ if Id.List.mem id not_on_hyp || not (Termops.occur_term expr (get_type decl))
then None else Some id) hyps in
let to_revert_constr = List.rev_map mkVar to_revert in
let type_of_expr = pf_unsafe_type_of g expr in
@@ -685,11 +689,13 @@ let mkDestructEq :
to_revert_constr in
pf_typel new_hyps (fun _ ->
observe_tclTHENLIST (str "mkDestructEq")
- [Simple.generalize new_hyps;
+ [generalize new_hyps;
(fun g2 ->
- Proofview.V82.of_tactic (change_in_concl None
- (fun patvars sigma ->
- pattern_occs [Locus.AllOccurrencesBut [1], expr] (pf_env g2) sigma (pf_concl g2))) g2);
+ let changefun patvars = { run = fun sigma ->
+ let redfun = pattern_occs [Locus.AllOccurrencesBut [1], expr] in
+ redfun.Reductionops.e_redfun (pf_env g2) sigma (pf_concl g2)
+ } in
+ Proofview.V82.of_tactic (change_in_concl None changefun) g2);
Proofview.V82.of_tactic (simplest_case expr)]), to_revert
@@ -897,10 +903,10 @@ let make_rewrite expr_info l hp max =
[observe_tac(str "make_rewrite finalize") (
(* tclORELSE( h_reflexivity) *)
(observe_tclTHENLIST (str "make_rewrite")[
- simpl_iter Locusops.onConcl;
+ Proofview.V82.of_tactic (simpl_iter Locusops.onConcl);
observe_tac (str "unfold functional")
- (unfold_in_concl[(Locus.OnlyOccurrences [1],
- evaluable_of_global_reference expr_info.func)]);
+ (Proofview.V82.of_tactic (unfold_in_concl[(Locus.OnlyOccurrences [1],
+ evaluable_of_global_reference expr_info.func)]));
(list_rewrite true
(List.map (fun e -> mkVar e,true) expr_info.eqs));
@@ -1110,7 +1116,7 @@ let termination_proof_header is_mes input_type ids args_id relation
[observe_tac (str "generalize")
(onNLastHypsId (nargs+1)
(tclMAP (fun id ->
- tclTHEN (Tactics.Simple.generalize [mkVar id]) (clear [id]))
+ tclTHEN (Tactics.generalize [mkVar id]) (clear [id]))
))
;
observe_tac (str "fix") (fix (Some hrec) (nargs+1));
@@ -1248,7 +1254,7 @@ let clear_goals =
then Termops.pop b'
else if b' == b then t
else mkProd(na,t',b')
- | _ -> map_constr clear_goal t
+ | _ -> Term.map_constr clear_goal t
in
List.map clear_goal
@@ -1300,7 +1306,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp
let hid = next_ident_away_in_goal h_id (pf_ids_of_hyps gls) in
observe_tclTHENLIST (str "")
[
- Simple.generalize [lemma];
+ generalize [lemma];
Proofview.V82.of_tactic (Simple.intro hid);
(fun g ->
let ids = pf_ids_of_hyps g in
@@ -1327,10 +1333,10 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp
tclFIRST[
tclTHEN
(Proofview.V82.of_tactic (eapply_with_bindings (mkVar (List.nth !lid !h_num), NoBindings)))
- e_assumption;
+ (Proofview.V82.of_tactic e_assumption);
Eauto.eauto_with_bases
(true,5)
- [Evd.empty,Lazy.force refl_equal]
+ [{ Tacexpr.delayed = fun _ sigma -> Sigma.here (Lazy.force refl_equal) sigma}]
[Hints.Hint_db.empty empty_transparent_state false]
]
)
@@ -1420,7 +1426,7 @@ let start_equation (f:global_reference) (term_f:global_reference)
let x = n_x_id ids nargs in
observe_tac (str "start_equation") (observe_tclTHENLIST (str "start_equation") [
h_intros x;
- unfold_in_concl [(Locus.AllOccurrences, evaluable_of_global_reference f)];
+ Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, evaluable_of_global_reference f)]);
observe_tac (str "simplest_case")
(Proofview.V82.of_tactic (simplest_case (mkApp (terminate_constr,
Array.of_list (List.map mkVar x)))));
@@ -1484,7 +1490,7 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num
let env = Global.env() in
let evd = ref (Evd.from_env env) in
let function_type = interp_type_evars env evd type_of_f in
- let env = push_named (function_name,None,function_type) env in
+ let env = push_named (Context.Named.Declaration.LocalAssum (function_name,function_type)) env in
(* Pp.msgnl (str "function type := " ++ Printer.pr_lconstr function_type); *)
let ty = interp_type_evars env evd ~impls:rec_impls eq in
let evm, nf = Evarutil.nf_evars_and_universes !evd in
@@ -1492,7 +1498,7 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num
let function_type = nf function_type in
(* Pp.msgnl (str "lemma type := " ++ Printer.pr_lconstr equation_lemma_type ++ fnl ()); *)
let res_vars,eq' = decompose_prod equation_lemma_type in
- let env_eq' = Environ.push_rel_context (List.map (fun (x,y) -> (x,None,y)) res_vars) env in
+ let env_eq' = Environ.push_rel_context (List.map (fun (x,y) -> LocalAssum (x,y)) res_vars) env in
let eq' = nf_zeta env_eq' eq' in
let res =
(* Pp.msgnl (str "res_var :=" ++ Printer.pr_lconstr_env (push_rel_context (List.map (function (x,t) -> (x,None,t)) res_vars) env) eq'); *)
@@ -1510,7 +1516,7 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num
let functional_id = add_suffix function_name "_F" in
let term_id = add_suffix function_name "_terminate" in
let functional_ref = declare_fun functional_id (IsDefinition Decl_kinds.Definition) ~ctx:(snd (Evd.universe_context evm)) res in
- let env_with_pre_rec_args = push_rel_context(List.map (function (x,t) -> (x,None,t)) pre_rec_args) env in
+ let env_with_pre_rec_args = push_rel_context(List.map (function (x,t) -> LocalAssum (x,t)) pre_rec_args) env in
let relation =
fst (*FIXME*)(interp_constr
env_with_pre_rec_args
diff --git a/plugins/micromega/Psatz.v b/plugins/micromega/Psatz.v
index a461b26a0..ba1f8956e 100644
--- a/plugins/micromega/Psatz.v
+++ b/plugins/micromega/Psatz.v
@@ -69,7 +69,7 @@ Ltac xpsatz dom d :=
end in tac.
Tactic Notation "psatz" constr(dom) int_or_var(n) := xpsatz dom n.
-Tactic Notation "psatz" constr(dom) := xpsatz dom ltac:-1.
+Tactic Notation "psatz" constr(dom) := xpsatz dom ltac:(-1).
Ltac psatzl dom :=
let tac := lazymatch dom with
@@ -96,6 +96,14 @@ Ltac psatzl dom :=
Ltac lra :=
first [ psatzl R | psatzl Q ].
+Ltac nra :=
+ unfold Rdiv in * ;
+ xnra ;
+ abstract
+ (intros __wit __varmap __ff ;
+ change (Tauto.eval_f (Reval_formula (@find R 0%R __varmap)) __ff) ;
+ apply (RTautoChecker_sound __ff __wit); vm_compute ; reflexivity).
+
(* Local Variables: *)
diff --git a/plugins/micromega/certificate.ml b/plugins/micromega/certificate.ml
index a5fceb562..1561c811c 100644
--- a/plugins/micromega/certificate.ml
+++ b/plugins/micromega/certificate.ml
@@ -63,82 +63,82 @@ let r_spec = z_spec
let dev_form n_spec p =
let rec dev_form p =
match p with
- | Mc.PEc z -> Poly.constant (n_spec.number_to_num z)
- | Mc.PEX v -> Poly.variable (C2Ml.positive v)
- | Mc.PEmul(p1,p2) ->
- let p1 = dev_form p1 in
- let p2 = dev_form p2 in
- Poly.product p1 p2
- | Mc.PEadd(p1,p2) -> Poly.addition (dev_form p1) (dev_form p2)
- | Mc.PEopp p -> Poly.uminus (dev_form p)
- | Mc.PEsub(p1,p2) -> Poly.addition (dev_form p1) (Poly.uminus (dev_form p2))
- | Mc.PEpow(p,n) ->
- let p = dev_form p in
- let n = C2Ml.n n in
- let rec pow n =
- if Int.equal n 0
- then Poly.constant (n_spec.number_to_num n_spec.unit)
- else Poly.product p (pow (n-1)) in
- pow n in
- dev_form p
+ | Mc.PEc z -> Poly.constant (n_spec.number_to_num z)
+ | Mc.PEX v -> Poly.variable (C2Ml.positive v)
+ | Mc.PEmul(p1,p2) ->
+ let p1 = dev_form p1 in
+ let p2 = dev_form p2 in
+ Poly.product p1 p2
+ | Mc.PEadd(p1,p2) -> Poly.addition (dev_form p1) (dev_form p2)
+ | Mc.PEopp p -> Poly.uminus (dev_form p)
+ | Mc.PEsub(p1,p2) -> Poly.addition (dev_form p1) (Poly.uminus (dev_form p2))
+ | Mc.PEpow(p,n) ->
+ let p = dev_form p in
+ let n = C2Ml.n n in
+ let rec pow n =
+ if Int.equal n 0
+ then Poly.constant (n_spec.number_to_num n_spec.unit)
+ else Poly.product p (pow (n-1)) in
+ pow n in
+ dev_form p
let monomial_to_polynomial mn =
Monomial.fold
(fun v i acc ->
- let v = Ml2C.positive v in
- let mn = if Int.equal i 1 then Mc.PEX v else Mc.PEpow (Mc.PEX v ,Ml2C.n i) in
- if Pervasives.(=) acc (Mc.PEc (Mc.Zpos Mc.XH)) (** FIXME *)
- then mn
- else Mc.PEmul(mn,acc))
- mn
- (Mc.PEc (Mc.Zpos Mc.XH))
+ let v = Ml2C.positive v in
+ let mn = if Int.equal i 1 then Mc.PEX v else Mc.PEpow (Mc.PEX v ,Ml2C.n i) in
+ if Pervasives.(=) acc (Mc.PEc (Mc.Zpos Mc.XH)) (** FIXME *)
+ then mn
+ else Mc.PEmul(mn,acc))
+ mn
+ (Mc.PEc (Mc.Zpos Mc.XH))
let list_to_polynomial vars l =
assert (List.for_all (fun x -> ceiling_num x =/ x) l);
let var x = monomial_to_polynomial (List.nth vars x) in
-
+
let rec xtopoly p i = function
| [] -> p
| c::l -> if c =/ (Int 0) then xtopoly p (i+1) l
- else let c = Mc.PEc (Ml2C.bigint (numerator c)) in
- let mn =
- if Pervasives.(=) c (Mc.PEc (Mc.Zpos Mc.XH))
- then var i
- else Mc.PEmul (c,var i) in
- let p' = if Pervasives.(=) p (Mc.PEc Mc.Z0) then mn else
- Mc.PEadd (mn, p) in
- xtopoly p' (i+1) l in
-
- xtopoly (Mc.PEc Mc.Z0) 0 l
+ else let c = Mc.PEc (Ml2C.bigint (numerator c)) in
+ let mn =
+ if Pervasives.(=) c (Mc.PEc (Mc.Zpos Mc.XH))
+ then var i
+ else Mc.PEmul (c,var i) in
+ let p' = if Pervasives.(=) p (Mc.PEc Mc.Z0) then mn else
+ Mc.PEadd (mn, p) in
+ xtopoly p' (i+1) l in
+
+ xtopoly (Mc.PEc Mc.Z0) 0 l
let rec fixpoint f x =
let y' = f x in
- if Pervasives.(=) y' x then y'
- else fixpoint f y'
+ if Pervasives.(=) y' x then y'
+ else fixpoint f y'
let rec_simpl_cone n_spec e =
let simpl_cone =
Mc.simpl_cone n_spec.zero n_spec.unit n_spec.mult n_spec.eqb in
let rec rec_simpl_cone = function
- | Mc.PsatzMulE(t1, t2) ->
- simpl_cone (Mc.PsatzMulE (rec_simpl_cone t1, rec_simpl_cone t2))
- | Mc.PsatzAdd(t1,t2) ->
- simpl_cone (Mc.PsatzAdd (rec_simpl_cone t1, rec_simpl_cone t2))
- | x -> simpl_cone x in
- rec_simpl_cone e
-
-
+ | Mc.PsatzMulE(t1, t2) ->
+ simpl_cone (Mc.PsatzMulE (rec_simpl_cone t1, rec_simpl_cone t2))
+ | Mc.PsatzAdd(t1,t2) ->
+ simpl_cone (Mc.PsatzAdd (rec_simpl_cone t1, rec_simpl_cone t2))
+ | x -> simpl_cone x in
+ rec_simpl_cone e
+
+
let simplify_cone n_spec c = fixpoint (rec_simpl_cone n_spec) c
type cone_prod =
- Const of cone
- | Ideal of cone *cone
- | Mult of cone * cone
- | Other of cone
+ Const of cone
+| Ideal of cone *cone
+| Mult of cone * cone
+| Other of cone
and cone = Mc.zWitness
@@ -147,32 +147,32 @@ let factorise_linear_cone c =
let rec cone_list c l =
match c with
- | Mc.PsatzAdd (x,r) -> cone_list r (x::l)
- | _ -> c :: l in
-
+ | Mc.PsatzAdd (x,r) -> cone_list r (x::l)
+ | _ -> c :: l in
+
let factorise c1 c2 =
match c1 , c2 with
- | Mc.PsatzMulC(x,y) , Mc.PsatzMulC(x',y') ->
- if Pervasives.(=) x x' then Some (Mc.PsatzMulC(x, Mc.PsatzAdd(y,y'))) else None
- | Mc.PsatzMulE(x,y) , Mc.PsatzMulE(x',y') ->
- if Pervasives.(=) x x' then Some (Mc.PsatzMulE(x, Mc.PsatzAdd(y,y'))) else None
- | _ -> None in
-
+ | Mc.PsatzMulC(x,y) , Mc.PsatzMulC(x',y') ->
+ if Pervasives.(=) x x' then Some (Mc.PsatzMulC(x, Mc.PsatzAdd(y,y'))) else None
+ | Mc.PsatzMulE(x,y) , Mc.PsatzMulE(x',y') ->
+ if Pervasives.(=) x x' then Some (Mc.PsatzMulE(x, Mc.PsatzAdd(y,y'))) else None
+ | _ -> None in
+
let rec rebuild_cone l pending =
match l with
- | [] -> (match pending with
- | None -> Mc.PsatzZ
- | Some p -> p
- )
- | e::l ->
- (match pending with
- | None -> rebuild_cone l (Some e)
- | Some p -> (match factorise p e with
- | None -> Mc.PsatzAdd(p, rebuild_cone l (Some e))
- | Some f -> rebuild_cone l (Some f) )
- ) in
+ | [] -> (match pending with
+ | None -> Mc.PsatzZ
+ | Some p -> p
+ )
+ | e::l ->
+ (match pending with
+ | None -> rebuild_cone l (Some e)
+ | Some p -> (match factorise p e with
+ | None -> Mc.PsatzAdd(p, rebuild_cone l (Some e))
+ | Some f -> rebuild_cone l (Some f) )
+ ) in
- (rebuild_cone (List.sort Pervasives.compare (cone_list c [])) None)
+ (rebuild_cone (List.sort Pervasives.compare (cone_list c [])) None)
@@ -199,28 +199,28 @@ open Mfourier
let constrain_monomial mn l =
let coeffs = List.fold_left (fun acc p -> (Poly.get mn p)::acc) [] l in
- if Pervasives.(=) mn Monomial.const
- then
- { coeffs = Vect.from_list ((Big_int unit_big_int):: (List.rev coeffs)) ;
- op = Eq ;
- cst = Big_int zero_big_int }
- else
- { coeffs = Vect.from_list ((Big_int zero_big_int):: (List.rev coeffs)) ;
- op = Eq ;
- cst = Big_int zero_big_int }
+ if Pervasives.(=) mn Monomial.const
+ then
+ { coeffs = Vect.from_list ((Big_int unit_big_int):: (List.rev coeffs)) ;
+ op = Eq ;
+ cst = Big_int zero_big_int }
+ else
+ { coeffs = Vect.from_list ((Big_int zero_big_int):: (List.rev coeffs)) ;
+ op = Eq ;
+ cst = Big_int zero_big_int }
-
+
let positivity l =
let rec xpositivity i l =
match l with
- | [] -> []
- | (_,Mc.Equal)::l -> xpositivity (i+1) l
- | (_,_)::l ->
- {coeffs = Vect.update (i+1) (fun _ -> Int 1) Vect.null ;
- op = Ge ;
- cst = Int 0 } :: (xpositivity (i+1) l)
+ | [] -> []
+ | (_,Mc.Equal)::l -> xpositivity (i+1) l
+ | (_,_)::l ->
+ {coeffs = Vect.update (i+1) (fun _ -> Int 1) Vect.null ;
+ op = Ge ;
+ cst = Int 0 } :: (xpositivity (i+1) l)
in
- xpositivity 0 l
+ xpositivity 0 l
let string_of_op = function
@@ -241,23 +241,23 @@ let build_linear_system l =
let monomials =
List.fold_left (fun acc p ->
- Poly.fold (fun m _ acc -> MonSet.add m acc) p acc)
- (MonSet.singleton Monomial.const) l'
+ Poly.fold (fun m _ acc -> MonSet.add m acc) p acc)
+ (MonSet.singleton Monomial.const) l'
in (* For each monomial, compute a constraint *)
let s0 =
MonSet.fold (fun mn res -> (constrain_monomial mn l')::res) monomials [] in
- (* I need at least something strictly positive *)
+ (* I need at least something strictly positive *)
let strict = {
coeffs = Vect.from_list ((Big_int unit_big_int)::
(List.map (fun (x,y) ->
- match y with Mc.Strict ->
- Big_int unit_big_int
- | _ -> Big_int zero_big_int) l));
+ match y with Mc.Strict ->
+ Big_int unit_big_int
+ | _ -> Big_int zero_big_int) l));
op = Ge ; cst = Big_int unit_big_int } in
(* Add the positivity constraint *)
- {coeffs = Vect.from_list ([Big_int unit_big_int]) ;
- op = Ge ;
- cst = Big_int zero_big_int}::(strict::(positivity l)@s0)
+ {coeffs = Vect.from_list ([Big_int unit_big_int]) ;
+ op = Ge ;
+ cst = Big_int zero_big_int}::(strict::(positivity l)@s0)
let big_int_to_z = Ml2C.bigint
@@ -266,32 +266,32 @@ let big_int_to_z = Ml2C.bigint
-- at a lower layer, certificates are using nums... *)
let make_certificate n_spec (cert,li) =
let bint_to_cst = n_spec.bigint_to_number in
- match cert with
- | [] -> failwith "empty_certificate"
- | e::cert' ->
-(* let cst = match compare_big_int e zero_big_int with
- | 0 -> Mc.PsatzZ
- | 1 -> Mc.PsatzC (bint_to_cst e)
- | _ -> failwith "positivity error"
- in *)
- let rec scalar_product cert l =
- match cert with
- | [] -> Mc.PsatzZ
- | c::cert ->
- match l with
- | [] -> failwith "make_certificate(1)"
- | i::l ->
- let r = scalar_product cert l in
- match compare_big_int c zero_big_int with
- | -1 -> Mc.PsatzAdd (
- Mc.PsatzMulC (Mc.Pc ( bint_to_cst c), Mc.PsatzIn (Ml2C.nat i)),
- r)
- | 0 -> r
- | _ -> Mc.PsatzAdd (
- Mc.PsatzMulE (Mc.PsatzC (bint_to_cst c), Mc.PsatzIn (Ml2C.nat i)),
- r) in
- (factorise_linear_cone
- (simplify_cone n_spec (scalar_product cert' li)))
+ match cert with
+ | [] -> failwith "empty_certificate"
+ | e::cert' ->
+ (* let cst = match compare_big_int e zero_big_int with
+ | 0 -> Mc.PsatzZ
+ | 1 -> Mc.PsatzC (bint_to_cst e)
+ | _ -> failwith "positivity error"
+ in *)
+ let rec scalar_product cert l =
+ match cert with
+ | [] -> Mc.PsatzZ
+ | c::cert ->
+ match l with
+ | [] -> failwith "make_certificate(1)"
+ | i::l ->
+ let r = scalar_product cert l in
+ match compare_big_int c zero_big_int with
+ | -1 -> Mc.PsatzAdd (
+ Mc.PsatzMulC (Mc.Pc ( bint_to_cst c), Mc.PsatzIn (Ml2C.nat i)),
+ r)
+ | 0 -> r
+ | _ -> Mc.PsatzAdd (
+ Mc.PsatzMulE (Mc.PsatzC (bint_to_cst c), Mc.PsatzIn (Ml2C.nat i)),
+ r) in
+ (factorise_linear_cone
+ (simplify_cone n_spec (scalar_product cert' li)))
exception Found of Monomial.t
@@ -301,91 +301,154 @@ exception Strict
module MonMap = Map.Make(Monomial)
let primal l =
- let vr = ref 0 in
-
- let vect_of_poly map p =
- Poly.fold (fun mn vl (map,vect) ->
- if Pervasives.(=) mn Monomial.const
- then (map,vect)
- else
- let (mn,m) = try (MonMap.find mn map,map) with Not_found -> let res = (!vr, MonMap.add mn !vr map) in incr vr ; res in
- (m,if Int.equal (sign_num vl) 0 then vect else (mn,vl)::vect)) p (map,[]) in
-
- let op_op = function Mc.NonStrict -> Ge |Mc.Equal -> Eq | _ -> raise Strict in
+ let vr = ref 0 in
+
+ let vect_of_poly map p =
+ Poly.fold (fun mn vl (map,vect) ->
+ if Pervasives.(=) mn Monomial.const
+ then (map,vect)
+ else
+ let (mn,m) = try (MonMap.find mn map,map) with Not_found -> let res = (!vr, MonMap.add mn !vr map) in incr vr ; res in
+ (m,if Int.equal (sign_num vl) 0 then vect else (mn,vl)::vect)) p (map,[]) in
+
+ let op_op = function Mc.NonStrict -> Ge |Mc.Equal -> Eq | _ -> raise Strict in
- let cmp x y = Int.compare (fst x) (fst y) in
+ let cmp x y = Int.compare (fst x) (fst y) in
- snd (List.fold_right (fun (p,op) (map,l) ->
- let (mp,vect) = vect_of_poly map p in
- let cstr = {coeffs = List.sort cmp vect; op = op_op op ; cst = minus_num (Poly.get Monomial.const p)} in
+ snd (List.fold_right (fun (p,op) (map,l) ->
+ let (mp,vect) = vect_of_poly map p in
+ let cstr = {coeffs = List.sort cmp vect; op = op_op op ; cst = minus_num (Poly.get Monomial.const p)} in
- (mp,cstr::l)) l (MonMap.empty,[]))
+ (mp,cstr::l)) l (MonMap.empty,[]))
let dual_raw_certificate (l: (Poly.t * Mc.op1) list) =
-(* List.iter (fun (p,op) -> Printf.fprintf stdout "%a %s 0\n" Poly.pp p (string_of_op op) ) l ; *)
-
+ (* List.iter (fun (p,op) -> Printf.fprintf stdout "%a %s 0\n" Poly.pp p (string_of_op op) ) l ; *)
+
let sys = build_linear_system l in
- try
- match Fourier.find_point sys with
- | Inr _ -> None
- | Inl cert -> Some (rats_to_ints (Vect.to_list cert))
- (* should not use rats_to_ints *)
- with x when Errors.noncritical x ->
- if debug
- then (Printf.printf "raw certificate %s" (Printexc.to_string x);
- flush stdout) ;
- None
+ try
+ match Fourier.find_point sys with
+ | Inr _ -> None
+ | Inl cert -> Some (rats_to_ints (Vect.to_list cert))
+ (* should not use rats_to_ints *)
+ with x when Errors.noncritical x ->
+ if debug
+ then (Printf.printf "raw certificate %s" (Printexc.to_string x);
+ flush stdout) ;
+ None
let raw_certificate l =
- try
- let p = primal l in
- match Fourier.find_point p with
- | Inr prf ->
- if debug then Printf.printf "AProof : %a\n" pp_proof prf ;
- let cert = List.map (fun (x,n) -> x+1,n) (fst (List.hd (Proof.mk_proof p prf))) in
- if debug then Printf.printf "CProof : %a" Vect.pp_vect cert ;
- Some (rats_to_ints (Vect.to_list cert))
- | Inl _ -> None
- with Strict ->
+ try
+ let p = primal l in
+ match Fourier.find_point p with
+ | Inr prf ->
+ if debug then Printf.printf "AProof : %a\n" pp_proof prf ;
+ let cert = List.map (fun (x,n) -> x+1,n) (fst (List.hd (Proof.mk_proof p prf))) in
+ if debug then Printf.printf "CProof : %a" Vect.pp_vect cert ;
+ Some (rats_to_ints (Vect.to_list cert))
+ | Inl _ -> None
+ with Strict ->
(* Fourier elimination should handle > *)
- dual_raw_certificate l
+ dual_raw_certificate l
let simple_linear_prover l =
let (lc,li) = List.split l in
- match raw_certificate lc with
- | None -> None (* No certificate *)
- | Some cert -> Some (cert,li)
-
+ match raw_certificate lc with
+ | None -> None (* No certificate *)
+ | Some cert -> Some (cert,li)
+
let linear_prover n_spec l =
- let build_system n_spec l =
- let li = List.combine l (interval 0 (List.length l -1)) in
- let (l1,l') = List.partition
- (fun (x,_) -> if Pervasives.(=) (snd x) Mc.NonEqual then true else false) li in
- List.map
- (fun ((x,y),i) -> match y with
- Mc.NonEqual -> failwith "cannot happen"
- | y -> ((dev_form n_spec x, y),i)) l' in
- let l' = build_system n_spec l in
- simple_linear_prover (*n_spec*) l'
+ let build_system n_spec l =
+ let li = List.combine l (interval 0 (List.length l -1)) in
+ let (l1,l') = List.partition
+ (fun (x,_) -> if Pervasives.(=) (snd x) Mc.NonEqual then true else false) li in
+ List.map
+ (fun ((x,y),i) -> match y with
+ Mc.NonEqual -> failwith "cannot happen"
+ | y -> ((dev_form n_spec x, y),i)) l' in
+ let l' = build_system n_spec l in
+ simple_linear_prover (*n_spec*) l'
let linear_prover n_spec l =
try linear_prover n_spec l
with x when Errors.noncritical x ->
- (print_string (Printexc.to_string x); None)
+ (print_string (Printexc.to_string x); None)
+
+let compute_max_nb_cstr l d =
+ let len = List.length l in
+ max len (max d (len * d))
+
+let linear_prover_with_cert prfdepth spec l =
+ max_nb_cstr := compute_max_nb_cstr l prfdepth ;
+ match linear_prover spec l with
+ | None -> None
+ | Some cert -> Some (make_certificate spec cert)
+
+let nlinear_prover prfdepth (sys: (Mc.q Mc.pExpr * Mc.op1) list) =
+ LinPoly.MonT.clear ();
+ max_nb_cstr := compute_max_nb_cstr sys prfdepth ;
+ (* Assign a proof to the initial hypotheses *)
+ let sys = mapi (fun c i -> (c,Mc.PsatzIn (Ml2C.nat i))) sys in
+
+
+ (* Add all the product of hypotheses *)
+ let prod = all_pairs (fun ((c,o),p) ((c',o'),p') ->
+ ((Mc.PEmul(c,c') , Mc.opMult o o') , Mc.PsatzMulE(p,p'))) sys in
+
+ (* Only filter those have a meaning *)
+ let prod = List.fold_left (fun l ((c,o),p) ->
+ match o with
+ | None -> l
+ | Some o -> ((c,o),p) :: l) [] prod in
+
+ let sys = sys @ prod in
+
+ let square =
+ (* Collect the squares and state that they are positive *)
+ let pols = List.map (fun ((p,_),_) -> dev_form q_spec p) sys in
+ let square =
+ List.fold_left (fun acc p ->
+ Poly.fold
+ (fun m _ acc ->
+ match Monomial.sqrt m with
+ | None -> acc
+ | Some s -> MonMap.add s m acc) p acc) MonMap.empty pols in
-let linear_prover_with_cert spec l =
- match linear_prover spec l with
- | None -> None
- | Some cert -> Some (make_certificate spec cert)
+ let pol_of_mon m =
+ Monomial.fold (fun x v p -> Mc.PEmul(Mc.PEpow(Mc.PEX(Ml2C.positive x),Ml2C.n v),p)) m (Mc.PEc q_spec.unit) in
+
+ let norm0 =
+ Mc.norm q_spec.zero q_spec.unit Mc.qplus Mc.qmult Mc.qminus Mc.qopp Mc.qeq_bool in
+
+
+ MonMap.fold (fun s m acc -> ((pol_of_mon m , Mc.NonStrict), Mc.PsatzSquare(norm0 (pol_of_mon s)))::acc) square [] in
+
+ let sys = sys @ square in
+
+
+ (* Call the linear prover without the proofs *)
+ let sys_no_prf = List.map fst sys in
+ match linear_prover q_spec sys_no_prf with
+ | None -> None
+ | Some cert ->
+ let cert = make_certificate q_spec cert in
+ let rec map_psatz = function
+ | Mc.PsatzIn n -> snd (List.nth sys (C2Ml.nat n))
+ | Mc.PsatzSquare c -> Mc.PsatzSquare c
+ | Mc.PsatzMulC(c,p) -> Mc.PsatzMulC(c, map_psatz p)
+ | Mc.PsatzMulE(p1,p2) -> Mc.PsatzMulE(map_psatz p1,map_psatz p2)
+ | Mc.PsatzAdd(p1,p2) -> Mc.PsatzAdd(map_psatz p1,map_psatz p2)
+ | Mc.PsatzC c -> Mc.PsatzC c
+ | Mc.PsatzZ -> Mc.PsatzZ in
+ Some (map_psatz cert)
@@ -395,11 +458,11 @@ let make_linear_system l =
(Poly.constant (Int 0)) l' in
let monomials = Poly.fold
(fun mn _ l -> if Pervasives.(=) mn Monomial.const then l else mn::l) monomials [] in
- (List.map (fun (c,op) ->
- {coeffs = Vect.from_list (List.map (fun mn -> (Poly.get mn c)) monomials) ;
- op = op ;
- cst = minus_num ( (Poly.get Monomial.const c))}) l
- ,monomials)
+ (List.map (fun (c,op) ->
+ {coeffs = Vect.from_list (List.map (fun mn -> (Poly.get mn c)) monomials) ;
+ op = op ;
+ cst = minus_num ( (Poly.get Monomial.const c))}) l
+ ,monomials)
let pplus x y = Mc.PEadd(x,y)
@@ -413,7 +476,7 @@ let rec mem p x l =
let rec remove_assoc p x l =
match l with [] -> [] | e::l -> if p x (fst e) then
- remove_assoc p x l else e::(remove_assoc p x l)
+ remove_assoc p x l else e::(remove_assoc p x l)
let eq x y = Int.equal (Vect.compare x y) 0
@@ -424,39 +487,39 @@ let remove e l = List.fold_left (fun l x -> if eq x e then l else x::l) [] l
only searching for naive cutting planes *)
let develop_constraint z_spec (e,k) =
- match k with
- | Mc.NonStrict -> (dev_form z_spec e , Ge)
- | Mc.Equal -> (dev_form z_spec e , Eq)
- | _ -> assert false
+ match k with
+ | Mc.NonStrict -> (dev_form z_spec e , Ge)
+ | Mc.Equal -> (dev_form z_spec e , Eq)
+ | _ -> assert false
let op_of_op_compat = function
- | Ge -> Mc.NonStrict
- | Eq -> Mc.Equal
+ | Ge -> Mc.NonStrict
+ | Eq -> Mc.Equal
let integer_vector coeffs =
- let vars , coeffs = List.split coeffs in
- List.combine vars (List.map (fun x -> Big_int x) (rats_to_ints coeffs))
+ let vars , coeffs = List.split coeffs in
+ List.combine vars (List.map (fun x -> Big_int x) (rats_to_ints coeffs))
let integer_cstr {coeffs = coeffs ; op = op ; cst = cst } =
- let vars , coeffs = List.split coeffs in
- match rats_to_ints (cst::coeffs) with
- | cst :: coeffs ->
- {
- coeffs = List.combine vars (List.map (fun x -> Big_int x) coeffs) ;
- op = op ; cst = Big_int cst}
- | _ -> assert false
-
+ let vars , coeffs = List.split coeffs in
+ match rats_to_ints (cst::coeffs) with
+ | cst :: coeffs ->
+ {
+ coeffs = List.combine vars (List.map (fun x -> Big_int x) coeffs) ;
+ op = op ; cst = Big_int cst}
+ | _ -> assert false
+
let pexpr_of_cstr_compat var cstr =
- let {coeffs = coeffs ; op = op ; cst = cst } = integer_cstr cstr in
- try
- let expr = list_to_polynomial var (Vect.to_list coeffs) in
- let d = Ml2C.bigint (denominator cst) in
- let n = Ml2C.bigint (numerator cst) in
- (pplus (pmult (pconst d) expr) (popp (pconst n)), op_of_op_compat op)
- with Failure _ -> failwith "pexpr_of_cstr_compat"
+ let {coeffs = coeffs ; op = op ; cst = cst } = integer_cstr cstr in
+ try
+ let expr = list_to_polynomial var (Vect.to_list coeffs) in
+ let d = Ml2C.bigint (denominator cst) in
+ let n = Ml2C.bigint (numerator cst) in
+ (pplus (pmult (pconst d) expr) (popp (pconst n)), op_of_op_compat op)
+ with Failure _ -> failwith "pexpr_of_cstr_compat"
@@ -465,41 +528,41 @@ open Sos_types
let rec scale_term t =
match t with
- | Zero -> unit_big_int , Zero
- | Const n -> (denominator n) , Const (Big_int (numerator n))
- | Var n -> unit_big_int , Var n
- | Inv _ -> failwith "scale_term : not implemented"
- | Opp t -> let s, t = scale_term t in s, Opp t
- | Add(t1,t2) -> let s1,y1 = scale_term t1 and s2,y2 = scale_term t2 in
- let g = gcd_big_int s1 s2 in
- let s1' = div_big_int s1 g in
- let s2' = div_big_int s2 g in
- let e = mult_big_int g (mult_big_int s1' s2') in
- if Int.equal (compare_big_int e unit_big_int) 0
- then (unit_big_int, Add (y1,y2))
- else e, Add (Mul(Const (Big_int s2'), y1),
- Mul (Const (Big_int s1'), y2))
- | Sub _ -> failwith "scale term: not implemented"
- | Mul(y,z) -> let s1,y1 = scale_term y and s2,y2 = scale_term z in
- mult_big_int s1 s2 , Mul (y1, y2)
- | Pow(t,n) -> let s,t = scale_term t in
- power_big_int_positive_int s n , Pow(t,n)
- | _ -> failwith "scale_term : not implemented"
+ | Zero -> unit_big_int , Zero
+ | Const n -> (denominator n) , Const (Big_int (numerator n))
+ | Var n -> unit_big_int , Var n
+ | Inv _ -> failwith "scale_term : not implemented"
+ | Opp t -> let s, t = scale_term t in s, Opp t
+ | Add(t1,t2) -> let s1,y1 = scale_term t1 and s2,y2 = scale_term t2 in
+ let g = gcd_big_int s1 s2 in
+ let s1' = div_big_int s1 g in
+ let s2' = div_big_int s2 g in
+ let e = mult_big_int g (mult_big_int s1' s2') in
+ if Int.equal (compare_big_int e unit_big_int) 0
+ then (unit_big_int, Add (y1,y2))
+ else e, Add (Mul(Const (Big_int s2'), y1),
+ Mul (Const (Big_int s1'), y2))
+ | Sub _ -> failwith "scale term: not implemented"
+ | Mul(y,z) -> let s1,y1 = scale_term y and s2,y2 = scale_term z in
+ mult_big_int s1 s2 , Mul (y1, y2)
+ | Pow(t,n) -> let s,t = scale_term t in
+ power_big_int_positive_int s n , Pow(t,n)
+ | _ -> failwith "scale_term : not implemented"
let scale_term t =
let (s,t') = scale_term t in
- s,t'
+ s,t'
let get_index_of_ith_match f i l =
let rec get j res l =
match l with
- | [] -> failwith "bad index"
- | e::l -> if f e
- then
- (if Int.equal j i then res else get (j+1) (res+1) l )
- else get j (res+1) l in
- get 0 0 l
+ | [] -> failwith "bad index"
+ | e::l -> if f e
+ then
+ (if Int.equal j i then res else get (j+1) (res+1) l )
+ else get j (res+1) l in
+ get 0 0 l
let rec scale_certificate pos = match pos with
@@ -511,97 +574,97 @@ let rec scale_certificate pos = match pos with
| Rational_le n -> (denominator n) , Rational_le (Big_int (numerator n))
| Rational_lt n -> (denominator n) , Rational_lt (Big_int (numerator n))
| Square t -> let s,t' = scale_term t in
- mult_big_int s s , Square t'
+ mult_big_int s s , Square t'
| Eqmul (t, y) -> let s1,y1 = scale_term t and s2,y2 = scale_certificate y in
- mult_big_int s1 s2 , Eqmul (y1,y2)
+ mult_big_int s1 s2 , Eqmul (y1,y2)
| Sum (y, z) -> let s1,y1 = scale_certificate y
- and s2,y2 = scale_certificate z in
- let g = gcd_big_int s1 s2 in
- let s1' = div_big_int s1 g in
- let s2' = div_big_int s2 g in
- mult_big_int g (mult_big_int s1' s2'),
- Sum (Product(Rational_le (Big_int s2'), y1),
- Product (Rational_le (Big_int s1'), y2))
+ and s2,y2 = scale_certificate z in
+ let g = gcd_big_int s1 s2 in
+ let s1' = div_big_int s1 g in
+ let s2' = div_big_int s2 g in
+ mult_big_int g (mult_big_int s1' s2'),
+ Sum (Product(Rational_le (Big_int s2'), y1),
+ Product (Rational_le (Big_int s1'), y2))
| Product (y, z) ->
- let s1,y1 = scale_certificate y and s2,y2 = scale_certificate z in
- mult_big_int s1 s2 , Product (y1,y2)
+ let s1,y1 = scale_certificate y and s2,y2 = scale_certificate z in
+ mult_big_int s1 s2 , Product (y1,y2)
open Micromega
- let rec term_to_q_expr = function
- | Const n -> PEc (Ml2C.q n)
- | Zero -> PEc ( Ml2C.q (Int 0))
- | Var s -> PEX (Ml2C.index
- (int_of_string (String.sub s 1 (String.length s - 1))))
- | Mul(p1,p2) -> PEmul(term_to_q_expr p1, term_to_q_expr p2)
- | Add(p1,p2) -> PEadd(term_to_q_expr p1, term_to_q_expr p2)
- | Opp p -> PEopp (term_to_q_expr p)
- | Pow(t,n) -> PEpow (term_to_q_expr t,Ml2C.n n)
- | Sub(t1,t2) -> PEsub (term_to_q_expr t1, term_to_q_expr t2)
- | _ -> failwith "term_to_q_expr: not implemented"
-
- let term_to_q_pol e = Mc.norm_aux (Ml2C.q (Int 0)) (Ml2C.q (Int 1)) Mc.qplus Mc.qmult Mc.qminus Mc.qopp Mc.qeq_bool (term_to_q_expr e)
-
-
- let rec product l =
- match l with
- | [] -> Mc.PsatzZ
- | [i] -> Mc.PsatzIn (Ml2C.nat i)
- | i ::l -> Mc.PsatzMulE(Mc.PsatzIn (Ml2C.nat i), product l)
+let rec term_to_q_expr = function
+ | Const n -> PEc (Ml2C.q n)
+ | Zero -> PEc ( Ml2C.q (Int 0))
+ | Var s -> PEX (Ml2C.index
+ (int_of_string (String.sub s 1 (String.length s - 1))))
+ | Mul(p1,p2) -> PEmul(term_to_q_expr p1, term_to_q_expr p2)
+ | Add(p1,p2) -> PEadd(term_to_q_expr p1, term_to_q_expr p2)
+ | Opp p -> PEopp (term_to_q_expr p)
+ | Pow(t,n) -> PEpow (term_to_q_expr t,Ml2C.n n)
+ | Sub(t1,t2) -> PEsub (term_to_q_expr t1, term_to_q_expr t2)
+ | _ -> failwith "term_to_q_expr: not implemented"
+
+let term_to_q_pol e = Mc.norm_aux (Ml2C.q (Int 0)) (Ml2C.q (Int 1)) Mc.qplus Mc.qmult Mc.qminus Mc.qopp Mc.qeq_bool (term_to_q_expr e)
+
+
+let rec product l =
+ match l with
+ | [] -> Mc.PsatzZ
+ | [i] -> Mc.PsatzIn (Ml2C.nat i)
+ | i ::l -> Mc.PsatzMulE(Mc.PsatzIn (Ml2C.nat i), product l)
let q_cert_of_pos pos =
let rec _cert_of_pos = function
- Axiom_eq i -> Mc.PsatzIn (Ml2C.nat i)
+ Axiom_eq i -> Mc.PsatzIn (Ml2C.nat i)
| Axiom_le i -> Mc.PsatzIn (Ml2C.nat i)
| Axiom_lt i -> Mc.PsatzIn (Ml2C.nat i)
| Monoid l -> product l
| Rational_eq n | Rational_le n | Rational_lt n ->
- if Int.equal (compare_num n (Int 0)) 0 then Mc.PsatzZ else
- Mc.PsatzC (Ml2C.q n)
+ if Int.equal (compare_num n (Int 0)) 0 then Mc.PsatzZ else
+ Mc.PsatzC (Ml2C.q n)
| Square t -> Mc.PsatzSquare (term_to_q_pol t)
| Eqmul (t, y) -> Mc.PsatzMulC(term_to_q_pol t, _cert_of_pos y)
| Sum (y, z) -> Mc.PsatzAdd (_cert_of_pos y, _cert_of_pos z)
| Product (y, z) -> Mc.PsatzMulE (_cert_of_pos y, _cert_of_pos z) in
- simplify_cone q_spec (_cert_of_pos pos)
+ simplify_cone q_spec (_cert_of_pos pos)
- let rec term_to_z_expr = function
- | Const n -> PEc (Ml2C.bigint (big_int_of_num n))
- | Zero -> PEc ( Z0)
- | Var s -> PEX (Ml2C.index
- (int_of_string (String.sub s 1 (String.length s - 1))))
- | Mul(p1,p2) -> PEmul(term_to_z_expr p1, term_to_z_expr p2)
- | Add(p1,p2) -> PEadd(term_to_z_expr p1, term_to_z_expr p2)
- | Opp p -> PEopp (term_to_z_expr p)
- | Pow(t,n) -> PEpow (term_to_z_expr t,Ml2C.n n)
- | Sub(t1,t2) -> PEsub (term_to_z_expr t1, term_to_z_expr t2)
- | _ -> failwith "term_to_z_expr: not implemented"
+let rec term_to_z_expr = function
+ | Const n -> PEc (Ml2C.bigint (big_int_of_num n))
+ | Zero -> PEc ( Z0)
+ | Var s -> PEX (Ml2C.index
+ (int_of_string (String.sub s 1 (String.length s - 1))))
+ | Mul(p1,p2) -> PEmul(term_to_z_expr p1, term_to_z_expr p2)
+ | Add(p1,p2) -> PEadd(term_to_z_expr p1, term_to_z_expr p2)
+ | Opp p -> PEopp (term_to_z_expr p)
+ | Pow(t,n) -> PEpow (term_to_z_expr t,Ml2C.n n)
+ | Sub(t1,t2) -> PEsub (term_to_z_expr t1, term_to_z_expr t2)
+ | _ -> failwith "term_to_z_expr: not implemented"
- let term_to_z_pol e = Mc.norm_aux (Ml2C.z 0) (Ml2C.z 1) Mc.Z.add Mc.Z.mul Mc.Z.sub Mc.Z.opp Mc.zeq_bool (term_to_z_expr e)
+let term_to_z_pol e = Mc.norm_aux (Ml2C.z 0) (Ml2C.z 1) Mc.Z.add Mc.Z.mul Mc.Z.sub Mc.Z.opp Mc.zeq_bool (term_to_z_expr e)
let z_cert_of_pos pos =
let s,pos = (scale_certificate pos) in
let rec _cert_of_pos = function
- Axiom_eq i -> Mc.PsatzIn (Ml2C.nat i)
+ Axiom_eq i -> Mc.PsatzIn (Ml2C.nat i)
| Axiom_le i -> Mc.PsatzIn (Ml2C.nat i)
| Axiom_lt i -> Mc.PsatzIn (Ml2C.nat i)
| Monoid l -> product l
| Rational_eq n | Rational_le n | Rational_lt n ->
- if Int.equal (compare_num n (Int 0)) 0 then Mc.PsatzZ else
- Mc.PsatzC (Ml2C.bigint (big_int_of_num n))
+ if Int.equal (compare_num n (Int 0)) 0 then Mc.PsatzZ else
+ Mc.PsatzC (Ml2C.bigint (big_int_of_num n))
| Square t -> Mc.PsatzSquare (term_to_z_pol t)
| Eqmul (t, y) ->
- let is_unit =
- match t with
- | Const n -> n =/ Int 1
- | _ -> false in
- if is_unit
- then _cert_of_pos y
- else Mc.PsatzMulC(term_to_z_pol t, _cert_of_pos y)
+ let is_unit =
+ match t with
+ | Const n -> n =/ Int 1
+ | _ -> false in
+ if is_unit
+ then _cert_of_pos y
+ else Mc.PsatzMulC(term_to_z_pol t, _cert_of_pos y)
| Sum (y, z) -> Mc.PsatzAdd (_cert_of_pos y, _cert_of_pos z)
| Product (y, z) -> Mc.PsatzMulE (_cert_of_pos y, _cert_of_pos z) in
- simplify_cone z_spec (_cert_of_pos pos)
+ simplify_cone z_spec (_cert_of_pos pos)
(** All constraints (initial or derived) have an index and have a justification i.e., proof.
Given a constraint, all the coefficients are always integers.
@@ -612,116 +675,109 @@ open Num
open Big_int
open Polynomial
-(*module Mc = Micromega*)
-(*module Ml2C = Mutils.CamlToCoq
-module C2Ml = Mutils.CoqToCaml
-*)
-let debug = false
-
-
module Env =
struct
- type t = int list
+ type t = int list
- let id_of_hyp hyp l =
- let rec xid_of_hyp i l =
- match l with
- | [] -> failwith "id_of_hyp"
- | hyp'::l -> if Pervasives.(=) hyp hyp' then i else xid_of_hyp (i+1) l in
- xid_of_hyp 0 l
+ let id_of_hyp hyp l =
+ let rec xid_of_hyp i l =
+ match l with
+ | [] -> failwith "id_of_hyp"
+ | hyp'::l -> if Pervasives.(=) hyp hyp' then i else xid_of_hyp (i+1) l in
+ xid_of_hyp 0 l
end
let coq_poly_of_linpol (p,c) =
- let pol_of_mon m =
- Monomial.fold (fun x v p -> Mc.PEmul(Mc.PEpow(Mc.PEX(Ml2C.positive x),Ml2C.n v),p)) m (Mc.PEc (Mc.Zpos Mc.XH)) in
+ let pol_of_mon m =
+ Monomial.fold (fun x v p -> Mc.PEmul(Mc.PEpow(Mc.PEX(Ml2C.positive x),Ml2C.n v),p)) m (Mc.PEc (Mc.Zpos Mc.XH)) in
- List.fold_left (fun acc (x,v) ->
- let mn = LinPoly.MonT.retrieve x in
- Mc.PEadd(Mc.PEmul(Mc.PEc (Ml2C.bigint (numerator v)), pol_of_mon mn),acc)) (Mc.PEc (Ml2C.bigint (numerator c))) p
-
+ List.fold_left (fun acc (x,v) ->
+ let mn = LinPoly.MonT.retrieve x in
+ Mc.PEadd(Mc.PEmul(Mc.PEc (Ml2C.bigint (numerator v)), pol_of_mon mn),acc)) (Mc.PEc (Ml2C.bigint (numerator c))) p
+
let rec cmpl_prf_rule env = function
- | Hyp i | Def i -> Mc.PsatzIn (Ml2C.nat (Env.id_of_hyp i env))
- | Cst i -> Mc.PsatzC (Ml2C.bigint i)
- | Zero -> Mc.PsatzZ
- | MulPrf(p1,p2) -> Mc.PsatzMulE(cmpl_prf_rule env p1, cmpl_prf_rule env p2)
- | AddPrf(p1,p2) -> Mc.PsatzAdd(cmpl_prf_rule env p1 , cmpl_prf_rule env p2)
- | MulC(lp,p) -> let lp = Mc.norm0 (coq_poly_of_linpol lp) in
- Mc.PsatzMulC(lp,cmpl_prf_rule env p)
- | Square lp -> Mc.PsatzSquare (Mc.norm0 (coq_poly_of_linpol lp))
- | _ -> failwith "Cuts should already be compiled"
-
+ | Hyp i | Def i -> Mc.PsatzIn (Ml2C.nat (Env.id_of_hyp i env))
+ | Cst i -> Mc.PsatzC (Ml2C.bigint i)
+ | Zero -> Mc.PsatzZ
+ | MulPrf(p1,p2) -> Mc.PsatzMulE(cmpl_prf_rule env p1, cmpl_prf_rule env p2)
+ | AddPrf(p1,p2) -> Mc.PsatzAdd(cmpl_prf_rule env p1 , cmpl_prf_rule env p2)
+ | MulC(lp,p) -> let lp = Mc.norm0 (coq_poly_of_linpol lp) in
+ Mc.PsatzMulC(lp,cmpl_prf_rule env p)
+ | Square lp -> Mc.PsatzSquare (Mc.norm0 (coq_poly_of_linpol lp))
+ | _ -> failwith "Cuts should already be compiled"
+
let rec cmpl_proof env = function
- | Done -> Mc.DoneProof
- | Step(i,p,prf) ->
- begin
- match p with
- | CutPrf p' ->
- Mc.CutProof(cmpl_prf_rule env p', cmpl_proof (i::env) prf)
- | _ -> Mc.RatProof(cmpl_prf_rule env p,cmpl_proof (i::env) prf)
- end
- | Enum(i,p1,_,p2,l) ->
- Mc.EnumProof(cmpl_prf_rule env p1,cmpl_prf_rule env p2,List.map (cmpl_proof (i::env)) l)
+ | Done -> Mc.DoneProof
+ | Step(i,p,prf) ->
+ begin
+ match p with
+ | CutPrf p' ->
+ Mc.CutProof(cmpl_prf_rule env p', cmpl_proof (i::env) prf)
+ | _ -> Mc.RatProof(cmpl_prf_rule env p,cmpl_proof (i::env) prf)
+ end
+ | Enum(i,p1,_,p2,l) ->
+ Mc.EnumProof(cmpl_prf_rule env p1,cmpl_prf_rule env p2,List.map (cmpl_proof (i::env)) l)
let compile_proof env prf =
- let id = 1 + proof_max_id prf in
- let _,prf = normalise_proof id prf in
- if debug then Printf.fprintf stdout "compiled proof %a\n" output_proof prf;
- cmpl_proof env prf
+ let id = 1 + proof_max_id prf in
+ let _,prf = normalise_proof id prf in
+ if debug then Printf.fprintf stdout "compiled proof %a\n" output_proof prf;
+ cmpl_proof env prf
type prf_sys = (cstr_compat * prf_rule) list
let xlinear_prover sys =
- match Fourier.find_point sys with
- | Inr prf ->
- if debug then Printf.printf "AProof : %a\n" pp_proof prf ;
- let cert = (*List.map (fun (x,n) -> x+1,n)*) (fst (List.hd (Proof.mk_proof sys prf))) in
- if debug then Printf.printf "CProof : %a" Vect.pp_vect cert ;
- Some (rats_to_ints (Vect.to_list cert))
- | Inl _ -> None
+ match Fourier.find_point sys with
+ | Inr prf ->
+ if debug then Printf.printf "AProof : %a\n" pp_proof prf ;
+ let cert = (*List.map (fun (x,n) -> x+1,n)*) (fst (List.hd (Proof.mk_proof sys prf))) in
+ if debug then Printf.printf "CProof : %a" Vect.pp_vect cert ;
+ Some (rats_to_ints (Vect.to_list cert))
+ | Inl _ -> None
let output_num o n = output_string o (string_of_num n)
let output_bigint o n = output_string o (string_of_big_int n)
let proof_of_farkas prf cert =
-(* Printf.printf "\nproof_of_farkas %a , %a \n" (pp_list output_prf_rule) prf (pp_list output_bigint) cert ; *)
- let rec mk_farkas acc prf cert =
- match prf, cert with
- | _ , [] -> acc
- | [] , _ -> failwith "proof_of_farkas : not enough hyps"
- | p::prf,c::cert ->
- mk_farkas (add_proof (mul_proof c p) acc) prf cert in
- let res = mk_farkas Zero prf cert in
+ (* Printf.printf "\nproof_of_farkas %a , %a \n" (pp_list output_prf_rule) prf (pp_list output_bigint) cert ; *)
+ let rec mk_farkas acc prf cert =
+ match prf, cert with
+ | _ , [] -> acc
+ | [] , _ -> failwith "proof_of_farkas : not enough hyps"
+ | p::prf,c::cert ->
+ mk_farkas (add_proof (mul_proof c p) acc) prf cert in
+ let res = mk_farkas Zero prf cert in
(*Printf.printf "==> %a" output_prf_rule res ; *)
- res
+ res
let linear_prover sys =
- let (sysi,prfi) = List.split sys in
- match xlinear_prover sysi with
- | None -> None
- | Some cert -> Some (proof_of_farkas prfi cert)
+ let (sysi,prfi) = List.split sys in
+ match xlinear_prover sysi with
+ | None -> None
+ | Some cert -> Some (proof_of_farkas prfi cert)
let linear_prover =
- if debug
- then
- fun sys ->
- Printf.printf "<linear_prover"; flush stdout ;
- let res = linear_prover sys in
- Printf.printf ">"; flush stdout ;
- res
- else linear_prover
+ if debug
+ then
+ fun sys ->
+ Printf.printf "<linear_prover"; flush stdout ;
+ let res = linear_prover sys in
+ Printf.printf ">"; flush stdout ;
+ res
+ else linear_prover
@@ -733,11 +789,11 @@ let linear_prover =
*)
type checksat =
- | Tauto (* Tautology *)
- | Unsat of prf_rule (* Unsatisfiable *)
- | Cut of cstr_compat * prf_rule (* Cutting plane *)
- | Normalise of cstr_compat * prf_rule (* coefficients are relatively prime *)
-
+| Tauto (* Tautology *)
+| Unsat of prf_rule (* Unsatisfiable *)
+| Cut of cstr_compat * prf_rule (* Cutting plane *)
+| Normalise of cstr_compat * prf_rule (* coefficients are relatively prime *)
+
(** [check_sat]
- detects constraints that are not satisfiable;
@@ -745,83 +801,83 @@ type checksat =
*)
let check_sat (cstr,prf) =
- let {coeffs=coeffs ; op=op ; cst=cst} = cstr in
- match coeffs with
- | [] ->
- if eval_op op (Int 0) cst then Tauto else Unsat prf
- | _ ->
- let gcdi = (gcd_list (List.map snd coeffs)) in
- let gcd = Big_int gcdi in
- if eq_num gcd (Int 1)
- then Normalise(cstr,prf)
- else
- if Int.equal (sign_num (mod_num cst gcd)) 0
- then (* We can really normalise *)
- begin
- assert (sign_num gcd >=1 ) ;
- let cstr = {
- coeffs = List.map (fun (x,v) -> (x, v // gcd)) coeffs;
- op = op ; cst = cst // gcd
- } in
- Normalise(cstr,Gcd(gcdi,prf))
- (* Normalise(cstr,CutPrf prf)*)
- end
- else
- match op with
- | Eq -> Unsat (CutPrf prf)
- | Ge ->
- let cstr = {
- coeffs = List.map (fun (x,v) -> (x, v // gcd)) coeffs;
- op = op ; cst = ceiling_num (cst // gcd)
- } in Cut(cstr,CutPrf prf)
+ let {coeffs=coeffs ; op=op ; cst=cst} = cstr in
+ match coeffs with
+ | [] ->
+ if eval_op op (Int 0) cst then Tauto else Unsat prf
+ | _ ->
+ let gcdi = (gcd_list (List.map snd coeffs)) in
+ let gcd = Big_int gcdi in
+ if eq_num gcd (Int 1)
+ then Normalise(cstr,prf)
+ else
+ if Int.equal (sign_num (mod_num cst gcd)) 0
+ then (* We can really normalise *)
+ begin
+ assert (sign_num gcd >=1 ) ;
+ let cstr = {
+ coeffs = List.map (fun (x,v) -> (x, v // gcd)) coeffs;
+ op = op ; cst = cst // gcd
+ } in
+ Normalise(cstr,Gcd(gcdi,prf))
+ (* Normalise(cstr,CutPrf prf)*)
+ end
+ else
+ match op with
+ | Eq -> Unsat (CutPrf prf)
+ | Ge ->
+ let cstr = {
+ coeffs = List.map (fun (x,v) -> (x, v // gcd)) coeffs;
+ op = op ; cst = ceiling_num (cst // gcd)
+ } in Cut(cstr,CutPrf prf)
(** Proof generating pivoting over variable v *)
let pivot v (c1,p1) (c2,p2) =
- let {coeffs = v1 ; op = op1 ; cst = n1} = c1
- and {coeffs = v2 ; op = op2 ; cst = n2} = c2 in
+ let {coeffs = v1 ; op = op1 ; cst = n1} = c1
+ and {coeffs = v2 ; op = op2 ; cst = n2} = c2 in
(* Could factorise gcd... *)
- let xpivot cv1 cv2 =
- (
- {coeffs = Vect.add (Vect.mul cv1 v1) (Vect.mul cv2 v2) ;
- op = Proof.add_op op1 op2 ;
- cst = n1 */ cv1 +/ n2 */ cv2 },
+ let xpivot cv1 cv2 =
+ (
+ {coeffs = Vect.add (Vect.mul cv1 v1) (Vect.mul cv2 v2) ;
+ op = Proof.add_op op1 op2 ;
+ cst = n1 */ cv1 +/ n2 */ cv2 },
- AddPrf(mul_proof (numerator cv1) p1,mul_proof (numerator cv2) p2)) in
+ AddPrf(mul_proof (numerator cv1) p1,mul_proof (numerator cv2) p2)) in
+
+ match Vect.get v v1 , Vect.get v v2 with
+ | None , _ | _ , None -> None
+ | Some a , Some b ->
+ if Int.equal ((sign_num a) * (sign_num b)) (-1)
+ then
+ let cv1 = abs_num b
+ and cv2 = abs_num a in
+ Some (xpivot cv1 cv2)
+ else
+ if op1 == Eq
+ then
+ let cv1 = minus_num (b */ (Int (sign_num a)))
+ and cv2 = abs_num a in
+ Some (xpivot cv1 cv2)
+ else if op2 == Eq
+ then
+ let cv1 = abs_num b
+ and cv2 = minus_num (a */ (Int (sign_num b))) in
+ Some (xpivot cv1 cv2)
+ else None (* op2 could be Eq ... this might happen *)
- match Vect.get v v1 , Vect.get v v2 with
- | None , _ | _ , None -> None
- | Some a , Some b ->
- if Int.equal ((sign_num a) * (sign_num b)) (-1)
- then
- let cv1 = abs_num b
- and cv2 = abs_num a in
- Some (xpivot cv1 cv2)
- else
- if op1 == Eq
- then
- let cv1 = minus_num (b */ (Int (sign_num a)))
- and cv2 = abs_num a in
- Some (xpivot cv1 cv2)
- else if op2 == Eq
- then
- let cv1 = abs_num b
- and cv2 = minus_num (a */ (Int (sign_num b))) in
- Some (xpivot cv1 cv2)
- else None (* op2 could be Eq ... this might happen *)
-
exception FoundProof of prf_rule
let simpl_sys sys =
- List.fold_left (fun acc (c,p) ->
- match check_sat (c,p) with
- | Tauto -> acc
- | Unsat prf -> raise (FoundProof prf)
- | Cut(c,p) -> (c,p)::acc
- | Normalise (c,p) -> (c,p)::acc) [] sys
+ List.fold_left (fun acc (c,p) ->
+ match check_sat (c,p) with
+ | Tauto -> acc
+ | Unsat prf -> raise (FoundProof prf)
+ | Cut(c,p) -> (c,p)::acc
+ | Normalise (c,p) -> (c,p)::acc) [] sys
(** [ext_gcd a b] is the extended Euclid algorithm.
@@ -829,77 +885,77 @@ let simpl_sys sys =
Source: http://en.wikipedia.org/wiki/Extended_Euclidean_algorithm
*)
let rec ext_gcd a b =
- if Int.equal (sign_big_int b) 0
- then (unit_big_int,zero_big_int)
- else
- let (q,r) = quomod_big_int a b in
- let (s,t) = ext_gcd b r in
- (t, sub_big_int s (mult_big_int q t))
+ if Int.equal (sign_big_int b) 0
+ then (unit_big_int,zero_big_int)
+ else
+ let (q,r) = quomod_big_int a b in
+ let (s,t) = ext_gcd b r in
+ (t, sub_big_int s (mult_big_int q t))
let pp_ext_gcd a b =
- let a' = big_int_of_int a in
- let b' = big_int_of_int b in
-
- let (x,y) = ext_gcd a' b' in
- Printf.fprintf stdout "%s * %s + %s * %s = %s\n"
- (string_of_big_int x) (string_of_big_int a')
- (string_of_big_int y) (string_of_big_int b')
- (string_of_big_int (add_big_int (mult_big_int x a') (mult_big_int y b')))
+ let a' = big_int_of_int a in
+ let b' = big_int_of_int b in
+
+ let (x,y) = ext_gcd a' b' in
+ Printf.fprintf stdout "%s * %s + %s * %s = %s\n"
+ (string_of_big_int x) (string_of_big_int a')
+ (string_of_big_int y) (string_of_big_int b')
+ (string_of_big_int (add_big_int (mult_big_int x a') (mult_big_int y b')))
exception Result of (int * (proof * cstr_compat))
let split_equations psys =
- List.partition (fun (c,p) -> c.op == Eq)
+ List.partition (fun (c,p) -> c.op == Eq)
let extract_coprime (c1,p1) (c2,p2) =
- let rec exist2 vect1 vect2 =
- match vect1 , vect2 with
- | _ , [] | [], _ -> None
- | (v1,n1)::vect1' , (v2, n2) :: vect2' ->
- if Pervasives.(=) v1 v2
- then
- if Int.equal (compare_big_int (gcd_big_int (numerator n1) (numerator n2)) unit_big_int) 0
- then Some (v1,n1,n2)
- else
- exist2 vect1' vect2'
- else
- if v1 < v2
- then exist2 vect1' vect2
- else exist2 vect1 vect2' in
-
- if c1.op == Eq && c2.op == Eq
- then exist2 c1.coeffs c2.coeffs
- else None
+ let rec exist2 vect1 vect2 =
+ match vect1 , vect2 with
+ | _ , [] | [], _ -> None
+ | (v1,n1)::vect1' , (v2, n2) :: vect2' ->
+ if Pervasives.(=) v1 v2
+ then
+ if Int.equal (compare_big_int (gcd_big_int (numerator n1) (numerator n2)) unit_big_int) 0
+ then Some (v1,n1,n2)
+ else
+ exist2 vect1' vect2'
+ else
+ if v1 < v2
+ then exist2 vect1' vect2
+ else exist2 vect1 vect2' in
+
+ if c1.op == Eq && c2.op == Eq
+ then exist2 c1.coeffs c2.coeffs
+ else None
let extract2 pred l =
- let rec xextract2 rl l =
- match l with
- | [] -> (None,rl) (* Did not find *)
- | e::l ->
- match extract (pred e) l with
- | None,_ -> xextract2 (e::rl) l
- | Some (r,e'),l' -> Some (r,e,e'), List.rev_append rl l' in
-
- xextract2 [] l
+ let rec xextract2 rl l =
+ match l with
+ | [] -> (None,rl) (* Did not find *)
+ | e::l ->
+ match extract (pred e) l with
+ | None,_ -> xextract2 (e::rl) l
+ | Some (r,e'),l' -> Some (r,e,e'), List.rev_append rl l' in
+
+ xextract2 [] l
let extract_coprime_equation psys =
- extract2 extract_coprime psys
+ extract2 extract_coprime psys
let apply_and_normalise f psys =
- List.fold_left (fun acc pc' ->
- match f pc' with
- | None -> pc'::acc
- | Some pc' ->
- match check_sat pc' with
- | Tauto -> acc
- | Unsat prf -> raise (FoundProof prf)
- | Cut(c,p) -> (c,p)::acc
- | Normalise (c,p) -> (c,p)::acc
- ) [] psys
+ List.fold_left (fun acc pc' ->
+ match f pc' with
+ | None -> pc'::acc
+ | Some pc' ->
+ match check_sat pc' with
+ | Tauto -> acc
+ | Unsat prf -> raise (FoundProof prf)
+ | Cut(c,p) -> (c,p)::acc
+ | Normalise (c,p) -> (c,p)::acc
+ ) [] psys
@@ -908,314 +964,317 @@ let pivot_sys v pc psys = apply_and_normalise (pivot v pc) psys
let reduce_coprime psys =
- let oeq,sys = extract_coprime_equation psys in
- match oeq with
- | None -> None (* Nothing to do *)
- | Some((v,n1,n2),(c1,p1),(c2,p2) ) ->
- let (l1,l2) = ext_gcd (numerator n1) (numerator n2) in
- let l1' = Big_int l1 and l2' = Big_int l2 in
- let cstr =
- {coeffs = Vect.add (Vect.mul l1' c1.coeffs) (Vect.mul l2' c2.coeffs);
- op = Eq ;
- cst = (l1' */ c1.cst) +/ (l2' */ c2.cst)
- } in
- let prf = add_proof (mul_proof (numerator l1') p1) (mul_proof (numerator l2') p2) in
-
- Some (pivot_sys v (cstr,prf) ((c1,p1)::sys))
+ let oeq,sys = extract_coprime_equation psys in
+ match oeq with
+ | None -> None (* Nothing to do *)
+ | Some((v,n1,n2),(c1,p1),(c2,p2) ) ->
+ let (l1,l2) = ext_gcd (numerator n1) (numerator n2) in
+ let l1' = Big_int l1 and l2' = Big_int l2 in
+ let cstr =
+ {coeffs = Vect.add (Vect.mul l1' c1.coeffs) (Vect.mul l2' c2.coeffs);
+ op = Eq ;
+ cst = (l1' */ c1.cst) +/ (l2' */ c2.cst)
+ } in
+ let prf = add_proof (mul_proof (numerator l1') p1) (mul_proof (numerator l2') p2) in
+
+ Some (pivot_sys v (cstr,prf) ((c1,p1)::sys))
(** If there is an equation [eq] of the form 1.x + e = c, do a pivot over x with equation [eq] *)
let reduce_unary psys =
- let is_unary_equation (cstr,prf) =
- if cstr.op == Eq
- then
- try
- Some (fst (List.find (fun (_,n) -> n =/ (Int 1) || n=/ (Int (-1))) cstr.coeffs))
- with Not_found -> None
- else None in
-
- let (oeq,sys) = extract is_unary_equation psys in
- match oeq with
- | None -> None (* Nothing to do *)
- | Some(v,pc) ->
- Some(pivot_sys v pc sys)
+ let is_unary_equation (cstr,prf) =
+ if cstr.op == Eq
+ then
+ try
+ Some (fst (List.find (fun (_,n) -> n =/ (Int 1) || n=/ (Int (-1))) cstr.coeffs))
+ with Not_found -> None
+ else None in
+
+ let (oeq,sys) = extract is_unary_equation psys in
+ match oeq with
+ | None -> None (* Nothing to do *)
+ | Some(v,pc) ->
+ Some(pivot_sys v pc sys)
let reduce_non_lin_unary psys =
- let is_unary_equation (cstr,prf) =
- if cstr.op == Eq
- then
- try
- let x = fst (List.find (fun (x,n) -> (n =/ (Int 1) || n=/ (Int (-1))) && Monomial.is_var (LinPoly.MonT.retrieve x) ) cstr.coeffs) in
- let x' = LinPoly.MonT.retrieve x in
- if List.for_all (fun (y,_) -> Pervasives.(=) y x || Int.equal (snd (Monomial.div (LinPoly.MonT.retrieve y) x')) 0) cstr.coeffs
- then Some x
- else None
- with Not_found -> None
- else None in
-
-
- let (oeq,sys) = extract is_unary_equation psys in
- match oeq with
- | None -> None (* Nothing to do *)
- | Some(v,pc) ->
- Some(apply_and_normalise (LinPoly.pivot_eq v pc) sys)
+ let is_unary_equation (cstr,prf) =
+ if cstr.op == Eq
+ then
+ try
+ let x = fst (List.find (fun (x,n) -> (n =/ (Int 1) || n=/ (Int (-1))) && Monomial.is_var (LinPoly.MonT.retrieve x) ) cstr.coeffs) in
+ let x' = LinPoly.MonT.retrieve x in
+ if List.for_all (fun (y,_) -> Pervasives.(=) y x || Int.equal (snd (Monomial.div (LinPoly.MonT.retrieve y) x')) 0) cstr.coeffs
+ then Some x
+ else None
+ with Not_found -> None
+ else None in
+
+
+ let (oeq,sys) = extract is_unary_equation psys in
+ match oeq with
+ | None -> None (* Nothing to do *)
+ | Some(v,pc) ->
+ Some(apply_and_normalise (LinPoly.pivot_eq v pc) sys)
let reduce_var_change psys =
- let rec rel_prime vect =
- match vect with
- | [] -> None
- | (x,v)::vect ->
- let v = numerator v in
- try
- let (x',v') = List.find (fun (_,v') ->
- let v' = numerator v' in
- eq_big_int (gcd_big_int v v') unit_big_int) vect in
- Some ((x,v),(x',numerator v'))
- with Not_found -> rel_prime vect in
-
- let rel_prime (cstr,prf) = if cstr.op == Eq then rel_prime cstr.coeffs else None in
-
- let (oeq,sys) = extract rel_prime psys in
-
- match oeq with
- | None -> None
- | Some(((x,v),(x',v')),(c,p)) ->
- let (l1,l2) = ext_gcd v v' in
- let l1,l2 = Big_int l1 , Big_int l2 in
+ let rec rel_prime vect =
+ match vect with
+ | [] -> None
+ | (x,v)::vect ->
+ let v = numerator v in
+ try
+ let (x',v') = List.find (fun (_,v') ->
+ let v' = numerator v' in
+ eq_big_int (gcd_big_int v v') unit_big_int) vect in
+ Some ((x,v),(x',numerator v'))
+ with Not_found -> rel_prime vect in
+
+ let rel_prime (cstr,prf) = if cstr.op == Eq then rel_prime cstr.coeffs else None in
- let get v vect =
- match Vect.get v vect with
- | None -> Int 0
- | Some n -> n in
+ let (oeq,sys) = extract rel_prime psys in
+
+ match oeq with
+ | None -> None
+ | Some(((x,v),(x',v')),(c,p)) ->
+ let (l1,l2) = ext_gcd v v' in
+ let l1,l2 = Big_int l1 , Big_int l2 in
- let pivot_eq (c',p') =
- let {coeffs = coeffs ; op = op ; cst = cst} = c' in
- let vx = get x coeffs in
- let vx' = get x' coeffs in
- let m = minus_num (vx */ l1 +/ vx' */ l2) in
- Some ({coeffs =
- Vect.add (Vect.mul m c.coeffs) coeffs ; op = op ; cst = m */ c.cst +/ cst} ,
- AddPrf(MulC(([], m),p),p')) in
+ let get v vect =
+ match Vect.get v vect with
+ | None -> Int 0
+ | Some n -> n in
- Some (apply_and_normalise pivot_eq sys)
+ let pivot_eq (c',p') =
+ let {coeffs = coeffs ; op = op ; cst = cst} = c' in
+ let vx = get x coeffs in
+ let vx' = get x' coeffs in
+ let m = minus_num (vx */ l1 +/ vx' */ l2) in
+ Some ({coeffs =
+ Vect.add (Vect.mul m c.coeffs) coeffs ; op = op ; cst = m */ c.cst +/ cst} ,
+ AddPrf(MulC(([], m),p),p')) in
+ Some (apply_and_normalise pivot_eq sys)
- let reduce_pivot psys =
- let is_equation (cstr,prf) =
- if cstr.op == Eq
- then
- try
- Some (fst (List.hd cstr.coeffs))
- with Not_found -> None
- else None in
- let (oeq,sys) = extract is_equation psys in
- match oeq with
- | None -> None (* Nothing to do *)
- | Some(v,pc) ->
- if debug then
- Printf.printf "Bad news : loss of completeness %a=%s" Vect.pp_vect (fst pc).coeffs (string_of_num (fst pc).cst);
- Some(pivot_sys v pc sys)
+let reduce_pivot psys =
+ let is_equation (cstr,prf) =
+ if cstr.op == Eq
+ then
+ try
+ Some (fst (List.hd cstr.coeffs))
+ with Not_found -> None
+ else None in
+ let (oeq,sys) = extract is_equation psys in
+ match oeq with
+ | None -> None (* Nothing to do *)
+ | Some(v,pc) ->
+ if debug then
+ Printf.printf "Bad news : loss of completeness %a=%s" Vect.pp_vect (fst pc).coeffs (string_of_num (fst pc).cst);
+ Some(pivot_sys v pc sys)
- let iterate_until_stable f x =
- let rec iter x =
- match f x with
- | None -> x
- | Some x' -> iter x' in
- iter x
- let rec app_funs l x =
- match l with
- | [] -> None
- | f::fl ->
- match f x with
- | None -> app_funs fl x
- | Some x' -> Some x'
+let iterate_until_stable f x =
+ let rec iter x =
+ match f x with
+ | None -> x
+ | Some x' -> iter x' in
+ iter x
- let reduction_equations psys =
- iterate_until_stable (app_funs
- [reduce_unary ; reduce_coprime ;
- reduce_var_change (*; reduce_pivot*)]) psys
+let rec app_funs l x =
+ match l with
+ | [] -> None
+ | f::fl ->
+ match f x with
+ | None -> app_funs fl x
+ | Some x' -> Some x'
- let reduction_non_lin_equations psys =
- iterate_until_stable (app_funs
- [reduce_non_lin_unary (*; reduce_coprime ;
- reduce_var_change ; reduce_pivot *)]) psys
+let reduction_equations psys =
+ iterate_until_stable (app_funs
+ [reduce_unary ; reduce_coprime ;
+ reduce_var_change (*; reduce_pivot*)]) psys
+
+let reduction_non_lin_equations psys =
+ iterate_until_stable (app_funs
+ [reduce_non_lin_unary (*; reduce_coprime ;
+ reduce_var_change ; reduce_pivot *)]) psys
(** [get_bound sys] returns upon success an interval (lb,e,ub) with proofs *)
- let get_bound sys =
- let is_small (v,i) =
- match Itv.range i with
- | None -> false
- | Some i -> i <=/ (Int 1) in
-
- let select_best (x1,i1) (x2,i2) =
- if Itv.smaller_itv i1 i2
- then (x1,i1) else (x2,i2) in
+let get_bound sys =
+ let is_small (v,i) =
+ match Itv.range i with
+ | None -> false
+ | Some i -> i <=/ (Int 1) in
+
+ let select_best (x1,i1) (x2,i2) =
+ if Itv.smaller_itv i1 i2
+ then (x1,i1) else (x2,i2) in
(* For lia, there are no equations => these precautions are not needed *)
(* For nlia, there are equations => do not enumerate over equations! *)
- let all_planes sys =
- let (eq,ineq) = List.partition (fun c -> c.op == Eq) sys in
- match eq with
- | [] -> List.rev_map (fun c -> c.coeffs) ineq
- | _ ->
- List.fold_left (fun acc c ->
- if List.exists (fun c' -> Vect.equal c.coeffs c'.coeffs) eq
- then acc else c.coeffs ::acc) [] ineq in
-
- let smallest_interval =
- List.fold_left
- (fun acc vect ->
- if is_small acc
- then acc
- else
- match Fourier.optimise vect sys with
- | None -> acc
- | Some i ->
- if debug then Printf.printf "Found a new bound %a" Vect.pp_vect vect ;
- select_best (vect,i) acc) (Vect.null, (None,None)) (all_planes sys) in
- let smallest_interval =
- match smallest_interval
- with
- | (x,(Some i, Some j)) -> Some(i,x,j)
- | x -> None (* This should not be possible *)
- in
- match smallest_interval with
- | Some (lb,e,ub) ->
- let (lbn,lbd) = (sub_big_int (numerator lb) unit_big_int, denominator lb) in
- let (ubn,ubd) = (add_big_int unit_big_int (numerator ub) , denominator ub) in
- (match
+ let all_planes sys =
+ let (eq,ineq) = List.partition (fun c -> c.op == Eq) sys in
+ match eq with
+ | [] -> List.rev_map (fun c -> c.coeffs) ineq
+ | _ ->
+ List.fold_left (fun acc c ->
+ if List.exists (fun c' -> Vect.equal c.coeffs c'.coeffs) eq
+ then acc else c.coeffs ::acc) [] ineq in
+
+ let smallest_interval =
+ List.fold_left
+ (fun acc vect ->
+ if is_small acc
+ then acc
+ else
+ match Fourier.optimise vect sys with
+ | None -> acc
+ | Some i ->
+ if debug then Printf.printf "Found a new bound %a" Vect.pp_vect vect ;
+ select_best (vect,i) acc) (Vect.null, (None,None)) (all_planes sys) in
+ let smallest_interval =
+ match smallest_interval
+ with
+ | (x,(Some i, Some j)) -> Some(i,x,j)
+ | x -> None (* This should not be possible *)
+ in
+ match smallest_interval with
+ | Some (lb,e,ub) ->
+ let (lbn,lbd) = (sub_big_int (numerator lb) unit_big_int, denominator lb) in
+ let (ubn,ubd) = (add_big_int unit_big_int (numerator ub) , denominator ub) in
+ (match
(* x <= ub -> x > ub *)
- xlinear_prover ({coeffs = Vect.mul (Big_int ubd) e ; op = Ge ; cst = Big_int ubn} :: sys),
+ xlinear_prover ({coeffs = Vect.mul (Big_int ubd) e ; op = Ge ; cst = Big_int ubn} :: sys),
(* lb <= x -> lb > x *)
- xlinear_prover
- ({coeffs = Vect.mul (minus_num (Big_int lbd)) e ; op = Ge ; cst = minus_num (Big_int lbn)} :: sys)
- with
- | Some cub , Some clb -> Some(List.tl clb,(lb,e,ub), List.tl cub)
- | _ -> failwith "Interval without proof"
- )
- | None -> None
-
-
- let check_sys sys =
- List.for_all (fun (c,p) -> List.for_all (fun (_,n) -> sign_num n <> 0) c.coeffs) sys
-
-
- let xlia reduction_equations sys =
-
- let rec enum_proof (id:int) (sys:prf_sys) : proof option =
- if debug then (Printf.printf "enum_proof\n" ; flush stdout) ;
- assert (check_sys sys) ;
-
- let nsys,prf = List.split sys in
- match get_bound nsys with
- | None -> None (* Is the systeme really unbounded ? *)
- | Some(prf1,(lb,e,ub),prf2) ->
- if debug then Printf.printf "Found interval: %a in [%s;%s] -> " Vect.pp_vect e (string_of_num lb) (string_of_num ub) ;
- (match start_enum id e (ceiling_num lb) (floor_num ub) sys
- with
- | Some prfl ->
- Some(Enum(id,proof_of_farkas prf prf1,e, proof_of_farkas prf prf2,prfl))
- | None -> None
- )
-
- and start_enum id e clb cub sys =
- if clb >/ cub
- then Some []
- else
- let eq = {coeffs = e ; op = Eq ; cst = clb} in
- match aux_lia (id+1) ((eq, Def id) :: sys) with
- | None -> None
- | Some prf ->
- match start_enum id e (clb +/ (Int 1)) cub sys with
- | None -> None
- | Some l -> Some (prf::l)
-
- and aux_lia (id:int) (sys:prf_sys) : proof option =
- assert (check_sys sys) ;
- if debug then Printf.printf "xlia: %a \n" (pp_list (fun o (c,_) -> output_cstr o c)) sys ;
- try
- let sys = reduction_equations sys in
- if debug then
+ xlinear_prover
+ ({coeffs = Vect.mul (minus_num (Big_int lbd)) e ; op = Ge ; cst = minus_num (Big_int lbn)} :: sys)
+ with
+ | Some cub , Some clb -> Some(List.tl clb,(lb,e,ub), List.tl cub)
+ | _ -> failwith "Interval without proof"
+ )
+ | None -> None
+
+
+let check_sys sys =
+ List.for_all (fun (c,p) -> List.for_all (fun (_,n) -> sign_num n <> 0) c.coeffs) sys
+
+
+let xlia (can_enum:bool) reduction_equations sys =
+
+
+ let rec enum_proof (id:int) (sys:prf_sys) : proof option =
+ if debug then (Printf.printf "enum_proof\n" ; flush stdout) ;
+ assert (check_sys sys) ;
+
+ let nsys,prf = List.split sys in
+ match get_bound nsys with
+ | None -> None (* Is the systeme really unbounded ? *)
+ | Some(prf1,(lb,e,ub),prf2) ->
+ if debug then Printf.printf "Found interval: %a in [%s;%s] -> " Vect.pp_vect e (string_of_num lb) (string_of_num ub) ;
+ (match start_enum id e (ceiling_num lb) (floor_num ub) sys
+ with
+ | Some prfl ->
+ Some(Enum(id,proof_of_farkas prf prf1,e, proof_of_farkas prf prf2,prfl))
+ | None -> None
+ )
+
+ and start_enum id e clb cub sys =
+ if clb >/ cub
+ then Some []
+ else
+ let eq = {coeffs = e ; op = Eq ; cst = clb} in
+ match aux_lia (id+1) ((eq, Def id) :: sys) with
+ | None -> None
+ | Some prf ->
+ match start_enum id e (clb +/ (Int 1)) cub sys with
+ | None -> None
+ | Some l -> Some (prf::l)
+
+ and aux_lia (id:int) (sys:prf_sys) : proof option =
+ assert (check_sys sys) ;
+ if debug then Printf.printf "xlia: %a \n" (pp_list (fun o (c,_) -> output_cstr o c)) sys ;
+ try
+ let sys = reduction_equations sys in
+ if debug then
Printf.printf "after reduction: %a \n" (pp_list (fun o (c,_) -> output_cstr o c)) sys ;
- match linear_prover sys with
- | Some prf -> Some (Step(id,prf,Done))
- | None -> enum_proof id sys
- with FoundProof prf ->
+ match linear_prover sys with
+ | Some prf -> Some (Step(id,prf,Done))
+ | None -> if can_enum then enum_proof id sys else None
+ with FoundProof prf ->
(* [reduction_equations] can find a proof *)
- Some(Step(id,prf,Done)) in
+ Some(Step(id,prf,Done)) in
(* let sys' = List.map (fun (p,o) -> Mc.norm0 p , o) sys in*)
- let id = List.length sys in
- let orpf =
- try
- let sys = simpl_sys sys in
- aux_lia id sys
- with FoundProof pr -> Some(Step(id,pr,Done)) in
- match orpf with
- | None -> None
- | Some prf ->
+ let id = List.length sys in
+ let orpf =
+ try
+ let sys = simpl_sys sys in
+ aux_lia id sys
+ with FoundProof pr -> Some(Step(id,pr,Done)) in
+ match orpf with
+ | None -> None
+ | Some prf ->
(*Printf.printf "direct proof %a\n" output_proof prf ; *)
- let env = mapi (fun _ i -> i) sys in
- let prf = compile_proof env prf in
+ let env = mapi (fun _ i -> i) sys in
+ let prf = compile_proof env prf in
(*try
if Mc.zChecker sys' prf then Some prf else
raise Certificate.BadCertificate
with Failure s -> (Printf.printf "%s" s ; Some prf)
*) Some prf
-
-
- let cstr_compat_of_poly (p,o) =
- let (v,c) = LinPoly.linpol_of_pol p in
- {coeffs = v ; op = o ; cst = minus_num c }
-
-
- let lia sys =
- LinPoly.MonT.clear ();
- let sys = List.map (develop_constraint z_spec) sys in
- let (sys:cstr_compat list) = List.map cstr_compat_of_poly sys in
- let sys = mapi (fun c i -> (c,Hyp i)) sys in
- xlia reduction_equations sys
-
-
- let nlia sys =
- LinPoly.MonT.clear ();
- let sys = List.map (develop_constraint z_spec) sys in
- let sys = mapi (fun c i -> (c,Hyp i)) sys in
-
- let is_linear = List.for_all (fun ((p,_),_) -> Poly.is_linear p) sys in
-
- let collect_square =
- List.fold_left (fun acc ((p,_),_) -> Poly.fold
- (fun m _ acc ->
- match Monomial.sqrt m with
- | None -> acc
- | Some s -> MonMap.add s m acc) p acc) MonMap.empty sys in
- let sys = MonMap.fold (fun s m acc ->
- let s = LinPoly.linpol_of_pol (Poly.add s (Int 1) (Poly.constant (Int 0))) in
- let m = Poly.add m (Int 1) (Poly.constant (Int 0)) in
- ((m, Ge), (Square s))::acc) collect_square sys in
-
-(* List.iter (fun ((p,_),_) -> Printf.printf "square %a\n" Poly.pp p) gen_square*)
-
- let sys =
- if is_linear then sys
- else sys @ (all_sym_pairs (fun ((c,o),p) ((c',o'),p') ->
- ((Poly.product c c',opMult o o'), MulPrf(p,p'))) sys) in
+
- let sys = List.map (fun (c,p) -> cstr_compat_of_poly c,p) sys in
- assert (check_sys sys) ;
- xlia (if is_linear then reduction_equations else reduction_non_lin_equations) sys
+let cstr_compat_of_poly (p,o) =
+ let (v,c) = LinPoly.linpol_of_pol p in
+ {coeffs = v ; op = o ; cst = minus_num c }
+
+
+let lia (can_enum:bool) (prfdepth:int) sys =
+ LinPoly.MonT.clear ();
+ max_nb_cstr := compute_max_nb_cstr sys prfdepth ;
+ let sys = List.map (develop_constraint z_spec) sys in
+ let (sys:cstr_compat list) = List.map cstr_compat_of_poly sys in
+ let sys = mapi (fun c i -> (c,Hyp i)) sys in
+ xlia can_enum reduction_equations sys
+
+
+let nlia enum prfdepth sys =
+ LinPoly.MonT.clear ();
+ max_nb_cstr := compute_max_nb_cstr sys prfdepth;
+ let sys = List.map (develop_constraint z_spec) sys in
+ let sys = mapi (fun c i -> (c,Hyp i)) sys in
+
+ let is_linear = List.for_all (fun ((p,_),_) -> Poly.is_linear p) sys in
+
+ let collect_square =
+ List.fold_left (fun acc ((p,_),_) -> Poly.fold
+ (fun m _ acc ->
+ match Monomial.sqrt m with
+ | None -> acc
+ | Some s -> MonMap.add s m acc) p acc) MonMap.empty sys in
+ let sys = MonMap.fold (fun s m acc ->
+ let s = LinPoly.linpol_of_pol (Poly.add s (Int 1) (Poly.constant (Int 0))) in
+ let m = Poly.add m (Int 1) (Poly.constant (Int 0)) in
+ ((m, Ge), (Square s))::acc) collect_square sys in
+
+ (* List.iter (fun ((p,_),_) -> Printf.printf "square %a\n" Poly.pp p) gen_square*)
+
+ let sys =
+ if is_linear then sys
+ else sys @ (all_sym_pairs (fun ((c,o),p) ((c',o'),p') ->
+ ((Poly.product c c',opMult o o'), MulPrf(p,p'))) sys) in
+
+ let sys = List.map (fun (c,p) -> cstr_compat_of_poly c,p) sys in
+ assert (check_sys sys) ;
+ xlia enum (if is_linear then reduction_equations else reduction_non_lin_equations) sys
diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml
index cce0a7280..5fef6d3fc 100644
--- a/plugins/micromega/coq_micromega.ml
+++ b/plugins/micromega/coq_micromega.ml
@@ -18,6 +18,9 @@
open Pp
open Mutils
+open Proofview
+open Goptions
+open Proofview.Notations
(**
* Debug flag
@@ -37,6 +40,53 @@ let time str f x =
flush stdout);
res
+
+(* Limit the proof search *)
+
+let max_depth = max_int
+
+(* Search limit for provers over Q R *)
+let lra_proof_depth = ref max_depth
+
+
+(* Search limit for provers over Z *)
+let lia_enum = ref true
+let lia_proof_depth = ref max_depth
+
+let get_lia_option () =
+ (!lia_enum,!lia_proof_depth)
+
+let get_lra_option () =
+ !lra_proof_depth
+
+
+
+let _ =
+
+ let int_opt l vref =
+ {
+ optsync = true;
+ optdepr = false;
+ optname = List.fold_right (^) l "";
+ optkey = l ;
+ optread = (fun () -> Some !vref);
+ optwrite = (fun x -> vref := (match x with None -> max_depth | Some v -> v))
+ } in
+
+ let lia_enum_opt =
+ {
+ optsync = true;
+ optdepr = false;
+ optname = "Lia Enum";
+ optkey = ["Lia";"Enum"];
+ optread = (fun () -> !lia_enum);
+ optwrite = (fun x -> lia_enum := x)
+ } in
+ let _ = declare_int_option (int_opt ["Lra"; "Depth"] lra_proof_depth) in
+ let _ = declare_int_option (int_opt ["Lia"; "Depth"] lia_proof_depth) in
+ let _ = declare_bool_option lia_enum_opt in
+ ()
+
(**
* Initialize a tag type to the Tag module declaration (see Mutils).
*)
@@ -303,7 +353,8 @@ struct
let r_modules =
[["Coq";"Reals" ; "Rdefinitions"];
["Coq";"Reals" ; "Rpow_def"] ;
-]
+ ["Coq";"Reals" ; "Raxioms"] ;
+ ]
let z_modules = [["Coq";"ZArith";"BinInt"]]
@@ -359,6 +410,7 @@ struct
let coq_Qmake = lazy (constant "Qmake")
let coq_Rcst = lazy (constant "Rcst")
+
let coq_C0 = lazy (m_constant "C0")
let coq_C1 = lazy (m_constant "C1")
let coq_CQ = lazy (m_constant "CQ")
@@ -415,8 +467,9 @@ struct
let coq_Rdiv = lazy (r_constant "Rdiv")
let coq_Rinv = lazy (r_constant "Rinv")
let coq_Rpower = lazy (r_constant "pow")
+ let coq_IZR = lazy (r_constant "IZR")
let coq_IQR = lazy (constant "IQR")
- let coq_IZR = lazy (constant "IZR")
+
let coq_PEX = lazy (constant "PEX" )
let coq_PEc = lazy (constant"PEc")
@@ -994,7 +1047,7 @@ struct
coq_Rplus , (fun x y -> Mc.CPlus(x,y)) ;
coq_Rminus , (fun x y -> Mc.CMinus(x,y)) ;
coq_Rmult , (fun x y -> Mc.CMult(x,y)) ;
- coq_Rdiv , (fun x y -> Mc.CMult(x,Mc.CInv y)) ;
+ (* coq_Rdiv , (fun x y -> Mc.CMult(x,Mc.CInv y)) ;*)
]
let rec rconstant term =
@@ -1016,10 +1069,14 @@ struct
with
ParseError ->
match op with
- | op when Constr.equal op (Lazy.force coq_Rinv) -> Mc.CInv(rconstant args.(0))
- | op when Constr.equal op (Lazy.force coq_IQR) -> Mc.CQ (parse_q args.(0))
-(* | op when op = Lazy.force coq_IZR -> Mc.CZ (parse_z args.(0))*)
- | _ -> raise ParseError
+ | op when Constr.equal op (Lazy.force coq_Rinv) ->
+ let arg = rconstant args.(0) in
+ if Mc.qeq_bool (Mc.q_of_Rcst arg) {Mc.qnum = Mc.Z0 ; Mc.qden = Mc.XH}
+ then raise ParseError (* This is a division by zero -- no semantics *)
+ else Mc.CInv(arg)
+ | op when Constr.equal op (Lazy.force coq_IQR) -> Mc.CQ (parse_q args.(0))
+ | op when Constr.equal op (Lazy.force coq_IZR) -> Mc.CZ (parse_z args.(0))
+ | _ -> raise ParseError
end
| _ -> raise ParseError
@@ -1094,10 +1151,6 @@ struct
| N (a) -> Mc.N(f2f a)
| I(a,_,b) -> Mc.I(f2f a,f2f b)
- let is_prop t =
- match t with
- | Names.Anonymous -> true (* Not quite right *)
- | Names.Name x -> false
let mkC f1 f2 = C(f1,f2)
let mkD f1 f2 = D(f1,f2)
@@ -1121,6 +1174,11 @@ struct
(A(at,tg,t), env,Tag.next tg)
with e when Errors.noncritical e -> (X(t),env,tg) in
+ let is_prop term =
+ let ty = Typing.unsafe_type_of (Tacmach.pf_env gl) (Tacmach.project gl) term in
+ let sort = Typing.e_sort_of (Tacmach.pf_env gl) (ref (Tacmach.project gl)) ty in
+ Term.is_prop_sort sort in
+
let rec xparse_formula env tg term =
match kind_of_term term with
| App(l,rst) ->
@@ -1140,13 +1198,15 @@ struct
let g,env,tg = xparse_formula env tg b in
mkformula_binary mkIff term f g,env,tg
| _ -> parse_atom env tg term)
- | Prod(typ,a,b) when not (Termops.dependent (mkRel 1) b) ->
+ | Prod(typ,a,b) when not (Termops.dependent (mkRel 1) b)->
let f,env,tg = xparse_formula env tg a in
let g,env,tg = xparse_formula env tg b in
mkformula_binary mkI term f g,env,tg
| _ when eq_constr term (Lazy.force coq_True) -> (TT,env,tg)
| _ when eq_constr term (Lazy.force coq_False) -> (FF,env,tg)
- | _ -> X(term),env,tg in
+ | _ when is_prop term -> X(term),env,tg
+ | _ -> raise ParseError
+ in
xparse_formula env tg ((*Reductionops.whd_zeta*) term)
let dump_formula typ dump_atom f =
@@ -1377,50 +1437,57 @@ let rcst_domain_spec = lazy {
dump_proof = dump_psatz coq_Q dump_q
}
+open Proofview.Notations
+
+
(**
* Instanciate the current Coq goal with a Micromega formula, a varmap, and a
* witness.
*)
-
-
-let micromega_order_change spec cert cert_typ env ff : Tacmach.tactic =
+let micromega_order_change spec cert cert_typ env ff (*: unit Proofview.tactic*) =
let ids = Util.List.map_i (fun i _ -> (Names.Id.of_string ("__z"^(string_of_int i)))) 0 env in
let formula_typ = (Term.mkApp (Lazy.force coq_Cstr,[|spec.coeff|])) in
let ff = dump_formula formula_typ (dump_cstr spec.coeff spec.dump_coeff) ff in
let vm = dump_varmap (spec.typ) env in
- (* todo : directly generate the proof term - or generalize befor conversion? *)
- Tacticals.tclTHENSEQ [
- (fun gl ->
- Proofview.V82.of_tactic (Tactics.change_concl
- (set
- [
- ("__ff", ff, Term.mkApp(Lazy.force coq_Formula, [|formula_typ |]));
- ("__varmap", vm, Term.mkApp
- (Coqlib.gen_constant_in_modules "VarMap"
- [["Coq" ; "micromega" ; "VarMap"] ; ["VarMap"]] "t", [|spec.typ|]));
- ("__wit", cert, cert_typ)
- ]
- (Tacmach.pf_concl gl))) gl);
- Tactics.generalize env ;
- Tacticals.tclTHENSEQ (List.map (fun id -> Proofview.V82.of_tactic (Tactics.introduction id)) ids) ;
- ]
-
+ (* todo : directly generate the proof term - or generalize before conversion? *)
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
+ let gl = Tacmach.New.of_old (fun x -> x) gl in
+ Tacticals.New.tclTHENLIST
+ [
+ Tactics.change_concl
+ (set
+ [
+ ("__ff", ff, Term.mkApp(Lazy.force coq_Formula, [|formula_typ |]));
+ ("__varmap", vm, Term.mkApp
+ (Coqlib.gen_constant_in_modules "VarMap"
+ [["Coq" ; "micromega" ; "VarMap"] ; ["VarMap"]] "t", [|spec.typ|]));
+ ("__wit", cert, cert_typ)
+ ]
+ (Tacmach.pf_concl gl))
+ ;
+ Tactics.new_generalize env ;
+ Tacticals.New.tclTHENLIST (List.map (fun id -> (Tactics.introduction id)) ids)
+ ]
+ end }
(**
* The datastructures that aggregate prover attributes.
*)
-type ('a,'prf) prover = {
+type ('option,'a,'prf) prover = {
name : string ; (* name of the prover *)
- prover : 'a list -> 'prf option ; (* the prover itself *)
+ get_option : unit ->'option ; (* find the options of the prover *)
+ prover : 'option * 'a list -> 'prf option ; (* the prover itself *)
hyps : 'prf -> ISet.t ; (* extract the indexes of the hypotheses really used in the proof *)
compact : 'prf -> (int -> int) -> 'prf ; (* remap the hyp indexes according to function *)
pp_prf : out_channel -> 'prf -> unit ;(* pretting printing of proof *)
pp_f : out_channel -> 'a -> unit (* pretty printing of the formulas (polynomials)*)
}
+
+
(**
* Given a list of provers and a disjunction of atoms, find a proof of any of
* the atoms. Returns an (optional) pair of a proof and a prover
@@ -1430,7 +1497,7 @@ type ('a,'prf) prover = {
let find_witness provers polys1 =
let provers = List.map (fun p ->
(fun l ->
- match p.prover l with
+ match p.prover (p.get_option (),l) with
| None -> None
| Some prf -> Some(prf,p)) , p.name) provers in
try_any provers (List.map fst polys1)
@@ -1485,7 +1552,7 @@ let compact_proofs (cnf_ff: 'cst cnf) res (cnf_ff': 'cst cnf) =
let res = try prover.compact prf remap with x when Errors.noncritical x ->
if debug then Printf.fprintf stdout "Proof compaction %s" (Printexc.to_string x) ;
(* This should not happen -- this is the recovery plan... *)
- match prover.prover (List.map fst new_cl) with
+ match prover.prover (prover.get_option () ,List.map fst new_cl) with
| None -> failwith "proof compaction error"
| Some p -> p
in
@@ -1646,58 +1713,76 @@ let micromega_gen
(negate:'cst atom -> 'cst mc_cnf)
(normalise:'cst atom -> 'cst mc_cnf)
unsat deduce
- spec prover gl =
- let concl = Tacmach.pf_concl gl in
- let hyps = Tacmach.pf_hyps_types gl in
- try
- let (hyps,concl,env) = parse_goal gl parse_arith Env.empty hyps concl in
- let env = Env.elements env in
- let spec = Lazy.force spec in
-
+ spec prover =
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
+ let gl = Tacmach.New.of_old (fun x -> x) gl in
+ let concl = Tacmach.pf_concl gl in
+ let hyps = Tacmach.pf_hyps_types gl in
+ try
+ let (hyps,concl,env) = parse_goal gl parse_arith Env.empty hyps concl in
+ let env = Env.elements env in
+ let spec = Lazy.force spec in
+
match micromega_tauto negate normalise unsat deduce spec prover env hyps concl gl with
- | None -> Tacticals.tclFAIL 0 (Pp.str " Cannot find witness") gl
- | Some (ids,ff',res') ->
- (Tacticals.tclTHENSEQ
- [
- Tactics.generalize (List.map Term.mkVar ids) ;
- micromega_order_change spec res'
- (Term.mkApp(Lazy.force coq_list, [|spec.proof_typ|])) env ff'
- ]) gl
- with
- | ParseError -> Tacticals.tclFAIL 0 (Pp.str "Bad logical fragment") gl
- | CsdpNotFound -> flush stdout ; Pp.pp_flush () ;
- Tacticals.tclFAIL 0 (Pp.str
- (" Skipping what remains of this tactic: the complexity of the goal requires "
- ^ "the use of a specialized external tool called csdp. \n\n"
- ^ "Unfortunately Coq isn't aware of the presence of any \"csdp\" executable in the path. \n\n"
- ^ "Csdp packages are provided by some OS distributions; binaries and source code can be downloaded from https://projects.coin-or.org/Csdp")) gl
-
-
-
-let micromega_order_changer cert env ff gl =
- let coeff = Lazy.force coq_Rcst in
- let dump_coeff = dump_Rcst in
- let typ = Lazy.force coq_R in
- let cert_typ = (Term.mkApp(Lazy.force coq_list, [|Lazy.force coq_QWitness |])) in
+ | None -> Tacticals.New.tclFAIL 0 (Pp.str " Cannot find witness")
+ | Some (ids,ff',res') ->
+ (Tacticals.New.tclTHENLIST
+ [
+ Tactics.new_generalize (List.map Term.mkVar ids) ;
+ micromega_order_change spec res'
+ (Term.mkApp(Lazy.force coq_list, [|spec.proof_typ|])) env ff'
+ ])
+ with
+ | ParseError -> Tacticals.New.tclFAIL 0 (Pp.str "Bad logical fragment")
+ | Mfourier.TimeOut -> Tacticals.New.tclFAIL 0 (Pp.str "Timeout")
+ | CsdpNotFound -> flush stdout ; Pp.pp_flush () ;
+ Tacticals.New.tclFAIL 0 (Pp.str
+ (" Skipping what remains of this tactic: the complexity of the goal requires "
+ ^ "the use of a specialized external tool called csdp. \n\n"
+ ^ "Unfortunately Coq isn't aware of the presence of any \"csdp\" executable in the path. \n\n"
+ ^ "Csdp packages are provided by some OS distributions; binaries and source code can be downloaded from https://projects.coin-or.org/Csdp"))
+ end }
+
+let micromega_gen parse_arith
+ (negate:'cst atom -> 'cst mc_cnf)
+ (normalise:'cst atom -> 'cst mc_cnf)
+ unsat deduce
+ spec prover =
+ (micromega_gen parse_arith negate normalise unsat deduce spec prover)
+
+
+let micromega_order_changer cert env ff =
+ let ids = Util.List.map_i (fun i _ -> (Names.Id.of_string ("__z"^(string_of_int i)))) 0 env in
+ let coeff = Lazy.force coq_Rcst in
+ let dump_coeff = dump_Rcst in
+ let typ = Lazy.force coq_R in
+ let cert_typ = (Term.mkApp(Lazy.force coq_list, [|Lazy.force coq_QWitness |])) in
+
let formula_typ = (Term.mkApp (Lazy.force coq_Cstr,[| coeff|])) in
let ff = dump_formula formula_typ (dump_cstr coeff dump_coeff) ff in
let vm = dump_varmap (typ) env in
- Proofview.V82.of_tactic (Tactics.change_concl
- (set
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
+ let gl = Tacmach.New.of_old (fun x -> x) gl in
+ Tacticals.New.tclTHENLIST
[
- ("__ff", ff, Term.mkApp(Lazy.force coq_Formula, [|formula_typ |]));
- ("__varmap", vm, Term.mkApp
- (Coqlib.gen_constant_in_modules "VarMap"
- [["Coq" ; "micromega" ; "VarMap"] ; ["VarMap"]] "t", [|typ|]));
- ("__wit", cert, cert_typ)
+ (Tactics.change_concl
+ (set
+ [
+ ("__ff", ff, Term.mkApp(Lazy.force coq_Formula, [|formula_typ |]));
+ ("__varmap", vm, Term.mkApp
+ (Coqlib.gen_constant_in_modules "VarMap"
+ [["Coq" ; "micromega" ; "VarMap"] ; ["VarMap"]] "t", [|typ|]));
+ ("__wit", cert, cert_typ)
+ ]
+ (Tacmach.pf_concl gl)));
+ Tactics.new_generalize env ;
+ Tacticals.New.tclTHENLIST (List.map (fun id -> (Tactics.introduction id)) ids)
]
- (Tacmach.pf_concl gl)
- ))
- gl
+ end }
-let micromega_genr prover gl =
+let micromega_genr prover =
let parse_arith = parse_rarith in
let negate = Mc.rnegate in
let normalise = Mc.rnormalise in
@@ -1710,39 +1795,41 @@ let micromega_genr prover gl =
proof_typ = Lazy.force coq_QWitness ;
dump_proof = dump_psatz coq_Q dump_q
} in
-
- let concl = Tacmach.pf_concl gl in
- let hyps = Tacmach.pf_hyps_types gl in
- try
- let (hyps,concl,env) = parse_goal gl parse_arith Env.empty hyps concl in
- let env = Env.elements env in
- let spec = Lazy.force spec in
-
- let hyps' = List.map (fun (n,f) -> (n, map_atoms (Micromega.map_Formula Micromega.q_of_Rcst) f)) hyps in
- let concl' = map_atoms (Micromega.map_Formula Micromega.q_of_Rcst) concl in
-
- match micromega_tauto negate normalise unsat deduce spec prover env hyps' concl' gl with
- | None -> Tacticals.tclFAIL 0 (Pp.str " Cannot find witness") gl
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
+ let gl = Tacmach.New.of_old (fun x -> x) gl in
+ let concl = Tacmach.pf_concl gl in
+ let hyps = Tacmach.pf_hyps_types gl in
+ try
+ let (hyps,concl,env) = parse_goal gl parse_arith Env.empty hyps concl in
+ let env = Env.elements env in
+ let spec = Lazy.force spec in
+
+ let hyps' = List.map (fun (n,f) -> (n, map_atoms (Micromega.map_Formula Micromega.q_of_Rcst) f)) hyps in
+ let concl' = map_atoms (Micromega.map_Formula Micromega.q_of_Rcst) concl in
+
+ match micromega_tauto negate normalise unsat deduce spec prover env hyps' concl' gl with
+ | None -> Tacticals.New.tclFAIL 0 (Pp.str " Cannot find witness")
| Some (ids,ff',res') ->
let (ff,ids') = formula_hyps_concl
(List.filter (fun (n,_) -> List.mem n ids) hyps) concl in
-
- (Tacticals.tclTHENSEQ
+ (Tacticals.New.tclTHENLIST
[
- Tactics.generalize (List.map Term.mkVar ids) ;
+ Tactics.new_generalize (List.map Term.mkVar ids) ;
micromega_order_changer res' env (abstract_wrt_formula ff' ff)
- ]) gl
+ ])
with
- | ParseError -> Tacticals.tclFAIL 0 (Pp.str "Bad logical fragment") gl
- | CsdpNotFound -> flush stdout ; Pp.pp_flush () ;
- Tacticals.tclFAIL 0 (Pp.str
+ | Mfourier.TimeOut -> Tacticals.New.tclFAIL 0 (Pp.str "TimeOut")
+ | ParseError -> Tacticals.New.tclFAIL 0 (Pp.str "Bad logical fragment")
+ | CsdpNotFound -> flush stdout ; Pp.pp_flush () ;
+ 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")) gl
-
+ ^ "Csdp packages are provided by some OS distributions; binaries and source code can be downloaded from https://projects.coin-or.org/Csdp"))
+ end }
+let micromega_genr prover = (micromega_genr prover)
let lift_ratproof prover l =
@@ -1898,38 +1985,61 @@ let compact_pt pt f =
let lift_pexpr_prover p l = p (List.map (fun (e,o) -> Mc.denorm e , o) l)
-let linear_prover_Z = {
- name = "linear prover" ;
- prover = lift_ratproof (lift_pexpr_prover (Certificate.linear_prover_with_cert Certificate.z_spec)) ;
- hyps = hyps_of_pt ;
- compact = compact_pt ;
- pp_prf = pp_proof_term;
- pp_f = fun o x -> pp_pol pp_z o (fst x)
-}
+module CacheZ = PHashtable(struct
+ type prover_option = bool * int
+
+ type t = prover_option * ((Mc.z Mc.pol * Mc.op1) list)
+ let equal = (=)
+ let hash = Hashtbl.hash
+end)
+
+module CacheQ = PHashtable(struct
+ type t = int * ((Mc.q Mc.pol * Mc.op1) list)
+ let equal = (=)
+ let hash = Hashtbl.hash
+end)
+let memo_zlinear_prover = CacheZ.memo "lia.cache" (fun ((ce,b),s) -> lift_pexpr_prover (Certificate.lia ce b) s)
+let memo_nlia = CacheZ.memo "nlia.cache" (fun ((ce,b),s) -> lift_pexpr_prover (Certificate.nlia ce b) s)
+let memo_nra = CacheQ.memo "nra.cache" (fun (o,s) -> lift_pexpr_prover (Certificate.nlinear_prover o) s)
+
+
+
let linear_prover_Q = {
- name = "linear prover";
- prover = lift_pexpr_prover (Certificate.linear_prover_with_cert Certificate.q_spec) ;
- hyps = hyps_of_cone ;
- compact = compact_cone ;
- pp_prf = pp_psatz pp_q ;
- pp_f = fun o x -> pp_pol pp_q o (fst x)
+ name = "linear prover";
+ get_option = get_lra_option ;
+ prover = (fun (o,l) -> lift_pexpr_prover (Certificate.linear_prover_with_cert o Certificate.q_spec) l) ;
+ hyps = hyps_of_cone ;
+ compact = compact_cone ;
+ pp_prf = pp_psatz pp_q ;
+ pp_f = fun o x -> pp_pol pp_q o (fst x)
}
let linear_prover_R = {
name = "linear prover";
- prover = lift_pexpr_prover (Certificate.linear_prover_with_cert Certificate.q_spec) ;
+ get_option = get_lra_option ;
+ prover = (fun (o,l) -> lift_pexpr_prover (Certificate.linear_prover_with_cert o Certificate.q_spec) l) ;
hyps = hyps_of_cone ;
compact = compact_cone ;
pp_prf = pp_psatz pp_q ;
pp_f = fun o x -> pp_pol pp_q o (fst x)
}
+let nlinear_prover_R = {
+ name = "nra";
+ get_option = get_lra_option;
+ prover = memo_nra ;
+ hyps = hyps_of_cone ;
+ compact = compact_cone ;
+ pp_prf = pp_psatz pp_q ;
+ pp_f = fun o x -> pp_pol pp_q o (fst x)
+}
let non_linear_prover_Q str o = {
name = "real nonlinear prover";
- prover = call_csdpcert_q (str, o);
+ get_option = (fun () -> (str,o));
+ prover = (fun (o,l) -> call_csdpcert_q o l);
hyps = hyps_of_cone;
compact = compact_cone ;
pp_prf = pp_psatz pp_q ;
@@ -1938,7 +2048,8 @@ let non_linear_prover_Q str o = {
let non_linear_prover_R str o = {
name = "real nonlinear prover";
- prover = call_csdpcert_q (str, o);
+ get_option = (fun () -> (str,o));
+ prover = (fun (o,l) -> call_csdpcert_q o l);
hyps = hyps_of_cone;
compact = compact_cone;
pp_prf = pp_psatz pp_q;
@@ -1947,30 +2058,19 @@ let non_linear_prover_R str o = {
let non_linear_prover_Z str o = {
name = "real nonlinear prover";
- prover = lift_ratproof (call_csdpcert_z (str, o));
+ get_option = (fun () -> (str,o));
+ prover = (fun (o,l) -> lift_ratproof (call_csdpcert_z o) l);
hyps = hyps_of_pt;
compact = compact_pt;
pp_prf = pp_proof_term;
pp_f = fun o x -> pp_pol pp_z o (fst x)
}
-module CacheZ = PHashtable(struct
- type t = (Mc.z Mc.pol * Mc.op1) list
- let equal = Pervasives.(=)
- let hash = Hashtbl.hash
-end)
-
-let memo_zlinear_prover = CacheZ.memo "lia.cache" (lift_pexpr_prover Certificate.lia)
-let memo_nlia = CacheZ.memo "nlia.cache" (lift_pexpr_prover Certificate.nlia)
-
-(*let memo_zlinear_prover = (lift_pexpr_prover Lia.lia)*)
-(*let memo_zlinear_prover = CacheZ.memo "lia.cache" (lift_pexpr_prover Certificate.zlinear_prover)*)
-
-
let linear_Z = {
name = "lia";
- prover = memo_zlinear_prover ;
+ get_option = get_lia_option;
+ prover = memo_zlinear_prover ;
hyps = hyps_of_pt;
compact = compact_pt;
pp_prf = pp_proof_term;
@@ -1979,7 +2079,8 @@ let linear_Z = {
let nlinear_Z = {
name = "nlia";
- prover = memo_nlia ;
+ get_option = get_lia_option;
+ prover = memo_nlia ;
hyps = hyps_of_pt;
compact = compact_pt;
pp_prf = pp_proof_term;
@@ -2001,56 +2102,56 @@ let tauto_lia ff =
* solvers
*)
-let psatzl_Z gl =
+let psatzl_Z =
micromega_gen parse_zarith Mc.negate Mc.normalise Mc.zunsat Mc.zdeduce zz_domain_spec
- [ linear_prover_Z ] gl
+ [ linear_Z ]
-let psatzl_Q gl =
+let psatzl_Q =
micromega_gen parse_qarith Mc.qnegate Mc.qnormalise Mc.qunsat Mc.qdeduce qq_domain_spec
- [ linear_prover_Q ] gl
+ [ linear_prover_Q ]
-let psatz_Q i gl =
+let psatz_Q i =
micromega_gen parse_qarith Mc.qnegate Mc.qnormalise Mc.qunsat Mc.qdeduce qq_domain_spec
- [ non_linear_prover_Q "real_nonlinear_prover" (Some i) ] gl
-
-
-let psatzl_R gl =
- micromega_genr [ linear_prover_R ] gl
+ [ non_linear_prover_Q "real_nonlinear_prover" (Some i) ]
+let psatzl_R =
+ micromega_genr [ linear_prover_R ]
-let psatz_R i gl =
- micromega_genr [ non_linear_prover_R "real_nonlinear_prover" (Some i) ] gl
+let psatz_R i =
+ micromega_genr [ non_linear_prover_R "real_nonlinear_prover" (Some i) ]
-let psatz_Z i gl =
+let psatz_Z i =
micromega_gen parse_zarith Mc.negate Mc.normalise Mc.zunsat Mc.zdeduce zz_domain_spec
- [ non_linear_prover_Z "real_nonlinear_prover" (Some i) ] gl
+ [ non_linear_prover_Z "real_nonlinear_prover" (Some i) ]
-let sos_Z gl =
+let sos_Z =
micromega_gen parse_zarith Mc.negate Mc.normalise Mc.zunsat Mc.zdeduce zz_domain_spec
- [ non_linear_prover_Z "pure_sos" None ] gl
+ [ non_linear_prover_Z "pure_sos" None ]
-let sos_Q gl =
+let sos_Q =
micromega_gen parse_qarith Mc.qnegate Mc.qnormalise Mc.qunsat Mc.qdeduce qq_domain_spec
- [ non_linear_prover_Q "pure_sos" None ] gl
+ [ non_linear_prover_Q "pure_sos" None ]
-let sos_R gl =
- micromega_genr [ non_linear_prover_R "pure_sos" None ] gl
+let sos_R =
+ micromega_genr [ non_linear_prover_R "pure_sos" None ]
-let xlia gl =
+let xlia =
try
micromega_gen parse_zarith Mc.negate Mc.normalise Mc.zunsat Mc.zdeduce zz_domain_spec
- [ linear_Z ] gl
+ [ linear_Z ]
with reraise -> (*Printexc.print_backtrace stdout ;*) raise reraise
-let xnlia gl =
+let xnlia =
try
micromega_gen parse_zarith Mc.negate Mc.normalise Mc.zunsat Mc.zdeduce zz_domain_spec
- [ nlinear_Z ] gl
+ [ nlinear_Z ]
with reraise -> (*Printexc.print_backtrace stdout ;*) raise reraise
+let nra =
+ micromega_genr [ nlinear_prover_R ]
(* Local Variables: *)
diff --git a/plugins/micromega/g_micromega.ml4 b/plugins/micromega/g_micromega.ml4
index 75237aaa5..bca1c2feb 100644
--- a/plugins/micromega/g_micromega.ml4
+++ b/plugins/micromega/g_micromega.ml4
@@ -18,61 +18,57 @@
open Errors
open Misctypes
+open Stdarg
+open Constrarg
+open Pcoq.Prim
+open Pcoq.Tactic
DECLARE PLUGIN "micromega_plugin"
-let out_arg = function
- | ArgVar _ -> anomaly (Pp.str "Unevaluated or_var variable")
- | ArgArg x -> x
-
TACTIC EXTEND PsatzZ
-| [ "psatz_Z" int_or_var(i) ] -> [ Proofview.V82.tactic (Coq_micromega.psatz_Z (out_arg i)) ]
-| [ "psatz_Z" ] -> [ Proofview.V82.tactic (Coq_micromega.psatz_Z (-1)) ]
+| [ "psatz_Z" int_or_var(i) ] -> [ (Coq_micromega.psatz_Z i) ]
+| [ "psatz_Z" ] -> [ (Coq_micromega.psatz_Z (-1)) ]
END
TACTIC EXTEND Lia
-[ "xlia" ] -> [ Proofview.V82.tactic (Coq_micromega.xlia) ]
+[ "xlia" ] -> [ (Coq_micromega.xlia) ]
END
TACTIC EXTEND Nia
-[ "xnlia" ] -> [ Proofview.V82.tactic (Coq_micromega.xnlia) ]
+[ "xnlia" ] -> [ (Coq_micromega.xnlia) ]
END
-
+TACTIC EXTEND NRA
+[ "xnra" ] -> [ (Coq_micromega.nra)]
+END
TACTIC EXTEND Sos_Z
-| [ "sos_Z" ] -> [ Proofview.V82.tactic (Coq_micromega.sos_Z) ]
+| [ "sos_Z" ] -> [ (Coq_micromega.sos_Z) ]
END
TACTIC EXTEND Sos_Q
-| [ "sos_Q" ] -> [ Proofview.V82.tactic (Coq_micromega.sos_Q) ]
+| [ "sos_Q" ] -> [ (Coq_micromega.sos_Q) ]
END
TACTIC EXTEND Sos_R
-| [ "sos_R" ] -> [ Proofview.V82.tactic (Coq_micromega.sos_R) ]
-END
-
-(*
-TACTIC EXTEND Omicron
-[ "psatzl_Z" ] -> [ Proofview.V82.tactic (Coq_micromega.psatzl_Z) ]
+| [ "sos_R" ] -> [ (Coq_micromega.sos_R) ]
END
-*)
TACTIC EXTEND LRA_Q
-[ "psatzl_Q" ] -> [ Proofview.V82.tactic (Coq_micromega.psatzl_Q) ]
+[ "psatzl_Q" ] -> [ (Coq_micromega.psatzl_Q) ]
END
TACTIC EXTEND LRA_R
-[ "psatzl_R" ] -> [ Proofview.V82.tactic (Coq_micromega.psatzl_R) ]
+[ "psatzl_R" ] -> [ (Coq_micromega.psatzl_R) ]
END
TACTIC EXTEND PsatzR
-| [ "psatz_R" int_or_var(i) ] -> [ Proofview.V82.tactic (Coq_micromega.psatz_R (out_arg i)) ]
-| [ "psatz_R" ] -> [ Proofview.V82.tactic (Coq_micromega.psatz_R (-1)) ]
+| [ "psatz_R" int_or_var(i) ] -> [ (Coq_micromega.psatz_R i) ]
+| [ "psatz_R" ] -> [ (Coq_micromega.psatz_R (-1)) ]
END
TACTIC EXTEND PsatzQ
-| [ "psatz_Q" int_or_var(i) ] -> [ Proofview.V82.tactic (Coq_micromega.psatz_Q (out_arg i)) ]
-| [ "psatz_Q" ] -> [ Proofview.V82.tactic (Coq_micromega.psatz_Q (-1)) ]
+| [ "psatz_Q" int_or_var(i) ] -> [ (Coq_micromega.psatz_Q i) ]
+| [ "psatz_Q" ] -> [ (Coq_micromega.psatz_Q (-1)) ]
END
diff --git a/plugins/micromega/mfourier.ml b/plugins/micromega/mfourier.ml
index a36369d22..e22fe5843 100644
--- a/plugins/micromega/mfourier.ml
+++ b/plugins/micromega/mfourier.ml
@@ -98,12 +98,12 @@ module PSet = ISet
module System = Hashtbl.Make(Vect)
- type proof =
- | Hyp of int
- | Elim of var * proof * proof
- | And of proof * proof
-
+type proof =
+| Hyp of int
+| Elim of var * proof * proof
+| And of proof * proof
+let max_nb_cstr = ref max_int
type system = {
sys : cstr_info ref System.t ;
@@ -120,7 +120,7 @@ and cstr_info = {
(** A system of constraints has the form [\{sys = s ; vars = v\}].
[s] is a hashtable mapping a normalised vector to a [cstr_info] record where
- [bound] is an interval
- - [prf_idx] is the set of hypothese indexes (i.e. constraints in the initial system) used to obtain the current constraint.
+ - [prf_idx] is the set of hypothesis indexes (i.e. constraints in the initial system) used to obtain the current constraint.
In the initial system, each constraint is given an unique singleton proof_idx.
When a new constraint c is computed by a function f(c1,...,cn), its proof_idx is ISet.fold union (List.map (fun x -> x.proof_idx) [c1;...;cn]
- [pos] is the number of positive values of the vector
@@ -208,8 +208,7 @@ let merge_cstr_info i1 i2 =
*)
let xadd_cstr vect cstr_info sys =
- if debug && Int.equal (System.length sys mod 1000) 0 then (print_string "*" ; flush stdout) ;
- try
+ try
let info = System.find sys vect in
match merge_cstr_info cstr_info !info with
| None -> raise (SystemContradiction (And(cstr_info.prf, (!info).prf)))
@@ -217,6 +216,13 @@ let xadd_cstr vect cstr_info sys =
with
| Not_found -> System.replace sys vect (ref cstr_info)
+exception TimeOut
+
+let xadd_cstr vect cstr_info sys =
+ if debug && Int.equal (System.length sys mod 1000) 0 then (print_string "*" ; flush stdout) ;
+ if System.length sys < !max_nb_cstr
+ then xadd_cstr vect cstr_info sys
+ else raise TimeOut
type cstr_ext =
| Contradiction (** The constraint is contradictory.
@@ -866,7 +872,7 @@ let mk_proof hyps prf =
| Elim(v,prf1,prf2) ->
let prfsl = mk_proof prf1
and prfsr = mk_proof prf2 in
- (* I take only the pairs for which the elimination is meaningfull *)
+ (* I take only the pairs for which the elimination is meaningful *)
forall_pairs (pivot v) prfsl prfsr
| And(prf1,prf2) ->
let prfsl1 = mk_proof prf1
diff --git a/plugins/micromega/mutils.ml b/plugins/micromega/mutils.ml
index 2dd443f00..c13e8fc28 100644
--- a/plugins/micromega/mutils.ml
+++ b/plugins/micromega/mutils.ml
@@ -66,6 +66,15 @@ let all_sym_pairs f l =
| e::l -> xpairs (pair_with acc e l) l in
xpairs [] l
+let all_pairs f l =
+ let pair_with acc e l = List.fold_left (fun acc x -> (f e x) ::acc) acc l in
+
+ let rec xpairs acc l =
+ match l with
+ | [] -> acc
+ | e::lx -> xpairs (pair_with acc e l) lx in
+ xpairs [] l
+
let rec map3 f l1 l2 l3 =
diff --git a/plugins/nsatz/g_nsatz.ml4 b/plugins/nsatz/g_nsatz.ml4
new file mode 100644
index 000000000..0da630530
--- /dev/null
+++ b/plugins/nsatz/g_nsatz.ml4
@@ -0,0 +1,17 @@
+DECLARE PLUGIN "nsatz_plugin"
+
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i camlp4deps: "grammar/grammar.cma" i*)
+
+DECLARE PLUGIN "nsatz_plugin"
+
+TACTIC EXTEND nsatz_compute
+| [ "nsatz_compute" constr(lt) ] -> [ Proofview.V82.tactic (Nsatz.nsatz_compute lt) ]
+END
diff --git a/plugins/nsatz/nsatz.ml4 b/plugins/nsatz/nsatz.ml
index ced53d82f..ee1904a66 100644
--- a/plugins/nsatz/nsatz.ml4
+++ b/plugins/nsatz/nsatz.ml
@@ -6,8 +6,6 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i camlp4deps: "grammar/grammar.cma" i*)
-
open Errors
open Util
open Term
@@ -17,8 +15,6 @@ open Coqlib
open Num
open Utile
-DECLARE PLUGIN "nsatz_plugin"
-
(***********************************************************************
Operations on coefficients
*)
@@ -591,8 +587,4 @@ let nsatz_compute t =
error "nsatz cannot solve this problem" in
return_term lpol
-TACTIC EXTEND nsatz_compute
-| [ "nsatz_compute" constr(lt) ] -> [ Proofview.V82.tactic (nsatz_compute lt) ]
-END
-
diff --git a/plugins/nsatz/nsatz_plugin.mllib b/plugins/nsatz/nsatz_plugin.mllib
index a25e649d0..e991fb76f 100644
--- a/plugins/nsatz/nsatz_plugin.mllib
+++ b/plugins/nsatz/nsatz_plugin.mllib
@@ -2,4 +2,5 @@ Utile
Polynom
Ideal
Nsatz
+G_nsatz
Nsatz_plugin_mod
diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml
index 8a2a957cd..1f420cf6a 100644
--- a/plugins/omega/coq_omega.ml
+++ b/plugins/omega/coq_omega.ml
@@ -27,6 +27,8 @@ open Globnames
open Nametab
open Contradiction
open Misctypes
+open Proofview.Notations
+open Context.Named.Declaration
module OmegaSolver = Omega.MakeOmegaSolver (Bigint)
open OmegaSolver
@@ -34,9 +36,9 @@ open OmegaSolver
(* Added by JCF, 09/03/98 *)
let elim_id id =
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
simplest_elim (Tacmach.New.pf_global id gl)
- end
+ end }
let resolve_id id gl = Proofview.V82.of_tactic (apply (pf_global gl id)) gl
let timing timer_name f arg = f arg
@@ -926,15 +928,15 @@ let rec transform p t =
transform p
(mkApp (Lazy.force coq_Zplus,
[| t1; (mkApp (Lazy.force coq_Zopp, [| t2 |])) |])) in
- unfold sp_Zminus :: tac,t
+ Proofview.V82.of_tactic (unfold sp_Zminus) :: tac,t
| Kapp(Zsucc,[t1]) ->
let tac,t = transform p (mkApp (Lazy.force coq_Zplus,
[| t1; mk_integer one |])) in
- unfold sp_Zsucc :: tac,t
+ Proofview.V82.of_tactic (unfold sp_Zsucc) :: tac,t
| Kapp(Zpred,[t1]) ->
let tac,t = transform p (mkApp (Lazy.force coq_Zplus,
[| t1; mk_integer negone |])) in
- unfold sp_Zpred :: tac,t
+ Proofview.V82.of_tactic (unfold sp_Zpred) :: tac,t
| Kapp(Zmult,[t1;t2]) ->
let tac1,t1' = transform (P_APP 1 :: p) t1
and tac2,t2' = transform (P_APP 2 :: p) t2 in
@@ -1090,8 +1092,8 @@ let replay_history tactic_normalisation =
in
Tacticals.New.tclTHENS
(Tacticals.New.tclTHENLIST [
- Proofview.V82.tactic (unfold sp_Zle);
- Proofview.V82.tactic (simpl_in_concl);
+ unfold sp_Zle;
+ simpl_in_concl;
intro;
(absurd not_sup_sup) ])
[ assumption ; reflexivity ]
@@ -1134,10 +1136,10 @@ let replay_history tactic_normalisation =
(intros_using [id]);
(loop l) ];
Tacticals.New.tclTHENLIST [
- (Proofview.V82.tactic (unfold sp_Zgt));
- (Proofview.V82.tactic simpl_in_concl);
+ (unfold sp_Zgt);
+ simpl_in_concl;
reflexivity ] ];
- Tacticals.New.tclTHENLIST [ Proofview.V82.tactic (unfold sp_Zgt); Proofview.V82.tactic simpl_in_concl; reflexivity ]
+ Tacticals.New.tclTHENLIST [ unfold sp_Zgt; simpl_in_concl; reflexivity ]
];
Tacticals.New.tclTHEN (Proofview.V82.tactic (mk_then tac)) reflexivity ]
@@ -1159,18 +1161,18 @@ let replay_history tactic_normalisation =
[mkApp (Lazy.force coq_OMEGA4,
[| dd;kk;eq2;mkVar aux1; mkVar aux2 |])]);
Proofview.V82.tactic (clear [aux1;aux2]);
- Proofview.V82.tactic (unfold sp_not);
+ unfold sp_not;
(intros_using [aux]);
Proofview.V82.tactic (resolve_id aux);
Proofview.V82.tactic (mk_then tac);
assumption ] ;
Tacticals.New.tclTHENLIST [
- Proofview.V82.tactic (unfold sp_Zgt);
- Proofview.V82.tactic simpl_in_concl;
+ unfold sp_Zgt;
+ simpl_in_concl;
reflexivity ] ];
Tacticals.New.tclTHENLIST [
- Proofview.V82.tactic (unfold sp_Zgt);
- Proofview.V82.tactic simpl_in_concl;
+ unfold sp_Zgt;
+ simpl_in_concl;
reflexivity ] ]
| EXACT_DIVIDE (e1,k) :: l ->
let id = hyp_of_tag e1.id in
@@ -1207,8 +1209,8 @@ let replay_history tactic_normalisation =
(intros_using [id]);
(loop l) ];
Tacticals.New.tclTHENLIST [
- Proofview.V82.tactic (unfold sp_Zgt);
- Proofview.V82.tactic simpl_in_concl;
+ unfold sp_Zgt;
+ simpl_in_concl;
reflexivity ] ];
Tacticals.New.tclTHEN (Proofview.V82.tactic (mk_then tac)) reflexivity ]
| (MERGE_EQ(e3,e1,e2)) :: l ->
@@ -1328,12 +1330,12 @@ let replay_history tactic_normalisation =
(intros_using [id]);
(loop l) ];
Tacticals.New.tclTHENLIST [
- Proofview.V82.tactic (unfold sp_Zgt);
- Proofview.V82.tactic simpl_in_concl;
+ unfold sp_Zgt;
+ simpl_in_concl;
reflexivity ] ];
Tacticals.New.tclTHENLIST [
- Proofview.V82.tactic (unfold sp_Zgt);
- Proofview.V82.tactic simpl_in_concl;
+ unfold sp_Zgt;
+ simpl_in_concl;
reflexivity ] ]
| CONSTANT_NOT_NUL(e,k) :: l ->
Tacticals.New.tclTHEN (Proofview.V82.tactic (generalize_tac [mkVar (hyp_of_tag e)])) Equality.discrConcl
@@ -1342,9 +1344,9 @@ let replay_history tactic_normalisation =
| CONSTANT_NEG(e,k) :: l ->
Tacticals.New.tclTHENLIST [
Proofview.V82.tactic (generalize_tac [mkVar (hyp_of_tag e)]);
- Proofview.V82.tactic (unfold sp_Zle);
- Proofview.V82.tactic simpl_in_concl;
- Proofview.V82.tactic (unfold sp_not);
+ unfold sp_Zle;
+ simpl_in_concl;
+ unfold sp_not;
(intros_using [aux]);
Proofview.V82.tactic (resolve_id aux);
reflexivity
@@ -1416,7 +1418,7 @@ let reintroduce id =
open Proofview.Notations
let coq_omega =
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
clear_constr_tables ();
let hyps_types = Tacmach.New.pf_hyps_types gl in
let destructure_omega = Tacmach.New.of_old destructure_omega gl in
@@ -1464,12 +1466,12 @@ let coq_omega =
Tacticals.New.tclTHEN prelude (replay_history tactic_normalisation path)
with NO_CONTRADICTION -> Tacticals.New.tclZEROMSG (Pp.str"Omega can't solve this system")
end
- end
+ end }
let coq_omega = coq_omega
let nat_inject =
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
let is_conv = Tacmach.New.pf_apply Reductionops.is_conv gl in
let rec explore p t : unit Proofview.tactic =
try match destructurate_term t with
@@ -1603,7 +1605,7 @@ let nat_inject =
in
let hyps_types = Tacmach.New.pf_hyps_types gl in
loop (List.rev hyps_types)
- end
+ end }
let dec_binop = function
| Zne -> coq_dec_Zne
@@ -1673,46 +1675,47 @@ let onClearedName id tac =
(* so renaming may be necessary *)
Tacticals.New.tclTHEN
(Proofview.V82.tactic (tclTRY (clear [id])))
- (Proofview.Goal.nf_enter begin fun gl ->
+ (Proofview.Goal.nf_enter { enter = begin fun gl ->
let id = Tacmach.New.of_old (fresh_id [] id) gl in
Tacticals.New.tclTHEN (introduction id) (tac id)
- end)
+ end })
let onClearedName2 id tac =
Tacticals.New.tclTHEN
(Proofview.V82.tactic (tclTRY (clear [id])))
- (Proofview.Goal.nf_enter begin fun gl ->
+ (Proofview.Goal.nf_enter { enter = begin fun gl ->
let id1 = Tacmach.New.of_old (fresh_id [] (add_suffix id "_left")) gl in
let id2 = Tacmach.New.of_old (fresh_id [] (add_suffix id "_right")) gl in
Tacticals.New.tclTHENLIST [ introduction id1; introduction id2; tac id1 id2 ]
- end)
+ end })
let destructure_hyps =
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
let type_of = Tacmach.New.pf_unsafe_type_of gl in
let decidability = Tacmach.New.of_old decidability gl in
let pf_nf = Tacmach.New.of_old pf_nf gl in
let rec loop = function
| [] -> (Tacticals.New.tclTHEN nat_inject coq_omega)
- | (i,body,t)::lit ->
+ | decl::lit ->
+ let (i,_,t) = to_tuple decl in
begin try match destructurate_prop t with
| Kapp(False,[]) -> elim_id i
| Kapp((Zle|Zge|Zgt|Zlt|Zne),[t1;t2]) -> loop lit
| Kapp(Or,[t1;t2]) ->
(Tacticals.New.tclTHENS
(elim_id i)
- [ onClearedName i (fun i -> (loop ((i,None,t1)::lit)));
- onClearedName i (fun i -> (loop ((i,None,t2)::lit))) ])
+ [ onClearedName i (fun i -> (loop (LocalAssum (i,t1)::lit)));
+ onClearedName i (fun i -> (loop (LocalAssum (i,t2)::lit))) ])
| Kapp(And,[t1;t2]) ->
Tacticals.New.tclTHEN
(elim_id i)
(onClearedName2 i (fun i1 i2 ->
- loop ((i1,None,t1)::(i2,None,t2)::lit)))
+ loop (LocalAssum (i1,t1) :: LocalAssum (i2,t2) :: lit)))
| Kapp(Iff,[t1;t2]) ->
Tacticals.New.tclTHEN
(elim_id i)
(onClearedName2 i (fun i1 i2 ->
- loop ((i1,None,mkArrow t1 t2)::(i2,None,mkArrow t2 t1)::lit)))
+ loop (LocalAssum (i1,mkArrow t1 t2) :: LocalAssum (i2,mkArrow t2 t1) :: lit)))
| Kimp(t1,t2) ->
(* t1 and t2 might be in Type rather than Prop.
For t1, the decidability check will ensure being Prop. *)
@@ -1723,7 +1726,7 @@ let destructure_hyps =
Proofview.V82.tactic (generalize_tac [mkApp (Lazy.force coq_imp_simp,
[| t1; t2; d1; mkVar i|])]);
(onClearedName i (fun i ->
- (loop ((i,None,mk_or (mk_not t1) t2)::lit))))
+ (loop (LocalAssum (i,mk_or (mk_not t1) t2) :: lit))))
]
else
loop lit
@@ -1734,7 +1737,7 @@ let destructure_hyps =
Proofview.V82.tactic (generalize_tac
[mkApp (Lazy.force coq_not_or,[| t1; t2; mkVar i |])]);
(onClearedName i (fun i ->
- (loop ((i,None,mk_and (mk_not t1) (mk_not t2)):: lit))))
+ (loop (LocalAssum (i,mk_and (mk_not t1) (mk_not t2)) :: lit))))
]
| Kapp(And,[t1;t2]) ->
let d1 = decidability t1 in
@@ -1743,7 +1746,7 @@ let destructure_hyps =
[mkApp (Lazy.force coq_not_and,
[| t1; t2; d1; mkVar i |])]);
(onClearedName i (fun i ->
- (loop ((i,None,mk_or (mk_not t1) (mk_not t2))::lit))))
+ (loop (LocalAssum (i,mk_or (mk_not t1) (mk_not t2)) :: lit))))
]
| Kapp(Iff,[t1;t2]) ->
let d1 = decidability t1 in
@@ -1753,9 +1756,8 @@ let destructure_hyps =
[mkApp (Lazy.force coq_not_iff,
[| t1; t2; d1; d2; mkVar i |])]);
(onClearedName i (fun i ->
- (loop ((i,None,
- mk_or (mk_and t1 (mk_not t2))
- (mk_and (mk_not t1) t2))::lit))))
+ (loop (LocalAssum (i, mk_or (mk_and t1 (mk_not t2))
+ (mk_and (mk_not t1) t2)) :: lit))))
]
| Kimp(t1,t2) ->
(* t2 must be in Prop otherwise ~(t1->t2) wouldn't be ok.
@@ -1766,14 +1768,14 @@ let destructure_hyps =
[mkApp (Lazy.force coq_not_imp,
[| t1; t2; d1; mkVar i |])]);
(onClearedName i (fun i ->
- (loop ((i,None,mk_and t1 (mk_not t2)) :: lit))))
+ (loop (LocalAssum (i,mk_and t1 (mk_not t2)) :: lit))))
]
| Kapp(Not,[t]) ->
let d = decidability t in
Tacticals.New.tclTHENLIST [
Proofview.V82.tactic (generalize_tac
[mkApp (Lazy.force coq_not_not, [| t; d; mkVar i |])]);
- (onClearedName i (fun i -> (loop ((i,None,t)::lit))))
+ (onClearedName i (fun i -> (loop (LocalAssum (i,t) :: lit))))
]
| Kapp(op,[t1;t2]) ->
(try
@@ -1806,15 +1808,13 @@ let destructure_hyps =
match destructurate_type (pf_nf typ) with
| Kapp(Nat,_) ->
(Tacticals.New.tclTHEN
- (convert_hyp_no_check
- (i,body,
- (mkApp (Lazy.force coq_neq, [| t1;t2|]))))
+ (convert_hyp_no_check (set_type (mkApp (Lazy.force coq_neq, [| t1;t2|]))
+ decl))
(loop lit))
| Kapp(Z,_) ->
(Tacticals.New.tclTHEN
- (convert_hyp_no_check
- (i,body,
- (mkApp (Lazy.force coq_Zne, [| t1;t2|]))))
+ (convert_hyp_no_check (set_type (mkApp (Lazy.force coq_Zne, [| t1;t2|]))
+ decl))
(loop lit))
| _ -> loop lit
end
@@ -1828,17 +1828,17 @@ let destructure_hyps =
in
let hyps = Proofview.Goal.hyps gl in
loop hyps
- end
+ end }
let destructure_goal =
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
let concl = Proofview.Goal.concl gl in
let decidability = Tacmach.New.of_old decidability gl in
let rec loop t =
match destructurate_prop t with
| Kapp(Not,[t]) ->
(Tacticals.New.tclTHEN
- (Tacticals.New.tclTHEN (Proofview.V82.tactic (unfold sp_not)) intro)
+ (Tacticals.New.tclTHEN (unfold sp_not) intro)
destructure_hyps)
| Kimp(a,b) -> (Tacticals.New.tclTHEN intro (loop b))
| Kapp(False,[]) -> destructure_hyps
@@ -1855,7 +1855,7 @@ let destructure_goal =
Tacticals.New.tclTHEN goal_tac destructure_hyps
in
(loop concl)
- end
+ end }
let destructure_goal = destructure_goal
diff --git a/plugins/omega/g_omega.ml4 b/plugins/omega/g_omega.ml4
index c96b4a4f4..b314e0d85 100644
--- a/plugins/omega/g_omega.ml4
+++ b/plugins/omega/g_omega.ml4
@@ -17,15 +17,24 @@
DECLARE PLUGIN "omega_plugin"
+open Names
open Coq_omega
+open Constrarg
+open Pcoq.Prim
+
+let eval_tactic name =
+ let dp = DirPath.make (List.map Id.of_string ["PreOmega"; "omega"; "Coq"]) in
+ let kn = KerName.make2 (MPfile dp) (Label.make name) in
+ let tac = Tacenv.interp_ltac kn in
+ Tacinterp.eval_tactic tac
let omega_tactic l =
let tacs = List.map
(function
- | "nat" -> Tacinterp.interp <:tactic<zify_nat>>
- | "positive" -> Tacinterp.interp <:tactic<zify_positive>>
- | "N" -> Tacinterp.interp <:tactic<zify_N>>
- | "Z" -> Tacinterp.interp <:tactic<zify_op>>
+ | "nat" -> eval_tactic "zify_nat"
+ | "positive" -> eval_tactic "zify_positive"
+ | "N" -> eval_tactic "zify_N"
+ | "Z" -> eval_tactic "zify_op"
| s -> Errors.error ("No Omega knowledge base for type "^s))
(Util.List.sort_uniquize String.compare l)
in
diff --git a/plugins/quote/g_quote.ml4 b/plugins/quote/g_quote.ml4
index fdc5c2bbd..a15b0eb05 100644
--- a/plugins/quote/g_quote.ml4
+++ b/plugins/quote/g_quote.ml4
@@ -13,19 +13,22 @@ open Misctypes
open Tacexpr
open Geninterp
open Quote
+open Constrarg
+open Pcoq.Prim
+open Pcoq.Constr
+open Pcoq.Tactic
DECLARE PLUGIN "quote_plugin"
let loc = Loc.ghost
-let cont = (loc, Id.of_string "cont")
-let x = (loc, Id.of_string "x")
+let cont = Id.of_string "cont"
+let x = Id.of_string "x"
-let make_cont (k : glob_tactic_expr) (c : Constr.t) =
+let make_cont (k : Genarg.Val.t) (c : Constr.t) =
let c = Tacinterp.Value.of_constr c in
- let tac = TacCall (loc, ArgVar cont, [Reference (ArgVar x)]) in
- let tac = TacLetIn (false, [(cont, Tacexp k)], TacArg (loc, tac)) in
- let ist = { lfun = Id.Map.singleton (snd x) c; extra = TacStore.empty; } in
- Tacinterp.eval_tactic_ist ist tac
+ let tac = TacCall (loc, ArgVar (loc, cont), [Reference (ArgVar (loc, x))]) in
+ let ist = { lfun = Id.Map.add cont k (Id.Map.singleton x c); extra = TacStore.empty; } in
+ Tacinterp.eval_tactic_ist ist (TacArg (loc, tac))
TACTIC EXTEND quote
[ "quote" ident(f) ] -> [ quote f [] ]
diff --git a/plugins/quote/quote.ml b/plugins/quote/quote.ml
index ff6acf139..dbd7460e2 100644
--- a/plugins/quote/quote.ml
+++ b/plugins/quote/quote.ml
@@ -109,6 +109,7 @@ open Pattern
open Patternops
open Constr_matching
open Tacmach
+open Proofview.Notations
(*i*)
(*s First, we need to access some Coq constants
@@ -227,7 +228,7 @@ let compute_ivs f cs gl =
let (args3, body3) = decompose_lam body2 in
let nargs3 = List.length args3 in
let env = Proofview.Goal.env gl in
- let sigma = Proofview.Goal.sigma gl in
+ let sigma = Tacmach.New.project gl in
let is_conv = Reductionops.is_conv env sigma in
begin match decomp_term body3 with
| Case(_,p,c,lci) -> (* <p> Case c of c1 ... cn end *)
@@ -446,7 +447,7 @@ let quote_terms ivs lc =
yet. *)
let quote f lid =
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
let f = Tacmach.New.pf_global f gl in
let cl = List.map (fun id -> Tacmach.New.pf_global id gl) lid in
let ivs = compute_ivs f cl gl in
@@ -459,10 +460,10 @@ let quote f lid =
match ivs.variable_lhs with
| None -> Tactics.convert_concl (mkApp (f, [| p |])) DEFAULTcast
| Some _ -> Tactics.convert_concl (mkApp (f, [| vm; p |])) DEFAULTcast
- end
+ end }
let gen_quote cont c f lid =
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
let f = Tacmach.New.pf_global f gl in
let cl = List.map (fun id -> Tacmach.New.pf_global id gl) lid in
let ivs = compute_ivs f cl gl in
@@ -474,7 +475,7 @@ let gen_quote cont c f lid =
match ivs.variable_lhs with
| None -> cont (mkApp (f, [| p |]))
| Some _ -> cont (mkApp (f, [| vm; p |]))
- end
+ end }
(*i
diff --git a/plugins/romega/ReflOmegaCore.v b/plugins/romega/ReflOmegaCore.v
index b84cf2540..36511386a 100644
--- a/plugins/romega/ReflOmegaCore.v
+++ b/plugins/romega/ReflOmegaCore.v
@@ -1492,7 +1492,7 @@ with Simplify := match goal with
end.
Ltac prove_stable x th :=
- match constr:x with
+ match constr:(x) with
| ?X1 =>
unfold term_stable, X1; intros; Simplify; simpl;
apply th
diff --git a/plugins/romega/g_romega.ml4 b/plugins/romega/g_romega.ml4
index 0a99a26b3..61efa9f54 100644
--- a/plugins/romega/g_romega.ml4
+++ b/plugins/romega/g_romega.ml4
@@ -10,15 +10,24 @@
DECLARE PLUGIN "romega_plugin"
+open Names
open Refl_omega
+open Constrarg
+open Pcoq.Prim
+
+let eval_tactic name =
+ let dp = DirPath.make (List.map Id.of_string ["PreOmega"; "omega"; "Coq"]) in
+ let kn = KerName.make2 (MPfile dp) (Label.make name) in
+ let tac = Tacenv.interp_ltac kn in
+ Tacinterp.eval_tactic tac
let romega_tactic l =
let tacs = List.map
(function
- | "nat" -> Tacinterp.interp <:tactic<zify_nat>>
- | "positive" -> Tacinterp.interp <:tactic<zify_positive>>
- | "N" -> Tacinterp.interp <:tactic<zify_N>>
- | "Z" -> Tacinterp.interp <:tactic<zify_op>>
+ | "nat" -> eval_tactic "zify_nat"
+ | "positive" -> eval_tactic "zify_positive"
+ | "N" -> eval_tactic "zify_N"
+ | "Z" -> eval_tactic "zify_op"
| s -> Errors.error ("No ROmega knowledge base for type "^s))
(Util.List.sort_uniquize String.compare l)
in
diff --git a/plugins/romega/refl_omega.ml b/plugins/romega/refl_omega.ml
index 560e6a899..177c870b3 100644
--- a/plugins/romega/refl_omega.ml
+++ b/plugins/romega/refl_omega.ml
@@ -1285,7 +1285,7 @@ let resolution env full_reified_goal systems_list =
Proofview.V82.of_tactic (Tactics.change_concl reified) >>
Proofview.V82.of_tactic (Tactics.apply (app coq_do_omega [|decompose_tactic; normalization_trace|])) >>
show_goal >>
- Tactics.normalise_vm_in_concl >>
+ Proofview.V82.of_tactic (Tactics.normalise_vm_in_concl) >>
(*i Alternatives to the previous line:
- Normalisation without VM:
Tactics.normalise_in_concl
diff --git a/plugins/rtauto/refl_tauto.ml b/plugins/rtauto/refl_tauto.ml
index 9c22b5adb..2f3a3e551 100644
--- a/plugins/rtauto/refl_tauto.ml
+++ b/plugins/rtauto/refl_tauto.ml
@@ -13,6 +13,7 @@ open Util
open Term
open Tacmach
open Proof_search
+open Context.Named.Declaration
let force count lazc = incr count;Lazy.force lazc
@@ -128,9 +129,9 @@ let rec make_form atom_env gls term =
let rec make_hyps atom_env gls lenv = function
[] -> []
- | (_,Some body,typ)::rest ->
+ | LocalDef (_,body,typ)::rest ->
make_hyps atom_env gls (typ::body::lenv) rest
- | (id,None,typ)::rest ->
+ | LocalAssum (id,typ)::rest ->
let hrec=
make_hyps atom_env gls (typ::lenv) rest in
if List.exists (Termops.dependent (mkVar id)) lenv ||
diff --git a/plugins/rtauto/refl_tauto.mli b/plugins/rtauto/refl_tauto.mli
index c9e591bbd..9a14ac6c7 100644
--- a/plugins/rtauto/refl_tauto.mli
+++ b/plugins/rtauto/refl_tauto.mli
@@ -18,7 +18,7 @@ val make_hyps :
atom_env ->
Proof_type.goal Tacmach.sigma ->
Term.types list ->
- (Names.Id.t * Term.types option * Term.types) list ->
+ Context.Named.t ->
(Names.Id.t * Proof_search.form) list
val rtauto_tac : Proof_type.tactic
diff --git a/plugins/setoid_ring/ArithRing.v b/plugins/setoid_ring/ArithRing.v
index 04decbce1..5f5b97925 100644
--- a/plugins/setoid_ring/ArithRing.v
+++ b/plugins/setoid_ring/ArithRing.v
@@ -32,13 +32,13 @@ Qed.
Ltac natcst t :=
match isnatcst t with
true => constr:(N.of_nat t)
- | _ => constr:InitialRing.NotConstant
+ | _ => constr:(InitialRing.NotConstant)
end.
Ltac Ss_to_add f acc :=
match f with
| S ?f1 => Ss_to_add f1 (S acc)
- | _ => constr:(acc + f)%nat
+ | _ => constr:((acc + f)%nat)
end.
Ltac natprering :=
diff --git a/plugins/setoid_ring/InitialRing.v b/plugins/setoid_ring/InitialRing.v
index 8362c8c26..8fcc07716 100644
--- a/plugins/setoid_ring/InitialRing.v
+++ b/plugins/setoid_ring/InitialRing.v
@@ -612,32 +612,32 @@ End GEN_DIV.
Ltac inv_gen_phi_pos rI add mul t :=
let rec inv_cst t :=
match t with
- rI => constr:1%positive
- | (add rI rI) => constr:2%positive
- | (add rI (add rI rI)) => constr:3%positive
+ rI => constr:(1%positive)
+ | (add rI rI) => constr:(2%positive)
+ | (add rI (add rI rI)) => constr:(3%positive)
| (mul (add rI rI) ?p) => (* 2p *)
match inv_cst p with
- NotConstant => constr:NotConstant
- | 1%positive => constr:NotConstant (* 2*1 is not convertible to 2 *)
+ NotConstant => constr:(NotConstant)
+ | 1%positive => constr:(NotConstant) (* 2*1 is not convertible to 2 *)
| ?p => constr:(xO p)
end
| (add rI (mul (add rI rI) ?p)) => (* 1+2p *)
match inv_cst p with
- NotConstant => constr:NotConstant
- | 1%positive => constr:NotConstant
+ NotConstant => constr:(NotConstant)
+ | 1%positive => constr:(NotConstant)
| ?p => constr:(xI p)
end
- | _ => constr:NotConstant
+ | _ => constr:(NotConstant)
end in
inv_cst t.
(* The (partial) inverse of gen_phiNword *)
Ltac inv_gen_phiNword rO rI add mul opp t :=
match t with
- rO => constr:NwO
+ rO => constr:(NwO)
| _ =>
match inv_gen_phi_pos rI add mul t with
- NotConstant => constr:NotConstant
+ NotConstant => constr:(NotConstant)
| ?p => constr:(Npos p::nil)
end
end.
@@ -646,10 +646,10 @@ End GEN_DIV.
(* The inverse of gen_phiN *)
Ltac inv_gen_phiN rO rI add mul t :=
match t with
- rO => constr:0%N
+ rO => constr:(0%N)
| _ =>
match inv_gen_phi_pos rI add mul t with
- NotConstant => constr:NotConstant
+ NotConstant => constr:(NotConstant)
| ?p => constr:(Npos p)
end
end.
@@ -657,15 +657,15 @@ End GEN_DIV.
(* The inverse of gen_phiZ *)
Ltac inv_gen_phiZ rO rI add mul opp t :=
match t with
- rO => constr:0%Z
+ rO => constr:(0%Z)
| (opp ?p) =>
match inv_gen_phi_pos rI add mul p with
- NotConstant => constr:NotConstant
+ NotConstant => constr:(NotConstant)
| ?p => constr:(Zneg p)
end
| _ =>
match inv_gen_phi_pos rI add mul t with
- NotConstant => constr:NotConstant
+ NotConstant => constr:(NotConstant)
| ?p => constr:(Zpos p)
end
end.
@@ -681,7 +681,7 @@ Ltac inv_gen_phi rO rI cO cI t :=
end.
(* A simple tactic recognizing no constant *)
- Ltac inv_morph_nothing t := constr:NotConstant.
+ Ltac inv_morph_nothing t := constr:(NotConstant).
Ltac coerce_to_almost_ring set ext rspec :=
match type of rspec with
@@ -825,31 +825,31 @@ Ltac ring_elements set ext rspec pspec sspec dspec rk :=
(* Tactic for constant *)
Ltac isnatcst t :=
match t with
- O => constr:true
+ O => constr:(true)
| S ?p => isnatcst p
- | _ => constr:false
+ | _ => constr:(false)
end.
Ltac isPcst t :=
match t with
| xI ?p => isPcst p
| xO ?p => isPcst p
- | xH => constr:true
+ | xH => constr:(true)
(* nat -> positive *)
| Pos.of_succ_nat ?n => isnatcst n
- | _ => constr:false
+ | _ => constr:(false)
end.
Ltac isNcst t :=
match t with
- N0 => constr:true
+ N0 => constr:(true)
| Npos ?p => isPcst p
- | _ => constr:false
+ | _ => constr:(false)
end.
Ltac isZcst t :=
match t with
- Z0 => constr:true
+ Z0 => constr:(true)
| Zpos ?p => isPcst p
| Zneg ?p => isPcst p
(* injection nat -> Z *)
@@ -857,7 +857,7 @@ Ltac isZcst t :=
(* injection N -> Z *)
| Z.of_N ?n => isNcst n
(* *)
- | _ => constr:false
+ | _ => constr:(false)
end.
diff --git a/plugins/setoid_ring/NArithRing.v b/plugins/setoid_ring/NArithRing.v
index 6c1a79e4e..54e2789ba 100644
--- a/plugins/setoid_ring/NArithRing.v
+++ b/plugins/setoid_ring/NArithRing.v
@@ -15,7 +15,7 @@ Set Implicit Arguments.
Ltac Ncst t :=
match isNcst t with
true => t
- | _ => constr:NotConstant
+ | _ => constr:(NotConstant)
end.
Add Ring Nr : Nth (decidable Neqb_ok, constants [Ncst]).
diff --git a/plugins/setoid_ring/Ring.v b/plugins/setoid_ring/Ring.v
index a0844100c..77576cb93 100644
--- a/plugins/setoid_ring/Ring.v
+++ b/plugins/setoid_ring/Ring.v
@@ -36,9 +36,9 @@ Qed.
Ltac bool_cst t :=
let t := eval hnf in t in
match t with
- true => constr:true
- | false => constr:false
- | _ => constr:NotConstant
+ true => constr:(true)
+ | false => constr:(false)
+ | _ => constr:(NotConstant)
end.
Add Ring bool_ring : BoolTheory (decidable bool_eq_ok, constants [bool_cst]).
diff --git a/plugins/setoid_ring/ZArithRing.v b/plugins/setoid_ring/ZArithRing.v
index 914843727..23784cf33 100644
--- a/plugins/setoid_ring/ZArithRing.v
+++ b/plugins/setoid_ring/ZArithRing.v
@@ -17,14 +17,14 @@ Set Implicit Arguments.
Ltac Zcst t :=
match isZcst t with
true => t
- | _ => constr:NotConstant
+ | _ => constr:(NotConstant)
end.
Ltac isZpow_coef t :=
match t with
| Zpos ?p => isPcst p
- | Z0 => constr:true
- | _ => constr:false
+ | Z0 => constr:(true)
+ | _ => constr:(false)
end.
Notation N_of_Z := Z.to_N (only parsing).
@@ -32,7 +32,7 @@ Notation N_of_Z := Z.to_N (only parsing).
Ltac Zpow_tac t :=
match isZpow_coef t with
| true => constr:(N_of_Z t)
- | _ => constr:NotConstant
+ | _ => constr:(NotConstant)
end.
Ltac Zpower_neg :=
diff --git a/plugins/setoid_ring/g_newring.ml4 b/plugins/setoid_ring/g_newring.ml4
new file mode 100644
index 000000000..1ebb6e6b7
--- /dev/null
+++ b/plugins/setoid_ring/g_newring.ml4
@@ -0,0 +1,105 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i camlp4deps: "grammar/grammar.cma" i*)
+
+open Pp
+open Util
+open Libnames
+open Printer
+open Newring_ast
+open Newring
+open Stdarg
+open Constrarg
+open Pcoq.Prim
+open Pcoq.Constr
+open Pcoq.Tactic
+
+DECLARE PLUGIN "newring_plugin"
+
+TACTIC EXTEND protect_fv
+ [ "protect_fv" string(map) "in" ident(id) ] ->
+ [ Proofview.V82.tactic (protect_tac_in map id) ]
+| [ "protect_fv" string(map) ] ->
+ [ Proofview.V82.tactic (protect_tac map) ]
+END
+
+TACTIC EXTEND closed_term
+ [ "closed_term" constr(t) "[" ne_reference_list(l) "]" ] ->
+ [ Proofview.V82.tactic (closed_term t l) ]
+END
+
+VERNAC ARGUMENT EXTEND ring_mod
+ | [ "decidable" constr(eq_test) ] -> [ Ring_kind(Computational eq_test) ]
+ | [ "abstract" ] -> [ Ring_kind Abstract ]
+ | [ "morphism" constr(morph) ] -> [ Ring_kind(Morphism morph) ]
+ | [ "constants" "[" tactic(cst_tac) "]" ] -> [ Const_tac(CstTac cst_tac) ]
+ | [ "closed" "[" ne_global_list(l) "]" ] -> [ Const_tac(Closed l) ]
+ | [ "preprocess" "[" tactic(pre) "]" ] -> [ Pre_tac pre ]
+ | [ "postprocess" "[" tactic(post) "]" ] -> [ Post_tac post ]
+ | [ "setoid" constr(sth) constr(ext) ] -> [ Setoid(sth,ext) ]
+ | [ "sign" constr(sign_spec) ] -> [ Sign_spec sign_spec ]
+ | [ "power" constr(pow_spec) "[" ne_global_list(l) "]" ] ->
+ [ Pow_spec (Closed l, pow_spec) ]
+ | [ "power_tac" constr(pow_spec) "[" tactic(cst_tac) "]" ] ->
+ [ Pow_spec (CstTac cst_tac, pow_spec) ]
+ | [ "div" constr(div_spec) ] -> [ Div_spec div_spec ]
+END
+
+VERNAC ARGUMENT EXTEND ring_mods
+ | [ "(" ne_ring_mod_list_sep(mods, ",") ")" ] -> [ mods ]
+END
+
+VERNAC COMMAND EXTEND AddSetoidRing CLASSIFIED AS SIDEFF
+ | [ "Add" "Ring" ident(id) ":" constr(t) ring_mods_opt(l) ] ->
+ [ let l = match l with None -> [] | Some l -> l in
+ let (k,set,cst,pre,post,power,sign, div) = process_ring_mods l in
+ add_theory id (ic t) set k cst (pre,post) power sign div]
+ | [ "Print" "Rings" ] => [Vernac_classifier.classify_as_query] -> [
+ msg_notice (strbrk "The following ring structures have been declared:");
+ Spmap.iter (fun fn fi ->
+ msg_notice (hov 2
+ (Ppconstr.pr_id (Libnames.basename fn)++spc()++
+ str"with carrier "++ pr_constr fi.ring_carrier++spc()++
+ str"and equivalence relation "++ pr_constr fi.ring_req))
+ ) !from_name ]
+END
+
+TACTIC EXTEND ring_lookup
+| [ "ring_lookup" tactic0(f) "[" constr_list(lH) "]" ne_constr_list(lrt) ] ->
+ [ let (t,lr) = List.sep_last lrt in ring_lookup f lH lr t]
+END
+
+VERNAC ARGUMENT EXTEND field_mod
+ | [ ring_mod(m) ] -> [ Ring_mod m ]
+ | [ "completeness" constr(inj) ] -> [ Inject inj ]
+END
+
+VERNAC ARGUMENT EXTEND field_mods
+ | [ "(" ne_field_mod_list_sep(mods, ",") ")" ] -> [ mods ]
+END
+
+VERNAC COMMAND EXTEND AddSetoidField CLASSIFIED AS SIDEFF
+| [ "Add" "Field" ident(id) ":" constr(t) field_mods_opt(l) ] ->
+ [ let l = match l with None -> [] | Some l -> l in
+ let (k,set,inj,cst_tac,pre,post,power,sign,div) = process_field_mods l in
+ add_field_theory id (ic t) set k cst_tac inj (pre,post) power sign div]
+| [ "Print" "Fields" ] => [Vernac_classifier.classify_as_query] -> [
+ msg_notice (strbrk "The following field structures have been declared:");
+ Spmap.iter (fun fn fi ->
+ msg_notice (hov 2
+ (Ppconstr.pr_id (Libnames.basename fn)++spc()++
+ str"with carrier "++ pr_constr fi.field_carrier++spc()++
+ str"and equivalence relation "++ pr_constr fi.field_req))
+ ) !field_from_name ]
+END
+
+TACTIC EXTEND field_lookup
+| [ "field_lookup" tactic(f) "[" constr_list(lH) "]" ne_constr_list(lt) ] ->
+ [ let (t,l) = List.sep_last lt in field_lookup f lH l t ]
+END
diff --git a/plugins/setoid_ring/newring.ml4 b/plugins/setoid_ring/newring.ml
index e704c466e..7ef89b7a0 100644
--- a/plugins/setoid_ring/newring.ml4
+++ b/plugins/setoid_ring/newring.ml
@@ -6,8 +6,6 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i camlp4deps: "grammar/grammar.cma" i*)
-
open Pp
open Errors
open Util
@@ -30,8 +28,8 @@ open Declare
open Decl_kinds
open Entries
open Misctypes
-
-DECLARE PLUGIN "newring_plugin"
+open Newring_ast
+open Proofview.Notations
(****************************************************************************)
(* controlled reduction *)
@@ -99,19 +97,12 @@ let protect_red map env sigma c =
(mk_clos_but (lookup_map map c) (Esubst.subs_id 0) c);;
let protect_tac map =
- Tactics.reduct_option (protect_red map,DEFAULTcast) None ;;
+ Proofview.V82.of_tactic (Tactics.reduct_option (protect_red map,DEFAULTcast) None);;
let protect_tac_in map id =
- Tactics.reduct_option (protect_red map,DEFAULTcast) (Some(id, Locus.InHyp));;
+ Proofview.V82.of_tactic (Tactics.reduct_option (protect_red map,DEFAULTcast) (Some(id, Locus.InHyp)));;
-TACTIC EXTEND protect_fv
- [ "protect_fv" string(map) "in" ident(id) ] ->
- [ Proofview.V82.tactic (protect_tac_in map id) ]
-| [ "protect_fv" string(map) ] ->
- [ Proofview.V82.tactic (protect_tac map) ]
-END;;
-
(****************************************************************************)
let closed_term t l =
@@ -120,12 +111,6 @@ let closed_term t l =
if Quote.closed_under cs t then tclIDTAC else tclFAIL 0 (mt())
;;
-TACTIC EXTEND closed_term
- [ "closed_term" constr(t) "[" ne_reference_list(l) "]" ] ->
- [ Proofview.V82.tactic (closed_term t l) ]
-END
-;;
-
(* TACTIC EXTEND echo
| [ "echo" constr(t) ] ->
[ Pp.msg (Termops.print_constr t); Tacinterp.eval_tactic (TacId []) ]
@@ -143,11 +128,15 @@ let closed_term_ast l =
mltac_plugin = "newring_plugin";
mltac_tactic = "closed_term";
} in
+ let tacname = {
+ mltac_name = tacname;
+ mltac_index = 0;
+ } in
let l = List.map (fun gr -> ArgArg(Loc.ghost,gr)) l in
TacFun([Some(Id.of_string"t")],
TacML(Loc.ghost,tacname,
- [Genarg.in_gen (Genarg.glbwit Constrarg.wit_constr) (GVar(Loc.ghost,Id.of_string"t"),None);
- Genarg.in_gen (Genarg.glbwit (Genarg.wit_list Constrarg.wit_ref)) l]))
+ [TacGeneric (Genarg.in_gen (Genarg.glbwit Constrarg.wit_constr) (GVar(Loc.ghost,Id.of_string"t"),None));
+ TacGeneric (Genarg.in_gen (Genarg.glbwit (Genarg.wit_list Constrarg.wit_ref)) l)]))
(*
let _ = add_tacdef false ((Loc.ghost,Id.of_string"ring_closed_term"
*)
@@ -185,19 +174,20 @@ let ltac_call tac (args:glob_tactic_arg list) =
let ltac_lcall tac args =
TacArg(Loc.ghost,TacCall(Loc.ghost, ArgVar(Loc.ghost, Id.of_string tac),args))
-let ltac_letin (x, e1) e2 =
- TacLetIn(false,[(Loc.ghost,Id.of_string x),e1],e2)
-
-let ltac_apply (f:glob_tactic_expr) (args:glob_tactic_arg list) =
- Tacinterp.eval_tactic
- (ltac_letin ("F", Tacexp f) (ltac_lcall "F" args))
+let ltac_apply (f : Value.t) (args: Tacinterp.Value.t list) =
+ let fold arg (i, vars, lfun) =
+ let id = Id.of_string ("x" ^ string_of_int i) in
+ let x = Reference (ArgVar (Loc.ghost, id)) in
+ (succ i, x :: vars, Id.Map.add id arg lfun)
+ in
+ let (_, args, lfun) = List.fold_right fold args (0, [], Id.Map.empty) in
+ let lfun = Id.Map.add (Id.of_string "F") f lfun in
+ let ist = { (Tacinterp.default_ist ()) with Tacinterp.lfun = lfun; } in
+ Tacinterp.eval_tactic_ist ist (ltac_lcall "F" args)
let ltac_record flds =
TacFun([Some(Id.of_string"proj")], ltac_lcall "proj" flds)
-
-let carg c = TacDynamic(Loc.ghost,Pretyping.constr_in c)
-
let dummy_goal env sigma =
let (gl,_,sigma) =
Goal.V82.mk_goal sigma (named_context_val env) mkProp Evd.Store.empty in
@@ -207,20 +197,39 @@ let constr_of v = match Value.to_constr v with
| Some c -> c
| None -> failwith "Ring.exec_tactic: anomaly"
+let tactic_res = ref [||]
+
+let get_res =
+ let open Tacexpr in
+ let name = { mltac_plugin = "newring_plugin"; mltac_tactic = "get_res"; } in
+ let entry = { mltac_name = name; mltac_index = 0 } in
+ let tac args ist =
+ let n = Tacinterp.Value.cast (Genarg.topwit Stdarg.wit_int) (List.hd args) in
+ let init i = Id.Map.find (Id.of_string ("x" ^ string_of_int i)) ist.lfun in
+ tactic_res := Array.init n init;
+ Proofview.tclUNIT ()
+ in
+ Tacenv.register_ml_tactic name [| tac |];
+ entry
+
let exec_tactic env evd n f args =
+ let fold arg (i, vars, lfun) =
+ let id = Id.of_string ("x" ^ string_of_int i) in
+ let x = Reference (ArgVar (Loc.ghost, id)) in
+ (succ i, x :: vars, Id.Map.add id (Value.of_constr arg) lfun)
+ in
+ let (_, args, lfun) = List.fold_right fold args (0, [], Id.Map.empty) in
+ let ist = { (Tacinterp.default_ist ()) with Tacinterp.lfun = lfun; } in
+ (** Build the getter *)
let lid = List.init n (fun i -> Id.of_string("x"^string_of_int i)) in
- let res = ref [||] in
- let get_res ist =
- let l = List.map (fun id -> Id.Map.find id ist.lfun) lid in
- res := Array.of_list l;
- TacId[] in
- let getter =
- Tacexp(TacFun(List.map(fun id -> Some id) lid,
- Tacintern.glob_tactic(tacticIn get_res))) in
+ let n = Genarg.in_gen (Genarg.glbwit Stdarg.wit_int) n in
+ let get_res = TacML (Loc.ghost, get_res, [TacGeneric n]) in
+ let getter = Tacexp (TacFun (List.map (fun id -> Some id) lid, get_res)) in
+ (** Evaluate the whole result *)
let gl = dummy_goal env evd in
- let gls = Proofview.V82.of_tactic (Tacinterp.eval_tactic(ltac_call f (args@[getter]))) gl in
+ let gls = Proofview.V82.of_tactic (Tacinterp.eval_tactic_ist ist (ltac_call f (args@[getter]))) gl in
let evd, nf = Evarutil.nf_evars_and_universes (Refiner.project gls) in
- Array.map (fun x -> nf (constr_of x)) !res, snd (Evd.universe_context evd)
+ Array.map (fun x -> nf (constr_of x)) !tactic_res, snd (Evd.universe_context evd)
let stdlib_modules =
[["Coq";"Setoids";"Setoid"];
@@ -355,20 +364,6 @@ let _ = add_map "ring"
(****************************************************************************)
(* Ring database *)
-type ring_info =
- { ring_carrier : types;
- ring_req : constr;
- ring_setoid : constr;
- ring_ext : constr;
- ring_morph : constr;
- ring_th : constr;
- ring_cst_tac : glob_tactic_expr;
- ring_pow_tac : glob_tactic_expr;
- ring_lemma1 : constr;
- ring_lemma2 : constr;
- ring_pre_tac : glob_tactic_expr;
- ring_post_tac : glob_tactic_expr }
-
module Cmap = Map.Make(Constr)
let from_carrier = Summary.ref Cmap.empty ~name:"ring-tac-carrier-table"
@@ -527,8 +522,8 @@ let ring_equality env evd (r,add,mul,opp,req) =
match opp with
Some opp -> plapp evd coq_eq_morph [|r;add;mul;opp|]
| None -> plapp evd coq_eq_smorph [|r;add;mul|] in
- let setoid = Typing.solve_evars env evd setoid in
- let op_morph = Typing.solve_evars env evd op_morph in
+ let setoid = Typing.e_solve_evars env evd setoid in
+ let op_morph = Typing.e_solve_evars env evd op_morph in
(setoid,op_morph)
| _ ->
let setoid = setoid_of_relation (Global.env ()) evd r req in
@@ -600,13 +595,6 @@ let dest_morph env sigma m_spec =
(c,czero,cone,cadd,cmul,None,None,ceqb,phi)
| _ -> error "bad morphism structure"
-
-type 'constr coeff_spec =
- Computational of 'constr (* equality test *)
- | Abstract (* coeffs = Z *)
- | Morphism of 'constr (* general morphism *)
-
-
let reflect_coeff rkind =
(* We build an ill-typed terms on purpose... *)
match rkind with
@@ -614,10 +602,6 @@ let reflect_coeff rkind =
| Computational c -> lapp coq_comp [|c|]
| Morphism m -> lapp coq_morph [|m|]
-type cst_tac_spec =
- CstTac of raw_tactic_expr
- | Closed of reference list
-
let interp_cst_tac env sigma rk kind (zero,one,add,mul,opp) cst_tac =
match cst_tac with
Some (CstTac t) -> Tacintern.glob_tactic t
@@ -638,7 +622,7 @@ let make_hyp_list env evd lH =
(fun c l -> plapp evd coq_cons [|carrier; (make_hyp env evd c); l|]) lH
(plapp evd coq_nil [|carrier|])
in
- let l' = Typing.solve_evars env evd l in
+ let l' = Typing.e_solve_evars env evd l in
Evarutil.nf_evars_universes !evd l'
let interp_power env evd pow =
@@ -686,7 +670,7 @@ let add_theory name (sigma,rth) eqth morphth cst_tac (pre,post) power sign div =
let rk = reflect_coeff morphth in
let params,ctx =
exec_tactic env !evd 5 (zltac "ring_lemmas")
- (List.map carg[sth;ext;rth;pspec;sspec;dspec;rk]) in
+ [sth;ext;rth;pspec;sspec;dspec;rk] in
let lemma1 = params.(3) in
let lemma2 = params.(4) in
@@ -721,41 +705,12 @@ let add_theory name (sigma,rth) eqth morphth cst_tac (pre,post) power sign div =
ring_post_tac = posttac }) in
()
-type 'constr ring_mod =
- Ring_kind of 'constr coeff_spec
- | Const_tac of cst_tac_spec
- | Pre_tac of raw_tactic_expr
- | Post_tac of raw_tactic_expr
- | Setoid of Constrexpr.constr_expr * Constrexpr.constr_expr
- | Pow_spec of cst_tac_spec * Constrexpr.constr_expr
- (* Syntaxification tactic , correctness lemma *)
- | Sign_spec of Constrexpr.constr_expr
- | Div_spec of Constrexpr.constr_expr
-
-
let ic_coeff_spec = function
| Computational t -> Computational (ic_unsafe t)
| Morphism t -> Morphism (ic_unsafe t)
| Abstract -> Abstract
-VERNAC ARGUMENT EXTEND ring_mod
- | [ "decidable" constr(eq_test) ] -> [ Ring_kind(Computational eq_test) ]
- | [ "abstract" ] -> [ Ring_kind Abstract ]
- | [ "morphism" constr(morph) ] -> [ Ring_kind(Morphism morph) ]
- | [ "constants" "[" tactic(cst_tac) "]" ] -> [ Const_tac(CstTac cst_tac) ]
- | [ "closed" "[" ne_global_list(l) "]" ] -> [ Const_tac(Closed l) ]
- | [ "preprocess" "[" tactic(pre) "]" ] -> [ Pre_tac pre ]
- | [ "postprocess" "[" tactic(post) "]" ] -> [ Post_tac post ]
- | [ "setoid" constr(sth) constr(ext) ] -> [ Setoid(sth,ext) ]
- | [ "sign" constr(sign_spec) ] -> [ Sign_spec sign_spec ]
- | [ "power" constr(pow_spec) "[" ne_global_list(l) "]" ] ->
- [ Pow_spec (Closed l, pow_spec) ]
- | [ "power_tac" constr(pow_spec) "[" tactic(cst_tac) "]" ] ->
- [ Pow_spec (CstTac cst_tac, pow_spec) ]
- | [ "div" constr(div_spec) ] -> [ Div_spec div_spec ]
-END
-
let set_once s r v =
if Option.is_empty !r then r := Some v else error (s^" cannot be set twice")
@@ -780,20 +735,6 @@ let process_ring_mods l =
let k = match !kind with Some k -> k | None -> Abstract in
(k, !set, !cst_tac, !pre, !post, !power, !sign, !div)
-VERNAC COMMAND EXTEND AddSetoidRing CLASSIFIED AS SIDEFF
- | [ "Add" "Ring" ident(id) ":" constr(t) ring_mods(l) ] ->
- [ let (k,set,cst,pre,post,power,sign, div) = process_ring_mods l in
- add_theory id (ic t) set k cst (pre,post) power sign div]
- | [ "Print" "Rings" ] => [Vernac_classifier.classify_as_query] -> [
- msg_notice (strbrk "The following ring structures have been declared:");
- Spmap.iter (fun fn fi ->
- msg_notice (hov 2
- (Ppconstr.pr_id (Libnames.basename fn)++spc()++
- str"with carrier "++ pr_constr fi.ring_carrier++spc()++
- str"and equivalence relation "++ pr_constr fi.ring_req))
- ) !from_name ]
-END
-
(*****************************************************************************)
(* The tactics consist then only in a lookup in the ring database and
call the appropriate ltac. *)
@@ -807,7 +748,11 @@ let make_term_list env evd carrier rl =
let l = List.fold_right
(fun x l -> plapp evd coq_cons [|carrier;x;l|]) rl
(plapp evd coq_nil [|carrier|])
- in Typing.solve_evars env evd l
+ in Typing.e_solve_evars env evd l
+
+let carg = Tacinterp.Value.of_constr
+let tacarg expr =
+ Tacinterp.Value.of_closure (Tacinterp.default_ist ()) expr
let ltac_ring_structure e =
let req = carg e.ring_req in
@@ -815,18 +760,18 @@ let ltac_ring_structure e =
let ext = carg e.ring_ext in
let morph = carg e.ring_morph in
let th = carg e.ring_th in
- let cst_tac = Tacexp e.ring_cst_tac in
- let pow_tac = Tacexp e.ring_pow_tac in
+ let cst_tac = tacarg e.ring_cst_tac in
+ let pow_tac = tacarg e.ring_pow_tac in
let lemma1 = carg e.ring_lemma1 in
let lemma2 = carg e.ring_lemma2 in
- let pretac = Tacexp(TacFun([None],e.ring_pre_tac)) in
- let posttac = Tacexp(TacFun([None],e.ring_post_tac)) in
+ let pretac = tacarg (TacFun([None],e.ring_pre_tac)) in
+ let posttac = tacarg (TacFun([None],e.ring_post_tac)) in
[req;sth;ext;morph;th;cst_tac;pow_tac;
lemma1;lemma2;pretac;posttac]
-let ring_lookup (f:glob_tactic_expr) lH rl t =
- Proofview.Goal.enter begin fun gl ->
- let sigma = Proofview.Goal.sigma gl in
+let ring_lookup (f : Value.t) lH rl t =
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let sigma = Tacmach.New.project gl in
let env = Proofview.Goal.env gl in
try (* find_ring_strucure can raise an exception *)
let evdref = ref sigma in
@@ -837,14 +782,7 @@ let ring_lookup (f:glob_tactic_expr) lH rl t =
let ring = ltac_ring_structure e in
Proofview.tclTHEN (Proofview.Unsafe.tclEVARS !evdref) (ltac_apply f (ring@[lH;rl]))
with e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e
- end
-
-TACTIC EXTEND ring_lookup
-| [ "ring_lookup" tactic0(f) "[" constr_list(lH) "]" ne_constr_list(lrt) ] ->
- [ let (t,lr) = List.sep_last lrt in ring_lookup f lH lr t]
-END
-
-
+ end }
(***********************************************************************)
@@ -919,19 +857,6 @@ let dest_field env evd th_spec =
(Some true,r,zero,one,add,mul,None,None,div,inv,req,rth)
| _ -> error "bad field structure"
-type field_info =
- { field_carrier : types;
- field_req : constr;
- field_cst_tac : glob_tactic_expr;
- field_pow_tac : glob_tactic_expr;
- field_ok : constr;
- field_simpl_eq_ok : constr;
- field_simpl_ok : constr;
- field_simpl_eq_in_ok : constr;
- field_cond : constr;
- field_pre_tac : glob_tactic_expr;
- field_post_tac : glob_tactic_expr }
-
let field_from_carrier = Summary.ref Cmap.empty ~name:"field-tac-carrier-table"
let field_from_name = Summary.ref Spmap.empty ~name:"field-tac-name-table"
@@ -1034,7 +959,7 @@ let add_field_theory name (sigma,fth) eqth morphth cst_tac inj (pre,post) power
let rk = reflect_coeff morphth in
let params,ctx =
exec_tactic env !evd 9 (field_ltac"field_lemmas")
- (List.map carg[sth;ext;inv_m;fth;pspec;sspec;dspec;rk]) in
+ [sth;ext;inv_m;fth;pspec;sspec;dspec;rk] in
let lemma1 = params.(3) in
let lemma2 = params.(4) in
let lemma3 = params.(5) in
@@ -1078,15 +1003,6 @@ let add_field_theory name (sigma,fth) eqth morphth cst_tac inj (pre,post) power
field_pre_tac = pretac;
field_post_tac = posttac }) in ()
-type 'constr field_mod =
- Ring_mod of 'constr ring_mod
- | Inject of Constrexpr.constr_expr
-
-VERNAC ARGUMENT EXTEND field_mod
- | [ ring_mod(m) ] -> [ Ring_mod m ]
- | [ "completeness" constr(inj) ] -> [ Inject inj ]
-END
-
let process_field_mods l =
let kind = ref None in
let set = ref None in
@@ -1111,38 +1027,23 @@ let process_field_mods l =
let k = match !kind with Some k -> k | None -> Abstract in
(k, !set, !inj, !cst_tac, !pre, !post, !power, !sign, !div)
-VERNAC COMMAND EXTEND AddSetoidField CLASSIFIED AS SIDEFF
-| [ "Add" "Field" ident(id) ":" constr(t) field_mods(l) ] ->
- [ let (k,set,inj,cst_tac,pre,post,power,sign,div) = process_field_mods l in
- add_field_theory id (ic t) set k cst_tac inj (pre,post) power sign div]
-| [ "Print" "Fields" ] => [Vernac_classifier.classify_as_query] -> [
- msg_notice (strbrk "The following field structures have been declared:");
- Spmap.iter (fun fn fi ->
- msg_notice (hov 2
- (Ppconstr.pr_id (Libnames.basename fn)++spc()++
- str"with carrier "++ pr_constr fi.field_carrier++spc()++
- str"and equivalence relation "++ pr_constr fi.field_req))
- ) !field_from_name ]
-END
-
-
let ltac_field_structure e =
let req = carg e.field_req in
- let cst_tac = Tacexp e.field_cst_tac in
- let pow_tac = Tacexp e.field_pow_tac in
+ let cst_tac = tacarg e.field_cst_tac in
+ let pow_tac = tacarg e.field_pow_tac in
let field_ok = carg e.field_ok in
let field_simpl_ok = carg e.field_simpl_ok in
let field_simpl_eq_ok = carg e.field_simpl_eq_ok in
let field_simpl_eq_in_ok = carg e.field_simpl_eq_in_ok in
let cond_ok = carg e.field_cond in
- let pretac = Tacexp(TacFun([None],e.field_pre_tac)) in
- let posttac = Tacexp(TacFun([None],e.field_post_tac)) in
+ let pretac = tacarg (TacFun([None],e.field_pre_tac)) in
+ let posttac = tacarg (TacFun([None],e.field_post_tac)) in
[req;cst_tac;pow_tac;field_ok;field_simpl_ok;field_simpl_eq_ok;
field_simpl_eq_in_ok;cond_ok;pretac;posttac]
-let field_lookup (f:glob_tactic_expr) lH rl t =
- Proofview.Goal.enter begin fun gl ->
- let sigma = Proofview.Goal.sigma gl in
+let field_lookup (f : Value.t) lH rl t =
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let sigma = Tacmach.New.project gl in
let env = Proofview.Goal.env gl in
try
let evdref = ref sigma in
@@ -1153,10 +1054,4 @@ let field_lookup (f:glob_tactic_expr) lH rl t =
let field = ltac_field_structure e in
Proofview.tclTHEN (Proofview.Unsafe.tclEVARS !evdref) (ltac_apply f (field@[lH;rl]))
with e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e
- end
-
-
-TACTIC EXTEND field_lookup
-| [ "field_lookup" tactic(f) "[" constr_list(lH) "]" ne_constr_list(lt) ] ->
- [ let (t,l) = List.sep_last lt in field_lookup f lH l t ]
-END
+ end }
diff --git a/plugins/setoid_ring/newring.mli b/plugins/setoid_ring/newring.mli
new file mode 100644
index 000000000..07a1ae833
--- /dev/null
+++ b/plugins/setoid_ring/newring.mli
@@ -0,0 +1,78 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Names
+open Constr
+open Libnames
+open Globnames
+open Constrexpr
+open Tacexpr
+open Proof_type
+open Newring_ast
+
+val protect_tac_in : string -> Id.t -> tactic
+
+val protect_tac : string -> tactic
+
+val closed_term : constr -> global_reference list -> tactic
+
+val process_ring_mods :
+ constr_expr ring_mod list ->
+ constr coeff_spec * (constr * constr) option *
+ cst_tac_spec option * raw_tactic_expr option *
+ raw_tactic_expr option *
+ (cst_tac_spec * constr_expr) option *
+ constr_expr option * constr_expr option
+
+val add_theory :
+ Id.t ->
+ Evd.evar_map * constr ->
+ (constr * constr) option ->
+ constr coeff_spec ->
+ cst_tac_spec option ->
+ raw_tactic_expr option * raw_tactic_expr option ->
+ (cst_tac_spec * constr_expr) option ->
+ constr_expr option ->
+ constr_expr option -> unit
+
+val ic : constr_expr -> Evd.evar_map * constr
+
+val from_name : ring_info Spmap.t ref
+
+val ring_lookup :
+ Genarg.Val.t ->
+ constr list ->
+ constr list -> constr -> unit Proofview.tactic
+
+val process_field_mods :
+ constr_expr field_mod list ->
+ constr coeff_spec *
+ (constr * constr) option * constr option *
+ cst_tac_spec option * raw_tactic_expr option *
+ raw_tactic_expr option *
+ (cst_tac_spec * constr_expr) option *
+ constr_expr option * constr_expr option
+
+val add_field_theory :
+ Id.t ->
+ Evd.evar_map * constr ->
+ (constr * constr) option ->
+ constr coeff_spec ->
+ cst_tac_spec option ->
+ constr option ->
+ raw_tactic_expr option * raw_tactic_expr option ->
+ (cst_tac_spec * constr_expr) option ->
+ constr_expr option ->
+ constr_expr option -> unit
+
+val field_from_name : field_info Spmap.t ref
+
+val field_lookup :
+ Genarg.Val.t ->
+ constr list ->
+ constr list -> constr -> unit Proofview.tactic
diff --git a/plugins/setoid_ring/newring_ast.mli b/plugins/setoid_ring/newring_ast.mli
new file mode 100644
index 000000000..c26fcc8d1
--- /dev/null
+++ b/plugins/setoid_ring/newring_ast.mli
@@ -0,0 +1,63 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Constr
+open Libnames
+open Constrexpr
+open Tacexpr
+
+type 'constr coeff_spec =
+ Computational of 'constr (* equality test *)
+ | Abstract (* coeffs = Z *)
+ | Morphism of 'constr (* general morphism *)
+
+type cst_tac_spec =
+ CstTac of raw_tactic_expr
+ | Closed of reference list
+
+type 'constr ring_mod =
+ Ring_kind of 'constr coeff_spec
+ | Const_tac of cst_tac_spec
+ | Pre_tac of raw_tactic_expr
+ | Post_tac of raw_tactic_expr
+ | Setoid of constr_expr * constr_expr
+ | Pow_spec of cst_tac_spec * constr_expr
+ (* Syntaxification tactic , correctness lemma *)
+ | Sign_spec of constr_expr
+ | Div_spec of constr_expr
+
+type 'constr field_mod =
+ Ring_mod of 'constr ring_mod
+ | Inject of constr_expr
+
+type ring_info =
+ { ring_carrier : types;
+ ring_req : constr;
+ ring_setoid : constr;
+ ring_ext : constr;
+ ring_morph : constr;
+ ring_th : constr;
+ ring_cst_tac : glob_tactic_expr;
+ ring_pow_tac : glob_tactic_expr;
+ ring_lemma1 : constr;
+ ring_lemma2 : constr;
+ ring_pre_tac : glob_tactic_expr;
+ ring_post_tac : glob_tactic_expr }
+
+type field_info =
+ { field_carrier : types;
+ field_req : constr;
+ field_cst_tac : glob_tactic_expr;
+ field_pow_tac : glob_tactic_expr;
+ field_ok : constr;
+ field_simpl_eq_ok : constr;
+ field_simpl_ok : constr;
+ field_simpl_eq_in_ok : constr;
+ field_cond : constr;
+ field_pre_tac : glob_tactic_expr;
+ field_post_tac : glob_tactic_expr }
diff --git a/plugins/setoid_ring/newring_plugin.mllib b/plugins/setoid_ring/newring_plugin.mllib
index a98392f1e..7d6c49588 100644
--- a/plugins/setoid_ring/newring_plugin.mllib
+++ b/plugins/setoid_ring/newring_plugin.mllib
@@ -1,2 +1,3 @@
Newring
Newring_plugin_mod
+G_newring