summaryrefslogtreecommitdiff
path: root/plugins/funind
diff options
context:
space:
mode:
Diffstat (limited to 'plugins/funind')
-rw-r--r--plugins/funind/functional_principles_proofs.ml136
-rw-r--r--plugins/funind/functional_principles_types.ml68
-rw-r--r--plugins/funind/g_indfun.ml457
-rw-r--r--plugins/funind/glob_term_to_relation.ml68
-rw-r--r--plugins/funind/glob_termops.ml2
-rw-r--r--plugins/funind/indfun.ml83
-rw-r--r--plugins/funind/indfun.mli4
-rw-r--r--plugins/funind/indfun_common.ml20
-rw-r--r--plugins/funind/indfun_common.mli2
-rw-r--r--plugins/funind/invfun.ml111
-rw-r--r--plugins/funind/merge.ml65
-rw-r--r--plugins/funind/recdef.ml124
-rw-r--r--plugins/funind/recdef_plugin.mlpack (renamed from plugins/funind/recdef_plugin.mllib)1
13 files changed, 404 insertions, 337 deletions
diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml
index 169a7060..b0ffc775 100644
--- a/plugins/funind/functional_principles_proofs.ml
+++ b/plugins/funind/functional_principles_proofs.ml
@@ -1,20 +1,20 @@
open Printer
-open Errors
+open CErrors
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 *)
@@ -27,7 +27,7 @@ let observe strm =
let do_observe_tac s tac g =
try let v = tac g in (* msgnl (goal ++ fnl () ++ (str s)++(str " ")++(str "finished")); *) v
with e ->
- let e = Cerrors.process_vernac_interp_error e in
+ let e = ExplainErr.process_vernac_interp_error e in
let goal = begin try (Printer.pr_goal g) with _ -> assert false end in
msg_debug (str "observation "++ s++str " raised exception " ++
Errors.print e ++ str " on goal " ++ goal );
@@ -52,17 +52,17 @@ 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)
+ Feedback.msg_debug (hov 0 (lmsg ++ (str " raised exception " ++ CErrors.print e) ++ str " on goal" ++ fnl() ++ goal))
| None ->
begin
- Pp.msg_debug (str " from " ++ lmsg ++ str " on goal " ++ goal);
+ Feedback.msg_debug (str " from " ++ lmsg ++ str " on goal" ++ fnl() ++ goal);
end in
print_debug_queue None ;
end
let observe strm =
if do_observe ()
- then Pp.msg_debug strm
+ then Feedback.msg_debug strm
else ()
let do_observe_tac s tac g =
@@ -74,9 +74,9 @@ let do_observe_tac s tac g =
ignore(Stack.pop debug_queue);
v
with reraise ->
- let reraise = Errors.push reraise in
+ let reraise = CErrors.push reraise in
if not (Stack.is_empty debug_queue)
- then print_debug_queue (Some (fst (Cerrors.process_vernac_interp_error reraise)));
+ then print_debug_queue (Some (fst (ExplainErr.process_vernac_interp_error reraise)));
iraise reraise
let observe_tac_stream s tac g =
@@ -127,8 +127,7 @@ let finish_proof dynamic_infos g =
let refine c =
Tacmach.refine c
-let thin l =
- Tacmach.thin_no_check l
+let thin l = Proofview.V82.of_tactic (Tactics.clear l)
let eq_constr u v = eq_constr_nounivs u v
@@ -142,7 +141,7 @@ let is_trivial_eq t =
eq_constr t1 t2 && eq_constr a1 a2
| _ -> false
end
- with e when Errors.noncritical e -> false
+ with e when CErrors.noncritical e -> false
in
(* observe (str "is_trivial_eq " ++ Printer.pr_lconstr t ++ (if res then str " true" else str " false")); *)
res
@@ -168,7 +167,7 @@ let is_incompatible_eq t =
(eq_constr u1 u2 &&
incompatible_constructor_terms t1 t2)
| _ -> false
- with e when Errors.noncritical e -> false
+ with e when CErrors.noncritical e -> false
in
if res then observe (str "is_incompatible_eq " ++ Printer.pr_lconstr t);
res
@@ -224,12 +223,12 @@ let isAppConstruct ?(env=Global.env ()) t =
let nf_betaiotazeta = (* Reductionops.local_strong Reductionops.whd_betaiotazeta *)
let clos_norm_flags flgs env sigma t =
- Closure.norm_val (Closure.create_clos_infos flgs env) (Closure.inject (Reductionops.nf_evar sigma t)) in
- clos_norm_flags Closure.betaiotazeta Environ.empty_env Evd.empty
+ CClosure.norm_val (CClosure.create_clos_infos flgs env) (CClosure.inject (Reductionops.nf_evar sigma t)) in
+ clos_norm_flags CClosure.betaiotazeta Environ.empty_env Evd.empty
-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 );
@@ -255,7 +254,7 @@ let change_eq env sigma hyp_id (context:rel_context) x t end_of_type =
then
(jmeq_refl (),(args.(1),args.(0)),(args.(3),args.(2)),args.(0))
else nochange "not an equality"
- with e when Errors.noncritical e -> nochange "not an equality"
+ with e when CErrors.noncritical e -> nochange "not an equality"
in
if not ((closed0 (fst t1)) && (closed0 (snd t1)))then nochange "not a closed lhs";
let rec compute_substitution sub t1 t2 =
@@ -282,7 +281,7 @@ let change_eq env sigma hyp_id (context:rel_context) x t end_of_type =
List.fold_left2 compute_substitution sub args1 args2
end
else
- if (eq_constr t1 t2) then sub else nochange ~t':(make_refl_eq constructor (Reduction.whd_betadeltaiota env t1) t2) "cannot solve (diff)"
+ if (eq_constr t1 t2) then sub else nochange ~t':(make_refl_eq constructor (Reduction.whd_all env t1) t2) "cannot solve (diff)"
in
let sub = compute_substitution Int.Map.empty (snd t1) (snd t2) in
let sub = compute_substitution sub (fst t1) (fst t2) in
@@ -304,11 +303,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 +370,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 +535,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
@@ -626,8 +625,8 @@ let treat_new_case ptes_infos nb_prod continue_tac term dyn_infos =
let my_orelse tac1 tac2 g =
try
tac1 g
- with e when Errors.noncritical e ->
-(* observe (str "using snd tac since : " ++ Errors.print e); *)
+ with e when CErrors.noncritical e ->
+(* observe (str "using snd tac since : " ++ CErrors.print e); *)
tac2 g
let instanciate_hyps_with_args (do_prove:Id.t list -> tactic) hyps args_id =
@@ -705,9 +704,9 @@ let build_proof
in
tclTHENSEQ
[
- Simple.generalize (term_eq::(List.map mkVar dyn_infos.rec_hyps));
+ Proofview.V82.of_tactic (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 +735,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 +921,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,15 +934,15 @@ 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" *) (Proofview.V82.of_tactic (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
- (generalize (List.map mkVar idl))
+ (Proofview.V82.of_tactic (generalize (List.map mkVar idl)))
(thin idl)
let generate_equation_lemma evd fnames f fun_num nb_params nb_args rec_args_num =
@@ -1023,7 +1025,7 @@ let do_replace (evd:Evd.evar_map ref) params rec_arg_num rev_args_id f fun_num a
{finfos with
equation_lemma = Some (match Nametab.locate (qualid_of_ident equation_lemma_id) with
ConstRef c -> c
- | _ -> Errors.anomaly (Pp.str "Not a constant")
+ | _ -> CErrors.anomaly (Pp.str "Not a constant")
)
}
| _ -> ()
@@ -1044,7 +1046,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 +1072,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;
@@ -1086,7 +1085,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
match Global.body_of_constant const with
| Some body ->
Tacred.cbv_norm_flags
- (Closure.RedFlags.mkflags [Closure.RedFlags.fZETA])
+ (CClosure.RedFlags.mkflags [CClosure.RedFlags.fZETA])
(Global.env ())
(Evd.empty)
body
@@ -1120,11 +1119,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 +1164,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
@@ -1227,10 +1227,10 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
if this_fix_info.idx + 1 = 0
then tclIDTAC (* Someone tries to defined a principle on a fully parametric definition declared as a fixpoint (strange but ....) *)
else
- observe_tac_stream (str "h_fix " ++ int (this_fix_info.idx +1) ) (fix (Some this_fix_info.name) (this_fix_info.idx +1))
+ observe_tac_stream (str "h_fix " ++ int (this_fix_info.idx +1) ) (Proofview.V82.of_tactic (fix (Some this_fix_info.name) (this_fix_info.idx +1)))
else
- Tactics.mutual_fix this_fix_info.name (this_fix_info.idx + 1)
- other_fix_infos 0
+ Proofview.V82.of_tactic (Tactics.mutual_fix this_fix_info.name (this_fix_info.idx + 1)
+ other_fix_infos 0)
| _ -> anomaly (Pp.str "Not a valid information")
in
let first_tac : tactic = (* every operations until fix creations *)
@@ -1259,7 +1259,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 +1277,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 +1318,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 +1336,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 +1405,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 +1462,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 +1474,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 +1522,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 +1552,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 +1564,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 (Proofview.V82.of_tactic (Tactics.generalize (List.map mkVar l))) (Proofview.V82.of_tactic (clear l))
in
let fix_id = Nameops.out_name (fresh_id (Name hrec_id)) in
let prove_rec_arg_acc g =
@@ -1582,7 +1584,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 !!"
@@ -1608,7 +1610,7 @@ let prove_principle_for_gen
in
tclTHENSEQ
[
- generalize [lemma];
+ Proofview.V82.of_tactic (generalize [lemma]);
Proofview.V82.of_tactic (Simple.intro hid);
Proofview.V82.of_tactic (Elim.h_decompose_and (mkVar hid));
(fun g ->
@@ -1629,7 +1631,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
@@ -1640,7 +1642,7 @@ let prove_principle_for_gen
(* observe_tac "reverting" *) (revert (List.rev (acc_rec_arg_id::args_ids)));
(* (fun g -> observe (Printer.pr_goal (sig_it g) ++ fnl () ++ *)
(* str "fix arg num" ++ int (List.length args_ids + 1) ); tclIDTAC g); *)
- (* observe_tac "h_fix " *) (fix (Some fix_id) (List.length args_ids + 1));
+ (* observe_tac "h_fix " *) (Proofview.V82.of_tactic (fix (Some fix_id) (List.length args_ids + 1)));
(* (fun g -> observe (Printer.pr_goal (sig_it g) ++ fnl() ++ pr_lconstr_env (pf_env g ) (pf_unsafe_type_of g (mkVar fix_id) )); tclIDTAC g); *)
h_intros (List.rev (acc_rec_arg_id::args_ids));
Proofview.V82.of_tactic (Equality.rewriteLR (mkConst eq_ref));
@@ -1667,7 +1669,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 +1685,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 +1714,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 c47602bd..5e72b867 100644
--- a/plugins/funind/functional_principles_types.ml
+++ b/plugins/funind/functional_principles_types.ml
@@ -1,24 +1,25 @@
open Printer
-open Errors
+open CErrors
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
let observe s =
if do_observe ()
- then Pp.msg_debug s
+ then Feedback.msg_debug s
(*
Transform an inductive induction principle into
@@ -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)
@@ -291,7 +298,7 @@ let build_functional_principle (evd:Evd.evar_map ref) interactive_proof old_prin
(* let dur1 = System.time_difference tim1 tim2 in *)
(* Pp.msgnl (str ("Time to compute proof: ") ++ str (string_of_float dur1)); *)
(* end; *)
- get_proof_clean true, Ephemeron.create hook
+ get_proof_clean true, CEphemeron.create hook
end
@@ -351,7 +358,7 @@ let generate_functional_principle (evd: Evd.evar_map ref)
Don't forget to close the goal if an error is raised !!!!
*)
save false new_princ_name entry g_kind hook
- with e when Errors.noncritical e ->
+ with e when CErrors.noncritical e ->
begin
begin
try
@@ -363,7 +370,7 @@ let generate_functional_principle (evd: Evd.evar_map ref)
then Pfedit.delete_current_proof ()
else ()
else ()
- with e when Errors.noncritical e -> ()
+ with e when CErrors.noncritical e -> ()
end;
raise (Defining_principle e)
end
@@ -393,7 +400,7 @@ let get_funs_constant mp dp =
match Global.body_of_constant const with
| Some body ->
let body = Tacred.cbv_norm_flags
- (Closure.RedFlags.mkflags [Closure.RedFlags.fZETA])
+ (CClosure.RedFlags.mkflags [CClosure.RedFlags.fZETA])
(Global.env ())
(Evd.from_env (Global.env ()))
body
@@ -503,7 +510,7 @@ let make_scheme evd (fas : (pconstant*glob_sort) list) : Safe_typing.private_con
0
(prove_princ_for_struct evd false 0 (Array.of_list (List.map fst funs)))
(fun _ _ _ -> ())
- with e when Errors.noncritical e ->
+ with e when CErrors.noncritical e ->
begin
begin
try
@@ -515,7 +522,7 @@ let make_scheme evd (fas : (pconstant*glob_sort) list) : Safe_typing.private_con
then Pfedit.delete_current_proof ()
else ()
else ()
- with e when Errors.noncritical e -> ()
+ with e when CErrors.noncritical e -> ()
end;
raise (Defining_principle e)
end
@@ -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 a15e46bf..42e49031 100644
--- a/plugins/funind/g_indfun.ml4
+++ b/plugins/funind/g_indfun.ml4
@@ -9,15 +9,16 @@
open Compat
open Util
open Term
-open Vars
-open Names
open Pp
open Constrexpr
open Indfun_common
open Indfun
open Genarg
-open Tacticals
+open Constrarg
open Misctypes
+open Pcoq.Prim
+open Pcoq.Constr
+open Pcoq.Tactic
DECLARE PLUGIN "recdef_plugin"
@@ -55,10 +56,13 @@ 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
+ TYPED AS constr_with_bindings option
PRINTED BY pr_fun_ind_using_typed
RAW_TYPED AS constr_with_bindings_opt
RAW_PRINTED BY pr_fun_ind_using
@@ -86,9 +90,9 @@ let pr_intro_as_pat _prc _ _ pat =
let out_disjunctive = function
| loc, IntroAction (IntroOrAndPattern l) -> (loc,l)
- | _ -> Errors.error "Disjunctive or conjunctive intro pattern expected."
+ | _ -> CErrors.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
@@ -119,12 +123,12 @@ TACTIC EXTEND snewfunind
END
-let pr_constr_coma_sequence prc _ _ = prlist_with_sep pr_comma prc
+let pr_constr_comma_sequence prc _ _ = prlist_with_sep pr_comma prc
-ARGUMENT EXTEND constr_coma_sequence'
+ARGUMENT EXTEND constr_comma_sequence'
TYPED AS constr_list
- PRINTED BY pr_constr_coma_sequence
-| [ constr(c) "," constr_coma_sequence'(l) ] -> [ c::l ]
+ PRINTED BY pr_constr_comma_sequence
+| [ constr(c) "," constr_comma_sequence'(l) ] -> [ c::l ]
| [ constr(c) ] -> [ [c] ]
END
@@ -133,7 +137,7 @@ let pr_auto_using prc _prlc _prt = Pptactic.pr_auto_using prc
ARGUMENT EXTEND auto_using'
TYPED AS constr_list
PRINTED BY pr_auto_using
-| [ "using" constr_coma_sequence'(l) ] -> [ l ]
+| [ "using" constr_comma_sequence'(l) ] -> [ l ]
| [ ] -> [ [] ]
END
@@ -144,10 +148,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 ;
@@ -158,6 +162,11 @@ GEXTEND Gram
END
+let () =
+ let raw_printer _ _ _ (loc,body) = Ppvernac.pr_rec_definition body in
+ let printer _ _ _ _ = str "<Unavailable printer for rec_definition>" in
+ Pptactic.declare_extra_genarg_pprule wit_function_rec_definition_loc raw_printer printer printer
+
(* TASSI: n'importe quoi ! *)
VERNAC COMMAND EXTEND Function
["Function" ne_function_rec_definition_loc_list_sep(recsl,"with")]
@@ -186,18 +195,16 @@ END
let warning_error names e =
- let (e, _) = Cerrors.process_vernac_interp_error (e, Exninfo.null) in
+ let (e, _) = ExplainErr.process_vernac_interp_error (e, Exninfo.null) in
match e with
| Building_graph e ->
- Pp.msg_warning
- (str "Cannot define graph(s) for " ++
- h 1 (pr_enum Libnames.pr_reference names) ++
- if do_observe () then (spc () ++ Errors.print e) else mt ())
+ let names = pr_enum Libnames.pr_reference names in
+ let error = if do_observe () then (spc () ++ CErrors.print e) else mt () in
+ warn_cannot_define_graph (names,error)
| Defining_principle e ->
- Pp.msg_warning
- (str "Cannot define principle(s) for "++
- h 1 (pr_enum Libnames.pr_reference names) ++
- if do_observe () then Errors.print e else mt ())
+ let names = pr_enum Libnames.pr_reference names in
+ let error = if do_observe () then CErrors.print e else mt () in
+ warn_cannot_define_principle (names,error)
| _ -> raise e
@@ -220,15 +227,15 @@ VERNAC COMMAND EXTEND NewFunctionalScheme
;
try Functional_principles_types.build_scheme fas
with Functional_principles_types.No_graph_found ->
- Errors.error ("Cannot generate induction principle(s)")
- | e when Errors.noncritical e ->
+ CErrors.error ("Cannot generate induction principle(s)")
+ | e when CErrors.noncritical e ->
let names = List.map (fun (_,na,_) -> na) fas in
warning_error names e
end
| _ -> assert false (* we can only have non empty list *)
end
- | e when Errors.noncritical e ->
+ | e when CErrors.noncritical e ->
let names = List.map (fun (_,na,_) -> na) fas in
warning_error names e
end
diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml
index 5d92fca5..52179ae5 100644
--- a/plugins/funind/glob_term_to_relation.ml
+++ b/plugins/funind/glob_term_to_relation.ml
@@ -7,14 +7,14 @@ open Glob_term
open Glob_ops
open Globnames
open Indfun_common
-open Errors
+open CErrors
open Util
open Glob_termops
open Misctypes
let observe strm =
if do_observe ()
- then Pp.msg_debug strm
+ then Feedback.msg_debug strm
else ()
(*let observennl strm =
if do_observe ()
@@ -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
@@ -915,7 +921,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
observe (str "computing new type for eq : " ++ pr_glob_constr rt);
let t' =
try fst (Pretyping.understand env (Evd.from_env env) t)(*FIXME*)
- with e when Errors.noncritical e -> raise Continue
+ with e when CErrors.noncritical e -> raise Continue
in
let is_in_b = is_free_in id b in
let _keep_eq =
@@ -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
@@ -1212,13 +1217,13 @@ let compute_params_name relnames (args : (Name.t * Glob_term.glob_constr * bool)
if Array.for_all
(fun l ->
let (n',nt',is_defined') = List.nth l i in
- Name.equal n n' && Notation_ops.eq_glob_constr nt nt' && (is_defined : bool) == is_defined')
+ Name.equal n n' && glob_constr_eq nt nt' && (is_defined : bool) == is_defined')
rels_params
then
l := param::!l
)
rels_params.(0)
- with e when Errors.noncritical e ->
+ with e when CErrors.noncritical e ->
()
in
List.rev !l
@@ -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 =
@@ -1454,7 +1460,7 @@ let do_build_inductive
str "while trying to define"++ spc () ++
Ppvernac.pr_vernac (Vernacexpr.VernacInductive(false,Decl_kinds.Finite,repacked_rel_inds))
++ fnl () ++
- Errors.print reraise
+ CErrors.print reraise
in
observe msg;
raise reraise
@@ -1470,7 +1476,7 @@ let build_inductive evd funconstants funsargs returned_types rtl =
do_build_inductive evd funconstants funsargs returned_types rtl;
Detyping.print_universes := pu;
Constrextern.print_universes := cu
- with e when Errors.noncritical e ->
+ with e when CErrors.noncritical e ->
Detyping.print_universes := pu;
Constrextern.print_universes := cu;
raise (Building_graph e)
diff --git a/plugins/funind/glob_termops.ml b/plugins/funind/glob_termops.ml
index 291f835e..01e5ef7f 100644
--- a/plugins/funind/glob_termops.ml
+++ b/plugins/funind/glob_termops.ml
@@ -1,6 +1,6 @@
open Pp
open Glob_term
-open Errors
+open CErrors
open Util
open Names
open Decl_kinds
diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml
index 3dbd4380..18817f50 100644
--- a/plugins/funind/indfun.ml
+++ b/plugins/funind/indfun.ml
@@ -1,4 +1,5 @@
-open Errors
+open Context.Rel.Declaration
+open CErrors
open Util
open Names
open Term
@@ -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
@@ -130,6 +132,7 @@ let rec abstract_glob_constr c = function
| Constrexpr.LocalRawAssum (idl,k,t)::bl ->
List.fold_right (fun x b -> Constrexpr_ops.mkLambdaC([x],k,t,b)) idl
(abstract_glob_constr c bl)
+ | Constrexpr.LocalPattern _::bl -> assert false
let interp_casted_constr_with_implicits env sigma impls c =
Constrintern.intern_gen Pretyping.WithoutTypeConstraint env ~impls
@@ -152,7 +155,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 *)
@@ -212,6 +216,7 @@ let rec local_binders_length = function
| [] -> 0
| Constrexpr.LocalRawDef _::bl -> 1 + local_binders_length bl
| Constrexpr.LocalRawAssum (idl,_,_)::bl -> List.length idl + local_binders_length bl
+ | Constrexpr.LocalPattern _::bl -> assert false
let prepare_body ((name,_,args,types,_),_) rt =
let n = local_binders_length args in
@@ -220,7 +225,12 @@ let prepare_body ((name,_,args,types,_),_) rt =
(fun_args,rt')
let process_vernac_interp_error e =
- fst (Cerrors.process_vernac_interp_error (e, Exninfo.null))
+ fst (ExplainErr.process_vernac_interp_error (e, Exninfo.null))
+
+let warn_funind_cannot_build_inversion =
+ CWarnings.create ~name:"funind-cannot-build-inversion" ~category:"funind"
+ (fun e' -> strbrk "Cannot build inversion information" ++
+ if do_observe () then (fnl() ++ CErrors.print e') else mt ())
let derive_inversion fix_names =
try
@@ -262,16 +272,22 @@ let derive_inversion fix_names =
functional_induction
fix_names_as_constant
lind;
- with e when Errors.noncritical e ->
- let e' = process_vernac_interp_error e in
- msg_warning
- (str "Cannot build inversion information" ++
- if do_observe () then (fnl() ++ Errors.print e') else mt ())
- with e when Errors.noncritical e ->
+ with e when CErrors.noncritical e ->
let e' = process_vernac_interp_error e in
- msg_warning
- (str "Cannot build inversion information (early)" ++
- if do_observe () then (fnl() ++ Errors.print e') else mt ())
+ warn_funind_cannot_build_inversion e'
+ with e when CErrors.noncritical e ->
+ let e' = process_vernac_interp_error e in
+ warn_funind_cannot_build_inversion e'
+
+let warn_cannot_define_graph =
+ CWarnings.create ~name:"funind-cannot-define-graph" ~category:"funind"
+ (fun (names,error) -> strbrk "Cannot define graph(s) for " ++
+ h 1 names ++ error)
+
+let warn_cannot_define_principle =
+ CWarnings.create ~name:"funind-cannot-define-principle" ~category:"funind"
+ (fun (names,error) -> strbrk "Cannot define induction principle(s) for "++
+ h 1 names ++ error)
let warning_error names e =
let e = process_vernac_interp_error e in
@@ -279,33 +295,29 @@ let warning_error names e =
match e with
| ToShow e ->
let e = process_vernac_interp_error e in
- spc () ++ Errors.print e
+ spc () ++ CErrors.print e
| _ ->
if do_observe ()
then
let e = process_vernac_interp_error e in
- (spc () ++ Errors.print e)
+ (spc () ++ CErrors.print e)
else mt ()
in
match e with
| Building_graph e ->
- Pp.msg_warning
- (str "Cannot define graph(s) for " ++
- h 1 (prlist_with_sep (fun _ -> str","++spc ()) Ppconstr.pr_id names) ++
- e_explain e)
+ let names = prlist_with_sep (fun _ -> str","++spc ()) Ppconstr.pr_id names in
+ warn_cannot_define_graph (names,e_explain e)
| Defining_principle e ->
- Pp.msg_warning
- (str "Cannot define principle(s) for "++
- h 1 (prlist_with_sep (fun _ -> str","++spc ()) Ppconstr.pr_id names) ++
- e_explain e)
+ let names = prlist_with_sep (fun _ -> str","++spc ()) Ppconstr.pr_id names in
+ warn_cannot_define_principle (names,e_explain e)
| _ -> raise e
let error_error names e =
let e = process_vernac_interp_error e in
let e_explain e =
match e with
- | ToShow e -> spc () ++ Errors.print e
- | _ -> if do_observe () then (spc () ++ Errors.print e) else mt ()
+ | ToShow e -> spc () ++ CErrors.print e
+ | _ -> if do_observe () then (spc () ++ CErrors.print e) else mt ()
in
match e with
| Building_graph e ->
@@ -373,7 +385,7 @@ let generate_principle (evd:Evd.evar_map ref) pconstants on_error
Array.iter (add_Function is_general) funs_kn;
()
end
- with e when Errors.noncritical e ->
+ with e when CErrors.noncritical e ->
on_error names e
let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) list) =
@@ -463,7 +475,7 @@ let register_wf ?(is_mes=false) fname rec_impls wf_rel_expr wf_arg using_lemmas
functional_ref eq_ref rec_arg_num rec_arg_type nb_args relation
);
derive_inversion [fname]
- with e when Errors.noncritical e ->
+ with e when CErrors.noncritical e ->
(* No proof done *)
()
in
@@ -727,9 +739,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 +763,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
@@ -860,6 +870,7 @@ let make_graph (f_ref:global_reference) =
(fun (loc,n) ->
CRef(Libnames.Ident(loc, Nameops.out_name n),None))
nal
+ | Constrexpr.LocalPattern _ -> assert false
)
nal_tas
)
diff --git a/plugins/funind/indfun.mli b/plugins/funind/indfun.mli
index e7206914..1c27bdfa 100644
--- a/plugins/funind/indfun.mli
+++ b/plugins/funind/indfun.mli
@@ -1,5 +1,9 @@
open Misctypes
+val warn_cannot_define_graph : ?loc:Loc.t -> Pp.std_ppcmds * Pp.std_ppcmds -> unit
+
+val warn_cannot_define_principle : ?loc:Loc.t -> Pp.std_ppcmds * Pp.std_ppcmds -> unit
+
val do_generate_principle :
bool ->
(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) list ->
diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml
index aa47e261..f56e9241 100644
--- a/plugins/funind/indfun_common.ml
+++ b/plugins/funind/indfun_common.ml
@@ -49,7 +49,7 @@ let locate_constant ref =
let locate_with_msg msg f x =
try f x
- with Not_found -> raise (Errors.UserError("", msg))
+ with Not_found -> raise (CErrors.UserError("", msg))
let filter_map filter f =
@@ -73,7 +73,7 @@ let chop_rlambda_n =
| Glob_term.GLambda(_,name,k,t,b) -> chop_lambda_n ((name,t,false)::acc) (n-1) b
| Glob_term.GLetIn(_,name,v,b) -> chop_lambda_n ((name,v,true)::acc) (n-1) b
| _ ->
- raise (Errors.UserError("chop_rlambda_n",
+ raise (CErrors.UserError("chop_rlambda_n",
str "chop_rlambda_n: Not enough Lambdas"))
in
chop_lambda_n []
@@ -85,7 +85,7 @@ let chop_rprod_n =
else
match rt with
| Glob_term.GProd(_,name,k,t,b) -> chop_prod_n ((name,t)::acc) (n-1) b
- | _ -> raise (Errors.UserError("chop_rprod_n",str "chop_rprod_n: Not enough products"))
+ | _ -> raise (CErrors.UserError("chop_rprod_n",str "chop_rprod_n: Not enough products"))
in
chop_prod_n []
@@ -110,7 +110,7 @@ let const_of_id id =
in
try Constrintern.locate_reference princ_ref
with Not_found ->
- Errors.errorlabstrm "IndFun.const_of_id"
+ CErrors.errorlabstrm "IndFun.const_of_id"
(str "cannot find " ++ Nameops.pr_id id)
let def_of_const t =
@@ -163,7 +163,7 @@ let save with_clean id const (locality,_,kind) hook =
(locality, ConstRef kn)
in
if with_clean then Pfedit.delete_current_proof ();
- Ephemeron.iter_opt hook (fun f -> Lemmas.call_hook fix_exn f l r);
+ CEphemeron.iter_opt hook (fun f -> Lemmas.call_hook fix_exn f l r);
definition_message id
@@ -344,7 +344,7 @@ let pr_info f_info =
(try
Printer.pr_lconstr
(Global.type_of_global_unsafe (ConstRef f_info.function_constant))
- with e when Errors.noncritical e -> mt ()) ++ fnl () ++
+ with e when CErrors.noncritical e -> mt ()) ++ fnl () ++
str "equation_lemma := " ++ pr_ocst f_info.equation_lemma ++ fnl () ++
str "completeness_lemma :=" ++ pr_ocst f_info.completeness_lemma ++ fnl () ++
str "correctness_lemma := " ++ pr_ocst f_info.correctness_lemma ++ fnl () ++
@@ -371,7 +371,7 @@ let in_Function : function_info -> Libobject.obj =
let find_or_none id =
try Some
- (match Nametab.locate (qualid_of_ident id) with ConstRef c -> c | _ -> Errors.anomaly (Pp.str "Not a constant")
+ (match Nametab.locate (qualid_of_ident id) with ConstRef c -> c | _ -> CErrors.anomaly (Pp.str "Not a constant")
)
with Not_found -> None
@@ -399,7 +399,7 @@ let add_Function is_general f =
and prop_lemma = find_or_none (Nameops.add_suffix f_id "_ind")
and graph_ind =
match Nametab.locate (qualid_of_ident (mk_rel_id f_id))
- with | IndRef ind -> ind | _ -> Errors.anomaly (Pp.str "Not an inductive")
+ with | IndRef ind -> ind | _ -> CErrors.anomaly (Pp.str "Not an inductive")
in
let finfos =
{ function_constant = f;
@@ -476,13 +476,13 @@ let jmeq () =
try
Coqlib.check_required_library Coqlib.jmeq_module_name;
Coqlib.gen_constant "Function" ["Logic";"JMeq"] "JMeq"
- with e when Errors.noncritical e -> raise (ToShow e)
+ with e when CErrors.noncritical e -> raise (ToShow e)
let jmeq_refl () =
try
Coqlib.check_required_library Coqlib.jmeq_module_name;
Coqlib.gen_constant "Function" ["Logic";"JMeq"] "JMeq_refl"
- with e when Errors.noncritical e -> raise (ToShow e)
+ with e when CErrors.noncritical e -> raise (ToShow e)
let h_intros l =
tclMAP (fun x -> Proofview.V82.of_tactic (Tactics.Simple.intro x)) l
diff --git a/plugins/funind/indfun_common.mli b/plugins/funind/indfun_common.mli
index 23f1da1b..e5c756f5 100644
--- a/plugins/funind/indfun_common.mli
+++ b/plugins/funind/indfun_common.mli
@@ -47,7 +47,7 @@ val jmeq : unit -> Term.constr
val jmeq_refl : unit -> Term.constr
val save : bool -> Id.t -> Safe_typing.private_constants Entries.definition_entry -> Decl_kinds.goal_kind ->
- unit Lemmas.declaration_hook Ephemeron.key -> unit
+ unit Lemmas.declaration_hook CEphemeron.key -> unit
(* [get_proof_clean do_reduce] : returns the proof name, definition, kind and hook and
abort the proof
diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml
index a800c186..26fc88a6 100644
--- a/plugins/funind/invfun.ml
+++ b/plugins/funind/invfun.ml
@@ -5,9 +5,10 @@
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+
open Tacexpr
open Declarations
-open Errors
+open CErrors
open Util
open Names
open Term
@@ -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 *)
@@ -50,7 +53,7 @@ let pr_constr_with_binding prc (c,bl) : Pp.std_ppcmds =
let observe strm =
if do_observe ()
- then Pp.msg_debug strm
+ then Feedback.msg_debug strm
else ()
(*let observennl strm =
@@ -62,16 +65,16 @@ let observe strm =
let do_observe_tac s tac g =
let goal =
try Printer.pr_goal g
- with e when Errors.noncritical e -> assert false
+ with e when CErrors.noncritical e -> assert false
in
try
let v = tac g in
msgnl (goal ++ fnl () ++ s ++(str " ")++(str "finished")); v
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 );
+ let reraise = CErrors.push reraise in
+ let e = ExplainErr.process_vernac_interp_error reraise in
+ observe (hov 0 (str "observation "++ s++str " raised exception " ++
+ CErrors.iprint e ++ str " on goal" ++ fnl() ++ goal ));
iraise reraise;;
@@ -87,10 +90,11 @@ let observe_tac s tac g =
(* [nf_zeta] $\zeta$-normalization of a term *)
let nf_zeta =
- Reductionops.clos_norm_flags (Closure.RedFlags.mkflags [Closure.RedFlags.fZETA])
+ Reductionops.clos_norm_flags (CClosure.RedFlags.mkflags [CClosure.RedFlags.fZETA])
Environ.empty_env
Evd.empty
+let thin ids gl = Proofview.V82.of_tactic (Tactics.clear ids) gl
(* (\* [id_to_constr id] finds the term associated to [id] in the global environment *\) *)
(* let id_to_constr id = *)
@@ -133,18 +137,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 +177,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 +266,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
@@ -358,18 +365,18 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
observe_tac("h_intro_patterns ") (let l = (List.nth intro_pats (pred i)) in
match l with
| [] -> tclIDTAC
- | _ -> Proofview.V82.of_tactic (intro_patterns l));
+ | _ -> Proofview.V82.of_tactic (intro_patterns false l));
(* 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 +396,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 +414,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 +424,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 +461,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 (Proofview.V82.of_tactic (Tactics.generalize [mkVar id])) (thin [id])
| _ -> tclIDTAC
)
(pf_hyps g)
@@ -467,6 +475,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 +500,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 +540,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 +548,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 +563,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
@@ -568,7 +585,7 @@ let rec reflexivity_with_destruct_cases g =
observe_tac "reflexivity_with_destruct_cases" reflexivity_with_destruct_cases
]
| _ -> Proofview.V82.of_tactic reflexivity
- with e when Errors.noncritical e -> Proofview.V82.of_tactic reflexivity
+ with e when CErrors.noncritical e -> Proofview.V82.of_tactic reflexivity
in
let eq_ind = make_eq () in
let discr_inject =
@@ -662,10 +679,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 +708,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);
+ Proofview.V82.of_tactic (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 +754,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)]);
+ (Proofview.V82.of_tactic (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 +937,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])];
+ Proofview.V82.of_tactic (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 +981,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])];
+ Proofview.V82.of_tactic (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));
@@ -981,7 +998,7 @@ let invfun qhyp f =
let f =
match f with
| ConstRef f -> f
- | _ -> raise (Errors.UserError("",str "Not a function"))
+ | _ -> raise (CErrors.UserError("",str "Not a function"))
in
try
let finfos = find_Function_infos f in
diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml
index 87d7ca76..de4210af 100644
--- a/plugins/funind/merge.ml
+++ b/plugins/funind/merge.ml
@@ -11,7 +11,7 @@
open Globnames
open Tactics
open Indfun_common
-open Errors
+open CErrors
open Util
open Constrexpr
open Vernacexpr
@@ -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} *)
@@ -73,7 +73,7 @@ let ident_global_exist id =
let ans = CRef (Libnames.Ident (Loc.ghost,id), None) in
let _ = ignore (Constrintern.intern_constr (Global.env()) ans) in
true
- with e when Errors.noncritical e -> false
+ with e when CErrors.noncritical e -> false
(** [next_ident_fresh id] returns a fresh identifier (ie not linked in
global env) with base [id]. *)
@@ -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;
@@ -503,19 +504,19 @@ let rec merge_app c1 c2 id1 id2 shift filter_shift_stable =
let lnk = Array.append shift.lnk1 shift.lnk2 in
match c1 , c2 with
| GApp(_,f1, arr1), GApp(_,f2,arr2) when isVarf id1 f1 && isVarf id2 f2 ->
- let _ = prstr "\nICI1!\n";Pp.flush_all() in
+ let _ = prstr "\nICI1!\n" in
let args = filter_shift_stable lnk (arr1 @ arr2) in
GApp (Loc.ghost,GVar (Loc.ghost,shift.ident) , args)
| GApp(_,f1, arr1), GApp(_,f2,arr2) -> raise NoMerge
| GLetIn(_,nme,bdy,trm) , _ ->
- let _ = prstr "\nICI2!\n";Pp.flush_all() in
+ let _ = prstr "\nICI2!\n" in
let newtrm = merge_app trm c2 id1 id2 shift filter_shift_stable in
GLetIn(Loc.ghost,nme,bdy,newtrm)
| _, GLetIn(_,nme,bdy,trm) ->
- let _ = prstr "\nICI3!\n";Pp.flush_all() in
+ let _ = prstr "\nICI3!\n" in
let newtrm = merge_app c1 trm id1 id2 shift filter_shift_stable in
GLetIn(Loc.ghost,nme,bdy,newtrm)
- | _ -> let _ = prstr "\nICI4!\n";Pp.flush_all() in
+ | _ -> let _ = prstr "\nICI4!\n" in
raise NoMerge
let rec merge_app_unsafe c1 c2 shift filter_shift_stable =
@@ -526,14 +527,14 @@ let rec merge_app_unsafe c1 c2 shift filter_shift_stable =
GApp (Loc.ghost,GVar(Loc.ghost,shift.ident) , args)
(* FIXME: what if the function appears in the body of the let? *)
| GLetIn(_,nme,bdy,trm) , _ ->
- let _ = prstr "\nICI2 '!\n";Pp.flush_all() in
+ let _ = prstr "\nICI2 '!\n" in
let newtrm = merge_app_unsafe trm c2 shift filter_shift_stable in
GLetIn(Loc.ghost,nme,bdy,newtrm)
| _, GLetIn(_,nme,bdy,trm) ->
- let _ = prstr "\nICI3 '!\n";Pp.flush_all() in
+ let _ = prstr "\nICI3 '!\n" in
let newtrm = merge_app_unsafe c1 trm shift filter_shift_stable in
GLetIn(Loc.ghost,nme,bdy,newtrm)
- | _ -> let _ = prstr "\nICI4 '!\n";Pp.flush_all() in raise NoMerge
+ | _ -> let _ = prstr "\nICI4 '!\n" in raise NoMerge
@@ -784,10 +785,10 @@ let merge_inductive_body (shift:merge_infos) avoid (oib1:one_inductive_body)
let params1 =
try fst (glob_decompose_prod_n shift.nrecprms1 (List.hd lcstr1))
- with e when Errors.noncritical e -> [] in
+ with e when CErrors.noncritical e -> [] in
let params2 =
try fst (glob_decompose_prod_n shift.nrecprms2 (List.hd lcstr2))
- with e when Errors.noncritical e -> [] in
+ with e when CErrors.noncritical e -> [] in
let lcstr1 = List.combine (Array.to_list oib1.mind_consnames) lcstr1 in
let lcstr2 = List.combine (Array.to_list oib2.mind_consnames) lcstr2 in
@@ -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 065d0fe5..fa84e4dd 100644
--- a/plugins/funind/recdef.ml
+++ b/plugins/funind/recdef.ml
@@ -16,7 +16,7 @@ open Names
open Libnames
open Globnames
open Nameops
-open Errors
+open CErrors
open Util
open Tacticals
open Tacmach
@@ -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 *)
@@ -90,15 +92,15 @@ let const_of_ref = function
let nf_zeta env =
- Reductionops.clos_norm_flags (Closure.RedFlags.mkflags [Closure.RedFlags.fZETA])
+ Reductionops.clos_norm_flags (CClosure.RedFlags.mkflags [CClosure.RedFlags.fZETA])
env
Evd.empty
let nf_betaiotazeta = (* Reductionops.local_strong Reductionops.whd_betaiotazeta *)
let clos_norm_flags flgs env sigma t =
- Closure.norm_val (Closure.create_clos_infos flgs env) (Closure.inject (Reductionops.nf_evar sigma t)) in
- clos_norm_flags Closure.betaiotazeta Environ.empty_env Evd.empty
+ CClosure.norm_val (CClosure.create_clos_infos flgs env) (CClosure.inject (Reductionops.nf_evar sigma t)) in
+ clos_norm_flags CClosure.betaiotazeta Environ.empty_env Evd.empty
@@ -159,7 +161,7 @@ let rec n_x_id ids n =
let simpl_iter clause =
reduce
(Lazy
- {rBeta=true;rIota=true;rZeta= true; rDelta=false;
+ {rBeta=true;rMatch=true;rFix=true;rCofix=true;rZeta=true;rDelta=false;
rConst = [ EvalConstRef (const_of_ref (delayed_force iter_ref))]})
clause
@@ -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 =
@@ -206,23 +208,23 @@ let (declare_f : Id.t -> logical_kind -> constr list -> global_reference -> glob
(* Debugging mechanism *)
let debug_queue = Stack.create ()
-let rec print_debug_queue b e =
+let print_debug_queue b e =
if not (Stack.is_empty debug_queue)
then
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)
+ Feedback.msg_debug (hov 1 (lmsg ++ (str " raised exception " ++ CErrors.print e) ++ str " on goal" ++ fnl() ++ goal))
else
begin
- Pp.msg_debug (str " from " ++ lmsg ++ str " on goal " ++ goal);
+ Feedback.msg_debug (hov 1 (str " from " ++ lmsg ++ str " on goal"++fnl() ++ goal));
end;
(* print_debug_queue false e; *)
end
let observe strm =
if do_observe ()
- then Pp.msg_debug strm
+ then Feedback.msg_debug strm
else ()
@@ -236,9 +238,9 @@ let do_observe_tac s tac g =
ignore(Stack.pop debug_queue);
v
with reraise ->
- let reraise = Errors.push reraise in
+ let reraise = CErrors.push reraise in
if not (Stack.is_empty debug_queue)
- then print_debug_queue true (fst (Cerrors.process_vernac_interp_error reraise));
+ then print_debug_queue true (fst (ExplainErr.process_vernac_interp_error reraise));
iraise reraise
let observe_tac s tac g =
@@ -265,8 +267,8 @@ let observe_tclTHENLIST s tacl =
let tclUSER tac is_mes l g =
let clear_tac =
match l with
- | None -> clear []
- | Some l -> tclMAP (fun id -> tclTRY (clear [id])) (List.rev l)
+ | None -> tclIDTAC
+ | Some l -> tclMAP (fun id -> tclTRY (Proofview.V82.of_tactic (clear [id]))) (List.rev l)
in
observe_tclTHENLIST (str "tclUSER1")
[
@@ -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
@@ -397,7 +399,7 @@ let treat_case forbid_new_ids to_intros finalize_tac nb_lam e infos : tactic =
Proofview.V82.of_tactic (intro_using teq_id);
onLastHypId (fun heq ->
observe_tclTHENLIST (str "treat_case2")[
- thin to_intros;
+ Proofview.V82.of_tactic (clear to_intros);
h_intros to_intros;
(fun g' ->
let ty_teq = pf_unsafe_type_of g' (mkVar heq) in
@@ -439,7 +441,7 @@ let rec travel_aux jinfo continuation_tac (expr_info:constr infos) =
try
check_not_nested (expr_info.f_id::expr_info.forbidden_ids) expr_info.info;
jinfo.otherS () expr_info continuation_tac expr_info
- with e when Errors.noncritical e ->
+ with e when CErrors.noncritical e ->
errorlabstrm "Recdef.travel" (str "the term " ++ Printer.pr_lconstr expr_info.info ++ str " can not contain a recursive call to " ++ pr_id expr_info.f_id)
end
| Lambda(n,t,b) ->
@@ -447,7 +449,7 @@ let rec travel_aux jinfo continuation_tac (expr_info:constr infos) =
try
check_not_nested (expr_info.f_id::expr_info.forbidden_ids) expr_info.info;
jinfo.otherS () expr_info continuation_tac expr_info
- with e when Errors.noncritical e ->
+ with e when CErrors.noncritical e ->
errorlabstrm "Recdef.travel" (str "the term " ++ Printer.pr_lconstr expr_info.info ++ str " can not contain a recursive call to " ++ pr_id expr_info.f_id)
end
| Case(ci,t,a,l) ->
@@ -558,12 +560,12 @@ let rec destruct_bounds_aux infos (bound,hyple,rechyps) lbounds g =
Proofview.V82.of_tactic (simplest_elim(mkApp(delayed_force lt_n_O,[|s_max|])));
Proofview.V82.of_tactic default_full_auto];
observe_tclTHENLIST (str "destruct_bounds_aux2")[
- observe_tac (str "clearing k ") (clear [id]);
+ observe_tac (str "clearing k ") (Proofview.V82.of_tactic (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
@@ -587,7 +589,7 @@ let rec destruct_bounds_aux infos (bound,hyple,rechyps) lbounds g =
| (_,v_bound)::l ->
observe_tclTHENLIST (str "destruct_bounds_aux3")[
Proofview.V82.of_tactic (simplest_elim (mkVar v_bound));
- clear [v_bound];
+ Proofview.V82.of_tactic (clear [v_bound]);
tclDO 2 (Proofview.V82.of_tactic intro);
onNthHypId 1
(fun p_hyp ->
@@ -643,7 +645,7 @@ let terminate_letin (na,b,t,e) expr_info continuation_tac info =
try
check_not_nested (expr_info.f_id::expr_info.forbidden_ids) b;
true
- with e when Errors.noncritical e -> false
+ with e when CErrors.noncritical e -> false
in
if forbid
then
@@ -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;
+ [Proofview.V82.of_tactic (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
@@ -698,7 +704,7 @@ let terminate_case next_step (ci,a,t,l) expr_info continuation_tac infos g =
try
check_not_nested (expr_info.f_id::expr_info.forbidden_ids) a;
false
- with e when Errors.noncritical e ->
+ with e when CErrors.noncritical e ->
true
in
let a' = infos.info in
@@ -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));
@@ -942,7 +948,7 @@ let rec destruct_hex expr_info acc l =
| (v,hex)::l ->
observe_tclTHENLIST (str "destruct_hex")[
Proofview.V82.of_tactic (simplest_case (mkVar hex));
- clear [hex];
+ Proofview.V82.of_tactic (clear [hex]);
tclDO 2 (Proofview.V82.of_tactic intro);
onNthHypId 1 (fun hp ->
onNthHypId 2 (fun p ->
@@ -1110,10 +1116,10 @@ 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 (Proofview.V82.of_tactic (Tactics.generalize [mkVar id])) (Proofview.V82.of_tactic (clear [id])))
))
;
- observe_tac (str "fix") (fix (Some hrec) (nargs+1));
+ observe_tac (str "fix") (Proofview.V82.of_tactic (fix (Some hrec) (nargs+1)));
h_intros args_id;
Proofview.V82.of_tactic (Simple.intro wf_rec_arg);
observe_tac (str "tac") (tac wf_rec_arg hrec wf_rec_arg acc_inv)
@@ -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
@@ -1275,12 +1281,12 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp
| Some s -> s
| None ->
try add_suffix current_proof_name "_subproof"
- with e when Errors.noncritical e ->
+ with e when CErrors.noncritical e ->
anomaly (Pp.str "open_new_goal with an unamed theorem")
in
let na = next_global_ident_away name [] in
if Termops.occur_existential gls_type then
- Errors.error "\"abstract\" cannot handle existentials";
+ CErrors.error "\"abstract\" cannot handle existentials";
let hook _ _ =
let opacity =
let na_ref = Libnames.Ident (Loc.ghost,na) in
@@ -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];
+ Proofview.V82.of_tactic (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,29 +1516,31 @@ 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 relation =
- fst (*FIXME*)(interp_constr
- env_with_pre_rec_args
- (Evd.from_env env_with_pre_rec_args)
- r)
+ (* Refresh the global universes, now including those of _F *)
+ let evm = Evd.from_env (Global.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, evuctx =
+ interp_constr env_with_pre_rec_args evm r
in
+ let evm = Evd.from_ctx evuctx in
let tcc_lemma_name = add_suffix function_name "_tcc" in
let tcc_lemma_constr = ref None in
(* let _ = Pp.msgnl (str "relation := " ++ Printer.pr_lconstr_env env_with_pre_rec_args relation) in *)
let hook _ _ =
let term_ref = Nametab.locate (qualid_of_ident term_id) in
let f_ref = declare_f function_name (IsProof Lemma) arg_types term_ref in
- let _ = Table.extraction_inline true [Ident (Loc.ghost,term_id)] in
+ let _ = Extraction_plugin.Table.extraction_inline true [Ident (Loc.ghost,term_id)] in
(* message "start second proof"; *)
let stop =
try com_eqn (List.length res_vars) equation_id functional_ref f_ref term_ref (subst_var function_name equation_lemma_type);
false
- with e when Errors.noncritical e ->
+ with e when CErrors.noncritical e ->
begin
if do_observe ()
- then msg_debug (str "Cannot create equation Lemma " ++ Errors.print e)
- else anomaly (Pp.str "Cannot create equation Lemma")
+ then Feedback.msg_debug (str "Cannot create equation Lemma " ++ CErrors.print e)
+ else CErrors.errorlabstrm "Cannot create equation Lemma"
+ (str "Cannot create equation lemma." ++ spc () ++
+ str "This may be because the function is nested-recursive.")
;
true
end
diff --git a/plugins/funind/recdef_plugin.mllib b/plugins/funind/recdef_plugin.mlpack
index ec1f5436..2b443f2a 100644
--- a/plugins/funind/recdef_plugin.mllib
+++ b/plugins/funind/recdef_plugin.mlpack
@@ -8,4 +8,3 @@ Invfun
Indfun
Merge
G_indfun
-Recdef_plugin_mod