summaryrefslogtreecommitdiff
path: root/plugins/funind/indfun.ml
diff options
context:
space:
mode:
Diffstat (limited to 'plugins/funind/indfun.ml')
-rw-r--r--plugins/funind/indfun.ml450
1 files changed, 236 insertions, 214 deletions
diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml
index 18817f50..57863ee6 100644
--- a/plugins/funind/indfun.ml
+++ b/plugins/funind/indfun.ml
@@ -1,8 +1,9 @@
-open Context.Rel.Declaration
open CErrors
+open Sorts
open Util
open Names
-open Term
+open Constr
+open EConstr
open Pp
open Indfun_common
open Libnames
@@ -11,39 +12,42 @@ open Glob_term
open Declarations
open Misctypes
open Decl_kinds
-open Sigma.Notations
-let is_rec_info scheme_info =
+module RelDecl = Context.Rel.Declaration
+
+let is_rec_info sigma scheme_info =
let test_branche min acc decl =
acc || (
let new_branche =
- it_mkProd_or_LetIn mkProp (fst (decompose_prod_assum (get_type decl))) in
- let free_rels_in_br = Termops.free_rels new_branche in
+ it_mkProd_or_LetIn mkProp (fst (decompose_prod_assum sigma (RelDecl.get_type decl))) in
+ let free_rels_in_br = Termops.free_rels sigma new_branche in
let max = min + scheme_info.Tactics.npredicates in
Int.Set.exists (fun i -> i >= min && i< max) free_rels_in_br
)
in
List.fold_left_i test_branche 1 false (List.rev scheme_info.Tactics.branches)
-let choose_dest_or_ind scheme_info =
- Tactics.induction_destruct (is_rec_info scheme_info) false
+let choose_dest_or_ind scheme_info args =
+ Proofview.tclBIND Proofview.tclEVARMAP (fun sigma ->
+ Tactics.induction_destruct (is_rec_info sigma scheme_info) false args)
let functional_induction with_clean c princl pat =
let res =
- let f,args = decompose_app c in
fun g ->
+ let sigma = Tacmach.project g in
+ let f,args = decompose_app sigma c in
let princ,bindings, princ_type,g' =
match princl with
| None -> (* No principle is given let's find the good one *)
begin
- match kind_of_term f with
+ match EConstr.kind sigma f with
| Const (c',u) ->
let princ_option =
let finfo = (* we first try to find out a graph on f *)
try find_Function_infos c'
with Not_found ->
- errorlabstrm "" (str "Cannot find induction information on "++
- Printer.pr_lconstr (mkConst c') )
+ user_err (str "Cannot find induction information on "++
+ Printer.pr_leconstr_env (Tacmach.pf_env g) sigma (mkConst c') )
in
match Tacticals.elimination_sort_of_goal g with
| InProp -> finfo.prop_lemma
@@ -61,7 +65,7 @@ let functional_induction with_clean c princl pat =
(or f_rec, f_rect) i*)
let princ_name =
Indrec.make_elimination_ident
- (Label.to_id (con_label c'))
+ (Label.to_id (Constant.label c'))
(Tacticals.elimination_sort_of_goal g)
in
try
@@ -70,30 +74,41 @@ let functional_induction with_clean c princl pat =
(b,a)
(* mkConst(const_of_id princ_name ),g (\* FIXME *\) *)
with Not_found -> (* This one is neither defined ! *)
- errorlabstrm "" (str "Cannot find induction principle for "
- ++Printer.pr_lconstr (mkConst c') )
+ user_err (str "Cannot find induction principle for "
+ ++ Printer.pr_leconstr_env (Tacmach.pf_env g) sigma (mkConst c') )
in
- (princ,NoBindings, Tacmach.pf_unsafe_type_of g' princ,g')
- | _ -> raise (UserError("",str "functional induction must be used with a function" ))
+ let princ = EConstr.of_constr princ in
+ (princ,NoBindings,Tacmach.pf_unsafe_type_of g' princ,g')
+ | _ -> raise (UserError(None,str "functional induction must be used with a function" ))
end
| Some ((princ,binding)) ->
princ,binding,Tacmach.pf_unsafe_type_of g princ,g
in
- let princ_infos = Tactics.compute_elim_sig princ_type in
+ let sigma = Tacmach.project g' in
+ let princ_infos = Tactics.compute_elim_sig (Tacmach.project g') princ_type in
let args_as_induction_constr =
let c_list =
if princ_infos.Tactics.farg_in_concl
then [c] else []
in
+ if List.length args + List.length c_list = 0
+ then user_err Pp.(str "Cannot recognize a valid functional scheme" );
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 ({ Tacexpr.delayed = fun env sigma -> Sigma ((c,NoBindings), sigma, Sigma.refl) })),(None,pat),None))
- (args@c_list) encoded_pat_as_patlist
+ List.make (List.length args + List.length c_list - 1) None @ [pat]
+ in
+ List.map2
+ (fun c pat ->
+ ((None,
+ Ltac_plugin.Tacexpr.ElimOnConstr (fun env sigma -> (sigma,(c,NoBindings)))),
+ (None,pat),
+ None))
+ (args@c_list)
+ encoded_pat_as_patlist
in
let princ' = Some (princ,bindings) in
let princ_vars =
List.fold_right
- (fun a acc -> try Id.Set.add (destVar a) acc with DestKO -> acc)
+ (fun a acc -> try Id.Set.add (destVar sigma a) acc with DestKO -> acc)
args
Id.Set.empty
in
@@ -128,15 +143,14 @@ let functional_induction with_clean c princl pat =
let rec abstract_glob_constr c = function
| [] -> c
- | Constrexpr.LocalRawDef (x,b)::bl -> Constrexpr_ops.mkLetInC(x,b,abstract_glob_constr c bl)
- | Constrexpr.LocalRawAssum (idl,k,t)::bl ->
+ | Constrexpr.CLocalDef (x,b,t)::bl -> Constrexpr_ops.mkLetInC(x,b,t,abstract_glob_constr c bl)
+ | Constrexpr.CLocalAssum (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
+ | Constrexpr.CLocalPattern _::bl -> assert false
let interp_casted_constr_with_implicits env sigma impls c =
- Constrintern.intern_gen Pretyping.WithoutTypeConstraint env ~impls
- ~allow_patvar:false c
+ Constrintern.intern_gen Pretyping.WithoutTypeConstraint env sigma ~impls c
(*
Construct a fixpoint as a Glob_term
@@ -149,14 +163,14 @@ let build_newrecursive
let sigma = Evd.from_env env0 in
let (rec_sign,rec_impls) =
List.fold_left
- (fun (env,impls) (((_,recname),_),bl,arityc,_) ->
- let arityc = Constrexpr_ops.prod_constr_expr arityc bl in
+ (fun (env,impls) (({CAst.v=recname},_),bl,arityc,_) ->
+ let arityc = Constrexpr_ops.mkCProdN bl arityc in
let arity,ctx = Constrintern.interp_type env0 sigma arityc in
- 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
+ let evd = Evd.from_env env0 in
+ let evd, (_, (_, impls')) = Constrintern.interp_context_evars env evd bl in
+ let impl = Constrintern.compute_internalization_data env0 evd Constrintern.Recursive arity impls' in
let open Context.Named.Declaration in
- (Environ.push_named (LocalAssum (recname,arity)) env, Id.Map.add recname impl impls))
+ (EConstr.push_named (LocalAssum (recname,arity)) env, Id.Map.add recname impl impls))
(env0,Constrintern.empty_internalization_env) lnameargsardef in
let recdef =
(* Declare local notations *)
@@ -175,37 +189,41 @@ let build_newrecursive l =
match body_opt with
| Some body ->
(fixna,bll,ar,body)
- | None -> user_err_loc (Loc.ghost,"Function",str "Body of Function must be given")
+ | None -> user_err ~hdr:"Function" (str "Body of Function must be given")
) l
in
build_newrecursive l'
+let error msg = user_err Pp.(str msg)
+
(* Checks whether or not the mutual bloc is recursive *)
let is_rec names =
let names = List.fold_right Id.Set.add names Id.Set.empty in
let check_id id names = Id.Set.mem id names in
- let rec lookup names = function
- | GVar(_,id) -> check_id id names
+ let rec lookup names gt = match DAst.get gt with
+ | GVar(id) -> check_id id names
| GRef _ | GEvar _ | GPatVar _ | GSort _ | GHole _ -> false
- | GCast(_,b,_) -> lookup names b
+ | GCast(b,_) -> lookup names b
| GRec _ -> error "GRec not handled"
- | GIf(_,b,_,lhs,rhs) ->
+ | GIf(b,_,lhs,rhs) ->
(lookup names b) || (lookup names lhs) || (lookup names rhs)
- | GLetIn(_,na,t,b) | GLambda(_,na,_,t,b) | GProd(_,na,_,t,b) ->
- lookup names t || lookup (Nameops.name_fold Id.Set.remove na names) b
- | GLetTuple(_,nal,_,t,b) -> lookup names t ||
+ | GProd(na,_,t,b) | GLambda(na,_,t,b) ->
+ lookup names t || lookup (Nameops.Name.fold_right Id.Set.remove na names) b
+ | GLetIn(na,b,t,c) ->
+ lookup names b || Option.cata (lookup names) true t || lookup (Nameops.Name.fold_right Id.Set.remove na names) c
+ | GLetTuple(nal,_,t,b) -> lookup names t ||
lookup
(List.fold_left
- (fun acc na -> Nameops.name_fold Id.Set.remove na acc)
+ (fun acc na -> Nameops.Name.fold_right Id.Set.remove na acc)
names
nal
)
b
- | GApp(_,f,args) -> List.exists (lookup names) (f::args)
- | GCases(_,_,_,el,brl) ->
+ | GApp(f,args) -> List.exists (lookup names) (f::args)
+ | GCases(_,_,el,brl) ->
List.exists (fun (e,_) -> lookup names e) el ||
List.exists (lookup_br names) brl
- and lookup_br names (_,idl,_,rt) =
+ and lookup_br names {CAst.v=(idl,_,rt)} =
let new_names = List.fold_right Id.Set.remove idl names in
lookup new_names rt
in
@@ -214,9 +232,9 @@ let is_rec names =
let rec local_binders_length = function
(* Assume that no `{ ... } contexts occur *)
| [] -> 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
+ | Constrexpr.CLocalDef _::bl -> 1 + local_binders_length bl
+ | Constrexpr.CLocalAssum (idl,_,_)::bl -> List.length idl + local_binders_length bl
+ | Constrexpr.CLocalPattern _::bl -> assert false
let prepare_body ((name,_,args,types,_),_) rt =
let n = local_binders_length args in
@@ -242,7 +260,9 @@ let derive_inversion fix_names =
let evd,c =
Evd.fresh_global
(Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident id)) in
- evd, destConst c::l
+ let c = EConstr.of_constr c in
+ let (cst, u) = destConst evd c in
+ evd, (cst, EInstance.kind evd u) :: l
)
fix_names
(evd',[])
@@ -262,14 +282,14 @@ let derive_inversion fix_names =
(Global.env ()) evd
(Constrintern.locate_reference (Libnames.qualid_of_ident (mk_rel_id id)))
in
- evd,(fst (destInd id))::l
+ let id = EConstr.of_constr id in
+ evd,(fst (destInd evd id))::l
)
fix_names
(evd',[])
in
Invfun.derive_correctness
Functional_principles_types.make_scheme
- functional_induction
fix_names_as_constant
lind;
with e when CErrors.noncritical e ->
@@ -321,7 +341,7 @@ let error_error names e =
in
match e with
| Building_graph e ->
- errorlabstrm ""
+ user_err
(str "Cannot define graph(s) for " ++
h 1 (prlist_with_sep (fun _ -> str","++spc ()) Ppconstr.pr_id names) ++
e_explain e)
@@ -329,9 +349,9 @@ let error_error names e =
let generate_principle (evd:Evd.evar_map ref) pconstants on_error
is_general do_built (fix_rec_l:(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) list) recdefs interactive_proof
- (continue_proof : int -> Names.constant array -> Term.constr array -> int ->
+ (continue_proof : int -> Names.Constant.t array -> EConstr.constr array -> int ->
Tacmach.tactic) : unit =
- let names = List.map (function (((_, name),_),_,_,_,_),_ -> name) fix_rec_l in
+ let names = List.map (function (({CAst.v=name},_),_,_,_,_),_ -> name) fix_rec_l in
let fun_bodies = List.map2 prepare_body fix_rec_l recdefs in
let funs_args = List.map fst fun_bodies in
let funs_types = List.map (function ((_,_,_,types,_),_) -> types) fix_rec_l in
@@ -344,7 +364,7 @@ let generate_principle (evd:Evd.evar_map ref) pconstants on_error
(*i The next call to mk_rel_id is valid since we have just construct the graph
Ensures by : do_built
i*)
- let f_R_mut = Ident (Loc.ghost,mk_rel_id (List.nth names 0)) in
+ let f_R_mut = CAst.make @@ Ident (mk_rel_id (List.nth names 0)) in
let ind_kn =
fst (locate_with_msg
(pr_reference f_R_mut++str ": Not an inductive type!")
@@ -352,7 +372,7 @@ let generate_principle (evd:Evd.evar_map ref) pconstants on_error
f_R_mut)
in
let fname_kn (((fname,_),_,_,_,_),_) =
- let f_ref = Ident fname in
+ let f_ref = CAst.map (fun n -> Ident n) fname in
locate_with_msg
(pr_reference f_ref++str ": Not an inductive type!")
locate_constant
@@ -367,7 +387,8 @@ let generate_principle (evd:Evd.evar_map ref) pconstants on_error
let evd = ref (Evd.from_env env) in
let evd',uprinc = Evd.fresh_global env !evd princ in
let _ = evd := evd' in
- let princ_type = Typing.e_type_of ~refresh:true env evd uprinc in
+ let princ_type = Typing.e_type_of ~refresh:true env evd (EConstr.of_constr uprinc) in
+ let princ_type = EConstr.Unsafe.to_constr princ_type in
Functional_principles_types.generate_functional_principle
evd
interactive_proof
@@ -390,33 +411,40 @@ let generate_principle (evd:Evd.evar_map ref) pconstants on_error
let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) list) =
match fixpoint_exprl with
- | [(((_,fname),pl),_,bl,ret_type,body),_] when not is_rec ->
- let body = match body with | Some body -> body | None -> user_err_loc (Loc.ghost,"Function",str "Body of Function must be given") in
- Command.do_definition
+ | [(({CAst.v=fname},pl),_,bl,ret_type,body),_] when not is_rec ->
+ let body = match body with | Some body -> body | None -> user_err ~hdr:"Function" (str "Body of Function must be given") in
+ ComDefinition.do_definition
+ ~program_mode:false
fname
(Decl_kinds.Global,(Flags.is_universe_polymorphism ()),Decl_kinds.Definition) pl
bl None body (Some ret_type) (Lemmas.mk_hook (fun _ _ -> ()));
let evd,rev_pconstants =
List.fold_left
- (fun (evd,l) ((((_,fname),_),_,_,_,_),_) ->
+ (fun (evd,l) ((({CAst.v=fname},_),_,_,_,_),_) ->
let evd,c =
Evd.fresh_global
(Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident fname)) in
- evd,((destConst c)::l)
+ let c = EConstr.of_constr c in
+ let (cst, u) = destConst evd c in
+ let u = EInstance.kind evd u in
+ evd,((cst, u) :: l)
)
(Evd.from_env (Global.env ()),[])
fixpoint_exprl
in
evd,List.rev rev_pconstants
| _ ->
- Command.do_fixpoint Global (Flags.is_universe_polymorphism ()) fixpoint_exprl;
+ ComFixpoint.do_fixpoint Global (Flags.is_universe_polymorphism ()) fixpoint_exprl;
let evd,rev_pconstants =
List.fold_left
- (fun (evd,l) ((((_,fname),_),_,_,_,_),_) ->
+ (fun (evd,l) ((({CAst.v=fname},_),_,_,_,_),_) ->
let evd,c =
Evd.fresh_global
(Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident fname)) in
- evd,((destConst c)::l)
+ let c = EConstr.of_constr c in
+ let (cst, u) = destConst evd c in
+ let u = EInstance.kind evd u in
+ evd,((cst, u) :: l)
)
(Evd.from_env (Global.env ()),[])
fixpoint_exprl
@@ -426,7 +454,7 @@ let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexp
let generate_correction_proof_wf f_ref tcc_lemma_ref
is_mes functional_ref eq_ref rec_arg_num rec_arg_type nb_args relation
- (_: int) (_:Names.constant array) (_:Term.constr array) (_:int) : Tacmach.tactic =
+ (_: int) (_:Names.Constant.t array) (_:EConstr.constr array) (_:int) : Tacmach.tactic =
Functional_principles_proofs.prove_principle_for_gen
(f_ref,functional_ref,eq_ref)
tcc_lemma_ref is_mes rec_arg_num rec_arg_type relation
@@ -435,11 +463,11 @@ let generate_correction_proof_wf f_ref tcc_lemma_ref
let register_wf ?(is_mes=false) fname rec_impls wf_rel_expr wf_arg using_lemmas args ret_type body
pre_hook
=
- let type_of_f = Constrexpr_ops.prod_constr_expr ret_type args in
+ let type_of_f = Constrexpr_ops.mkCProdN args ret_type in
let rec_arg_num =
let names =
List.map
- snd
+ CAst.(with_val (fun x -> x))
(Constrexpr_ops.names_of_local_assums args)
in
match wf_arg with
@@ -451,22 +479,21 @@ let register_wf ?(is_mes=false) fname rec_impls wf_rel_expr wf_arg using_lemmas
in
let unbounded_eq =
let f_app_args =
- Constrexpr.CAppExpl
- (Loc.ghost,
- (None,(Ident (Loc.ghost,fname)),None) ,
+ CAst.make @@ Constrexpr.CAppExpl(
+ (None,CAst.make @@ Ident fname,None) ,
(List.map
(function
- | _,Anonymous -> assert false
- | _,Name e -> (Constrexpr_ops.mkIdentC e)
+ | {CAst.v=Anonymous} -> assert false
+ | {CAst.v=Name e} -> (Constrexpr_ops.mkIdentC e)
)
(Constrexpr_ops.names_of_local_assums args)
)
)
in
- Constrexpr.CApp (Loc.ghost,(None,Constrexpr_ops.mkRefC (Qualid (Loc.ghost,(qualid_of_string "Logic.eq")))),
+ CAst.make @@ Constrexpr.CApp ((None,Constrexpr_ops.mkRefC (CAst.make @@ Qualid (qualid_of_string "Logic.eq"))),
[(f_app_args,None);(body,None)])
in
- let eq = Constrexpr_ops.prod_constr_expr unbounded_eq args in
+ let eq = Constrexpr_ops.mkCProdN args unbounded_eq in
let hook ((f_ref,_) as fconst) tcc_lemma_ref (functional_ref,_) (eq_ref,_) rec_arg_num rec_arg_type
nb_args relation =
try
@@ -495,7 +522,7 @@ let register_mes fname rec_impls wf_mes_expr wf_rel_expr_opt wf_arg using_lemmas
| None ->
begin
match args with
- | [Constrexpr.LocalRawAssum ([(_,Name x)],k,t)] -> t,x
+ | [Constrexpr.CLocalAssum ([{CAst.v=Name x}],k,t)] -> t,x
| _ -> error "Recursive argument must be specified"
end
| Some wf_args ->
@@ -503,15 +530,15 @@ let register_mes fname rec_impls wf_mes_expr wf_rel_expr_opt wf_arg using_lemmas
match
List.find
(function
- | Constrexpr.LocalRawAssum(l,k,t) ->
+ | Constrexpr.CLocalAssum(l,k,t) ->
List.exists
- (function (_,Name id) -> Id.equal id wf_args | _ -> false)
+ (function {CAst.v=Name id} -> Id.equal id wf_args | _ -> false)
l
| _ -> false
)
args
with
- | Constrexpr.LocalRawAssum(_,k,t) -> t,wf_args
+ | Constrexpr.CLocalAssum(_,k,t) -> t,wf_args
| _ -> assert false
with Not_found -> assert false
in
@@ -520,13 +547,13 @@ let register_mes fname rec_impls wf_mes_expr wf_rel_expr_opt wf_arg using_lemmas
| None ->
let ltof =
let make_dir l = DirPath.make (List.rev_map Id.of_string l) in
- Libnames.Qualid (Loc.ghost,Libnames.qualid_of_path
+ CAst.make @@ Libnames.Qualid (Libnames.qualid_of_path
(Libnames.make_path (make_dir ["Arith";"Wf_nat"]) (Id.of_string "ltof")))
in
let fun_from_mes =
let applied_mes =
Constrexpr_ops.mkAppC(wf_mes_expr,[Constrexpr_ops.mkIdentC wf_arg]) in
- Constrexpr_ops.mkLambdaC ([(Loc.ghost,Name wf_arg)],Constrexpr_ops.default_binder_kind,wf_arg_type,applied_mes)
+ Constrexpr_ops.mkLambdaC ([CAst.make @@ Name wf_arg],Constrexpr_ops.default_binder_kind,wf_arg_type,applied_mes)
in
let wf_rel_from_mes =
Constrexpr_ops.mkAppC(Constrexpr_ops.mkRefC ltof,[wf_arg_type;fun_from_mes])
@@ -537,7 +564,7 @@ let register_mes fname rec_impls wf_mes_expr wf_rel_expr_opt wf_arg using_lemmas
let a = Names.Id.of_string "___a" in
let b = Names.Id.of_string "___b" in
Constrexpr_ops.mkLambdaC(
- [Loc.ghost,Name a;Loc.ghost,Name b],
+ [CAst.make @@ Name a; CAst.make @@ Name b],
Constrexpr.Default Explicit,
wf_arg_type,
Constrexpr_ops.mkAppC(wf_rel_expr,
@@ -557,60 +584,54 @@ let map_option f = function
| Some v -> Some (f v)
open Constrexpr
-open Topconstr
-let make_assoc assoc l1 l2 =
- let fold assoc a b = match a, b with
- | (_, Name na), (_, Name id) -> Id.Map.add na id assoc
- | _, _ -> assoc
- in
- List.fold_left2 fold assoc l1 l2
-
-let rec rebuild_bl (aux,assoc) bl typ =
- match bl,typ with
- | [], _ -> (List.rev aux,replace_vars_constr_expr assoc typ,assoc)
- | (Constrexpr.LocalRawAssum(nal,bk,_))::bl',typ ->
- rebuild_nal (aux,assoc) bk bl' nal (List.length nal) typ
- | (Constrexpr.LocalRawDef(na,_))::bl',Constrexpr.CLetIn(_,_,nat,typ') ->
- rebuild_bl ((Constrexpr.LocalRawDef(na,replace_vars_constr_expr assoc nat)::aux),assoc)
+let rec rebuild_bl aux bl typ =
+ match bl,typ with
+ | [], _ -> List.rev aux,typ
+ | (CLocalAssum(nal,bk,_))::bl',typ ->
+ rebuild_nal aux bk bl' nal typ
+ | (CLocalDef(na,_,_))::bl',{ CAst.v = CLetIn(_,nat,ty,typ') } ->
+ rebuild_bl (Constrexpr.CLocalDef(na,nat,ty)::aux)
bl' typ'
| _ -> assert false
- and rebuild_nal (aux,assoc) bk bl' nal lnal typ =
+and rebuild_nal aux bk bl' nal typ =
match nal,typ with
- | [], _ -> rebuild_bl (aux,assoc) bl' typ
- | _,CProdN(_,[],typ) -> rebuild_nal (aux,assoc) bk bl' nal lnal typ
- | _,CProdN(_,(nal',bk',nal't)::rest,typ') ->
- let lnal' = List.length nal' in
- if lnal' >= lnal
- then
- let old_nal',new_nal' = List.chop lnal nal' in
- let nassoc = make_assoc assoc old_nal' nal in
- let assum = LocalRawAssum(nal,bk,replace_vars_constr_expr assoc nal't) in
- rebuild_bl ((assum :: aux), nassoc) bl'
- (if List.is_empty new_nal' && List.is_empty rest
- then typ'
- else if List.is_empty new_nal'
- then CProdN(Loc.ghost,rest,typ')
- else CProdN(Loc.ghost,((new_nal',bk',nal't)::rest),typ'))
- else
- let captured_nal,non_captured_nal = List.chop lnal' nal in
- let nassoc = make_assoc assoc nal' captured_nal in
- let assum = LocalRawAssum(captured_nal,bk,replace_vars_constr_expr assoc nal't) in
- rebuild_nal ((assum :: aux), nassoc)
- bk bl' non_captured_nal (lnal - lnal') (CProdN(Loc.ghost,rest,typ'))
- | _ -> assert false
-
-let rebuild_bl (aux,assoc) bl typ = rebuild_bl (aux,assoc) bl typ
+ | _,{ CAst.v = CProdN([],typ) } -> rebuild_nal aux bk bl' nal typ
+ | [], _ -> rebuild_bl aux bl' typ
+ | na::nal,{ CAst.v = CProdN(CLocalAssum(na'::nal',bk',nal't)::rest,typ') } ->
+ if Name.equal (na.CAst.v) (na'.CAst.v) || Name.is_anonymous (na'.CAst.v)
+ then
+ let assum = CLocalAssum([na],bk,nal't) in
+ let new_rest = if nal' = [] then rest else (CLocalAssum(nal',bk',nal't)::rest) in
+ rebuild_nal
+ (assum::aux)
+ bk
+ bl'
+ nal
+ (CAst.make @@ CProdN(new_rest,typ'))
+ else
+ let assum = CLocalAssum([na'],bk,nal't) in
+ let new_rest = if nal' = [] then rest else (CLocalAssum(nal',bk',nal't)::rest) in
+ rebuild_nal
+ (assum::aux)
+ bk
+ bl'
+ (na::nal)
+ (CAst.make @@ CProdN(new_rest,typ'))
+ | _ ->
+ assert false
+
+let rebuild_bl aux bl typ = rebuild_bl aux bl typ
let recompute_binder_list (fixpoint_exprl : (Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) list) =
- let fixl,ntns = Command.extract_fixpoint_components false fixpoint_exprl in
- let ((_,_,typel),_,ctx,_) = Command.interp_fixpoint fixl ntns in
+ let fixl,ntns = ComFixpoint.extract_fixpoint_components false fixpoint_exprl in
+ let ((_,_,typel),_,ctx,_) = ComFixpoint.interp_fixpoint ~cofix:false fixl ntns in
let constr_expr_typel =
- with_full_print (List.map (Constrextern.extern_constr false (Global.env ()) (Evd.from_ctx ctx))) typel in
+ with_full_print (List.map (fun c -> Constrextern.extern_constr false (Global.env ()) (Evd.from_ctx ctx) (EConstr.of_constr c))) typel in
let fixpoint_exprl_with_new_bl =
List.map2 (fun ((lna,(rec_arg_opt,rec_order),bl,ret_typ,opt_body),notation_list) fix_typ ->
- let new_bl',new_ret_type,_ = rebuild_bl ([],Id.Map.empty) bl fix_typ in
+ let new_bl',new_ret_type = rebuild_bl [] bl fix_typ in
(((lna,(rec_arg_opt,rec_order),new_bl',new_ret_type,opt_body),notation_list):(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list))
)
fixpoint_exprl constr_expr_typel
@@ -624,13 +645,13 @@ let do_generate_principle pconstants on_error register_built interactive_proof
let _is_struct =
match fixpoint_exprl with
| [((_,(wf_x,Constrexpr.CWfRec wf_rel),_,_,_),_) as fixpoint_expr] ->
- let (((((_,name),pl),_,args,types,body)),_) as fixpoint_expr =
+ let (((({CAst.v=name},pl),_,args,types,body)),_) as fixpoint_expr =
match recompute_binder_list [fixpoint_expr] with
| [e] -> e
| _ -> assert false
in
let fixpoint_exprl = [fixpoint_expr] in
- let body = match body with | Some body -> body | None -> user_err_loc (Loc.ghost,"Function",str "Body of Function must be given") in
+ let body = match body with | Some body -> body | None -> user_err ~hdr:"Function" (str "Body of Function must be given") in
let recdefs,rec_impls = build_newrecursive fixpoint_exprl in
let using_lemmas = [] in
let pre_hook pconstants =
@@ -645,10 +666,10 @@ let do_generate_principle pconstants on_error register_built interactive_proof
true
in
if register_built
- then register_wf name rec_impls wf_rel (map_option snd wf_x) using_lemmas args types body pre_hook;
+ then register_wf name rec_impls wf_rel (map_option (fun x -> x.CAst.v) wf_x) using_lemmas args types body pre_hook;
false
|[((_,(wf_x,Constrexpr.CMeasureRec(wf_mes,wf_rel_opt)),_,_,_),_) as fixpoint_expr] ->
- let (((((_,name),_),_,args,types,body)),_) as fixpoint_expr =
+ let (((({CAst.v=name},_),_,args,types,body)),_) as fixpoint_expr =
match recompute_binder_list [fixpoint_expr] with
| [e] -> e
| _ -> assert false
@@ -656,7 +677,7 @@ let do_generate_principle pconstants on_error register_built interactive_proof
let fixpoint_exprl = [fixpoint_expr] in
let recdefs,rec_impls = build_newrecursive fixpoint_exprl in
let using_lemmas = [] in
- let body = match body with | Some body -> body | None -> user_err_loc (Loc.ghost,"Function",str "Body of Function must be given") in
+ let body = match body with | Some body -> body | None -> user_err ~hdr:"Function" (str "Body of Function must be given") in
let pre_hook pconstants =
generate_principle
(ref (Evd.from_env (Global.env ())))
@@ -669,7 +690,7 @@ let do_generate_principle pconstants on_error register_built interactive_proof
true
in
if register_built
- then register_mes name rec_impls wf_mes wf_rel_opt (map_option snd wf_x) using_lemmas args types body pre_hook;
+ then register_mes name rec_impls wf_mes wf_rel_opt (map_option (fun x -> x.CAst.v) wf_x) using_lemmas args types body pre_hook;
true
| _ ->
List.iter (function ((_na,(_,ord),_args,_body,_type),_not) ->
@@ -682,7 +703,7 @@ let do_generate_principle pconstants on_error register_built interactive_proof
fixpoint_exprl;
let fixpoint_exprl = recompute_binder_list fixpoint_exprl in
let fix_names =
- List.map (function ((((_,name),_),_,_,_,_),_) -> name) fixpoint_exprl
+ List.map (function ((({CAst.v=name},_),_,_,_,_),_) -> name) fixpoint_exprl
in
(* ok all the expressions are structural *)
let recdefs,rec_impls = build_newrecursive fixpoint_exprl in
@@ -708,67 +729,69 @@ let do_generate_principle pconstants on_error register_built interactive_proof
in
()
-let rec add_args id new_args b =
- match b with
- | CRef (r,_) ->
- begin match r with
- | Libnames.Ident(loc,fname) when Id.equal fname id ->
- CAppExpl(Loc.ghost,(None,r,None),new_args)
+let rec add_args id new_args = CAst.map (function
+ | CRef (r,_) as b ->
+ begin match r with
+ | {CAst.v=Libnames.Ident fname} when Id.equal fname id ->
+ CAppExpl((None,r,None),new_args)
| _ -> b
end
- | CFix _ | CCoFix _ -> anomaly ~label:"add_args " (Pp.str "todo")
- | CProdN(loc,nal,b1) ->
- CProdN(loc,
- List.map (fun (nal,k,b2) -> (nal,k,add_args id new_args b2)) nal,
+ | CFix _ | CCoFix _ -> anomaly ~label:"add_args " (Pp.str "todo.")
+ | CProdN(nal,b1) ->
+ CProdN(List.map (function CLocalAssum (nal,k,b2) -> CLocalAssum (nal,k,add_args id new_args b2)
+ | CLocalDef (na,b1,t) -> CLocalDef (na,add_args id new_args b1,Option.map (add_args id new_args) t)
+ | CLocalPattern _ -> user_err (Pp.str "pattern with quote not allowed here.")) nal,
add_args id new_args b1)
- | CLambdaN(loc,nal,b1) ->
- CLambdaN(loc,
- List.map (fun (nal,k,b2) -> (nal,k,add_args id new_args b2)) nal,
+ | CLambdaN(nal,b1) ->
+ CLambdaN(List.map (function CLocalAssum (nal,k,b2) -> CLocalAssum (nal,k,add_args id new_args b2)
+ | CLocalDef (na,b1,t) -> CLocalDef (na,add_args id new_args b1,Option.map (add_args id new_args) t)
+ | CLocalPattern _ -> user_err (Pp.str "pattern with quote not allowed here.")) nal,
add_args id new_args b1)
- | CLetIn(loc,na,b1,b2) ->
- CLetIn(loc,na,add_args id new_args b1,add_args id new_args b2)
- | CAppExpl(loc,(pf,r,us),exprl) ->
+ | CLetIn(na,b1,t,b2) ->
+ CLetIn(na,add_args id new_args b1,Option.map (add_args id new_args) t,add_args id new_args b2)
+ | CAppExpl((pf,r,us),exprl) ->
begin
match r with
- | Libnames.Ident(loc,fname) when Id.equal fname id ->
- CAppExpl(loc,(pf,r,us),new_args@(List.map (add_args id new_args) exprl))
- | _ -> CAppExpl(loc,(pf,r,us),List.map (add_args id new_args) exprl)
+ | {CAst.v=Libnames.Ident fname} when Id.equal fname id ->
+ CAppExpl((pf,r,us),new_args@(List.map (add_args id new_args) exprl))
+ | _ -> CAppExpl((pf,r,us),List.map (add_args id new_args) exprl)
end
- | CApp(loc,(pf,b),bl) ->
- CApp(loc,(pf,add_args id new_args b),
+ | CApp((pf,b),bl) ->
+ CApp((pf,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,
+ | CCases(sty,b_option,cel,cal) ->
+ CCases(sty,Option.map (add_args id new_args) b_option,
List.map (fun (b,na,b_option) ->
add_args id new_args b,
na, b_option) cel,
- List.map (fun (loc,cpl,e) -> (loc,cpl,add_args id new_args e)) cal
+ List.map CAst.(map (fun (cpl,e) -> (cpl,add_args id new_args e))) cal
)
- | CLetTuple(loc,nal,(na,b_option),b1,b2) ->
- CLetTuple(loc,nal,(na,Option.map (add_args id new_args) b_option),
+ | CLetTuple(nal,(na,b_option),b1,b2) ->
+ CLetTuple(nal,(na,Option.map (add_args id new_args) b_option),
add_args id new_args b1,
add_args id new_args b2
)
- | CIf(loc,b1,(na,b_option),b2,b3) ->
- CIf(loc,add_args id new_args b1,
+ | CIf(b1,(na,b_option),b2,b3) ->
+ CIf(add_args id new_args b1,
(na,Option.map (add_args id new_args) b_option),
add_args id new_args b2,
add_args id new_args b3
)
- | CHole _ -> b
- | CPatVar _ -> b
- | CEvar _ -> b
- | CSort _ -> b
- | CCast(loc,b1,b2) ->
- CCast(loc,add_args id new_args b1,
+ | CHole _
+ | CPatVar _
+ | CEvar _
+ | CPrim _
+ | CSort _ as b -> b
+ | CCast(b1,b2) ->
+ CCast(add_args id new_args b1,
Miscops.map_cast_type (add_args id new_args) b2)
- | 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
- | CDelimiters _ -> anomaly ~label:"add_args " (Pp.str "CDelimiters")
+ | CRecord pars ->
+ CRecord (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.")
+ | CDelimiters _ -> anomaly ~label:"add_args " (Pp.str "CDelimiters.")
+ )
exception Stop of Constrexpr.constr_expr
@@ -779,8 +802,8 @@ let rec chop_n_arrow n t =
if n <= 0
then t (* If we have already removed all the arrows then return the type *)
else (* If not we check the form of [t] *)
- match t with
- | Constrexpr.CProdN(_,nal_ta',t') -> (* If we have a forall, to result are possible :
+ match t.CAst.v with
+ | Constrexpr.CProdN(nal_ta',t') -> (* If we have a forall, two results are possible :
either we need to discard more than the number of arrows contained
in this product declaration then we just recall [chop_n_arrow] on
the remaining number of arrow to chop and [t'] we discard it and
@@ -792,104 +815,103 @@ let rec chop_n_arrow n t =
let new_n =
let rec aux (n:int) = function
[] -> n
- | (nal,k,t'')::nal_ta' ->
+ | CLocalAssum(nal,k,t'')::nal_ta' ->
let nal_l = List.length nal in
if n >= nal_l
then
aux (n - nal_l) nal_ta'
else
- let new_t' =
- Constrexpr.CProdN(Loc.ghost,
- ((snd (List.chop n nal)),k,t'')::nal_ta',t')
+ let new_t' = CAst.make @@
+ Constrexpr.CProdN(
+ CLocalAssum((snd (List.chop n nal)),k,t'')::nal_ta',t')
in
raise (Stop new_t')
+ | _ -> anomaly (Pp.str "Not enough products.")
in
aux n nal_ta'
in
chop_n_arrow new_n t'
with Stop t -> t
end
- | _ -> anomaly (Pp.str "Not enough products")
+ | _ -> anomaly (Pp.str "Not enough products.")
-let rec get_args b t : Constrexpr.local_binder list *
+let rec get_args b t : Constrexpr.local_binder_expr list *
Constrexpr.constr_expr * Constrexpr.constr_expr =
- match b with
- | Constrexpr.CLambdaN (loc, (nal_ta), b') ->
+ match b.CAst.v with
+ | Constrexpr.CLambdaN (CLocalAssum(nal,k,ta) as d::rest, b') ->
begin
- let n =
- (List.fold_left (fun n (nal,_,_) ->
- n+List.length nal) 0 nal_ta )
- in
- let nal_tas,b'',t'' = get_args b' (chop_n_arrow n t) in
- (List.map (fun (nal,k,ta) ->
- (Constrexpr.LocalRawAssum (nal,k,ta))) nal_ta)@nal_tas, b'',t''
+ let n = List.length nal in
+ let nal_tas,b'',t'' = get_args (CAst.make ?loc:b.CAst.loc @@ Constrexpr.CLambdaN (rest,b')) (chop_n_arrow n t) in
+ d :: nal_tas, b'',t''
end
+ | Constrexpr.CLambdaN ([], b) -> [],b,t
| _ -> [],b,t
let make_graph (f_ref:global_reference) =
let c,c_body =
match f_ref with
- | ConstRef c ->
- begin try c,Global.lookup_constant c
- with Not_found ->
- raise (UserError ("",str "Cannot find " ++ Printer.pr_lconstr (mkConst c)) )
- end
- | _ -> raise (UserError ("", str "Not a function reference") )
+ | ConstRef c ->
+ begin try c,Global.lookup_constant c
+ with Not_found ->
+ let sigma, env = Pfedit.get_current_context () in
+ raise (UserError (None,str "Cannot find " ++ Printer.pr_leconstr_env env sigma (mkConst c)) )
+ end
+ | _ -> raise (UserError (None, str "Not a function reference") )
in
(match Global.body_of_constant_body c_body with
- | None -> error "Cannot build a graph over an axiom !"
- | Some body ->
+ | None -> error "Cannot build a graph over an axiom!"
+ | Some (body, _) ->
let env = Global.env () in
let sigma = Evd.from_env env in
let extern_body,extern_type =
with_full_print (fun () ->
- (Constrextern.extern_constr false env sigma body,
+ (Constrextern.extern_constr false env sigma (EConstr.of_constr body),
Constrextern.extern_type false env sigma
- ((*FIXME*) Typeops.type_of_constant_type env c_body.const_type)
+ (EConstr.of_constr (*FIXME*) c_body.const_type)
)
)
()
in
let (nal_tas,b,t) = get_args extern_body extern_type in
let expr_list =
- match b with
- | Constrexpr.CFix(loc,l_id,fixexprl) ->
+ match b.CAst.v with
+ | Constrexpr.CFix(l_id,fixexprl) ->
let l =
List.map
(fun (id,(n,recexp),bl,t,b) ->
- let loc, rec_id = Option.get n in
+ let { CAst.loc; v=rec_id } = Option.get n in
let new_args =
List.flatten
(List.map
(function
- | Constrexpr.LocalRawDef (na,_)-> []
- | Constrexpr.LocalRawAssum (nal,_,_) ->
+ | Constrexpr.CLocalDef (na,_,_)-> []
+ | Constrexpr.CLocalAssum (nal,_,_) ->
List.map
- (fun (loc,n) ->
- CRef(Libnames.Ident(loc, Nameops.out_name n),None))
+ (fun {CAst.loc;v=n} -> CAst.make ?loc @@
+ CRef(CAst.make ?loc @@ Libnames.Ident(Nameops.Name.get_id n),None))
nal
- | Constrexpr.LocalPattern _ -> assert false
+ | Constrexpr.CLocalPattern _ -> assert false
)
nal_tas
)
in
- let b' = add_args (snd id) new_args b in
- ((((id,None), ( Some (Loc.ghost,rec_id),CStructRec),nal_tas@bl,t,Some b'),[]):(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list))
+ let b' = add_args id.CAst.v new_args b in
+ ((((id,None), ( Some CAst.(make rec_id),CStructRec),nal_tas@bl,t,Some b'),[]):(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list))
)
fixexprl
in
l
| _ ->
- let id = Label.to_id (con_label c) in
- [(((Loc.ghost,id),None),(None,Constrexpr.CStructRec),nal_tas,t,Some b),[]]
+ let id = Label.to_id (Constant.label c) in
+ [((CAst.make id,None),(None,Constrexpr.CStructRec),nal_tas,t,Some b),[]]
in
- let mp,dp,_ = repr_con c in
+ let mp,dp,_ = Constant.repr3 c in
do_generate_principle [c,Univ.Instance.empty] error_error false false expr_list;
(* We register the infos *)
List.iter
- (fun ((((_,id),_),_,_,_,_),_) -> add_Function false (make_con mp dp (Label.of_id id)))
+ (fun ((({CAst.v=id},_),_,_,_,_),_) -> add_Function false (Constant.make3 mp dp (Label.of_id id)))
expr_list)
let do_generate_principle = do_generate_principle [] warning_error true