From 9043add656177eeac1491a73d2f3ab92bec0013c Mon Sep 17 00:00:00 2001 From: Benjamin Barenblat Date: Sat, 29 Dec 2018 14:31:27 -0500 Subject: Imported Upstream version 8.8.2 --- vernac/assumptions.ml | 356 +++++++ vernac/assumptions.mli | 33 + vernac/auto_ind_decl.ml | 993 +++++++++++++++++ vernac/auto_ind_decl.mli | 44 + vernac/class.ml | 325 ++++++ vernac/class.mli | 50 + vernac/classes.ml | 438 ++++++++ vernac/classes.mli | 69 ++ vernac/comAssumption.ml | 182 ++++ vernac/comAssumption.mli | 36 + vernac/comDefinition.ml | 134 +++ vernac/comDefinition.mli | 32 + vernac/comFixpoint.ml | 353 +++++++ vernac/comFixpoint.mli | 95 ++ vernac/comInductive.ml | 453 ++++++++ vernac/comInductive.mli | 67 ++ vernac/comProgramFixpoint.ml | 342 ++++++ vernac/comProgramFixpoint.mli | 12 + vernac/declareDef.ml | 66 ++ vernac/declareDef.mli | 24 + vernac/explainErr.ml | 128 +++ vernac/explainErr.mli | 23 + vernac/himsg.ml | 1317 +++++++++++++++++++++++ vernac/himsg.mli | 46 + vernac/indschemes.ml | 517 +++++++++ vernac/indschemes.mli | 53 + vernac/lemmas.ml | 535 ++++++++++ vernac/lemmas.mli | 72 ++ vernac/locality.ml | 68 ++ vernac/locality.mli | 40 + vernac/metasyntax.ml | 1551 +++++++++++++++++++++++++++ vernac/metasyntax.mli | 63 ++ vernac/mltop.ml | 473 +++++++++ vernac/mltop.mli | 102 ++ vernac/obligations.ml | 1196 +++++++++++++++++++++ vernac/obligations.mli | 111 ++ vernac/proof_using.ml | 192 ++++ vernac/proof_using.mli | 23 + vernac/record.ml | 620 +++++++++++ vernac/record.mli | 35 + vernac/search.ml | 380 +++++++ vernac/search.mli | 85 ++ vernac/topfmt.ml | 334 ++++++ vernac/topfmt.mli | 62 ++ vernac/vernac.mllib | 26 + vernac/vernacentries.ml | 2349 +++++++++++++++++++++++++++++++++++++++++ vernac/vernacentries.mli | 42 + vernac/vernacinterp.ml | 87 ++ vernac/vernacinterp.mli | 30 + vernac/vernacprop.ml | 60 ++ vernac/vernacprop.mli | 28 + vernac/vernacstate.ml | 43 + vernac/vernacstate.mli | 21 + 53 files changed, 14816 insertions(+) create mode 100644 vernac/assumptions.ml create mode 100644 vernac/assumptions.mli create mode 100644 vernac/auto_ind_decl.ml create mode 100644 vernac/auto_ind_decl.mli create mode 100644 vernac/class.ml create mode 100644 vernac/class.mli create mode 100644 vernac/classes.ml create mode 100644 vernac/classes.mli create mode 100644 vernac/comAssumption.ml create mode 100644 vernac/comAssumption.mli create mode 100644 vernac/comDefinition.ml create mode 100644 vernac/comDefinition.mli create mode 100644 vernac/comFixpoint.ml create mode 100644 vernac/comFixpoint.mli create mode 100644 vernac/comInductive.ml create mode 100644 vernac/comInductive.mli create mode 100644 vernac/comProgramFixpoint.ml create mode 100644 vernac/comProgramFixpoint.mli create mode 100644 vernac/declareDef.ml create mode 100644 vernac/declareDef.mli create mode 100644 vernac/explainErr.ml create mode 100644 vernac/explainErr.mli create mode 100644 vernac/himsg.ml create mode 100644 vernac/himsg.mli create mode 100644 vernac/indschemes.ml create mode 100644 vernac/indschemes.mli create mode 100644 vernac/lemmas.ml create mode 100644 vernac/lemmas.mli create mode 100644 vernac/locality.ml create mode 100644 vernac/locality.mli create mode 100644 vernac/metasyntax.ml create mode 100644 vernac/metasyntax.mli create mode 100644 vernac/mltop.ml create mode 100644 vernac/mltop.mli create mode 100644 vernac/obligations.ml create mode 100644 vernac/obligations.mli create mode 100644 vernac/proof_using.ml create mode 100644 vernac/proof_using.mli create mode 100644 vernac/record.ml create mode 100644 vernac/record.mli create mode 100644 vernac/search.ml create mode 100644 vernac/search.mli create mode 100644 vernac/topfmt.ml create mode 100644 vernac/topfmt.mli create mode 100644 vernac/vernac.mllib create mode 100644 vernac/vernacentries.ml create mode 100644 vernac/vernacentries.mli create mode 100644 vernac/vernacinterp.ml create mode 100644 vernac/vernacinterp.mli create mode 100644 vernac/vernacprop.ml create mode 100644 vernac/vernacprop.mli create mode 100644 vernac/vernacstate.ml create mode 100644 vernac/vernacstate.mli (limited to 'vernac') diff --git a/vernac/assumptions.ml b/vernac/assumptions.ml new file mode 100644 index 00000000..742ff3a9 --- /dev/null +++ b/vernac/assumptions.ml @@ -0,0 +1,356 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* raise Not_found + | (l, SFBmodule mb) :: _ when Label.equal l lab -> mb + | _ :: fields -> search_mod_label lab fields + +let rec search_cst_label lab = function + | [] -> raise Not_found + | (l, SFBconst cb) :: _ when Label.equal l lab -> cb + | _ :: fields -> search_cst_label lab fields + +let rec search_mind_label lab = function + | [] -> raise Not_found + | (l, SFBmind mind) :: _ when Label.equal l lab -> mind + | _ :: fields -> search_mind_label lab fields + +(* TODO: using [empty_delta_resolver] below is probably slightly incorrect. But: + a) I don't see currently what should be used instead + b) this shouldn't be critical for Print Assumption. At worse some + constants will have a canonical name which is non-canonical, + leading to failures in [Global.lookup_constant], but our own + [lookup_constant] should work. +*) + +let rec fields_of_functor f subs mp0 args = function + |NoFunctor a -> f subs mp0 args a + |MoreFunctor (mbid,_,e) -> + match args with + | [] -> assert false (* we should only encounter applied functors *) + | mpa :: args -> + let subs = join (map_mbid mbid mpa empty_delta_resolver (*TODO*)) subs in + fields_of_functor f subs mp0 args e + +let rec lookup_module_in_impl mp = + match mp with + | MPfile _ -> Global.lookup_module mp + | MPbound _ -> Global.lookup_module mp + | MPdot (mp',lab') -> + if ModPath.equal mp' (Global.current_modpath ()) then + Global.lookup_module mp + else + let fields = memoize_fields_of_mp mp' in + search_mod_label lab' fields + +and memoize_fields_of_mp mp = + try MPmap.find mp !modcache + with Not_found -> + let l = fields_of_mp mp in + modcache := MPmap.add mp l !modcache; + l + +and fields_of_mp mp = + let mb = lookup_module_in_impl mp in + let fields,inner_mp,subs = fields_of_mb empty_subst mb [] in + let subs = + if ModPath.equal inner_mp mp then subs + else add_mp inner_mp mp mb.mod_delta subs + in + Modops.subst_structure subs fields + +and fields_of_mb subs mb args = match mb.mod_expr with + |Algebraic expr -> fields_of_expression subs mb.mod_mp args expr + |Struct sign -> fields_of_signature subs mb.mod_mp args sign + |Abstract|FullStruct -> fields_of_signature subs mb.mod_mp args mb.mod_type + +(** The Abstract case above corresponds to [Declare Module] *) + +and fields_of_signature x = + fields_of_functor + (fun subs mp0 args struc -> + assert (List.is_empty args); + (struc, mp0, subs)) x + +and fields_of_expr subs mp0 args = function + |MEident mp -> + let mb = lookup_module_in_impl (subst_mp subs mp) in + fields_of_mb subs mb args + |MEapply (me1,mp2) -> fields_of_expr subs mp0 (mp2::args) me1 + |MEwith _ -> assert false (* no 'with' in [mod_expr] *) + +and fields_of_expression x = fields_of_functor fields_of_expr x + +let lookup_constant_in_impl cst fallback = + try + let mp,dp,lab = KerName.repr (Constant.canonical cst) in + let fields = memoize_fields_of_mp mp in + (* A module found this way is necessarily closed, in particular + our constant cannot be in an opened section : *) + search_cst_label lab fields + with Not_found -> + (* Either: + - The module part of the constant isn't registered yet : + we're still in it, so the [constant_body] found earlier + (if any) was a true axiom. + - The label has not been found in the structure. This is an error *) + match fallback with + | Some cb -> cb + | None -> anomaly (str "Print Assumption: unknown constant " ++ Constant.print cst ++ str ".") + +let lookup_constant cst = + try + let cb = Global.lookup_constant cst in + if Declareops.constant_has_body cb then cb + else lookup_constant_in_impl cst (Some cb) + with Not_found -> lookup_constant_in_impl cst None + +let lookup_mind_in_impl mind = + try + let mp,dp,lab = KerName.repr (MutInd.canonical mind) in + let fields = memoize_fields_of_mp mp in + search_mind_label lab fields + with Not_found -> + anomaly (str "Print Assumption: unknown inductive " ++ MutInd.print mind ++ str ".") + +let lookup_mind mind = + try Global.lookup_mind mind + with Not_found -> lookup_mind_in_impl mind + +(** Graph traversal of an object, collecting on the way the dependencies of + traversed objects *) + +let label_of = function + | ConstRef kn -> pi3 (Constant.repr3 kn) + | IndRef (kn,_) + | ConstructRef ((kn,_),_) -> pi3 (MutInd.repr3 kn) + | VarRef id -> Label.of_id id + +let fold_constr_with_full_binders g f n acc c = + let open Context.Rel.Declaration in + match Constr.kind c with + | Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ | Construct _ -> acc + | Cast (c,_, t) -> f n (f n acc c) t + | Prod (na,t,c) -> f (g (LocalAssum (na,t)) n) (f n acc t) c + | Lambda (na,t,c) -> f (g (LocalAssum (na,t)) n) (f n acc t) c + | LetIn (na,b,t,c) -> f (g (LocalDef (na,b,t)) n) (f n (f n acc b) t) c + | App (c,l) -> Array.fold_left (f n) (f n acc c) l + | Proj (p,c) -> f n acc c + | Evar (_,l) -> Array.fold_left (f n) acc l + | Case (_,p,c,bl) -> Array.fold_left (f n) (f n (f n acc p) c) bl + | Fix (_,(lna,tl,bl)) -> + let n' = CArray.fold_left2 (fun c n t -> g (LocalAssum (n,t)) c) n lna tl in + let fd = Array.map2 (fun t b -> (t,b)) tl bl in + Array.fold_left (fun acc (t,b) -> f n' (f n acc t) b) acc fd + | CoFix (_,(lna,tl,bl)) -> + let n' = CArray.fold_left2 (fun c n t -> g (LocalAssum (n,t)) c) n lna tl in + let fd = Array.map2 (fun t b -> (t,b)) tl bl in + Array.fold_left (fun acc (t,b) -> f n' (f n acc t) b) acc fd + +let rec traverse current ctx accu t = match Constr.kind t with +| Var id -> + let body () = id |> Global.lookup_named |> NamedDecl.get_value in + traverse_object accu body (VarRef id) +| Const (kn, _) -> + let body () = Option.map fst (Global.body_of_constant_body (lookup_constant kn)) in + traverse_object accu body (ConstRef kn) +| Ind ((mind, _) as ind, _) -> + traverse_inductive accu mind (IndRef ind) +| Construct (((mind, _), _) as cst, _) -> + traverse_inductive accu mind (ConstructRef cst) +| Meta _ | Evar _ -> assert false +| Case (_,oty,c,[||]) -> + (* non dependent match on an inductive with no constructors *) + begin match Constr.(kind oty, kind c) with + | Lambda(_,_,oty), Const (kn, _) + when Vars.noccurn 1 oty && + not (Declareops.constant_has_body (lookup_constant kn)) -> + let body () = Option.map fst (Global.body_of_constant_body (lookup_constant kn)) in + traverse_object + ~inhabits:(current,ctx,Vars.subst1 mkProp oty) accu body (ConstRef kn) + | _ -> + fold_constr_with_full_binders + Context.Rel.add (traverse current) ctx accu t + end +| _ -> fold_constr_with_full_binders + Context.Rel.add (traverse current) ctx accu t + +and traverse_object ?inhabits (curr, data, ax2ty) body obj = + let data, ax2ty = + let already_in = Refmap_env.mem obj data in + match body () with + | None -> + let data = + if not already_in then Refmap_env.add obj Refset_env.empty data else data in + let ax2ty = + if Option.is_empty inhabits then ax2ty else + let ty = Option.get inhabits in + try let l = Refmap_env.find obj ax2ty in Refmap_env.add obj (ty::l) ax2ty + with Not_found -> Refmap_env.add obj [ty] ax2ty in + data, ax2ty + | Some body -> + if already_in then data, ax2ty else + let contents,data,ax2ty = + traverse (label_of obj) Context.Rel.empty + (Refset_env.empty,data,ax2ty) body in + Refmap_env.add obj contents data, ax2ty + in + (Refset_env.add obj curr, data, ax2ty) + +(** Collects the references occurring in the declaration of mutual inductive + definitions. All the constructors and names of a mutual inductive + definition share exactly the same dependencies. Also, there is no explicit + dependency between mutually defined inductives and constructors. *) +and traverse_inductive (curr, data, ax2ty) mind obj = + let firstind_ref = (IndRef (mind, 0)) in + let label = label_of obj in + let data, ax2ty = + (* Invariant : I_0 \in data iff I_i \in data iff c_ij \in data + where I_0, I_1, ... are in the same mutual definition and c_ij + are all their constructors. *) + if Refmap_env.mem firstind_ref data then data, ax2ty else + let mib = lookup_mind mind in + (* Collects references of parameters *) + let param_ctx = mib.mind_params_ctxt in + let nparam = List.length param_ctx in + let accu = + traverse_context label Context.Rel.empty + (Refset_env.empty, data, ax2ty) param_ctx + in + (* Build the context of all arities *) + let arities_ctx = + let global_env = Global.env () in + Array.fold_left (fun accu oib -> + let pspecif = Univ.in_punivs (mib, oib) in + let ind_type = Inductive.type_of_inductive global_env pspecif in + let ind_name = Name oib.mind_typename in + Context.Rel.add (Context.Rel.Declaration.LocalAssum (ind_name, ind_type)) accu) + Context.Rel.empty mib.mind_packets + in + (* For each inductive, collects references in their arity and in the type + of constructors*) + let (contents, data, ax2ty) = Array.fold_left (fun accu oib -> + let arity_wo_param = + List.rev (List.skipn nparam (List.rev oib.mind_arity_ctxt)) + in + let accu = + traverse_context + label param_ctx accu arity_wo_param + in + Array.fold_left (fun accu cst_typ -> + let param_ctx, cst_typ_wo_param = Term.decompose_prod_n_assum nparam cst_typ in + let ctx = Context.(Rel.fold_outside Context.Rel.add ~init:arities_ctx param_ctx) in + traverse label ctx accu cst_typ_wo_param) + accu oib.mind_user_lc) + accu mib.mind_packets + in + (* Maps all these dependencies to inductives and constructors*) + let data = Array.fold_left_i (fun n data oib -> + let ind = (mind, n) in + let data = Refmap_env.add (IndRef ind) contents data in + Array.fold_left_i (fun k data _ -> + Refmap_env.add (ConstructRef (ind, k+1)) contents data + ) data oib.mind_consnames) data mib.mind_packets + in + data, ax2ty + in + (Refset_env.add obj curr, data, ax2ty) + +(** Collects references in a rel_context. *) +and traverse_context current ctx accu ctxt = + snd (Context.Rel.fold_outside (fun decl (ctx, accu) -> + match decl with + | Context.Rel.Declaration.LocalDef (_,c,t) -> + let accu = traverse current ctx (traverse current ctx accu t) c in + let ctx = Context.Rel.add decl ctx in + ctx, accu + | Context.Rel.Declaration.LocalAssum (_,t) -> + let accu = traverse current ctx accu t in + let ctx = Context.Rel.add decl ctx in + ctx, accu) ctxt ~init:(ctx, accu)) + +let traverse current t = + let () = modcache := MPmap.empty in + traverse current Context.Rel.empty (Refset_env.empty, Refmap_env.empty, Refmap_env.empty) t + +(** Hopefully bullet-proof function to recover the type of a constant. It just + ignores all the universe stuff. There are many issues that can arise when + considering terms out of any valid environment, so use with caution. *) +let type_of_constant cb = cb.Declarations.const_type + +let assumptions ?(add_opaque=false) ?(add_transparent=false) st gr t = + let (idts, knst) = st in + (** Only keep the transitive dependencies *) + let (_, graph, ax2ty) = traverse (label_of gr) t in + let fold obj _ accu = match obj with + | VarRef id -> + let decl = Global.lookup_named id in + if is_local_assum decl then + let t = Context.Named.Declaration.get_type decl in + ContextObjectMap.add (Variable id) t accu + else accu + | ConstRef kn -> + let cb = lookup_constant kn in + let accu = + if cb.const_typing_flags.check_guarded then accu + else + let l = try Refmap_env.find obj ax2ty with Not_found -> [] in + ContextObjectMap.add (Axiom (Guarded kn, l)) Constr.mkProp accu + in + if not (Declareops.constant_has_body cb) || not cb.const_typing_flags.check_universes then + let t = type_of_constant cb in + let l = try Refmap_env.find obj ax2ty with Not_found -> [] in + ContextObjectMap.add (Axiom (Constant kn,l)) t accu + else if add_opaque && (Declareops.is_opaque cb || not (Cpred.mem kn knst)) then + let t = type_of_constant cb in + ContextObjectMap.add (Opaque kn) t accu + else if add_transparent then + let t = type_of_constant cb in + ContextObjectMap.add (Transparent kn) t accu + else + accu + | IndRef (m,_) | ConstructRef ((m,_),_) -> + let mind = lookup_mind m in + if mind.mind_typing_flags.check_guarded then + accu + else + let l = try Refmap_env.find obj ax2ty with Not_found -> [] in + ContextObjectMap.add (Axiom (Positive m, l)) Constr.mkProp accu + in + Refmap_env.fold fold graph ContextObjectMap.empty diff --git a/vernac/assumptions.mli b/vernac/assumptions.mli new file mode 100644 index 00000000..7e13f8f2 --- /dev/null +++ b/vernac/assumptions.mli @@ -0,0 +1,33 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* constr -> + (Refset_env.t * Refset_env.t Refmap_env.t * + (Label.t * Context.Rel.t * types) list Refmap_env.t) + +(** Collects all the assumptions (optionally including opaque definitions) + on which a term relies (together with their type). The above warning of + {!traverse} also applies. *) +val assumptions : + ?add_opaque:bool -> ?add_transparent:bool -> transparent_state -> + global_reference -> constr -> types ContextObjectMap.t diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml new file mode 100644 index 00000000..0404d015 --- /dev/null +++ b/vernac/auto_ind_decl.ml @@ -0,0 +1,993 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* [] + | t::q -> t::(kick_last q) + | [] -> failwith "kick_last" +and aux = function + | (0,l') -> l' + | (n,h::t) -> aux (n-1,t) + | _ -> failwith "quick_chop" + in + if n > (List.length l) then failwith "quick_chop args" + else kick_last (aux (n,l) ) + +let deconstruct_type t = + let l,r = decompose_prod t in + (List.rev_map snd l)@[r] + +exception EqNotFound of inductive * inductive +exception EqUnknown of string +exception UndefinedCst of string +exception InductiveWithProduct +exception InductiveWithSort +exception ParameterWithoutEquality of global_reference +exception NonSingletonProp of inductive +exception DecidabilityMutualNotSupported +exception NoDecidabilityCoInductive + +let constr_of_global g = lazy (Universes.constr_of_global g) + +(* Some pre declaration of constant we are going to use *) +let bb = constr_of_global Coqlib.glob_bool + +let andb_prop = fun _ -> Universes.constr_of_global (Coqlib.build_bool_type()).Coqlib.andb_prop + +let andb_true_intro = fun _ -> + Universes.constr_of_global + (Coqlib.build_bool_type()).Coqlib.andb_true_intro + +let tt = constr_of_global Coqlib.glob_true + +let ff = constr_of_global Coqlib.glob_false + +let eq = constr_of_global Coqlib.glob_eq + +let sumbool () = Universes.constr_of_global (Coqlib.build_coq_sumbool ()) + +let andb = fun _ -> Universes.constr_of_global (Coqlib.build_bool_type()).Coqlib.andb + +let induct_on c = induction false None c None None + +let destruct_on c = destruct false None c None None + +let destruct_on_using c id = + destruct false None c + (Some (CAst.make @@ IntroOrPattern [[CAst.make @@ IntroNaming IntroAnonymous]; + [CAst.make @@ IntroNaming (IntroIdentifier id)]])) + None + +let destruct_on_as c l = + destruct false None c (Some (CAst.make l)) None + +let inj_flags = Some { + Equality.keep_proof_equalities = true; (* necessary *) + injection_in_context = true; (* does not matter here *) + Equality.injection_pattern_l2r_order = true; (* does not matter here *) + } + +let my_discr_tac = Equality.discr_tac false None +let my_inj_tac x = Equality.inj inj_flags None false None (EConstr.mkVar x,NoBindings) + +(* reconstruct the inductive with the correct de Bruijn indexes *) +let mkFullInd (ind,u) n = + let mib = Global.lookup_mind (fst ind) in + let nparams = mib.mind_nparams in + let nparrec = mib.mind_nparams_rec in + (* params context divided *) + let lnonparrec,lnamesparrec = + context_chop (nparams-nparrec) mib.mind_params_ctxt in + if nparrec > 0 + then mkApp (mkIndU (ind,u), + Array.of_list(Context.Rel.to_extended_list mkRel (nparrec+n) lnamesparrec)) + else mkIndU (ind,u) + +let check_bool_is_defined () = + try let _ = Global.type_of_global_in_context (Global.env ()) Coqlib.glob_bool in () + with e when CErrors.noncritical e -> raise (UndefinedCst "bool") + +let beq_scheme_kind_aux = ref (fun _ -> failwith "Undefined") + +let build_beq_scheme mode kn = + check_bool_is_defined (); + (* fetching global env *) + let env = Global.env() in + (* fetching the mutual inductive body *) + let mib = Global.lookup_mind kn in + (* number of inductives in the mutual *) + let nb_ind = Array.length mib.mind_packets in + (* number of params in the type *) + let nparams = mib.mind_nparams in + let nparrec = mib.mind_nparams_rec in + (* params context divided *) + let lnonparrec,lnamesparrec = + context_chop (nparams-nparrec) mib.mind_params_ctxt in + (* predef coq's boolean type *) + (* rec name *) + let rec_name i =(Id.to_string (Array.get mib.mind_packets i).mind_typename)^ + "_eqrec" + in + (* construct the "fun A B ... N, eqA eqB eqC ... N => fixpoint" part *) + let create_input c = + let myArrow u v = mkArrow u (lift 1 v) + and eqName = function + | Name s -> Id.of_string ("eq_"^(Id.to_string s)) + | Anonymous -> Id.of_string "eq_A" + in + let ext_rel_list = Context.Rel.to_extended_list mkRel 0 lnamesparrec in + let lift_cnt = ref 0 in + let eqs_typ = List.map (fun aa -> + let a = lift !lift_cnt aa in + incr lift_cnt; + myArrow a (myArrow a (Lazy.force bb)) + ) ext_rel_list in + + let eq_input = List.fold_left2 + ( fun a b decl -> (* mkLambda(n,b,a) ) *) + (* here I leave the Naming thingy so that the type of + the function is more readable for the user *) + mkNamedLambda (eqName (RelDecl.get_name decl)) b a ) + c (List.rev eqs_typ) lnamesparrec + in + List.fold_left (fun a decl ->(* mkLambda(n,t,a)) eq_input rel_list *) + (* Same here , hoping the auto renaming will do something good ;) *) + mkNamedLambda + (match RelDecl.get_name decl with Name s -> s | Anonymous -> Id.of_string "A") + (RelDecl.get_type decl) a) eq_input lnamesparrec + in + let make_one_eq cur = + let u = Univ.Instance.empty in + let ind = (kn,cur),u (* FIXME *) in + (* current inductive we are working on *) + let cur_packet = mib.mind_packets.(snd (fst ind)) in + (* Inductive toto : [rettyp] := *) + let rettyp = Inductive.type_of_inductive env ((mib,cur_packet),u) in + (* split rettyp in a list without the non rec params and the last -> + e.g. Inductive vec (A:Set) : nat -> Set := ... will do [nat] *) + let rettyp_l = quick_chop nparrec (deconstruct_type rettyp) in + (* give a type A, this function tries to find the equality on A declared + previously *) + (* nlist = the number of args (A , B , ... ) + eqA = the de Bruijn index of the first eq param + ndx = how much to translate due to the 2nd Case + *) + let compute_A_equality rel_list nlist eqA ndx t = + let lifti = ndx in + let sigma = Evd.empty (** FIXME *) in + let rec aux c = + let (c,a) = Reductionops.whd_betaiota_stack Evd.empty c in + match EConstr.kind sigma c with + | Rel x -> mkRel (x-nlist+ndx), Safe_typing.empty_private_constants + | Var x -> + let eid = Id.of_string ("eq_"^(Id.to_string x)) in + let () = + try ignore (Environ.lookup_named eid env) + with Not_found -> raise (ParameterWithoutEquality (VarRef x)) + in + mkVar eid, Safe_typing.empty_private_constants + | Cast (x,_,_) -> aux (EConstr.applist (x,a)) + | App _ -> assert false + | Ind ((kn',i as ind'),u) (*FIXME: universes *) -> + if MutInd.equal kn kn' then mkRel(eqA-nlist-i+nb_ind-1), Safe_typing.empty_private_constants + else begin + try + let eq, eff = + let c, eff = find_scheme ~mode (!beq_scheme_kind_aux()) (kn',i) in + mkConst c, eff in + let eqa, eff = + let eqa, effs = List.split (List.map aux a) in + Array.of_list eqa, + List.fold_left Safe_typing.concat_private eff (List.rev effs) + in + let args = + Array.append + (Array.of_list (List.map (fun x -> lift lifti (EConstr.Unsafe.to_constr x)) a)) eqa in + if Int.equal (Array.length args) 0 then eq, eff + else mkApp (eq, args), eff + with Not_found -> raise(EqNotFound (ind', fst ind)) + end + | Sort _ -> raise InductiveWithSort + | Prod _ -> raise InductiveWithProduct + | Lambda _-> raise (EqUnknown "abstraction") + | LetIn _ -> raise (EqUnknown "let-in") + | Const (kn, u) -> + let u = EConstr.EInstance.kind sigma u in + (match Environ.constant_opt_value_in env (kn, u) with + | None -> raise (ParameterWithoutEquality (ConstRef kn)) + | Some c -> aux (EConstr.applist (EConstr.of_constr c,a))) + | Proj _ -> raise (EqUnknown "projection") + | Construct _ -> raise (EqUnknown "constructor") + | Case _ -> raise (EqUnknown "match") + | CoFix _ -> raise (EqUnknown "cofix") + | Fix _ -> raise (EqUnknown "fix") + | Meta _ -> raise (EqUnknown "meta-variable") + | Evar _ -> raise (EqUnknown "existential variable") + in + aux t + in + (* construct the predicate for the Case part*) + let do_predicate rel_list n = + List.fold_left (fun a b -> mkLambda(Anonymous,b,a)) + (mkLambda (Anonymous, + mkFullInd ind (n+3+(List.length rettyp_l)+nb_ind-1), + (Lazy.force bb))) + (List.rev rettyp_l) in + (* make_one_eq *) + (* do the [| C1 ... => match Y with ... end + ... + Cn => match Y with ... end |] part *) + let ci = make_case_info env (fst ind) MatchStyle in + let constrs n = get_constructors env (make_ind_family (ind, + Context.Rel.to_extended_list mkRel (n+nb_ind-1) mib.mind_params_ctxt)) in + let constrsi = constrs (3+nparrec) in + let n = Array.length constrsi in + let ar = Array.make n (Lazy.force ff) in + let eff = ref Safe_typing.empty_private_constants in + for i=0 to n-1 do + let nb_cstr_args = List.length constrsi.(i).cs_args in + let ar2 = Array.make n (Lazy.force ff) in + let constrsj = constrs (3+nparrec+nb_cstr_args) in + for j=0 to n-1 do + if Int.equal i j then + ar2.(j) <- let cc = (match nb_cstr_args with + | 0 -> Lazy.force tt + | _ -> let eqs = Array.make nb_cstr_args (Lazy.force tt) in + for ndx = 0 to nb_cstr_args-1 do + let cc = RelDecl.get_type (List.nth constrsi.(i).cs_args ndx) in + let eqA, eff' = compute_A_equality rel_list + nparrec + (nparrec+3+2*nb_cstr_args) + (nb_cstr_args+ndx+1) + (EConstr.of_constr cc) + in + eff := Safe_typing.concat_private eff' !eff; + Array.set eqs ndx + (mkApp (eqA, + [|mkRel (ndx+1+nb_cstr_args);mkRel (ndx+1)|] + )) + done; + Array.fold_left + (fun a b -> mkApp (andb(),[|b;a|])) + (eqs.(0)) + (Array.sub eqs 1 (nb_cstr_args - 1)) + ) + in + (List.fold_left (fun a decl -> mkLambda (RelDecl.get_name decl, RelDecl.get_type decl, a)) cc + (constrsj.(j).cs_args) + ) + else ar2.(j) <- (List.fold_left (fun a decl -> + mkLambda (RelDecl.get_name decl, RelDecl.get_type decl, a)) (Lazy.force ff) (constrsj.(j).cs_args) ) + done; + + ar.(i) <- (List.fold_left (fun a decl -> mkLambda (RelDecl.get_name decl, RelDecl.get_type decl, a)) + (mkCase (ci,do_predicate rel_list nb_cstr_args, + mkVar (Id.of_string "Y") ,ar2)) + (constrsi.(i).cs_args)) + done; + mkNamedLambda (Id.of_string "X") (mkFullInd ind (nb_ind-1+1)) ( + mkNamedLambda (Id.of_string "Y") (mkFullInd ind (nb_ind-1+2)) ( + mkCase (ci, do_predicate rel_list 0,mkVar (Id.of_string "X"),ar))), + !eff + in (* build_beq_scheme *) + let names = Array.make nb_ind Anonymous and + types = Array.make nb_ind mkSet and + cores = Array.make nb_ind mkSet in + let eff = ref Safe_typing.empty_private_constants in + let u = Univ.Instance.empty in + for i=0 to (nb_ind-1) do + names.(i) <- Name (Id.of_string (rec_name i)); + types.(i) <- mkArrow (mkFullInd ((kn,i),u) 0) + (mkArrow (mkFullInd ((kn,i),u) 1) (Lazy.force bb)); + let c, eff' = make_one_eq i in + cores.(i) <- c; + eff := Safe_typing.concat_private eff' !eff + done; + (Array.init nb_ind (fun i -> + let kelim = Inductive.elim_sorts (mib,mib.mind_packets.(i)) in + if not (Sorts.List.mem InSet kelim) then + raise (NonSingletonProp (kn,i)); + let fix = match mib.mind_finite with + | CoFinite -> + raise NoDecidabilityCoInductive; + | Finite -> + mkFix (((Array.make nb_ind 0),i),(names,types,cores)) + | BiFinite -> + (** If the inductive type is not recursive, the fixpoint is not + used, so let's replace it with garbage *) + let subst = List.init nb_ind (fun _ -> mkProp) in + Vars.substl subst cores.(i) + in + create_input fix), + UState.make (Global.universes ())), + !eff + +let beq_scheme_kind = declare_mutual_scheme_object "_beq" build_beq_scheme + +let _ = beq_scheme_kind_aux := fun () -> beq_scheme_kind + +(* This function tryies to get the [inductive] between a constr + the constr should be Ind i or App(Ind i,[|args|]) +*) +let destruct_ind sigma c = + let open EConstr in + try let u,v = destApp sigma c in + let indc = destInd sigma u in + indc,v + with DestKO -> let indc = destInd sigma c in + indc,[||] + +(* + In the following, avoid is the list of names to avoid. + If the args of the Inductive type are A1 ... An + then avoid should be + [| lb_An ... lb _A1 (resp. bl_An ... bl_A1) + eq_An .... eq_A1 An ... A1 |] +so from Ai we can find the correct eq_Ai bl_ai or lb_ai +*) +(* used in the leib -> bool side*) +let do_replace_lb mode lb_scheme_key aavoid narg p q = + let open EConstr in + let avoid = Array.of_list aavoid in + let do_arg sigma v offset = + try + let x = narg*offset in + let s = destVar sigma v in + let n = Array.length avoid in + let rec find i = + if Id.equal avoid.(n-i) s then avoid.(n-i-x) + else (if i + (* if this happen then the args have to be already declared as a + Parameter*) + ( + let mp,dir,lbl = Constant.repr3 (fst (destConst sigma v)) in + mkConst (Constant.make3 mp dir (Label.make ( + if Int.equal offset 1 then ("eq_"^(Label.to_string lbl)) + else ((Label.to_string lbl)^"_lb") + ))) + ) + in + Proofview.Goal.enter begin fun gl -> + let type_of_pq = Tacmach.New.pf_unsafe_type_of gl p in + let sigma = Tacmach.New.project gl in + let env = Tacmach.New.pf_env gl in + let u,v = destruct_ind sigma type_of_pq + in let lb_type_of_p = + try + let c, eff = find_scheme ~mode lb_scheme_key (fst u) (*FIXME*) in + Proofview.tclUNIT (mkConst c, eff) + with Not_found -> + (* spiwack: the format of this error message should probably + be improved. *) + let err_msg = + (str "Leibniz->boolean:" ++ + str "You have to declare the" ++ + str "decidability over " ++ + Printer.pr_econstr_env env sigma type_of_pq ++ + str " first.") + in + Tacticals.New.tclZEROMSG err_msg + in + lb_type_of_p >>= fun (lb_type_of_p,eff) -> + Proofview.tclEVARMAP >>= fun sigma -> + let lb_args = Array.append (Array.append + (Array.map (fun x -> x) v) + (Array.map (fun x -> do_arg sigma x 1) v)) + (Array.map (fun x -> do_arg sigma x 2) v) + in let app = if Array.is_empty lb_args + then lb_type_of_p else mkApp (lb_type_of_p,lb_args) + in + Tacticals.New.tclTHENLIST [ + Proofview.tclEFFECTS eff; + Equality.replace p q ; apply app ; Auto.default_auto] + end + +(* used in the bool -> leib side *) +let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt = + let open EConstr in + let avoid = Array.of_list aavoid in + let do_arg sigma v offset = + try + let x = narg*offset in + let s = destVar sigma v in + let n = Array.length avoid in + let rec find i = + if Id.equal avoid.(n-i) s then avoid.(n-i-x) + else (if i + (* if this happen then the args have to be already declared as a + Parameter*) + ( + let mp,dir,lbl = Constant.repr3 (fst (destConst sigma v)) in + mkConst (Constant.make3 mp dir (Label.make ( + if Int.equal offset 1 then ("eq_"^(Label.to_string lbl)) + else ((Label.to_string lbl)^"_bl") + ))) + ) + in + + let rec aux l1 l2 = + match (l1,l2) with + | (t1::q1,t2::q2) -> + Proofview.Goal.enter begin fun gl -> + let tt1 = Tacmach.New.pf_unsafe_type_of gl t1 in + let sigma = Tacmach.New.project gl in + let env = Tacmach.New.pf_env gl in + if EConstr.eq_constr sigma t1 t2 then aux q1 q2 + else ( + let u,v = try destruct_ind sigma tt1 + (* trick so that the good sequence is returned*) + with e when CErrors.noncritical e -> indu,[||] + in if eq_ind (fst u) ind + then Tacticals.New.tclTHENLIST [Equality.replace t1 t2; Auto.default_auto ; aux q1 q2 ] + else ( + let bl_t1, eff = + try + let c, eff = find_scheme bl_scheme_key (fst u) (*FIXME*) in + mkConst c, eff + with Not_found -> + (* spiwack: the format of this error message should probably + be improved. *) + let err_msg = + (str "boolean->Leibniz:" ++ + str "You have to declare the" ++ + str "decidability over " ++ + Printer.pr_econstr_env env sigma tt1 ++ + str " first.") + in + user_err err_msg + in let bl_args = + Array.append (Array.append + (Array.map (fun x -> x) v) + (Array.map (fun x -> do_arg sigma x 1) v)) + (Array.map (fun x -> do_arg sigma x 2) v ) + in + let app = if Array.is_empty bl_args + then bl_t1 else mkApp (bl_t1,bl_args) + in + Tacticals.New.tclTHENLIST [ + Proofview.tclEFFECTS eff; + Equality.replace_by t1 t2 + (Tacticals.New.tclTHEN (apply app) (Auto.default_auto)) ; + aux q1 q2 ] + ) + ) + end + | ([],[]) -> Proofview.tclUNIT () + | _ -> Tacticals.New.tclZEROMSG (str "Both side of the equality must have the same arity.") + in + Proofview.tclEVARMAP >>= fun sigma -> + begin try Proofview.tclUNIT (destApp sigma lft) + with DestKO -> Tacticals.New.tclZEROMSG (str "replace failed.") + end >>= fun (ind1,ca1) -> + begin try Proofview.tclUNIT (destApp sigma rgt) + with DestKO -> Tacticals.New.tclZEROMSG (str "replace failed.") + end >>= fun (ind2,ca2) -> + begin try Proofview.tclUNIT (fst (destInd sigma ind1)) + with DestKO -> + begin try Proofview.tclUNIT (fst (fst (destConstruct sigma ind1))) + with DestKO -> Tacticals.New.tclZEROMSG (str "The expected type is an inductive one.") + end + end >>= fun (sp1,i1) -> + begin try Proofview.tclUNIT (fst (destInd sigma ind2)) + with DestKO -> + begin try Proofview.tclUNIT (fst (fst (destConstruct sigma ind2))) + with DestKO -> Tacticals.New.tclZEROMSG (str "The expected type is an inductive one.") + end + end >>= fun (sp2,i2) -> + if not (MutInd.equal sp1 sp2) || not (Int.equal i1 i2) + then Tacticals.New.tclZEROMSG (str "Eq should be on the same type") + else aux (Array.to_list ca1) (Array.to_list ca2) + +(* + create, from a list of ids [i1,i2,...,in] the list + [(in,eq_in,in_bl,in_al),,...,(i1,eq_i1,i1_bl_i1_al )] +*) +let list_id l = List.fold_left ( fun a decl -> let s' = + match RelDecl.get_name decl with + Name s -> Id.to_string s + | Anonymous -> "A" in + (Id.of_string s',Id.of_string ("eq_"^s'), + Id.of_string (s'^"_bl"), + Id.of_string (s'^"_lb")) + ::a + ) [] l +(* + build the right eq_I A B.. N eq_A .. eq_N +*) +let eqI ind l = + let list_id = list_id l in + let eA = Array.of_list((List.map (fun (s,_,_,_) -> mkVar s) list_id)@ + (List.map (fun (_,seq,_,_)-> mkVar seq) list_id )) + and e, eff = + try let c, eff = find_scheme beq_scheme_kind ind in mkConst c, eff + with Not_found -> user_err ~hdr:"AutoIndDecl.eqI" + (str "The boolean equality on " ++ MutInd.print (fst ind) ++ str " is needed."); + in (if Array.equal Constr.equal eA [||] then e else mkApp(e,eA)), eff + +(**********************************************************************) +(* Boolean->Leibniz *) + +open Namegen + +let compute_bl_goal ind lnamesparrec nparrec = + let eqI, eff = eqI ind lnamesparrec in + let list_id = list_id lnamesparrec in + let avoid = List.fold_right (Nameops.Name.fold_right (fun id l -> Id.Set.add id l)) (List.map RelDecl.get_name lnamesparrec) Id.Set.empty in + let create_input c = + let x = next_ident_away (Id.of_string "x") avoid and + y = next_ident_away (Id.of_string "y") avoid in + let bl_typ = List.map (fun (s,seq,_,_) -> + mkNamedProd x (mkVar s) ( + mkNamedProd y (mkVar s) ( + mkArrow + ( mkApp(Lazy.force eq,[|(Lazy.force bb);mkApp(mkVar seq,[|mkVar x;mkVar y|]);(Lazy.force tt)|])) + ( mkApp(Lazy.force eq,[|mkVar s;mkVar x;mkVar y|])) + )) + ) list_id in + let bl_input = List.fold_left2 ( fun a (s,_,sbl,_) b -> + mkNamedProd sbl b a + ) c (List.rev list_id) (List.rev bl_typ) in + let eqs_typ = List.map (fun (s,_,_,_) -> + mkProd(Anonymous,mkVar s,mkProd(Anonymous,mkVar s,(Lazy.force bb))) + ) list_id in + let eq_input = List.fold_left2 ( fun a (s,seq,_,_) b -> + mkNamedProd seq b a + ) bl_input (List.rev list_id) (List.rev eqs_typ) in + List.fold_left (fun a decl -> mkNamedProd + (match RelDecl.get_name decl with Name s -> s | Anonymous -> next_ident_away (Id.of_string "A") avoid) + (RelDecl.get_type decl) a) eq_input lnamesparrec + in + let n = next_ident_away (Id.of_string "x") avoid and + m = next_ident_away (Id.of_string "y") avoid in + let u = Univ.Instance.empty in + create_input ( + mkNamedProd n (mkFullInd (ind,u) nparrec) ( + mkNamedProd m (mkFullInd (ind,u) (nparrec+1)) ( + mkArrow + (mkApp(Lazy.force eq,[|(Lazy.force bb);mkApp(eqI,[|mkVar n;mkVar m|]);(Lazy.force tt)|])) + (mkApp(Lazy.force eq,[|mkFullInd (ind,u) (nparrec+3);mkVar n;mkVar m|])) + ))), eff + +let compute_bl_tact mode bl_scheme_key ind lnamesparrec nparrec = + let list_id = list_id lnamesparrec in + let avoid = ref [] in + let first_intros = + ( List.map (fun (s,_,_,_) -> s ) list_id ) @ + ( List.map (fun (_,seq,_,_ ) -> seq) list_id ) @ + ( List.map (fun (_,_,sbl,_ ) -> sbl) list_id ) + in + let fresh_id s gl = + let fresh = fresh_id_in_env (Id.Set.of_list !avoid) s (Proofview.Goal.env gl) in + avoid := fresh::(!avoid); fresh + in + Proofview.Goal.enter begin fun gl -> + let fresh_first_intros = List.map (fun id -> fresh_id id gl) first_intros in + let freshn = fresh_id (Id.of_string "x") gl in + let freshm = fresh_id (Id.of_string "y") gl in + let freshz = fresh_id (Id.of_string "Z") gl in + (* try with *) + Tacticals.New.tclTHENLIST [ intros_using fresh_first_intros; + intro_using freshn ; + induct_on (EConstr.mkVar freshn); + intro_using freshm; + destruct_on (EConstr.mkVar freshm); + intro_using freshz; + intros; + Tacticals.New.tclTRY ( + Tacticals.New.tclORELSE reflexivity my_discr_tac + ); + simpl_in_hyp (freshz,Locus.InHyp); +(* +repeat ( apply andb_prop in z;let z1:= fresh "Z" in destruct z as [z1 z]). +*) + Tacticals.New.tclREPEAT ( + Tacticals.New.tclTHENLIST [ + Simple.apply_in freshz (EConstr.of_constr (andb_prop())); + Proofview.Goal.enter begin fun gl -> + let fresht = fresh_id (Id.of_string "Z") gl in + destruct_on_as (EConstr.mkVar freshz) + (IntroOrPattern [[CAst.make @@ IntroNaming (IntroIdentifier fresht); + CAst.make @@ IntroNaming (IntroIdentifier freshz)]]) + end + ]); +(* + Ci a1 ... an = Ci b1 ... bn + replace bi with ai; auto || replace bi with ai by apply typeofbi_prod ; auto +*) + Proofview.Goal.enter begin fun gl -> + let concl = Proofview.Goal.concl gl in + let sigma = Tacmach.New.project gl in + match EConstr.kind sigma concl with + | App (c,ca) -> ( + match EConstr.kind sigma c with + | Ind (indeq, u) -> + if eq_gr (IndRef indeq) Coqlib.glob_eq + then + Tacticals.New.tclTHEN + (do_replace_bl mode bl_scheme_key ind + (!avoid) + nparrec (ca.(2)) + (ca.(1))) + Auto.default_auto + else + Tacticals.New.tclZEROMSG (str "Failure while solving Boolean->Leibniz.") + | _ -> Tacticals.New.tclZEROMSG (str" Failure while solving Boolean->Leibniz.") + ) + | _ -> Tacticals.New.tclZEROMSG (str "Failure while solving Boolean->Leibniz.") + end + + ] + end + +let bl_scheme_kind_aux = ref (fun _ -> failwith "Undefined") + +let side_effect_of_mode = function + | Declare.UserAutomaticRequest -> false + | Declare.InternalTacticRequest -> true + | Declare.UserIndividualRequest -> false + +let make_bl_scheme mode mind = + let mib = Global.lookup_mind mind in + if not (Int.equal (Array.length mib.mind_packets) 1) then + user_err + (str "Automatic building of boolean->Leibniz lemmas not supported"); + let ind = (mind,0) in + let nparams = mib.mind_nparams in + let nparrec = mib.mind_nparams_rec in + let lnonparrec,lnamesparrec = (* TODO subst *) + context_chop (nparams-nparrec) mib.mind_params_ctxt in + let bl_goal, eff = compute_bl_goal ind lnamesparrec nparrec in + let ctx = UState.make (Global.universes ()) in + let side_eff = side_effect_of_mode mode in + let bl_goal = EConstr.of_constr bl_goal in + let (ans, _, ctx) = Pfedit.build_by_tactic ~side_eff (Global.env()) ctx bl_goal + (compute_bl_tact mode (!bl_scheme_kind_aux()) (ind, EConstr.EInstance.empty) lnamesparrec nparrec) + in + ([|ans|], ctx), eff + +let bl_scheme_kind = declare_mutual_scheme_object "_dec_bl" make_bl_scheme + +let _ = bl_scheme_kind_aux := fun () -> bl_scheme_kind + +(**********************************************************************) +(* Leibniz->Boolean *) + +let compute_lb_goal ind lnamesparrec nparrec = + let list_id = list_id lnamesparrec in + let eq = Lazy.force eq and tt = Lazy.force tt and bb = Lazy.force bb in + let avoid = List.fold_right (Nameops.Name.fold_right (fun id l -> Id.Set.add id l)) (List.map RelDecl.get_name lnamesparrec) Id.Set.empty in + let eqI, eff = eqI ind lnamesparrec in + let create_input c = + let x = next_ident_away (Id.of_string "x") avoid and + y = next_ident_away (Id.of_string "y") avoid in + let lb_typ = List.map (fun (s,seq,_,_) -> + mkNamedProd x (mkVar s) ( + mkNamedProd y (mkVar s) ( + mkArrow + ( mkApp(eq,[|mkVar s;mkVar x;mkVar y|])) + ( mkApp(eq,[|bb;mkApp(mkVar seq,[|mkVar x;mkVar y|]);tt|])) + )) + ) list_id in + let lb_input = List.fold_left2 ( fun a (s,_,_,slb) b -> + mkNamedProd slb b a + ) c (List.rev list_id) (List.rev lb_typ) in + let eqs_typ = List.map (fun (s,_,_,_) -> + mkProd(Anonymous,mkVar s,mkProd(Anonymous,mkVar s,bb)) + ) list_id in + let eq_input = List.fold_left2 ( fun a (s,seq,_,_) b -> + mkNamedProd seq b a + ) lb_input (List.rev list_id) (List.rev eqs_typ) in + List.fold_left (fun a decl -> mkNamedProd + (match (RelDecl.get_name decl) with Name s -> s | Anonymous -> Id.of_string "A") + (RelDecl.get_type decl) a) eq_input lnamesparrec + in + let n = next_ident_away (Id.of_string "x") avoid and + m = next_ident_away (Id.of_string "y") avoid in + let u = Univ.Instance.empty in + create_input ( + mkNamedProd n (mkFullInd (ind,u) nparrec) ( + mkNamedProd m (mkFullInd (ind,u) (nparrec+1)) ( + mkArrow + (mkApp(eq,[|mkFullInd (ind,u) (nparrec+2);mkVar n;mkVar m|])) + (mkApp(eq,[|bb;mkApp(eqI,[|mkVar n;mkVar m|]);tt|])) + ))), eff + +let compute_lb_tact mode lb_scheme_key ind lnamesparrec nparrec = + let list_id = list_id lnamesparrec in + let avoid = ref [] in + let first_intros = + ( List.map (fun (s,_,_,_) -> s ) list_id ) @ + ( List.map (fun (_,seq,_,_) -> seq) list_id ) @ + ( List.map (fun (_,_,_,slb) -> slb) list_id ) + in + let fresh_id s gl = + let fresh = fresh_id_in_env (Id.Set.of_list !avoid) s (Proofview.Goal.env gl) in + avoid := fresh::(!avoid); fresh + in + Proofview.Goal.enter begin fun gl -> + let fresh_first_intros = List.map (fun id -> fresh_id id gl) first_intros in + let freshn = fresh_id (Id.of_string "x") gl in + let freshm = fresh_id (Id.of_string "y") gl in + let freshz = fresh_id (Id.of_string "Z") gl in + (* try with *) + Tacticals.New.tclTHENLIST [ intros_using fresh_first_intros; + intro_using freshn ; + induct_on (EConstr.mkVar freshn); + intro_using freshm; + destruct_on (EConstr.mkVar freshm); + intro_using freshz; + intros; + Tacticals.New.tclTRY ( + Tacticals.New.tclORELSE reflexivity my_discr_tac + ); + my_inj_tac freshz; + intros; simpl_in_concl; + Auto.default_auto; + Tacticals.New.tclREPEAT ( + Tacticals.New.tclTHENLIST [apply (EConstr.of_constr (andb_true_intro())); + simplest_split ;Auto.default_auto ] + ); + Proofview.Goal.enter begin fun gls -> + let concl = Proofview.Goal.concl gls in + let sigma = Tacmach.New.project gl in + (* assume the goal to be eq (eq_type ...) = true *) + match EConstr.kind sigma concl with + | App(c,ca) -> (match (EConstr.kind sigma ca.(1)) with + | App(c',ca') -> + let n = Array.length ca' in + do_replace_lb mode lb_scheme_key + (!avoid) + nparrec + ca'.(n-2) ca'.(n-1) + | _ -> + Tacticals.New.tclZEROMSG (str "Failure while solving Leibniz->Boolean.") + ) + | _ -> + Tacticals.New.tclZEROMSG (str "Failure while solving Leibniz->Boolean.") + end + ] + end + +let lb_scheme_kind_aux = ref (fun () -> failwith "Undefined") + +let make_lb_scheme mode mind = + let mib = Global.lookup_mind mind in + if not (Int.equal (Array.length mib.mind_packets) 1) then + user_err + (str "Automatic building of Leibniz->boolean lemmas not supported"); + let ind = (mind,0) in + let nparams = mib.mind_nparams in + let nparrec = mib.mind_nparams_rec in + let lnonparrec,lnamesparrec = + context_chop (nparams-nparrec) mib.mind_params_ctxt in + let lb_goal, eff = compute_lb_goal ind lnamesparrec nparrec in + let ctx = UState.make (Global.universes ()) in + let side_eff = side_effect_of_mode mode in + let lb_goal = EConstr.of_constr lb_goal in + let (ans, _, ctx) = Pfedit.build_by_tactic ~side_eff (Global.env()) ctx lb_goal + (compute_lb_tact mode (!lb_scheme_kind_aux()) ind lnamesparrec nparrec) + in + ([|ans|], ctx), eff + +let lb_scheme_kind = declare_mutual_scheme_object "_dec_lb" make_lb_scheme + +let _ = lb_scheme_kind_aux := fun () -> lb_scheme_kind + +(**********************************************************************) +(* Decidable equality *) + +let check_not_is_defined () = + try ignore (Coqlib.build_coq_not ()) + with e when CErrors.noncritical e -> raise (UndefinedCst "not") + +(* {n=m}+{n<>m} part *) +let compute_dec_goal ind lnamesparrec nparrec = + check_not_is_defined (); + let eq = Lazy.force eq and tt = Lazy.force tt and bb = Lazy.force bb in + let list_id = list_id lnamesparrec in + let avoid = List.fold_right (Nameops.Name.fold_right (fun id l -> Id.Set.add id l)) (List.map RelDecl.get_name lnamesparrec) Id.Set.empty in + let create_input c = + let x = next_ident_away (Id.of_string "x") avoid and + y = next_ident_away (Id.of_string "y") avoid in + let lb_typ = List.map (fun (s,seq,_,_) -> + mkNamedProd x (mkVar s) ( + mkNamedProd y (mkVar s) ( + mkArrow + ( mkApp(eq,[|mkVar s;mkVar x;mkVar y|])) + ( mkApp(eq,[|bb;mkApp(mkVar seq,[|mkVar x;mkVar y|]);tt|])) + )) + ) list_id in + let bl_typ = List.map (fun (s,seq,_,_) -> + mkNamedProd x (mkVar s) ( + mkNamedProd y (mkVar s) ( + mkArrow + ( mkApp(eq,[|bb;mkApp(mkVar seq,[|mkVar x;mkVar y|]);tt|])) + ( mkApp(eq,[|mkVar s;mkVar x;mkVar y|])) + )) + ) list_id in + + let lb_input = List.fold_left2 ( fun a (s,_,_,slb) b -> + mkNamedProd slb b a + ) c (List.rev list_id) (List.rev lb_typ) in + let bl_input = List.fold_left2 ( fun a (s,_,sbl,_) b -> + mkNamedProd sbl b a + ) lb_input (List.rev list_id) (List.rev bl_typ) in + + let eqs_typ = List.map (fun (s,_,_,_) -> + mkProd(Anonymous,mkVar s,mkProd(Anonymous,mkVar s,bb)) + ) list_id in + let eq_input = List.fold_left2 ( fun a (s,seq,_,_) b -> + mkNamedProd seq b a + ) bl_input (List.rev list_id) (List.rev eqs_typ) in + List.fold_left (fun a decl -> mkNamedProd + (match RelDecl.get_name decl with Name s -> s | Anonymous -> Id.of_string "A") + (RelDecl.get_type decl) a) eq_input lnamesparrec + in + let n = next_ident_away (Id.of_string "x") avoid and + m = next_ident_away (Id.of_string "y") avoid in + let eqnm = mkApp(eq,[|mkFullInd ind (2*nparrec+2);mkVar n;mkVar m|]) in + create_input ( + mkNamedProd n (mkFullInd ind (2*nparrec)) ( + mkNamedProd m (mkFullInd ind (2*nparrec+1)) ( + mkApp(sumbool(),[|eqnm;mkApp (Universes.constr_of_global @@ Coqlib.build_coq_not(),[|eqnm|])|]) + ) + ) + ) + +let compute_dec_tact ind lnamesparrec nparrec = + let eq = Lazy.force eq and tt = Lazy.force tt + and ff = Lazy.force ff and bb = Lazy.force bb in + let list_id = list_id lnamesparrec in + let eqI, eff = eqI ind lnamesparrec in + let avoid = ref [] in + let eqtrue x = mkApp(eq,[|bb;x;tt|]) in + let eqfalse x = mkApp(eq,[|bb;x;ff|]) in + let first_intros = + ( List.map (fun (s,_,_,_) -> s ) list_id ) @ + ( List.map (fun (_,seq,_,_) -> seq) list_id ) @ + ( List.map (fun (_,_,sbl,_) -> sbl) list_id ) @ + ( List.map (fun (_,_,_,slb) -> slb) list_id ) + in + let fresh_id s gl = + let fresh = fresh_id_in_env (Id.Set.of_list !avoid) s (Proofview.Goal.env gl) in + avoid := fresh::(!avoid); fresh + in + Proofview.Goal.enter begin fun gl -> + let fresh_first_intros = List.map (fun id -> fresh_id id gl) first_intros in + let freshn = fresh_id (Id.of_string "x") gl in + let freshm = fresh_id (Id.of_string "y") gl in + let freshH = fresh_id (Id.of_string "H") gl in + let eqbnm = mkApp(eqI,[|mkVar freshn;mkVar freshm|]) in + let arfresh = Array.of_list fresh_first_intros in + let xargs = Array.sub arfresh 0 (2*nparrec) in + begin try + let c, eff = find_scheme bl_scheme_kind ind in + Proofview.tclUNIT (mkConst c,eff) with + Not_found -> + Tacticals.New.tclZEROMSG (str "Error during the decidability part, boolean to leibniz equality is required.") + end >>= fun (blI,eff') -> + begin try + let c, eff = find_scheme lb_scheme_kind ind in + Proofview.tclUNIT (mkConst c,eff) with + Not_found -> + Tacticals.New.tclZEROMSG (str "Error during the decidability part, leibniz to boolean equality is required.") + end >>= fun (lbI,eff'') -> + let eff = (Safe_typing.concat_private eff'' (Safe_typing.concat_private eff' eff)) in + Tacticals.New.tclTHENLIST [ + Proofview.tclEFFECTS eff; + intros_using fresh_first_intros; + intros_using [freshn;freshm]; + (*we do this so we don't have to prove the same goal twice *) + assert_by (Name freshH) (EConstr.of_constr ( + mkApp(sumbool(),[|eqtrue eqbnm; eqfalse eqbnm|]) + )) + (Tacticals.New.tclTHEN (destruct_on (EConstr.of_constr eqbnm)) Auto.default_auto); + + Proofview.Goal.enter begin fun gl -> + let freshH2 = fresh_id (Id.of_string "H") gl in + Tacticals.New.tclTHENS (destruct_on_using (EConstr.mkVar freshH) freshH2) [ + (* left *) + Tacticals.New.tclTHENLIST [ + simplest_left; + apply (EConstr.of_constr (mkApp(blI,Array.map(fun x->mkVar x) xargs))); + Auto.default_auto + ] + ; + + (*right *) + Proofview.Goal.enter begin fun gl -> + let freshH3 = fresh_id (Id.of_string "H") gl in + Tacticals.New.tclTHENLIST [ + simplest_right ; + unfold_constr (Lazy.force Coqlib.coq_not_ref); + intro; + Equality.subst_all (); + assert_by (Name freshH3) + (EConstr.of_constr (mkApp(eq,[|bb;mkApp(eqI,[|mkVar freshm;mkVar freshm|]);tt|]))) + (Tacticals.New.tclTHENLIST [ + apply (EConstr.of_constr (mkApp(lbI,Array.map (fun x->mkVar x) xargs))); + Auto.default_auto + ]); + Equality.general_rewrite_bindings_in true + Locus.AllOccurrences true false + (List.hd !avoid) + ((EConstr.mkVar (List.hd (List.tl !avoid))), + NoBindings + ) + true; + my_discr_tac + ] + end + ] + end + ] + end + +let make_eq_decidability mode mind = + let mib = Global.lookup_mind mind in + if not (Int.equal (Array.length mib.mind_packets) 1) then + raise DecidabilityMutualNotSupported; + let ind = (mind,0) in + let nparams = mib.mind_nparams in + let nparrec = mib.mind_nparams_rec in + let u = Univ.Instance.empty in + let ctx = UState.make (Global.universes ()) in + let lnonparrec,lnamesparrec = + context_chop (nparams-nparrec) mib.mind_params_ctxt in + let side_eff = side_effect_of_mode mode in + let (ans, _, ctx) = Pfedit.build_by_tactic ~side_eff (Global.env()) ctx + (EConstr.of_constr (compute_dec_goal (ind,u) lnamesparrec nparrec)) + (compute_dec_tact ind lnamesparrec nparrec) + in + ([|ans|], ctx), Safe_typing.empty_private_constants + +let eq_dec_scheme_kind = + declare_mutual_scheme_object "_eq_dec" make_eq_decidability + +(* The eq_dec_scheme proofs depend on the equality and discr tactics + but the inj tactics, that comes with discr, depends on the + eq_dec_scheme... *) + +let _ = Equality.set_eq_dec_scheme_kind eq_dec_scheme_kind diff --git a/vernac/auto_ind_decl.mli b/vernac/auto_ind_decl.mli new file mode 100644 index 00000000..5cc783df --- /dev/null +++ b/vernac/auto_ind_decl.mli @@ -0,0 +1,44 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* + (Printer.pr_global g ++ str" is already a coercion") + | NotAFunction -> + (Printer.pr_global g ++ str" is not a function") + | NoSource (Some cl) -> + (str "Cannot recognize " ++ pr_class cl ++ str " as a source class of " + ++ Printer.pr_global g) + | NoSource None -> + (str ": cannot find the source class of " ++ Printer.pr_global g) + | ForbiddenSourceClass cl -> + pr_class cl ++ str " cannot be a source class" + | NoTarget -> + (str"Cannot find the target class") + | WrongTarget (clt,cl) -> + (str"Found target class " ++ pr_class cl ++ + str " instead of " ++ pr_class clt) + | NotAClass ref -> + (str "Type of " ++ Printer.pr_global ref ++ + str " does not end with a sort") + +(* Verifications pour l'ajout d'une classe *) + +let check_reference_arity ref = + let env = Global.env () in + let c, _ = Global.type_of_global_in_context env ref in + if not (Reductionops.is_arity env Evd.empty (EConstr.of_constr c)) (** FIXME *) then + raise (CoercionError (NotAClass ref)) + +let check_arity = function + | CL_FUN | CL_SORT -> () + | CL_CONST cst -> check_reference_arity (ConstRef cst) + | CL_PROJ cst -> check_reference_arity (ConstRef cst) + | CL_SECVAR id -> check_reference_arity (VarRef id) + | CL_IND kn -> check_reference_arity (IndRef kn) + +(* Coercions *) + +(* check that the computed target is the provided one *) +let check_target clt = function + | Some cl when not (cl_typ_eq cl clt) -> raise (CoercionError (WrongTarget(clt,cl))) + | _ -> () + +(* condition d'heritage uniforme *) + +let uniform_cond sigma ctx lt = + List.for_all2eq (EConstr.eq_constr sigma) + lt (Context.Rel.to_extended_list EConstr.mkRel 0 ctx) + +let class_of_global = function + | ConstRef sp -> + if Environ.is_projection sp (Global.env ()) + then CL_PROJ sp else CL_CONST sp + | IndRef sp -> CL_IND sp + | VarRef id -> CL_SECVAR id + | ConstructRef _ as c -> + user_err ~hdr:"class_of_global" + (str "Constructors, such as " ++ Printer.pr_global c ++ + str ", cannot be used as a class.") + +(* +lp est la liste (inverse'e) des arguments de la coercion +ids est le nom de la classe source +sps_opt est le sp de la classe source dans le cas des structures +retourne: +la classe source +nbre d'arguments de la classe +le constr de la class +la liste des variables dont depend la classe source +l'indice de la classe source dans la liste lp +*) + +let get_source lp source = + let open Context.Rel.Declaration in + match source with + | None -> + (* Take the latest non let-in argument *) + let rec aux = function + | [] -> raise Not_found + | LocalDef _ :: lt -> aux lt + | LocalAssum (_,t1) :: lt -> + let cl1,u1,lv1 = find_class_type Evd.empty (EConstr.of_constr t1) in + cl1,lt,lv1,1 + in aux lp + | Some cl -> + (* Take the first argument that matches *) + let rec aux acc = function + | [] -> raise Not_found + | LocalDef _ as decl :: lt -> aux (decl::acc) lt + | LocalAssum (_,t1) as decl :: lt -> + try + let cl1,u1,lv1 = find_class_type Evd.empty (EConstr.of_constr t1) in + if cl_typ_eq cl cl1 then cl1,acc,lv1,Context.Rel.nhyps lt+1 + else raise Not_found + with Not_found -> aux (decl::acc) lt + in aux [] (List.rev lp) + +let get_target t ind = + if (ind > 1) then + CL_FUN + else + match pi1 (find_class_type Evd.empty (EConstr.of_constr t)) with + | CL_CONST p when Environ.is_projection p (Global.env ()) -> + CL_PROJ p + | x -> x + +let strength_of_cl = function + | CL_CONST kn -> `GLOBAL + | CL_SECVAR id -> `LOCAL + | _ -> `GLOBAL + +let strength_of_global = function + | VarRef _ -> `LOCAL + | _ -> `GLOBAL + +let get_strength stre ref cls clt = + let stres = strength_of_cl cls in + let stret = strength_of_cl clt in + let stref = strength_of_global ref in + strength_min [stre;stres;stret;stref] + +let ident_key_of_class = function + | CL_FUN -> "Funclass" + | CL_SORT -> "Sortclass" + | CL_CONST sp | CL_PROJ sp -> Label.to_string (Constant.label sp) + | CL_IND (sp,_) -> Label.to_string (MutInd.label sp) + | CL_SECVAR id -> Id.to_string id + +(* Identity coercion *) + +let error_not_transparent source = + user_err ~hdr:"build_id_coercion" + (pr_class source ++ str " must be a transparent constant.") + +let build_id_coercion idf_opt source poly = + let env = Global.env () in + let sigma = Evd.from_env env in + let sigma, vs = match source with + | CL_CONST sp -> Evd.fresh_global env sigma (ConstRef sp) + | _ -> error_not_transparent source in + let c = match constant_opt_value_in env (destConst vs) with + | Some c -> c + | None -> error_not_transparent source in + let lams,t = decompose_lam_assum c in + let val_f = + it_mkLambda_or_LetIn + (mkLambda (Name Namegen.default_dependent_ident, + applistc vs (Context.Rel.to_extended_list mkRel 0 lams), + mkRel 1)) + lams + in + let typ_f = + List.fold_left (fun d c -> Term.mkProd_wo_LetIn c d) + (mkProd (Anonymous, applistc vs (Context.Rel.to_extended_list mkRel 0 lams), lift 1 t)) + lams + in + (* juste pour verification *) + let _ = + if not + (Reductionops.is_conv_leq env sigma + (Typing.unsafe_type_of env sigma (EConstr.of_constr val_f)) (EConstr.of_constr typ_f)) + then + user_err (strbrk + "Cannot be defined as coercion (maybe a bad number of arguments).") + in + let idf = + match idf_opt with + | Some idf -> idf + | None -> + let cl,u,_ = find_class_type sigma (EConstr.of_constr t) in + Id.of_string ("Id_"^(ident_key_of_class source)^"_"^ + (ident_key_of_class cl)) + in + let univs = Evd.const_univ_entry ~poly sigma in + let constr_entry = (* Cast is necessary to express [val_f] is identity *) + DefinitionEntry + (definition_entry ~types:typ_f ~univs + ~inline:true (mkCast (val_f, DEFAULTcast, typ_f))) + in + let decl = (constr_entry, IsDefinition IdentityCoercion) in + let kn = declare_constant idf decl in + ConstRef kn + +let check_source = function +| Some (CL_FUN as s) -> raise (CoercionError (ForbiddenSourceClass s)) +| _ -> () + +(* +nom de la fonction coercion +strength de f +nom de la classe source (optionnel) +sp de la classe source (dans le cas des structures) +nom de la classe target (optionnel) +booleen "coercion identite'?" + +lorque source est None alors target est None aussi. +*) + +let warn_uniform_inheritance = + CWarnings.create ~name:"uniform-inheritance" ~category:"typechecker" + (fun g -> + Printer.pr_global g ++ + strbrk" does not respect the uniform inheritance condition") + +let add_new_coercion_core coef stre poly source target isid = + check_source source; + let t, _ = Global.type_of_global_in_context (Global.env ()) coef in + if coercion_exists coef then raise (CoercionError AlreadyExists); + let lp,tg = decompose_prod_assum t in + let llp = List.length lp in + if Int.equal llp 0 then raise (CoercionError NotAFunction); + let (cls,ctx,lvs,ind) = + try + get_source lp source + with Not_found -> + raise (CoercionError (NoSource source)) + in + check_source (Some cls); + if not (uniform_cond Evd.empty (** FIXME - for when possibly called with unresolved evars in the future *) + ctx lvs) then + warn_uniform_inheritance coef; + let clt = + try + get_target tg ind + with Not_found -> + raise (CoercionError NoTarget) + in + check_target clt target; + check_arity cls; + check_arity clt; + let local = match get_strength stre coef cls clt with + | `LOCAL -> true + | `GLOBAL -> false + in + declare_coercion coef ~local ~isid ~src:cls ~target:clt ~params:(List.length lvs) + + +let try_add_new_coercion_core ref ~local c d e f = + try add_new_coercion_core ref (loc_of_bool local) c d e f + with CoercionError e -> + user_err ~hdr:"try_add_new_coercion_core" + (explain_coercion_error ref e ++ str ".") + +let try_add_new_coercion ref ~local poly = + try_add_new_coercion_core ref ~local poly None None false + +let try_add_new_coercion_subclass cl ~local poly = + let coe_ref = build_id_coercion None cl poly in + try_add_new_coercion_core coe_ref ~local poly (Some cl) None true + +let try_add_new_coercion_with_target ref ~local poly ~source ~target = + try_add_new_coercion_core ref ~local poly (Some source) (Some target) false + +let try_add_new_identity_coercion id ~local poly ~source ~target = + let ref = build_id_coercion (Some id) source poly in + try_add_new_coercion_core ref ~local poly (Some source) (Some target) true + +let try_add_new_coercion_with_source ref ~local poly ~source = + try_add_new_coercion_core ref ~local poly (Some source) None false + +let add_coercion_hook poly local ref = + let stre = match local with + | Local -> true + | Global -> false + | Discharge -> assert false + in + let () = try_add_new_coercion ref ~local:stre poly in + let msg = pr_global_env Id.Set.empty ref ++ str " is now a coercion" in + Flags.if_verbose Feedback.msg_info msg + +let add_coercion_hook poly = Lemmas.mk_hook (add_coercion_hook poly) + +let add_subclass_hook poly local ref = + let stre = match local with + | Local -> true + | Global -> false + | Discharge -> assert false + in + let cl = class_of_global ref in + try_add_new_coercion_subclass cl ~local:stre poly + +let add_subclass_hook poly = Lemmas.mk_hook (add_subclass_hook poly) diff --git a/vernac/class.mli b/vernac/class.mli new file mode 100644 index 00000000..33d31fe1 --- /dev/null +++ b/vernac/class.mli @@ -0,0 +1,50 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* local:bool -> + Decl_kinds.polymorphic -> + source:cl_typ -> target:cl_typ -> unit + +(** [try_add_new_coercion ref s] declares [ref], assumed to be of type + [(x1:T1)...(xn:Tn)src->tg], as a coercion from [src] to [tg] *) +val try_add_new_coercion : global_reference -> local:bool -> + Decl_kinds.polymorphic -> unit + +(** [try_add_new_coercion_subclass cst s] expects that [cst] denotes a + transparent constant which unfolds to some class [tg]; it declares + an identity coercion from [cst] to [tg], named something like + ["Id_cst_tg"] *) +val try_add_new_coercion_subclass : cl_typ -> local:bool -> + Decl_kinds.polymorphic -> unit + +(** [try_add_new_coercion_with_source ref s src] declares [ref] as a coercion + from [src] to [tg] where the target is inferred from the type of [ref] *) +val try_add_new_coercion_with_source : global_reference -> local:bool -> + Decl_kinds.polymorphic -> source:cl_typ -> unit + +(** [try_add_new_identity_coercion id s src tg] enriches the + environment with a new definition of name [id] declared as an + identity coercion from [src] to [tg] *) +val try_add_new_identity_coercion : Id.t -> local:bool -> + Decl_kinds.polymorphic -> source:cl_typ -> target:cl_typ -> unit + +val add_coercion_hook : Decl_kinds.polymorphic -> unit Lemmas.declaration_hook + +val add_subclass_hook : Decl_kinds.polymorphic -> unit Lemmas.declaration_hook + +val class_of_global : global_reference -> cl_typ diff --git a/vernac/classes.ml b/vernac/classes.ml new file mode 100644 index 00000000..76d427ad --- /dev/null +++ b/vernac/classes.ml @@ -0,0 +1,438 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* !refine_instance); + Goptions.optwrite = (fun b -> refine_instance := b) +} + +let typeclasses_db = "typeclass_instances" + +let set_typeclass_transparency c local b = + Hints.add_hints local [typeclasses_db] + (Hints.HintsTransparencyEntry ([c], b)) + +let _ = + Hook.set Typeclasses.add_instance_hint_hook + (fun inst path local info poly -> + let inst' = match inst with IsConstr c -> Hints.IsConstr (EConstr.of_constr c, Univ.ContextSet.empty) + | IsGlobal gr -> Hints.IsGlobRef gr + in + let info = + let open Vernacexpr in + { info with hint_pattern = + Option.map + (Constrintern.intern_constr_pattern (Global.env()) Evd.(from_env Global.(env()))) + info.hint_pattern } in + Flags.silently (fun () -> + Hints.add_hints local [typeclasses_db] + (Hints.HintsResolveEntry + [info, poly, false, Hints.PathHints path, inst'])) ()); + Hook.set Typeclasses.set_typeclass_transparency_hook set_typeclass_transparency; + Hook.set Typeclasses.classes_transparent_state_hook + (fun () -> Hints.Hint_db.transparent_state (Hints.searchtable_map typeclasses_db)) + +(** TODO: add subinstances *) +let existing_instance glob g info = + let c = global g in + let info = Option.default Hints.empty_hint_info info in + let instance, _ = Global.type_of_global_in_context (Global.env ()) c in + let _, r = Term.decompose_prod_assum instance in + match class_of_constr Evd.empty (EConstr.of_constr r) with + | Some (_, ((tc,u), _)) -> add_instance (new_instance tc info glob c) + | None -> user_err ?loc:g.CAst.loc + ~hdr:"declare_instance" + (Pp.str "Constant does not build instances of a declared type class.") + +let mismatched_params env n m = mismatched_ctx_inst env Parameters n m +let mismatched_props env n m = mismatched_ctx_inst env Properties n m + +(* Declare everything in the parameters as implicit, and the class instance as well *) + +let type_ctx_instance env sigma ctx inst subst = + let open Vars in + let rec aux (sigma, subst, instctx) l = function + decl :: ctx -> + let t' = substl subst (RelDecl.get_type decl) in + let (sigma, c'), l = + match decl with + | LocalAssum _ -> interp_casted_constr_evars env sigma (List.hd l) t', List.tl l + | LocalDef (_,b,_) -> (sigma, substl subst b), l + in + let d = RelDecl.get_name decl, Some c', t' in + aux (sigma, c' :: subst, d :: instctx) l ctx + | [] -> sigma, subst + in aux (sigma, subst, []) inst (List.rev ctx) + +let id_of_class cl = + match cl.cl_impl with + | ConstRef kn -> let _,_,l = Constant.repr3 kn in Label.to_id l + | IndRef (kn,i) -> + let mip = (Environ.lookup_mind kn (Global.env ())).Declarations.mind_packets in + mip.(0).Declarations.mind_typename + | _ -> assert false + +open Pp + +let instance_hook k info global imps ?hook cst = + Impargs.maybe_declare_manual_implicits false cst ~enriching:false imps; + Typeclasses.declare_instance (Some info) (not global) cst; + (match hook with Some h -> h cst | None -> ()) + +let declare_instance_constant k info global imps ?hook id decl poly sigma term termtype = + let kind = IsDefinition Instance in + let sigma = + let env = Global.env () in + let levels = Univ.LSet.union (Univops.universes_of_constr env termtype) + (Univops.universes_of_constr env term) in + Evd.restrict_universe_context sigma levels + in + let uctx = Evd.check_univ_decl ~poly sigma decl in + let entry = + Declare.definition_entry ~types:termtype ~univs:uctx term + in + let cdecl = (DefinitionEntry entry, kind) in + let kn = Declare.declare_constant id cdecl in + Declare.definition_message id; + Declare.declare_univ_binders (ConstRef kn) (Evd.universe_binders sigma); + instance_hook k info global imps ?hook (ConstRef kn); + id + +let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) + ~program_mode poly ctx (instid, bk, cl) props ?(generalize=true) + ?(tac:unit Proofview.tactic option) ?hook pri = + let env = Global.env() in + let ({CAst.loc;v=instid}, pl) = instid in + let sigma, decl = Univdecls.interp_univ_decl_opt env pl in + let tclass, ids = + match bk with + | Decl_kinds.Implicit -> + Implicit_quantifiers.implicit_application Id.Set.empty ~allow_partial:false + (fun avoid (clname, _) -> + match clname with + | Some cl -> + let t = CAst.make @@ CHole (None, Misctypes.IntroAnonymous, None) in + t, avoid + | None -> failwith ("new instance: under-applied typeclass")) + cl + | Explicit -> cl, Id.Set.empty + in + let tclass = + if generalize then CAst.make @@ CGeneralization (Implicit, Some AbsPi, tclass) + else tclass + in + let sigma, k, u, cty, ctx', ctx, len, imps, subst = + let sigma, (impls, ((env', ctx), imps)) = interp_context_evars env sigma ctx in + let sigma, (c', imps') = interp_type_evars_impls ~impls env' sigma tclass in + let len = List.length ctx in + let imps = imps @ Impargs.lift_implicits len imps' in + let ctx', c = decompose_prod_assum sigma c' in + let ctx'' = ctx' @ ctx in + let (k, u), args = Typeclasses.dest_class_app (push_rel_context ctx'' env) sigma c in + let u_s = EInstance.kind sigma u in + let cl = Typeclasses.typeclass_univ_instance (k, u_s) in + let args = List.map of_constr args in + let cl_context = List.map (Termops.map_rel_decl of_constr) (snd cl.cl_context) in + let _, args = + List.fold_right (fun decl (args, args') -> + match decl with + | LocalAssum _ -> (List.tl args, List.hd args :: args') + | LocalDef (_,b,_) -> (args, Vars.substl args' b :: args')) + cl_context (args, []) + in + sigma, cl, u, c', ctx', ctx, len, imps, args + in + let id = + match instid with + Name id -> + let sp = Lib.make_path id in + if Nametab.exists_cci sp then + user_err ~hdr:"new_instance" (Id.print id ++ Pp.str " already exists."); + id + | Anonymous -> + let i = Nameops.add_suffix (id_of_class k) "_instance_0" in + Namegen.next_global_ident_away i (Termops.vars_of_env env) + in + let env' = push_rel_context ctx env in + let sigma = Evarutil.nf_evar_map sigma in + let sigma = resolve_typeclasses ~filter:Typeclasses.all_evars ~fail:true env sigma in + if abstract then + begin + let subst = List.fold_left2 + (fun subst' s decl -> if is_local_assum decl then s :: subst' else subst') + [] subst (snd k.cl_context) + in + let (_, ty_constr) = instance_constructor (k,u) subst in + let termtype = it_mkProd_or_LetIn ty_constr (ctx' @ ctx) in + let sigma,_ = Evarutil.nf_evars_and_universes sigma in + Pretyping.check_evars env Evd.empty sigma termtype; + let univs = Evd.check_univ_decl ~poly sigma decl in + let termtype = to_constr sigma termtype in + let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest id + (ParameterEntry + (None,(termtype,univs),None), Decl_kinds.IsAssumption Decl_kinds.Logical) + in + Declare.declare_univ_binders (ConstRef cst) (Evd.universe_binders sigma); + instance_hook k pri global imps ?hook (ConstRef cst); id + end + else ( + let props = + match props with + | Some (true, { CAst.v = CRecord fs }) -> + if List.length fs > List.length k.cl_props then + mismatched_props env' (List.map snd fs) k.cl_props; + Some (Inl fs) + | Some (_, t) -> Some (Inr t) + | None -> + if program_mode then Some (Inl []) + else None + in + let subst, sigma = + match props with + | None -> + (if List.is_empty k.cl_props then Some (Inl subst) else None), sigma + | Some (Inr term) -> + let sigma, c = interp_casted_constr_evars env' sigma term cty in + Some (Inr (c, subst)), sigma + | Some (Inl props) -> + let get_id = CAst.map (function + | Ident id' -> id' + | Qualid id' -> snd (repr_qualid id')) + in + let props, rest = + List.fold_left + (fun (props, rest) decl -> + if is_local_assum decl then + try + let is_id (id', _) = match RelDecl.get_name decl, get_id id' with + | Name id, {CAst.v=id'} -> Id.equal id id' + | Anonymous, _ -> false + in + let (loc_mid, c) = + List.find is_id rest + in + let rest' = + List.filter (fun v -> not (is_id v)) rest + in + let {CAst.loc;v=mid} = get_id loc_mid in + List.iter (fun (n, _, x) -> + if Name.equal n (Name mid) then + Option.iter (fun x -> Dumpglob.add_glob ?loc (ConstRef x)) x) + k.cl_projs; + c :: props, rest' + with Not_found -> + ((CAst.make @@ CHole (None(* Some Evar_kinds.GoalEvar *), Misctypes.IntroAnonymous, None)) :: props), rest + else props, rest) + ([], props) k.cl_props + in + match rest with + | (n, _) :: _ -> + unbound_method env' k.cl_impl (get_id n) + | _ -> + let kcl_props = List.map (Termops.map_rel_decl of_constr) k.cl_props in + let sigma, res = type_ctx_instance (push_rel_context ctx' env') sigma kcl_props props subst in + Some (Inl res), sigma + in + let term, termtype = + match subst with + | None -> let termtype = it_mkProd_or_LetIn cty ctx in + None, termtype + | Some (Inl subst) -> + let subst = List.fold_left2 + (fun subst' s decl -> if is_local_assum decl then s :: subst' else subst') + [] subst (k.cl_props @ snd k.cl_context) + in + let (app, ty_constr) = instance_constructor (k,u) subst in + let termtype = it_mkProd_or_LetIn ty_constr (ctx' @ ctx) in + let term = it_mkLambda_or_LetIn (Option.get app) (ctx' @ ctx) in + Some term, termtype + | Some (Inr (def, subst)) -> + let termtype = it_mkProd_or_LetIn cty ctx in + let term = it_mkLambda_or_LetIn def ctx in + Some term, termtype + in + let sigma = Evarutil.nf_evar_map sigma in + let sigma = Typeclasses.resolve_typeclasses ~filter:Typeclasses.no_goals_or_obligations ~fail:true env sigma in + (* Try resolving fields that are typeclasses automatically. *) + let sigma = Typeclasses.resolve_typeclasses ~filter:Typeclasses.all_evars ~fail:false env sigma in + let sigma = Evarutil.nf_evar_map_undefined sigma in + (* Beware of this step, it is required as to minimize universes. *) + let sigma, _nf = Evarutil.nf_evar_map_universes sigma in + (* Check that the type is free of evars now. *) + Pretyping.check_evars env Evd.empty sigma termtype; + let termtype = to_constr sigma termtype in + let term = Option.map (to_constr sigma) term in + if not (Evd.has_undefined sigma) && not (Option.is_empty term) then + declare_instance_constant k pri global imps ?hook id decl + poly sigma (Option.get term) termtype + else if program_mode || refine || Option.is_empty term then begin + let kind = Decl_kinds.Global, poly, Decl_kinds.DefinitionBody Decl_kinds.Instance in + if program_mode then + let hook vis gr _ = + let cst = match gr with ConstRef kn -> kn | _ -> assert false in + Impargs.declare_manual_implicits false gr ~enriching:false [imps]; + Typeclasses.declare_instance (Some pri) (not global) (ConstRef cst) + in + let obls, constr, typ = + match term with + | Some t -> + let obls, _, constr, typ = + Obligations.eterm_obligations env id sigma 0 t termtype + in obls, Some constr, typ + | None -> [||], None, termtype + in + let hook = Lemmas.mk_hook hook in + let ctx = Evd.evar_universe_context sigma in + ignore (Obligations.add_definition id ?term:constr + ~univdecl:decl typ ctx ~kind:(Global,poly,Instance) ~hook obls); + id + else + (Flags.silently + (fun () -> + (* spiwack: it is hard to reorder the actions to do + the pretyping after the proof has opened. As a + consequence, we use the low-level primitives to code + the refinement manually.*) + let gls = List.rev (Evd.future_goals sigma) in + let sigma = Evd.reset_future_goals sigma in + Lemmas.start_proof id ~pl:decl kind sigma (EConstr.of_constr termtype) + (Lemmas.mk_hook + (fun _ -> instance_hook k pri global imps ?hook)); + (* spiwack: I don't know what to do with the status here. *) + if not (Option.is_empty term) then + let init_refine = + Tacticals.New.tclTHENLIST [ + Refine.refine ~typecheck:false (fun sigma -> (sigma,EConstr.of_constr (Option.get term))); + Proofview.Unsafe.tclNEWGOALS (CList.map Proofview.with_empty_state gls); + Tactics.New.reduce_after_refine; + ] + in + ignore (Pfedit.by init_refine) + else if Flags.is_auto_intros () then + ignore (Pfedit.by (Tacticals.New.tclDO len Tactics.intro)); + (match tac with Some tac -> ignore (Pfedit.by tac) | None -> ())) (); + id) + end + else CErrors.user_err Pp.(str "Unsolved obligations remaining.")) + +let named_of_rel_context l = + let open Vars in + let acc, ctx = + List.fold_right + (fun decl (subst, ctx) -> + let id = match RelDecl.get_name decl with Anonymous -> invalid_arg "named_of_rel_context" | Name id -> id in + let d = match decl with + | LocalAssum (_,t) -> id, None, substl subst t + | LocalDef (_,b,t) -> id, Some (substl subst b), substl subst t in + (mkVar id :: subst, d :: ctx)) + l ([], []) + in ctx + +let context poly l = + let env = Global.env() in + let sigma = Evd.from_env env in + let sigma, (_, ((env', fullctx), impls)) = interp_context_evars env sigma l in + (* Note, we must use the normalized evar from now on! *) + let sigma,_ = Evarutil.nf_evars_and_universes sigma in + let ce t = Pretyping.check_evars env Evd.empty sigma t in + let () = List.iter (fun decl -> Context.Rel.Declaration.iter_constr ce decl) fullctx in + let ctx = + try named_of_rel_context fullctx + with e when CErrors.noncritical e -> + user_err Pp.(str "Anonymous variables not allowed in contexts.") + in + let univs = + let uctx = Evd.universe_context_set sigma in + match ctx with + | [] -> assert false + | [_] -> + if poly + then Polymorphic_const_entry (Univ.ContextSet.to_context uctx) + else Monomorphic_const_entry uctx + | _::_::_ -> + if Lib.sections_are_opened () + then + begin + Declare.declare_universe_context poly uctx; + if poly then Polymorphic_const_entry Univ.UContext.empty + else Monomorphic_const_entry Univ.ContextSet.empty + end + else if poly + then Polymorphic_const_entry (Univ.ContextSet.to_context uctx) + else + begin + Declare.declare_universe_context poly uctx; + Monomorphic_const_entry Univ.ContextSet.empty + end + in + let fn status (id, b, t) = + let b, t = Option.map (to_constr sigma) b, to_constr sigma t in + if Lib.is_modtype () && not (Lib.sections_are_opened ()) then + (* Declare the universe context once *) + let decl = match b with + | None -> + (ParameterEntry (None,(t,univs),None), IsAssumption Logical) + | Some b -> + let entry = Declare.definition_entry ~univs ~types:t b in + (DefinitionEntry entry, IsAssumption Logical) + in + let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest id decl in + match class_of_constr sigma (of_constr t) with + | Some (rels, ((tc,_), args) as _cl) -> + add_instance (Typeclasses.new_instance tc Hints.empty_hint_info false (ConstRef cst)); + status + (* declare_subclasses (ConstRef cst) cl *) + | None -> status + else + let test (x, _) = match x with + | ExplByPos (_, Some id') -> Id.equal id id' + | _ -> false + in + let impl = List.exists test impls in + let decl = (Discharge, poly, Definitional) in + let nstatus = match b with + | None -> + pi3 (ComAssumption.declare_assumption false decl (t, univs) Universes.empty_binders [] impl + Vernacexpr.NoInline (CAst.make id)) + | Some b -> + let decl = (Discharge, poly, Definition) in + let entry = Declare.definition_entry ~univs ~types:t b in + let hook = Lemmas.mk_hook (fun _ gr -> gr) in + let _ = DeclareDef.declare_definition id decl entry Universes.empty_binders [] hook in + Lib.sections_are_opened () || Lib.is_modtype_strict () + in + status && nstatus + in + List.fold_left fn true (List.rev ctx) diff --git a/vernac/classes.mli b/vernac/classes.mli new file mode 100644 index 00000000..0342c840 --- /dev/null +++ b/vernac/classes.mli @@ -0,0 +1,69 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* constr_expr list -> Context.Rel.t -> 'a + +val mismatched_props : env -> constr_expr list -> Context.Rel.t -> 'a + +(** Instance declaration *) + +val existing_instance : bool -> reference -> Vernacexpr.hint_info_expr option -> unit +(** globality, reference, optional priority and pattern information *) + +val declare_instance_constant : + typeclass -> + Vernacexpr.hint_info_expr -> (** priority *) + bool -> (** globality *) + Impargs.manual_explicitation list -> (** implicits *) + ?hook:(Globnames.global_reference -> unit) -> + Id.t -> (** name *) + Univdecls.universe_decl -> + bool -> (* polymorphic *) + Evd.evar_map -> (* Universes *) + Constr.t -> (** body *) + Constr.types -> (** type *) + Names.Id.t + +val new_instance : + ?abstract:bool -> (** Not abstract by default. *) + ?global:bool -> (** Not global by default. *) + ?refine:bool -> (** Allow refinement *) + program_mode:bool -> + Decl_kinds.polymorphic -> + local_binder_expr list -> + Vernacexpr.typeclass_constraint -> + (bool * constr_expr) option -> + ?generalize:bool -> + ?tac:unit Proofview.tactic -> + ?hook:(Globnames.global_reference -> unit) -> + Vernacexpr.hint_info_expr -> + Id.t + +(** Setting opacity *) + +val set_typeclass_transparency : evaluable_global_reference -> bool -> bool -> unit + +(** For generation on names based on classes only *) + +val id_of_class : typeclass -> Id.t + +(** Context command *) + +(** returns [false] if, for lack of section, it declares an assumption + (unless in a module type). *) +val context : Decl_kinds.polymorphic -> local_binder_expr list -> bool diff --git a/vernac/comAssumption.ml b/vernac/comAssumption.ml new file mode 100644 index 00000000..6a590758 --- /dev/null +++ b/vernac/comAssumption.ml @@ -0,0 +1,182 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* !axiom_into_instance); + optwrite = (:=) axiom_into_instance; } + +let should_axiom_into_instance = function + | Discharge -> + (* The typeclass behaviour of Variable and Context doesn't depend + on section status *) + true + | Global | Local -> !axiom_into_instance + +let declare_assumption is_coe (local,p,kind) (c,ctx) pl imps impl nl {CAst.v=ident} = +match local with +| Discharge when Lib.sections_are_opened () -> + let ctx = match ctx with + | Monomorphic_const_entry ctx -> ctx + | Polymorphic_const_entry ctx -> Univ.ContextSet.of_context ctx + in + let decl = (Lib.cwd(), SectionLocalAssum ((c,ctx),p,impl), IsAssumption kind) in + let _ = declare_variable ident decl in + let () = assumption_message ident in + let () = + if not !Flags.quiet && Proof_global.there_are_pending_proofs () then + Feedback.msg_info (str"Variable" ++ spc () ++ Id.print ident ++ + strbrk " is not visible from current goals") + in + let r = VarRef ident in + let () = Typeclasses.declare_instance None true r in + let () = if is_coe then Class.try_add_new_coercion r ~local:true false in + (r,Univ.Instance.empty,true) + +| Global | Local | Discharge -> + let do_instance = should_axiom_into_instance local in + let local = DeclareDef.get_locality ident ~kind:"axiom" local in + let inl = match nl with + | NoInline -> None + | DefaultInline -> Some (Flags.get_inline_level()) + | InlineAt i -> Some i + in + let decl = (ParameterEntry (None,(c,ctx),inl), IsAssumption kind) in + let kn = declare_constant ident ~local decl in + let gr = ConstRef kn in + let () = maybe_declare_manual_implicits false gr imps in + let () = Declare.declare_univ_binders gr pl in + let () = assumption_message ident in + let () = if do_instance then Typeclasses.declare_instance None false gr in + let () = if is_coe then Class.try_add_new_coercion gr ~local p in + let inst = match ctx with + | Polymorphic_const_entry ctx -> Univ.UContext.instance ctx + | Monomorphic_const_entry _ -> Univ.Instance.empty + in + (gr,inst,Lib.is_modtype_strict ()) + +let interp_assumption sigma env impls bl c = + let c = mkCProdN ?loc:(local_binders_loc bl) bl c in + let sigma, (ty, impls) = interp_type_evars_impls env sigma ~impls c in + sigma, (ty, impls) + +(* When monomorphic the universe constraints are declared with the first declaration only. *) +let next_uctx = + let empty_uctx = Monomorphic_const_entry Univ.ContextSet.empty in + function + | Polymorphic_const_entry _ as uctx -> uctx + | Monomorphic_const_entry _ -> empty_uctx + +let declare_assumptions idl is_coe k (c,uctx) pl imps nl = + let refs, status, _ = + List.fold_left (fun (refs,status,uctx) id -> + let ref',u',status' = + declare_assumption is_coe k (c,uctx) pl imps false nl id in + (ref',u')::refs, status' && status, next_uctx uctx) + ([],true,uctx) idl + in + List.rev refs, status + + +let maybe_error_many_udecls = function + | ({CAst.loc;v=id}, Some _) -> + user_err ?loc ~hdr:"many_universe_declarations" + Pp.(str "When declaring multiple axioms in one command, " ++ + str "only the first is allowed a universe binder " ++ + str "(which will be shared by the whole block).") + | (_, None) -> () + +let process_assumptions_udecls kind l = + let udecl, first_id = match l with + | (coe, ((id, udecl)::rest, c))::rest' -> + List.iter maybe_error_many_udecls rest; + List.iter (fun (coe, (idl, c)) -> List.iter maybe_error_many_udecls idl) rest'; + udecl, id + | (_, ([], _))::_ | [] -> assert false + in + let () = match kind, udecl with + | (Discharge, _, _), Some _ when Lib.sections_are_opened () -> + let loc = first_id.CAst.loc in + let msg = Pp.str "Section variables cannot be polymorphic." in + user_err ?loc msg + | _ -> () + in + udecl, List.map (fun (coe, (idl, c)) -> coe, (List.map fst idl, c)) l + +let do_assumptions kind nl l = + let open Context.Named.Declaration in + let env = Global.env () in + let udecl, l = process_assumptions_udecls kind l in + let sigma, udecl = Univdecls.interp_univ_decl_opt env udecl in + let l = + if pi2 kind (* poly *) then + (* Separate declarations so that A B : Type puts A and B in different levels. *) + List.fold_right (fun (is_coe,(idl,c)) acc -> + List.fold_right (fun id acc -> + (is_coe, ([id], c)) :: acc) idl acc) + l [] + else l + in + (* We intepret all declarations in the same evar_map, i.e. as a telescope. *) + let (sigma,_,_),l = List.fold_left_map (fun (sigma,env,ienv) (is_coe,(idl,c)) -> + let sigma,(t,imps) = interp_assumption sigma env ienv [] c in + let env = + EConstr.push_named_context (List.map (fun {CAst.v=id} -> LocalAssum (id,t)) idl) env in + let ienv = List.fold_right (fun {CAst.v=id} ienv -> + let impls = compute_internalization_data env sigma Variable t imps in + Id.Map.add id impls ienv) idl ienv in + ((sigma,env,ienv),((is_coe,idl),t,imps))) + (sigma,env,empty_internalization_env) l + in + let sigma = solve_remaining_evars all_and_fail_flags env sigma Evd.empty in + (* The universe constraints come from the whole telescope. *) + let sigma = Evd.minimize_universes sigma in + let nf_evar c = EConstr.to_constr sigma c in + let uvars, l = List.fold_left_map (fun uvars (coe,t,imps) -> + let t = nf_evar t in + let uvars = Univ.LSet.union uvars (Univops.universes_of_constr env t) in + uvars, (coe,t,imps)) + Univ.LSet.empty l + in + let sigma = Evd.restrict_universe_context sigma uvars in + let uctx = Evd.check_univ_decl ~poly:(pi2 kind) sigma udecl in + let ubinders = Evd.universe_binders sigma in + pi2 (List.fold_left (fun (subst,status,uctx) ((is_coe,idl),t,imps) -> + let t = replace_vars subst t in + let refs, status' = declare_assumptions idl is_coe kind (t,uctx) ubinders imps nl in + let subst' = List.map2 + (fun {CAst.v=id} (c,u) -> (id, Universes.constr_of_global_univ (c,u))) + idl refs + in + subst'@subst, status' && status, next_uctx uctx) + ([], true, uctx) l) diff --git a/vernac/comAssumption.mli b/vernac/comAssumption.mli new file mode 100644 index 00000000..56e32437 --- /dev/null +++ b/vernac/comAssumption.mli @@ -0,0 +1,36 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* + Vernacexpr.inline -> (ident_decl list * constr_expr) with_coercion list -> bool + +(************************************************************************) +(** Internal API *) +(************************************************************************) + +(** Exported for Classes *) + +(** returns [false] if the assumption is neither local to a section, + nor in a module type and meant to be instantiated. *) +val declare_assumption : coercion_flag -> assumption_kind -> + types in_constant_universes_entry -> + Universes.universe_binders -> Impargs.manual_implicits -> + bool (** implicit *) -> Vernacexpr.inline -> variable CAst.t -> + global_reference * Univ.Instance.t * bool diff --git a/vernac/comDefinition.ml b/vernac/comDefinition.ml new file mode 100644 index 00000000..b18a60a1 --- /dev/null +++ b/vernac/comDefinition.ml @@ -0,0 +1,134 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* + mkLambda (x,t,under_binders (push_rel (LocalAssum (x,t)) env) sigma f (n-1) c) + | LetIn (x,b,t,c) -> + mkLetIn (x,b,t,under_binders (push_rel (LocalDef (x,b,t)) env) sigma f (n-1) c) + | _ -> assert false + +let red_constant_entry n ce sigma = function + | None -> ce + | Some red -> + let proof_out = ce.const_entry_body in + let env = Global.env () in + let (redfun, _) = reduction_of_red_expr env red in + let redfun env sigma c = + let (_, c) = redfun env sigma c in + EConstr.Unsafe.to_constr c + in + { ce with const_entry_body = Future.chain proof_out + (fun ((body,ctx),eff) -> (under_binders env sigma redfun n body,ctx),eff) } + +let warn_implicits_in_term = + CWarnings.create ~name:"implicits-in-term" ~category:"implicits" + (fun () -> + strbrk "Implicit arguments declaration relies on type." ++ spc () ++ + strbrk "The term declares more implicits than the type here.") + +let check_imps ~impsty ~impsbody = + let b = + try + List.for_all (fun (key, (va:bool*bool*bool)) -> + (* Pervasives.(=) is OK for this type *) + Pervasives.(=) (List.assoc_f Impargs.explicitation_eq key impsty) va) + impsbody + with Not_found -> false + in + if not b then warn_implicits_in_term () + +let interp_definition pl bl poly red_option c ctypopt = + let open EConstr in + let env = Global.env() in + (* Explicitly bound universes and constraints *) + let evd, decl = Univdecls.interp_univ_decl_opt env pl in + (* Build the parameters *) + let evd, (impls, ((env_bl, ctx), imps1)) = interp_context_evars env evd bl in + (* Build the type *) + let evd, tyopt = Option.fold_left_map + (interp_type_evars_impls ~impls env_bl) + evd ctypopt + in + (* Build the body, and merge implicits from parameters and from type/body *) + let evd, c, imps, tyopt = + match tyopt with + | None -> + let evd, (c, impsbody) = interp_constr_evars_impls ~impls env_bl evd c in + evd, c, imps1@Impargs.lift_implicits (Context.Rel.nhyps ctx) impsbody, None + | Some (ty, impsty) -> + let evd, (c, impsbody) = interp_casted_constr_evars_impls ~impls env_bl evd c ty in + check_imps ~impsty ~impsbody; + evd, c, imps1@Impargs.lift_implicits (Context.Rel.nhyps ctx) impsty, Some ty + in + (* universe minimization *) + let evd = Evd.minimize_universes evd in + (* Substitute evars and universes, and add parameters. + Note: in program mode some evars may remain. *) + let ctx = List.map (EConstr.to_rel_decl evd) ctx in + let c = Term.it_mkLambda_or_LetIn (EConstr.to_constr evd c) ctx in + let tyopt = Option.map (fun ty -> Term.it_mkProd_or_LetIn (EConstr.to_constr evd ty) ctx) tyopt in + (* Keep only useful universes. *) + let uvars_fold uvars c = + Univ.LSet.union uvars (universes_of_constr env evd (of_constr c)) + in + let uvars = List.fold_left uvars_fold Univ.LSet.empty (Option.List.cons tyopt [c]) in + let evd = Evd.restrict_universe_context evd uvars in + (* Check we conform to declared universes *) + let uctx = Evd.check_univ_decl ~poly evd decl in + (* We're done! *) + let ce = definition_entry ?types:tyopt ~univs:uctx c in + (red_constant_entry (Context.Rel.length ctx) ce evd red_option, evd, decl, imps) + +let check_definition (ce, evd, _, imps) = + check_evars_are_solved (Global.env ()) evd Evd.empty; + ce + +let do_definition ~program_mode ident k univdecl bl red_option c ctypopt hook = + let (ce, evd, univdecl, imps as def) = + interp_definition univdecl bl (pi2 k) red_option c ctypopt + in + if program_mode then + let env = Global.env () in + let (c,ctx), sideff = Future.force ce.const_entry_body in + assert(Safe_typing.empty_private_constants = sideff); + assert(Univ.ContextSet.is_empty ctx); + let typ = match ce.const_entry_type with + | Some t -> t + | None -> EConstr.to_constr evd (Retyping.get_type_of env evd (EConstr.of_constr c)) + in + Obligations.check_evars env evd; + let obls, _, c, cty = + Obligations.eterm_obligations env ident evd 0 c typ + in + let ctx = Evd.evar_universe_context evd in + let hook = Lemmas.mk_hook (fun l r _ -> Lemmas.call_hook (fun exn -> exn) hook l r) in + ignore(Obligations.add_definition + ident ~term:c cty ctx ~univdecl ~implicits:imps ~kind:k ~hook obls) + else let ce = check_definition def in + ignore(DeclareDef.declare_definition ident k ce (Evd.universe_binders evd) imps + (Lemmas.mk_hook + (fun l r -> Lemmas.call_hook (fun exn -> exn) hook l r;r))) diff --git a/vernac/comDefinition.mli b/vernac/comDefinition.mli new file mode 100644 index 00000000..6f81c457 --- /dev/null +++ b/vernac/comDefinition.mli @@ -0,0 +1,32 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* + Id.t -> definition_kind -> universe_decl_expr option -> + local_binder_expr list -> red_expr option -> constr_expr -> + constr_expr option -> unit Lemmas.declaration_hook -> unit + +(************************************************************************) +(** Internal API *) +(************************************************************************) + +(** Not used anywhere. *) +val interp_definition : + universe_decl_expr option -> local_binder_expr list -> polymorphic -> red_expr option -> constr_expr -> + constr_expr option -> Safe_typing.private_constants definition_entry * Evd.evar_map * + Univdecls.universe_decl * Impargs.manual_implicits diff --git a/vernac/comFixpoint.ml b/vernac/comFixpoint.ml new file mode 100644 index 00000000..1c141c01 --- /dev/null +++ b/vernac/comFixpoint.ml @@ -0,0 +1,353 @@ +open Pp +open CErrors +open Util +open Constr +open Vars +open Termops +open Declare +open Names +open Constrexpr +open Constrexpr_ops +open Constrintern +open Decl_kinds +open Pretyping +open Evarutil +open Evarconv +open Misctypes + +module RelDecl = Context.Rel.Declaration + +(* 3c| Fixpoints and co-fixpoints *) + +(* An (unoptimized) function that maps preorders to partial orders... + + Input: a list of associations (x,[y1;...;yn]), all yi distincts + and different of x, meaning x<=y1, ..., x<=yn + + Output: a list of associations (x,Inr [y1;...;yn]), collecting all + distincts yi greater than x, _or_, (x, Inl y) meaning that + x is in the same class as y (in which case, x occurs + nowhere else in the association map) + + partial_order : ('a * 'a list) list -> ('a * ('a,'a list) union) list +*) + +let rec partial_order cmp = function + | [] -> [] + | (x,xge)::rest -> + let rec browse res xge' = function + | [] -> + let res = List.map (function + | (z, Inr zge) when List.mem_f cmp x zge -> + (z, Inr (List.union cmp zge xge')) + | r -> r) res in + (x,Inr xge')::res + | y::xge -> + let rec link y = + try match List.assoc_f cmp y res with + | Inl z -> link z + | Inr yge -> + if List.mem_f cmp x yge then + let res = List.remove_assoc_f cmp y res in + let res = List.map (function + | (z, Inl t) -> + if cmp t y then (z, Inl x) else (z, Inl t) + | (z, Inr zge) -> + if List.mem_f cmp y zge then + (z, Inr (List.add_set cmp x (List.remove cmp y zge))) + else + (z, Inr zge)) res in + browse ((y,Inl x)::res) xge' (List.union cmp xge (List.remove cmp x yge)) + else + browse res (List.add_set cmp y (List.union cmp xge' yge)) xge + with Not_found -> browse res (List.add_set cmp y xge') xge + in link y + in browse (partial_order cmp rest) [] xge + +let non_full_mutual_message x xge y yge isfix rest = + let reason = + if Id.List.mem x yge then + Id.print y ++ str " depends on " ++ Id.print x ++ strbrk " but not conversely" + else if Id.List.mem y xge then + Id.print x ++ str " depends on " ++ Id.print y ++ strbrk " but not conversely" + else + Id.print y ++ str " and " ++ Id.print x ++ strbrk " are not mutually dependent" in + let e = if List.is_empty rest then reason else strbrk "e.g., " ++ reason in + let k = if isfix then "fixpoint" else "cofixpoint" in + let w = + if isfix + then strbrk "Well-foundedness check may fail unexpectedly." ++ fnl() + else mt () in + strbrk "Not a fully mutually defined " ++ str k ++ fnl () ++ + str "(" ++ e ++ str ")." ++ fnl () ++ w + +let warn_non_full_mutual = + CWarnings.create ~name:"non-full-mutual" ~category:"fixpoints" + (fun (x,xge,y,yge,isfix,rest) -> + non_full_mutual_message x xge y yge isfix rest) + +let check_mutuality env evd isfix fixl = + let names = List.map fst fixl in + let preorder = + List.map (fun (id,def) -> + (id, List.filter (fun id' -> not (Id.equal id id') && occur_var env evd id' (EConstr.of_constr def)) names)) + fixl in + let po = partial_order Id.equal preorder in + match List.filter (function (_,Inr _) -> true | _ -> false) po with + | (x,Inr xge)::(y,Inr yge)::rest -> + warn_non_full_mutual (x,xge,y,yge,isfix,rest) + | _ -> () + +type structured_fixpoint_expr = { + fix_name : Id.t; + fix_univs : universe_decl_expr option; + fix_annot : lident option; + fix_binders : local_binder_expr list; + fix_body : constr_expr option; + fix_type : constr_expr +} + +let interp_fix_context ~cofix env sigma fix = + let before, after = if not cofix then split_at_annot fix.fix_binders fix.fix_annot else [], fix.fix_binders in + let sigma, (impl_env, ((env', ctx), imps)) = interp_context_evars env sigma before in + let sigma, (impl_env', ((env'', ctx'), imps')) = interp_context_evars ~impl_env ~shift:(Context.Rel.nhyps ctx) env' sigma after in + let annot = Option.map (fun _ -> List.length (assums_of_rel_context ctx)) fix.fix_annot in + sigma, ((env'', ctx' @ ctx), (impl_env',imps @ imps'), annot) + +let interp_fix_ccl sigma impls (env,_) fix = + interp_type_evars_impls ~impls env sigma fix.fix_type + +let interp_fix_body env_rec sigma impls (_,ctx) fix ccl = + let open EConstr in + Option.cata (fun body -> + let env = push_rel_context ctx env_rec in + let sigma, body = interp_casted_constr_evars env sigma ~impls body ccl in + sigma, Some (it_mkLambda_or_LetIn body ctx)) (sigma, None) fix.fix_body + +let build_fix_type (_,ctx) ccl = EConstr.it_mkProd_or_LetIn ccl ctx + +let prepare_recursive_declaration fixnames fixtypes fixdefs = + let defs = List.map (subst_vars (List.rev fixnames)) fixdefs in + let names = List.map (fun id -> Name id) fixnames in + (Array.of_list names, Array.of_list fixtypes, Array.of_list defs) + +(* Jump over let-bindings. *) + +let compute_possible_guardness_evidences (ctx,_,recindex) = + (* A recursive index is characterized by the number of lambdas to + skip before finding the relevant inductive argument *) + match recindex with + | Some i -> [i] + | None -> + (* If recursive argument was not given by user, we try all args. + An earlier approach was to look only for inductive arguments, + but doing it properly involves delta-reduction, and it finally + doesn't seem to worth the effort (except for huge mutual + fixpoints ?) *) + List.interval 0 (Context.Rel.nhyps ctx - 1) + +type recursive_preentry = + Id.t list * constr option list * types list + +(* Wellfounded definition *) + +let contrib_name = "Program" +let subtac_dir = [contrib_name] +let tactics_module = subtac_dir @ ["Tactics"] + +let init_constant dir s sigma = + Evarutil.new_global sigma (Coqlib.coq_reference "Command" dir s) + +let fix_proto = init_constant tactics_module "fix_proto" + +let interp_recursive ~program_mode ~cofix fixl notations = + let open Context.Named.Declaration in + let open EConstr in + let env = Global.env() in + let fixnames = List.map (fun fix -> fix.fix_name) fixl in + + (* Interp arities allowing for unresolved types *) + let all_universes = + List.fold_right (fun sfe acc -> + match sfe.fix_univs , acc with + | None , acc -> acc + | x , None -> x + | Some ls , Some us -> + let lsu = ls.univdecl_instance and usu = us.univdecl_instance in + if not (CList.for_all2eq (fun x y -> Id.equal x.CAst.v y.CAst.v) lsu usu) then + user_err Pp.(str "(co)-recursive definitions should all have the same universe binders"); + Some us) fixl None in + let sigma, decl = Univdecls.interp_univ_decl_opt env all_universes in + let sigma, (fixctxs, fiximppairs, fixannots) = + on_snd List.split3 @@ + List.fold_left_map (fun sigma -> interp_fix_context env sigma ~cofix) sigma fixl in + let fixctximpenvs, fixctximps = List.split fiximppairs in + let sigma, (fixccls,fixcclimps) = + on_snd List.split @@ + List.fold_left3_map interp_fix_ccl sigma fixctximpenvs fixctxs fixl in + let fixtypes = List.map2 build_fix_type fixctxs fixccls in + let fixtypes = List.map (fun c -> nf_evar sigma c) fixtypes in + let fiximps = List.map3 + (fun ctximps cclimps (_,ctx) -> ctximps@(Impargs.lift_implicits (Context.Rel.nhyps ctx) cclimps)) + fixctximps fixcclimps fixctxs in + let sigma, rec_sign = + List.fold_left2 + (fun (sigma, env') id t -> + if program_mode then + let sigma, sort = Typing.type_of ~refresh:true env sigma t in + let sigma, fixprot = + try + let sigma, h_term = fix_proto sigma in + let app = mkApp (h_term, [|sort; t|]) in + let _evd = ref sigma in + let res = Typing.e_solve_evars env _evd app in + !_evd, res + with e when CErrors.noncritical e -> sigma, t + in + sigma, LocalAssum (id,fixprot) :: env' + else sigma, LocalAssum (id,t) :: env') + (sigma,[]) fixnames fixtypes + in + let env_rec = push_named_context rec_sign env in + + (* Get interpretation metadatas *) + let impls = compute_internalization_env env sigma Recursive fixnames fixtypes fiximps in + + (* Interp bodies with rollback because temp use of notations/implicit *) + let sigma, fixdefs = + Metasyntax.with_syntax_protection (fun () -> + List.iter (Metasyntax.set_notation_for_interpretation env_rec impls) notations; + List.fold_left4_map + (fun sigma fixctximpenv -> interp_fix_body env_rec sigma (Id.Map.fold Id.Map.add fixctximpenv impls)) + sigma fixctximpenvs fixctxs fixl fixccls) + () in + + (* Instantiate evars and check all are resolved *) + let sigma = solve_unif_constraints_with_heuristics env_rec sigma in + let sigma, _ = nf_evars_and_universes sigma in + let fixdefs = List.map (fun c -> Option.map EConstr.(to_constr sigma) c) fixdefs in + let fixtypes = List.map EConstr.(to_constr sigma) fixtypes in + let fixctxs = List.map (fun (_,ctx) -> ctx) fixctxs in + + (* Build the fix declaration block *) + (env,rec_sign,decl,sigma), (fixnames,fixdefs,fixtypes), List.combine3 fixctxs fiximps fixannots + +let check_recursive isfix env evd (fixnames,fixdefs,_) = + check_evars_are_solved env evd Evd.empty; + if List.for_all Option.has_some fixdefs then begin + let fixdefs = List.map Option.get fixdefs in + check_mutuality env evd isfix (List.combine fixnames fixdefs) + end + +let interp_fixpoint ~cofix l ntns = + let (env,_,pl,evd),fix,info = interp_recursive ~program_mode:false ~cofix l ntns in + check_recursive true env evd fix; + (fix,pl,Evd.evar_universe_context evd,info) + +let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) indexes ntns = + if List.exists Option.is_empty fixdefs then + (* Some bodies to define by proof *) + let thms = + List.map3 (fun id t (ctx,imps,_) -> (id,(EConstr.of_constr t,(List.map RelDecl.get_name ctx,imps)))) + fixnames fixtypes fiximps in + let init_tac = + Some (List.map (Option.cata (EConstr.of_constr %> Tactics.exact_no_check) Tacticals.New.tclIDTAC) + fixdefs) in + let evd = Evd.from_ctx ctx in + Lemmas.start_proof_with_initialization (local,poly,DefinitionBody Fixpoint) + evd pl (Some(false,indexes,init_tac)) thms None (Lemmas.mk_hook (fun _ _ -> ())) + else begin + (* We shortcut the proof process *) + let fixdefs = List.map Option.get fixdefs in + let fixdecls = prepare_recursive_declaration fixnames fixtypes fixdefs in + let env = Global.env() in + let indexes = search_guard env indexes fixdecls in + let fiximps = List.map (fun (n,r,p) -> r) fiximps in + let vars = Univops.universes_of_constr env (mkFix ((indexes,0),fixdecls)) in + let fixdecls = + List.map_i (fun i _ -> mkFix ((indexes,i),fixdecls)) 0 fixnames in + let evd = Evd.from_ctx ctx in + let evd = Evd.restrict_universe_context evd vars in + let ctx = Evd.check_univ_decl ~poly evd pl in + let pl = Evd.universe_binders evd in + let fixdecls = List.map Safe_typing.mk_pure_proof fixdecls in + ignore (List.map4 (DeclareDef.declare_fix (local, poly, Fixpoint) pl ctx) + fixnames fixdecls fixtypes fiximps); + (* Declare the recursive definitions *) + fixpoint_message (Some indexes) fixnames; + end; + (* Declare notations *) + List.iter (Metasyntax.add_notation_interpretation (Global.env())) ntns + +let declare_cofixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) ntns = + if List.exists Option.is_empty fixdefs then + (* Some bodies to define by proof *) + let thms = + List.map3 (fun id t (ctx,imps,_) -> (id,(EConstr.of_constr t,(List.map RelDecl.get_name ctx,imps)))) + fixnames fixtypes fiximps in + let init_tac = + Some (List.map (Option.cata (EConstr.of_constr %> Tactics.exact_no_check) Tacticals.New.tclIDTAC) + fixdefs) in + let evd = Evd.from_ctx ctx in + Lemmas.start_proof_with_initialization (Global,poly, DefinitionBody CoFixpoint) + evd pl (Some(true,[],init_tac)) thms None (Lemmas.mk_hook (fun _ _ -> ())) + else begin + (* We shortcut the proof process *) + let fixdefs = List.map Option.get fixdefs in + let fixdecls = prepare_recursive_declaration fixnames fixtypes fixdefs in + let fixdecls = List.map_i (fun i _ -> mkCoFix (i,fixdecls)) 0 fixnames in + let env = Global.env () in + let vars = Univops.universes_of_constr env (List.hd fixdecls) in + let fixdecls = List.map Safe_typing.mk_pure_proof fixdecls in + let fiximps = List.map (fun (len,imps,idx) -> imps) fiximps in + let evd = Evd.from_ctx ctx in + let evd = Evd.restrict_universe_context evd vars in + let ctx = Evd.check_univ_decl ~poly evd pl in + let pl = Evd.universe_binders evd in + ignore (List.map4 (DeclareDef.declare_fix (local, poly, CoFixpoint) pl ctx) + fixnames fixdecls fixtypes fiximps); + (* Declare the recursive definitions *) + cofixpoint_message fixnames + end; + (* Declare notations *) + List.iter (Metasyntax.add_notation_interpretation (Global.env())) ntns + +let extract_decreasing_argument limit = function + | (na,CStructRec) -> na + | (na,_) when not limit -> na + | _ -> user_err Pp.(str + "Only structural decreasing is supported for a non-Program Fixpoint") + +let extract_fixpoint_components limit l = + let fixl, ntnl = List.split l in + let fixl = List.map (fun (({CAst.v=id},pl),ann,bl,typ,def) -> + let ann = extract_decreasing_argument limit ann in + {fix_name = id; fix_annot = ann; fix_univs = pl; + fix_binders = bl; fix_body = def; fix_type = typ}) fixl in + fixl, List.flatten ntnl + +let extract_cofixpoint_components l = + let fixl, ntnl = List.split l in + List.map (fun (({CAst.v=id},pl),bl,typ,def) -> + {fix_name = id; fix_annot = None; fix_univs = pl; + fix_binders = bl; fix_body = def; fix_type = typ}) fixl, + List.flatten ntnl + +let check_safe () = + let open Declarations in + let flags = Environ.typing_flags (Global.env ()) in + flags.check_universes && flags.check_guarded + +let do_fixpoint local poly l = + let fixl, ntns = extract_fixpoint_components true l in + let (_, _, _, info as fix) = interp_fixpoint ~cofix:false fixl ntns in + let possible_indexes = + List.map compute_possible_guardness_evidences info in + declare_fixpoint local poly fix possible_indexes ntns; + if not (check_safe ()) then Feedback.feedback Feedback.AddedAxiom else () + +let do_cofixpoint local poly l = + let fixl,ntns = extract_cofixpoint_components l in + let cofix = interp_fixpoint ~cofix:true fixl ntns in + declare_cofixpoint local poly cofix ntns; + if not (check_safe ()) then Feedback.feedback Feedback.AddedAxiom else () diff --git a/vernac/comFixpoint.mli b/vernac/comFixpoint.mli new file mode 100644 index 00000000..36c2993a --- /dev/null +++ b/vernac/comFixpoint.mli @@ -0,0 +1,95 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* polymorphic -> (fixpoint_expr * decl_notation list) list -> unit + +val do_cofixpoint : + (* When [false], assume guarded. *) + locality -> polymorphic -> (cofixpoint_expr * decl_notation list) list -> unit + +(************************************************************************) +(** Internal API *) +(************************************************************************) + +type structured_fixpoint_expr = { + fix_name : Id.t; + fix_univs : Constrexpr.universe_decl_expr option; + fix_annot : Misctypes.lident option; + fix_binders : local_binder_expr list; + fix_body : constr_expr option; + fix_type : constr_expr +} + +(** Typing global fixpoints and cofixpoint_expr *) + +(** Exported for Program *) +val interp_recursive : + (* Misc arguments *) + program_mode:bool -> cofix:bool -> + (* Notations of the fixpoint / should that be folded in the previous argument? *) + structured_fixpoint_expr list -> decl_notation list -> + + (* env / signature / univs / evar_map *) + (Environ.env * EConstr.named_context * Univdecls.universe_decl * Evd.evar_map) * + (* names / defs / types *) + (Id.t list * Constr.constr option list * Constr.types list) * + (* ctx per mutual def / implicits / struct annotations *) + (EConstr.rel_context * Impargs.manual_explicitation list * int option) list + +(** Exported for Funind *) + +(** Extracting the semantical components out of the raw syntax of + (co)fixpoints declarations *) + +val extract_fixpoint_components : bool -> + (fixpoint_expr * decl_notation list) list -> + structured_fixpoint_expr list * decl_notation list + +val extract_cofixpoint_components : + (cofixpoint_expr * decl_notation list) list -> + structured_fixpoint_expr list * decl_notation list + +type recursive_preentry = + Id.t list * constr option list * types list + +val interp_fixpoint : + cofix:bool -> + structured_fixpoint_expr list -> decl_notation list -> + recursive_preentry * Univdecls.universe_decl * UState.t * + (EConstr.rel_context * Impargs.manual_implicits * int option) list + +(** Registering fixpoints and cofixpoints in the environment *) +(** [Not used so far] *) +val declare_fixpoint : + locality -> polymorphic -> + recursive_preentry * Univdecls.universe_decl * UState.t * + (Context.Rel.t * Impargs.manual_implicits * int option) list -> + Proof_global.lemma_possible_guards -> decl_notation list -> unit + +val declare_cofixpoint : locality -> polymorphic -> + recursive_preentry * Univdecls.universe_decl * UState.t * + (Context.Rel.t * Impargs.manual_implicits * int option) list -> + decl_notation list -> unit + +(** Very private function, do not use *) +val compute_possible_guardness_evidences : + ('a, 'b) Context.Rel.pt * 'c * int option -> int list diff --git a/vernac/comInductive.ml b/vernac/comInductive.ml new file mode 100644 index 00000000..05c40dbd --- /dev/null +++ b/vernac/comInductive.ml @@ -0,0 +1,453 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* function + | CProdN (bl,c) -> CProdN (bl,complete_conclusion a cs c) + | CLetIn (na,b,t,c) -> CLetIn (na,b,t,complete_conclusion a cs c) + | CHole (k, _, _) -> + let (has_no_args,name,params) = a in + if not has_no_args then + user_err ?loc + (strbrk"Cannot infer the non constant arguments of the conclusion of " + ++ Id.print cs ++ str "."); + let args = List.map (fun id -> CAst.(make ?loc @@ CRef(make ?loc @@ Ident id,None))) params in + CAppExpl ((None,CAst.make ?loc @@ Ident name,None),List.rev args) + | c -> c + ) + +let push_types env idl tl = + List.fold_left2 (fun env id t -> EConstr.push_rel (LocalAssum (Name id,t)) env) + env idl tl + +type structured_one_inductive_expr = { + ind_name : Id.t; + ind_univs : universe_decl_expr option; + ind_arity : constr_expr; + ind_lc : (Id.t * constr_expr) list +} + +type structured_inductive_expr = + local_binder_expr list * structured_one_inductive_expr list + +let minductive_message warn = function + | [] -> user_err Pp.(str "No inductive definition.") + | [x] -> (Id.print x ++ str " is defined" ++ + if warn then str " as a non-primitive record" else mt()) + | l -> hov 0 (prlist_with_sep pr_comma Id.print l ++ + spc () ++ str "are defined") + +let check_all_names_different indl = + let ind_names = List.map (fun ind -> ind.ind_name) indl in + let cstr_names = List.map_append (fun ind -> List.map fst ind.ind_lc) indl in + let l = List.duplicates Id.equal ind_names in + let () = match l with + | [] -> () + | t :: _ -> raise (InductiveError (SameNamesTypes t)) + in + let l = List.duplicates Id.equal cstr_names in + let () = match l with + | [] -> () + | c :: _ -> raise (InductiveError (SameNamesConstructors (List.hd l))) + in + let l = List.intersect Id.equal ind_names cstr_names in + match l with + | [] -> () + | _ -> raise (InductiveError (SameNamesOverlap l)) + +let mk_mltype_data sigma env assums arity indname = + let is_ml_type = is_sort env sigma arity in + (is_ml_type,indname,assums) + +let prepare_param = function + | LocalAssum (na,t) -> Name.get_id na, LocalAssumEntry t + | LocalDef (na,b,_) -> Name.get_id na, LocalDefEntry b + +(** Make the arity conclusion flexible to avoid generating an upper bound universe now, + only if the universe does not appear anywhere else. + This is really a hack to stay compatible with the semantics of template polymorphic + inductives which are recognized when a "Type" appears at the end of the conlusion in + the source syntax. *) + +let rec check_anonymous_type ind = + let open Glob_term in + match DAst.get ind with + | GSort (GType []) -> true + | GProd ( _, _, _, e) + | GLetIn (_, _, _, e) + | GLambda (_, _, _, e) + | GApp (e, _) + | GCast (e, _) -> check_anonymous_type e + | _ -> false + +let make_conclusion_flexible sigma ty poly = + if poly && Term.isArity ty then + let _, concl = Term.destArity ty in + match concl with + | Type u -> + (match Univ.universe_level u with + | Some u -> + Evd.make_flexible_variable sigma ~algebraic:true u + | None -> sigma) + | _ -> sigma + else sigma + +let is_impredicative env u = + u = Prop Null || (is_impredicative_set env && u = Prop Pos) + +let interp_ind_arity env sigma ind = + let c = intern_gen IsType env sigma ind.ind_arity in + let impls = Implicit_quantifiers.implicits_of_glob_constr ~with_products:true c in + let sigma,t = understand_tcc env sigma ~expected_type:IsType c in + let pseudo_poly = check_anonymous_type c in + let () = if not (Reductionops.is_arity env sigma t) then + user_err ?loc:(constr_loc ind.ind_arity) (str "Not an arity") + in + sigma, (t, pseudo_poly, impls) + +let interp_cstrs env sigma impls mldata arity ind = + let cnames,ctyps = List.split ind.ind_lc in + (* Complete conclusions of constructor types if given in ML-style syntax *) + let ctyps' = List.map2 (complete_conclusion mldata) cnames ctyps in + (* Interpret the constructor types *) + let sigma, (ctyps'', cimpls) = + on_snd List.split @@ + List.fold_left_map (fun sigma l -> + on_snd (on_fst EConstr.Unsafe.to_constr) @@ + interp_type_evars_impls env sigma ~impls l) sigma ctyps' in + sigma, (cnames, ctyps'', cimpls) + +let sign_level env evd sign = + fst (List.fold_right + (fun d (lev,env) -> + match d with + | LocalDef _ -> lev, push_rel d env + | LocalAssum _ -> + let s = destSort (Reduction.whd_all env + (EConstr.Unsafe.to_constr (nf_evar evd (Retyping.get_type_of env evd (EConstr.of_constr (RelDecl.get_type d)))))) + in + let u = univ_of_sort s in + (Univ.sup u lev, push_rel d env)) + sign (Univ.type0m_univ,env)) + +let sup_list min = List.fold_left Univ.sup min + +let extract_level env evd min tys = + let sorts = List.map (fun ty -> + let ctx, concl = Reduction.dest_prod_assum env ty in + sign_level env evd (LocalAssum (Anonymous, concl) :: ctx)) tys + in sup_list min sorts + +let is_flexible_sort evd u = + match Univ.Universe.level u with + | Some l -> Evd.is_flexible_level evd l + | None -> false + +let inductive_levels env evd poly arities inds = + let destarities = List.map (fun x -> x, Reduction.dest_arity env x) arities in + let levels = List.map (fun (x,(ctx,a)) -> + if a = Prop Null then None + else Some (univ_of_sort a)) destarities + in + let cstrs_levels, min_levels, sizes = + CList.split3 + (List.map2 (fun (_,tys,_) (arity,(ctx,du)) -> + let len = List.length tys in + let minlev = Sorts.univ_of_sort du in + let minlev = + if len > 1 && not (is_impredicative env du) then + Univ.sup minlev Univ.type0_univ + else minlev + in + let minlev = + (** Indices contribute. *) + if Indtypes.is_indices_matter () && List.length ctx > 0 then ( + let ilev = sign_level env evd ctx in + Univ.sup ilev minlev) + else minlev + in + let clev = extract_level env evd minlev tys in + (clev, minlev, len)) inds destarities) + in + (* Take the transitive closure of the system of constructors *) + (* level constraints and remove the recursive dependencies *) + let levels' = Universes.solve_constraints_system (Array.of_list levels) + (Array.of_list cstrs_levels) (Array.of_list min_levels) + in + let evd, arities = + CList.fold_left3 (fun (evd, arities) cu (arity,(ctx,du)) len -> + if is_impredicative env du then + (** Any product is allowed here. *) + evd, arity :: arities + else (** If in a predicative sort, or asked to infer the type, + we take the max of: + - indices (if in indices-matter mode) + - constructors + - Type(1) if there is more than 1 constructor + *) + (** Constructors contribute. *) + let evd = + if Sorts.is_set du then + if not (Evd.check_leq evd cu Univ.type0_univ) then + raise (Indtypes.InductiveError Indtypes.LargeNonPropInductiveNotInType) + else evd + else evd + (* Evd.set_leq_sort env evd (Type cu) du *) + in + let evd = + if len >= 2 && Univ.is_type0m_univ cu then + (** "Polymorphic" type constraint and more than one constructor, + should not land in Prop. Add constraint only if it would + land in Prop directly (no informative arguments as well). *) + Evd.set_leq_sort env evd (Prop Pos) du + else evd + in + let duu = Sorts.univ_of_sort du in + let evd = + if not (Univ.is_small_univ duu) && Univ.Universe.equal cu duu then + if is_flexible_sort evd duu && not (Evd.check_leq evd Univ.type0_univ duu) then + Evd.set_eq_sort env evd (Prop Null) du + else evd + else Evd.set_eq_sort env evd (Type cu) du + in + (evd, arity :: arities)) + (evd,[]) (Array.to_list levels') destarities sizes + in evd, List.rev arities + +let check_named {CAst.loc;v=na} = match na with +| Name _ -> () +| Anonymous -> + let msg = str "Parameters must be named." in + user_err ?loc msg + + +let check_param = function +| CLocalDef (na, _, _) -> check_named na +| CLocalAssum (nas, Default _, _) -> List.iter check_named nas +| CLocalAssum (nas, Generalized _, _) -> () +| CLocalPattern {CAst.loc} -> + Loc.raise ?loc (Stream.Error "pattern with quote not allowed here") + +let interp_mutual_inductive (paramsl,indl) notations cum poly prv finite = + check_all_names_different indl; + List.iter check_param paramsl; + let env0 = Global.env() in + let pl = (List.hd indl).ind_univs in + let sigma, decl = Univdecls.interp_univ_decl_opt env0 pl in + let sigma, (impls, ((env_params, ctx_params), userimpls)) = + interp_context_evars env0 sigma paramsl + in + let indnames = List.map (fun ind -> ind.ind_name) indl in + + (* Names of parameters as arguments of the inductive type (defs removed) *) + let assums = List.filter is_local_assum ctx_params in + let params = List.map (RelDecl.get_name %> Name.get_id) assums in + + (* Interpret the arities *) + let sigma, arities = List.fold_left_map (fun sigma -> interp_ind_arity env_params sigma) sigma indl in + + let fullarities = List.map (fun (c, _, _) -> EConstr.it_mkProd_or_LetIn c ctx_params) arities in + let env_ar = push_types env0 indnames fullarities in + let env_ar_params = EConstr.push_rel_context ctx_params env_ar in + + (* Compute interpretation metadatas *) + let indimpls = List.map (fun (_, _, impls) -> userimpls @ + lift_implicits (Context.Rel.nhyps ctx_params) impls) arities in + let arities = List.map pi1 arities and aritypoly = List.map pi2 arities in + let impls = compute_internalization_env env0 sigma ~impls (Inductive (params,true)) indnames fullarities indimpls in + let ntn_impls = compute_internalization_env env0 sigma (Inductive (params,true)) indnames fullarities indimpls in + let mldatas = List.map2 (mk_mltype_data sigma env_params params) arities indnames in + + let sigma, constructors = + Metasyntax.with_syntax_protection (fun () -> + (* Temporary declaration of notations and scopes *) + List.iter (Metasyntax.set_notation_for_interpretation env_params ntn_impls) notations; + (* Interpret the constructor types *) + List.fold_left3_map (fun sigma -> interp_cstrs env_ar_params sigma impls) sigma mldatas arities indl) + () in + + (* Try further to solve evars, and instantiate them *) + let sigma = solve_remaining_evars all_and_fail_flags env_params sigma Evd.empty in + (* Compute renewed arities *) + let sigma, nf = nf_evars_and_universes sigma in + let constructors = List.map (fun (idl,cl,impsl) -> (idl,List.map nf cl,impsl)) constructors in + let arities = List.map EConstr.(to_constr sigma) arities in + let sigma = List.fold_left2 (fun sigma ty poly -> make_conclusion_flexible sigma ty poly) sigma arities aritypoly in + let sigma, arities = inductive_levels env_ar_params sigma poly arities constructors in + let sigma, nf' = nf_evars_and_universes sigma in + let arities = List.map nf' arities in + let constructors = List.map (fun (idl,cl,impsl) -> (idl,List.map nf' cl,impsl)) constructors in + let ctx_params = List.map Termops.(map_rel_decl (EConstr.to_constr sigma)) ctx_params in + let uctx = Evd.check_univ_decl ~poly sigma decl in + List.iter (fun c -> check_evars env_params Evd.empty sigma (EConstr.of_constr c)) arities; + Context.Rel.iter (fun c -> check_evars env0 Evd.empty sigma (EConstr.of_constr c)) ctx_params; + List.iter (fun (_,ctyps,_) -> + List.iter (fun c -> check_evars env_ar_params Evd.empty sigma (EConstr.of_constr c)) ctyps) + constructors; + + (* Build the inductive entries *) + let entries = List.map4 (fun ind arity template (cnames,ctypes,cimpls) -> { + mind_entry_typename = ind.ind_name; + mind_entry_arity = arity; + mind_entry_template = template; + mind_entry_consnames = cnames; + mind_entry_lc = ctypes + }) indl arities aritypoly constructors in + let impls = + let len = Context.Rel.nhyps ctx_params in + List.map2 (fun indimpls (_,_,cimpls) -> + indimpls, List.map (fun impls -> + userimpls @ (lift_implicits len impls)) cimpls) indimpls constructors + in + let univs = + match uctx with + | Polymorphic_const_entry uctx -> + if cum then + Cumulative_ind_entry (Univ.CumulativityInfo.from_universe_context uctx) + else Polymorphic_ind_entry uctx + | Monomorphic_const_entry uctx -> + Monomorphic_ind_entry uctx + in + (* Build the mutual inductive entry *) + let mind_ent = + { mind_entry_params = List.map prepare_param ctx_params; + mind_entry_record = None; + mind_entry_finite = finite; + mind_entry_inds = entries; + mind_entry_private = if prv then Some false else None; + mind_entry_universes = univs; + } + in + (if poly && cum then + InferCumulativity.infer_inductive env_ar mind_ent + else mind_ent), Evd.universe_binders sigma, impls + +(* Very syntactical equality *) +let eq_local_binders bl1 bl2 = + List.equal local_binder_eq bl1 bl2 + +let extract_coercions indl = + let mkqid (_,({CAst.v=id},_)) = qualid_of_ident id in + let extract lc = List.filter (fun (iscoe,_) -> iscoe) lc in + List.map mkqid (List.flatten(List.map (fun (_,_,_,lc) -> extract lc) indl)) + +let extract_params indl = + let paramsl = List.map (fun (_,params,_,_) -> params) indl in + match paramsl with + | [] -> anomaly (Pp.str "empty list of inductive types.") + | params::paramsl -> + if not (List.for_all (eq_local_binders params) paramsl) then user_err Pp.(str + "Parameters should be syntactically the same for each inductive type."); + params + +let extract_inductive indl = + List.map (fun (({CAst.v=indname},pl),_,ar,lc) -> { + ind_name = indname; ind_univs = pl; + ind_arity = Option.cata (fun x -> x) (CAst.make @@ CSort (GType [])) ar; + ind_lc = List.map (fun (_,({CAst.v=id},t)) -> (id,t)) lc + }) indl + +let extract_mutual_inductive_declaration_components indl = + let indl,ntnl = List.split indl in + let params = extract_params indl in + let coes = extract_coercions indl in + let indl = extract_inductive indl in + (params,indl), coes, List.flatten ntnl + +let is_recursive mie = + let rec is_recursive_constructor lift typ = + match Constr.kind typ with + | Prod (_,arg,rest) -> + not (EConstr.Vars.noccurn Evd.empty (** FIXME *) lift (EConstr.of_constr arg)) || + is_recursive_constructor (lift+1) rest + | LetIn (na,b,t,rest) -> is_recursive_constructor (lift+1) rest + | _ -> false + in + match mie.mind_entry_inds with + | [ind] -> + let nparams = List.length mie.mind_entry_params in + List.exists (fun t -> is_recursive_constructor (nparams+1) t) ind.mind_entry_lc + | _ -> false + +let declare_mutual_inductive_with_eliminations mie pl impls = + (* spiwack: raises an error if the structure is supposed to be non-recursive, + but isn't *) + begin match mie.mind_entry_finite with + | Declarations.BiFinite when is_recursive mie -> + if Option.has_some mie.mind_entry_record then + user_err Pp.(str "Records declared with the keywords Record or Structure cannot be recursive. You can, however, define recursive records using the Inductive or CoInductive command.") + else + user_err Pp.(str ("Types declared with the keyword Variant cannot be recursive. Recursive types are defined with the Inductive and CoInductive command.")) + | _ -> () + end; + let names = List.map (fun e -> e.mind_entry_typename) mie.mind_entry_inds in + let (_, kn), prim = declare_mind mie in + let mind = Global.mind_of_delta_kn kn in + List.iteri (fun i (indimpls, constrimpls) -> + let ind = (mind,i) in + let gr = IndRef ind in + maybe_declare_manual_implicits false gr indimpls; + Declare.declare_univ_binders gr pl; + List.iteri + (fun j impls -> + maybe_declare_manual_implicits false + (ConstructRef (ind, succ j)) impls) + constrimpls) + impls; + let warn_prim = match mie.mind_entry_record with Some (Some _) -> not prim | _ -> false in + Flags.if_verbose Feedback.msg_info (minductive_message warn_prim names); + if mie.mind_entry_private == None + then declare_default_schemes mind; + mind + +type one_inductive_impls = + Impargs.manual_explicitation list (* for inds *)* + Impargs.manual_explicitation list list (* for constrs *) + +let do_mutual_inductive indl cum poly prv finite = + let indl,coes,ntns = extract_mutual_inductive_declaration_components indl in + (* Interpret the types *) + let mie,pl,impls = interp_mutual_inductive indl ntns cum poly prv finite in + (* Declare the mutual inductive block with its associated schemes *) + ignore (declare_mutual_inductive_with_eliminations mie pl impls); + (* Declare the possible notations of inductive types *) + List.iter (Metasyntax.add_notation_interpretation (Global.env ())) ntns; + (* Declare the coercions *) + List.iter (fun qid -> Class.try_add_new_coercion (locate qid) ~local:false poly) coes; + (* If positivity is assumed declares itself as unsafe. *) + if Environ.deactivated_guard (Global.env ()) then Feedback.feedback Feedback.AddedAxiom else () diff --git a/vernac/comInductive.mli b/vernac/comInductive.mli new file mode 100644 index 00000000..83393572 --- /dev/null +++ b/vernac/comInductive.mli @@ -0,0 +1,67 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* cumulative_inductive_flag -> + polymorphic -> private_flag -> Declarations.recursivity_kind -> unit + +(************************************************************************) +(** Internal API *) +(************************************************************************) + +(** Exported for Record and Funind *) + +(** Registering a mutual inductive definition together with its + associated schemes *) + +type one_inductive_impls = + Impargs.manual_implicits (** for inds *)* + Impargs.manual_implicits list (** for constrs *) + +val declare_mutual_inductive_with_eliminations : + mutual_inductive_entry -> Universes.universe_binders -> one_inductive_impls list -> + MutInd.t + +(** Exported for Funind *) + +(** Extracting the semantical components out of the raw syntax of mutual + inductive declarations *) + +type structured_one_inductive_expr = { + ind_name : Id.t; + ind_univs : universe_decl_expr option; + ind_arity : constr_expr; + ind_lc : (Id.t * constr_expr) list +} + +type structured_inductive_expr = + local_binder_expr list * structured_one_inductive_expr list + +val extract_mutual_inductive_declaration_components : + (one_inductive_expr * decl_notation list) list -> + structured_inductive_expr * (*coercions:*) qualid list * decl_notation list + +(** Typing mutual inductive definitions *) + +val interp_mutual_inductive : + structured_inductive_expr -> decl_notation list -> cumulative_inductive_flag -> + polymorphic -> private_flag -> Declarations.recursivity_kind -> + mutual_inductive_entry * Universes.universe_binders * one_inductive_impls list diff --git a/vernac/comProgramFixpoint.ml b/vernac/comProgramFixpoint.ml new file mode 100644 index 00000000..b95741ca --- /dev/null +++ b/vernac/comProgramFixpoint.ml @@ -0,0 +1,342 @@ +open Pp +open CErrors +open Util +open Constr +open Entries +open Vars +open Declare +open Names +open Libnames +open Globnames +open Nameops +open Constrexpr +open Constrexpr_ops +open Constrintern +open Decl_kinds +open Evarutil +open Context.Rel.Declaration +open ComFixpoint + +module RelDecl = Context.Rel.Declaration + +(* Wellfounded definition *) + +open Coqlib + +let contrib_name = "Program" +let subtac_dir = [contrib_name] +let fixsub_module = subtac_dir @ ["Wf"] +(* let tactics_module = subtac_dir @ ["Tactics"] *) + +let init_reference dir s () = Coqlib.coq_reference "Command" dir s +let init_constant dir s sigma = + Evarutil.new_global sigma (Coqlib.coq_reference "Command" dir s) + +let make_ref l s = init_reference l s +(* let fix_proto = init_constant tactics_module "fix_proto" *) +let fix_sub_ref = make_ref fixsub_module "Fix_sub" +let measure_on_R_ref = make_ref fixsub_module "MR" +let well_founded = init_constant ["Init"; "Wf"] "well_founded" +let mkSubset sigma name typ prop = + let open EConstr in + let sigma, app_h = Evarutil.new_global sigma (delayed_force build_sigma).typ in + sigma, mkApp (app_h, [| typ; mkLambda (name, typ, prop) |]) + +let sigT = Lazy.from_fun build_sigma_type + +let make_qref s = CAst.make @@ Qualid (qualid_of_string s) +let lt_ref = make_qref "Init.Peano.lt" + +let rec telescope sigma l = + let open EConstr in + let open Vars in + match l with + | [] -> assert false + | [LocalAssum (n, t)] -> + sigma, t, [LocalDef (n, mkRel 1, t)], mkRel 1 + | LocalAssum (n, t) :: tl -> + let sigma, ty, tys, (k, constr) = + List.fold_left + (fun (sigma, ty, tys, (k, constr)) decl -> + let t = RelDecl.get_type decl in + let pred = mkLambda (RelDecl.get_name decl, t, ty) in + let sigma, ty = Evarutil.new_global sigma (Lazy.force sigT).typ in + let sigma, intro = Evarutil.new_global sigma (Lazy.force sigT).intro in + let sigty = mkApp (ty, [|t; pred|]) in + let intro = mkApp (intro, [|lift k t; lift k pred; mkRel k; constr|]) in + (sigma, sigty, pred :: tys, (succ k, intro))) + (sigma, t, [], (2, mkRel 1)) tl + in + let sigma, last, subst = List.fold_right2 + (fun pred decl (sigma, prev, subst) -> + let t = RelDecl.get_type decl in + let sigma, p1 = Evarutil.new_global sigma (Lazy.force sigT).proj1 in + let sigma, p2 = Evarutil.new_global sigma (Lazy.force sigT).proj2 in + let proj1 = applist (p1, [t; pred; prev]) in + let proj2 = applist (p2, [t; pred; prev]) in + (sigma, lift 1 proj2, LocalDef (get_name decl, proj1, t) :: subst)) + (List.rev tys) tl (sigma, mkRel 1, []) + in sigma, ty, (LocalDef (n, last, t) :: subst), constr + + | LocalDef (n, b, t) :: tl -> + let sigma, ty, subst, term = telescope sigma tl in + sigma, ty, (LocalDef (n, b, t) :: subst), lift 1 term + +let nf_evar_context sigma ctx = + List.map (map_constr (fun c -> Evarutil.nf_evar sigma c)) ctx + +let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = + let open EConstr in + let open Vars in + let lift_rel_context n l = Termops.map_rel_context_with_binders (liftn n) l in + Coqlib.check_required_library ["Coq";"Program";"Wf"]; + let env = Global.env() in + let sigma, decl = Univdecls.interp_univ_decl_opt env pl in + let sigma, (_, ((env', binders_rel), impls)) = interp_context_evars env sigma bl in + let len = List.length binders_rel in + let top_env = push_rel_context binders_rel env in + let sigma, top_arity = interp_type_evars top_env sigma arityc in + let full_arity = it_mkProd_or_LetIn top_arity binders_rel in + let sigma, argtyp, letbinders, make = telescope sigma binders_rel in + let argname = Id.of_string "recarg" in + let arg = LocalAssum (Name argname, argtyp) in + let binders = letbinders @ [arg] in + let binders_env = push_rel_context binders_rel env in + let sigma, (rel, _) = interp_constr_evars_impls env sigma r in + let relty = Typing.unsafe_type_of env sigma rel in + let relargty = + let error () = + user_err ?loc:(constr_loc r) + ~hdr:"Command.build_wellfounded" + (Printer.pr_econstr_env env sigma rel ++ str " is not an homogeneous binary relation.") + in + try + let ctx, ar = Reductionops.splay_prod_n env sigma 2 relty in + match ctx, EConstr.kind sigma ar with + | [LocalAssum (_,t); LocalAssum (_,u)], Sort s + when Sorts.is_prop (ESorts.kind sigma s) && Reductionops.is_conv env sigma t u -> t + | _, _ -> error () + with e when CErrors.noncritical e -> error () + in + let sigma, measure = interp_casted_constr_evars binders_env sigma measure relargty in + let sigma, wf_rel, wf_rel_fun, measure_fn = + let measure_body, measure = + it_mkLambda_or_LetIn measure letbinders, + it_mkLambda_or_LetIn measure binders + in + let sigma, comb = Evarutil.new_global sigma (delayed_force measure_on_R_ref) in + let wf_rel = mkApp (comb, [| argtyp; relargty; rel; measure |]) in + let wf_rel_fun x y = + mkApp (rel, [| subst1 x measure_body; + subst1 y measure_body |]) + in sigma, wf_rel, wf_rel_fun, measure + in + let sigma, wf_term = well_founded sigma in + let wf_proof = mkApp (wf_term, [| argtyp ; wf_rel |]) in + let argid' = Id.of_string (Id.to_string argname ^ "'") in + let wfarg sigma len = + let sigma, ss_term = mkSubset sigma (Name argid') argtyp (wf_rel_fun (mkRel 1) (mkRel (len + 1))) in + sigma, LocalAssum (Name argid', ss_term) + in + let sigma, intern_bl = + let sigma, wfa = wfarg sigma 1 in + sigma, wfa :: [arg] + in + let _intern_env = push_rel_context intern_bl env in + let sigma, proj = Evarutil.new_global sigma (delayed_force build_sigma).Coqlib.proj1 in + let wfargpred = mkLambda (Name argid', argtyp, wf_rel_fun (mkRel 1) (mkRel 3)) in + let projection = (* in wfarg :: arg :: before *) + mkApp (proj, [| argtyp ; wfargpred ; mkRel 1 |]) + in + let top_arity_let = it_mkLambda_or_LetIn top_arity letbinders in + let intern_arity = substl [projection] top_arity_let in + (* substitute the projection of wfarg for something, + now intern_arity is in wfarg :: arg *) + let sigma, wfa = wfarg sigma 1 in + let intern_fun_arity_prod = it_mkProd_or_LetIn intern_arity [wfa] in + let intern_fun_binder = LocalAssum (Name (add_suffix recname "'"), intern_fun_arity_prod) in + let sigma, curry_fun = + let wfpred = mkLambda (Name argid', argtyp, wf_rel_fun (mkRel 1) (mkRel (2 * len + 4))) in + let sigma, intro = Evarutil.new_global sigma (delayed_force build_sigma).Coqlib.intro in + let arg = mkApp (intro, [| argtyp; wfpred; lift 1 make; mkRel 1 |]) in + let app = mkApp (mkRel (2 * len + 2 (* recproof + orig binders + current binders *)), [| arg |]) in + let rcurry = mkApp (rel, [| measure; lift len measure |]) in + let lam = LocalAssum (Name (Id.of_string "recproof"), rcurry) in + let body = it_mkLambda_or_LetIn app (lam :: binders_rel) in + let ty = it_mkProd_or_LetIn (lift 1 top_arity) (lam :: binders_rel) in + sigma, LocalDef (Name recname, body, ty) + in + let fun_bl = intern_fun_binder :: [arg] in + let lift_lets = lift_rel_context 1 letbinders in + let sigma, intern_body = + let ctx = LocalAssum (Name recname, get_type curry_fun) :: binders_rel in + let (r, l, impls, scopes) = + Constrintern.compute_internalization_data env sigma + Constrintern.Recursive full_arity impls + in + let newimpls = Id.Map.singleton recname + (r, l, impls @ [(Some (Id.of_string "recproof", Impargs.Manual, (true, false)))], + scopes @ [None]) in + interp_casted_constr_evars (push_rel_context ctx env) sigma + ~impls:newimpls body (lift 1 top_arity) + in + let intern_body_lam = it_mkLambda_or_LetIn intern_body (curry_fun :: lift_lets @ fun_bl) in + let prop = mkLambda (Name argname, argtyp, top_arity_let) in + (* XXX: Previous code did parallel evdref update, so possible old + weak ordering semantics may bite here. *) + let sigma, def = + let sigma, h_a_term = Evarutil.new_global sigma (delayed_force fix_sub_ref) in + let sigma, h_e_term = Evarutil.new_evar env sigma + ~src:(Loc.tag @@ Evar_kinds.QuestionMark (Evar_kinds.Define false,Anonymous)) wf_proof in + sigma, mkApp (h_a_term, [| argtyp ; wf_rel ; h_e_term; prop |]) + in + let _evd = ref sigma in + let def = Typing.e_solve_evars env _evd def in + let sigma = !_evd in + let sigma = Evarutil.nf_evar_map sigma in + let def = mkApp (def, [|intern_body_lam|]) in + let binders_rel = nf_evar_context sigma binders_rel in + let binders = nf_evar_context sigma binders in + let top_arity = Evarutil.nf_evar sigma top_arity in + let hook, recname, typ = + if List.length binders_rel > 1 then + let name = add_suffix recname "_func" in + (* XXX: Mutating the evar_map in the hook! *) + (* XXX: Likely the sigma is out of date when the hook is called .... *) + let hook sigma l gr _ = + let sigma, h_body = Evarutil.new_global sigma gr in + let body = it_mkLambda_or_LetIn (mkApp (h_body, [|make|])) binders_rel in + let ty = it_mkProd_or_LetIn top_arity binders_rel in + let ty = EConstr.Unsafe.to_constr ty in + let univs = Evd.check_univ_decl ~poly sigma decl in + (*FIXME poly? *) + let ce = definition_entry ~types:ty ~univs (EConstr.to_constr sigma body) in + (** FIXME: include locality *) + let c = Declare.declare_constant recname (DefinitionEntry ce, IsDefinition Definition) in + let gr = ConstRef c in + let () = Universes.register_universe_binders gr (Evd.universe_binders sigma) in + if Impargs.is_implicit_args () || not (List.is_empty impls) then + Impargs.declare_manual_implicits false gr [impls] + in + let typ = it_mkProd_or_LetIn top_arity binders in + hook, name, typ + else + let typ = it_mkProd_or_LetIn top_arity binders_rel in + let hook sigma l gr _ = + if Impargs.is_implicit_args () || not (List.is_empty impls) then + Impargs.declare_manual_implicits false gr [impls] + in hook, recname, typ + in + (* XXX: Capturing sigma here... bad bad *) + let hook = Lemmas.mk_hook (hook sigma) in + let fullcoqc = EConstr.to_constr sigma def in + let fullctyp = EConstr.to_constr sigma typ in + Obligations.check_evars env sigma; + let evars, _, evars_def, evars_typ = + Obligations.eterm_obligations env recname sigma 0 fullcoqc fullctyp + in + let ctx = Evd.evar_universe_context sigma in + ignore(Obligations.add_definition recname ~term:evars_def ~univdecl:decl + evars_typ ctx evars ~hook) + +let out_def = function + | Some def -> def + | None -> user_err Pp.(str "Program Fixpoint needs defined bodies.") + +let collect_evars_of_term evd c ty = + let evars = Evar.Set.union (Evd.evars_of_term c) (Evd.evars_of_term ty) in + Evar.Set.fold (fun ev acc -> Evd.add acc ev (Evd.find_undefined evd ev)) + evars (Evd.from_ctx (Evd.evar_universe_context evd)) + +let do_program_recursive local poly fixkind fixl ntns = + let cofix = fixkind = Obligations.IsCoFixpoint in + let (env, rec_sign, pl, evd), fix, info = + interp_recursive ~cofix ~program_mode:true fixl ntns + in + (* Program-specific code *) + (* Get the interesting evars, those that were not instanciated *) + let evd = Typeclasses.resolve_typeclasses ~filter:Typeclasses.no_goals ~fail:true env evd in + (* Solve remaining evars *) + let evd = nf_evar_map_undefined evd in + let collect_evars id def typ imps = + (* Generalize by the recursive prototypes *) + let def = + EConstr.to_constr evd (Termops.it_mkNamedLambda_or_LetIn (EConstr.of_constr def) rec_sign) + and typ = + EConstr.to_constr evd (Termops.it_mkNamedProd_or_LetIn (EConstr.of_constr typ) rec_sign) + in + let evm = collect_evars_of_term evd def typ in + let evars, _, def, typ = + Obligations.eterm_obligations env id evm + (List.length rec_sign) def typ + in (id, def, typ, imps, evars) + in + let (fixnames,fixdefs,fixtypes) = fix in + let fiximps = List.map pi2 info in + let fixdefs = List.map out_def fixdefs in + let defs = List.map4 collect_evars fixnames fixdefs fixtypes fiximps in + let () = if not cofix then begin + let possible_indexes = List.map ComFixpoint.compute_possible_guardness_evidences info in + let fixdecls = Array.of_list (List.map (fun x -> Name x) fixnames), + Array.of_list fixtypes, + Array.of_list (List.map (subst_vars (List.rev fixnames)) fixdefs) + in + let indexes = + Pretyping.search_guard (Global.env ()) possible_indexes fixdecls in + List.iteri (fun i _ -> + Inductive.check_fix env + ((indexes,i),fixdecls)) + fixl + end in + let ctx = Evd.evar_universe_context evd in + let kind = match fixkind with + | Obligations.IsFixpoint _ -> (local, poly, Fixpoint) + | Obligations.IsCoFixpoint -> (local, poly, CoFixpoint) + in + Obligations.add_mutual_definitions defs ~kind ~univdecl:pl ctx ntns fixkind + +let do_program_fixpoint local poly l = + let g = List.map (fun ((_,wf,_,_,_),_) -> wf) l in + match g, l with + | [(n, CWfRec r)], [((({CAst.v=id},pl),_,bl,typ,def),ntn)] -> + let recarg = + match n with + | Some n -> mkIdentC n.CAst.v + | None -> + user_err ~hdr:"do_program_fixpoint" + (str "Recursive argument required for well-founded fixpoints") + in build_wellfounded (id, pl, n, bl, typ, out_def def) poly r recarg ntn + + | [(n, CMeasureRec (m, r))], [((({CAst.v=id},pl),_,bl,typ,def),ntn)] -> + build_wellfounded (id, pl, n, bl, typ, out_def def) poly + (Option.default (CAst.make @@ CRef (lt_ref,None)) r) m ntn + + | _, _ when List.for_all (fun (n, ro) -> ro == CStructRec) g -> + let fixl,ntns = extract_fixpoint_components true l in + let fixkind = Obligations.IsFixpoint g in + do_program_recursive local poly fixkind fixl ntns + + | _, _ -> + user_err ~hdr:"do_program_fixpoint" + (str "Well-founded fixpoints not allowed in mutually recursive blocks") + +let extract_cofixpoint_components l = + let fixl, ntnl = List.split l in + List.map (fun (({CAst.v=id},pl),bl,typ,def) -> + {fix_name = id; fix_annot = None; fix_univs = pl; + fix_binders = bl; fix_body = def; fix_type = typ}) fixl, + List.flatten ntnl + +let check_safe () = + let open Declarations in + let flags = Environ.typing_flags (Global.env ()) in + flags.check_universes && flags.check_guarded + +let do_fixpoint local poly l = + do_program_fixpoint local poly l; + if not (check_safe ()) then Feedback.feedback Feedback.AddedAxiom else () + +let do_cofixpoint local poly l = + let fixl,ntns = extract_cofixpoint_components l in + do_program_recursive local poly Obligations.IsCoFixpoint fixl ntns; + if not (check_safe ()) then Feedback.feedback Feedback.AddedAxiom else () diff --git a/vernac/comProgramFixpoint.mli b/vernac/comProgramFixpoint.mli new file mode 100644 index 00000000..943cb8ef --- /dev/null +++ b/vernac/comProgramFixpoint.mli @@ -0,0 +1,12 @@ +open Decl_kinds +open Vernacexpr + +(** Special Fixpoint handling when command is activated. *) + +val do_fixpoint : + (* When [false], assume guarded. *) + locality -> polymorphic -> (fixpoint_expr * decl_notation list) list -> unit + +val do_cofixpoint : + (* When [false], assume guarded. *) + locality -> polymorphic -> (cofixpoint_expr * decl_notation list) list -> unit diff --git a/vernac/declareDef.ml b/vernac/declareDef.ml new file mode 100644 index 00000000..77177dfa --- /dev/null +++ b/vernac/declareDef.ml @@ -0,0 +1,66 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* + strbrk "Section definition " ++ + Names.Id.print ident ++ strbrk " is not visible from current goals") + +let warn_local_declaration = + CWarnings.create ~name:"local-declaration" ~category:"scope" + Pp.(fun (id,kind) -> + Names.Id.print id ++ strbrk " is declared as a local " ++ str kind) + +let get_locality id ~kind = function +| Discharge -> + (** If a Let is defined outside a section, then we consider it as a local definition *) + warn_local_declaration (id,kind); + true +| Local -> true +| Global -> false + +let declare_global_definition ident ce local k pl imps = + let local = get_locality ident ~kind:"definition" local in + let kn = declare_constant ident ~local (DefinitionEntry ce, IsDefinition k) in + let gr = ConstRef kn in + let () = maybe_declare_manual_implicits false gr imps in + let () = Declare.declare_univ_binders gr pl in + let () = definition_message ident in + gr + +let declare_definition ident (local, p, k) ce pl imps hook = + let fix_exn = Future.fix_exn_of ce.const_entry_body in + let r = match local with + | Discharge when Lib.sections_are_opened () -> + let c = SectionLocalDef ce in + let _ = declare_variable ident (Lib.cwd(), c, IsDefinition k) in + let () = definition_message ident in + let gr = VarRef ident in + let () = maybe_declare_manual_implicits false gr imps in + let () = Declare.declare_univ_binders gr pl in + let () = if Proof_global.there_are_pending_proofs () then + warn_definition_not_visible ident + in + gr + | Discharge | Local | Global -> + declare_global_definition ident ce local k pl imps in + Lemmas.call_hook fix_exn hook local r + +let declare_fix ?(opaque = false) (_,poly,_ as kind) pl univs f ((def,_),eff) t imps = + let ce = definition_entry ~opaque ~types:t ~univs ~eff def in + declare_definition f kind ce pl imps (Lemmas.mk_hook (fun _ r -> r)) + diff --git a/vernac/declareDef.mli b/vernac/declareDef.mli new file mode 100644 index 00000000..010874e2 --- /dev/null +++ b/vernac/declareDef.mli @@ -0,0 +1,24 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* kind:string -> Decl_kinds.locality -> bool + +val declare_definition : Id.t -> definition_kind -> + Safe_typing.private_constants Entries.definition_entry -> Universes.universe_binders -> Impargs.manual_implicits -> + Globnames.global_reference Lemmas.declaration_hook -> Globnames.global_reference + +val declare_fix : ?opaque:bool -> definition_kind -> + Universes.universe_binders -> Entries.constant_universes_entry -> + Id.t -> Safe_typing.private_constants Entries.proof_output -> + Constr.types -> Impargs.manual_implicits -> + Globnames.global_reference diff --git a/vernac/explainErr.ml b/vernac/explainErr.ml new file mode 100644 index 00000000..f9167f96 --- /dev/null +++ b/vernac/explainErr.ml @@ -0,0 +1,128 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* hov 0 (str "Syntax error: " ++ str txt ++ str ".") + | Token.Error txt -> hov 0 (str "Syntax error: " ++ str txt ++ str ".") + | CLexer.Error.E err -> hov 0 (str (CLexer.Error.to_string err)) + | Sys_error msg -> hov 0 (str "System error: " ++ guill msg) + | Out_of_memory -> hov 0 (str "Out of memory.") + | Stack_overflow -> hov 0 (str "Stack overflow.") + | Dynlink.Error e -> hov 0 (str "Dynlink error: " ++ str Dynlink.(error_message e)) + | Timeout -> hov 0 (str "Timeout!") + | Sys.Break -> hov 0 (fnl () ++ str "User interrupt.") + (* Exceptions with pre-evaluated error messages *) + | EvaluatedError (msg,None) -> msg + | EvaluatedError (msg,Some reraise) -> msg ++ CErrors.print reraise + (* Otherwise, not handled here *) + | _ -> raise CErrors.Unhandled + +let _ = CErrors.register_handler explain_exn_default + + +(** Pre-explain a vernac interpretation error *) + +let wrap_vernac_error (exn, info) strm = (EvaluatedError (strm, None), info) + +let process_vernac_interp_error exn = match fst exn with + | Univ.UniverseInconsistency i -> + let msg = + if !Constrextern.print_universes then + str "." ++ spc() ++ + Univ.explain_universe_inconsistency Universes.pr_with_global_universes i + else + mt() in + wrap_vernac_error exn (str "Universe inconsistency" ++ msg ++ str ".") + | TypeError(ctx,te) -> + let te = Himsg.map_ptype_error EConstr.of_constr te in + wrap_vernac_error exn (Himsg.explain_type_error ctx Evd.empty te) + | PretypeError(ctx,sigma,te) -> + wrap_vernac_error exn (Himsg.explain_pretype_error ctx sigma te) + | Typeclasses_errors.TypeClassError(env, te) -> + wrap_vernac_error exn (Himsg.explain_typeclass_error env te) + | InductiveError e -> + wrap_vernac_error exn (Himsg.explain_inductive_error e) + | Modops.ModuleTypingError e -> + wrap_vernac_error exn (Himsg.explain_module_error e) + | Modintern.ModuleInternalizationError e -> + wrap_vernac_error exn (Himsg.explain_module_internalization_error e) + | RecursionSchemeError e -> + wrap_vernac_error exn (Himsg.explain_recursion_scheme_error e) + | Cases.PatternMatchingError (env,sigma,e) -> + wrap_vernac_error exn (Himsg.explain_pattern_matching_error env sigma e) + | Tacred.ReductionTacticError e -> + wrap_vernac_error exn (Himsg.explain_reduction_tactic_error e) + | Logic.RefinerError (env, sigma, e) -> + wrap_vernac_error exn (Himsg.explain_refiner_error env sigma e) + | Nametab.GlobalizationError q -> + wrap_vernac_error exn + (str "The reference" ++ spc () ++ Libnames.pr_qualid q ++ + spc () ++ str "was not found" ++ + spc () ++ str "in the current" ++ spc () ++ str "environment.") + | Refiner.FailError (i,s) -> + let s = Lazy.force s in + wrap_vernac_error exn + (str "Tactic failure" ++ + (if Pp.ismt s then s else str ": " ++ s) ++ + if Int.equal i 0 then str "." else str " (level " ++ int i ++ str").") + | AlreadyDeclared msg -> + wrap_vernac_error exn (msg ++ str ".") + | _ -> + exn + +let rec strip_wrapping_exceptions = function + | Logic_monad.TacticFailure e -> + strip_wrapping_exceptions e + | exc -> exc + +let additional_error_info = ref [] + +let register_additional_error_info f = + additional_error_info := f :: !additional_error_info + +let process_vernac_interp_error ?(allow_uncaught=true) (exc, info) = + let exc = strip_wrapping_exceptions exc in + let e = process_vernac_interp_error (exc, info) in + let () = + if not allow_uncaught && not (CErrors.handled (fst e)) then + let (e, info) = e in + let msg = str "Uncaught exception " ++ str (Printexc.to_string e) ++ str "." in + let err = CErrors.make_anomaly msg in + Util.iraise (err, info) + in + let e' = + try Some (CList.find_map (fun f -> f e) !additional_error_info) + with _ -> None + in + let add_loc_opt ?loc info = Option.cata (fun l -> Loc.add_loc info l) info loc in + match e' with + | None -> e + | Some (loc, None) -> (fst e, add_loc_opt ?loc (snd e)) + | Some (loc, Some msg) -> + (EvaluatedError (msg, Some (fst e)), add_loc_opt ?loc (snd e)) diff --git a/vernac/explainErr.mli b/vernac/explainErr.mli new file mode 100644 index 00000000..b54912a1 --- /dev/null +++ b/vernac/explainErr.mli @@ -0,0 +1,23 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* Util.iexn -> Util.iexn + +(** General explain function. Should not be used directly now, + see instead function [Errors.print] and variants *) + +val explain_exn_default : exn -> Pp.t + +val register_additional_error_info : (Util.iexn -> (Pp.t option Loc.located) option) -> unit diff --git a/vernac/himsg.ml b/vernac/himsg.ml new file mode 100644 index 00000000..2d2f1cd2 --- /dev/null +++ b/vernac/himsg.ml @@ -0,0 +1,1317 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* + l := (Vars.substl !l c') :: !l; + env + | _ -> + let t = Vars.substl !l (RelDecl.get_type decl) in + let decl = decl |> RelDecl.map_name (named_hd env sigma t) |> RelDecl.map_value (Vars.substl !l) |> RelDecl.set_type t in + l := (mkRel 1) :: List.map (Vars.lift 1) !l; + push_rel decl env + in + let env = process_rel_context contract_context env in + (env, List.map (Vars.substl !l) lc) + +let contract2 env sigma a b = match contract env sigma [a;b] with + | env, [a;b] -> env,a,b | _ -> assert false + +let contract3 env sigma a b c = match contract env sigma [a;b;c] with + | env, [a;b;c] -> env,a,b,c | _ -> assert false + +let contract4 env sigma a b c d = match contract env sigma [a;b;c;d] with + | env, [a;b;c;d] -> (env,a,b,c),d | _ -> assert false + +let contract1_vect env sigma a v = + match contract env sigma (a :: Array.to_list v) with + | env, a::l -> env,a,Array.of_list l + | _ -> assert false + +let rec contract3' env sigma a b c = function + | OccurCheck (evk,d) -> + let x,d = contract4 env sigma a b c d in x,OccurCheck(evk, d) + | NotClean ((evk,args),env',d) -> + let env',d,args = contract1_vect env' sigma d args in + contract3 env sigma a b c,NotClean((evk,args),env',d) + | ConversionFailed (env',t1,t2) -> + let (env',t1,t2) = contract2 env' sigma t1 t2 in + contract3 env sigma a b c, ConversionFailed (env',t1,t2) + | NotSameArgSize | NotSameHead | NoCanonicalStructure + | MetaOccurInBody _ | InstanceNotSameType _ | ProblemBeyondCapabilities + | UnifUnivInconsistency _ as x -> contract3 env sigma a b c, x + | CannotSolveConstraint ((pb,env',t,u),x) -> + let t = EConstr.of_constr t in + let u = EConstr.of_constr u in + let env',t,u = contract2 env' sigma t u in + let t = EConstr.Unsafe.to_constr t in + let u = EConstr.Unsafe.to_constr u in + let y,x = contract3' env sigma a b c x in + y,CannotSolveConstraint ((pb,env',t,u),x) + +(** Ad-hoc reductions *) + +let j_nf_betaiotaevar env sigma j = + { uj_val = j.uj_val; + uj_type = Reductionops.nf_betaiota env sigma j.uj_type } + +let jv_nf_betaiotaevar env sigma jl = + Array.map (fun j -> j_nf_betaiotaevar env sigma j) jl + +(** Printers *) + +let pr_lconstr_env e s c = quote (pr_lconstr_env e s c) +let pr_leconstr_env e s c = quote (pr_leconstr_env e s c) +let pr_ljudge_env e s c = let v,t = pr_ljudge_env e s c in (quote v,quote t) + +(** A canonisation procedure for constr such that comparing there + externalisation catches more equalities *) +let canonize_constr sigma c = + (* replaces all the names in binders by [dn] ("default name"), + ensures that [alpha]-equivalent terms will have the same + externalisation. *) + let open EConstr in + let dn = Name.Anonymous in + let rec canonize_binders c = + match EConstr.kind sigma c with + | Prod (_,t,b) -> mkProd(dn,t,b) + | Lambda (_,t,b) -> mkLambda(dn,t,b) + | LetIn (_,u,t,b) -> mkLetIn(dn,u,t,b) + | _ -> EConstr.map sigma canonize_binders c + in + canonize_binders c + +(** Tries to realize when the two terms, albeit different are printed the same. *) +let display_eq ~flags env sigma t1 t2 = + (* terms are canonized, then their externalisation is compared syntactically *) + let open Constrextern in + let t1 = canonize_constr sigma t1 in + let t2 = canonize_constr sigma t2 in + let ct1 = Flags.with_options flags (fun () -> extern_constr false env sigma t1) () in + let ct2 = Flags.with_options flags (fun () -> extern_constr false env sigma t2) () in + Constrexpr_ops.constr_expr_eq ct1 ct2 + +(** This function adds some explicit printing flags if the two arguments are + printed alike. *) +let rec pr_explicit_aux env sigma t1 t2 = function +| [] -> + (** no specified flags: default. *) + (quote (Printer.pr_leconstr_env env sigma t1), quote (Printer.pr_leconstr_env env sigma t2)) +| flags :: rem -> + let equal = display_eq ~flags env sigma t1 t2 in + if equal then + (** The two terms are the same from the user point of view *) + pr_explicit_aux env sigma t1 t2 rem + else + let open Constrextern in + let ct1 = Flags.with_options flags (fun () -> extern_constr false env sigma t1) () + in + let ct2 = Flags.with_options flags (fun () -> extern_constr false env sigma t2) () + in + quote (Ppconstr.pr_lconstr_expr ct1), quote (Ppconstr.pr_lconstr_expr ct2) + +let explicit_flags = + let open Constrextern in + [ []; (** First, try with the current flags *) + [print_implicits]; (** Then with implicit *) + [print_universes]; (** Then with universes *) + [print_universes; print_implicits]; (** With universes AND implicits *) + [print_implicits; print_coercions; print_no_symbol]; (** Then more! *) + [print_universes; print_implicits; print_coercions; print_no_symbol] (** and more! *) ] + +let pr_explicit env sigma t1 t2 = + pr_explicit_aux env sigma t1 t2 explicit_flags + +let pr_db env i = + try + match env |> lookup_rel i |> get_name with + | Name id -> Id.print id + | Anonymous -> str "<>" + with Not_found -> str "UNBOUND_REL_" ++ int i + +let explain_unbound_rel env sigma n = + let pe = pr_ne_context_of (str "In environment") env sigma in + str "Unbound reference: " ++ pe ++ + str "The reference " ++ int n ++ str " is free." + +let explain_unbound_var env v = + let var = Id.print v in + str "No such section variable or assumption: " ++ var ++ str "." + +let explain_not_type env sigma j = + let pe = pr_ne_context_of (str "In environment") env sigma in + let pc,pt = pr_ljudge_env env sigma j in + pe ++ str "The term" ++ brk(1,1) ++ pc ++ spc () ++ + str "has type" ++ spc () ++ pt ++ spc () ++ + str "which should be Set, Prop or Type." + +let explain_bad_assumption env sigma j = + let pe = pr_ne_context_of (str "In environment") env sigma in + let pc,pt = pr_ljudge_env env sigma j in + pe ++ str "Cannot declare a variable or hypothesis over the term" ++ + brk(1,1) ++ pc ++ spc () ++ str "of type" ++ spc () ++ pt ++ spc () ++ + str "because this term is not a type." + +let explain_reference_variables sigma id c = + (* c is intended to be a global reference *) + let pc = pr_global (fst (Termops.global_of_constr sigma c)) in + pc ++ strbrk " depends on the variable " ++ Id.print id ++ + strbrk " which is not declared in the context." + +let rec pr_disjunction pr = function + | [a] -> pr a + | [a;b] -> pr a ++ str " or" ++ spc () ++ pr b + | a::l -> pr a ++ str "," ++ spc () ++ pr_disjunction pr l + | [] -> assert false + +let pr_puniverses f env (c,u) = + f env c ++ + (if Flags.is_universe_polymorphism () && not (Univ.Instance.is_empty u) then + str"(*" ++ Univ.Instance.pr Universes.pr_with_global_universes u ++ str"*)" + else mt()) + +let explain_elim_arity env sigma ind sorts c pj okinds = + let open EConstr in + let env = make_all_name_different env sigma in + let pi = pr_inductive env (fst ind) in + let pc = pr_leconstr_env env sigma c in + let msg = match okinds with + | Some(kp,ki,explanation) -> + let pki = pr_sort_family ki in + let pkp = pr_sort_family kp in + let explanation = match explanation with + | NonInformativeToInformative -> + "proofs can be eliminated only to build proofs" + | StrongEliminationOnNonSmallType -> + "strong elimination on non-small inductive types leads to paradoxes" + | WrongArity -> + "wrong arity" in + let ppar = pr_disjunction (fun s -> quote (pr_sort_family s)) sorts in + let ppt = pr_leconstr_env env sigma (snd (decompose_prod_assum sigma pj.uj_type)) in + hov 0 + (str "the return type has sort" ++ spc () ++ ppt ++ spc () ++ + str "while it" ++ spc () ++ str "should be " ++ ppar ++ str ".") ++ + fnl () ++ + hov 0 + (str "Elimination of an inductive object of sort " ++ + pki ++ brk(1,0) ++ + str "is not allowed on a predicate in sort " ++ pkp ++ fnl () ++ + str "because" ++ spc () ++ str explanation ++ str ".") + | None -> + str "ill-formed elimination predicate." + in + hov 0 ( + str "Incorrect elimination of" ++ spc () ++ pc ++ spc () ++ + str "in the inductive type" ++ spc () ++ quote pi ++ str ":") ++ + fnl () ++ msg + +let explain_case_not_inductive env sigma cj = + let env = make_all_name_different env sigma in + let pc = pr_leconstr_env env sigma cj.uj_val in + let pct = pr_leconstr_env env sigma cj.uj_type in + match EConstr.kind sigma cj.uj_type with + | Evar _ -> + str "Cannot infer a type for this expression." + | _ -> + str "The term" ++ brk(1,1) ++ pc ++ spc () ++ + str "has type" ++ brk(1,1) ++ pct ++ spc () ++ + str "which is not a (co-)inductive type." + +let explain_number_branches env sigma cj expn = + let env = make_all_name_different env sigma in + let pc = pr_leconstr_env env sigma cj.uj_val in + let pct = pr_leconstr_env env sigma cj.uj_type in + str "Matching on term" ++ brk(1,1) ++ pc ++ spc () ++ + str "of type" ++ brk(1,1) ++ pct ++ spc () ++ + str "expects " ++ int expn ++ str " branches." + +let explain_ill_formed_branch env sigma c ci actty expty = + let simp t = Reductionops.nf_betaiota env sigma t in + let env = make_all_name_different env sigma in + let pc = pr_leconstr_env env sigma c in + let pa, pe = pr_explicit env sigma (simp actty) (simp expty) in + strbrk "In pattern-matching on term" ++ brk(1,1) ++ pc ++ + spc () ++ strbrk "the branch for constructor" ++ spc () ++ + quote (pr_puniverses pr_constructor env ci) ++ + spc () ++ str "has type" ++ brk(1,1) ++ pa ++ spc () ++ + str "which should be" ++ brk(1,1) ++ pe ++ str "." + +let explain_generalization env sigma (name,var) j = + let pe = pr_ne_context_of (str "In environment") env sigma in + let pv = pr_letype_env env sigma var in + let (pc,pt) = pr_ljudge_env (push_rel_assum (name,var) env) sigma j in + pe ++ str "Cannot generalize" ++ brk(1,1) ++ pv ++ spc () ++ + str "over" ++ brk(1,1) ++ pc ++ str "," ++ spc () ++ + str "it has type" ++ spc () ++ pt ++ + spc () ++ str "which should be Set, Prop or Type." + +let explain_unification_error env sigma p1 p2 = function + | None -> mt() + | Some e -> + let rec aux p1 p2 = function + | OccurCheck (evk,rhs) -> + [str "cannot define " ++ quote (pr_existential_key sigma evk) ++ + strbrk " with term " ++ pr_leconstr_env env sigma rhs ++ + strbrk " that would depend on itself"] + | NotClean ((evk,args),env,c) -> + [str "cannot instantiate " ++ quote (pr_existential_key sigma evk) + ++ strbrk " because " ++ pr_leconstr_env env sigma c ++ + strbrk " is not in its scope" ++ + (if Array.is_empty args then mt() else + strbrk ": available arguments are " ++ + pr_sequence (pr_leconstr_env env sigma) (List.rev (Array.to_list args)))] + | NotSameArgSize | NotSameHead | NoCanonicalStructure -> + (* Error speaks from itself *) [] + | ConversionFailed (env,t1,t2) -> + let t1 = Reductionops.nf_betaiota env sigma t1 in + let t2 = Reductionops.nf_betaiota env sigma t2 in + if EConstr.eq_constr sigma t1 p1 && EConstr.eq_constr sigma t2 p2 then [] else + let env = make_all_name_different env sigma in + if not (EConstr.eq_constr sigma t1 p1) || not (EConstr.eq_constr sigma t2 p2) then + let t1, t2 = pr_explicit env sigma t1 t2 in + [str "cannot unify " ++ t1 ++ strbrk " and " ++ t2] + else [] + | MetaOccurInBody evk -> + [str "instance for " ++ quote (pr_existential_key sigma evk) ++ + strbrk " refers to a metavariable - please report your example" ++ + strbrk "at " ++ str Coq_config.wwwbugtracker ++ str "."] + | InstanceNotSameType (evk,env,t,u) -> + let t, u = pr_explicit env sigma t u in + [str "unable to find a well-typed instantiation for " ++ + quote (pr_existential_key sigma evk) ++ + strbrk ": cannot ensure that " ++ + t ++ strbrk " is a subtype of " ++ u] + | UnifUnivInconsistency p -> + if !Constrextern.print_universes then + [str "universe inconsistency: " ++ + Univ.explain_universe_inconsistency Universes.pr_with_global_universes p] + else + [str "universe inconsistency"] + | CannotSolveConstraint ((pb,env,t,u),e) -> + let t = EConstr.of_constr t in + let u = EConstr.of_constr u in + let env = make_all_name_different env sigma in + (strbrk "cannot satisfy constraint " ++ pr_leconstr_env env sigma t ++ + str " == " ++ pr_leconstr_env env sigma u) + :: aux t u e + | ProblemBeyondCapabilities -> + [] + in + match aux p1 p2 e with + | [] -> mt () + | l -> spc () ++ str "(" ++ + prlist_with_sep pr_semicolon (fun x -> x) l ++ str ")" + +let explain_actual_type env sigma j t reason = + let env = make_all_name_different env sigma in + let j = j_nf_betaiotaevar env sigma j in + let t = Reductionops.nf_betaiota env sigma t in + (** Actually print *) + let pe = pr_ne_context_of (str "In environment") env sigma in + let pc = pr_leconstr_env env sigma (Environ.j_val j) in + let (pt, pct) = pr_explicit env sigma t (Environ.j_type j) in + let ppreason = explain_unification_error env sigma j.uj_type t reason in + pe ++ + hov 0 ( + str "The term" ++ brk(1,1) ++ pc ++ spc () ++ + str "has type" ++ brk(1,1) ++ pct ++ spc () ++ + str "while it is expected to have type" ++ brk(1,1) ++ pt ++ + ppreason ++ str ".") + +let explain_cant_apply_bad_type env sigma (n,exptyp,actualtyp) rator randl = + let randl = jv_nf_betaiotaevar env sigma randl in + let actualtyp = Reductionops.nf_betaiota env sigma actualtyp in + let env = make_all_name_different env sigma in + let actualtyp, exptyp = pr_explicit env sigma actualtyp exptyp in + let nargs = Array.length randl in +(* let pe = pr_ne_context_of (str "in environment") env sigma in*) + let pr,prt = pr_ljudge_env env sigma rator in + let term_string1 = str (String.plural nargs "term") in + let term_string2 = + if nargs>1 then str "The " ++ pr_nth n ++ str " term" else str "This term" + in + let appl = prvect_with_sep fnl + (fun c -> + let pc,pct = pr_ljudge_env env sigma c in + hov 2 (pc ++ spc () ++ str ": " ++ pct)) randl + in + str "Illegal application: " ++ (* pe ++ *) fnl () ++ + str "The term" ++ brk(1,1) ++ pr ++ spc () ++ + str "of type" ++ brk(1,1) ++ prt ++ spc () ++ + str "cannot be applied to the " ++ term_string1 ++ fnl () ++ + str " " ++ v 0 appl ++ fnl () ++ term_string2 ++ str " has type" ++ + brk(1,1) ++ actualtyp ++ spc () ++ + str "which should be coercible to" ++ brk(1,1) ++ + exptyp ++ str "." + +let explain_cant_apply_not_functional env sigma rator randl = + let env = make_all_name_different env sigma in + let nargs = Array.length randl in +(* let pe = pr_ne_context_of (str "in environment") env sigma in*) + let pr = pr_leconstr_env env sigma rator.uj_val in + let prt = pr_leconstr_env env sigma rator.uj_type in + let appl = prvect_with_sep fnl + (fun c -> + let pc = pr_leconstr_env env sigma c.uj_val in + let pct = pr_leconstr_env env sigma c.uj_type in + hov 2 (pc ++ spc () ++ str ": " ++ pct)) randl + in + str "Illegal application (Non-functional construction): " ++ + (* pe ++ *) fnl () ++ + str "The expression" ++ brk(1,1) ++ pr ++ spc () ++ + str "of type" ++ brk(1,1) ++ prt ++ spc () ++ + str "cannot be applied to the " ++ str (String.plural nargs "term") ++ + fnl () ++ str " " ++ v 0 appl + +let explain_unexpected_type env sigma actual_type expected_type = + let pract, prexp = pr_explicit env sigma actual_type expected_type in + str "Found type" ++ spc () ++ pract ++ spc () ++ + str "where" ++ spc () ++ prexp ++ str " was expected." + +let explain_not_product env sigma c = + let c = EConstr.to_constr sigma c in + let pr = pr_lconstr_env env sigma c in + str "The type of this term is a product" ++ spc () ++ + str "while it is expected to be" ++ + (if Constr.is_Type c then str " a sort" else (brk(1,1) ++ pr)) ++ str "." + +(* TODO: use the names *) +(* (co)fixpoints *) +let explain_ill_formed_rec_body env sigma err names i fixenv vdefj = + let pr_lconstr_env env sigma c = pr_leconstr_env env sigma c in + let prt_name i = + match names.(i) with + Name id -> str "Recursive definition of " ++ Id.print id + | Anonymous -> str "The " ++ pr_nth i ++ str " definition" in + + let st = match err with + + (* Fixpoint guard errors *) + | NotEnoughAbstractionInFixBody -> + str "Not enough abstractions in the definition" + | RecursionNotOnInductiveType c -> + str "Recursive definition on" ++ spc () ++ pr_lconstr_env env sigma c ++ + spc () ++ str "which should be a recursive inductive type" + | RecursionOnIllegalTerm(j,(arg_env, arg),le,lt) -> + let arg_env = make_all_name_different arg_env sigma in + let called = + match names.(j) with + Name id -> Id.print id + | Anonymous -> str "the " ++ pr_nth i ++ str " definition" in + let pr_db x = quote (pr_db env x) in + let vars = + match (lt,le) with + ([],[]) -> assert false + | ([],[x]) -> str "a subterm of " ++ pr_db x + | ([],_) -> str "a subterm of the following variables: " ++ + pr_sequence pr_db le + | ([x],_) -> pr_db x + | _ -> + str "one of the following variables: " ++ + pr_sequence pr_db lt in + str "Recursive call to " ++ called ++ spc () ++ + strbrk "has principal argument equal to" ++ spc () ++ + pr_lconstr_env arg_env sigma arg ++ strbrk " instead of " ++ vars + + | NotEnoughArgumentsForFixCall j -> + let called = + match names.(j) with + Name id -> Id.print id + | Anonymous -> str "the " ++ pr_nth i ++ str " definition" in + str "Recursive call to " ++ called ++ str " has not enough arguments" + + (* CoFixpoint guard errors *) + | CodomainNotInductiveType c -> + str "The codomain is" ++ spc () ++ pr_lconstr_env env sigma c ++ spc () ++ + str "which should be a coinductive type" + | NestedRecursiveOccurrences -> + str "Nested recursive occurrences" + | UnguardedRecursiveCall c -> + str "Unguarded recursive call in" ++ spc () ++ pr_lconstr_env env sigma c + | RecCallInTypeOfAbstraction c -> + str "Recursive call forbidden in the domain of an abstraction:" ++ + spc () ++ pr_lconstr_env env sigma c + | RecCallInNonRecArgOfConstructor c -> + str "Recursive call on a non-recursive argument of constructor" ++ + spc () ++ pr_lconstr_env env sigma c + | RecCallInTypeOfDef c -> + str "Recursive call forbidden in the type of a recursive definition" ++ + spc () ++ pr_lconstr_env env sigma c + | RecCallInCaseFun c -> + str "Invalid recursive call in a branch of" ++ + spc () ++ pr_lconstr_env env sigma c + | RecCallInCaseArg c -> + str "Invalid recursive call in the argument of \"match\" in" ++ spc () ++ + pr_lconstr_env env sigma c + | RecCallInCasePred c -> + str "Invalid recursive call in the \"return\" clause of \"match\" in" ++ + spc () ++ pr_lconstr_env env sigma c + | NotGuardedForm c -> + str "Sub-expression " ++ pr_lconstr_env env sigma c ++ + strbrk " not in guarded form (should be a constructor," ++ + strbrk " an abstraction, a match, a cofix or a recursive call)" + | ReturnPredicateNotCoInductive c -> + str "The return clause of the following pattern matching should be" ++ + strbrk " a coinductive type:" ++ + spc () ++ pr_lconstr_env env sigma c + in + prt_name i ++ str " is ill-formed." ++ fnl () ++ + pr_ne_context_of (str "In environment") env sigma ++ + st ++ str "." ++ fnl () ++ + (try (* May fail with unresolved globals. *) + let fixenv = make_all_name_different fixenv sigma in + let pvd = pr_lconstr_env fixenv sigma vdefj.(i).uj_val in + str"Recursive definition is:" ++ spc () ++ pvd ++ str "." + with e when CErrors.noncritical e -> mt ()) + +let explain_ill_typed_rec_body env sigma i names vdefj vargs = + let env = make_all_name_different env sigma in + let pvd = pr_leconstr_env env sigma vdefj.(i).uj_val in + let pvdt, pv = pr_explicit env sigma vdefj.(i).uj_type vargs.(i) in + str "The " ++ + (match vdefj with [|_|] -> mt () | _ -> pr_nth (i+1) ++ spc ()) ++ + str "recursive definition" ++ spc () ++ pvd ++ spc () ++ + str "has type" ++ spc () ++ pvdt ++ spc () ++ + str "while it should be" ++ spc () ++ pv ++ str "." + +let explain_cant_find_case_type env sigma c = + let env = make_all_name_different env sigma in + let pe = pr_leconstr_env env sigma c in + str "Cannot infer the return type of pattern-matching on" ++ ws 1 ++ + pe ++ str "." + +let explain_occur_check env sigma ev rhs = + let env = make_all_name_different env sigma in + let pt = pr_leconstr_env env sigma rhs in + str "Cannot define " ++ pr_existential_key sigma ev ++ str " with term" ++ + brk(1,1) ++ pt ++ spc () ++ str "that would depend on itself." + +let pr_trailing_ne_context_of env sigma = + if List.is_empty (Environ.rel_context env) && + List.is_empty (Environ.named_context env) + then str "." + else (str " in environment:"++ pr_context_unlimited env sigma) + +let rec explain_evar_kind env sigma evk ty = function + | Evar_kinds.NamedHole id -> + strbrk "the existential variable named " ++ Id.print id + | Evar_kinds.QuestionMark _ -> + strbrk "this placeholder of type " ++ ty + | Evar_kinds.CasesType false -> + strbrk "the type of this pattern-matching problem" + | Evar_kinds.CasesType true -> + strbrk "a subterm of type " ++ ty ++ + strbrk " in the type of this pattern-matching problem" + | Evar_kinds.BinderType (Name id) -> + strbrk "the type of " ++ Id.print id + | Evar_kinds.BinderType Anonymous -> + strbrk "the type of this anonymous binder" + | Evar_kinds.ImplicitArg (c,(n,ido),b) -> + let id = Option.get ido in + strbrk "the implicit parameter " ++ Id.print id ++ spc () ++ str "of" ++ + spc () ++ Nametab.pr_global_env Id.Set.empty c ++ + strbrk " whose type is " ++ ty + | Evar_kinds.InternalHole -> strbrk "an internal placeholder of type " ++ ty + | Evar_kinds.TomatchTypeParameter (tyi,n) -> + strbrk "the " ++ pr_nth n ++ + strbrk " argument of the inductive type (" ++ pr_inductive env tyi ++ + strbrk ") of this term" + | Evar_kinds.GoalEvar -> + strbrk "an existential variable of type " ++ ty + | Evar_kinds.ImpossibleCase -> + strbrk "the type of an impossible pattern-matching clause" + | Evar_kinds.MatchingVar _ -> + assert false + | Evar_kinds.VarInstance id -> + strbrk "an instance of type " ++ ty ++ + str " for the variable " ++ Id.print id + | Evar_kinds.SubEvar evk' -> + let evi = Evd.find sigma evk' in + let pc = match evi.evar_body with + | Evar_defined c -> pr_leconstr_env env sigma (EConstr.of_constr c) + | Evar_empty -> assert false in + let ty' = EConstr.of_constr evi.evar_concl in + pr_existential_key sigma evk ++ str " of type " ++ ty ++ + str " in the partial instance " ++ pc ++ + str " found for " ++ explain_evar_kind env sigma evk' + (pr_leconstr_env env sigma ty') (snd evi.evar_source) + +let explain_typeclass_resolution env sigma evi k = + match Typeclasses.class_of_constr sigma (EConstr.of_constr evi.evar_concl) with + | Some _ -> + let env = Evd.evar_filtered_env evi in + fnl () ++ str "Could not find an instance for " ++ + pr_lconstr_env env sigma evi.evar_concl ++ + pr_trailing_ne_context_of env sigma + | _ -> mt() + +let explain_placeholder_kind env sigma c e = + match e with + | Some (SeveralInstancesFound n) -> + strbrk " (several distinct possible type class instances found)" + | None -> + match Typeclasses.class_of_constr sigma (EConstr.of_constr c) with + | Some _ -> strbrk " (no type class instance found)" + | _ -> mt () + +let explain_unsolvable_implicit env sigma evk explain = + let evi = Evarutil.nf_evar_info sigma (Evd.find_undefined sigma evk) in + let env = Evd.evar_filtered_env evi in + let type_of_hole = pr_lconstr_env env sigma evi.evar_concl in + let pe = pr_trailing_ne_context_of env sigma in + strbrk "Cannot infer " ++ + explain_evar_kind env sigma evk type_of_hole (snd evi.evar_source) ++ + explain_placeholder_kind env sigma evi.evar_concl explain ++ pe + +let explain_var_not_found env id = + str "The variable" ++ spc () ++ Id.print id ++ + spc () ++ str "was not found" ++ + spc () ++ str "in the current" ++ spc () ++ str "environment" ++ str "." + +let explain_wrong_case_info env (ind,u) ci = + let pi = pr_inductive (Global.env()) ind in + if eq_ind ci.ci_ind ind then + str "Pattern-matching expression on an object of inductive type" ++ + spc () ++ pi ++ spc () ++ str "has invalid information." + else + let pc = pr_inductive (Global.env()) ci.ci_ind in + str "A term of inductive type" ++ spc () ++ pi ++ spc () ++ + str "was given to a pattern-matching expression on the inductive type" ++ + spc () ++ pc ++ str "." + +let explain_cannot_unify env sigma m n e = + let env = make_all_name_different env sigma in + let pm, pn = pr_explicit env sigma m n in + let ppreason = explain_unification_error env sigma m n e in + let pe = pr_ne_context_of (str "In environment") env sigma in + pe ++ str "Unable to unify" ++ brk(1,1) ++ pm ++ spc () ++ + str "with" ++ brk(1,1) ++ pn ++ ppreason ++ str "." + +let explain_cannot_unify_local env sigma m n subn = + let pm = pr_leconstr_env env sigma m in + let pn = pr_leconstr_env env sigma n in + let psubn = pr_leconstr_env env sigma subn in + str "Unable to unify" ++ brk(1,1) ++ pm ++ spc () ++ + str "with" ++ brk(1,1) ++ pn ++ spc () ++ str "as" ++ brk(1,1) ++ + psubn ++ str " contains local variables." + +let explain_refiner_cannot_generalize env sigma ty = + str "Cannot find a well-typed generalisation of the goal with type: " ++ + pr_leconstr_env env sigma ty ++ str "." + +let explain_no_occurrence_found env sigma c id = + let c = EConstr.to_constr sigma c in + str "Found no subterm matching " ++ pr_lconstr_env env sigma c ++ + str " in " ++ + (match id with + | Some id -> Id.print id + | None -> str"the current goal") ++ str "." + +let explain_cannot_unify_binding_type env sigma m n = + let pm = pr_leconstr_env env sigma m in + let pn = pr_leconstr_env env sigma n in + str "This binding has type" ++ brk(1,1) ++ pm ++ spc () ++ + str "which should be unifiable with" ++ brk(1,1) ++ pn ++ str "." + +let explain_cannot_find_well_typed_abstraction env sigma p l e = + let p = EConstr.to_constr sigma p in + str "Abstracting over the " ++ + str (String.plural (List.length l) "term") ++ spc () ++ + hov 0 (pr_enum (fun c -> pr_lconstr_env env sigma (EConstr.to_constr sigma c)) l) ++ spc () ++ + str "leads to a term" ++ spc () ++ pr_lconstr_goal_style_env env sigma p ++ + spc () ++ str "which is ill-typed." ++ + (match e with None -> mt () | Some e -> fnl () ++ str "Reason is: " ++ e) + +let explain_wrong_abstraction_type env sigma na abs expected result = + let abs = EConstr.to_constr sigma abs in + let expected = EConstr.to_constr sigma expected in + let result = EConstr.to_constr sigma result in + let ppname = match na with Name id -> Id.print id ++ spc () | _ -> mt () in + str "Cannot instantiate metavariable " ++ ppname ++ strbrk "of type " ++ + pr_lconstr_env env sigma expected ++ strbrk " with abstraction " ++ + pr_lconstr_env env sigma abs ++ strbrk " of incompatible type " ++ + pr_lconstr_env env sigma result ++ str "." + +let explain_abstraction_over_meta _ m n = + strbrk "Too complex unification problem: cannot find a solution for both " ++ + Name.print m ++ spc () ++ str "and " ++ Name.print n ++ str "." + +let explain_non_linear_unification env sigma m t = + let t = EConstr.to_constr sigma t in + strbrk "Cannot unambiguously instantiate " ++ + Name.print m ++ str ":" ++ + strbrk " which would require to abstract twice on " ++ + pr_lconstr_env env sigma t ++ str "." + +let explain_unsatisfied_constraints env sigma cst = + strbrk "Unsatisfied constraints: " ++ + Univ.pr_constraints (Termops.pr_evd_level sigma) cst ++ + spc () ++ str "(maybe a bugged tactic)." + +let explain_type_error env sigma err = + let env = make_all_name_different env sigma in + match err with + | UnboundRel n -> + explain_unbound_rel env sigma n + | UnboundVar v -> + explain_unbound_var env v + | NotAType j -> + explain_not_type env sigma j + | BadAssumption c -> + explain_bad_assumption env sigma c + | ReferenceVariables (id,c) -> + explain_reference_variables sigma id c + | ElimArity (ind, aritylst, c, pj, okinds) -> + explain_elim_arity env sigma ind aritylst c pj okinds + | CaseNotInductive cj -> + explain_case_not_inductive env sigma cj + | NumberBranches (cj, n) -> + explain_number_branches env sigma cj n + | IllFormedBranch (c, i, actty, expty) -> + explain_ill_formed_branch env sigma c i actty expty + | Generalization (nvar, c) -> + explain_generalization env sigma nvar c + | ActualType (j, pt) -> + explain_actual_type env sigma j pt None + | CantApplyBadType (t, rator, randl) -> + explain_cant_apply_bad_type env sigma t rator randl + | CantApplyNonFunctional (rator, randl) -> + explain_cant_apply_not_functional env sigma rator randl + | IllFormedRecBody (err, lna, i, fixenv, vdefj) -> + explain_ill_formed_rec_body env sigma err lna i fixenv vdefj + | IllTypedRecBody (i, lna, vdefj, vargs) -> + explain_ill_typed_rec_body env sigma i lna vdefj vargs + | WrongCaseInfo (ind,ci) -> + explain_wrong_case_info env ind ci + | UnsatisfiedConstraints cst -> + explain_unsatisfied_constraints env sigma cst + +let pr_position (cl,pos) = + let clpos = match cl with + | None -> str " of the goal" + | Some (id,Locus.InHyp) -> str " of hypothesis " ++ Id.print id + | Some (id,Locus.InHypTypeOnly) -> str " of the type of hypothesis " ++ Id.print id + | Some (id,Locus.InHypValueOnly) -> str " of the body of hypothesis " ++ Id.print id in + int pos ++ clpos + +let explain_cannot_unify_occurrences env sigma nested ((cl2,pos2),t2) ((cl1,pos1),t1) e = + if nested then + str "Found nested occurrences of the pattern at positions " ++ + int pos1 ++ strbrk " and " ++ pr_position (cl2,pos2) ++ str "." + else + let ppreason = match e with + | None -> mt() + | Some (c1,c2,e) -> + explain_unification_error env sigma c1 c2 (Some e) + in + str "Found incompatible occurrences of the pattern" ++ str ":" ++ + spc () ++ str "Matched term " ++ pr_lconstr_env env sigma (EConstr.to_constr sigma t2) ++ + strbrk " at position " ++ pr_position (cl2,pos2) ++ + strbrk " is not compatible with matched term " ++ + pr_lconstr_env env sigma (EConstr.to_constr sigma t1) ++ strbrk " at position " ++ + pr_position (cl1,pos1) ++ ppreason ++ str "." + +let pr_constraints printenv env sigma evars cstrs = + let (ev, evi) = Evar.Map.choose evars in + if Evar.Map.for_all (fun ev' evi' -> + eq_named_context_val evi.evar_hyps evi'.evar_hyps) evars + then + let l = Evar.Map.bindings evars in + let env' = reset_with_named_context evi.evar_hyps env in + let pe = + if printenv then + pr_ne_context_of (str "In environment:") env' sigma + else mt () + in + let evs = + prlist + (fun (ev, evi) -> fnl () ++ pr_existential_key sigma ev ++ + str " : " ++ pr_lconstr_env env' sigma evi.evar_concl ++ fnl ()) l + in + h 0 (pe ++ evs ++ pr_evar_constraints sigma cstrs) + else + let filter evk _ = Evar.Map.mem evk evars in + pr_evar_map_filter ~with_univs:false filter sigma + +let explain_unsatisfiable_constraints env sigma constr comp = + let (_, constraints) = Evd.extract_all_conv_pbs sigma in + let undef = Evd.undefined_map sigma in + (** Only keep evars that are subject to resolution and members of the given + component. *) + let is_kept evk evi = match comp with + | None -> Typeclasses.is_resolvable evi + | Some comp -> Typeclasses.is_resolvable evi && Evar.Set.mem evk comp + in + let undef = + let m = Evar.Map.filter is_kept undef in + if Evar.Map.is_empty m then undef + else m + in + match constr with + | None -> + str "Unable to satisfy the following constraints:" ++ fnl () ++ + pr_constraints true env sigma undef constraints + | Some (ev, k) -> + let cstr = + let remaining = Evar.Map.remove ev undef in + if not (Evar.Map.is_empty remaining) then + str "With the following constraints:" ++ fnl () ++ + pr_constraints false env sigma remaining constraints + else mt () + in + let info = Evar.Map.find ev undef in + explain_typeclass_resolution env sigma info k ++ fnl () ++ cstr + +let explain_pretype_error env sigma err = + let env = Evardefine.env_nf_betaiotaevar sigma env in + let env = make_all_name_different env sigma in + match err with + | CantFindCaseType c -> explain_cant_find_case_type env sigma c + | ActualTypeNotCoercible (j,t,e) -> + let {uj_val = c; uj_type = actty} = j in + let (env, c, actty, expty), e = contract3' env sigma c actty t e in + let j = {uj_val = c; uj_type = actty} in + explain_actual_type env sigma j expty (Some e) + | UnifOccurCheck (ev,rhs) -> explain_occur_check env sigma ev rhs + | UnsolvableImplicit (evk,exp) -> explain_unsolvable_implicit env sigma evk exp + | VarNotFound id -> explain_var_not_found env id + | UnexpectedType (actual,expect) -> + let env, actual, expect = contract2 env sigma actual expect in + explain_unexpected_type env sigma actual expect + | NotProduct c -> explain_not_product env sigma c + | CannotUnify (m,n,e) -> + let env, m, n = contract2 env sigma m n in + explain_cannot_unify env sigma m n e + | CannotUnifyLocal (m,n,sn) -> explain_cannot_unify_local env sigma m n sn + | CannotGeneralize ty -> explain_refiner_cannot_generalize env sigma ty + | NoOccurrenceFound (c, id) -> explain_no_occurrence_found env sigma c id + | CannotUnifyBindingType (m,n) -> explain_cannot_unify_binding_type env sigma m n + | CannotFindWellTypedAbstraction (p,l,e) -> + explain_cannot_find_well_typed_abstraction env sigma p l + (Option.map (fun (env',e) -> explain_type_error env' sigma e) e) + | WrongAbstractionType (n,a,t,u) -> + explain_wrong_abstraction_type env sigma n a t u + | AbstractionOverMeta (m,n) -> explain_abstraction_over_meta env m n + | NonLinearUnification (m,c) -> explain_non_linear_unification env sigma m c + | TypingError t -> explain_type_error env sigma t + | CannotUnifyOccurrences (b,c1,c2,e) -> explain_cannot_unify_occurrences env sigma b c1 c2 e + | UnsatisfiableConstraints (c,comp) -> explain_unsatisfiable_constraints env sigma c comp +(* Module errors *) + +open Modops + +let explain_not_match_error = function + | InductiveFieldExpected _ -> + strbrk "an inductive definition is expected" + | DefinitionFieldExpected -> + strbrk "a definition is expected" + | ModuleFieldExpected -> + strbrk "a module is expected" + | ModuleTypeFieldExpected -> + strbrk "a module type is expected" + | NotConvertibleInductiveField id | NotConvertibleConstructorField id -> + str "types given to " ++ Id.print id ++ str " differ" + | NotConvertibleBodyField -> + str "the body of definitions differs" + | NotConvertibleTypeField (env, typ1, typ2) -> + str "expected type" ++ spc () ++ + quote (Printer.safe_pr_lconstr_env env Evd.empty typ2) ++ spc () ++ + str "but found type" ++ spc () ++ + quote (Printer.safe_pr_lconstr_env env Evd.empty typ1) + | NotSameConstructorNamesField -> + str "constructor names differ" + | NotSameInductiveNameInBlockField -> + str "inductive types names differ" + | FiniteInductiveFieldExpected isfinite -> + str "type is expected to be " ++ + str (if isfinite then "coinductive" else "inductive") + | InductiveNumbersFieldExpected n -> + str "number of inductive types differs" + | InductiveParamsNumberField n -> + str "inductive type has not the right number of parameters" + | RecordFieldExpected isrecord -> + str "type is expected " ++ str (if isrecord then "" else "not ") ++ + str "to be a record" + | RecordProjectionsExpected nal -> + (if List.length nal >= 2 then str "expected projection names are " + else str "expected projection name is ") ++ + pr_enum (function Name id -> Id.print id | _ -> str "_") nal + | NotEqualInductiveAliases -> + str "Aliases to inductive types do not match" + | CumulativeStatusExpected b -> + let status b = if b then str"cumulative" else str"non-cumulative" in + str "a " ++ status b ++ str" declaration was expected, but a " ++ + status (not b) ++ str" declaration was found" + | PolymorphicStatusExpected b -> + let status b = if b then str"polymorphic" else str"monomorphic" in + str "a " ++ status b ++ str" declaration was expected, but a " ++ + status (not b) ++ str" declaration was found" + | IncompatibleInstances -> + str"polymorphic universe instances do not match" + | IncompatibleUniverses incon -> + str"the universe constraints are inconsistent: " ++ + Univ.explain_universe_inconsistency Universes.pr_with_global_universes incon + | IncompatiblePolymorphism (env, t1, t2) -> + str "conversion of polymorphic values generates additional constraints: " ++ + quote (Printer.safe_pr_lconstr_env env Evd.empty t1) ++ spc () ++ + str "compared to " ++ spc () ++ + quote (Printer.safe_pr_lconstr_env env Evd.empty t2) + | IncompatibleConstraints cst -> + str " the expected (polymorphic) constraints do not imply " ++ + let cst = Univ.AUContext.instantiate (Univ.AUContext.instance cst) cst in + quote (Univ.pr_constraints (Termops.pr_evd_level Evd.empty) cst) + +let explain_signature_mismatch l spec why = + str "Signature components for label " ++ Label.print l ++ + str " do not match:" ++ spc () ++ explain_not_match_error why ++ str "." + +let explain_label_already_declared l = + str "The label " ++ Label.print l ++ str " is already declared." + +let explain_application_to_not_path _ = + strbrk "A module cannot be applied to another module application or " ++ + strbrk "with-expression; you must give a name to the intermediate result " ++ + strbrk "module first." + +let explain_not_a_functor () = + str "Application of a non-functor." + +let explain_is_a_functor () = + str "Illegal use of a functor." + +let explain_incompatible_module_types mexpr1 mexpr2 = + let open Declarations in + let rec get_arg = function + | NoFunctor _ -> 0 + | MoreFunctor (_, _, ty) -> succ (get_arg ty) + in + let len1 = get_arg mexpr1.mod_type in + let len2 = get_arg mexpr2.mod_type in + if len1 <> len2 then + str "Incompatible module types: module expects " ++ int len2 ++ + str " arguments, found " ++ int len1 ++ str "." + else str "Incompatible module types." + +let explain_not_equal_module_paths mp1 mp2 = + str "Non equal modules." + +let explain_no_such_label l = + str "No such label " ++ Label.print l ++ str "." + +let explain_incompatible_labels l l' = + str "Opening and closing labels are not the same: " ++ + Label.print l ++ str " <> " ++ Label.print l' ++ str "!" + +let explain_not_a_module s = + quote (str s) ++ str " is not a module." + +let explain_not_a_module_type s = + quote (str s) ++ str " is not a module type." + +let explain_not_a_constant l = + quote (Label.print l) ++ str " is not a constant." + +let explain_incorrect_label_constraint l = + str "Incorrect constraint for label " ++ + quote (Label.print l) ++ str "." + +let explain_generative_module_expected l = + str "The module " ++ Label.print l ++ str " is not generative." ++ + strbrk " Only components of generative modules can be changed" ++ + strbrk " using the \"with\" construct." + +let explain_label_missing l s = + str "The field " ++ Label.print l ++ str " is missing in " + ++ str s ++ str "." + +let explain_include_restricted_functor mp = + let q = Nametab.shortest_qualid_of_module mp in + str "Cannot include the functor " ++ Libnames.pr_qualid q ++ + strbrk " since it has a restricted signature. " ++ + strbrk "You may name first an instance of this functor, and include it." + +let explain_module_error = function + | SignatureMismatch (l,spec,err) -> explain_signature_mismatch l spec err + | LabelAlreadyDeclared l -> explain_label_already_declared l + | ApplicationToNotPath mexpr -> explain_application_to_not_path mexpr + | NotAFunctor -> explain_not_a_functor () + | IsAFunctor -> explain_is_a_functor () + | IncompatibleModuleTypes (m1,m2) -> explain_incompatible_module_types m1 m2 + | NotEqualModulePaths (mp1,mp2) -> explain_not_equal_module_paths mp1 mp2 + | NoSuchLabel l -> explain_no_such_label l + | IncompatibleLabels (l1,l2) -> explain_incompatible_labels l1 l2 + | NotAModule s -> explain_not_a_module s + | NotAModuleType s -> explain_not_a_module_type s + | NotAConstant l -> explain_not_a_constant l + | IncorrectWithConstraint l -> explain_incorrect_label_constraint l + | GenerativeModuleExpected l -> explain_generative_module_expected l + | LabelMissing (l,s) -> explain_label_missing l s + | IncludeRestrictedFunctor mp -> explain_include_restricted_functor mp + +(* Module internalization errors *) + +(* +let explain_declaration_not_path _ = + str "Declaration is not a path." + +*) + +let explain_not_module_nor_modtype s = + quote (str s) ++ str " is not a module or module type." + +let explain_incorrect_with_in_module () = + str "The syntax \"with\" is not allowed for modules." + +let explain_incorrect_module_application () = + str "Illegal application to a module type." + +open Modintern + +let explain_module_internalization_error = function + | NotAModuleNorModtype s -> explain_not_module_nor_modtype s + | IncorrectWithInModule -> explain_incorrect_with_in_module () + | IncorrectModuleApplication -> explain_incorrect_module_application () + +(* Typeclass errors *) + +let explain_not_a_class env c = + let c = EConstr.to_constr Evd.empty c in + pr_constr_env env Evd.empty c ++ str" is not a declared type class." + +let explain_unbound_method env cid { CAst.v = id } = + str "Unbound method name " ++ Id.print (id) ++ spc () ++ + str"of class" ++ spc () ++ pr_global cid ++ str "." + +let pr_constr_exprs exprs = + hv 0 (List.fold_right + (fun d pps -> ws 2 ++ Ppconstr.pr_constr_expr d ++ pps) + exprs (mt ())) + +let explain_mismatched_contexts env c i j = + str"Mismatched contexts while declaring instance: " ++ brk (1,1) ++ + hov 1 (str"Expected:" ++ brk (1, 1) ++ pr_rel_context env Evd.empty j) ++ + fnl () ++ brk (1,1) ++ + hov 1 (str"Found:" ++ brk (1, 1) ++ pr_constr_exprs i) + +let explain_typeclass_error env = function + | NotAClass c -> explain_not_a_class env c + | UnboundMethod (cid, id) -> explain_unbound_method env cid id + | MismatchedContextInstance (c,i,j) -> explain_mismatched_contexts env c i j + +(* Refiner errors *) + +let explain_refiner_bad_type env sigma arg ty conclty = + str "Refiner was given an argument" ++ brk(1,1) ++ + pr_lconstr_env env sigma arg ++ spc () ++ + str "of type" ++ brk(1,1) ++ pr_lconstr_env env sigma ty ++ spc () ++ + str "instead of" ++ brk(1,1) ++ pr_lconstr_env env sigma conclty ++ str "." + +let explain_refiner_unresolved_bindings l = + str "Unable to find an instance for the " ++ + str (String.plural (List.length l) "variable") ++ spc () ++ + prlist_with_sep pr_comma Name.print l ++ str"." + +let explain_refiner_cannot_apply env sigma t harg = + str "In refiner, a term of type" ++ brk(1,1) ++ + pr_lconstr_env env sigma t ++ spc () ++ str "could not be applied to" ++ brk(1,1) ++ + pr_lconstr_env env sigma harg ++ str "." + +let explain_refiner_not_well_typed env sigma c = + str "The term " ++ pr_lconstr_env env sigma c ++ str " is not well-typed." + +let explain_intro_needs_product () = + str "Introduction tactics needs products." + +let explain_does_not_occur_in env sigma c hyp = + str "The term" ++ spc () ++ pr_lconstr_env env sigma c ++ spc () ++ + str "does not occur in" ++ spc () ++ Id.print hyp ++ str "." + +let explain_non_linear_proof env sigma c = + str "Cannot refine with term" ++ brk(1,1) ++ pr_lconstr_env env sigma c ++ + spc () ++ str "because a metavariable has several occurrences." + +let explain_meta_in_type env sigma c = + str "In refiner, a meta appears in the type " ++ brk(1,1) ++ pr_leconstr_env env sigma c ++ + str " of another meta" + +let explain_no_such_hyp id = + str "No such hypothesis: " ++ Id.print id + +let explain_refiner_error env sigma = function + | BadType (arg,ty,conclty) -> explain_refiner_bad_type env sigma arg ty conclty + | UnresolvedBindings t -> explain_refiner_unresolved_bindings t + | CannotApply (t,harg) -> explain_refiner_cannot_apply env sigma t harg + | NotWellTyped c -> explain_refiner_not_well_typed env sigma c + | IntroNeedsProduct -> explain_intro_needs_product () + | DoesNotOccurIn (c,hyp) -> explain_does_not_occur_in env sigma c hyp + | NonLinearProof c -> explain_non_linear_proof env sigma c + | MetaInType c -> explain_meta_in_type env sigma c + | NoSuchHyp id -> explain_no_such_hyp id + +(* Inductive errors *) + +let error_non_strictly_positive env c v = + let pc = pr_lconstr_env env Evd.empty c in + let pv = pr_lconstr_env env Evd.empty v in + str "Non strictly positive occurrence of " ++ pv ++ str " in" ++ + brk(1,1) ++ pc ++ str "." + +let error_ill_formed_inductive env c v = + let pc = pr_lconstr_env env Evd.empty c in + let pv = pr_lconstr_env env Evd.empty v in + str "Not enough arguments applied to the " ++ pv ++ + str " in" ++ brk(1,1) ++ pc ++ str "." + +let error_ill_formed_constructor env id c v nparams nargs = + let pv = pr_lconstr_env env Evd.empty v in + let atomic = Int.equal (nb_prod Evd.empty (EConstr.of_constr c)) (** FIXME *) 0 in + str "The type of constructor" ++ brk(1,1) ++ Id.print id ++ brk(1,1) ++ + str "is not valid;" ++ brk(1,1) ++ + strbrk (if atomic then "it must be " else "its conclusion must be ") ++ + pv ++ + (* warning: because of implicit arguments it is difficult to say which + parameters must be explicitly given *) + (if not (Int.equal nparams 0) then + strbrk " applied to its " ++ str (String.plural nparams "parameter") + else + mt()) ++ + (if not (Int.equal nargs 0) then + str (if not (Int.equal nparams 0) then " and" else " applied") ++ + strbrk " to some " ++ str (String.plural nargs "argument") + else + mt()) ++ str "." + +let pr_ltype_using_barendregt_convention_env env c = + (* Use goal_concl_style as an approximation of Barendregt's convention (?) *) + quote (pr_goal_concl_style_env env Evd.empty (EConstr.of_constr c)) + +let error_bad_ind_parameters env c n v1 v2 = + let pc = pr_ltype_using_barendregt_convention_env env c in + let pv1 = pr_lconstr_env env Evd.empty v1 in + let pv2 = pr_lconstr_env env Evd.empty v2 in + str "Last occurrence of " ++ pv2 ++ str " must have " ++ pv1 ++ + str " as " ++ pr_nth n ++ str " argument in" ++ brk(1,1) ++ pc ++ str "." + +let error_same_names_types id = + str "The name" ++ spc () ++ Id.print id ++ spc () ++ + str "is used more than once." + +let error_same_names_constructors id = + str "The constructor name" ++ spc () ++ Id.print id ++ spc () ++ + str "is used more than once." + +let error_same_names_overlap idl = + strbrk "The following names are used both as type names and constructor " ++ + str "names:" ++ spc () ++ + prlist_with_sep pr_comma Id.print idl ++ str "." + +let error_not_an_arity env c = + str "The type" ++ spc () ++ pr_lconstr_env env Evd.empty c ++ spc () ++ + str "is not an arity." + +let error_bad_entry () = + str "Bad inductive definition." + +let error_large_non_prop_inductive_not_in_type () = + str "Large non-propositional inductive types must be in Type." + +(* Recursion schemes errors *) + +let error_not_allowed_case_analysis isrec kind i = + str (if isrec then "Induction" else "Case analysis") ++ + strbrk " on sort " ++ pr_sort Evd.empty kind ++ + strbrk " is not allowed for inductive definition " ++ + pr_inductive (Global.env()) (fst i) ++ str "." + +let error_not_allowed_dependent_analysis isrec i = + str "Dependent " ++ str (if isrec then "induction" else "case analysis") ++ + strbrk " is not allowed for inductive definition " ++ + pr_inductive (Global.env()) i ++ str "." + +let error_not_mutual_in_scheme ind ind' = + if eq_ind ind ind' then + str "The inductive type " ++ pr_inductive (Global.env()) ind ++ + str " occurs twice." + else + str "The inductive types " ++ pr_inductive (Global.env()) ind ++ spc () ++ + str "and" ++ spc () ++ pr_inductive (Global.env()) ind' ++ spc () ++ + str "are not mutually defined." + +(* Inductive constructions errors *) + +let explain_inductive_error = function + | NonPos (env,c,v) -> error_non_strictly_positive env c v + | NotEnoughArgs (env,c,v) -> error_ill_formed_inductive env c v + | NotConstructor (env,id,c,v,n,m) -> + error_ill_formed_constructor env id c v n m + | NonPar (env,c,n,v1,v2) -> error_bad_ind_parameters env c n v1 v2 + | SameNamesTypes id -> error_same_names_types id + | SameNamesConstructors id -> error_same_names_constructors id + | SameNamesOverlap idl -> error_same_names_overlap idl + | NotAnArity (env, c) -> error_not_an_arity env c + | BadEntry -> error_bad_entry () + | LargeNonPropInductiveNotInType -> + error_large_non_prop_inductive_not_in_type () + +(* Recursion schemes errors *) + +let explain_recursion_scheme_error = function + | NotAllowedCaseAnalysis (isrec,k,i) -> + error_not_allowed_case_analysis isrec k i + | NotMutualInScheme (ind,ind')-> error_not_mutual_in_scheme ind ind' + | NotAllowedDependentAnalysis (isrec, i) -> + error_not_allowed_dependent_analysis isrec i + +(* Pattern-matching errors *) + +let explain_bad_pattern env sigma cstr ty = + let ty = EConstr.to_constr sigma ty in + let env = make_all_name_different env sigma in + let pt = pr_lconstr_env env sigma ty in + let pc = pr_constructor env cstr in + str "Found the constructor " ++ pc ++ brk(1,1) ++ + str "while matching a term of type " ++ pt ++ brk(1,1) ++ + str "which is not an inductive type." + +let explain_bad_constructor env cstr ind = + let pi = pr_inductive env ind in +(* let pc = pr_constructor env cstr in*) + let pt = pr_inductive env (inductive_of_constructor cstr) in + str "Found a constructor of inductive type " ++ pt ++ brk(1,1) ++ + str "while a constructor of " ++ pi ++ brk(1,1) ++ + str "is expected." + +let decline_string n s = + if Int.equal n 0 then str "no " ++ str s ++ str "s" + else if Int.equal n 1 then str "1 " ++ str s + else (int n ++ str " " ++ str s ++ str "s") + +let explain_wrong_numarg_constructor env cstr n = + str "The constructor " ++ pr_constructor env cstr ++ + str " (in type " ++ pr_inductive env (inductive_of_constructor cstr) ++ + str ") expects " ++ decline_string n "argument" ++ str "." + +let explain_wrong_numarg_inductive env ind n = + str "The inductive type " ++ pr_inductive env ind ++ + str " expects " ++ decline_string n "argument" ++ str "." + +let explain_unused_clause env pats = +(* Without localisation + let s = if List.length pats > 1 then "s" else "" in + (str ("Unused clause with pattern"^s) ++ spc () ++ + hov 0 (pr_sequence pr_cases_pattern pats) ++ str ")") +*) + str "This clause is redundant." + +let explain_non_exhaustive env pats = + str "Non exhaustive pattern-matching: no clause found for " ++ + str (String.plural (List.length pats) "pattern") ++ + spc () ++ hov 0 (prlist_with_sep pr_comma pr_cases_pattern pats) + +let explain_cannot_infer_predicate env sigma typs = + let inj c = EConstr.to_constr sigma c in + let typs = Array.map_to_list (fun (c1, c2) -> (inj c1, inj c2)) typs in + let env = make_all_name_different env sigma in + let pr_branch (cstr,typ) = + let cstr,_ = decompose_app cstr in + str "For " ++ pr_lconstr_env env sigma cstr ++ str ": " ++ pr_lconstr_env env sigma typ + in + str "Unable to unify the types found in the branches:" ++ + spc () ++ hov 0 (prlist_with_sep fnl pr_branch typs) + +let explain_pattern_matching_error env sigma = function + | BadPattern (c,t) -> + explain_bad_pattern env sigma c t + | BadConstructor (c,ind) -> + explain_bad_constructor env c ind + | WrongNumargConstructor (c,n) -> + explain_wrong_numarg_constructor env c n + | WrongNumargInductive (c,n) -> + explain_wrong_numarg_inductive env c n + | UnusedClause tms -> + explain_unused_clause env tms + | NonExhaustive tms -> + explain_non_exhaustive env tms + | CannotInferPredicate typs -> + explain_cannot_infer_predicate env sigma typs + +let map_pguard_error f = function +| NotEnoughAbstractionInFixBody -> NotEnoughAbstractionInFixBody +| RecursionNotOnInductiveType c -> RecursionNotOnInductiveType (f c) +| RecursionOnIllegalTerm (n, (env, c), l1, l2) -> RecursionOnIllegalTerm (n, (env, f c), l1, l2) +| NotEnoughArgumentsForFixCall n -> NotEnoughArgumentsForFixCall n +| CodomainNotInductiveType c -> CodomainNotInductiveType (f c) +| NestedRecursiveOccurrences -> NestedRecursiveOccurrences +| UnguardedRecursiveCall c -> UnguardedRecursiveCall (f c) +| RecCallInTypeOfAbstraction c -> RecCallInTypeOfAbstraction (f c) +| RecCallInNonRecArgOfConstructor c -> RecCallInNonRecArgOfConstructor (f c) +| RecCallInTypeOfDef c -> RecCallInTypeOfDef (f c) +| RecCallInCaseFun c -> RecCallInCaseFun (f c) +| RecCallInCaseArg c -> RecCallInCaseArg (f c) +| RecCallInCasePred c -> RecCallInCasePred (f c) +| NotGuardedForm c -> NotGuardedForm (f c) +| ReturnPredicateNotCoInductive c -> ReturnPredicateNotCoInductive (f c) + +let map_ptype_error f = function +| UnboundRel n -> UnboundRel n +| UnboundVar id -> UnboundVar id +| NotAType j -> NotAType (on_judgment f j) +| BadAssumption j -> BadAssumption (on_judgment f j) +| ReferenceVariables (id, c) -> ReferenceVariables (id, f c) +| ElimArity (pi, dl, c, j, ar) -> ElimArity (pi, dl, f c, on_judgment f j, ar) +| CaseNotInductive j -> CaseNotInductive (on_judgment f j) +| WrongCaseInfo (pi, ci) -> WrongCaseInfo (pi, ci) +| NumberBranches (j, n) -> NumberBranches (on_judgment f j, n) +| IllFormedBranch (c, pc, t1, t2) -> IllFormedBranch (f c, pc, f t1, f t2) +| Generalization ((na, t), j) -> Generalization ((na, f t), on_judgment f j) +| ActualType (j, t) -> ActualType (on_judgment f j, f t) +| CantApplyBadType ((n, c1, c2), j, vj) -> + CantApplyBadType ((n, f c1, f c2), on_judgment f j, Array.map (on_judgment f) vj) +| CantApplyNonFunctional (j, jv) -> CantApplyNonFunctional (on_judgment f j, Array.map (on_judgment f) jv) +| IllFormedRecBody (ge, na, n, env, jv) -> + IllFormedRecBody (map_pguard_error f ge, na, n, env, Array.map (on_judgment f) jv) +| IllTypedRecBody (n, na, jv, t) -> + IllTypedRecBody (n, na, Array.map (on_judgment f) jv, Array.map f t) +| UnsatisfiedConstraints g -> UnsatisfiedConstraints g + +let explain_reduction_tactic_error = function + | Tacred.InvalidAbstraction (env,sigma,c,(env',e)) -> + let e = map_ptype_error EConstr.of_constr e in + str "The abstracted term" ++ spc () ++ + quote (pr_goal_concl_style_env env sigma c) ++ + spc () ++ str "is not well typed." ++ fnl () ++ + explain_type_error env' Evd.empty e diff --git a/vernac/himsg.mli b/vernac/himsg.mli new file mode 100644 index 00000000..0e20d18c --- /dev/null +++ b/vernac/himsg.mli @@ -0,0 +1,46 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* Evd.evar_map -> type_error -> Pp.t + +val explain_pretype_error : env -> Evd.evar_map -> pretype_error -> Pp.t + +val explain_inductive_error : inductive_error -> Pp.t + +val explain_typeclass_error : env -> typeclass_error -> Pp.t + +val explain_recursion_scheme_error : recursion_scheme_error -> Pp.t + +val explain_refiner_error : env -> Evd.evar_map -> refiner_error -> Pp.t + +val explain_pattern_matching_error : + env -> Evd.evar_map -> pattern_matching_error -> Pp.t + +val explain_reduction_tactic_error : + Tacred.reduction_tactic_error -> Pp.t + +val explain_module_error : Modops.module_typing_error -> Pp.t + +val explain_module_internalization_error : + Modintern.module_internalization_error -> Pp.t + +val map_pguard_error : ('c -> 'd) -> 'c pguard_error -> 'd pguard_error +val map_ptype_error : ('c -> 'd) -> ('c, 'c) ptype_error -> ('d, 'd) ptype_error diff --git a/vernac/indschemes.ml b/vernac/indschemes.ml new file mode 100644 index 00000000..49414862 --- /dev/null +++ b/vernac/indschemes.ml @@ -0,0 +1,517 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* !elim_flag) ; + optwrite = (fun b -> elim_flag := b) } + +let bifinite_elim_flag = ref false +let _ = + declare_bool_option + { optdepr = false; + optname = "automatic declaration of induction schemes for non-recursive types"; + optkey = ["Nonrecursive";"Elimination";"Schemes"]; + optread = (fun () -> !bifinite_elim_flag) ; + optwrite = (fun b -> bifinite_elim_flag := b) } + +let case_flag = ref false +let _ = + declare_bool_option + { optdepr = false; + optname = "automatic declaration of case analysis schemes"; + optkey = ["Case";"Analysis";"Schemes"]; + optread = (fun () -> !case_flag) ; + optwrite = (fun b -> case_flag := b) } + +let eq_flag = ref false +let _ = + declare_bool_option + { optdepr = false; + optname = "automatic declaration of boolean equality"; + optkey = ["Boolean";"Equality";"Schemes"]; + optread = (fun () -> !eq_flag) ; + optwrite = (fun b -> eq_flag := b) } + +let is_eq_flag () = !eq_flag + +let eq_dec_flag = ref false +let _ = + declare_bool_option + { optdepr = false; + optname = "automatic declaration of decidable equality"; + optkey = ["Decidable";"Equality";"Schemes"]; + optread = (fun () -> !eq_dec_flag) ; + optwrite = (fun b -> eq_dec_flag := b) } + +let rewriting_flag = ref false +let _ = + declare_bool_option + { optdepr = false; + optname ="automatic declaration of rewriting schemes for equality types"; + optkey = ["Rewriting";"Schemes"]; + optread = (fun () -> !rewriting_flag) ; + optwrite = (fun b -> rewriting_flag := b) } + +(* Util *) + +let define id internal ctx c t = + let f = declare_constant ~internal in + let univs = + if Flags.is_universe_polymorphism () + then Polymorphic_const_entry (Evd.to_universe_context ctx) + else Monomorphic_const_entry (Evd.universe_context_set ctx) + in + let kn = f id + (DefinitionEntry + { const_entry_body = c; + const_entry_secctx = None; + const_entry_type = t; + const_entry_universes = univs; + const_entry_opaque = false; + const_entry_inline_code = false; + const_entry_feedback = None; + }, + Decl_kinds.IsDefinition Scheme) in + definition_message id; + kn + +(* Boolean equality *) + +let declare_beq_scheme_gen internal names kn = + ignore (define_mutual_scheme beq_scheme_kind internal names kn) + +let alarm what internal msg = + let debug = false in + match internal with + | UserAutomaticRequest + | InternalTacticRequest -> + (if debug then + Feedback.msg_debug + (hov 0 msg ++ fnl () ++ what ++ str " not defined.")); None + | _ -> Some msg + +let try_declare_scheme what f internal names kn = + try f internal names kn + with e -> + let e = CErrors.push e in + let msg = match fst e with + | ParameterWithoutEquality cst -> + alarm what internal + (str "Boolean equality not found for parameter " ++ Printer.pr_global cst ++ + str".") + | InductiveWithProduct -> + alarm what internal + (str "Unable to decide equality of functional arguments.") + | InductiveWithSort -> + alarm what internal + (str "Unable to decide equality of type arguments.") + | NonSingletonProp ind -> + alarm what internal + (str "Cannot extract computational content from proposition " ++ + quote (Printer.pr_inductive (Global.env()) ind) ++ str ".") + | EqNotFound (ind',ind) -> + alarm what internal + (str "Boolean equality on " ++ + quote (Printer.pr_inductive (Global.env()) ind') ++ + strbrk " is missing.") + | UndefinedCst s -> + alarm what internal + (strbrk "Required constant " ++ str s ++ str " undefined.") + | AlreadyDeclared msg -> + alarm what internal (msg ++ str ".") + | DecidabilityMutualNotSupported -> + alarm what internal + (str "Decidability lemma for mutual inductive types not supported.") + | EqUnknown s -> + alarm what internal + (str "Found unsupported " ++ str s ++ str " while building Boolean equality.") + | NoDecidabilityCoInductive -> + alarm what internal + (str "Scheme Equality is only for inductive types.") + | e when CErrors.noncritical e -> + alarm what internal + (str "Unexpected error during scheme creation: " ++ CErrors.print e) + | _ -> iraise e + in + match msg with + | None -> () + | Some msg -> iraise (UserError (None, msg), snd e) + +let beq_scheme_msg mind = + let mib = Global.lookup_mind mind in + (* TODO: mutual inductive case *) + str "Boolean equality on " ++ + pr_enum (fun ind -> quote (Printer.pr_inductive (Global.env()) ind)) + (List.init (Array.length mib.mind_packets) (fun i -> (mind,i))) + +let declare_beq_scheme_with l kn = + try_declare_scheme (beq_scheme_msg kn) declare_beq_scheme_gen UserIndividualRequest l kn + +let try_declare_beq_scheme kn = + (* TODO: handle Fix, eventually handle + proof-irrelevance; improve decidability by depending on decidability + for the parameters rather than on the bl and lb properties *) + try_declare_scheme (beq_scheme_msg kn) declare_beq_scheme_gen UserAutomaticRequest [] kn + +let declare_beq_scheme = declare_beq_scheme_with [] + +(* Case analysis schemes *) +let declare_one_case_analysis_scheme ind = + let (mib,mip) = Global.lookup_inductive ind in + let kind = inductive_sort_family mip in + let dep = + if kind == InProp then case_scheme_kind_from_prop + else if not (Inductiveops.has_dependent_elim mib) then + case_scheme_kind_from_type + else case_dep_scheme_kind_from_type in + let kelim = elim_sorts (mib,mip) in + (* in case the inductive has a type elimination, generates only one + induction scheme, the other ones share the same code with the + apropriate type *) + if Sorts.List.mem InType kelim then + ignore (define_individual_scheme dep UserAutomaticRequest None ind) + +(* Induction/recursion schemes *) + +let kinds_from_prop = + [InType,rect_scheme_kind_from_prop; + InProp,ind_scheme_kind_from_prop; + InSet,rec_scheme_kind_from_prop] + +let kinds_from_type = + [InType,rect_dep_scheme_kind_from_type; + InProp,ind_dep_scheme_kind_from_type; + InSet,rec_dep_scheme_kind_from_type] + +let nondep_kinds_from_type = + [InType,rect_scheme_kind_from_type; + InProp,ind_scheme_kind_from_type; + InSet,rec_scheme_kind_from_type] + +let declare_one_induction_scheme ind = + let (mib,mip) = Global.lookup_inductive ind in + let kind = inductive_sort_family mip in + let from_prop = kind == InProp in + let depelim = Inductiveops.has_dependent_elim mib in + let kelim = elim_sorts (mib,mip) in + let elims = + List.map_filter (fun (sort,kind) -> + if Sorts.List.mem sort kelim then Some kind else None) + (if from_prop then kinds_from_prop + else if depelim then kinds_from_type + else nondep_kinds_from_type) in + List.iter (fun kind -> ignore (define_individual_scheme kind UserAutomaticRequest None ind)) + elims + +let declare_induction_schemes kn = + let mib = Global.lookup_mind kn in + if mib.mind_finite <> Declarations.CoFinite then begin + for i = 0 to Array.length mib.mind_packets - 1 do + declare_one_induction_scheme (kn,i); + done; + end + +(* Decidable equality *) + +let declare_eq_decidability_gen internal names kn = + let mib = Global.lookup_mind kn in + if mib.mind_finite <> Declarations.CoFinite then + ignore (define_mutual_scheme eq_dec_scheme_kind internal names kn) + +let eq_dec_scheme_msg ind = (* TODO: mutual inductive case *) + str "Decidable equality on " ++ quote (Printer.pr_inductive (Global.env()) ind) + +let declare_eq_decidability_scheme_with l kn = + try_declare_scheme (eq_dec_scheme_msg (kn,0)) + declare_eq_decidability_gen UserIndividualRequest l kn + +let try_declare_eq_decidability kn = + try_declare_scheme (eq_dec_scheme_msg (kn,0)) + declare_eq_decidability_gen UserAutomaticRequest [] kn + +let declare_eq_decidability = declare_eq_decidability_scheme_with [] + +let ignore_error f x = + try ignore (f x) with e when CErrors.noncritical e -> () + +let declare_rewriting_schemes ind = + if Hipattern.is_inductive_equality ind then begin + ignore (define_individual_scheme rew_r2l_scheme_kind UserAutomaticRequest None ind); + ignore (define_individual_scheme rew_r2l_dep_scheme_kind UserAutomaticRequest None ind); + ignore (define_individual_scheme rew_r2l_forward_dep_scheme_kind + UserAutomaticRequest None ind); + (* These ones expect the equality to be symmetric; the first one also *) + (* needs eq *) + ignore_error (define_individual_scheme rew_l2r_scheme_kind UserAutomaticRequest None) ind; + ignore_error + (define_individual_scheme rew_l2r_dep_scheme_kind UserAutomaticRequest None) ind; + ignore_error + (define_individual_scheme rew_l2r_forward_dep_scheme_kind UserAutomaticRequest None) ind + end + +let warn_cannot_build_congruence = + CWarnings.create ~name:"cannot-build-congruence" ~category:"schemes" + (fun () -> + strbrk "Cannot build congruence scheme because eq is not found") + +let declare_congr_scheme ind = + if Hipattern.is_equality_type Evd.empty (EConstr.of_constr (mkInd ind)) (** FIXME *) then begin + if + try Coqlib.check_required_library Coqlib.logic_module_name; true + with e when CErrors.noncritical e -> false + then + ignore (define_individual_scheme congr_scheme_kind UserAutomaticRequest None ind) + else + warn_cannot_build_congruence () + end + +let declare_sym_scheme ind = + if Hipattern.is_inductive_equality ind then + (* Expect the equality to be symmetric *) + ignore_error (define_individual_scheme sym_scheme_kind UserAutomaticRequest None) ind + +(* Scheme command *) + +let smart_global_inductive y = smart_global_inductive y +let rec split_scheme l = + let env = Global.env() in + match l with + | [] -> [],[] + | (Some id,t)::q -> let l1,l2 = split_scheme q in + ( match t with + | InductionScheme (x,y,z) -> ((id,x,smart_global_inductive y,z)::l1),l2 + | CaseScheme (x,y,z) -> ((id,x,smart_global_inductive y,z)::l1),l2 + | EqualityScheme x -> l1,((Some id,smart_global_inductive x)::l2) + ) +(* + if no name has been provided, we build one from the types of the ind +requested +*) + | (None,t)::q -> + let l1,l2 = split_scheme q in + let names inds recs isdep y z = + let ind = smart_global_inductive y in + let sort_of_ind = inductive_sort_family (snd (lookup_mind_specif env ind)) in + let suffix = ( + match sort_of_ind with + | InProp -> + if isdep then (match z with + | InProp -> inds ^ "_dep" + | InSet -> recs ^ "_dep" + | InType -> recs ^ "t_dep") + else ( match z with + | InProp -> inds + | InSet -> recs + | InType -> recs ^ "t" ) + | _ -> + if isdep then (match z with + | InProp -> inds + | InSet -> recs + | InType -> recs ^ "t" ) + else (match z with + | InProp -> inds ^ "_nodep" + | InSet -> recs ^ "_nodep" + | InType -> recs ^ "t_nodep") + ) in + let newid = add_suffix (basename_of_global (IndRef ind)) suffix in + let newref = CAst.make newid in + ((newref,isdep,ind,z)::l1),l2 + in + match t with + | CaseScheme (x,y,z) -> names "_case" "_case" x y z + | InductionScheme (x,y,z) -> names "_ind" "_rec" x y z + | EqualityScheme x -> l1,((None,smart_global_inductive x)::l2) + +let do_mutual_induction_scheme ?(force_mutual=false) lnamedepindsort = + let lrecnames = List.map (fun ({CAst.v},_,_,_) -> v) lnamedepindsort + and env0 = Global.env() in + let sigma, lrecspec, _ = + List.fold_right + (fun (_,dep,ind,sort) (evd, l, inst) -> + let evd, indu, inst = + match inst with + | None -> + let _, ctx = Global.type_of_global_in_context env0 (IndRef ind) in + let u, ctx = Universes.fresh_instance_from ctx None in + let evd = Evd.from_ctx (UState.of_context_set ctx) in + evd, (ind,u), Some u + | Some ui -> evd, (ind, ui), inst + in + (evd, (indu,dep,sort) :: l, inst)) + lnamedepindsort (Evd.from_env env0,[],None) + in + let sigma, listdecl = Indrec.build_mutual_induction_scheme env0 sigma ~force_mutual lrecspec in + let declare decl fi lrecref = + let decltype = Retyping.get_type_of env0 sigma (EConstr.of_constr decl) in + let decltype = EConstr.to_constr sigma decltype in + let proof_output = Future.from_val ((decl,Univ.ContextSet.empty),Safe_typing.empty_private_constants) in + let cst = define fi UserIndividualRequest sigma proof_output (Some decltype) in + ConstRef cst :: lrecref + in + let _ = List.fold_right2 declare listdecl lrecnames [] in + fixpoint_message None lrecnames + +let get_common_underlying_mutual_inductive = function + | [] -> assert false + | (id,(mind,i as ind))::l as all -> + match List.filter (fun (_,(mind',_)) -> not (MutInd.equal mind mind')) l with + | (_,ind')::_ -> + raise (RecursionSchemeError (NotMutualInScheme (ind,ind'))) + | [] -> + if not (List.distinct_f Int.compare (List.map snd (List.map snd all))) + then user_err Pp.(str "A type occurs twice"); + mind, + List.map_filter + (function (Some id,(_,i)) -> Some (i,id.CAst.v) | (None,_) -> None) all + +let do_scheme l = + let ischeme,escheme = split_scheme l in +(* we want 1 kind of scheme at a time so we check if the user +tried to declare different schemes at once *) + if not (List.is_empty ischeme) && not (List.is_empty escheme) + then + user_err Pp.(str "Do not declare equality and induction scheme at the same time.") + else ( + if not (List.is_empty ischeme) then do_mutual_induction_scheme ischeme + else + let mind,l = get_common_underlying_mutual_inductive escheme in + declare_beq_scheme_with l mind; + declare_eq_decidability_scheme_with l mind + ) + +(**********************************************************************) +(* Combined scheme *) +(* Matthieu Sozeau, Dec 2006 *) + +let list_split_rev_at index l = + let rec aux i acc = function + hd :: tl when Int.equal i index -> acc, tl + | hd :: tl -> aux (succ i) (hd :: acc) tl + | [] -> failwith "List.split_when: Invalid argument" + in aux 0 [] l + +let fold_left' f = function + [] -> invalid_arg "fold_left'" + | hd :: tl -> List.fold_left f hd tl + +let mk_coq_and sigma = Evarutil.new_global sigma (Coqlib.build_coq_and ()) +let mk_coq_conj sigma = Evarutil.new_global sigma (Coqlib.build_coq_conj ()) + +let build_combined_scheme env schemes = + let evdref = ref (Evd.from_env env) in + let defs = List.map (fun cst -> + let evd, c = Evd.fresh_constant_instance env !evdref cst in + evdref := evd; (c, Typeops.type_of_constant_in env c)) schemes in + let find_inductive ty = + let (ctx, arity) = decompose_prod ty in + let (_, last) = List.hd ctx in + match Constr.kind last with + | App (ind, args) -> + let ind = destInd ind in + let (_,spec) = Inductive.lookup_mind_specif env (fst ind) in + ctx, ind, spec.mind_nrealargs + | _ -> ctx, destInd last, 0 + in + let (c, t) = List.hd defs in + let ctx, ind, nargs = find_inductive t in + (* Number of clauses, including the predicates quantification *) + let prods = nb_prod !evdref (EConstr.of_constr t) - (nargs + 1) in + let sigma, coqand = mk_coq_and !evdref in + let sigma, coqconj = mk_coq_conj sigma in + let () = evdref := sigma in + let relargs = rel_vect 0 prods in + let concls = List.rev_map + (fun (cst, t) -> + mkApp(mkConstU cst, relargs), + snd (decompose_prod_n prods t)) defs in + let concl_bod, concl_typ = + fold_left' + (fun (accb, acct) (cst, x) -> + mkApp (EConstr.to_constr !evdref coqconj, [| x; acct; cst; accb |]), + mkApp (EConstr.to_constr !evdref coqand, [| x; acct |])) concls + in + let ctx, _ = + list_split_rev_at prods + (List.rev_map (fun (x, y) -> LocalAssum (x, y)) ctx) in + let typ = List.fold_left (fun d c -> Term.mkProd_wo_LetIn c d) concl_typ ctx in + let body = it_mkLambda_or_LetIn concl_bod ctx in + (!evdref, body, typ) + +let do_combined_scheme name schemes = + let open CAst in + let csts = + List.map (fun {CAst.loc;v} -> + let qualid = qualid_of_ident v in + try Nametab.locate_constant qualid + with Not_found -> user_err ?loc Pp.(pr_qualid qualid ++ str " is not declared.")) + schemes + in + let sigma,body,typ = build_combined_scheme (Global.env ()) csts in + let proof_output = Future.from_val ((body,Univ.ContextSet.empty),Safe_typing.empty_private_constants) in + ignore (define name.v UserIndividualRequest sigma proof_output (Some typ)); + fixpoint_message None [name.v] + +(**********************************************************************) + +let map_inductive_block f kn n = for i=0 to n-1 do f (kn,i) done + +let declare_default_schemes kn = + let mib = Global.lookup_mind kn in + let n = Array.length mib.mind_packets in + if !elim_flag && (mib.mind_finite <> Declarations.BiFinite || !bifinite_elim_flag) + && mib.mind_typing_flags.check_guarded then + declare_induction_schemes kn; + if !case_flag then map_inductive_block declare_one_case_analysis_scheme kn n; + if is_eq_flag() then try_declare_beq_scheme kn; + if !eq_dec_flag then try_declare_eq_decidability kn; + if !rewriting_flag then map_inductive_block declare_congr_scheme kn n; + if !rewriting_flag then map_inductive_block declare_sym_scheme kn n; + if !rewriting_flag then map_inductive_block declare_rewriting_schemes kn n diff --git a/vernac/indschemes.mli b/vernac/indschemes.mli new file mode 100644 index 00000000..43915788 --- /dev/null +++ b/vernac/indschemes.mli @@ -0,0 +1,53 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* unit + +val declare_eq_decidability : MutInd.t -> unit + +(** Build and register a congruence scheme for an equality-like inductive type *) + +val declare_congr_scheme : inductive -> unit + +(** Build and register rewriting schemes for an equality-like inductive type *) + +val declare_rewriting_schemes : inductive -> unit + +(** Mutual Minimality/Induction scheme. + [force_mutual] forces the construction of eliminators having the same predicates and + methods even if some of the inductives are not recursive. + By default it is [false] and some of the eliminators are defined as simple case analysis. + *) + +val do_mutual_induction_scheme : ?force_mutual:bool -> + (Misctypes.lident * bool * inductive * Sorts.family) list -> unit + +(** Main calls to interpret the Scheme command *) + +val do_scheme : (Misctypes.lident option * scheme) list -> unit + +(** Combine a list of schemes into a conjunction of them *) + +val build_combined_scheme : env -> Constant.t list -> Evd.evar_map * constr * types + +val do_combined_scheme : Misctypes.lident -> Misctypes.lident list -> unit + +(** Hook called at each inductive type definition *) + +val declare_default_schemes : MutInd.t -> unit diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml new file mode 100644 index 00000000..30dd6ec7 --- /dev/null +++ b/vernac/lemmas.ml @@ -0,0 +1,535 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* Globnames.global_reference -> 'a +let mk_hook hook = hook +let call_hook fix_exn hook l c = + try hook l c + with e when CErrors.noncritical e -> + let e = CErrors.push e in + iraise (fix_exn e) + +(* Support for mutually proved theorems *) + +let retrieve_first_recthm uctx = function + | VarRef id -> + (NamedDecl.get_value (Global.lookup_named id),variable_opacity id) + | ConstRef cst -> + let cb = Global.lookup_constant cst in + (* we get the right order somehow but surely it could be enforced in a better way *) + let uctx = UState.context uctx in + let inst = Univ.UContext.instance uctx in + let map (c, ctx) = Vars.subst_instance_constr inst c in + (Option.map map (Global.body_of_constant_body cb), is_opaque cb) + | _ -> assert false + +let adjust_guardness_conditions const = function + | [] -> const (* Not a recursive statement *) + | possible_indexes -> + (* Try all combinations... not optimal *) + let env = Global.env() in + { const with const_entry_body = + Future.chain const.const_entry_body + (fun ((body, ctx), eff) -> + match Constr.kind body with + | Fix ((nv,0),(_,_,fixdefs as fixdecls)) -> +(* let possible_indexes = + List.map2 (fun i c -> match i with Some i -> i | None -> + List.interval 0 (List.length ((lam_assum c)))) + lemma_guard (Array.to_list fixdefs) in +*) + let add c cb e = + let exists c e = + try ignore(Environ.lookup_constant c e); true + with Not_found -> false in + if exists c e then e else Environ.add_constant c cb e in + let env = List.fold_left (fun env { eff } -> + match eff with + | SEsubproof (c, cb,_) -> add c cb env + | SEscheme (l,_) -> + List.fold_left (fun e (_,c,cb,_) -> add c cb e) env l) + env (Safe_typing.side_effects_of_private_constants eff) in + let indexes = + search_guard env + possible_indexes fixdecls in + (mkFix ((indexes,0),fixdecls), ctx), eff + | _ -> (body, ctx), eff) } + +let find_mutually_recursive_statements sigma thms = + let n = List.length thms in + let inds = List.map (fun (id,(t,impls)) -> + let (hyps,ccl) = EConstr.decompose_prod_assum sigma t in + let x = (id,(t,impls)) in + let whnf_hyp_hds = EConstr.map_rel_context_in_env + (fun env c -> fst (Reductionops.whd_all_stack env sigma c)) + (Global.env()) hyps in + let ind_hyps = + List.flatten (List.map_i (fun i decl -> + let t = RelDecl.get_type decl in + match EConstr.kind sigma t with + | Ind ((kn,_ as ind),u) when + let mind = Global.lookup_mind kn in + mind.mind_finite <> Declarations.CoFinite -> + [ind,x,i] + | _ -> + []) 0 (List.rev (List.filter Context.Rel.Declaration.is_local_assum whnf_hyp_hds))) in + let ind_ccl = + let cclenv = EConstr.push_rel_context hyps (Global.env()) in + let whnf_ccl,_ = whd_all_stack cclenv Evd.empty ccl in + match EConstr.kind sigma whnf_ccl with + | Ind ((kn,_ as ind),u) when + let mind = Global.lookup_mind kn in + Int.equal mind.mind_ntypes n && mind.mind_finite == Declarations.CoFinite -> + [ind,x,0] + | _ -> + [] in + ind_hyps,ind_ccl) thms in + let inds_hyps,ind_ccls = List.split inds in + let of_same_mutind ((kn,_),_,_) = function ((kn',_),_,_) -> MutInd.equal kn kn' in + (* Check if all conclusions are coinductive in the same type *) + (* (degenerated cartesian product since there is at most one coind ccl) *) + let same_indccl = + List.cartesians_filter (fun hyp oks -> + if List.for_all (of_same_mutind hyp) oks + then Some (hyp::oks) else None) [] ind_ccls in + let ordered_same_indccl = + List.filter (List.for_all_i (fun i ((kn,j),_,_) -> Int.equal i j) 0) same_indccl in + (* Check if some hypotheses are inductive in the same type *) + let common_same_indhyp = + List.cartesians_filter (fun hyp oks -> + if List.for_all (of_same_mutind hyp) oks + then Some (hyp::oks) else None) [] inds_hyps in + let ordered_inds,finite,guard = + match ordered_same_indccl, common_same_indhyp with + | indccl::rest, _ -> + assert (List.is_empty rest); + (* One occ. of common coind ccls and no common inductive hyps *) + if not (List.is_empty common_same_indhyp) then + Flags.if_verbose Feedback.msg_info (str "Assuming mutual coinductive statements."); + flush_all (); + indccl, true, [] + | [], _::_ -> + let () = match same_indccl with + | ind :: _ -> + if List.distinct_f ind_ord (List.map pi1 ind) + then + Flags.if_verbose Feedback.msg_info + (strbrk + ("Coinductive statements do not follow the order of "^ + "definition, assuming the proof to be by induction.")); + flush_all () + | _ -> () + in + let possible_guards = List.map (List.map pi3) inds_hyps in + (* assume the largest indices as possible *) + List.last common_same_indhyp, false, possible_guards + | _, [] -> + user_err Pp.(str + ("Cannot find common (mutual) inductive premises or coinductive" ^ + " conclusions in the statements.")) + in + (finite,guard,None), ordered_inds + +let look_for_possibly_mutual_statements sigma = function + | [id,(t,impls)] -> + (* One non recursively proved theorem *) + None,[id,(t,impls)],None + | _::_ as thms -> + (* More than one statement and/or an explicit decreasing mark: *) + (* we look for a common inductive hyp or a common coinductive conclusion *) + let recguard,ordered_inds = find_mutually_recursive_statements sigma thms in + let thms = List.map pi2 ordered_inds in + Some recguard,thms, Some (List.map (fun (_,_,i) -> succ i) ordered_inds) + | [] -> anomaly (Pp.str "Empty list of theorems.") + +(* Saving a goal *) + +let save ?export_seff id const uctx do_guard (locality,poly,kind) hook = + let fix_exn = Future.fix_exn_of const.Entries.const_entry_body in + try + let const = adjust_guardness_conditions const do_guard in + let k = Kindops.logical_kind_of_goal_kind kind in + let should_suggest = const.const_entry_opaque && Option.is_empty const.const_entry_secctx in + let l,r = match locality with + | Discharge when Lib.sections_are_opened () -> + let c = SectionLocalDef const in + let _ = declare_variable id (Lib.cwd(), c, k) in + let () = if should_suggest + then Proof_using.suggest_variable (Global.env ()) id + in + (Local, VarRef id) + | Local | Global | Discharge -> + let local = match locality with + | Local | Discharge -> true + | Global -> false + in + let kn = + declare_constant ?export_seff id ~local (DefinitionEntry const, k) in + let () = if should_suggest + then Proof_using.suggest_constant (Global.env ()) kn + in + (locality, ConstRef kn) + in + definition_message id; + Declare.declare_univ_binders r (UState.universe_binders uctx); + call_hook (fun exn -> exn) hook l r + with e when CErrors.noncritical e -> + let e = CErrors.push e in + iraise (fix_exn e) + +let default_thm_id = Id.of_string "Unnamed_thm" + +let fresh_name_for_anonymous_theorem () = + let avoid = Id.Set.of_list (Proof_global.get_all_proof_names ()) in + next_global_ident_away default_thm_id avoid + +let check_name_freshness locality {CAst.loc;v=id} : unit = + (* We check existence here: it's a bit late at Qed time *) + if Nametab.exists_cci (Lib.make_path id) || is_section_variable id || + locality == Global && Nametab.exists_cci (Lib.make_path_except_section id) + then + user_err ?loc (Id.print id ++ str " already exists.") + +let save_remaining_recthms (locality,p,kind) norm univs body opaq i (id,(t_i,(_,imps))) = + let t_i = norm t_i in + match body with + | None -> + (match locality with + | Discharge -> + let impl = false in (* copy values from Vernacentries *) + let k = IsAssumption Conjectural in + let univs = match univs with + | Polymorphic_const_entry univs -> + (* What is going on here? *) + Univ.ContextSet.of_context univs + | Monomorphic_const_entry univs -> univs + in + let c = SectionLocalAssum ((t_i, univs),p,impl) in + let _ = declare_variable id (Lib.cwd(),c,k) in + (Discharge, VarRef id,imps) + | Local | Global -> + let k = IsAssumption Conjectural in + let local = match locality with + | Local -> true + | Global -> false + | Discharge -> assert false + in + let decl = (ParameterEntry (None,(t_i,univs),None), k) in + let kn = declare_constant id ~local decl in + (locality,ConstRef kn,imps)) + | Some body -> + let body = norm body in + let k = Kindops.logical_kind_of_goal_kind kind in + let rec body_i t = match Constr.kind t with + | Fix ((nv,0),decls) -> mkFix ((nv,i),decls) + | CoFix (0,decls) -> mkCoFix (i,decls) + | LetIn(na,t1,ty,t2) -> mkLetIn (na,t1,ty, body_i t2) + | Lambda(na,ty,t) -> mkLambda(na,ty,body_i t) + | App (t, args) -> mkApp (body_i t, args) + | _ -> + let sigma, env = Pfedit.get_current_context () in + anomaly Pp.(str "Not a proof by induction: " ++ Printer.pr_constr_env env sigma body ++ str ".") in + let body_i = body_i body in + match locality with + | Discharge -> + let const = definition_entry ~types:t_i ~opaque:opaq ~univs body_i in + let c = SectionLocalDef const in + let _ = declare_variable id (Lib.cwd(), c, k) in + (Discharge,VarRef id,imps) + | Local | Global -> + let local = match locality with + | Local -> true + | Global -> false + | Discharge -> assert false + in + let const = + Declare.definition_entry ~types:t_i ~univs ~opaque:opaq body_i + in + let kn = declare_constant id ~local (DefinitionEntry const, k) in + (locality,ConstRef kn,imps) + +let save_hook = ref ignore +let set_save_hook f = save_hook := f + +let save_named ?export_seff proof = + let id,const,uctx,do_guard,persistence,hook = proof in + save ?export_seff id const uctx do_guard persistence hook + +let check_anonymity id save_ident = + if not (String.equal (atompart_of_id id) (Id.to_string (default_thm_id))) then + user_err Pp.(str "This command can only be used for unnamed theorem.") + +let save_anonymous ?export_seff proof save_ident = + let id,const,uctx,do_guard,persistence,hook = proof in + check_anonymity id save_ident; + save ?export_seff save_ident const uctx do_guard persistence hook + +(* Admitted *) + +let warn_let_as_axiom = + CWarnings.create ~name:"let-as-axiom" ~category:"vernacular" + (fun id -> strbrk "Let definition" ++ spc () ++ Id.print id ++ + spc () ++ strbrk "declared as an axiom.") + +let admit (id,k,e) pl hook () = + let kn = declare_constant id (ParameterEntry e, IsAssumption Conjectural) in + let () = match k with + | Global, _, _ -> () + | Local, _, _ | Discharge, _, _ -> warn_let_as_axiom id + in + let () = assumption_message id in + Declare.declare_univ_binders (ConstRef kn) pl; + call_hook (fun exn -> exn) hook Global (ConstRef kn) + +(* Starting a goal *) + +let start_hook = ref ignore +let set_start_hook = (:=) start_hook + + +let get_proof proof do_guard hook opacity = + let (id,(const,univs,persistence)) = + Pfedit.cook_this_proof proof + in + id,{const with const_entry_opaque = opacity},univs,do_guard,persistence,hook + +let universe_proof_terminator compute_guard hook = + let open Proof_global in + make_terminator begin function + | Admitted (id,k,pe,ctx) -> + admit (id,k,pe) (UState.universe_binders ctx) (hook (Some ctx)) (); + Feedback.feedback Feedback.AddedAxiom + | Proved (opaque,idopt,proof) -> + let is_opaque, export_seff = match opaque with + | Vernacexpr.Transparent -> false, true + | Vernacexpr.Opaque -> true, false + in + let proof = get_proof proof compute_guard + (hook (Some (proof.Proof_global.universes))) is_opaque in + begin match idopt with + | None -> save_named ~export_seff proof + | Some { CAst.v = id } -> save_anonymous ~export_seff proof id + end + end + +let standard_proof_terminator compute_guard hook = + universe_proof_terminator compute_guard (fun _ -> hook) + +let start_proof id ?pl kind sigma ?terminator ?sign c ?init_tac ?(compute_guard=[]) hook = + let terminator = match terminator with + | None -> standard_proof_terminator compute_guard hook + | Some terminator -> terminator compute_guard hook + in + let sign = + match sign with + | Some sign -> sign + | None -> initialize_named_context_for_proof () + in + !start_hook c; + Pfedit.start_proof id ?pl kind sigma sign c ?init_tac terminator + +let start_proof_univs id ?pl kind sigma ?terminator ?sign c ?init_tac ?(compute_guard=[]) hook = + let terminator = match terminator with + | None -> universe_proof_terminator compute_guard hook + | Some terminator -> terminator compute_guard hook + in + let sign = + match sign with + | Some sign -> sign + | None -> initialize_named_context_for_proof () + in + !start_hook c; + Pfedit.start_proof id ?pl kind sigma sign c ?init_tac terminator + +let rec_tac_initializer finite guard thms snl = + if finite then + match List.map (fun (id,(t,_)) -> (id,t)) thms with + | (id,_)::l -> Tactics.mutual_cofix id l 0 + | _ -> assert false + else + (* nl is dummy: it will be recomputed at Qed-time *) + let nl = match snl with + | None -> List.map succ (List.map List.last guard) + | Some nl -> nl + in match List.map2 (fun (id,(t,_)) n -> (id,n, t)) thms nl with + | (id,n,_)::l -> Tactics.mutual_fix id n l 0 + | _ -> assert false + +let start_proof_with_initialization kind sigma decl recguard thms snl hook = + let intro_tac (_, (_, (ids, _))) = + Tacticals.New.tclMAP (function + | Name id -> Tactics.intro_mustbe_force id + | Anonymous -> Tactics.intro) (List.rev ids) in + let init_tac,guard = match recguard with + | Some (finite,guard,init_tac) -> + let rec_tac = rec_tac_initializer finite guard thms snl in + Some (match init_tac with + | None -> + if Flags.is_auto_intros () then + Tacticals.New.tclTHENS rec_tac (List.map intro_tac thms) + else + rec_tac + | Some tacl -> + Tacticals.New.tclTHENS rec_tac + (if Flags.is_auto_intros () then + List.map2 (fun tac thm -> Tacticals.New.tclTHEN tac (intro_tac thm)) tacl thms + else + tacl)),guard + | None -> + let () = match thms with [_] -> () | _ -> assert false in + (if Flags.is_auto_intros () then Some (intro_tac (List.hd thms)) else None), [] in + match thms with + | [] -> anomaly (Pp.str "No proof to start.") + | (id,(t,(_,imps)))::other_thms -> + let hook ctx strength ref = + let ctx = match ctx with + | None -> UState.empty + | Some ctx -> ctx + in + let other_thms_data = + if List.is_empty other_thms then [] else + (* there are several theorems defined mutually *) + let body,opaq = retrieve_first_recthm ctx ref in + let norm c = EConstr.to_constr (Evd.from_ctx ctx) c in + let body = Option.map EConstr.of_constr body in + let uctx = UState.check_univ_decl ~poly:(pi2 kind) ctx decl in + List.map_i (save_remaining_recthms kind norm uctx body opaq) 1 other_thms in + let thms_data = (strength,ref,imps)::other_thms_data in + List.iter (fun (strength,ref,imps) -> + maybe_declare_manual_implicits false ref imps; + call_hook (fun exn -> exn) hook strength ref) thms_data in + start_proof_univs id ~pl:decl kind sigma t ?init_tac (fun ctx -> mk_hook (hook ctx)) ~compute_guard:guard + +let start_proof_com ?inference_hook kind thms hook = + let env0 = Global.env () in + let decl = fst (List.hd thms) in + let evd, decl = Univdecls.interp_univ_decl_opt env0 (snd decl) in + let evd, thms = List.fold_left_map (fun evd ((id, _), (bl, t)) -> + let evd, (impls, ((env, ctx), imps)) = interp_context_evars env0 evd bl in + let evd, (t', imps') = interp_type_evars_impls ~impls env evd t in + let flags = all_and_fail_flags in + let flags = { flags with use_hook = inference_hook } in + let evd = solve_remaining_evars flags env evd Evd.empty in + let ids = List.map RelDecl.get_name ctx in + check_name_freshness (pi1 kind) id; + (* XXX: The nf_evar is critical !! *) + evd, (id.CAst.v, + (Evarutil.nf_evar evd (EConstr.it_mkProd_or_LetIn t' ctx), + (ids, imps @ lift_implicits (Context.Rel.nhyps ctx) imps')))) + evd thms in + let recguard,thms,snl = look_for_possibly_mutual_statements evd thms in + let evd, _nf = Evarutil.nf_evars_and_universes evd in + (* XXX: This nf_evar is critical too!! We are normalizing twice if + you look at the previous lines... *) + let thms = List.map (fun (n, (t, info)) -> (n, (nf_evar evd t, info))) thms in + let () = + let open Misctypes in + if not (decl.univdecl_extensible_instance && decl.univdecl_extensible_constraints) then + ignore (Evd.check_univ_decl ~poly:(pi2 kind) evd decl) + in + let evd = + if pi2 kind then evd + else (* We fix the variables to ensure they won't be lowered to Set *) + Evd.fix_undefined_variables evd + in + start_proof_with_initialization kind evd decl recguard thms snl hook + +(* Saving a proof *) + +let keep_admitted_vars = ref true + +let _ = + let open Goptions in + declare_bool_option + { optdepr = false; + optname = "keep section variables in admitted proofs"; + optkey = ["Keep"; "Admitted"; "Variables"]; + optread = (fun () -> !keep_admitted_vars); + optwrite = (fun b -> keep_admitted_vars := b) } + +let save_proof ?proof = function + | Vernacexpr.Admitted -> + let pe = + let open Proof_global in + match proof with + | Some ({ id; entries; persistence = k; universes }, _) -> + if List.length entries <> 1 then + user_err Pp.(str "Admitted does not support multiple statements"); + let { const_entry_secctx; const_entry_type } = List.hd entries in + if const_entry_type = None then + user_err Pp.(str "Admitted requires an explicit statement"); + let typ = Option.get const_entry_type in + let ctx = UState.const_univ_entry ~poly:(pi2 k) universes in + let sec_vars = if !keep_admitted_vars then const_entry_secctx else None in + Admitted(id, k, (sec_vars, (typ, ctx), None), universes) + | None -> + let pftree = Proof_global.give_me_the_proof () in + let id, k, typ = Pfedit.current_proof_statement () in + let typ = EConstr.Unsafe.to_constr typ in + let universes = Proof.initial_euctx pftree in + (* This will warn if the proof is complete *) + let pproofs, _univs = + Proof_global.return_proof ~allow_partial:true () in + let sec_vars = + if not !keep_admitted_vars then None + else match Proof_global.get_used_variables(), pproofs with + | Some _ as x, _ -> x + | None, (pproof, _) :: _ -> + let env = Global.env () in + let ids_typ = Environ.global_vars_set env typ in + let ids_def = Environ.global_vars_set env pproof in + Some (Environ.keep_hyps env (Id.Set.union ids_typ ids_def)) + | _ -> None in + let decl = Proof_global.get_universe_decl () in + let poly = pi2 k in + let ctx = UState.check_univ_decl ~poly universes decl in + Admitted(id,k,(sec_vars, (typ, ctx), None), universes) + in + Proof_global.apply_terminator (Proof_global.get_terminator ()) pe + | Vernacexpr.Proved (is_opaque,idopt) -> + let (proof_obj,terminator) = + match proof with + | None -> + Proof_global.close_proof ~keep_body_ucst_separate:false (fun x -> x) + | Some proof -> proof + in + (* if the proof is given explicitly, nothing has to be deleted *) + if Option.is_empty proof then Proof_global.discard_current (); + Proof_global.(apply_terminator terminator (Proved (is_opaque,idopt,proof_obj))) + +(* Miscellaneous *) +let get_current_context () = Pfedit.get_current_context () + diff --git a/vernac/lemmas.mli b/vernac/lemmas.mli new file mode 100644 index 00000000..ad4c278e --- /dev/null +++ b/vernac/lemmas.mli @@ -0,0 +1,72 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* Globnames.global_reference -> 'a) -> 'a declaration_hook + +val call_hook : + Future.fix_exn -> 'a declaration_hook -> Decl_kinds.locality -> Globnames.global_reference -> 'a + +(** A hook start_proof calls on the type of the definition being started *) +val set_start_hook : (EConstr.types -> unit) -> unit + +val start_proof : Id.t -> ?pl:Univdecls.universe_decl -> goal_kind -> Evd.evar_map -> + ?terminator:(Proof_global.lemma_possible_guards -> unit declaration_hook -> Proof_global.proof_terminator) -> + ?sign:Environ.named_context_val -> EConstr.types -> + ?init_tac:unit Proofview.tactic -> ?compute_guard:Proof_global.lemma_possible_guards -> + unit declaration_hook -> unit + +val start_proof_univs : Id.t -> ?pl:Univdecls.universe_decl -> goal_kind -> Evd.evar_map -> + ?terminator:(Proof_global.lemma_possible_guards -> (UState.t option -> unit declaration_hook) -> Proof_global.proof_terminator) -> + ?sign:Environ.named_context_val -> EConstr.types -> + ?init_tac:unit Proofview.tactic -> ?compute_guard:Proof_global.lemma_possible_guards -> + (UState.t option -> unit declaration_hook) -> unit + +val start_proof_com : + ?inference_hook:Pretyping.inference_hook -> + goal_kind -> Vernacexpr.proof_expr list -> + unit declaration_hook -> unit + +val start_proof_with_initialization : + goal_kind -> Evd.evar_map -> Univdecls.universe_decl -> + (bool * Proof_global.lemma_possible_guards * unit Proofview.tactic list option) option -> + (Id.t (* name of thm *) * + (EConstr.types (* type of thm *) * (Name.t list (* names to pre-introduce *) * Impargs.manual_explicitation list))) list + -> int list option -> unit declaration_hook -> unit + +val universe_proof_terminator : + Proof_global.lemma_possible_guards -> + (UState.t option -> unit declaration_hook) -> + Proof_global.proof_terminator + +val standard_proof_terminator : + Proof_global.lemma_possible_guards -> unit declaration_hook -> + Proof_global.proof_terminator + +val fresh_name_for_anonymous_theorem : unit -> Id.t + +(** {6 ... } *) + +(** A hook the next three functions pass to cook_proof *) +val set_save_hook : (Proof.t -> unit) -> unit + +val save_proof : ?proof:Proof_global.closed_proof -> Vernacexpr.proof_end -> unit + + +(** [get_current_context ()] returns the evar context and env of the + current open proof if any, otherwise returns the empty evar context + and the current global env *) + +val get_current_context : unit -> Evd.evar_map * Environ.env +[@@ocaml.deprecated "please use [Pfedit.get_current_context]"] diff --git a/vernac/locality.ml b/vernac/locality.ml new file mode 100644 index 00000000..21be73b3 --- /dev/null +++ b/vernac/locality.ml @@ -0,0 +1,68 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* Local + | false -> Global + + +(** Positioning locality for commands supporting discharging and export + outside of modules *) + +(* For commands whose default is to discharge and export: + Global is the default and is neutral; + Local in a section deactivates discharge, + Local not in a section deactivates export *) +let make_non_locality = function Some false -> false | _ -> true + +let make_locality = function Some true -> true | _ -> false + +let enforce_locality_exp locality_flag discharge = + match locality_flag, discharge with + | Some b, NoDischarge -> local_of_bool b + | None, NoDischarge -> Global + | None, DoDischarge -> Discharge + | Some true, DoDischarge -> CErrors.user_err Pp.(str "Local not allowed in this case") + | Some false, DoDischarge -> CErrors.user_err Pp.(str "Global not allowed in this case") + +let enforce_locality locality_flag = + make_locality locality_flag + +(* For commands whose default is to not discharge but to export: + Global in sections forces discharge, Global not in section is the default; + Local in sections is the default, Local not in section forces non-export *) + +let make_section_locality = + function Some b -> b | None -> Lib.sections_are_opened () + +let enforce_section_locality locality_flag = + make_section_locality locality_flag + +(** Positioning locality for commands supporting export but not discharge *) + +(* For commands whose default is to export (if not in section): + Global in sections is forbidden, Global not in section is neutral; + Local in sections is the default, Local not in section forces non-export *) + +let make_module_locality = function + | Some false -> + if Lib.sections_are_opened () then + CErrors.user_err Pp.(str + "This command does not support the Global option in sections."); + false + | Some true -> true + | None -> false + +let enforce_module_locality locality_flag = + make_module_locality locality_flag diff --git a/vernac/locality.mli b/vernac/locality.mli new file mode 100644 index 00000000..3c63c821 --- /dev/null +++ b/vernac/locality.mli @@ -0,0 +1,40 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* bool +val make_non_locality : bool option -> bool +val enforce_locality_exp : bool option -> Decl_kinds.discharge -> Decl_kinds.locality +val enforce_locality : bool option -> bool + +(** For commands whose default is to not discharge but to export: + Global in sections forces discharge, Global not in section is the default; + Local in sections is the default, Local not in section forces non-export *) + +val make_section_locality : bool option -> bool +val enforce_section_locality : bool option -> bool + +(** * Positioning locality for commands supporting export but not discharge *) + +(** For commands whose default is to export (if not in section): + Global in sections is forbidden, Global not in section is neutral; + Local in sections is the default, Local not in section forces non-export *) + +val make_module_locality : bool option -> bool +val enforce_module_locality : bool option -> bool diff --git a/vernac/metasyntax.ml b/vernac/metasyntax.ml new file mode 100644 index 00000000..8c9d8f6b --- /dev/null +++ b/vernac/metasyntax.ml @@ -0,0 +1,1551 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* obj = + declare_object {(default_object "TOKEN") with + open_function = (fun i o -> if Int.equal i 1 then cache_token o); + cache_function = cache_token; + subst_function = Libobject.ident_subst_function; + classify_function = (fun o -> Substitute o)} + +let add_token_obj s = Lib.add_anonymous_leaf (inToken s) + +(**********************************************************************) +(* Printing grammar entries *) + +let entry_buf = Buffer.create 64 + +let pr_entry e = + let () = Buffer.clear entry_buf in + let ft = Format.formatter_of_buffer entry_buf in + let () = Pcoq.Gram.entry_print ft e in + str (Buffer.contents entry_buf) + +let pr_registered_grammar name = + let gram = try Some (Pcoq.find_grammars_by_name name) with Not_found -> None in + match gram with + | None -> user_err Pp.(str "Unknown or unprintable grammar entry.") + | Some entries -> + let pr_one (Pcoq.AnyEntry e) = + str "Entry " ++ str (Pcoq.Gram.Entry.name e) ++ str " is" ++ fnl () ++ + pr_entry e + in + prlist pr_one entries + +let pr_grammar = function + | "constr" | "operconstr" | "binder_constr" -> + str "Entry constr is" ++ fnl () ++ + pr_entry Pcoq.Constr.constr ++ + str "and lconstr is" ++ fnl () ++ + pr_entry Pcoq.Constr.lconstr ++ + str "where binder_constr is" ++ fnl () ++ + pr_entry Pcoq.Constr.binder_constr ++ + str "and operconstr is" ++ fnl () ++ + pr_entry Pcoq.Constr.operconstr + | "pattern" -> + pr_entry Pcoq.Constr.pattern + | "vernac" -> + str "Entry vernac_control is" ++ fnl () ++ + pr_entry Pcoq.Vernac_.vernac_control ++ + str "Entry command is" ++ fnl () ++ + pr_entry Pcoq.Vernac_.command ++ + str "Entry syntax is" ++ fnl () ++ + pr_entry Pcoq.Vernac_.syntax ++ + str "Entry gallina is" ++ fnl () ++ + pr_entry Pcoq.Vernac_.gallina ++ + str "Entry gallina_ext is" ++ fnl () ++ + pr_entry Pcoq.Vernac_.gallina_ext + | name -> pr_registered_grammar name + +(**********************************************************************) +(* Parse a format (every terminal starting with a letter or a single + quote (except a single quote alone) must be quoted) *) + +let parse_format ({CAst.loc;v=str} : Misctypes.lstring) = + let len = String.length str in + (* TODO: update the line of the location when the string contains newlines *) + let make_loc i j = Option.map (Loc.shift_loc (i+1) (j-len)) loc in + let push_token loc a = function + | (i,cur)::l -> (i,(loc,a)::cur)::l + | [] -> assert false in + let push_white i n l = + if Int.equal n 0 then l else push_token (make_loc i (i+n)) (UnpTerminal (String.make n ' ')) l in + let close_box start stop b = function + | (_,a)::(_::_ as l) -> push_token (make_loc start stop) (UnpBox (b,a)) l + | [a] -> user_err ?loc:(make_loc start stop) Pp.(str "Non terminated box in format.") + | [] -> assert false in + let close_quotation start i = + if i < len && str.[i] == '\'' then + if (Int.equal (i+1) len || str.[i+1] == ' ') + then i+1 + else user_err ?loc:(make_loc (i+1) (i+1)) Pp.(str "Space expected after quoted expression.") + else + user_err ?loc:(make_loc start (i-1)) Pp.(str "Beginning of quoted expression expected to be ended by a quote.") in + let rec spaces n i = + if i < len && str.[i] == ' ' then spaces (n+1) (i+1) + else n in + let rec nonspaces quoted n i = + if i < len && str.[i] != ' ' then + if str.[i] == '\'' && quoted && + (i+1 >= len || str.[i+1] == ' ') + then if Int.equal n 0 then user_err ?loc:(make_loc (i-1) i) Pp.(str "Empty quoted token.") else n + else nonspaces quoted (n+1) (i+1) + else + if quoted then user_err ?loc:(make_loc i i) Pp.(str "Spaces are not allowed in (quoted) symbols.") + else n in + let rec parse_non_format i = + let n = nonspaces false 0 i in + push_token (make_loc i (i+n-1)) (UnpTerminal (String.sub str i n)) (parse_token 1 (i+n)) + and parse_quoted n i = + if i < len then match str.[i] with + (* Parse " // " *) + | '/' when i+1 < len && str.[i+1] == '/' -> + (* We discard the useless n spaces... *) + push_token (make_loc (i-n) (i+1)) (UnpCut PpFnl) + (parse_token 1 (close_quotation i (i+2))) + (* Parse " .. / .. " *) + | '/' when i+1 < len -> + let p = spaces 0 (i+1) in + push_token (make_loc (i-n) (i+p)) (UnpCut (PpBrk (n,p))) + (parse_token 1 (close_quotation i (i+p+1))) + | c -> + (* The spaces are real spaces *) + push_white i n (match c with + | '[' -> + if i+1 < len then match str.[i+1] with + (* Parse " [h .. ", *) + | 'h' when i+1 <= len && str.[i+2] == 'v' -> + (parse_box i (fun n -> PpHVB n) (i+3)) + (* Parse " [v .. ", *) + | 'v' -> + parse_box i (fun n -> PpVB n) (i+2) + (* Parse " [ .. ", *) + | ' ' | '\'' -> + parse_box i (fun n -> PpHOVB n) (i+1) + | _ -> user_err ?loc:(make_loc i i) Pp.(str "\"v\", \"hv\", \" \" expected after \"[\" in format.") + else user_err ?loc:(make_loc i i) Pp.(str "\"v\", \"hv\" or \" \" expected after \"[\" in format.") + (* Parse "]" *) + | ']' -> + ((i,[]) :: parse_token 1 (close_quotation i (i+1))) + (* Parse a non formatting token *) + | c -> + let n = nonspaces true 0 i in + push_token (make_loc i (i+n-1)) (UnpTerminal (String.sub str (i-1) (n+2))) + (parse_token 1 (close_quotation i (i+n)))) + else + if Int.equal n 0 then [] + else user_err ?loc:(make_loc (len-n) len) Pp.(str "Ending spaces non part of a format annotation.") + and parse_box start box i = + let n = spaces 0 i in + close_box start (i+n-1) (box n) (parse_token 1 (close_quotation i (i+n))) + and parse_token k i = + let n = spaces 0 i in + let i = i+n in + if i < len then match str.[i] with + (* Parse a ' *) + | '\'' when i+1 >= len || str.[i+1] == ' ' -> + push_white (i-n) (n-k) (push_token (make_loc i (i+1)) (UnpTerminal "'") (parse_token 1 (i+1))) + (* Parse the beginning of a quoted expression *) + | '\'' -> + parse_quoted (n-k) (i+1) + (* Otherwise *) + | _ -> + push_white (i-n) (n-k) (parse_non_format i) + else push_white (i-n) n [(len,[])] + in + if not (String.is_empty str) then + match parse_token 0 0 with + | [_,l] -> l + | (i,_)::_ -> user_err ?loc:(make_loc i i) Pp.(str "Box closed without being opened.") + | [] -> assert false + else + [] + +(***********************) +(* Analyzing notations *) + +(* Interpret notations with a recursive component *) + +let out_nt = function NonTerminal x -> x | _ -> assert false + +let msg_expected_form_of_recursive_notation = + "In the notation, the special symbol \"..\" must occur in\na configuration of the form \"x symbs .. symbs y\"." + +let rec find_pattern nt xl = function + | Break n as x :: l, Break n' :: l' when Int.equal n n' -> + find_pattern nt (x::xl) (l,l') + | Terminal s as x :: l, Terminal s' :: l' when String.equal s s' -> + find_pattern nt (x::xl) (l,l') + | [], NonTerminal x' :: l' -> + (out_nt nt,x',List.rev xl),l' + | _, Break s :: _ | Break s :: _, _ -> + user_err Pp.(str ("A break occurs on one side of \"..\" but not on the other side.")) + | _, Terminal s :: _ | Terminal s :: _, _ -> + user_err ~hdr:"Metasyntax.find_pattern" + (str "The token \"" ++ str s ++ str "\" occurs on one side of \"..\" but not on the other side.") + | _, [] -> + user_err Pp.(str msg_expected_form_of_recursive_notation) + | ((SProdList _ | NonTerminal _) :: _), _ | _, (SProdList _ :: _) -> + anomaly (Pp.str "Only Terminal or Break expected on left, non-SProdList on right.") + +let rec interp_list_parser hd = function + | [] -> [], List.rev hd + | NonTerminal id :: tl when Id.equal id ldots_var -> + if List.is_empty hd then user_err Pp.(str msg_expected_form_of_recursive_notation); + let hd = List.rev hd in + let ((x,y,sl),tl') = find_pattern (List.hd hd) [] (List.tl hd,tl) in + let xyl,tl'' = interp_list_parser [] tl' in + (* We remember each pair of variable denoting a recursive part to *) + (* remove the second copy of it afterwards *) + (x,y)::xyl, SProdList (x,sl) :: tl'' + | (Terminal _ | Break _) as s :: tl -> + if List.is_empty hd then + let yl,tl' = interp_list_parser [] tl in + yl, s :: tl' + else + interp_list_parser (s::hd) tl + | NonTerminal _ as x :: tl -> + let xyl,tl' = interp_list_parser [x] tl in + xyl, List.rev_append hd tl' + | SProdList _ :: _ -> anomaly (Pp.str "Unexpected SProdList in interp_list_parser.") + + +(* Find non-terminal tokens of notation *) + +(* To protect alphabetic tokens and quotes from being seen as variables *) +let quote_notation_token x = + let n = String.length x in + let norm = CLexer.is_ident x in + if (n > 0 && norm) || (n > 2 && x.[0] == '\'') then "'"^x^"'" + else x + +let is_numeral symbs = + match List.filter (function Break _ -> false | _ -> true) symbs with + | ([Terminal "-"; Terminal x] | [Terminal x]) -> + (try let _ = Bigint.of_string x in true with Failure _ -> false) + | _ -> + false + +let rec get_notation_vars onlyprint = function + | [] -> [] + | NonTerminal id :: sl -> + let vars = get_notation_vars onlyprint sl in + if Id.equal id ldots_var then vars else + (* don't check for nonlinearity if printing only, see Bug 5526 *) + if not onlyprint && Id.List.mem id vars then + user_err ~hdr:"Metasyntax.get_notation_vars" + (str "Variable " ++ Id.print id ++ str " occurs more than once.") + else id::vars + | (Terminal _ | Break _) :: sl -> get_notation_vars onlyprint sl + | SProdList _ :: _ -> assert false + +let analyze_notation_tokens ~onlyprint ntn = + let l = decompose_raw_notation ntn in + let vars = get_notation_vars onlyprint l in + let recvars,l = interp_list_parser [] l in + recvars, List.subtract Id.equal vars (List.map snd recvars), l + +let error_not_same_scope x y = + user_err ~hdr:"Metasyntax.error_not_name_scope" + (str "Variables " ++ Id.print x ++ str " and " ++ Id.print y ++ str " must be in the same scope.") + +(**********************************************************************) +(* Build pretty-printing rules *) + +let prec_assoc = function + | RightA -> (L,E) + | LeftA -> (E,L) + | NonA -> (L,L) + +let precedence_of_position_and_level from = function + | NumLevel n, BorderProd (_,None) -> n, Prec n + | NumLevel n, BorderProd (b,Some a) -> + n, let (lp,rp) = prec_assoc a in if b == Left then lp else rp + | NumLevel n, InternalProd -> n, Prec n + | NextLevel, _ -> from, L + +let precedence_of_entry_type from = function + | ETConstr x | ETConstrAsBinder (_,x) -> precedence_of_position_and_level from x + | ETPattern (_,n) -> let n = match n with None -> 0 | Some n -> n in n, Prec n + | _ -> 0, E (* should not matter *) + +(* Some breaking examples *) +(* "x = y" : "x /1 = y" (breaks before any symbol) *) +(* "x =S y" : "x /1 =S /1 y" (protect from confusion; each side for symmetry)*) +(* "+ {" : "+ {" may breaks reversibility without space but oth. not elegant *) +(* "x y" : "x spc y" *) +(* "{ x } + { y }" : "{ x } / + { y }" *) +(* "< x , y > { z , t }" : "< x , / y > / { z , / t }" *) + +let starts_with_left_bracket s = + let l = String.length s in not (Int.equal l 0) && + (s.[0] == '{' || s.[0] == '[' || s.[0] == '(') + +let ends_with_right_bracket s = + let l = String.length s in not (Int.equal l 0) && + (s.[l-1] == '}' || s.[l-1] == ']' || s.[l-1] == ')') + +let is_left_bracket s = + starts_with_left_bracket s && not (ends_with_right_bracket s) + +let is_right_bracket s = + not (starts_with_left_bracket s) && ends_with_right_bracket s + +let is_comma s = + let l = String.length s in not (Int.equal l 0) && + (s.[0] == ',' || s.[0] == ';') + +let is_operator s = + let l = String.length s in not (Int.equal l 0) && + (s.[0] == '+' || s.[0] == '*' || s.[0] == '=' || + s.[0] == '-' || s.[0] == '/' || s.[0] == '<' || s.[0] == '>' || + s.[0] == '@' || s.[0] == '\\' || s.[0] == '&' || s.[0] == '~' || s.[0] == '$') + +let is_non_terminal = function + | NonTerminal _ | SProdList _ -> true + | _ -> false + +let is_next_non_terminal b = function +| [] -> b +| pr :: _ -> is_non_terminal pr + +let is_next_terminal = function Terminal _ :: _ -> true | _ -> false + +let is_next_break = function Break _ :: _ -> true | _ -> false + +let add_break n l = (None,UnpCut (PpBrk(n,0))) :: l + +let add_break_if_none n b = function + | (_,UnpCut (PpBrk _)) :: _ as l -> l + | [] when not b -> [] + | l -> (None,UnpCut (PpBrk(n,0))) :: l + +let check_open_binder isopen sl m = + let pr_token = function + | Terminal s -> str s + | Break n -> str "␣" + | _ -> assert false + in + if isopen && not (List.is_empty sl) then + user_err (str "as " ++ Id.print m ++ + str " is a non-closed binder, no such \"" ++ + prlist_with_sep spc pr_token sl + ++ strbrk "\" is allowed to occur.") + +let unparsing_metavar i from typs = + let x = List.nth typs (i-1) in + let prec = snd (precedence_of_entry_type from x) in + match x with + | ETConstr _ | ETConstrAsBinder _ | ETReference | ETBigint -> + UnpMetaVar (i,prec) + | ETPattern _ -> + UnpBinderMetaVar (i,prec) + | ETName -> + UnpBinderMetaVar (i,Prec 0) + | ETBinder isopen -> + assert false + | ETOther _ -> failwith "TODO" + +(* Heuristics for building default printing rules *) + +let index_id id l = List.index Id.equal id l + +let make_hunks etyps symbols from = + let vars,typs = List.split etyps in + let rec make b = function + | NonTerminal m :: prods -> + let i = index_id m vars in + let u = unparsing_metavar i from typs in + if is_next_non_terminal b prods then + (None, u) :: add_break_if_none 1 b (make b prods) + else + (None, u) :: make_with_space b prods + | Terminal s :: prods + when (* true to simulate presence of non-terminal *) b || List.exists is_non_terminal prods -> + if (is_comma s || is_operator s) then + (* Always a breakable space after comma or separator *) + (None, UnpTerminal s) :: add_break_if_none 1 b (make b prods) + else if is_right_bracket s && is_next_terminal prods then + (* Always no space after right bracked, but possibly a break *) + (None, UnpTerminal s) :: add_break_if_none 0 b (make b prods) + else if is_left_bracket s && is_next_non_terminal b prods then + (None, UnpTerminal s) :: make b prods + else if not (is_next_break prods) then + (* Add rigid space, no break, unless user asked for something *) + (None, UnpTerminal (s^" ")) :: make b prods + else + (* Rely on user spaces *) + (None, UnpTerminal s) :: make b prods + + | Terminal s :: prods -> + (* Separate but do not cut a trailing sequence of terminal *) + (match prods with + | Terminal _ :: _ -> (None,UnpTerminal (s^" ")) :: make b prods + | _ -> (None,UnpTerminal s) :: make b prods) + + | Break n :: prods -> + add_break n (make b prods) + + | SProdList (m,sl) :: prods -> + let i = index_id m vars in + let typ = List.nth typs (i-1) in + let _,prec = precedence_of_entry_type from typ in + let sl' = + (* If no separator: add a break *) + if List.is_empty sl then add_break 1 [] + (* We add NonTerminal for simulation but remove it afterwards *) + else make true sl in + let hunk = match typ with + | ETConstr _ -> UnpListMetaVar (i,prec,List.map snd sl') + | ETBinder isopen -> + check_open_binder isopen sl m; + UnpBinderListMetaVar (i,isopen,List.map snd sl') + | _ -> assert false in + (None, hunk) :: make_with_space b prods + + | [] -> [] + + and make_with_space b prods = + match prods with + | Terminal s' :: prods'-> + if is_operator s' then + (* A rigid space before operator and a breakable after *) + (None,UnpTerminal (" "^s')) :: add_break_if_none 1 b (make b prods') + else if is_comma s' then + (* No space whatsoever before comma *) + make b prods + else if is_right_bracket s' then + make b prods + else + (* A breakable space between any other two terminals *) + add_break_if_none 1 b (make b prods) + | (NonTerminal _ | SProdList _) :: _ -> + (* A breakable space before a non-terminal *) + add_break_if_none 1 b (make b prods) + | Break _ :: _ -> + (* Rely on user wish *) + make b prods + | [] -> [] + + in make false symbols + +(* Build default printing rules from explicit format *) + +let error_format ?loc () = user_err ?loc Pp.(str "The format does not match the notation.") + +let warn_format_break = + CWarnings.create ~name:"notation-both-format-and-spaces" ~category:"parsing" + (fun () -> + strbrk "Discarding format implicitly indicated by multiple spaces in notation because an explicit format modifier is given.") + +let rec split_format_at_ldots hd = function + | (loc,UnpTerminal s) :: fmt when String.equal s (Id.to_string Notation_ops.ldots_var) -> loc, List.rev hd, fmt + | u :: fmt -> + check_no_ldots_in_box u; + split_format_at_ldots (u::hd) fmt + | [] -> raise Exit + +and check_no_ldots_in_box = function + | (_,UnpBox (_,fmt)) -> + (try + let loc,_,_ = split_format_at_ldots [] fmt in + user_err ?loc Pp.(str ("The special symbol \"..\" must occur at the same formatting depth than the variables of which it is the ellipse.")) + with Exit -> ()) + | _ -> () + +let error_not_same ?loc () = + user_err ?loc Pp.(str "The format is not the same on the right- and left-hand sides of the special token \"..\".") + +let find_prod_list_loc sfmt fmt = + (* [fmt] is some [UnpTerminal x :: sfmt @ UnpTerminal ".." :: sfmt @ UnpTerminal y :: rest] *) + if List.is_empty sfmt then + (* No separators; we highlight the sequence "x .." *) + Loc.merge_opt (fst (List.hd fmt)) (fst (List.hd (List.tl fmt))) + else + (* A separator; we highlight the separating sequence *) + Loc.merge_opt (fst (List.hd sfmt)) (fst (List.last sfmt)) + +let skip_var_in_recursive_format = function + | (_,UnpTerminal s) :: sl (* skip first var *) when not (List.for_all (fun c -> c = " ") (String.explode s)) -> + (* To do, though not so important: check that the names match + the names in the notation *) + sl + | (loc,_) :: _ -> error_not_same ?loc () + | [] -> assert false + +let read_recursive_format sl fmt = + (* Turn [[UnpTerminal s :: some-list @ UnpTerminal ".." :: same-some-list @ UnpTerminal s' :: rest] *) + (* into [(some-list,rest)] *) + let get_head fmt = + let sl = skip_var_in_recursive_format fmt in + try split_format_at_ldots [] sl with Exit -> error_not_same ?loc:(fst (List.last (if sl = [] then fmt else sl))) () in + let rec get_tail = function + | (loc,a) :: sepfmt, (_,b) :: fmt when Pervasives.(=) a b -> get_tail (sepfmt, fmt) (* FIXME *) + | [], tail -> skip_var_in_recursive_format tail + | (loc,_) :: _, ([] | (_,UnpTerminal _) :: _)-> error_not_same ?loc () + | _, (loc,_)::_ -> error_not_same ?loc () in + let loc, slfmt, fmt = get_head fmt in + slfmt, get_tail (slfmt, fmt) + +let hunks_of_format (from,(vars,typs)) symfmt = + let rec aux = function + | symbs, (_,(UnpTerminal s' as u)) :: fmt + when String.equal s' (String.make (String.length s') ' ') -> + let symbs, l = aux (symbs,fmt) in symbs, u :: l + | Terminal s :: symbs, (_,UnpTerminal s') :: fmt + when String.equal s (String.drop_simple_quotes s') -> + let symbs, l = aux (symbs,fmt) in symbs, UnpTerminal s :: l + | NonTerminal s :: symbs, (_,UnpTerminal s') :: fmt when Id.equal s (Id.of_string s') -> + let i = index_id s vars in + let symbs, l = aux (symbs,fmt) in symbs, unparsing_metavar i from typs :: l + | symbs, (_,UnpBox (a,b)) :: fmt -> + let symbs', b' = aux (symbs,b) in + let symbs', l = aux (symbs',fmt) in + symbs', UnpBox (a,List.map (fun x -> (None,x)) b') :: l + | symbs, (_,(UnpCut _ as u)) :: fmt -> + let symbs, l = aux (symbs,fmt) in symbs, u :: l + | SProdList (m,sl) :: symbs, fmt -> + let i = index_id m vars in + let typ = List.nth typs (i-1) in + let _,prec = precedence_of_entry_type from typ in + let loc_slfmt,rfmt = read_recursive_format sl fmt in + let sl, slfmt = aux (sl,loc_slfmt) in + if not (List.is_empty sl) then error_format ?loc:(find_prod_list_loc loc_slfmt fmt) (); + let symbs, l = aux (symbs,rfmt) in + let hunk = match typ with + | ETConstr _ -> UnpListMetaVar (i,prec,slfmt) + | ETBinder isopen -> + check_open_binder isopen sl m; + UnpBinderListMetaVar (i,isopen,slfmt) + | _ -> assert false in + symbs, hunk :: l + | symbs, [] -> symbs, [] + | Break _ :: symbs, fmt -> warn_format_break (); aux (symbs,fmt) + | _, fmt -> error_format ?loc:(fst (List.hd fmt)) () + in + match aux symfmt with + | [], l -> l + | _ -> error_format () + +(**********************************************************************) +(* Build parsing rules *) + +let assoc_of_type n (_,typ) = precedence_of_entry_type n typ + +let is_not_small_constr = function + ETProdConstr _ -> true + | ETProdOther("constr","binder_constr") -> true + | _ -> false + +let rec define_keywords_aux = function + | GramConstrNonTerminal(e,Some _) as n1 :: GramConstrTerminal(IDENT k) :: l + when is_not_small_constr e -> + Flags.if_verbose Feedback.msg_info (str "Identifier '" ++ str k ++ str "' now a keyword"); + CLexer.add_keyword k; + n1 :: GramConstrTerminal(KEYWORD k) :: define_keywords_aux l + | n :: l -> n :: define_keywords_aux l + | [] -> [] + + (* Ensure that IDENT articulation terminal symbols are keywords *) +let define_keywords = function + | GramConstrTerminal(IDENT k)::l -> + Flags.if_verbose Feedback.msg_info (str "Identifier '" ++ str k ++ str "' now a keyword"); + CLexer.add_keyword k; + GramConstrTerminal(KEYWORD k) :: define_keywords_aux l + | l -> define_keywords_aux l + +let distribute a ll = List.map (fun l -> a @ l) ll + + (* Expand LIST1(t,sep);sep;t;...;t (with the trailing pattern + occurring p times, possibly p=0) into the combination of + t;sep;t;...;t;sep;t (p+1 times) + t;sep;t;...;t;sep;t;sep;t (p+2 times) + ... + t;sep;t;...;t;sep;t;...;t;sep;t (p+n times) + t;sep;t;...;t;sep;t;...;t;sep;t;LIST1(t,sep) *) + +let expand_list_rule typ tkl x n p ll = + let camlp5_message_name = Some (add_suffix x ("_"^string_of_int n)) in + let main = GramConstrNonTerminal (ETProdConstr typ, camlp5_message_name) in + let tks = List.map (fun x -> GramConstrTerminal x) tkl in + let rec aux i hds ll = + if i < p then aux (i+1) (main :: tks @ hds) ll + else if Int.equal i (p+n) then + let hds = + GramConstrListMark (p+n,true,p) :: hds + @ [GramConstrNonTerminal (ETProdConstrList (typ,tkl), Some x)] in + distribute hds ll + else + distribute (GramConstrListMark (i+1,false,p) :: hds @ [main]) ll @ + aux (i+1) (main :: tks @ hds) ll in + aux 0 [] ll + +let is_constr_typ typ x etyps = + match List.assoc x etyps with + | ETConstr typ' | ETConstrAsBinder (_,typ') -> typ = typ' + | _ -> false + +let include_possible_similar_trailing_pattern typ etyps sl l = + let rec aux n = function + | Terminal s :: sl, Terminal s'::l' when s = s' -> aux n (sl,l') + | [], NonTerminal x ::l' when is_constr_typ typ x etyps -> try_aux n l' + | _ -> raise Exit + and try_aux n l = + try aux (n+1) (sl,l) + with Exit -> n,l in + try_aux 0 l + +let prod_entry_type = function + | ETName -> ETProdName + | ETReference -> ETProdReference + | ETBigint -> ETProdBigint + | ETBinder _ -> assert false (* See check_binder_type *) + | ETConstr p | ETConstrAsBinder (_,p) -> ETProdConstr p + | ETPattern (_,n) -> ETProdPattern (match n with None -> 0 | Some n -> n) + | ETOther (s,t) -> ETProdOther (s,t) + +let make_production etyps symbols = + let rec aux = function + | [] -> [[]] + | NonTerminal m :: l -> + let typ = List.assoc m etyps in + distribute [GramConstrNonTerminal (prod_entry_type typ, Some m)] (aux l) + | Terminal s :: l -> + distribute [GramConstrTerminal (CLexer.terminal s)] (aux l) + | Break _ :: l -> + aux l + | SProdList (x,sl) :: l -> + let tkl = List.flatten + (List.map (function Terminal s -> [CLexer.terminal s] + | Break _ -> [] + | _ -> anomaly (Pp.str "Found a non terminal token in recursive notation separator.")) sl) in + match List.assoc x etyps with + | ETConstr typ -> + let p,l' = include_possible_similar_trailing_pattern typ etyps sl l in + expand_list_rule typ tkl x 1 p (aux l') + | ETBinder o -> + check_open_binder o sl x; + let typ = if o then (assert (tkl = []); ETBinderOpen) else ETBinderClosed tkl in + distribute + [GramConstrNonTerminal (ETProdBinderList typ, Some x)] (aux l) + | _ -> + user_err Pp.(str "Components of recursive patterns in notation must be terms or binders.") in + let prods = aux symbols in + List.map define_keywords prods + +let rec find_symbols c_current c_next c_last = function + | [] -> [] + | NonTerminal id :: sl -> + let prec = if not (List.is_empty sl) then c_current else c_last in + (id, prec) :: (find_symbols c_next c_next c_last sl) + | Terminal s :: sl -> find_symbols c_next c_next c_last sl + | Break n :: sl -> find_symbols c_current c_next c_last sl + | SProdList (x,_) :: sl' -> + (x,c_next)::(find_symbols c_next c_next c_last sl') + +let border = function + | (_,ETConstr(_,BorderProd (_,a))) :: _ -> a + | (_,(ETConstrAsBinder(_,(_,BorderProd (_,a))))) :: _ -> a + | _ -> None + +let recompute_assoc typs = + match border typs, border (List.rev typs) with + | Some LeftA, Some RightA -> assert false + | Some LeftA, _ -> Some LeftA + | _, Some RightA -> Some RightA + | _ -> None + +(**************************************************************************) +(* Registration of syntax extensions (parsing/printing, no interpretation)*) + +let pr_arg_level from (lev,typ) = + let pplev = function + | (n,L) when Int.equal n from -> str "at next level" + | (n,E) -> str "at level " ++ int n + | (n,L) -> str "at level below " ++ int n + | (n,Prec m) when Int.equal m n -> str "at level " ++ int n + | (n,_) -> str "Unknown level" in + Ppvernac.pr_set_entry_type (fun _ -> (*TO CHECK*) mt()) typ ++ + (match typ with + | ETConstr _ | ETConstrAsBinder _ | ETPattern _ -> spc () ++ pplev lev + | _ -> mt ()) + +let pr_level ntn (from,args,typs) = + str "at level " ++ int from ++ spc () ++ str "with arguments" ++ spc() ++ + prlist_with_sep pr_comma (pr_arg_level from) (List.combine args typs) + +let error_incompatible_level ntn oldprec prec = + user_err + (str "Notation " ++ qstring ntn ++ str " is already defined" ++ spc() ++ + pr_level ntn oldprec ++ + spc() ++ str "while it is now required to be" ++ spc() ++ + pr_level ntn prec ++ str ".") + +let error_parsing_incompatible_level ntn ntn' oldprec prec = + user_err + (str "Notation " ++ qstring ntn ++ str " relies on a parsing rule for " ++ qstring ntn' ++ spc() ++ + str " which is already defined" ++ spc() ++ + pr_level ntn oldprec ++ + spc() ++ str "while it is now required to be" ++ spc() ++ + pr_level ntn prec ++ str ".") + +type syntax_extension = { + synext_level : Notation_term.level; + synext_notation : notation; + synext_notgram : notation_grammar; + synext_unparsing : unparsing list; + synext_extra : (string * string) list; + synext_compat : Flags.compat_version option; +} + +let is_active_compat = function +| None -> true +| Some v -> 0 <= Flags.version_compare v !Flags.compat_version + +type syntax_extension_obj = locality_flag * syntax_extension + +let check_and_extend_constr_grammar ntn rule = + try + let ntn_for_grammar = rule.notgram_notation in + if String.equal ntn ntn_for_grammar then raise Not_found; + let prec = rule.notgram_level in + let oldprec = Notation.level_of_notation ntn_for_grammar in + if not (Notation.level_eq prec oldprec) then error_parsing_incompatible_level ntn ntn_for_grammar oldprec prec; + with Not_found -> + Egramcoq.extend_constr_grammar rule + +let cache_one_syntax_extension se = + let ntn = se.synext_notation in + let prec = se.synext_level in + let onlyprint = se.synext_notgram.notgram_onlyprinting in + try + let oldprec = Notation.level_of_notation ~onlyprint ntn in + if not (Notation.level_eq prec oldprec) then error_incompatible_level ntn oldprec prec; + with Not_found -> + if is_active_compat se.synext_compat then begin + (* Reserve the notation level *) + Notation.declare_notation_level ntn prec ~onlyprint; + (* Declare the parsing rule *) + if not onlyprint then List.iter (check_and_extend_constr_grammar ntn) se.synext_notgram.notgram_rules; + (* Declare the notation rule *) + Notation.declare_notation_rule ntn + ~extra:se.synext_extra (se.synext_unparsing, pi1 prec) se.synext_notgram + end + +let cache_syntax_extension (_, (_, sy)) = + cache_one_syntax_extension sy + +let subst_parsing_rule subst x = x + +let subst_printing_rule subst x = x + +let subst_syntax_extension (subst, (local, sy)) = + (local, { sy with + synext_notgram = { sy.synext_notgram with notgram_rules = List.map (subst_parsing_rule subst) sy.synext_notgram.notgram_rules }; + synext_unparsing = subst_printing_rule subst sy.synext_unparsing; + }) + +let classify_syntax_definition (local, _ as o) = + if local then Dispose else Substitute o + +let inSyntaxExtension : syntax_extension_obj -> obj = + declare_object {(default_object "SYNTAX-EXTENSION") with + open_function = (fun i o -> if Int.equal i 1 then cache_syntax_extension o); + cache_function = cache_syntax_extension; + subst_function = subst_syntax_extension; + classify_function = classify_syntax_definition} + +(**************************************************************************) +(* Precedences *) + +(* Interpreting user-provided modifiers *) + +(* XXX: We could move this to the parser itself *) +module NotationMods = struct + +type notation_modifier = { + assoc : gram_assoc option; + level : int option; + etyps : (Id.t * simple_constr_prod_entry_key) list; + + (* common to syn_data below *) + only_parsing : bool; + only_printing : bool; + compat : Flags.compat_version option; + format : Misctypes.lstring option; + extra : (string * string) list; +} + +let default = { + assoc = None; + level = None; + etyps = []; + only_parsing = false; + only_printing = false; + compat = None; + format = None; + extra = []; +} + +end + +let interp_modifiers modl = let open NotationMods in + let rec interp acc = function + | [] -> acc + | SetEntryType (s,typ) :: l -> + let id = Id.of_string s in + if Id.List.mem_assoc id acc.etyps then + user_err ~hdr:"Metasyntax.interp_modifiers" + (str s ++ str " is already assigned to an entry or constr level."); + interp { acc with etyps = (id,typ) :: acc.etyps; } l + | SetItemLevel ([],n) :: l -> + interp acc l + | SetItemLevelAsBinder ([],_,_) :: l -> + interp acc l + | SetItemLevel (s::idl,n) :: l -> + let id = Id.of_string s in + if Id.List.mem_assoc id acc.etyps then + user_err ~hdr:"Metasyntax.interp_modifiers" + (str s ++ str " is already assigned to an entry or constr level."); + let typ = ETConstr (Some n) in + interp { acc with etyps = (id,typ)::acc.etyps; } (SetItemLevel (idl,n)::l) + | SetItemLevelAsBinder (s::idl,bk,n) :: l -> + let id = Id.of_string s in + if Id.List.mem_assoc id acc.etyps then + user_err ~hdr:"Metasyntax.interp_modifiers" + (str s ++ str " is already assigned to an entry or constr level."); + let typ = ETConstrAsBinder (bk,n) in + interp { acc with etyps = (id,typ)::acc.etyps; } (SetItemLevelAsBinder (idl,bk,n)::l) + | SetLevel n :: l -> + interp { acc with level = Some n; } l + | SetAssoc a :: l -> + if not (Option.is_empty acc.assoc) then user_err Pp.(str "An associativity is given more than once."); + interp { acc with assoc = Some a; } l + | SetOnlyParsing :: l -> + interp { acc with only_parsing = true; } l + | SetOnlyPrinting :: l -> + interp { acc with only_printing = true; } l + | SetCompatVersion v :: l -> + interp { acc with compat = Some v; } l + | SetFormat ("text",s) :: l -> + if not (Option.is_empty acc.format) then user_err Pp.(str "A format is given more than once."); + interp { acc with format = Some s; } l + | SetFormat (k,{CAst.v=s}) :: l -> + interp { acc with extra = (k,s)::acc.extra; } l + in interp default modl + +let check_infix_modifiers modifiers = + let t = (interp_modifiers modifiers).NotationMods.etyps in + if not (List.is_empty t) then + user_err Pp.(str "Explicit entry level or type unexpected in infix notation.") + +let check_useless_entry_types recvars mainvars etyps = + let vars = let (l1,l2) = List.split recvars in l1@l2@mainvars in + match List.filter (fun (x,etyp) -> not (List.mem x vars)) etyps with + | (x,_)::_ -> user_err ~hdr:"Metasyntax.check_useless_entry_types" + (Id.print x ++ str " is unbound in the notation.") + | _ -> () + +let check_binder_type recvars etyps = + let l1,l2 = List.split recvars in + let l = l1@l2 in + List.iter (function + | (x,ETBinder b) when not (List.mem x l) -> + CErrors.user_err (str (if b then "binder" else "closed binder") ++ + strbrk " is only for use in recursive notations for binders.") + | _ -> ()) etyps + +let not_a_syntax_modifier = function +| SetOnlyParsing -> true +| SetOnlyPrinting -> true +| SetCompatVersion _ -> true +| _ -> false + +let no_syntax_modifiers mods = List.for_all not_a_syntax_modifier mods + +let is_only_parsing mods = + let test = function SetOnlyParsing -> true | _ -> false in + List.exists test mods + +let is_only_printing mods = + let test = function SetOnlyPrinting -> true | _ -> false in + List.exists test mods + +let get_compat_version mods = + let test = function SetCompatVersion v -> Some v | _ -> None in + try Some (List.find_map test mods) with Not_found -> None + +(* Compute precedences from modifiers (or find default ones) *) + +let set_entry_type etyps (x,typ) = + let typ = try + match List.assoc x etyps, typ with + | ETConstr (Some n), (_,BorderProd (left,_)) -> + ETConstr (n,BorderProd (left,None)) + | ETConstr (Some n), (_,InternalProd) -> ETConstr (n,InternalProd) + | ETConstrAsBinder (bk, Some n), (_,BorderProd (left,_)) -> + ETConstrAsBinder (bk, (n,BorderProd (left,None))) + | ETConstrAsBinder (bk, Some n), (_,InternalProd) -> + ETConstrAsBinder (bk, (n,InternalProd)) + | ETPattern (b,n), _ -> ETPattern (b,n) + | (ETName | ETBigint | ETReference | ETBinder _ | ETOther _ as x), _ -> x + | ETConstr None, _ -> ETConstr typ + | ETConstrAsBinder (bk,None), _ -> ETConstrAsBinder (bk,typ) + with Not_found -> ETConstr typ + in (x,typ) + +let join_auxiliary_recursive_types recvars etyps = + List.fold_right (fun (x,y) typs -> + let xtyp = try Some (List.assoc x etyps) with Not_found -> None in + let ytyp = try Some (List.assoc y etyps) with Not_found -> None in + match xtyp,ytyp with + | None, None -> typs + | Some _, None -> typs + | None, Some ytyp -> (x,ytyp)::typs + | Some xtyp, Some ytyp when Pervasives.(=) xtyp ytyp -> typs (* FIXME *) + | Some xtyp, Some ytyp -> + user_err + (strbrk "In " ++ Id.print x ++ str " .. " ++ Id.print y ++ + strbrk ", both ends have incompatible types.")) + recvars etyps + +let internalization_type_of_entry_type = function + | ETBinder _ -> NtnInternTypeOnlyBinder + | ETConstr _ | ETConstrAsBinder _ | ETBigint | ETReference + | ETName | ETPattern _ | ETOther _ -> NtnInternTypeAny + +let set_internalization_type typs = + List.map (fun (_, e) -> internalization_type_of_entry_type e) typs + +let make_internalization_vars recvars mainvars typs = + let maintyps = List.combine mainvars typs in + let extratyps = List.map (fun (x,y) -> (y,List.assoc x maintyps)) recvars in + maintyps @ extratyps + +let make_interpretation_type isrec isonlybinding = function + | ETConstr _ -> + if isrec then NtnTypeConstrList else + if isonlybinding then + (* Parsed as constr, but interpreted as a binder: default is to parse it as an ident only *) + NtnTypeBinder (NtnBinderParsedAsConstr AsIdent) + else NtnTypeConstr + | ETConstrAsBinder (bk,_) -> NtnTypeBinder (NtnBinderParsedAsConstr bk) + | ETName -> NtnTypeBinder NtnParsedAsIdent + | ETPattern (ppstrict,_) -> NtnTypeBinder (NtnParsedAsPattern ppstrict) (* Parsed as ident/pattern, primarily interpreted as binder; maybe strict at printing *) + | ETBigint | ETReference | ETOther _ -> NtnTypeConstr + | ETBinder _ -> + if isrec then NtnTypeBinderList + else anomaly Pp.(str "Type binder is only for use in recursive notations for binders.") + +let make_interpretation_vars recvars allvars typs = + let eq_subscope (sc1, l1) (sc2, l2) = + Option.equal String.equal sc1 sc2 && + List.equal String.equal l1 l2 + in + let check (x, y) = + let (_,scope1) = Id.Map.find x allvars in + let (_,scope2) = Id.Map.find y allvars in + if not (eq_subscope scope1 scope2) then error_not_same_scope x y + in + let () = List.iter check recvars in + let useless_recvars = List.map snd recvars in + let mainvars = + Id.Map.filter (fun x _ -> not (Id.List.mem x useless_recvars)) allvars in + Id.Map.mapi (fun x (isonlybinding, sc) -> + (sc, make_interpretation_type (Id.List.mem_assoc x recvars) isonlybinding (Id.List.assoc x typs))) mainvars + +let check_rule_productivity l = + if List.for_all (function NonTerminal _ | Break _ -> true | _ -> false) l then + user_err Pp.(str "A notation must include at least one symbol."); + if (match l with SProdList _ :: _ -> true | _ -> false) then + user_err Pp.(str "A recursive notation must start with at least one symbol.") + +let warn_notation_bound_to_variable = + CWarnings.create ~name:"notation-bound-to-variable" ~category:"parsing" + (fun () -> + strbrk "This notation will not be used for printing as it is bound to a single variable.") + +let warn_non_reversible_notation = + CWarnings.create ~name:"non-reversible-notation" ~category:"parsing" + (function + | APrioriReversible -> assert false + | HasLtac -> + strbrk "This notation contains Ltac expressions: it will not be used for printing." + | NonInjective ids -> + let n = List.length ids in + strbrk (String.plural n "Variable") ++ spc () ++ pr_enum Id.print ids ++ spc () ++ + strbrk (if n > 1 then "do" else "does") ++ + str " not occur in the right-hand side." ++ spc() ++ + strbrk "The notation will not be used for printing as it is not reversible.") + +let is_not_printable onlyparse reversibility = function +| NVar _ -> + if not onlyparse then warn_notation_bound_to_variable (); + true +| _ -> + if not onlyparse && reversibility <> APrioriReversible then + (warn_non_reversible_notation reversibility; true) + else onlyparse + + +let find_precedence lev etyps symbols onlyprint = + let first_symbol = + let rec aux = function + | Break _ :: t -> aux t + | h :: t -> Some h + | [] -> None in + aux symbols in + let last_is_terminal () = + let rec aux b = function + | Break _ :: t -> aux b t + | Terminal _ :: t -> aux true t + | _ :: t -> aux false t + | [] -> b in + aux false symbols in + match first_symbol with + | None -> [],0 + | Some (NonTerminal x) -> + let test () = + if onlyprint then + if Option.is_empty lev then + user_err Pp.(str "Explicit level needed in only-printing mode when the level of the leftmost non-terminal is given.") + else [],Option.get lev + else + user_err Pp.(str "The level of the leftmost non-terminal cannot be changed.") in + (try match List.assoc x etyps with + | ETConstr _ -> test () + | ETConstrAsBinder (_,Some _) -> test () + | (ETName | ETBigint | ETReference) -> + begin match lev with + | None -> + ([Feedback.msg_info ?loc:None ,strbrk "Setting notation at level 0."],0) + | Some 0 -> + ([],0) + | _ -> + user_err Pp.(str "A notation starting with an atomic expression must be at level 0.") + end + | (ETPattern _ | ETBinder _ | ETOther _ | ETConstrAsBinder _) -> + (* Give a default ? *) + if Option.is_empty lev then + user_err Pp.(str "Need an explicit level.") + else [],Option.get lev + with Not_found -> + if Option.is_empty lev then + user_err Pp.(str "A left-recursive notation must have an explicit level.") + else [],Option.get lev) + | Some (Terminal _) when last_is_terminal () -> + if Option.is_empty lev then + ([Feedback.msg_info ?loc:None ,strbrk "Setting notation at level 0."], 0) + else [],Option.get lev + | Some _ -> + if Option.is_empty lev then user_err Pp.(str "Cannot determine the level."); + [],Option.get lev + +let check_curly_brackets_notation_exists () = + try let _ = Notation.level_of_notation "{ _ }" in () + with Not_found -> + user_err Pp.(str "Notations involving patterns of the form \"{ _ }\" are treated \n\ +specially and require that the notation \"{ _ }\" is already reserved.") + +(* Remove patterns of the form "{ _ }", unless it is the "{ _ }" notation *) +let remove_curly_brackets l = + let rec skip_break acc = function + | Break _ as br :: l -> skip_break (br::acc) l + | l -> List.rev acc, l in + let rec aux deb = function + | [] -> [] + | Terminal "{" as t1 :: l -> + let br,next = skip_break [] l in + (match next with + | NonTerminal _ as x :: l' -> + let br',next' = skip_break [] l' in + (match next' with + | Terminal "}" as t2 :: l'' -> + if deb && List.is_empty l'' then [t1;x;t2] else begin + check_curly_brackets_notation_exists (); + x :: aux false l'' + end + | l1 -> t1 :: br @ x :: br' @ aux false l1) + | l0 -> t1 :: aux false l0) + | x :: l -> x :: aux false l + in aux true l + +module SynData = struct + + type subentry_types = (Id.t * (production_level * production_position) constr_entry_key_gen) list + + (* XXX: Document *) + type syn_data = { + + (* Notation name and location *) + info : notation * notation_location; + + (* Fields coming from the vernac-level modifiers *) + only_parsing : bool; + only_printing : bool; + compat : Flags.compat_version option; + format : Misctypes.lstring option; + extra : (string * string) list; + + (* XXX: Callback to printing, must remove *) + msgs : ((Pp.t -> unit) * Pp.t) list; + + (* Fields for internalization *) + recvars : (Id.t * Id.t) list; + mainvars : Id.List.elt list; + intern_typs : notation_var_internalization_type list; + + (* Notation data for parsing *) + level : level; + pa_syntax_data : subentry_types * symbol list; + pp_syntax_data : subentry_types * symbol list; + not_data : notation * (* notation *) + level * (* level, precedence, types *) + bool; (* needs_squash *) + } + +end + +let find_subentry_types n assoc etyps symbols = + let innerlevel = NumLevel 200 in + let typs = + find_symbols + (NumLevel n,BorderProd(Left,assoc)) + (innerlevel,InternalProd) + (NumLevel n,BorderProd(Right,assoc)) + symbols in + let sy_typs = List.map (set_entry_type etyps) typs in + let prec = List.map (assoc_of_type n) sy_typs in + sy_typs, prec + +let compute_syntax_data df modifiers = + let open SynData in + let open NotationMods in + let mods = interp_modifiers modifiers in + let onlyprint = mods.only_printing in + let onlyparse = mods.only_parsing in + if onlyprint && onlyparse then user_err (str "A notation cannot be both 'only printing' and 'only parsing'."); + let assoc = Option.append mods.assoc (Some NonA) in + let (recvars,mainvars,symbols) = analyze_notation_tokens ~onlyprint df in + let _ = check_useless_entry_types recvars mainvars mods.etyps in + let _ = check_binder_type recvars mods.etyps in + + (* Notations for interp and grammar *) + let ntn_for_interp = make_notation_key symbols in + let symbols_for_grammar = remove_curly_brackets symbols in + let need_squash = not (List.equal Notation.symbol_eq symbols symbols_for_grammar) in + let ntn_for_grammar = if need_squash then make_notation_key symbols_for_grammar else ntn_for_interp in + if not onlyprint then check_rule_productivity symbols_for_grammar; + let msgs,n = find_precedence mods.level mods.etyps symbols onlyprint in + (* To globalize... *) + let etyps = join_auxiliary_recursive_types recvars mods.etyps in + let sy_typs, prec = + find_subentry_types n assoc etyps symbols in + let sy_typs_for_grammar, prec_for_grammar = + if need_squash then + find_subentry_types n assoc etyps symbols_for_grammar + else + sy_typs, prec in + let i_typs = set_internalization_type sy_typs in + let pa_sy_data = (sy_typs_for_grammar,symbols_for_grammar) in + let pp_sy_data = (sy_typs,symbols) in + let sy_fulldata = (ntn_for_grammar,(n,prec_for_grammar,List.map snd sy_typs_for_grammar),need_squash) in + let df' = ((Lib.library_dp(),Lib.current_dirpath true),df) in + let i_data = ntn_for_interp, df' in + + (* Return relevant data for interpretation and for parsing/printing *) + { info = i_data; + + only_parsing = mods.only_parsing; + only_printing = mods.only_printing; + compat = mods.compat; + format = mods.format; + extra = mods.extra; + + msgs; + + recvars; + mainvars; + intern_typs = i_typs; + + level = (n,prec,List.map snd sy_typs); + pa_syntax_data = pa_sy_data; + pp_syntax_data = pp_sy_data; + not_data = sy_fulldata; + } + +let compute_pure_syntax_data df mods = + let open SynData in + let sd = compute_syntax_data df mods in + let msgs = + if sd.only_parsing then + (Feedback.msg_warning ?loc:None, + strbrk "The only parsing modifier has no effect in Reserved Notation.")::sd.msgs + else sd.msgs in + { sd with msgs } + +(**********************************************************************) +(* Registration of notations interpretation *) + +type notation_obj = { + notobj_local : bool; + notobj_scope : scope_name option; + notobj_interp : interpretation; + notobj_onlyparse : bool; + notobj_onlyprint : bool; + notobj_compat : Flags.compat_version option; + notobj_notation : notation * notation_location; +} + +let load_notation _ (_, nobj) = + Option.iter Notation.declare_scope nobj.notobj_scope + +let open_notation i (_, nobj) = + let scope = nobj.notobj_scope in + let (ntn, df) = nobj.notobj_notation in + let pat = nobj.notobj_interp in + let onlyprint = nobj.notobj_onlyprint in + let fresh = not (Notation.exists_notation_in_scope scope ntn onlyprint pat) in + let active = is_active_compat nobj.notobj_compat in + if Int.equal i 1 && fresh && active then begin + (* Declare the interpretation *) + let () = Notation.declare_notation_interpretation ntn scope pat df ~onlyprint in + (* Declare the uninterpretation *) + if not nobj.notobj_onlyparse then + Notation.declare_uninterpretation (NotationRule (scope, ntn)) pat + end + +let cache_notation o = + load_notation 1 o; + open_notation 1 o + +let subst_notation (subst, nobj) = + { nobj with notobj_interp = subst_interpretation subst nobj.notobj_interp; } + +let classify_notation nobj = + if nobj.notobj_local then Dispose else Substitute nobj + +let inNotation : notation_obj -> obj = + declare_object {(default_object "NOTATION") with + open_function = open_notation; + cache_function = cache_notation; + subst_function = subst_notation; + load_function = load_notation; + classify_function = classify_notation} + +(**********************************************************************) + +let with_lib_stk_protection f x = + let fs = Lib.freeze ~marshallable:`No in + try let a = f x in Lib.unfreeze fs; a + with reraise -> + let reraise = CErrors.push reraise in + let () = Lib.unfreeze fs in + iraise reraise + +let with_syntax_protection f x = + with_lib_stk_protection + (Pcoq.with_grammar_rule_protection + (with_notation_protection f)) x + +(**********************************************************************) +(* Recovering existing syntax *) + +exception NoSyntaxRule + +let recover_notation_syntax ntn = + try + let prec = Notation.level_of_notation ~onlyprint:true ntn (* Be as little restrictive as possible *) in + let pp_rule,_ = Notation.find_notation_printing_rule ntn in + let pp_extra_rules = Notation.find_notation_extra_printing_rules ntn in + let pa_rule = Notation.find_notation_parsing_rules ntn in + { synext_level = prec; + synext_notation = ntn; + synext_notgram = pa_rule; + synext_unparsing = pp_rule; + synext_extra = pp_extra_rules; + synext_compat = None; + } + with Not_found -> + raise NoSyntaxRule + +let recover_squash_syntax sy = + let sq = recover_notation_syntax "{ _ }" in + sy :: sq.synext_notgram.notgram_rules + +(**********************************************************************) +(* Main entry point for building parsing and printing rules *) + +let make_pa_rule level (typs,symbols) ntn need_squash = + let assoc = recompute_assoc typs in + let prod = make_production typs symbols in + let sy = { + notgram_level = level; + notgram_assoc = assoc; + notgram_notation = ntn; + notgram_prods = prod; + } in + (* By construction, the rule for "{ _ }" is declared, but we need to + redeclare it because the file where it is declared needs not be open + when the current file opens (especially in presence of -nois) *) + if need_squash then recover_squash_syntax sy else [sy] + +let make_pp_rule level (typs,symbols) fmt = + match fmt with + | None -> + let hunks = make_hunks typs symbols level in + if List.exists (function _,(UnpCut (PpBrk _) | UnpListMetaVar _) -> true | _ -> false) hunks then + [UnpBox (PpHOVB 0,hunks)] + else + (* Optimization to work around what seems an ocaml Format bug (see Mantis #7804/#7807) *) + List.map snd hunks (* drop locations which are dummy *) + | Some fmt -> + hunks_of_format (level, List.split typs) (symbols, parse_format fmt) + +(* let make_syntax_rules i_typs (ntn,prec,need_squash) sy_data fmt extra onlyprint compat = *) +let make_syntax_rules (sd : SynData.syn_data) = let open SynData in + let ntn_for_grammar, prec_for_grammar, need_squash = sd.not_data in + let pa_rule = make_pa_rule prec_for_grammar sd.pa_syntax_data ntn_for_grammar need_squash in + let pp_rule = make_pp_rule (pi1 sd.level) sd.pp_syntax_data sd.format in { + synext_level = sd.level; + synext_notation = fst sd.info; + synext_notgram = { notgram_onlyprinting = sd.only_printing; notgram_rules = pa_rule }; + synext_unparsing = pp_rule; + synext_extra = sd.extra; + synext_compat = sd.compat; + } + +(**********************************************************************) +(* Main functions about notations *) + +let to_map l = + let fold accu (x, v) = Id.Map.add x v accu in + List.fold_left fold Id.Map.empty l + +let add_notation_in_scope local df env c mods scope = + let open SynData in + let sd = compute_syntax_data df mods in + (* Prepare the interpretation *) + (* Prepare the parsing and printing rules *) + let sy_rules = make_syntax_rules sd in + let i_vars = make_internalization_vars sd.recvars sd.mainvars sd.intern_typs in + let nenv = { + ninterp_var_type = to_map i_vars; + ninterp_rec_vars = to_map sd.recvars; + } in + let (acvars, ac, reversibility) = interp_notation_constr env nenv c in + let interp = make_interpretation_vars sd.recvars acvars (fst sd.pa_syntax_data) in + let map (x, _) = try Some (x, Id.Map.find x interp) with Not_found -> None in + let onlyparse = is_not_printable sd.only_parsing reversibility ac in + let notation = { + notobj_local = local; + notobj_scope = scope; + notobj_interp = (List.map_filter map i_vars, ac); + (** Order is important here! *) + notobj_onlyparse = onlyparse; + notobj_onlyprint = sd.only_printing; + notobj_compat = sd.compat; + notobj_notation = sd.info; + } in + (* Ready to change the global state *) + Flags.if_verbose (List.iter (fun (f,x) -> f x)) sd.msgs; + Lib.add_anonymous_leaf (inSyntaxExtension (local, sy_rules)); + Lib.add_anonymous_leaf (inNotation notation); + sd.info + +let add_notation_interpretation_core local df env ?(impls=empty_internalization_env) c scope onlyparse onlyprint compat = + let (recvars,mainvars,symbs) = analyze_notation_tokens ~onlyprint df in + (* Recover types of variables and pa/pp rules; redeclare them if needed *) + let i_typs, onlyprint = if not (is_numeral symbs) then begin + let sy = recover_notation_syntax (make_notation_key symbs) in + let () = Lib.add_anonymous_leaf (inSyntaxExtension (local,sy)) in + (** If the only printing flag has been explicitly requested, put it back *) + let onlyprint = onlyprint || sy.synext_notgram.notgram_onlyprinting in + pi3 sy.synext_level, onlyprint + end else [], false in + (* Declare interpretation *) + let path = (Lib.library_dp(), Lib.current_dirpath true) in + let df' = (make_notation_key symbs, (path,df)) in + let i_vars = make_internalization_vars recvars mainvars (List.map internalization_type_of_entry_type i_typs) in + let nenv = { + ninterp_var_type = to_map i_vars; + ninterp_rec_vars = to_map recvars; + } in + let (acvars, ac, reversibility) = interp_notation_constr env ~impls nenv c in + let interp = make_interpretation_vars recvars acvars (List.combine mainvars i_typs) in + let map (x, _) = try Some (x, Id.Map.find x interp) with Not_found -> None in + let onlyparse = is_not_printable onlyparse reversibility ac in + let notation = { + notobj_local = local; + notobj_scope = scope; + notobj_interp = (List.map_filter map i_vars, ac); + (** Order is important here! *) + notobj_onlyparse = onlyparse; + notobj_onlyprint = onlyprint; + notobj_compat = compat; + notobj_notation = df'; + } in + Lib.add_anonymous_leaf (inNotation notation); + df' + +(* Notations without interpretation (Reserved Notation) *) + +let add_syntax_extension local ({CAst.loc;v=df},mods) = let open SynData in + let psd = compute_pure_syntax_data df mods in + let sy_rules = make_syntax_rules {psd with compat = None} in + Flags.if_verbose (List.iter (fun (f,x) -> f x)) psd.msgs; + Lib.add_anonymous_leaf (inSyntaxExtension(local,sy_rules)) + +(* Notations with only interpretation *) + +let add_notation_interpretation env ({CAst.loc;v=df},c,sc) = + let df' = add_notation_interpretation_core false df env c sc false false None in + Dumpglob.dump_notation (loc,df') sc true + +let set_notation_for_interpretation env impls ({CAst.v=df},c,sc) = + (try ignore + (Flags.silently (fun () -> add_notation_interpretation_core false df env ~impls c sc false false None) ()); + with NoSyntaxRule -> + user_err Pp.(str "Parsing rule for this notation has to be previously declared.")); + Option.iter (fun sc -> Notation.open_close_scope (false,true,sc)) sc + +(* Main entry point *) + +let add_notation local env c ({CAst.loc;v=df},modifiers) sc = + let df' = + if no_syntax_modifiers modifiers then + (* No syntax data: try to rely on a previously declared rule *) + let onlyparse = is_only_parsing modifiers in + let onlyprint = is_only_printing modifiers in + let compat = get_compat_version modifiers in + try add_notation_interpretation_core local df env c sc onlyparse onlyprint compat + with NoSyntaxRule -> + (* Try to determine a default syntax rule *) + add_notation_in_scope local df env c modifiers sc + else + (* Declare both syntax and interpretation *) + add_notation_in_scope local df env c modifiers sc + in + Dumpglob.dump_notation (loc,df') sc true + +let add_notation_extra_printing_rule df k v = + let notk = + let _,_, symbs = analyze_notation_tokens ~onlyprint:true df in + make_notation_key symbs in + Notation.add_notation_extra_printing_rule notk k v + +(* Infix notations *) + +let inject_var x = CAst.make @@ CRef (CAst.make @@ Ident (Id.of_string x),None) + +let add_infix local env ({CAst.loc;v=inf},modifiers) pr sc = + check_infix_modifiers modifiers; + (* check the precedence *) + let metas = [inject_var "x"; inject_var "y"] in + let c = mkAppC (pr,metas) in + let df = CAst.make ?loc @@ "x "^(quote_notation_token inf)^" y" in + add_notation local env c (df,modifiers) sc + +(**********************************************************************) +(* Delimiters and classes bound to scopes *) + +type scope_command = + | ScopeDelim of string + | ScopeClasses of scope_class list + | ScopeRemove + +let load_scope_command _ (_,(scope,dlm)) = + Notation.declare_scope scope + +let open_scope_command i (_,(scope,o)) = + if Int.equal i 1 then + match o with + | ScopeDelim dlm -> Notation.declare_delimiters scope dlm + | ScopeClasses cl -> List.iter (Notation.declare_scope_class scope) cl + | ScopeRemove -> Notation.remove_delimiters scope + +let cache_scope_command o = + load_scope_command 1 o; + open_scope_command 1 o + +let subst_scope_command (subst,(scope,o as x)) = match o with + | ScopeClasses cl -> + let cl' = List.map_filter (subst_scope_class subst) cl in + let cl' = + if List.for_all2eq (==) cl cl' then cl + else cl' in + scope, ScopeClasses cl' + | _ -> x + +let inScopeCommand : scope_name * scope_command -> obj = + declare_object {(default_object "DELIMITERS") with + cache_function = cache_scope_command; + open_function = open_scope_command; + load_function = load_scope_command; + subst_function = subst_scope_command; + classify_function = (fun obj -> Substitute obj)} + +let add_delimiters scope key = + Lib.add_anonymous_leaf (inScopeCommand(scope,ScopeDelim key)) + +let remove_delimiters scope = + Lib.add_anonymous_leaf (inScopeCommand(scope,ScopeRemove)) + +let add_class_scope scope cl = + Lib.add_anonymous_leaf (inScopeCommand(scope,ScopeClasses cl)) + +(* Check if abbreviation to a name and avoid early insertion of + maximal implicit arguments *) +let try_interp_name_alias = function + | [], { CAst.v = CRef (ref,_) } -> intern_reference ref + | _ -> raise Not_found + +let add_syntactic_definition env ident (vars,c) local onlyparse = + let vars,reversibility,pat = + try [], APrioriReversible, NRef (try_interp_name_alias (vars,c)) + with Not_found -> + let fold accu id = Id.Map.add id NtnInternTypeAny accu in + let i_vars = List.fold_left fold Id.Map.empty vars in + let nenv = { + ninterp_var_type = i_vars; + ninterp_rec_vars = Id.Map.empty; + } in + let nvars, pat, reversibility = interp_notation_constr env nenv c in + let map id = let (_,sc) = Id.Map.find id nvars in (id, sc) in + List.map map vars, reversibility, pat + in + let onlyparse = match onlyparse with + | None when (is_not_printable false reversibility pat) -> Some Flags.Current + | p -> p + in + Syntax_def.declare_syntactic_definition local ident onlyparse (vars,pat) diff --git a/vernac/metasyntax.mli b/vernac/metasyntax.mli new file mode 100644 index 00000000..a6c12e08 --- /dev/null +++ b/vernac/metasyntax.mli @@ -0,0 +1,63 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* unit + +(** Adding a (constr) notation in the environment*) + +val add_infix : locality_flag -> env -> (lstring * syntax_modifier list) -> + constr_expr -> scope_name option -> unit + +val add_notation : locality_flag -> env -> constr_expr -> + (lstring * syntax_modifier list) -> scope_name option -> unit + +val add_notation_extra_printing_rule : string -> string -> string -> unit + +(** Declaring delimiter keys and default scopes *) + +val add_delimiters : scope_name -> string -> unit +val remove_delimiters : scope_name -> unit +val add_class_scope : scope_name -> scope_class list -> unit + +(** Add only the interpretation of a notation that already has pa/pp rules *) + +val add_notation_interpretation : + env -> (lstring * constr_expr * scope_name option) -> unit + +(** Add a notation interpretation for supporting the "where" clause *) + +val set_notation_for_interpretation : env -> Constrintern.internalization_env -> + (lstring * constr_expr * scope_name option) -> unit + +(** Add only the parsing/printing rule of a notation *) + +val add_syntax_extension : + locality_flag -> (lstring * syntax_modifier list) -> unit + +(** Add a syntactic definition (as in "Notation f := ...") *) + +val add_syntactic_definition : env -> Id.t -> Id.t list * constr_expr -> + bool -> Flags.compat_version option -> unit + +(** Print the Camlp5 state of a grammar *) + +val pr_grammar : string -> Pp.t + +val check_infix_modifiers : syntax_modifier list -> unit + +val with_syntax_protection : ('a -> 'b) -> 'a -> 'b diff --git a/vernac/mltop.ml b/vernac/mltop.ml new file mode 100644 index 00000000..343b0925 --- /dev/null +++ b/vernac/mltop.ml @@ -0,0 +1,473 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* unit; + use_file : string -> unit; + add_dir : string -> unit; + ml_loop : unit -> unit } + +(* Determines the behaviour of Coq with respect to ML files (compiled + or not) *) +type kind_load = + | WithTop of toplevel + | WithoutTop + +(* Must be always initialized *) +let load = ref WithoutTop + +(* Are we in a native version of Coq? *) +let is_native = Dynlink.is_native + +(* Sets and initializes a toplevel (if any) *) +let set_top toplevel = load := + WithTop toplevel; + Nativelib.load_obj := toplevel.load_obj + +(* Removes the toplevel (if any) *) +let remove () = + load := WithoutTop; + Nativelib.load_obj := (fun x -> () : string -> unit) + +(* Tests if an Ocaml toplevel runs under Coq *) +let is_ocaml_top () = + match !load with + | WithTop _ -> true + |_ -> false + +(* Tests if we can load ML files *) +let has_dynlink = Coq_config.has_natdynlink || not is_native + +(* Runs the toplevel loop of Ocaml *) +let ocaml_toploop () = + match !load with + | WithTop t -> Printexc.catch t.ml_loop () + | _ -> () + +(* Try to interpret load_obj's (internal) errors *) +let report_on_load_obj_error exc = + let x = Obj.repr exc in + (* Try an horrible (fragile) hack to report on Symtable dynlink errors *) + (* (we follow ocaml's Printexc.to_string decoding of exceptions) *) + if Obj.is_block x && String.equal (Obj.magic (Obj.field (Obj.field x 0) 0)) "Symtable.Error" + then + let err_block = Obj.field x 1 in + if Int.equal (Obj.tag err_block) 0 then + (* Symtable.Undefined_global of string *) + str "reference to undefined global " ++ + str (Obj.magic (Obj.field err_block 0)) + else str (Printexc.to_string exc) + else str (Printexc.to_string exc) + +(* Dynamic loading of .cmo/.cma *) + +let ml_load s = + match !load with + | WithTop t -> + (try t.load_obj s; s + with + | e when CErrors.noncritical e -> + let e = CErrors.push e in + match fst e with + | (UserError _ | Failure _ | Not_found as u) -> Exninfo.iraise (u, snd e) + | exc -> + let msg = report_on_load_obj_error exc in + user_err ~hdr:"Mltop.load_object" (str"Cannot link ml-object " ++ + str s ++ str" to Coq code (" ++ msg ++ str ").")) + | WithoutTop -> + try + Dynlink.loadfile s; s + with Dynlink.Error a -> + user_err ~hdr:"Mltop.load_object" + (strbrk "while loading " ++ str s ++ + strbrk ": " ++ str (Dynlink.error_message a)) + +let dir_ml_load s = + match !load with + | WithTop _ -> ml_load s + | WithoutTop -> + let warn = not !Flags.quiet in + let _,gname = find_file_in_path ~warn !coq_mlpath_copy s in + ml_load gname + +(* Dynamic interpretation of .ml *) +let dir_ml_use s = + match !load with + | WithTop t -> t.use_file s + | _ -> + let moreinfo = + if Dynlink.is_native then " Loading ML code works only in bytecode." + else "" + in + user_err ~hdr:"Mltop.dir_ml_use" (str "Could not load ML code." ++ str moreinfo) + +(* Adds a path to the ML paths *) +let add_ml_dir s = + match !load with + | WithTop t -> t.add_dir s; keep_copy_mlpath s + | WithoutTop when has_dynlink -> keep_copy_mlpath s + | _ -> () + +(* For Rec Add ML Path (-R) *) +let add_rec_ml_dir unix_path = + List.iter (fun (lp,_) -> add_ml_dir lp) (all_subdirs ~unix_path) + +(* Adding files to Coq and ML loadpath *) + +let warn_cannot_use_directory = + CWarnings.create ~name:"cannot-use-directory" ~category:"filesystem" + (fun d -> + str "Directory " ++ str d ++ + strbrk " cannot be used as a Coq identifier (skipped)") + +let convert_string d = + try Names.Id.of_string d + with UserError _ -> + let d = Unicode.escaped_if_non_utf8 d in + warn_cannot_use_directory d; + raise Exit + +let warn_cannot_open_path = + CWarnings.create ~name:"cannot-open-path" ~category:"filesystem" + (fun unix_path -> str "Cannot open " ++ str unix_path) + +type add_ml = AddNoML | AddTopML | AddRecML + +type vo_path_spec = { + unix_path : string; + coq_path : Names.DirPath.t; + implicit : bool; + has_ml : add_ml; +} + +type coq_path_spec = + | VoPath of vo_path_spec + | MlPath of string + +type coq_path = { + path_spec: coq_path_spec; + recursive: bool; +} + +let add_vo_path ~recursive lp = + let unix_path = lp.unix_path in + let implicit = lp.implicit in + if exists_dir unix_path then + let dirs = if recursive then all_subdirs ~unix_path else [] in + let prefix = Names.DirPath.repr lp.coq_path in + let convert_dirs (lp, cp) = + try + let path = List.rev_map convert_string cp @ prefix in + Some (lp, Names.DirPath.make path) + with Exit -> None + in + let dirs = List.map_filter convert_dirs dirs in + let () = match lp.has_ml with + | AddNoML -> () + | AddTopML -> add_ml_dir unix_path + | AddRecML -> List.iter (fun (lp,_) -> add_ml_dir lp) dirs in + let add (path, dir) = + Loadpath.add_load_path path ~implicit dir in + let () = List.iter add dirs in + Loadpath.add_load_path unix_path ~implicit lp.coq_path + else + warn_cannot_open_path unix_path + +let add_coq_path { recursive; path_spec } = match path_spec with + | VoPath lp -> + add_vo_path ~recursive lp + | MlPath dir -> + if recursive then add_rec_ml_dir dir else add_ml_dir dir + +(* convertit un nom quelconque en nom de fichier ou de module *) +let mod_of_name name = + if Filename.check_suffix name ".cmo" then + Filename.chop_suffix name ".cmo" + else + name + +let get_ml_object_suffix name = + if Filename.check_suffix name ".cmo" then + Some ".cmo" + else if Filename.check_suffix name ".cma" then + Some ".cma" + else if Filename.check_suffix name ".cmxs" then + Some ".cmxs" + else + None + +let file_of_name name = + let suffix = get_ml_object_suffix name in + let fail s = + user_err ~hdr:"Mltop.load_object" + (str"File not found on loadpath : " ++ str s ++ str"\n" ++ + str"Loadpath: " ++ str(String.concat ":" !coq_mlpath_copy)) in + if not (Filename.is_relative name) then + if Sys.file_exists name then name else fail name + else if is_native then + let name = match suffix with + | Some ((".cmo"|".cma") as suffix) -> + (Filename.chop_suffix name suffix) ^ ".cmxs" + | Some ".cmxs" -> name + | _ -> name ^ ".cmxs" + in + if is_in_path !coq_mlpath_copy name then name else fail name + else + let (full, base) = match suffix with + | Some ".cmo" | Some ".cma" -> true, name + | Some ".cmxs" -> false, Filename.chop_suffix name ".cmxs" + | _ -> false, name + in + if full then + if is_in_path !coq_mlpath_copy base then base else fail base + else + let name = base ^ ".cma" in + if is_in_path !coq_mlpath_copy name then name else + let name = base ^ ".cmo" in + if is_in_path !coq_mlpath_copy name then name else + fail (base ^ ".cm[ao]") + +(** Is the ML code of the standard library placed into loadable plugins + or statically compiled into coqtop ? For the moment this choice is + made according to the presence of native dynlink : even if bytecode + coqtop could always load plugins, we prefer to have uniformity between + bytecode and native versions. *) + +(* [known_loaded_module] contains the names of the loaded ML modules + * (linked or loaded with load_object). It is used not to load a + * module twice. It is NOT the list of ML modules Coq knows. *) + +let known_loaded_modules = ref String.Map.empty + +let add_known_module mname path = + if not (String.Map.mem mname !known_loaded_modules) || + String.Map.find mname !known_loaded_modules = None then + known_loaded_modules := String.Map.add mname path !known_loaded_modules + +let module_is_known mname = + String.Map.mem mname !known_loaded_modules + +let known_module_path mname = + String.Map.find mname !known_loaded_modules + +(** A plugin is just an ML module with an initialization function. *) + +let known_loaded_plugins = ref String.Map.empty + +let add_known_plugin init name = + add_known_module name None; + known_loaded_plugins := String.Map.add name init !known_loaded_plugins + +let init_known_plugins () = + String.Map.iter (fun _ f -> f()) !known_loaded_plugins + +(** Registering functions to be used at caching time, that is when the Declare + ML module command is issued. *) + +let cache_objs = ref String.Map.empty + +let declare_cache_obj f name = + let objs = try String.Map.find name !cache_objs with Not_found -> [] in + let objs = f :: objs in + cache_objs := String.Map.add name objs !cache_objs + +let perform_cache_obj name = + let objs = try String.Map.find name !cache_objs with Not_found -> [] in + let objs = List.rev objs in + List.iter (fun f -> f ()) objs + +(** ml object = ml module or plugin *) + +let init_ml_object mname = + try String.Map.find mname !known_loaded_plugins () + with Not_found -> () + +let load_ml_object mname ?path fname= + let path = match path with + | None -> dir_ml_load fname + | Some p -> ml_load p in + add_known_module mname (Some path); + init_ml_object mname; + path + +let dir_ml_load m = ignore(dir_ml_load m) +let add_known_module m = add_known_module m None +let load_ml_object_raw fname = dir_ml_load (file_of_name fname) +let load_ml_objects_raw_rex rex = + List.iter (fun (_,fp) -> + let name = file_of_name (Filename.basename fp) in + try dir_ml_load name + with e -> prerr_endline (Printexc.to_string e)) + (System.where_in_path_rex !coq_mlpath_copy rex) + +(* Summary of declared ML Modules *) + +(* List and not String.Set because order is important: most recent first. *) + +let loaded_modules = ref [] +let get_loaded_modules () = List.rev !loaded_modules +let add_loaded_module md path = + if not (List.mem_assoc md !loaded_modules) then + loaded_modules := (md,path) :: !loaded_modules +let reset_loaded_modules () = loaded_modules := [] + +let if_verbose_load verb f name ?path fname = + if not verb then f name ?path fname + else + let info = str "[Loading ML file " ++ str fname ++ str " ..." in + try + let path = f name ?path fname in + Feedback.msg_info (info ++ str " done]"); + path + with reraise -> + Feedback.msg_info (info ++ str " failed]"); + raise reraise + +(** Load a module for the first time (i.e. dynlink it) + or simulate its reload (i.e. doing nothing except maybe + an initialization function). *) + +let trigger_ml_object verb cache reinit ?path name = + if module_is_known name then begin + if reinit then init_ml_object name; + add_loaded_module name (known_module_path name); + if cache then perform_cache_obj name + end else if not has_dynlink then + user_err ~hdr:"Mltop.trigger_ml_object" + (str "Dynamic link not supported (module " ++ str name ++ str ")") + else begin + let file = file_of_name (Option.default name path) in + let path = + if_verbose_load (verb && not !Flags.quiet) load_ml_object name ?path file in + add_loaded_module name (Some path); + if cache then perform_cache_obj name + end + +let load_ml_object n m = ignore(load_ml_object n m) + +let unfreeze_ml_modules x = + reset_loaded_modules (); + List.iter + (fun (name,path) -> trigger_ml_object false false false ?path name) x + +let _ = + Summary.declare_ml_modules_summary + { Summary.freeze_function = (fun _ -> get_loaded_modules ()); + Summary.unfreeze_function = unfreeze_ml_modules; + Summary.init_function = reset_loaded_modules } + +(* Liboject entries of declared ML Modules *) + +type ml_module_object = { + mlocal : Vernacexpr.locality_flag; + mnames : string list +} + +let cache_ml_objects (_,{mnames=mnames}) = + let iter obj = trigger_ml_object true true true obj in + List.iter iter mnames + +let load_ml_objects _ (_,{mnames=mnames}) = + let iter obj = trigger_ml_object true false true obj in + List.iter iter mnames + +let classify_ml_objects ({mlocal=mlocal} as o) = + if mlocal then Dispose else Substitute o + +let inMLModule : ml_module_object -> obj = + declare_object + {(default_object "ML-MODULE") with + cache_function = cache_ml_objects; + load_function = load_ml_objects; + subst_function = (fun (_,o) -> o); + classify_function = classify_ml_objects } + +let declare_ml_modules local l = + let l = List.map mod_of_name l in + Lib.add_anonymous_leaf ~cache_first:false (inMLModule {mlocal=local; mnames=l}) + +let print_ml_path () = + let l = !coq_mlpath_copy in + str"ML Load Path:" ++ fnl () ++ str" " ++ + hv 0 (prlist_with_sep fnl str l) + +(* Printing of loaded ML modules *) + +let print_ml_modules () = + let l = get_loaded_modules () in + str"Loaded ML Modules: " ++ pr_vertical_list str (List.map fst l) + +let print_gc () = + let stat = Gc.stat () in + let msg = + str "minor words: " ++ real stat.Gc.minor_words ++ fnl () ++ + str "promoted words: " ++ real stat.Gc.promoted_words ++ fnl () ++ + str "major words: " ++ real stat.Gc.major_words ++ fnl () ++ + str "minor_collections: " ++ int stat.Gc.minor_collections ++ fnl () ++ + str "major_collections: " ++ int stat.Gc.major_collections ++ fnl () ++ + str "heap_words: " ++ int stat.Gc.heap_words ++ fnl () ++ + str "heap_chunks: " ++ int stat.Gc.heap_chunks ++ fnl () ++ + str "live_words: " ++ int stat.Gc.live_words ++ fnl () ++ + str "live_blocks: " ++ int stat.Gc.live_blocks ++ fnl () ++ + str "free_words: " ++ int stat.Gc.free_words ++ fnl () ++ + str "free_blocks: " ++ int stat.Gc.free_blocks ++ fnl () ++ + str "largest_free: " ++ int stat.Gc.largest_free ++ fnl () ++ + str "fragments: " ++ int stat.Gc.fragments ++ fnl () ++ + str "compactions: " ++ int stat.Gc.compactions ++ fnl () ++ + str "top_heap_words: " ++ int stat.Gc.top_heap_words ++ fnl () ++ + str "stack_size: " ++ int stat.Gc.stack_size + in + hv 0 msg diff --git a/vernac/mltop.mli b/vernac/mltop.mli new file mode 100644 index 00000000..da195f4f --- /dev/null +++ b/vernac/mltop.mli @@ -0,0 +1,102 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* unit; + use_file : string -> unit; + add_dir : string -> unit; + ml_loop : unit -> unit } + +(** Sets and initializes a toplevel (if any) *) +val set_top : toplevel -> unit + +(** Are we in a native version of Coq? *) +val is_native : bool + +(** Removes the toplevel (if any) *) +val remove : unit -> unit + +(** Tests if an Ocaml toplevel runs under Coq *) +val is_ocaml_top : unit -> bool + +(** Starts the Ocaml toplevel loop *) +val ocaml_toploop : unit -> unit + +(** {5 ML Dynlink} *) + +(** Tests if we can load ML files *) +val has_dynlink : bool + +(** Dynamic loading of .cmo *) +val dir_ml_load : string -> unit + +(** Dynamic interpretation of .ml *) +val dir_ml_use : string -> unit + +(** Adds a path to the Coq and ML paths *) +type add_ml = AddNoML | AddTopML | AddRecML + +type vo_path_spec = { + unix_path : string; (* Filesystem path contaning vo/ml files *) + coq_path : Names.DirPath.t; (* Coq prefix for the path *) + implicit : bool; (* [implicit = true] avoids having to qualify with [coq_path] *) + has_ml : add_ml; (* If [has_ml] is true, the directory will also be search for plugins *) +} + +type coq_path_spec = + | VoPath of vo_path_spec + | MlPath of string + +type coq_path = { + path_spec: coq_path_spec; + recursive: bool; +} + +val add_coq_path : coq_path -> unit + +(** List of modules linked to the toplevel *) +val add_known_module : string -> unit +val module_is_known : string -> bool +val load_ml_object : string -> string -> unit +val load_ml_object_raw : string -> unit +val load_ml_objects_raw_rex : Str.regexp -> unit + +(** {5 Initialization functions} *) + +(** Declare a plugin and its initialization function. + A plugin is just an ML module with an initialization function. + Adding a known plugin implies adding it as a known ML module. + The initialization function is granted to be called after Coq is fully + bootstrapped, even if the plugin is statically linked with the toplevel *) +val add_known_plugin : (unit -> unit) -> string -> unit + +(** Calls all initialization functions in a non-specified order *) +val init_known_plugins : unit -> unit + +(** Register a callback that will be called when the module is declared with + the Declare ML Module command. This is useful to define Coq objects at that + time only. Several functions can be defined for one module; they will be + called in the order of declaration, and after the ML module has been + properly initialized. *) +val declare_cache_obj : (unit -> unit) -> string -> unit + +(** {5 Declaring modules} *) + +val declare_ml_modules : Vernacexpr.locality_flag -> string list -> unit + +(** {5 Utilities} *) + +val print_ml_path : unit -> Pp.t +val print_ml_modules : unit -> Pp.t +val print_gc : unit -> Pp.t diff --git a/vernac/obligations.ml b/vernac/obligations.ml new file mode 100644 index 00000000..765d6851 --- /dev/null +++ b/vernac/obligations.ml @@ -0,0 +1,1196 @@ +open Printf +open Libobject +open Entries +open Decl_kinds +open Declare + +(** + - Get types of existentials ; + - Flatten dependency tree (prefix order) ; + - Replace existentials by de Bruijn indices in term, applied to the right arguments ; + - Apply term prefixed by quantification on "existentials". +*) + +open Term +open Constr +open Vars +open Names +open Evd +open Pp +open CErrors +open Util + +module NamedDecl = Context.Named.Declaration + +let get_fix_exn, stm_get_fix_exn = Hook.make () + +let succfix (depth, fixrels) = + (succ depth, List.map succ fixrels) + +let check_evars env evm = + Evar.Map.iter + (fun key evi -> + let (loc,k) = evar_source key evm in + match k with + | Evar_kinds.QuestionMark _ + | Evar_kinds.ImplicitArg (_,_,false) -> () + | _ -> + Pretype_errors.error_unsolvable_implicit ?loc env evm key None) + (Evd.undefined_map evm) + +type oblinfo = + { ev_name: int * Id.t; + ev_hyps: Context.Named.t; + ev_status: bool * Evar_kinds.obligation_definition_status; + ev_chop: int option; + ev_src: Evar_kinds.t Loc.located; + ev_typ: types; + ev_tac: unit Proofview.tactic option; + ev_deps: Int.Set.t } + +(** Substitute evar references in t using de Bruijn indices, + where n binders were passed through. *) + +let subst_evar_constr evs n idf t = + let seen = ref Int.Set.empty in + let transparent = ref Id.Set.empty in + let evar_info id = List.assoc_f Evar.equal id evs in + let rec substrec (depth, fixrels) c = match Constr.kind c with + | Evar (k, args) -> + let { ev_name = (id, idstr) ; + ev_hyps = hyps ; ev_chop = chop } = + try evar_info k + with Not_found -> + anomaly ~label:"eterm" (Pp.str "existential variable " ++ int (Evar.repr k) ++ str " not found.") + in + seen := Int.Set.add id !seen; + (* Evar arguments are created in inverse order, + and we must not apply to defined ones (i.e. LetIn's) + *) + let args = + let n = match chop with None -> 0 | Some c -> c in + let (l, r) = List.chop n (List.rev (Array.to_list args)) in + List.rev r + in + let args = + let rec aux hyps args acc = + let open Context.Named.Declaration in + match hyps, args with + (LocalAssum _ :: tlh), (c :: tla) -> + aux tlh tla ((substrec (depth, fixrels) c) :: acc) + | (LocalDef _ :: tlh), (_ :: tla) -> + aux tlh tla acc + | [], [] -> acc + | _, _ -> acc (*failwith "subst_evars: invalid argument"*) + in aux hyps args [] + in + if List.exists + (fun x -> match Constr.kind x with + | Rel n -> Int.List.mem n fixrels + | _ -> false) args + then + transparent := Id.Set.add idstr !transparent; + mkApp (idf idstr, Array.of_list args) + | Fix _ -> + Constr.map_with_binders succfix substrec (depth, 1 :: fixrels) c + | _ -> Constr.map_with_binders succfix substrec (depth, fixrels) c + in + let t' = substrec (0, []) t in + t', !seen, !transparent + + +(** Substitute variable references in t using de Bruijn indices, + where n binders were passed through. *) +let subst_vars acc n t = + let var_index id = Util.List.index Id.equal id acc in + let rec substrec depth c = match Constr.kind c with + | Var v -> (try mkRel (depth + (var_index v)) with Not_found -> c) + | _ -> Constr.map_with_binders succ substrec depth c + in + substrec 0 t + +(** Rewrite type of an evar ([ H1 : t1, ... Hn : tn |- concl ]) + to a product : forall H1 : t1, ..., forall Hn : tn, concl. + Changes evars and hypothesis references to variable references. +*) +let etype_of_evar evs hyps concl = + let open Context.Named.Declaration in + let rec aux acc n = function + decl :: tl -> + let t', s, trans = subst_evar_constr evs n mkVar (NamedDecl.get_type decl) in + let t'' = subst_vars acc 0 t' in + let rest, s', trans' = aux (NamedDecl.get_id decl :: acc) (succ n) tl in + let s' = Int.Set.union s s' in + let trans' = Id.Set.union trans trans' in + (match decl with + | LocalDef (id,c,_) -> + let c', s'', trans'' = subst_evar_constr evs n mkVar c in + let c' = subst_vars acc 0 c' in + mkNamedProd_or_LetIn (LocalDef (id, c', t'')) rest, + Int.Set.union s'' s', + Id.Set.union trans'' trans' + | LocalAssum (id,_) -> + mkNamedProd_or_LetIn (LocalAssum (id, t'')) rest, s', trans') + | [] -> + let t', s, trans = subst_evar_constr evs n mkVar concl in + subst_vars acc 0 t', s, trans + in aux [] 0 (List.rev hyps) + +let trunc_named_context n ctx = + let len = List.length ctx in + List.firstn (len - n) ctx + +let rec chop_product n t = + let pop t = Vars.lift (-1) t in + if Int.equal n 0 then Some t + else + match Constr.kind t with + | Prod (_, _, b) -> if noccurn 1 b then chop_product (pred n) (pop b) else None + | _ -> None + +let evar_dependencies evm oev = + let one_step deps = + Evar.Set.fold (fun ev s -> + let evi = Evd.find evm ev in + let deps' = evars_of_filtered_evar_info evi in + if Evar.Set.mem oev deps' then + invalid_arg ("Ill-formed evar map: cycle detected for evar " ^ Pp.string_of_ppcmds @@ Evar.print oev) + else Evar.Set.union deps' s) + deps deps + in + let rec aux deps = + let deps' = one_step deps in + if Evar.Set.equal deps deps' then deps + else aux deps' + in aux (Evar.Set.singleton oev) + +let move_after (id, ev, deps as obl) l = + let rec aux restdeps = function + | (id', _, _) as obl' :: tl -> + let restdeps' = Evar.Set.remove id' restdeps in + if Evar.Set.is_empty restdeps' then + obl' :: obl :: tl + else obl' :: aux restdeps' tl + | [] -> [obl] + in aux (Evar.Set.remove id deps) l + +let sort_dependencies evl = + let rec aux l found list = + match l with + | (id, ev, deps) as obl :: tl -> + let found' = Evar.Set.union found (Evar.Set.singleton id) in + if Evar.Set.subset deps found' then + aux tl found' (obl :: list) + else aux (move_after obl tl) found list + | [] -> List.rev list + in aux evl Evar.Set.empty [] + +open Environ + +let eterm_obligations env name evm fs ?status t ty = + (* 'Serialize' the evars *) + let nc = Environ.named_context env in + let nc_len = Context.Named.length nc in + let evm = Evarutil.nf_evar_map_undefined evm in + let evl = Evarutil.non_instantiated evm in + let evl = Evar.Map.bindings evl in + let evl = List.map (fun (id, ev) -> (id, ev, evar_dependencies evm id)) evl in + let sevl = sort_dependencies evl in + let evl = List.map (fun (id, ev, _) -> id, ev) sevl in + let evn = + let i = ref (-1) in + List.rev_map (fun (id, ev) -> incr i; + (id, (!i, Id.of_string + (Id.to_string name ^ "_obligation_" ^ string_of_int (succ !i))), + ev)) evl + in + let evts = + (* Remove existential variables in types and build the corresponding products *) + List.fold_right + (fun (id, (n, nstr), ev) l -> + let hyps = Evd.evar_filtered_context ev in + let hyps = trunc_named_context nc_len hyps in + let evtyp, deps, transp = etype_of_evar l hyps ev.evar_concl in + let evtyp, hyps, chop = + match chop_product fs evtyp with + | Some t -> t, trunc_named_context fs hyps, fs + | None -> evtyp, hyps, 0 + in + let loc, k = evar_source id evm in + let status = match k with + | Evar_kinds.QuestionMark (o,_) -> o + | _ -> match status with + | Some o -> o + | None -> Evar_kinds.Define (not (Program.get_proofs_transparency ())) + in + let force_status, status, chop = match status with + | Evar_kinds.Define true as stat -> + if not (Int.equal chop fs) then true, Evar_kinds.Define false, None + else false, stat, Some chop + | s -> false, s, None + in + let info = { ev_name = (n, nstr); + ev_hyps = hyps; ev_status = force_status, status; ev_chop = chop; + ev_src = loc, k; ev_typ = evtyp ; ev_deps = deps; ev_tac = None } + in (id, info) :: l) + evn [] + in + let t', _, transparent = (* Substitute evar refs in the term by variables *) + subst_evar_constr evts 0 mkVar t + in + let ty, _, _ = subst_evar_constr evts 0 mkVar ty in + let evars = + List.map (fun (ev, info) -> + let { ev_name = (_, name); ev_status = force_status, status; + ev_src = src; ev_typ = typ; ev_deps = deps; ev_tac = tac } = info + in + let force_status, status = match status with + | Evar_kinds.Define true when Id.Set.mem name transparent -> + true, Evar_kinds.Define false + | _ -> force_status, status + in name, typ, src, (force_status, status), deps, tac) evts + in + let evnames = List.map (fun (ev, info) -> ev, snd info.ev_name) evts in + let evmap f c = pi1 (subst_evar_constr evts 0 f c) in + Array.of_list (List.rev evars), (evnames, evmap), t', ty + +let tactics_module = ["Program";"Tactics"] +let safe_init_constant md name () = + Coqlib.check_required_library ("Coq"::md); + Universes.constr_of_global (Coqlib.coq_reference "Obligations" md name) +let hide_obligation = safe_init_constant tactics_module "obligation" + +let pperror cmd = CErrors.user_err ~hdr:"Program" cmd +let error s = pperror (str s) + +let reduce c = + EConstr.Unsafe.to_constr (Reductionops.clos_norm_flags CClosure.betaiota (Global.env ()) Evd.empty (EConstr.of_constr c)) + +exception NoObligations of Id.t option + +let explain_no_obligations = function + Some ident -> str "No obligations for program " ++ Id.print ident + | None -> str "No obligations remaining" + +type obligation_info = + (Names.Id.t * types * Evar_kinds.t Loc.located * + (bool * Evar_kinds.obligation_definition_status) + * Int.Set.t * unit Proofview.tactic option) array + +type 'a obligation_body = + | DefinedObl of 'a + | TermObl of constr + +type obligation = + { obl_name : Id.t; + obl_type : types; + obl_location : Evar_kinds.t Loc.located; + obl_body : pconstant obligation_body option; + obl_status : bool * Evar_kinds.obligation_definition_status; + obl_deps : Int.Set.t; + obl_tac : unit Proofview.tactic option; + } + +type obligations = (obligation array * int) + +type fixpoint_kind = + | IsFixpoint of (Misctypes.lident option * Constrexpr.recursion_order_expr) list + | IsCoFixpoint + +type notations = (Misctypes.lstring * Constrexpr.constr_expr * Notation_term.scope_name option) list + +type program_info_aux = { + prg_name: Id.t; + prg_body: constr; + prg_type: constr; + prg_ctx: UState.t; + prg_univdecl: Univdecls.universe_decl; + prg_obligations: obligations; + prg_deps : Id.t list; + prg_fixkind : fixpoint_kind option ; + prg_implicits : (Constrexpr.explicitation * (bool * bool * bool)) list; + prg_notations : notations ; + prg_kind : definition_kind; + prg_reduce : constr -> constr; + prg_hook : (UState.t -> unit) Lemmas.declaration_hook; + prg_opaque : bool; + prg_sign: named_context_val; +} + +type program_info = program_info_aux CEphemeron.key + +let get_info x = + try CEphemeron.get x + with CEphemeron.InvalidKey -> + CErrors.anomaly Pp.(str "Program obligation can't be accessed by a worker.") + +let assumption_message = Declare.assumption_message + +let default_tactic = ref (Proofview.tclUNIT ()) + +(* true = hide obligations *) +let hide_obligations = ref false + +let set_hide_obligations = (:=) hide_obligations +let get_hide_obligations () = !hide_obligations + +open Goptions +let _ = + declare_bool_option + { optdepr = false; + optname = "Hidding of Program obligations"; + optkey = ["Hide";"Obligations"]; + optread = get_hide_obligations; + optwrite = set_hide_obligations; } + +let shrink_obligations = ref true + +let set_shrink_obligations = (:=) shrink_obligations +let get_shrink_obligations () = !shrink_obligations + +let _ = + declare_bool_option + { optdepr = true; (* remove in 8.8 *) + optname = "Shrinking of Program obligations"; + optkey = ["Shrink";"Obligations"]; + optread = get_shrink_obligations; + optwrite = set_shrink_obligations; } + +let evar_of_obligation o = make_evar (Global.named_context_val ()) o.obl_type + +let get_obligation_body expand obl = + match obl.obl_body with + | None -> None + | Some c -> + if expand && snd obl.obl_status == Evar_kinds.Expand then + match c with + | DefinedObl pc -> Some (constant_value_in (Global.env ()) pc) + | TermObl c -> Some c + else (match c with + | DefinedObl pc -> Some (mkConstU pc) + | TermObl c -> Some c) + +let obl_substitution expand obls deps = + Int.Set.fold + (fun x acc -> + let xobl = obls.(x) in + match get_obligation_body expand xobl with + | None -> acc + | Some oblb -> (xobl.obl_name, (xobl.obl_type, oblb)) :: acc) + deps [] + +let subst_deps expand obls deps t = + let osubst = obl_substitution expand obls deps in + (Vars.replace_vars (List.map (fun (n, (_, b)) -> n, b) osubst) t) + +let rec prod_app t n = + match Constr.kind (EConstr.Unsafe.to_constr (Termops.strip_outer_cast Evd.empty (EConstr.of_constr t))) (** FIXME *) with + | Prod (_,_,b) -> subst1 n b + | LetIn (_, b, t, b') -> prod_app (subst1 b b') n + | _ -> + user_err ~hdr:"prod_app" + (str"Needed a product, but didn't find one" ++ fnl ()) + + +(* prod_appvect T [| a1 ; ... ; an |] -> (T a1 ... an) *) +let prod_applist t nL = List.fold_left prod_app t nL + +let replace_appvars subst = + let rec aux c = + let f, l = decompose_app c in + if isVar f then + try + let c' = List.map (Constr.map aux) l in + let (t, b) = Id.List.assoc (destVar f) subst in + mkApp (delayed_force hide_obligation, + [| prod_applist t c'; applistc b c' |]) + with Not_found -> Constr.map aux c + else Constr.map aux c + in Constr.map aux + +let subst_prog expand obls ints prg = + let subst = obl_substitution expand obls ints in + if get_hide_obligations () then + (replace_appvars subst prg.prg_body, + replace_appvars subst ((* Termops.refresh_universes *) prg.prg_type)) + else + let subst' = List.map (fun (n, (_, b)) -> n, b) subst in + (Vars.replace_vars subst' prg.prg_body, + Vars.replace_vars subst' ((* Termops.refresh_universes *) prg.prg_type)) + +let subst_deps_obl obls obl = + let t' = subst_deps true obls obl.obl_deps obl.obl_type in + { obl with obl_type = t' } + +module ProgMap = Id.Map + +let map_replace k v m = ProgMap.add k (CEphemeron.create v) (ProgMap.remove k m) + +let map_keys m = ProgMap.fold (fun k _ l -> k :: l) m [] + +let from_prg, program_tcc_summary_tag = + Summary.ref_tag ProgMap.empty ~name:"program-tcc-table" + +let close sec = + if not (ProgMap.is_empty !from_prg) then + let keys = map_keys !from_prg in + user_err ~hdr:"Program" + (str "Unsolved obligations when closing " ++ str sec ++ str":" ++ spc () ++ + prlist_with_sep spc (fun x -> Id.print x) keys ++ + (str (if Int.equal (List.length keys) 1 then " has " else " have ") ++ + str "unsolved obligations")) + +let input : program_info ProgMap.t -> obj = + declare_object + { (default_object "Program state") with + cache_function = (fun (na, pi) -> from_prg := pi); + load_function = (fun _ (_, pi) -> from_prg := pi); + discharge_function = (fun _ -> close "section"; None); + classify_function = (fun _ -> close "module"; Dispose) } + +open Evd + +let progmap_remove prg = + Lib.add_anonymous_leaf (input (ProgMap.remove prg.prg_name !from_prg)) + +let progmap_add n prg = + Lib.add_anonymous_leaf (input (ProgMap.add n prg !from_prg)) + +let progmap_replace prg' = + Lib.add_anonymous_leaf (input (map_replace prg'.prg_name prg' !from_prg)) + +let rec intset_to = function + -1 -> Int.Set.empty + | n -> Int.Set.add n (intset_to (pred n)) + +let subst_body expand prg = + let obls, _ = prg.prg_obligations in + let ints = intset_to (pred (Array.length obls)) in + subst_prog expand obls ints prg + +let declare_definition prg = + let body, typ = subst_body true prg in + let nf = Universes.nf_evars_and_universes_opt_subst (fun x -> None) + (UState.subst prg.prg_ctx) in + let opaque = prg.prg_opaque in + let fix_exn = Hook.get get_fix_exn () in + let typ = nf typ in + let body = nf body in + let env = Global.env () in + let uvars = Univ.LSet.union + (Univops.universes_of_constr env typ) + (Univops.universes_of_constr env body) in + let uctx = UState.restrict prg.prg_ctx uvars in + let univs = UState.check_univ_decl ~poly:(pi2 prg.prg_kind) uctx prg.prg_univdecl in + let ce = definition_entry ~fix_exn ~opaque ~types:typ ~univs body in + let () = progmap_remove prg in + let ubinders = UState.universe_binders uctx in + DeclareDef.declare_definition prg.prg_name + prg.prg_kind ce ubinders prg.prg_implicits + (Lemmas.mk_hook (fun l r -> Lemmas.call_hook fix_exn prg.prg_hook l r uctx; r)) + +let rec lam_index n t acc = + match Constr.kind t with + | Lambda (Name n', _, _) when Id.equal n n' -> + acc + | Lambda (_, _, b) -> + lam_index n b (succ acc) + | _ -> raise Not_found + +let compute_possible_guardness_evidences (n,_) fixbody fixtype = + match n with + | Some { CAst.loc; v = n } -> [lam_index n fixbody 0] + | None -> + (* If recursive argument was not given by user, we try all args. + An earlier approach was to look only for inductive arguments, + but doing it properly involves delta-reduction, and it finally + doesn't seem to worth the effort (except for huge mutual + fixpoints ?) *) + let m = Termops.nb_prod Evd.empty (EConstr.of_constr fixtype) (** FIXME *) in + let ctx = fst (decompose_prod_n_assum m fixtype) in + List.map_i (fun i _ -> i) 0 ctx + +let mk_proof c = ((c, Univ.ContextSet.empty), Safe_typing.empty_private_constants) + +let declare_mutual_definition l = + let len = List.length l in + let first = List.hd l in + let fixdefs, fixtypes, fiximps = + List.split3 + (List.map (fun x -> + let subs, typ = (subst_body true x) in + let env = Global.env () in + let sigma = Evd.from_ctx x.prg_ctx in + let term = snd (Reductionops.splay_lam_n env sigma len (EConstr.of_constr subs)) in + let typ = snd (Reductionops.splay_prod_n env sigma len (EConstr.of_constr typ)) in + let term = EConstr.to_constr sigma term in + let typ = EConstr.to_constr sigma typ in + x.prg_reduce term, x.prg_reduce typ, x.prg_implicits) l) + in +(* let fixdefs = List.map reduce_fix fixdefs in *) + let fixkind = Option.get first.prg_fixkind in + let arrrec, recvec = Array.of_list fixtypes, Array.of_list fixdefs in + let fixdecls = (Array.of_list (List.map (fun x -> Name x.prg_name) l), arrrec, recvec) in + let (local,poly,kind) = first.prg_kind in + let fixnames = first.prg_deps in + let opaque = first.prg_opaque in + let kind = if fixkind != IsCoFixpoint then Fixpoint else CoFixpoint in + let indexes, fixdecls = + match fixkind with + | IsFixpoint wfl -> + let possible_indexes = + List.map3 compute_possible_guardness_evidences + wfl fixdefs fixtypes in + let indexes = + Pretyping.search_guard (Global.env()) + possible_indexes fixdecls in + Some indexes, + List.map_i (fun i _ -> + mk_proof (mkFix ((indexes,i),fixdecls))) 0 l + | IsCoFixpoint -> + None, + List.map_i (fun i _ -> + mk_proof (mkCoFix (i,fixdecls))) 0 l + in + (* Declare the recursive definitions *) + let univs = UState.const_univ_entry ~poly first.prg_ctx in + let fix_exn = Hook.get get_fix_exn () in + let kns = List.map4 (DeclareDef.declare_fix ~opaque (local, poly, kind) Universes.empty_binders univs) + fixnames fixdecls fixtypes fiximps in + (* Declare notations *) + List.iter (Metasyntax.add_notation_interpretation (Global.env())) first.prg_notations; + Declare.recursive_message (fixkind != IsCoFixpoint) indexes fixnames; + let gr = List.hd kns in + Lemmas.call_hook fix_exn first.prg_hook local gr first.prg_ctx; + List.iter progmap_remove l; gr + +let decompose_lam_prod c ty = + let open Context.Rel.Declaration in + let rec aux ctx c ty = + match Constr.kind c, Constr.kind ty with + | LetIn (x, b, t, c), LetIn (x', b', t', ty) + when Constr.equal b b' && Constr.equal t t' -> + let ctx' = Context.Rel.add (LocalDef (x,b',t')) ctx in + aux ctx' c ty + | _, LetIn (x', b', t', ty) -> + let ctx' = Context.Rel.add (LocalDef (x',b',t')) ctx in + aux ctx' (lift 1 c) ty + | LetIn (x, b, t, c), _ -> + let ctx' = Context.Rel.add (LocalDef (x,b,t)) ctx in + aux ctx' c (lift 1 ty) + | Lambda (x, b, t), Prod (x', b', t') + (* By invariant, must be convertible *) -> + let ctx' = Context.Rel.add (LocalAssum (x,b')) ctx in + aux ctx' t t' + | Cast (c, _, _), _ -> aux ctx c ty + | _, _ -> ctx, c, ty + in aux Context.Rel.empty c ty + +let shrink_body c ty = + let ctx, b, ty = + match ty with + | None -> + let ctx, b = decompose_lam_assum c in + ctx, b, None + | Some ty -> + let ctx, b, ty = decompose_lam_prod c ty in + ctx, b, Some ty + in + let b', ty', n, args = + List.fold_left (fun (b, ty, i, args) decl -> + if noccurn 1 b && Option.cata (noccurn 1) true ty then + subst1 mkProp b, Option.map (subst1 mkProp) ty, succ i, args + else + let open Context.Rel.Declaration in + let args = if is_local_assum decl then mkRel i :: args else args in + mkLambda_or_LetIn decl b, Option.map (mkProd_or_LetIn decl) ty, + succ i, args) + (b, ty, 1, []) ctx + in ctx, b', ty', Array.of_list args + +let unfold_entry cst = Hints.HintsUnfoldEntry [EvalConstRef cst] + +let add_hint local prg cst = + Hints.add_hints local [Id.to_string prg.prg_name] (unfold_entry cst) + +let it_mkLambda_or_LetIn_or_clean t ctx = + let open Context.Rel.Declaration in + let fold t decl = + if is_local_assum decl then mkLambda_or_LetIn decl t + else + if noccurn 1 t then subst1 mkProp t + else mkLambda_or_LetIn decl t + in + Context.Rel.fold_inside fold ctx ~init:t + +let declare_obligation prg obl body ty uctx = + let body = prg.prg_reduce body in + let ty = Option.map prg.prg_reduce ty in + match obl.obl_status with + | _, Evar_kinds.Expand -> false, { obl with obl_body = Some (TermObl body) } + | force, Evar_kinds.Define opaque -> + let opaque = not force && opaque in + let poly = pi2 prg.prg_kind in + let ctx, body, ty, args = + if get_shrink_obligations () && not poly then + shrink_body body ty else [], body, ty, [||] + in + let body = ((body,Univ.ContextSet.empty),Safe_typing.empty_private_constants) in + let ce = + { const_entry_body = Future.from_val ~fix_exn:(fun x -> x) body; + const_entry_secctx = None; + const_entry_type = ty; + const_entry_universes = uctx; + const_entry_opaque = opaque; + const_entry_inline_code = false; + const_entry_feedback = None; + } in + (** ppedrot: seems legit to have obligations as local *) + let constant = Declare.declare_constant obl.obl_name ~local:true + (DefinitionEntry ce,IsProof Property) + in + if not opaque then add_hint (Locality.make_section_locality None) prg constant; + definition_message obl.obl_name; + let body = match uctx with + | Polymorphic_const_entry uctx -> + Some (DefinedObl (constant, Univ.UContext.instance uctx)) + | Monomorphic_const_entry _ -> + Some (TermObl (it_mkLambda_or_LetIn_or_clean (mkApp (mkConst constant, args)) ctx)) + in + true, { obl with obl_body = body } + +let init_prog_info ?(opaque = false) sign n udecl b t ctx deps fixkind + notations obls impls kind reduce hook = + let obls', b = + match b with + | None -> + assert(Int.equal (Array.length obls) 0); + let n = Nameops.add_suffix n "_obligation" in + [| { obl_name = n; obl_body = None; + obl_location = Loc.tag Evar_kinds.InternalHole; obl_type = t; + obl_status = false, Evar_kinds.Expand; obl_deps = Int.Set.empty; + obl_tac = None } |], + mkVar n + | Some b -> + Array.mapi + (fun i (n, t, l, o, d, tac) -> + { obl_name = n ; obl_body = None; + obl_location = l; obl_type = t; obl_status = o; + obl_deps = d; obl_tac = tac }) + obls, b + in + let ctx = UState.make_flexible_nonalgebraic ctx in + { prg_name = n ; prg_body = b; prg_type = reduce t; + prg_ctx = ctx; prg_univdecl = udecl; + prg_obligations = (obls', Array.length obls'); + prg_deps = deps; prg_fixkind = fixkind ; prg_notations = notations ; + prg_implicits = impls; prg_kind = kind; prg_reduce = reduce; + prg_hook = hook; prg_opaque = opaque; + prg_sign = sign } + +let map_cardinal m = + let i = ref 0 in + ProgMap.iter (fun _ v -> + if snd (CEphemeron.get v).prg_obligations > 0 then incr i) m; + !i + +exception Found of program_info + +let map_first m = + try + ProgMap.iter (fun _ v -> + if snd (CEphemeron.get v).prg_obligations > 0 then + raise (Found v)) m; + assert(false) + with Found x -> x + +let get_prog name = + let prg_infos = !from_prg in + match name with + Some n -> + (try get_info (ProgMap.find n prg_infos) + with Not_found -> raise (NoObligations (Some n))) + | None -> + (let n = map_cardinal prg_infos in + match n with + 0 -> raise (NoObligations None) + | 1 -> get_info (map_first prg_infos) + | _ -> + let progs = Id.Set.elements (ProgMap.domain prg_infos) in + let prog = List.hd progs in + let progs = prlist_with_sep pr_comma Id.print progs in + user_err + (str "More than one program with unsolved obligations: " ++ progs + ++ str "; use the \"of\" clause to specify, as in \"Obligation 1 of " ++ Id.print prog ++ str "\"")) + +let get_any_prog () = + let prg_infos = !from_prg in + let n = map_cardinal prg_infos in + if n > 0 then get_info (map_first prg_infos) + else raise (NoObligations None) + +let get_prog_err n = + try get_prog n with NoObligations id -> pperror (explain_no_obligations id) + +let get_any_prog_err () = + try get_any_prog () with NoObligations id -> pperror (explain_no_obligations id) + +let obligations_solved prg = Int.equal (snd prg.prg_obligations) 0 + +let all_programs () = + ProgMap.fold (fun k p l -> p :: l) !from_prg [] + +type progress = + | Remain of int + | Dependent + | Defined of global_reference + +let obligations_message rem = + if rem > 0 then + if Int.equal rem 1 then + Flags.if_verbose Feedback.msg_info (int rem ++ str " obligation remaining") + else + Flags.if_verbose Feedback.msg_info (int rem ++ str " obligations remaining") + else + Flags.if_verbose Feedback.msg_info (str "No more obligations remaining") + +let update_obls prg obls rem = + let prg' = { prg with prg_obligations = (obls, rem) } in + progmap_replace prg'; + obligations_message rem; + if rem > 0 then Remain rem + else ( + match prg'.prg_deps with + | [] -> + let kn = declare_definition prg' in + progmap_remove prg'; + Defined kn + | l -> + let progs = List.map (fun x -> get_info (ProgMap.find x !from_prg)) prg'.prg_deps in + if List.for_all (fun x -> obligations_solved x) progs then + let kn = declare_mutual_definition progs in + Defined kn + else Dependent) + +let is_defined obls x = not (Option.is_empty obls.(x).obl_body) + +let deps_remaining obls deps = + Int.Set.fold + (fun x acc -> + if is_defined obls x then acc + else x :: acc) + deps [] + +let dependencies obls n = + let res = ref Int.Set.empty in + Array.iteri + (fun i obl -> + if not (Int.equal i n) && Int.Set.mem n obl.obl_deps then + res := Int.Set.add i !res) + obls; + !res + +let goal_kind poly = Decl_kinds.Local, poly, Decl_kinds.DefinitionBody Decl_kinds.Definition + +let goal_proof_kind poly = Decl_kinds.Local, poly, Decl_kinds.Proof Decl_kinds.Lemma + +let kind_of_obligation poly o = + match o with + | Evar_kinds.Define false | Evar_kinds.Expand -> goal_kind poly + | _ -> goal_proof_kind poly + +let not_transp_msg = + str "Obligation should be transparent but was declared opaque." ++ spc () ++ + str"Use 'Defined' instead." + +let err_not_transp () = pperror not_transp_msg + +let rec string_of_list sep f = function + [] -> "" + | x :: [] -> f x + | x :: ((y :: _) as tl) -> f x ^ sep ^ string_of_list sep f tl + +(* Solve an obligation using tactics, return the corresponding proof term *) + +let solve_by_tac name evi t poly ctx = + let id = name in + let concl = EConstr.of_constr evi.evar_concl in + (* spiwack: the status is dropped. *) + let (entry,_,ctx') = Pfedit.build_constant_by_tactic + id ~goal_kind:(goal_kind poly) ctx evi.evar_hyps concl (Tacticals.New.tclCOMPLETE t) in + let env = Global.env () in + let entry = Safe_typing.inline_private_constants_in_definition_entry env entry in + let body, () = Future.force entry.const_entry_body in + let ctx' = Evd.merge_context_set ~sideff:true Evd.univ_rigid (Evd.from_ctx ctx') (snd body) in + Inductiveops.control_only_guard env ctx' (EConstr.of_constr (fst body)); + (fst body), entry.const_entry_type, Evd.evar_universe_context ctx' + +let obligation_terminator name num guard hook auto pf = + let open Proof_global in + let term = Lemmas.universe_proof_terminator guard hook in + match pf with + | Admitted _ -> apply_terminator term pf + | Proved (opq, id, proof) -> + let (_, (entry, uctx, _)) = Pfedit.cook_this_proof proof in + let env = Global.env () in + let entry = Safe_typing.inline_private_constants_in_definition_entry env entry in + let ty = entry.Entries.const_entry_type in + let (body, cstr), () = Future.force entry.Entries.const_entry_body in + let sigma = Evd.from_ctx uctx in + let sigma = Evd.merge_context_set ~sideff:true Evd.univ_rigid sigma cstr in + Inductiveops.control_only_guard (Global.env ()) sigma (EConstr.of_constr body); + (** Declare the obligation ourselves and drop the hook *) + let prg = get_info (ProgMap.find name !from_prg) in + (** Ensure universes are substituted properly in body and type *) + let body = EConstr.to_constr sigma (EConstr.of_constr body) in + let ty = Option.map (fun x -> EConstr.to_constr sigma (EConstr.of_constr x)) ty in + let ctx = Evd.evar_universe_context sigma in + let obls, rem = prg.prg_obligations in + let obl = obls.(num) in + let status = + match obl.obl_status, opq with + | (_, Evar_kinds.Expand), Vernacexpr.Opaque -> err_not_transp () + | (true, _), Vernacexpr.Opaque -> err_not_transp () + | (false, _), Vernacexpr.Opaque -> Evar_kinds.Define true + | (_, Evar_kinds.Define true), Vernacexpr.Transparent -> + Evar_kinds.Define false + | (_, status), Vernacexpr.Transparent -> status + in + let obl = { obl with obl_status = false, status } in + let ctx = + if pi2 prg.prg_kind then ctx + else UState.union prg.prg_ctx ctx + in + let uctx = UState.const_univ_entry ~poly:(pi2 prg.prg_kind) ctx in + let (defined, obl) = declare_obligation prg obl body ty uctx in + let obls = Array.copy obls in + let _ = obls.(num) <- obl in + let prg_ctx = + if pi2 (prg.prg_kind) then (* Polymorphic *) + (** We merge the new universes and constraints of the + polymorphic obligation with the existing ones *) + UState.union prg.prg_ctx ctx + else + (** The first obligation, if defined, + declares the univs of the constant, + each subsequent obligation declares its own additional + universes and constraints if any *) + if defined then UState.make (Global.universes ()) + else ctx + in + let prg = { prg with prg_ctx } in + try + ignore (update_obls prg obls (pred rem)); + if pred rem > 0 then + begin + let deps = dependencies obls num in + if not (Int.Set.is_empty deps) then + ignore (auto (Some name) None deps) + end + with e when CErrors.noncritical e -> + let e = CErrors.push e in + pperror (CErrors.iprint (ExplainErr.process_vernac_interp_error e)) + +let obligation_hook prg obl num auto ctx' _ gr = + let obls, rem = prg.prg_obligations in + let cst = match gr with ConstRef cst -> cst | _ -> assert false in + let transparent = evaluable_constant cst (Global.env ()) in + let () = match obl.obl_status with + (true, Evar_kinds.Expand) + | (true, Evar_kinds.Define true) -> + if not transparent then err_not_transp () + | _ -> () +in + let ctx' = match ctx' with None -> prg.prg_ctx | Some ctx' -> ctx' in + let inst, ctx' = + if not (pi2 prg.prg_kind) (* Not polymorphic *) then + (* The universe context was declared globally, we continue + from the new global environment. *) + let evd = Evd.from_env (Global.env ()) in + let ctx' = Evd.merge_universe_subst evd (Evd.universe_subst (Evd.from_ctx ctx')) in + Univ.Instance.empty, Evd.evar_universe_context ctx' + else + (* We get the right order somehow, but surely it could be enforced in a clearer way. *) + let uctx = UState.context ctx' in + Univ.UContext.instance uctx, ctx' + in + let obl = { obl with obl_body = Some (DefinedObl (cst, inst)) } in + let () = if transparent then add_hint true prg cst in + let obls = Array.copy obls in + let _ = obls.(num) <- obl in + let prg = { prg with prg_ctx = ctx' } in + let () = + try ignore (update_obls prg obls (pred rem)) + with e when CErrors.noncritical e -> + let e = CErrors.push e in + pperror (CErrors.iprint (ExplainErr.process_vernac_interp_error e)) + in + if pred rem > 0 then begin + let deps = dependencies obls num in + if not (Int.Set.is_empty deps) then + ignore (auto (Some prg.prg_name) None deps) + end + +let rec solve_obligation prg num tac = + let user_num = succ num in + let obls, rem = prg.prg_obligations in + let obl = obls.(num) in + let remaining = deps_remaining obls obl.obl_deps in + let () = + if not (Option.is_empty obl.obl_body) then + pperror (str "Obligation" ++ spc () ++ int user_num ++ str "already" ++ spc() ++ str "solved."); + if not (List.is_empty remaining) then + pperror (str "Obligation " ++ int user_num ++ str " depends on obligation(s) " + ++ str (string_of_list ", " (fun x -> string_of_int (succ x)) remaining)); + in + let obl = subst_deps_obl obls obl in + let kind = kind_of_obligation (pi2 prg.prg_kind) (snd obl.obl_status) in + let evd = Evd.from_ctx prg.prg_ctx in + let evd = Evd.update_sigma_env evd (Global.env ()) in + let auto n tac oblset = auto_solve_obligations n ~oblset tac in + let terminator guard hook = + Proof_global.make_terminator + (obligation_terminator prg.prg_name num guard hook auto) in + let hook ctx = Lemmas.mk_hook (obligation_hook prg obl num auto ctx) in + let () = Lemmas.start_proof_univs ~sign:prg.prg_sign obl.obl_name kind evd (EConstr.of_constr obl.obl_type) ~terminator hook in + let _ = Pfedit.by !default_tactic in + Option.iter (fun tac -> Proof_global.set_endline_tactic tac) tac + +and obligation (user_num, name, typ) tac = + let num = pred user_num in + let prg = get_prog_err name in + let obls, rem = prg.prg_obligations in + if num >= 0 && num < Array.length obls then + let obl = obls.(num) in + match obl.obl_body with + None -> solve_obligation prg num tac + | Some r -> error "Obligation already solved" + else error (sprintf "Unknown obligation number %i" (succ num)) + + +and solve_obligation_by_tac prg obls i tac = + let obl = obls.(i) in + match obl.obl_body with + | Some _ -> None + | None -> + try + if List.is_empty (deps_remaining obls obl.obl_deps) then + let obl = subst_deps_obl obls obl in + let tac = + match tac with + | Some t -> t + | None -> + match obl.obl_tac with + | Some t -> t + | None -> !default_tactic + in + let evd = Evd.from_ctx prg.prg_ctx in + let evd = Evd.update_sigma_env evd (Global.env ()) in + let t, ty, ctx = + solve_by_tac obl.obl_name (evar_of_obligation obl) tac + (pi2 prg.prg_kind) (Evd.evar_universe_context evd) + in + let uctx = if pi2 prg.prg_kind + then Polymorphic_const_entry (UState.context ctx) + else Monomorphic_const_entry (UState.context_set ctx) + in + let prg = {prg with prg_ctx = ctx} in + let def, obl' = declare_obligation prg obl t ty uctx in + obls.(i) <- obl'; + if def && not (pi2 prg.prg_kind) then ( + (* Declare the term constraints with the first obligation only *) + let evd = Evd.from_env (Global.env ()) in + let evd = Evd.merge_universe_subst evd (Evd.universe_subst (Evd.from_ctx ctx)) in + let ctx' = Evd.evar_universe_context evd in + Some {prg with prg_ctx = ctx'}) + else Some prg + else None + with e when CErrors.noncritical e -> + let (e, _) = CErrors.push e in + match e with + | Refiner.FailError (_, s) -> + user_err ?loc:(fst obl.obl_location) ~hdr:"solve_obligation" (Lazy.force s) + | e -> None (* FIXME really ? *) + +and solve_prg_obligations prg ?oblset tac = + let obls, rem = prg.prg_obligations in + let rem = ref rem in + let obls' = Array.copy obls in + let set = ref Int.Set.empty in + let p = match oblset with + | None -> (fun _ -> true) + | Some s -> set := s; + (fun i -> Int.Set.mem i !set) + in + let prgref = ref prg in + let _ = + Array.iteri (fun i x -> + if p i then + match solve_obligation_by_tac !prgref obls' i tac with + | None -> () + | Some prg' -> + prgref := prg'; + let deps = dependencies obls i in + (set := Int.Set.union !set deps; + decr rem)) + obls' + in + update_obls !prgref obls' !rem + +and solve_obligations n tac = + let prg = get_prog_err n in + solve_prg_obligations prg tac + +and solve_all_obligations tac = + ProgMap.iter (fun k v -> ignore(solve_prg_obligations (get_info v) tac)) !from_prg + +and try_solve_obligation n prg tac = + let prg = get_prog prg in + let obls, rem = prg.prg_obligations in + let obls' = Array.copy obls in + match solve_obligation_by_tac prg obls' n tac with + | Some prg' -> ignore(update_obls prg' obls' (pred rem)) + | None -> () + +and try_solve_obligations n tac = + try ignore (solve_obligations n tac) with NoObligations _ -> () + +and auto_solve_obligations n ?oblset tac : progress = + Flags.if_verbose Feedback.msg_info (str "Solving obligations automatically..."); + try solve_prg_obligations (get_prog_err n) ?oblset tac with NoObligations _ -> Dependent + +open Pp +let show_obligations_of_prg ?(msg=true) prg = + let n = prg.prg_name in + let obls, rem = prg.prg_obligations in + let showed = ref 5 in + if msg then Feedback.msg_info (int rem ++ str " obligation(s) remaining: "); + Array.iteri (fun i x -> + match x.obl_body with + | None -> + if !showed > 0 then ( + decr showed; + let x = subst_deps_obl obls x in + Feedback.msg_info (str "Obligation" ++ spc() ++ int (succ i) ++ spc () ++ + str "of" ++ spc() ++ Id.print n ++ str ":" ++ spc () ++ + hov 1 (Printer.pr_constr_env (Global.env ()) Evd.empty x.obl_type ++ + str "." ++ fnl ()))) + | Some _ -> ()) + obls + +let show_obligations ?(msg=true) n = + let progs = match n with + | None -> all_programs () + | Some n -> + try [ProgMap.find n !from_prg] + with Not_found -> raise (NoObligations (Some n)) + in List.iter (fun x -> show_obligations_of_prg ~msg (get_info x)) progs + +let show_term n = + let prg = get_prog_err n in + let n = prg.prg_name in + (Id.print n ++ spc () ++ str":" ++ spc () ++ + Printer.pr_constr_env (Global.env ()) Evd.empty prg.prg_type ++ spc () ++ str ":=" ++ fnl () + ++ Printer.pr_constr_env (Global.env ()) Evd.empty prg.prg_body) + +let add_definition n ?term t ctx ?(univdecl=Univdecls.default_univ_decl) + ?(implicits=[]) ?(kind=Global,false,Definition) ?tactic + ?(reduce=reduce) ?(hook=Lemmas.mk_hook (fun _ _ _ -> ())) ?(opaque = false) obls = + let sign = Decls.initialize_named_context_for_proof () in + let info = Id.print n ++ str " has type-checked" in + let prg = init_prog_info sign ~opaque n univdecl term t ctx [] None [] obls implicits kind reduce hook in + let obls,_ = prg.prg_obligations in + if Int.equal (Array.length obls) 0 then ( + Flags.if_verbose Feedback.msg_info (info ++ str "."); + let cst = declare_definition prg in + Defined cst) + else ( + let len = Array.length obls in + let _ = Flags.if_verbose Feedback.msg_info (info ++ str ", generating " ++ int len ++ str (String.plural len " obligation")) in + progmap_add n (CEphemeron.create prg); + let res = auto_solve_obligations (Some n) tactic in + match res with + | Remain rem -> Flags.if_verbose (fun () -> show_obligations ~msg:false (Some n)) (); res + | _ -> res) + +let add_mutual_definitions l ctx ?(univdecl=Univdecls.default_univ_decl) ?tactic + ?(kind=Global,false,Definition) ?(reduce=reduce) + ?(hook=Lemmas.mk_hook (fun _ _ _ -> ())) ?(opaque = false) notations fixkind = + let sign = Decls.initialize_named_context_for_proof () in + let deps = List.map (fun (n, b, t, imps, obls) -> n) l in + List.iter + (fun (n, b, t, imps, obls) -> + let prg = init_prog_info sign ~opaque n univdecl (Some b) t ctx deps (Some fixkind) + notations obls imps kind reduce hook + in progmap_add n (CEphemeron.create prg)) l; + let _defined = + List.fold_left (fun finished x -> + if finished then finished + else + let res = auto_solve_obligations (Some x) tactic in + match res with + | Defined _ -> + (* If one definition is turned into a constant, + the whole block is defined. *) true + | _ -> false) + false deps + in () + +let admit_prog prg = + let obls, rem = prg.prg_obligations in + let obls = Array.copy obls in + Array.iteri + (fun i x -> + match x.obl_body with + | None -> + let x = subst_deps_obl obls x in + let ctx = Monomorphic_const_entry (UState.context_set prg.prg_ctx) in + let kn = Declare.declare_constant x.obl_name ~local:true + (ParameterEntry (None,(x.obl_type,ctx),None), IsAssumption Conjectural) + in + assumption_message x.obl_name; + obls.(i) <- { x with obl_body = Some (DefinedObl (kn, Univ.Instance.empty)) } + | Some _ -> ()) + obls; + ignore(update_obls prg obls 0) + +let rec admit_all_obligations () = + let prg = try Some (get_any_prog ()) with NoObligations _ -> None in + match prg with + | None -> () + | Some prg -> + admit_prog prg; + admit_all_obligations () + +let admit_obligations n = + match n with + | None -> admit_all_obligations () + | Some _ -> + let prg = get_prog_err n in + admit_prog prg + +let next_obligation n tac = + let prg = match n with + | None -> get_any_prog_err () + | Some _ -> get_prog_err n + in + let obls, rem = prg.prg_obligations in + let is_open _ x = Option.is_empty x.obl_body && List.is_empty (deps_remaining obls x.obl_deps) in + let i = match Array.findi is_open obls with + | Some i -> i + | None -> anomaly (Pp.str "Could not find a solvable obligation.") + in + solve_obligation prg i tac + +let init_program () = + Coqlib.check_required_library Coqlib.datatypes_module_name; + Coqlib.check_required_library ["Coq";"Init";"Specif"]; + Coqlib.check_required_library ["Coq";"Program";"Tactics"] + +let set_program_mode c = + if c then + if !Flags.program_mode then () + else begin + init_program (); + Flags.program_mode := true; + end diff --git a/vernac/obligations.mli b/vernac/obligations.mli new file mode 100644 index 00000000..cc2cacd8 --- /dev/null +++ b/vernac/obligations.mli @@ -0,0 +1,111 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* Exninfo.iexn -> Exninfo.iexn) Hook.t + +val check_evars : env -> evar_map -> unit + +val evar_dependencies : evar_map -> Evar.t -> Evar.Set.t +val sort_dependencies : (Evar.t * evar_info * Evar.Set.t) list -> (Evar.t * evar_info * Evar.Set.t) list + +(* env, id, evars, number of function prototypes to try to clear from + evars contexts, object and type *) +val eterm_obligations : env -> Id.t -> evar_map -> int -> + ?status:Evar_kinds.obligation_definition_status -> constr -> types -> + (Id.t * types * Evar_kinds.t Loc.located * + (bool * Evar_kinds.obligation_definition_status) * Int.Set.t * + unit Proofview.tactic option) array + (* Existential key, obl. name, type as product, + location of the original evar, associated tactic, + status and dependencies as indexes into the array *) + * ((Evar.t * Id.t) list * ((Id.t -> constr) -> constr -> constr)) * + constr * types + (* Translations from existential identifiers to obligation identifiers + and for terms with existentials to closed terms, given a + translation from obligation identifiers to constrs, new term, new type *) + +type obligation_info = + (Id.t * types * Evar_kinds.t Loc.located * + (bool * Evar_kinds.obligation_definition_status) * Int.Set.t * unit Proofview.tactic option) array + (* ident, type, location, (opaque or transparent, expand or define), + dependencies, tactic to solve it *) + +type progress = (* Resolution status of a program *) + | Remain of int (* n obligations remaining *) + | Dependent (* Dependent on other definitions *) + | Defined of global_reference (* Defined as id *) + +val default_tactic : unit Proofview.tactic ref + +val add_definition : Names.Id.t -> ?term:constr -> types -> + UState.t -> + ?univdecl:Univdecls.universe_decl -> (* Universe binders and constraints *) + ?implicits:(Constrexpr.explicitation * (bool * bool * bool)) list -> + ?kind:Decl_kinds.definition_kind -> + ?tactic:unit Proofview.tactic -> + ?reduce:(constr -> constr) -> + ?hook:(UState.t -> unit) Lemmas.declaration_hook -> ?opaque:bool -> obligation_info -> progress + +type notations = + (Misctypes.lstring * Constrexpr.constr_expr * Notation_term.scope_name option) list + +type fixpoint_kind = + | IsFixpoint of (Misctypes.lident option * Constrexpr.recursion_order_expr) list + | IsCoFixpoint + +val add_mutual_definitions : + (Names.Id.t * constr * types * + (Constrexpr.explicitation * (bool * bool * bool)) list * obligation_info) list -> + UState.t -> + ?univdecl:Univdecls.universe_decl -> (* Universe binders and constraints *) + ?tactic:unit Proofview.tactic -> + ?kind:Decl_kinds.definition_kind -> + ?reduce:(constr -> constr) -> + ?hook:(UState.t -> unit) Lemmas.declaration_hook -> ?opaque:bool -> + notations -> + fixpoint_kind -> unit + +val obligation : int * Names.Id.t option * Constrexpr.constr_expr option -> + Genarg.glob_generic_argument option -> unit + +val next_obligation : Names.Id.t option -> Genarg.glob_generic_argument option -> unit + +val solve_obligations : Names.Id.t option -> unit Proofview.tactic option -> progress +(* Number of remaining obligations to be solved for this program *) + +val solve_all_obligations : unit Proofview.tactic option -> unit + +val try_solve_obligation : int -> Names.Id.t option -> unit Proofview.tactic option -> unit + +val try_solve_obligations : Names.Id.t option -> unit Proofview.tactic option -> unit + +val show_obligations : ?msg:bool -> Names.Id.t option -> unit + +val show_term : Names.Id.t option -> Pp.t + +val admit_obligations : Names.Id.t option -> unit + +exception NoObligations of Names.Id.t option + +val explain_no_obligations : Names.Id.t option -> Pp.t + +val set_program_mode : bool -> unit + +type program_info +val program_tcc_summary_tag : program_info Id.Map.t Summary.Dyn.tag diff --git a/vernac/proof_using.ml b/vernac/proof_using.ml new file mode 100644 index 00000000..f8b085f3 --- /dev/null +++ b/vernac/proof_using.ml @@ -0,0 +1,192 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* known_names := x :: !known_names); + classify_function = (fun _ -> Dispose); + discharge_function = (fun _ -> None) + } + +let rec close_fwd e s = + let s' = + List.fold_left (fun s decl -> + let vb = match decl with + | LocalAssum _ -> Id.Set.empty + | LocalDef (_,b,_) -> global_vars_set e b + in + let vty = global_vars_set e (NamedDecl.get_type decl) in + let vbty = Id.Set.union vb vty in + if Id.Set.exists (fun v -> Id.Set.mem v s) vbty + then Id.Set.add (NamedDecl.get_id decl) (Id.Set.union s vbty) else s) + s (named_context e) + in + if Id.Set.equal s s' then s else close_fwd e s' + +let set_of_type env ty = + List.fold_left (fun acc ty -> + Id.Set.union (global_vars_set env ty) acc) + Id.Set.empty ty + +let full_set env = + List.fold_right Id.Set.add (List.map NamedDecl.get_id (named_context env)) Id.Set.empty + +let rec process_expr env e ty = + let rec aux = function + | SsEmpty -> Id.Set.empty + | SsType -> set_of_type env ty + | SsSingl { CAst.v = id } -> set_of_id env id + | SsUnion(e1,e2) -> Id.Set.union (aux e1) (aux e2) + | SsSubstr(e1,e2) -> Id.Set.diff (aux e1) (aux e2) + | SsCompl e -> Id.Set.diff (full_set env) (aux e) + | SsFwdClose e -> close_fwd env (aux e) + in + aux e + +and set_of_id env id = + if Id.to_string id = "All" then + List.fold_right Id.Set.add (List.map NamedDecl.get_id (named_context env)) Id.Set.empty + else if CList.mem_assoc_f Id.equal id !known_names then + process_expr env (CList.assoc_f Id.equal id !known_names) [] + else Id.Set.singleton id + +let process_expr env e ty = + let v_ty = set_of_type env ty in + let s = Id.Set.union v_ty (process_expr env e ty) in + Id.Set.elements s + +let name_set id expr = Lib.add_anonymous_leaf (in_nameset (id,expr)) + +let minimize_hyps env ids = + let rec aux ids = + let ids' = + Id.Set.fold (fun id alive -> + let impl_by_id = + Id.Set.remove id (really_needed env (Id.Set.singleton id)) in + if Id.Set.is_empty impl_by_id then alive + else Id.Set.diff alive impl_by_id) + ids ids in + if Id.Set.equal ids ids' then ids else aux ids' + in + aux ids + +let remove_ids_and_lets env s ids = + let not_ids id = not (Id.Set.mem id ids) in + let no_body id = named_body id env = None in + let deps id = really_needed env (Id.Set.singleton id) in + (Id.Set.filter (fun id -> + not_ids id && + (no_body id || + Id.Set.exists not_ids (Id.Set.filter no_body (deps id)))) s) + +let record_proof_using expr = + Aux_file.record_in_aux "suggest_proof_using" expr + +(* Variables in [skip] come from after the definition, so don't count + for "All". Used in the variable case since the env contains the + variable itself. *) +let suggest_common env ppid used ids_typ skip = + let module S = Id.Set in + let open Pp in + let print x = Feedback.msg_debug x in + let pr_set parens s = + let wrap ppcmds = + if parens && S.cardinal s > 1 then str "(" ++ ppcmds ++ str ")" + else ppcmds in + wrap (prlist_with_sep (fun _ -> str" ") Id.print (S.elements s)) in + + let needed = minimize_hyps env (remove_ids_and_lets env used ids_typ) in + let all_needed = really_needed env needed in + let all = List.fold_left (fun all d -> S.add (NamedDecl.get_id d) all) + S.empty (named_context env) + in + let all = S.diff all skip in + let fwd_typ = close_fwd env ids_typ in + if !Flags.debug then begin + print (str "All " ++ pr_set false all); + print (str "Type " ++ pr_set false ids_typ); + print (str "needed " ++ pr_set false needed); + print (str "all_needed " ++ pr_set false all_needed); + print (str "Type* " ++ pr_set false fwd_typ); + end; + let valid_exprs = ref [] in + let valid e = valid_exprs := e :: !valid_exprs in + if S.is_empty needed then valid (str "Type"); + if S.equal all_needed fwd_typ then valid (str "Type*"); + if S.equal all all_needed then valid(str "All"); + valid (pr_set false needed); + Feedback.msg_info ( + str"The proof of "++ ppid ++ spc() ++ + str "should start with one of the following commands:"++spc()++ + v 0 ( + prlist_with_sep cut (fun x->str"Proof using " ++x++ str". ") !valid_exprs)); + if !Flags.record_aux_file + then + let s = string_of_ppcmds (prlist_with_sep (fun _ -> str";") (fun x->x) !valid_exprs) in + record_proof_using s + +let suggest_proof_using = ref false + +let _ = + Goptions.declare_bool_option + { Goptions.optdepr = false; + Goptions.optname = "suggest Proof using"; + Goptions.optkey = ["Suggest";"Proof";"Using"]; + Goptions.optread = (fun () -> !suggest_proof_using); + Goptions.optwrite = ((:=) suggest_proof_using) } + +let suggest_constant env kn = + if !suggest_proof_using + then begin + let open Declarations in + let body = lookup_constant kn env in + let used = Id.Set.of_list @@ List.map NamedDecl.get_id body.const_hyps in + let ids_typ = global_vars_set env body.const_type in + suggest_common env (Printer.pr_constant env kn) used ids_typ Id.Set.empty + end + +let suggest_variable env id = + if !suggest_proof_using + then begin + match lookup_named id env with + | LocalDef (_,body,typ) -> + let ids_typ = global_vars_set env typ in + let ids_body = global_vars_set env body in + let used = Id.Set.union ids_body ids_typ in + suggest_common env (Id.print id) used ids_typ (Id.Set.singleton id) + | LocalAssum _ -> assert false + end + +let value = ref None + +let using_to_string us = Pp.string_of_ppcmds (Ppvernac.pr_using us) +let using_from_string us = Pcoq.Gram.(entry_parse G_vernac.section_subset_expr (parsable (Stream.of_string us))) + +let _ = + Goptions.declare_stringopt_option + { Goptions.optdepr = false; + Goptions.optname = "default value for Proof using"; + Goptions.optkey = ["Default";"Proof";"Using"]; + Goptions.optread = (fun () -> Option.map using_to_string !value); + Goptions.optwrite = (fun b -> value := Option.map using_from_string b); + } + +let get_default_proof_using () = !value diff --git a/vernac/proof_using.mli b/vernac/proof_using.mli new file mode 100644 index 00000000..7d1110aa --- /dev/null +++ b/vernac/proof_using.mli @@ -0,0 +1,23 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* Vernacexpr.section_subset_expr -> Constr.types list -> + Names.Id.t list + +val name_set : Names.Id.t -> Vernacexpr.section_subset_expr -> unit + +val suggest_constant : Environ.env -> Names.Constant.t -> unit + +val suggest_variable : Environ.env -> Names.Id.t -> unit + +val get_default_proof_using : unit -> Vernacexpr.section_subset_expr option diff --git a/vernac/record.ml b/vernac/record.ml new file mode 100644 index 00000000..78e68e8a --- /dev/null +++ b/vernac/record.ml @@ -0,0 +1,620 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* !primitive_flag) ; + optwrite = (fun b -> primitive_flag := b) } + +let typeclasses_strict = ref false +let _ = + declare_bool_option + { optdepr = false; + optname = "strict typeclass resolution"; + optkey = ["Typeclasses";"Strict";"Resolution"]; + optread = (fun () -> !typeclasses_strict); + optwrite = (fun b -> typeclasses_strict := b); } + +let typeclasses_unique = ref false +let _ = + declare_bool_option + { optdepr = false; + optname = "unique typeclass instances"; + optkey = ["Typeclasses";"Unique";"Instances"]; + optread = (fun () -> !typeclasses_unique); + optwrite = (fun b -> typeclasses_unique := b); } + +let interp_fields_evars env sigma impls_env nots l = + List.fold_left2 + (fun (env, sigma, uimpls, params, impls) no ({CAst.loc;v=i}, b, t) -> + let sigma, (t', impl) = interp_type_evars_impls env sigma ~impls t in + let sigma, b' = + Option.cata (fun x -> on_snd (fun x -> Some (fst x)) @@ + interp_casted_constr_evars_impls env sigma ~impls x t') (sigma,None) b in + let impls = + match i with + | Anonymous -> impls + | Name id -> Id.Map.add id (compute_internalization_data env sigma Constrintern.Method t' impl) impls + in + let d = match b' with + | None -> LocalAssum (i,t') + | Some b' -> LocalDef (i,b',t') + in + List.iter (Metasyntax.set_notation_for_interpretation env impls) no; + (EConstr.push_rel d env, sigma, impl :: uimpls, d::params, impls)) + (env, sigma, [], [], impls_env) nots l + +let compute_constructor_level evars env l = + List.fold_right (fun d (env, univ) -> + let univ = + if is_local_assum d then + let s = Retyping.get_sort_of env evars (RelDecl.get_type d) in + Univ.sup (univ_of_sort s) univ + else univ + in (EConstr.push_rel d env, univ)) + l (env, Univ.type0m_univ) + +let binder_of_decl = function + | Vernacexpr.AssumExpr(n,t) -> (n,None,t) + | Vernacexpr.DefExpr(n,c,t) -> + (n,Some c, match t with Some c -> c + | None -> CAst.make ?loc:n.CAst.loc @@ CHole (None, Misctypes.IntroAnonymous, None)) + +let binders_of_decls = List.map binder_of_decl + +let typecheck_params_and_fields finite def id poly pl t ps nots fs = + let env0 = Global.env () in + let sigma, decl = Univdecls.interp_univ_decl_opt env0 pl in + let _ = + let error bk {CAst.loc; v=name} = + match bk, name with + | Default _, Anonymous -> + user_err ?loc ~hdr:"record" (str "Record parameters must be named") + | _ -> () + in + List.iter + (function CLocalDef (b, _, _) -> error default_binder_kind b + | CLocalAssum (ls, bk, ce) -> List.iter (error bk) ls + | CLocalPattern {CAst.loc} -> + Loc.raise ?loc (Stream.Error "pattern with quote not allowed in record parameters")) ps + in + let sigma, (impls_env, ((env1,newps), imps)) = interp_context_evars env0 sigma ps in + let sigma, typ, sort, template = match t with + | Some t -> + let env = EConstr.push_rel_context newps env0 in + let poly = + match t with + | { CAst.v = CSort (Misctypes.GType []) } -> true | _ -> false in + let sigma, s = interp_type_evars env sigma ~impls:empty_internalization_env t in + let sred = Reductionops.whd_allnolet env sigma s in + (match EConstr.kind sigma sred with + | Sort s' -> + let s' = EConstr.ESorts.kind sigma s' in + (if poly then + match Evd.is_sort_variable sigma s' with + | Some l -> + let sigma = Evd.make_flexible_variable sigma ~algebraic:true l in + sigma, s, s', true + | None -> + sigma, s, s', false + else sigma, s, s', false) + | _ -> user_err ?loc:(constr_loc t) (str"Sort expected.")) + | None -> + let uvarkind = Evd.univ_flexible_alg in + let sigma, s = Evd.new_sort_variable uvarkind sigma in + sigma, EConstr.mkSort s, s, true + in + let arity = EConstr.it_mkProd_or_LetIn typ newps in + let env_ar = EConstr.push_rel_context newps (EConstr.push_rel (LocalAssum (Name id,arity)) env0) in + let assums = List.filter is_local_assum newps in + let params = List.map (RelDecl.get_name %> Name.get_id) assums in + let ty = Inductive (params,(finite != Declarations.BiFinite)) in + let impls_env = compute_internalization_env env0 sigma ~impls:impls_env ty [id] [arity] [imps] in + let env2,sigma,impls,newfs,data = + interp_fields_evars env_ar sigma impls_env nots (binders_of_decls fs) + in + let sigma = + Pretyping.solve_remaining_evars Pretyping.all_and_fail_flags env_ar sigma Evd.empty in + let sigma, typ = + let _, univ = compute_constructor_level sigma env_ar newfs in + if not def && (Sorts.is_prop sort || + (Sorts.is_set sort && is_impredicative_set env0)) then + sigma, typ + else + let sigma = Evd.set_leq_sort env_ar sigma (Type univ) sort in + if Univ.is_small_univ univ && + Option.cata (Evd.is_flexible_level sigma) false (Evd.is_sort_variable sigma sort) then + (* We can assume that the level in aritysort is not constrained + and clear it, if it is flexible *) + Evd.set_eq_sort env_ar sigma (Prop Pos) sort, + EConstr.mkSort (Sorts.sort_of_univ univ) + else sigma, typ + in + let sigma, _ = Evarutil.nf_evars_and_universes sigma in + let newfs = List.map (EConstr.to_rel_decl sigma) newfs in + let newps = List.map (EConstr.to_rel_decl sigma) newps in + let typ = EConstr.to_constr sigma typ in + let ce t = Pretyping.check_evars env0 Evd.empty sigma (EConstr.of_constr t) in + let univs = Evd.check_univ_decl ~poly sigma decl in + let ubinders = Evd.universe_binders sigma in + List.iter (iter_constr ce) (List.rev newps); + List.iter (iter_constr ce) (List.rev newfs); + ubinders, univs, typ, template, imps, newps, impls, newfs + +let degenerate_decl decl = + let id = match RelDecl.get_name decl with + | Name id -> id + | Anonymous -> anomaly (Pp.str "Unnamed record variable.") in + match decl with + | LocalAssum (_,t) -> (id, LocalAssumEntry t) + | LocalDef (_,b,_) -> (id, LocalDefEntry b) + +type record_error = + | MissingProj of Id.t * Id.t list + | BadTypedProj of Id.t * env * Type_errors.type_error + +let warn_cannot_define_projection = + CWarnings.create ~name:"cannot-define-projection" ~category:"records" + (fun msg -> hov 0 msg) + +(* If a projection is not definable, we throw an error if the user +asked it to be a coercion. Otherwise, we just print an info +message. The user might still want to name the field of the record. *) +let warning_or_error coe indsp err = + let st = match err with + | MissingProj (fi,projs) -> + let s,have = if List.length projs > 1 then "s","were" else "","was" in + (Id.print fi ++ + strbrk" cannot be defined because the projection" ++ str s ++ spc () ++ + prlist_with_sep pr_comma Id.print projs ++ spc () ++ str have ++ + strbrk " not defined.") + | BadTypedProj (fi,ctx,te) -> + match te with + | ElimArity (_,_,_,_,Some (_,_,NonInformativeToInformative)) -> + (Id.print fi ++ + strbrk" cannot be defined because it is informative and " ++ + Printer.pr_inductive (Global.env()) indsp ++ + strbrk " is not.") + | ElimArity (_,_,_,_,Some (_,_,StrongEliminationOnNonSmallType)) -> + (Id.print fi ++ + strbrk" cannot be defined because it is large and " ++ + Printer.pr_inductive (Global.env()) indsp ++ + strbrk " is not.") + | _ -> + (Id.print fi ++ strbrk " cannot be defined because it is not typable.") + in + if coe then user_err ~hdr:"structure" st; + warn_cannot_define_projection (hov 0 st) + +type field_status = + | NoProjection of Name.t + | Projection of constr + +exception NotDefinable of record_error + +(* This replaces previous projection bodies in current projection *) +(* Undefined projs are collected and, at least one undefined proj occurs *) +(* in the body of current projection then the latter can not be defined *) +(* [c] is defined in ctxt [[params;fields]] and [l] is an instance of *) +(* [[fields]] defined in ctxt [[params;x:ind]] *) +let subst_projection fid l c = + let lv = List.length l in + let bad_projs = ref [] in + let rec substrec depth c = match Constr.kind c with + | Rel k -> + (* We are in context [[params;fields;x:ind;...depth...]] *) + if k <= depth+1 then + c + else if k-depth-1 <= lv then + match List.nth l (k-depth-2) with + | Projection t -> lift depth t + | NoProjection (Name id) -> bad_projs := id :: !bad_projs; mkRel k + | NoProjection Anonymous -> + user_err (str "Field " ++ Id.print fid ++ + str " depends on the " ++ pr_nth (k-depth-1) ++ str + " field which has no name.") + else + mkRel (k-lv) + | _ -> Constr.map_with_binders succ substrec depth c + in + let c' = lift 1 c in (* to get [c] defined in ctxt [[params;fields;x:ind]] *) + let c'' = substrec 0 c' in + if not (List.is_empty !bad_projs) then + raise (NotDefinable (MissingProj (fid,List.rev !bad_projs))); + c'' + +let instantiate_possibly_recursive_type indu paramdecls fields = + let subst = List.map_i (fun i _ -> mkRel i) 1 paramdecls in + Termops.substl_rel_context (subst@[mkIndU indu]) fields + +let warn_non_primitive_record = + CWarnings.create ~name:"non-primitive-record" ~category:"record" + (fun (env,indsp) -> + (hov 0 (str "The record " ++ Printer.pr_inductive env indsp ++ + strbrk" could not be defined as a primitive record"))) + +(* We build projections *) +let declare_projections indsp ctx ?(kind=StructureComponent) binder_name coers ubinders fieldimpls fields = + let env = Global.env() in + let (mib,mip) = Global.lookup_inductive indsp in + let poly = Declareops.inductive_is_polymorphic mib in + let u = match ctx with + | Polymorphic_const_entry ctx -> Univ.UContext.instance ctx + | Monomorphic_const_entry ctx -> Univ.Instance.empty + in + let paramdecls = Inductive.inductive_paramdecls (mib, u) in + let indu = indsp, u in + let r = mkIndU (indsp,u) in + let rp = applist (r, Context.Rel.to_extended_list mkRel 0 paramdecls) in + let paramargs = Context.Rel.to_extended_list mkRel 1 paramdecls in (*def in [[params;x:rp]]*) + let x = Name binder_name in + let fields = instantiate_possibly_recursive_type indu paramdecls fields in + let lifted_fields = Termops.lift_rel_context 1 fields in + let primitive = + if !primitive_flag then + let is_primitive = + match mib.mind_record with + | Some (Some _) -> true + | Some None | None -> false + in + if not is_primitive then + warn_non_primitive_record (env,indsp); + is_primitive + else false + in + let (_,_,kinds,sp_projs,_) = + List.fold_left3 + (fun (nfi,i,kinds,sp_projs,subst) coe decl impls -> + let fi = RelDecl.get_name decl in + let ti = RelDecl.get_type decl in + let (sp_projs,i,subst) = + match fi with + | Anonymous -> + (None::sp_projs,i,NoProjection fi::subst) + | Name fid -> try + let kn, term = + if is_local_assum decl && primitive then + (** Already defined in the kernel silently *) + let gr = Nametab.locate (Libnames.qualid_of_ident fid) in + let kn = destConstRef gr in + Declare.definition_message fid; + Universes.register_universe_binders gr ubinders; + kn, mkProj (Projection.make kn false,mkRel 1) + else + let ccl = subst_projection fid subst ti in + let body = match decl with + | LocalDef (_,ci,_) -> subst_projection fid subst ci + | LocalAssum _ -> + (* [ccl] is defined in context [params;x:rp] *) + (* [ccl'] is defined in context [params;x:rp;x:rp] *) + let ccl' = liftn 1 2 ccl in + let p = mkLambda (x, lift 1 rp, ccl') in + let branch = it_mkLambda_or_LetIn (mkRel nfi) lifted_fields in + let ci = Inductiveops.make_case_info env indsp LetStyle in + mkCase (ci, p, mkRel 1, [|branch|]) + in + let proj = + it_mkLambda_or_LetIn (mkLambda (x,rp,body)) paramdecls in + let projtyp = + it_mkProd_or_LetIn (mkProd (x,rp,ccl)) paramdecls in + try + let entry = { + const_entry_body = + Future.from_val (Safe_typing.mk_pure_proof proj); + const_entry_secctx = None; + const_entry_type = Some projtyp; + const_entry_universes = ctx; + const_entry_opaque = false; + const_entry_inline_code = false; + const_entry_feedback = None } in + let k = (DefinitionEntry entry,IsDefinition kind) in + let kn = declare_constant ~internal:InternalTacticRequest fid k in + let constr_fip = + let proj_args = (*Rel 1 refers to "x"*) paramargs@[mkRel 1] in + applist (mkConstU (kn,u),proj_args) + in + Declare.definition_message fid; + Universes.register_universe_binders (ConstRef kn) ubinders; + kn, constr_fip + with Type_errors.TypeError (ctx,te) -> + raise (NotDefinable (BadTypedProj (fid,ctx,te))) + in + let refi = ConstRef kn in + Impargs.maybe_declare_manual_implicits false refi impls; + if coe then begin + let cl = Class.class_of_global (IndRef indsp) in + Class.try_add_new_coercion_with_source refi ~local:false poly ~source:cl + end; + let i = if is_local_assum decl then i+1 else i in + (Some kn::sp_projs, i, Projection term::subst) + with NotDefinable why -> + warning_or_error coe indsp why; + (None::sp_projs,i,NoProjection fi::subst) in + (nfi-1,i,(fi, is_local_assum decl)::kinds,sp_projs,subst)) + (List.length fields,0,[],[],[]) coers (List.rev fields) (List.rev fieldimpls) + in (kinds,sp_projs) + +open Typeclasses + +let declare_structure finite ubinders univs id idbuild paramimpls params arity template + fieldimpls fields ?(kind=StructureComponent) ?name is_coe coers = + let nparams = List.length params and nfields = List.length fields in + let args = Context.Rel.to_extended_list mkRel nfields params in + let ind = applist (mkRel (1+nparams+nfields), args) in + let type_constructor = it_mkProd_or_LetIn ind fields in + let template, ctx = + match univs with + | Monomorphic_ind_entry ctx -> + template, Monomorphic_const_entry Univ.ContextSet.empty + | Polymorphic_ind_entry ctx -> + false, Polymorphic_const_entry ctx + | Cumulative_ind_entry cumi -> + false, Polymorphic_const_entry (Univ.CumulativityInfo.univ_context cumi) + in + let binder_name = + match name with + | None -> Id.of_string (Unicode.lowercase_first_char (Id.to_string id)) + | Some n -> n + in + let mie_ind = + { mind_entry_typename = id; + mind_entry_arity = arity; + mind_entry_template = template; + mind_entry_consnames = [idbuild]; + mind_entry_lc = [type_constructor] } + in + let mie = + { mind_entry_params = List.map degenerate_decl params; + mind_entry_record = Some (if !primitive_flag then Some binder_name else None); + mind_entry_finite = finite; + mind_entry_inds = [mie_ind]; + mind_entry_private = None; + mind_entry_universes = univs; + } + in + let mie = InferCumulativity.infer_inductive (Global.env ()) mie in + let kn = ComInductive.declare_mutual_inductive_with_eliminations mie ubinders [(paramimpls,[])] in + let rsp = (kn,0) in (* This is ind path of idstruc *) + let cstr = (rsp,1) in + let kinds,sp_projs = declare_projections rsp ctx ~kind binder_name coers ubinders fieldimpls fields in + let build = ConstructRef cstr in + let poly = match ctx with | Polymorphic_const_entry _ -> true | Monomorphic_const_entry _ -> false in + let () = if is_coe then Class.try_add_new_coercion build ~local:false poly in + Recordops.declare_structure(rsp,cstr,List.rev kinds,List.rev sp_projs); + rsp + +let implicits_of_context ctx = + List.map_i (fun i name -> + let explname = + match name with + | Name n -> Some n + | Anonymous -> None + in ExplByPos (i, explname), (true, true, true)) + 1 (List.rev (Anonymous :: (List.map RelDecl.get_name ctx))) + +let declare_class finite def cum ubinders univs id idbuild paramimpls params arity + template fieldimpls fields ?(kind=StructureComponent) is_coe coers priorities = + let fieldimpls = + (* Make the class implicit in the projections, and the params if applicable. *) + let len = List.length params in + let impls = implicits_of_context params in + List.map (fun x -> impls @ Impargs.lift_implicits (succ len) x) fieldimpls + in + let binder_name = Namegen.next_ident_away (snd id) (Termops.vars_of_env (Global.env())) in + let impl, projs = + match fields with + | [LocalAssum (Name proj_name, field) | LocalDef (Name proj_name, _, field)] when def -> + let class_body = it_mkLambda_or_LetIn field params in + let class_type = it_mkProd_or_LetIn arity params in + let class_entry = + Declare.definition_entry ~types:class_type ~univs class_body in + let cst = Declare.declare_constant (snd id) + (DefinitionEntry class_entry, IsDefinition Definition) + in + let cstu = (cst, match univs with + | Polymorphic_const_entry univs -> Univ.UContext.instance univs + | Monomorphic_const_entry _ -> Univ.Instance.empty) + in + let inst_type = appvectc (mkConstU cstu) + (Termops.rel_vect 0 (List.length params)) in + let proj_type = + it_mkProd_or_LetIn (mkProd(Name binder_name, inst_type, lift 1 field)) params in + let proj_body = + it_mkLambda_or_LetIn (mkLambda (Name binder_name, inst_type, mkRel 1)) params in + let proj_entry = Declare.definition_entry ~types:proj_type ~univs proj_body in + let proj_cst = Declare.declare_constant proj_name + (DefinitionEntry proj_entry, IsDefinition Definition) + in + let cref = ConstRef cst in + Impargs.declare_manual_implicits false cref [paramimpls]; + Universes.register_universe_binders cref ubinders; + Impargs.declare_manual_implicits false (ConstRef proj_cst) [List.hd fieldimpls]; + Universes.register_universe_binders (ConstRef proj_cst) ubinders; + Classes.set_typeclass_transparency (EvalConstRef cst) false false; + let sub = match List.hd coers with + | Some b -> Some ((if b then Backward else Forward), List.hd priorities) + | None -> None + in + cref, [Name proj_name, sub, Some proj_cst] + | _ -> + let univs = + match univs with + | Polymorphic_const_entry univs -> + if cum then + Cumulative_ind_entry (Univ.CumulativityInfo.from_universe_context univs) + else + Polymorphic_ind_entry univs + | Monomorphic_const_entry univs -> + Monomorphic_ind_entry univs + in + let ind = declare_structure Declarations.BiFinite ubinders univs (snd id) idbuild paramimpls + params arity template fieldimpls fields + ~kind:Method ~name:binder_name false (List.map (fun _ -> false) fields) + in + let coers = List.map2 (fun coe pri -> + Option.map (fun b -> + if b then Backward, pri else Forward, pri) coe) + coers priorities + in + let l = List.map3 (fun decl b y -> RelDecl.get_name decl, b, y) + (List.rev fields) coers (Recordops.lookup_projections ind) + in IndRef ind, l + in + let ctx_context = + List.map (fun decl -> + match Typeclasses.class_of_constr Evd.empty (EConstr.of_constr (RelDecl.get_type decl)) with + | Some (_, ((cl,_), _)) -> Some cl.cl_impl + | None -> None) + params, params + in + let univs, ctx_context, fields = + match univs with + | Polymorphic_const_entry univs -> + let usubst, auctx = Univ.abstract_universes univs in + let usubst = Univ.make_instance_subst usubst in + let map c = Vars.subst_univs_level_constr usubst c in + let fields = Context.Rel.map map fields in + let ctx_context = on_snd (fun d -> Context.Rel.map map d) ctx_context in + auctx, ctx_context, fields + | Monomorphic_const_entry _ -> + Univ.AUContext.empty, ctx_context, fields + in + let k = + { cl_univs = univs; + cl_impl = impl; + cl_strict = !typeclasses_strict; + cl_unique = !typeclasses_unique; + cl_context = ctx_context; + cl_props = fields; + cl_projs = projs } + in + add_class k; impl + + +let add_constant_class cst = + let ty, univs = Global.type_of_global_in_context (Global.env ()) (ConstRef cst) in + let ctx, arity = decompose_prod_assum ty in + let tc = + { cl_univs = univs; + cl_impl = ConstRef cst; + cl_context = (List.map (const None) ctx, ctx); + cl_props = [LocalAssum (Anonymous, arity)]; + cl_projs = []; + cl_strict = !typeclasses_strict; + cl_unique = !typeclasses_unique + } + in add_class tc; + set_typeclass_transparency (EvalConstRef cst) false false + +let add_inductive_class ind = + let mind, oneind = Global.lookup_inductive ind in + let k = + let ctx = oneind.mind_arity_ctxt in + let univs = Declareops.inductive_polymorphic_context mind in + let env = push_context ~strict:false (Univ.AUContext.repr univs) (Global.env ()) in + let env = push_rel_context ctx env in + let inst = Univ.make_abstract_instance univs in + let ty = Inductive.type_of_inductive env ((mind, oneind), inst) in + { cl_univs = univs; + cl_impl = IndRef ind; + cl_context = List.map (const None) ctx, ctx; + cl_props = [LocalAssum (Anonymous, ty)]; + cl_projs = []; + cl_strict = !typeclasses_strict; + cl_unique = !typeclasses_unique } + in add_class k + +let declare_existing_class g = + match g with + | ConstRef x -> add_constant_class x + | IndRef x -> add_inductive_class x + | _ -> user_err ~hdr:"declare_existing_class" + (Pp.str"Unsupported class type, only constants and inductives are allowed") + +open Vernacexpr + +(* [fs] corresponds to fields and [ps] to parameters; [coers] is a + list telling if the corresponding fields must me declared as coercions + or subinstances. *) +let definition_structure (kind,cum,poly,finite,(is_coe,({CAst.loc;v=idstruc},pl)),ps,cfs,idbuild,s) = + let cfs,notations = List.split cfs in + let cfs,priorities = List.split cfs in + let coers,fs = List.split cfs in + let extract_name acc = function + Vernacexpr.AssumExpr({CAst.v=Name id},_) -> id::acc + | Vernacexpr.DefExpr ({CAst.v=Name id},_,_) -> id::acc + | _ -> acc in + let allnames = idstruc::(List.fold_left extract_name [] fs) in + let () = match List.duplicates Id.equal allnames with + | [] -> () + | id :: _ -> user_err (str "Two objects have the same name" ++ spc () ++ quote (Id.print id)) + in + let isnot_class = match kind with Class false -> false | _ -> true in + if isnot_class && List.exists (fun opt -> not (Option.is_empty opt)) priorities then + user_err Pp.(str "Priorities only allowed for type class substructures"); + (* Now, younger decl in params and fields is on top *) + let pl, univs, arity, template, implpars, params, implfs, fields = + States.with_state_protection (fun () -> + typecheck_params_and_fields finite (kind = Class true) idstruc poly pl s ps notations fs) () in + match kind with + | Class def -> + let priorities = List.map (fun id -> {hint_priority = id; hint_pattern = None}) priorities in + declare_class finite def cum pl univs (loc,idstruc) idbuild + implpars params arity template implfs fields is_coe coers priorities + | _ -> + let implfs = List.map + (fun impls -> implpars @ Impargs.lift_implicits + (succ (List.length params)) impls) implfs + in + let univs = + match univs with + | Polymorphic_const_entry univs -> + if cum then + Cumulative_ind_entry (Univ.CumulativityInfo.from_universe_context univs) + else + Polymorphic_ind_entry univs + | Monomorphic_const_entry univs -> + Monomorphic_ind_entry univs + in + let ind = declare_structure finite pl univs idstruc + idbuild implpars params arity template implfs + fields is_coe (List.map (fun coe -> not (Option.is_empty coe)) coers) in + IndRef ind diff --git a/vernac/record.mli b/vernac/record.mli new file mode 100644 index 00000000..992da2aa --- /dev/null +++ b/vernac/record.mli @@ -0,0 +1,35 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* + Entries.constant_universes_entry -> + ?kind:Decl_kinds.definition_object_kind -> + Id.t -> + bool list -> + Universes.universe_binders -> + Impargs.manual_implicits list -> + Context.Rel.t -> + (Name.t * bool) list * Constant.t option list + +val definition_structure : + inductive_kind * Decl_kinds.cumulative_inductive_flag * Decl_kinds.polymorphic * + Declarations.recursivity_kind * ident_decl with_coercion * local_binder_expr list * + (local_decl_expr with_instance with_priority with_notation) list * + Id.t * constr_expr option -> global_reference + +val declare_existing_class : global_reference -> unit diff --git a/vernac/search.ml b/vernac/search.ml new file mode 100644 index 00000000..a2a4fb40 --- /dev/null +++ b/vernac/search.ml @@ -0,0 +1,380 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* env -> constr -> bool +type display_function = global_reference -> env -> constr -> unit + +(* This option restricts the output of [SearchPattern ...], +[SearchAbout ...], etc. to the names of the symbols matching the +query, separated by a newline. This type of output is useful for +editors (like emacs), to generate a list of completion candidates +without having to parse thorugh the types of all symbols. *) + +type glob_search_about_item = + | GlobSearchSubPattern of constr_pattern + | GlobSearchString of string + +module SearchBlacklist = + Goptions.MakeStringTable + (struct + let key = ["Search";"Blacklist"] + let title = "Current search blacklist : " + let member_message s b = + str "Search blacklist does " ++ (if b then mt () else str "not ") ++ str "include " ++ str s + end) + +(* The functions iter_constructors and iter_declarations implement the behavior + needed for the Coq searching commands. + These functions take as first argument the procedure + that will be called to treat each entry. This procedure receives the name + of the object, the assumptions that will make it possible to print its type, + and the constr term that represent its type. *) + +let iter_constructors indsp u fn env nconstr = + for i = 1 to nconstr do + let typ = Inductiveops.type_of_constructor env ((indsp, i), u) in + fn (ConstructRef (indsp, i)) env typ + done + +let iter_named_context_name_type f = + List.iter (fun decl -> f (NamedDecl.get_id decl) (NamedDecl.get_type decl)) + +(* General search over hypothesis of a goal *) +let iter_hypothesis glnum (fn : global_reference -> env -> constr -> unit) = + let env = Global.env () in + let iter_hyp idh typ = fn (VarRef idh) env typ in + let evmap,e = Pfedit.get_goal_context glnum in + let pfctxt = named_context e in + iter_named_context_name_type iter_hyp pfctxt + +(* General search over declarations *) +let iter_declarations (fn : global_reference -> env -> constr -> unit) = + let env = Global.env () in + let iter_obj (sp, kn) lobj = match object_tag lobj with + | "VARIABLE" -> + begin try + let decl = Global.lookup_named (basename sp) in + fn (VarRef (NamedDecl.get_id decl)) env (NamedDecl.get_type decl) + with Not_found -> (* we are in a section *) () end + | "CONSTANT" -> + let cst = Global.constant_of_delta_kn kn in + let gr = ConstRef cst in + let (typ, _) = Global.type_of_global_in_context (Global.env ()) gr in + fn gr env typ + | "INDUCTIVE" -> + let mind = Global.mind_of_delta_kn kn in + let mib = Global.lookup_mind mind in + let iter_packet i mip = + let ind = (mind, i) in + let u = Univ.make_abstract_instance (Declareops.inductive_polymorphic_context mib) in + let i = (ind, u) in + let typ = Inductiveops.type_of_inductive env i in + let () = fn (IndRef ind) env typ in + let len = Array.length mip.mind_user_lc in + iter_constructors ind u fn env len + in + Array.iteri iter_packet mib.mind_packets + | _ -> () + in + try Declaremods.iter_all_segments iter_obj + with Not_found -> () + +let generic_search glnumopt fn = + (match glnumopt with + | None -> () + | Some glnum -> iter_hypothesis glnum fn); + iter_declarations fn + +(** This module defines a preference on constrs in the form of a + [compare] function (preferred constr must be big for this + functions, so preferences such as small constr must use a reversed + order). This priority will be used to order search results and + propose first results which are more likely to be relevant to the + query, this is why the type [t] contains the other elements + required of a search. *) +module ConstrPriority = struct + + (* The priority is memoised here. Because of the very localised use + of this module, it is not worth it making a convenient interface. *) + type t = + Globnames.global_reference * Environ.env * Constr.t * priority + and priority = int + + module ConstrSet = CSet.Make(Constr) + + (** A measure of the size of a term *) + let rec size t = + Constr.fold (fun s t -> 1 + s + size t) 0 t + + (** Set of the "symbols" (definitions, inductives, constructors) + which appear in a term. *) + let rec symbols acc t = + let open Constr in + match kind t with + | Const _ | Ind _ | Construct _ -> ConstrSet.add t acc + | _ -> Constr.fold symbols acc t + + (** The number of distinct "symbols" (see {!symbols}) which appear + in a term. *) + let num_symbols t = + ConstrSet.(cardinal (symbols empty t)) + + let priority t : priority = + -(3*(num_symbols t) + size t) + + let compare (_,_,_,p1) (_,_,_,p2) = + Pervasives.compare p1 p2 +end + +module PriorityQueue = Heap.Functional(ConstrPriority) + +let rec iter_priority_queue q fn = + (* use an option to make the function tail recursive. Will be + obsoleted with Ocaml 4.02 with the [match … with | exception …] + syntax. *) + let next = begin + try Some (PriorityQueue.maximum q) + with Heap.EmptyHeap -> None + end in + match next with + | Some (gref,env,t,_) -> + fn gref env t; + iter_priority_queue (PriorityQueue.remove q) fn + | None -> () + +let prioritize_search seq fn = + let acc = ref PriorityQueue.empty in + let iter gref env t = + let p = ConstrPriority.priority t in + acc := PriorityQueue.add (gref,env,t,p) !acc + in + let () = seq iter in + iter_priority_queue !acc fn + +(** Filters *) + +(** This function tries to see whether the conclusion matches a pattern. *) +(** FIXME: this is quite dummy, we may find a more efficient algorithm. *) +let rec pattern_filter pat ref env sigma typ = + let typ = Termops.strip_outer_cast sigma typ in + if Constr_matching.is_matching env sigma pat typ then true + else match EConstr.kind sigma typ with + | Prod (_, _, typ) + | LetIn (_, _, _, typ) -> pattern_filter pat ref env sigma typ + | _ -> false + +let rec head_filter pat ref env sigma typ = + let typ = Termops.strip_outer_cast sigma typ in + if Constr_matching.is_matching_head env sigma pat typ then true + else match EConstr.kind sigma typ with + | Prod (_, _, typ) + | LetIn (_, _, _, typ) -> head_filter pat ref env sigma typ + | _ -> false + +let full_name_of_reference ref = + let (dir,id) = repr_path (path_of_global ref) in + DirPath.to_string dir ^ "." ^ Id.to_string id + +(** Whether a reference is blacklisted *) +let blacklist_filter_aux () = + let l = SearchBlacklist.elements () in + fun ref env typ -> + let name = full_name_of_reference ref in + let is_not_bl str = not (String.string_contains ~where:name ~what:str) in + List.for_all is_not_bl l + +let module_filter (mods, outside) ref env typ = + let sp = path_of_global ref in + let sl = dirpath sp in + let is_outside md = not (is_dirpath_prefix_of md sl) in + let is_inside md = is_dirpath_prefix_of md sl in + if outside then List.for_all is_outside mods + else List.is_empty mods || List.exists is_inside mods + +let name_of_reference ref = Id.to_string (basename_of_global ref) + +let search_about_filter query gr env typ = match query with +| GlobSearchSubPattern pat -> + Constr_matching.is_matching_appsubterm ~closed:false env Evd.empty pat (EConstr.of_constr typ) +| GlobSearchString s -> + String.string_contains ~where:(name_of_reference gr) ~what:s + + +(** SearchPattern *) + +let search_pattern gopt pat mods pr_search = + let blacklist_filter = blacklist_filter_aux () in + let filter ref env typ = + module_filter mods ref env typ && + pattern_filter pat ref env Evd.empty (* FIXME *) (EConstr.of_constr typ) && + blacklist_filter ref env typ + in + let iter ref env typ = + if filter ref env typ then pr_search ref env typ + in + generic_search gopt iter + +(** SearchRewrite *) + +let eq = Coqlib.glob_eq + +let rewrite_pat1 pat = + PApp (PRef eq, [| PMeta None; pat; PMeta None |]) + +let rewrite_pat2 pat = + PApp (PRef eq, [| PMeta None; PMeta None; pat |]) + +let search_rewrite gopt pat mods pr_search = + let pat1 = rewrite_pat1 pat in + let pat2 = rewrite_pat2 pat in + let blacklist_filter = blacklist_filter_aux () in + let filter ref env typ = + module_filter mods ref env typ && + (pattern_filter pat1 ref env Evd.empty (* FIXME *) (EConstr.of_constr typ) || + pattern_filter pat2 ref env Evd.empty (EConstr.of_constr typ)) && + blacklist_filter ref env typ + in + let iter ref env typ = + if filter ref env typ then pr_search ref env typ + in + generic_search gopt iter + +(** Search *) + +let search_by_head gopt pat mods pr_search = + let blacklist_filter = blacklist_filter_aux () in + let filter ref env typ = + module_filter mods ref env typ && + head_filter pat ref env Evd.empty (* FIXME *) (EConstr.of_constr typ) && + blacklist_filter ref env typ + in + let iter ref env typ = + if filter ref env typ then pr_search ref env typ + in + generic_search gopt iter + +(** SearchAbout *) + +let search_about gopt items mods pr_search = + let blacklist_filter = blacklist_filter_aux () in + let filter ref env typ = + let eqb b1 b2 = if b1 then b2 else not b2 in + module_filter mods ref env typ && + List.for_all + (fun (b,i) -> eqb b (search_about_filter i ref env typ)) items && + blacklist_filter ref env typ + in + let iter ref env typ = + if filter ref env typ then pr_search ref env typ + in + generic_search gopt iter + +type search_constraint = + | Name_Pattern of Str.regexp + | Type_Pattern of Pattern.constr_pattern + | SubType_Pattern of Pattern.constr_pattern + | In_Module of Names.DirPath.t + | Include_Blacklist + +type 'a coq_object = { + coq_object_prefix : string list; + coq_object_qualid : string list; + coq_object_object : 'a; +} + +let interface_search = + let rec extract_flags name tpe subtpe mods blacklist = function + | [] -> (name, tpe, subtpe, mods, blacklist) + | (Name_Pattern regexp, b) :: l -> + extract_flags ((regexp, b) :: name) tpe subtpe mods blacklist l + | (Type_Pattern pat, b) :: l -> + extract_flags name ((pat, b) :: tpe) subtpe mods blacklist l + | (SubType_Pattern pat, b) :: l -> + extract_flags name tpe ((pat, b) :: subtpe) mods blacklist l + | (In_Module id, b) :: l -> + extract_flags name tpe subtpe ((id, b) :: mods) blacklist l + | (Include_Blacklist, b) :: l -> + extract_flags name tpe subtpe mods b l + in + fun ?glnum flags -> + let (name, tpe, subtpe, mods, blacklist) = + extract_flags [] [] [] [] false flags + in + let blacklist_filter = blacklist_filter_aux () in + let filter_function ref env constr = + let id = Names.Id.to_string (Nametab.basename_of_global ref) in + let path = Libnames.dirpath (Nametab.path_of_global ref) in + let toggle x b = if x then b else not b in + let match_name (regexp, flag) = + toggle (Str.string_match regexp id 0) flag + in + let match_type (pat, flag) = + toggle (Constr_matching.is_matching env Evd.empty pat (EConstr.of_constr constr)) flag + in + let match_subtype (pat, flag) = + toggle + (Constr_matching.is_matching_appsubterm ~closed:false + env Evd.empty pat (EConstr.of_constr constr)) flag + in + let match_module (mdl, flag) = + toggle (Libnames.is_dirpath_prefix_of mdl path) flag + in + List.for_all match_name name && + List.for_all match_type tpe && + List.for_all match_subtype subtpe && + List.for_all match_module mods && + (blacklist || blacklist_filter ref env constr) + in + let ans = ref [] in + let print_function ref env constr = + let fullpath = DirPath.repr (Nametab.dirpath_of_global ref) in + let qualid = Nametab.shortest_qualid_of_global Id.Set.empty ref in + let (shortpath, basename) = Libnames.repr_qualid qualid in + let shortpath = DirPath.repr shortpath in + (* [shortpath] is a suffix of [fullpath] and we're looking for the missing + prefix *) + let rec prefix full short accu = match full, short with + | _, [] -> + let full = List.rev_map Id.to_string full in + (full, accu) + | _ :: full, m :: short -> + prefix full short (Id.to_string m :: accu) + | _ -> assert false + in + let (prefix, qualid) = prefix fullpath shortpath [Id.to_string basename] in + let answer = { + coq_object_prefix = prefix; + coq_object_qualid = qualid; + coq_object_object = constr; + } in + ans := answer :: !ans; + in + let iter ref env typ = + if filter_function ref env typ then print_function ref env typ + in + let () = generic_search glnum iter in + !ans + +let blacklist_filter ref env typ = + blacklist_filter_aux () ref env typ diff --git a/vernac/search.mli b/vernac/search.mli new file mode 100644 index 00000000..a1fb7ed3 --- /dev/null +++ b/vernac/search.mli @@ -0,0 +1,85 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* env -> constr -> bool +type display_function = global_reference -> env -> constr -> unit + +(** {6 Generic filter functions} *) + +val blacklist_filter : filter_function +(** Check whether a reference is blacklisted. *) + +val module_filter : DirPath.t list * bool -> filter_function +(** Check whether a reference pertains or not to a set of modules *) + +val search_about_filter : glob_search_about_item -> filter_function +(** Check whether a reference matches a SearchAbout query. *) + +(** {6 Specialized search functions} + +[search_xxx gl pattern modinout] searches the hypothesis of the [gl]th +goal and the global environment for things matching [pattern] and +satisfying module exclude/include clauses of [modinout]. *) + +val search_by_head : int option -> constr_pattern -> DirPath.t list * bool + -> display_function -> unit +val search_rewrite : int option -> constr_pattern -> DirPath.t list * bool + -> display_function -> unit +val search_pattern : int option -> constr_pattern -> DirPath.t list * bool + -> display_function -> unit +val search_about : int option -> (bool * glob_search_about_item) list + -> DirPath.t list * bool -> display_function -> unit + +type search_constraint = + (** Whether the name satisfies a regexp (uses Ocaml Str syntax) *) + | Name_Pattern of Str.regexp + (** Whether the object type satisfies a pattern *) + | Type_Pattern of Pattern.constr_pattern + (** Whether some subtype of object type satisfies a pattern *) + | SubType_Pattern of Pattern.constr_pattern + (** Whether the object pertains to a module *) + | In_Module of Names.DirPath.t + (** Bypass the Search blacklist *) + | Include_Blacklist + +type 'a coq_object = { + coq_object_prefix : string list; + coq_object_qualid : string list; + coq_object_object : 'a; +} + +val interface_search : ?glnum:int -> (search_constraint * bool) list -> + constr coq_object list + +(** {6 Generic search function} *) + +val generic_search : int option -> display_function -> unit +(** This function iterates over all hypothesis of the goal numbered + [glnum] (if present) and all known declarations. *) + +(** {6 Search function modifiers} *) + +val prioritize_search : (display_function -> unit) -> display_function -> unit +(** [prioritize_search iter] iterates over the values of [iter] (seen + as a sequence of declarations), in a relevance order. This requires to + perform the entire iteration of [iter] before starting streaming. So + [prioritize_search] should not be used for low-latency streaming. *) diff --git a/vernac/topfmt.ml b/vernac/topfmt.ml new file mode 100644 index 00000000..1d720330 --- /dev/null +++ b/vernac/topfmt.ml @@ -0,0 +1,334 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* pp_global_params -> unit + * set the parameters of a formatter *) + +let set_gp ft gp = + Format.pp_set_margin ft gp.margin ; + Format.pp_set_max_indent ft gp.max_indent ; + Format.pp_set_max_boxes ft gp.max_depth ; + Format.pp_set_ellipsis_text ft gp.ellipsis + +let set_dflt_gp ft = set_gp ft dflt_gp + +let get_gp ft = + { margin = Format.pp_get_margin ft (); + max_indent = Format.pp_get_max_indent ft (); + max_depth = Format.pp_get_max_boxes ft (); + ellipsis = Format.pp_get_ellipsis_text ft () } + +(* with_fp : 'a pp_formatter_params -> Format.formatter + * returns of formatter for given formatter functions *) + +let with_fp chan out_function flush_function = + let ft = Format.make_formatter out_function flush_function in + Format.pp_set_formatter_out_channel ft chan; + ft + +(* Output on a channel ch *) + +let with_output_to ch = + let ft = with_fp ch (output_substring ch) (fun () -> flush ch) in + set_gp ft deep_gp; + ft + +let std_ft = ref Format.std_formatter +let _ = set_dflt_gp !std_ft + +let err_ft = ref Format.err_formatter +let _ = set_gp !err_ft deep_gp + +let deep_ft = ref (with_output_to stdout) +let _ = set_gp !deep_ft deep_gp + +(* For parametrization through vernacular *) +let default = Format.pp_get_max_boxes !std_ft () +let default_margin = Format.pp_get_margin !std_ft () + +let get_depth_boxes () = Some (Format.pp_get_max_boxes !std_ft ()) +let set_depth_boxes v = + Format.pp_set_max_boxes !std_ft (match v with None -> default | Some v -> v) + +let get_margin () = Some (Format.pp_get_margin !std_ft ()) +let set_margin v = + let v = match v with None -> default_margin | Some v -> v in + Format.pp_set_margin Format.str_formatter v; + Format.pp_set_margin !std_ft v; + Format.pp_set_margin !deep_ft v; + Format.pp_set_margin !err_ft v; + (* Heuristic, based on usage: the column on the right of max_indent + column is 20% of width, capped to 30 characters *) + let m = max (64 * v / 100) (v-30) in + Format.pp_set_max_indent Format.str_formatter m; + Format.pp_set_max_indent !std_ft m; + Format.pp_set_max_indent !deep_ft m; + Format.pp_set_max_indent !err_ft m + +(** Console display of feedback *) + +(** Default tags *) +module Tag = struct + + let error = "message.error" + let warning = "message.warning" + let debug = "message.debug" + +end + +let msgnl_with ?pre_hdr fmt strm = + pp_with fmt (strm ++ fnl ()); + Format.pp_print_flush fmt () + +module Emacs = struct + + (* Special chars for emacs, to detect warnings inside goal output *) + let quote_warning_start = "" + let quote_warning_end = "" + + let quote_info_start = "" + let quote_info_end = "" + + let quote_emacs q_start q_end msg = + hov 0 (seq [str q_start; brk(0,0); msg; brk(0,0); str q_end]) + + let quote_warning = quote_emacs quote_warning_start quote_warning_end + let quote_info = quote_emacs quote_info_start quote_info_end + +end + +let dbg_hdr = tag Tag.debug (str "Debug:") ++ spc () +let info_hdr = mt () +let warn_hdr = tag Tag.warning (str "Warning:") ++ spc () +let err_hdr = tag Tag.error (str "Error:") ++ spc () + +let make_body quoter info ?pre_hdr s = + pr_opt_no_spc (fun x -> x ++ fnl ()) pre_hdr ++ quoter (hov 0 (info ++ s)) + +(* The empty quoter *) +let noq x = x +(* Generic logger *) +let gen_logger dbg warn ?pre_hdr level msg = let open Feedback in match level with + | Debug -> msgnl_with !std_ft (make_body dbg dbg_hdr ?pre_hdr msg) + | Info -> msgnl_with !std_ft (make_body dbg info_hdr ?pre_hdr msg) + | Notice -> msgnl_with !std_ft (make_body noq info_hdr ?pre_hdr msg) + | Warning -> Flags.if_warn (fun () -> + msgnl_with !err_ft (make_body warn warn_hdr ?pre_hdr msg)) () + | Error -> msgnl_with !err_ft (make_body noq err_hdr ?pre_hdr msg) + +(** Standard loggers *) + +(* We provide a generic clear_log_backend callback for backends + wanting to do cleanup after the print. +*) +let std_logger_cleanup = ref (fun () -> ()) + +let std_logger ?pre_hdr level msg = + gen_logger (fun x -> x) (fun x -> x) ?pre_hdr level msg; + !std_logger_cleanup () + +(** Color logging. Moved from Ppstyle, it may need some more refactoring *) + +(* Tag map for terminal style *) +let default_tag_map () = let open Terminal in [ + (* Local to console toplevel *) + "message.error" , make ~bold:true ~fg_color:`WHITE ~bg_color:`RED () + ; "message.warning" , make ~bold:true ~fg_color:`WHITE ~bg_color:`YELLOW () + ; "message.debug" , make ~bold:true ~fg_color:`WHITE ~bg_color:`MAGENTA () + (* Coming from the printer *) + ; "constr.evar" , make ~fg_color:`LIGHT_BLUE () + ; "constr.keyword" , make ~bold:true () + ; "constr.type" , make ~bold:true ~fg_color:`YELLOW () + ; "constr.notation" , make ~fg_color:`WHITE () + (* ["constr"; "variable"] is not assigned *) + ; "constr.reference" , make ~fg_color:`LIGHT_GREEN () + ; "constr.path" , make ~fg_color:`LIGHT_MAGENTA () + ; "module.definition", make ~bold:true ~fg_color:`LIGHT_RED () + ; "module.keyword" , make ~bold:true () + ; "tactic.keyword" , make ~bold:true () + ; "tactic.primitive" , make ~fg_color:`LIGHT_GREEN () + ; "tactic.string" , make ~fg_color:`LIGHT_RED () + ] + +let tag_map = ref CString.Map.empty + +let init_tag_map styles = + let set accu (name, st) = CString.Map.add name st accu in + tag_map := List.fold_left set !tag_map styles + +let default_styles () = + init_tag_map (default_tag_map ()) + +let parse_color_config file = + let styles = Terminal.parse file in + init_tag_map styles + +let dump_tags () = CString.Map.bindings !tag_map + +(** Not thread-safe. We should put a lock somewhere if we print from + different threads. Do we? *) +let make_style_stack () = + (** Default tag is to reset everything *) + let empty = Terminal.make () in + let default_tag = Terminal.({ + fg_color = Some `DEFAULT; + bg_color = Some `DEFAULT; + bold = Some false; + italic = Some false; + underline = Some false; + negative = Some false; + prefix = None; + suffix = None; + }) + in + let style_stack = ref [] in + let peek () = match !style_stack with + | [] -> default_tag (** Anomalous case, but for robustness *) + | st :: _ -> st + in + let push tag = + let style = + try CString.Map.find tag !tag_map + with | Not_found -> empty + in + (** Use the merging of the latest tag and the one being currently pushed. + This may be useful if for instance the latest tag changes the background and + the current one the foreground, so that the two effects are additioned. *) + let style = Terminal.merge (peek ()) style in + style_stack := style :: !style_stack; + Terminal.eval style + in + let pop _ = match !style_stack with + | [] -> (** Something went wrong, we fallback *) + Terminal.eval default_tag + | _ :: rem -> style_stack := rem; + Terminal.eval (peek ()) + in + let clear () = style_stack := [] in + push, pop, clear + +let make_printing_functions () = + let empty = Terminal.make () in + let print_prefix ft tag = + let style = + try CString.Map.find tag !tag_map + with | Not_found -> empty + in + match style.Terminal.prefix with Some s -> Format.pp_print_string ft s | None -> () + in + let print_suffix ft tag = + let style = + try CString.Map.find tag !tag_map + with | Not_found -> empty + in + match style.Terminal.suffix with Some s -> Format.pp_print_string ft s | None -> () + in + print_prefix, print_suffix + +let init_terminal_output ~color = + let push_tag, pop_tag, clear_tag = make_style_stack () in + let print_prefix, print_suffix = make_printing_functions () in + let tag_handler ft = { + Format.mark_open_tag = push_tag; + Format.mark_close_tag = pop_tag; + Format.print_open_tag = print_prefix ft; + Format.print_close_tag = print_suffix ft; + } in + if color then + (* Use 0-length markers *) + begin + std_logger_cleanup := clear_tag; + Format.pp_set_mark_tags !std_ft true; + Format.pp_set_mark_tags !err_ft true + end + else + (* Use textual markers *) + begin + Format.pp_set_print_tags !std_ft true; + Format.pp_set_print_tags !err_ft true + end; + Format.pp_set_formatter_tag_functions !std_ft (tag_handler !std_ft); + Format.pp_set_formatter_tag_functions !err_ft (tag_handler !err_ft) + +(* Rules for emacs: + - Debug/info: emacs_quote_info + - Warning/Error: emacs_quote_err + - Notice: unquoted + *) +let emacs_logger = gen_logger Emacs.quote_info Emacs.quote_warning + +(* This is specific to the toplevel *) +let pr_loc loc = + let fname = loc.Loc.fname in + match fname with + | Loc.ToplevelInput -> + Loc.(str"Toplevel input, characters " ++ int loc.bp ++ + str"-" ++ int loc.ep ++ str":") + | Loc.InFile fname -> + Loc.(str"File " ++ str "\"" ++ str fname ++ str "\"" ++ + str", line " ++ int loc.line_nb ++ str", characters " ++ + int (loc.bp-loc.bol_pos) ++ str"-" ++ int (loc.ep-loc.bol_pos) ++ + str":") + +let print_err_exn ?extra any = + let (e, info) = CErrors.push any in + let loc = Loc.get_loc info in + let msg_loc = Option.cata pr_loc (mt ()) loc in + let pre_hdr = pr_opt_no_spc (fun x -> x) extra ++ msg_loc in + let msg = CErrors.iprint (e, info) ++ fnl () in + std_logger ~pre_hdr Feedback.Error msg + +let with_output_to_file fname func input = + let channel = open_out (String.concat "." [fname; "out"]) in + let old_fmt = !std_ft, !err_ft, !deep_ft in + let new_ft = Format.formatter_of_out_channel channel in + std_ft := new_ft; + err_ft := new_ft; + deep_ft := new_ft; + try + let output = func input in + std_ft := Util.pi1 old_fmt; + err_ft := Util.pi2 old_fmt; + deep_ft := Util.pi3 old_fmt; + close_out channel; + output + with reraise -> + let reraise = Backtrace.add_backtrace reraise in + std_ft := Util.pi1 old_fmt; + err_ft := Util.pi2 old_fmt; + deep_ft := Util.pi3 old_fmt; + close_out channel; + Exninfo.iraise reraise diff --git a/vernac/topfmt.mli b/vernac/topfmt.mli new file mode 100644 index 00000000..579b456a --- /dev/null +++ b/vernac/topfmt.mli @@ -0,0 +1,62 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* pp_global_params -> unit +val set_dflt_gp : Format.formatter -> unit +val get_gp : Format.formatter -> pp_global_params + +(** {6 Output functions of pretty-printing. } *) + +val with_output_to : out_channel -> Format.formatter + +val std_ft : Format.formatter ref +val err_ft : Format.formatter ref +val deep_ft : Format.formatter ref + +(** {6 For parametrization through vernacular. } *) + +val set_depth_boxes : int option -> unit +val get_depth_boxes : unit -> int option + +val set_margin : int option -> unit +val get_margin : unit -> int option + +(** Console display of feedback, we may add some location information *) +val std_logger : ?pre_hdr:Pp.t -> Feedback.level -> Pp.t -> unit +val emacs_logger : ?pre_hdr:Pp.t -> Feedback.level -> Pp.t -> unit + +(** Color output *) +val default_styles : unit -> unit +val parse_color_config : string -> unit +val dump_tags : unit -> (string * Terminal.style) list + +(** Initialization of interpretation of tags *) +val init_terminal_output : color:bool -> unit + +(** Error printing *) +(* To be deprecated when we can fully move to feedback-based error + printing. *) +val pr_loc : Loc.t -> Pp.t +val print_err_exn : ?extra:Pp.t -> exn -> unit + +(** [with_output_to_file file f x] executes [f x] with logging + redirected to a file [file] *) +val with_output_to_file : string -> ('a -> 'b) -> 'a -> 'b + diff --git a/vernac/vernac.mllib b/vernac/vernac.mllib new file mode 100644 index 00000000..f001b572 --- /dev/null +++ b/vernac/vernac.mllib @@ -0,0 +1,26 @@ +Vernacprop +Proof_using +Lemmas +Himsg +ExplainErr +Class +Locality +Metasyntax +Auto_ind_decl +Search +Indschemes +DeclareDef +Obligations +ComDefinition +ComAssumption +ComInductive +ComFixpoint +ComProgramFixpoint +Classes +Record +Assumptions +Vernacstate +Vernacinterp +Mltop +Topfmt +Vernacentries diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml new file mode 100644 index 00000000..fceb0bf3 --- /dev/null +++ b/vernac/vernacentries.ml @@ -0,0 +1,2349 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* Classops.CL_FUN + | SortClass -> Classops.CL_SORT + | RefClass r -> Class.class_of_global (Smartlocate.smart_global ~head:true r) + +let scope_class_of_qualid qid = + Notation.scope_class_of_class (cl_of_qualid qid) + +(*******************) +(* "Show" commands *) + +let show_proof () = + (* spiwack: this would probably be cooler with a bit of polishing. *) + let p = Proof_global.give_me_the_proof () in + let sigma, env = Pfedit.get_current_context () in + let pprf = Proof.partial_proof p in + Pp.prlist_with_sep Pp.fnl (Printer.pr_econstr_env env sigma) pprf + +let show_top_evars () = + (* spiwack: new as of Feb. 2010: shows goal evars in addition to non-goal evars. *) + let pfts = Proof_global.give_me_the_proof () in + let gls,_,shelf,givenup,sigma = Proof.proof pfts in + pr_evars_int sigma ~shelf ~givenup 1 (Evd.undefined_map sigma) + +let show_universes () = + let pfts = Proof_global.give_me_the_proof () in + let gls,_,_,_,sigma = Proof.proof pfts in + let ctx = Evd.universe_context_set (Evd.minimize_universes sigma) in + Termops.pr_evar_universe_context (Evd.evar_universe_context sigma) ++ fnl () ++ + str "Normalized constraints: " ++ Univ.pr_universe_context_set (Termops.pr_evd_level sigma) ctx + +(* Simulate the Intro(s) tactic *) +let show_intro all = + let open EConstr in + let pf = Proof_global.give_me_the_proof() in + let gls,_,_,_,sigma = Proof.proof pf in + if not (List.is_empty gls) then begin + let gl = {Evd.it=List.hd gls ; sigma = sigma; } in + let l,_= decompose_prod_assum sigma (Termops.strip_outer_cast sigma (pf_concl gl)) in + if all then + let lid = Tactics.find_intro_names l gl in + hov 0 (prlist_with_sep spc Id.print lid) + else if not (List.is_empty l) then + let n = List.last l in + Id.print (List.hd (Tactics.find_intro_names [n] gl)) + else mt () + end else mt () + +(** Prepare a "match" template for a given inductive type. + For each branch of the match, we list the constructor name + followed by enough pattern variables. + [Not_found] is raised if the given string isn't the qualid of + a known inductive type. *) + +(* + + HH notes in PR #679: + + The Show Match could also be made more robust, for instance in the + presence of let in the branch of a constructor. A + decompose_prod_assum would probably suffice for that, but then, it + is a Context.Rel.Declaration.t which needs to be matched and not + just a pair (name,type). + + Otherwise, this is OK. After all, the API on inductive types is not + so canonical in general, and in this simple case, working at the + low-level of mind_nf_lc seems reasonable (compared to working at the + higher-level of Inductiveops). + +*) + +let make_cases_aux glob_ref = + match glob_ref with + | Globnames.IndRef ind -> + let {Declarations.mind_nparams = np} , {Declarations.mind_nf_lc = tarr} = Global.lookup_inductive ind in + Util.Array.fold_right_i + (fun i typ l -> + let al = List.rev (fst (decompose_prod typ)) in + let al = Util.List.skipn np al in + let rec rename avoid = function + | [] -> [] + | (n,_)::l -> + let n' = Namegen.next_name_away_with_default (Id.to_string Namegen.default_dependent_ident) n avoid in + Id.to_string n' :: rename (Id.Set.add n' avoid) l in + let al' = rename Id.Set.empty al in + let consref = ConstructRef (ith_constructor_of_inductive ind (i + 1)) in + (Libnames.string_of_qualid (Nametab.shortest_qualid_of_global Id.Set.empty consref) :: al') :: l) + tarr [] + | _ -> raise Not_found + +let make_cases s = + let qualified_name = Libnames.qualid_of_string s in + let glob_ref = Nametab.locate qualified_name in + make_cases_aux glob_ref + +(** Textual display of a generic "match" template *) + +let show_match id = + let patterns = + try make_cases_aux (Nametab.global id) + with Not_found -> user_err Pp.(str "Unknown inductive type.") + in + let pr_branch l = + str "| " ++ hov 1 (prlist_with_sep spc str l) ++ str " =>" + in + v 1 (str "match # with" ++ fnl () ++ + prlist_with_sep fnl pr_branch patterns ++ fnl () ++ str "end" ++ fnl ()) + +(* "Print" commands *) + +let print_path_entry p = + let dir = DirPath.print (Loadpath.logical p) in + let path = str (CUnix.escaped_string_of_physical_path (Loadpath.physical p)) in + Pp.hov 2 (dir ++ spc () ++ path) + +let print_loadpath dir = + let l = Loadpath.get_load_paths () in + let l = match dir with + | None -> l + | Some dir -> + let filter p = is_dirpath_prefix_of dir (Loadpath.logical p) in + List.filter filter l + in + str "Logical Path / Physical path:" ++ fnl () ++ + prlist_with_sep fnl print_path_entry l + +let print_modules () = + let opened = Library.opened_libraries () + and loaded = Library.loaded_libraries () in + (* we intersect over opened to preserve the order of opened since *) + (* non-commutative operations (e.g. visibility) are done at import time *) + let loaded_opened = List.intersect DirPath.equal opened loaded + and only_loaded = List.subtract DirPath.equal loaded opened in + str"Loaded and imported library files: " ++ + pr_vertical_list DirPath.print loaded_opened ++ fnl () ++ + str"Loaded and not imported library files: " ++ + pr_vertical_list DirPath.print only_loaded + + +let print_module r = + let qid = qualid_of_reference r in + try + let globdir = Nametab.locate_dir qid.v in + match globdir with + DirModule { obj_dir; obj_mp; _ } -> + Printmod.print_module (Printmod.printable_body obj_dir) obj_mp + | _ -> raise Not_found + with + Not_found -> user_err (str"Unknown Module " ++ pr_qualid qid.v) + +let print_modtype r = + let qid = qualid_of_reference r in + try + let kn = Nametab.locate_modtype qid.v in + Printmod.print_modtype kn + with Not_found -> + (* Is there a module of this name ? If yes we display its type *) + try + let mp = Nametab.locate_module qid.v in + Printmod.print_module false mp + with Not_found -> + user_err (str"Unknown Module Type or Module " ++ pr_qualid qid.v) + +let print_namespace ns = + let ns = List.rev (Names.DirPath.repr ns) in + (* [match_dirpath], [match_modulpath] are helpers for [matches] + which checks whether a constant is in the namespace [ns]. *) + let rec match_dirpath ns = function + | [] -> Some ns + | id::dir -> + begin match match_dirpath ns dir with + | Some [] as y -> y + | Some (a::ns') -> + if Names.Id.equal a id then Some ns' + else None + | None -> None + end + in + let rec match_modulepath ns = function + | MPbound _ -> None (* Not a proper namespace. *) + | MPfile dir -> match_dirpath ns (Names.DirPath.repr dir) + | MPdot (mp,lbl) -> + let id = Names.Label.to_id lbl in + begin match match_modulepath ns mp with + | Some [] as y -> y + | Some (a::ns') -> + if Names.Id.equal a id then Some ns' + else None + | None -> None + end + in + (* [qualified_minus n mp] returns a list of qualifiers representing + [mp] except the [n] first (in the concrete syntax order). The + idea is that if [mp] matches [ns], then [qualified_minus mp + (length ns)] will be the correct representation of [mp] assuming + [ns] is imported. *) + (* precondition: [mp] matches some namespace of length [n] *) + let qualified_minus n mp = + let rec list_of_modulepath = function + | MPbound _ -> assert false (* MPbound never matches *) + | MPfile dir -> Names.DirPath.repr dir + | MPdot (mp,lbl) -> (Names.Label.to_id lbl)::(list_of_modulepath mp) + in + snd (Util.List.chop n (List.rev (list_of_modulepath mp))) + in + let print_list pr l = prlist_with_sep (fun () -> str".") pr l in + let print_kn kn = + (* spiwack: I'm ignoring the dirpath, is that bad? *) + let (mp,_,lbl) = Names.KerName.repr kn in + let qn = (qualified_minus (List.length ns) mp)@[Names.Label.to_id lbl] in + print_list Id.print qn + in + let print_constant k body = + (* FIXME: universes *) + let t = body.Declarations.const_type in + let sigma, env = Pfedit.get_current_context () in + print_kn k ++ str":" ++ spc() ++ Printer.pr_type_env env sigma t + in + let matches mp = match match_modulepath ns mp with + | Some [] -> true + | _ -> false in + let constants = (Environ.pre_env (Global.env ())).Pre_env.env_globals.Pre_env.env_constants in + let constants_in_namespace = + Cmap_env.fold (fun c (body,_) acc -> + let kn = Constant.user c in + if matches (KerName.modpath kn) then + acc++fnl()++hov 2 (print_constant kn body) + else + acc + ) constants (str"") + in + (print_list Id.print ns)++str":"++fnl()++constants_in_namespace + +let print_strategy r = + let open Conv_oracle in + let pr_level = function + | Expand -> str "expand" + | Level 0 -> str "transparent" + | Level n -> str "level" ++ spc() ++ int n + | Opaque -> str "opaque" + in + let pr_strategy (ref, lvl) = pr_global ref ++ str " : " ++ pr_level lvl in + let oracle = Environ.oracle (Global.env ()) in + match r with + | None -> + let fold key lvl (vacc, cacc) = match key with + | VarKey id -> ((VarRef id, lvl) :: vacc, cacc) + | ConstKey cst -> (vacc, (ConstRef cst, lvl) :: cacc) + | RelKey _ -> (vacc, cacc) + in + let var_lvl, cst_lvl = fold_strategy fold oracle ([], []) in + let var_msg = + if List.is_empty var_lvl then mt () + else str "Variable strategies" ++ fnl () ++ + hov 0 (prlist_with_sep fnl pr_strategy var_lvl) ++ fnl () + in + let cst_msg = + if List.is_empty cst_lvl then mt () + else str "Constant strategies" ++ fnl () ++ + hov 0 (prlist_with_sep fnl pr_strategy cst_lvl) + in + var_msg ++ cst_msg + | Some r -> + let r = Smartlocate.smart_global r in + let key = match r with + | VarRef id -> VarKey id + | ConstRef cst -> ConstKey cst + | IndRef _ | ConstructRef _ -> user_err Pp.(str "The reference is not unfoldable") + in + let lvl = get_strategy oracle key in + pr_strategy (r, lvl) + +let dump_universes_gen g s = + let output = open_out s in + let output_constraint, close = + if Filename.check_suffix s ".dot" || Filename.check_suffix s ".gv" then begin + (* the lazy unit is to handle errors while printing the first line *) + let init = lazy (Printf.fprintf output "digraph universes {\n") in + begin fun kind left right -> + let () = Lazy.force init in + match kind with + | Univ.Lt -> + Printf.fprintf output " \"%s\" -> \"%s\" [style=bold];\n" right left + | Univ.Le -> + Printf.fprintf output " \"%s\" -> \"%s\" [style=solid];\n" right left + | Univ.Eq -> + Printf.fprintf output " \"%s\" -> \"%s\" [style=dashed];\n" left right + end, begin fun () -> + if Lazy.is_val init then Printf.fprintf output "}\n"; + close_out output + end + end else begin + begin fun kind left right -> + let kind = match kind with + | Univ.Lt -> "<" + | Univ.Le -> "<=" + | Univ.Eq -> "=" + in Printf.fprintf output "%s %s %s ;\n" left kind right + end, (fun () -> close_out output) + end + in + try + UGraph.dump_universes output_constraint g; + close (); + str "Universes written to file \"" ++ str s ++ str "\"." + with reraise -> + let reraise = CErrors.push reraise in + close (); + iraise reraise + +(*********************) +(* "Locate" commands *) + +let locate_file f = + let file = Flags.silently Loadpath.locate_file f in + str file + +let msg_found_library = function + | Library.LibLoaded, fulldir, file -> + hov 0 (DirPath.print fulldir ++ strbrk " has been loaded from file " ++ str file) + | Library.LibInPath, fulldir, file -> + hov 0 (DirPath.print fulldir ++ strbrk " is bound to file " ++ str file) + +let err_unmapped_library ?loc ?from qid = + let dir = fst (repr_qualid qid) in + let prefix = match from with + | None -> str "." + | Some from -> + str " and prefix " ++ DirPath.print from ++ str "." + in + user_err ?loc + ~hdr:"locate_library" + (strbrk "Cannot find a physical path bound to logical path matching suffix " ++ + DirPath.print dir ++ prefix) + +let err_notfound_library ?loc ?from qid = + let prefix = match from with + | None -> str "." + | Some from -> + str " with prefix " ++ DirPath.print from ++ str "." + in + user_err ?loc ~hdr:"locate_library" + (strbrk "Unable to locate library " ++ pr_qualid qid ++ prefix) + +let print_located_library r = + let {loc;v=qid} = qualid_of_reference r in + try msg_found_library (Library.locate_qualified_library ~warn:false qid) + with + | Library.LibUnmappedDir -> err_unmapped_library ?loc qid + | Library.LibNotFound -> err_notfound_library ?loc qid + +let smart_global r = + let gr = Smartlocate.smart_global r in + Dumpglob.add_glob ?loc:r.loc gr; + gr + +let dump_global r = + try + let gr = Smartlocate.smart_global r in + Dumpglob.add_glob ?loc:r.loc gr + with e when CErrors.noncritical e -> () +(**********) +(* Syntax *) + +let vernac_syntax_extension atts infix l = + let local = enforce_module_locality atts.locality in + if infix then Metasyntax.check_infix_modifiers (snd l); + Metasyntax.add_syntax_extension local l + +let vernac_delimiters sc = function + | Some lr -> Metasyntax.add_delimiters sc lr + | None -> Metasyntax.remove_delimiters sc + +let vernac_bind_scope sc cll = + Metasyntax.add_class_scope sc (List.map scope_class_of_qualid cll) + +let vernac_open_close_scope ~atts (b,s) = + let local = enforce_section_locality atts.locality in + Notation.open_close_scope (local,b,s) + +let vernac_arguments_scope ~atts r scl = + let local = make_section_locality atts.locality in + Notation.declare_arguments_scope local (smart_global r) scl + +let vernac_infix ~atts = + let local = enforce_module_locality atts.locality in + Metasyntax.add_infix local (Global.env()) + +let vernac_notation ~atts = + let local = enforce_module_locality atts.locality in + Metasyntax.add_notation local (Global.env()) + +(***********) +(* Gallina *) + +let start_proof_and_print k l hook = + let inference_hook = + if Flags.is_program_mode () then + let hook env sigma ev = + let tac = !Obligations.default_tactic in + let evi = Evd.find sigma ev in + let evi = Evarutil.nf_evar_info sigma evi in + let env = Evd.evar_filtered_env evi in + try + let concl = EConstr.of_constr evi.Evd.evar_concl in + if not (Evarutil.is_ground_env sigma env && + Evarutil.is_ground_term sigma concl) + then raise Exit; + let c, _, ctx = + Pfedit.build_by_tactic env (Evd.evar_universe_context sigma) + concl (Tacticals.New.tclCOMPLETE tac) + in Evd.set_universe_context sigma ctx, EConstr.of_constr c + with Logic_monad.TacticFailure e when Logic.catchable_exception e -> + user_err Pp.(str "The statement obligations could not be resolved \ + automatically, write a statement definition first.") + in Some hook + else None + in + start_proof_com ?inference_hook k l hook + +let no_hook = Lemmas.mk_hook (fun _ _ -> ()) + +let vernac_definition_hook p = function +| Coercion -> Class.add_coercion_hook p +| CanonicalStructure -> + Lemmas.mk_hook (fun _ -> Recordops.declare_canonical_structure) +| SubClass -> Class.add_subclass_hook p +| _ -> no_hook + +let vernac_definition ~atts discharge kind ({loc;v=id}, pl) def = + let local = enforce_locality_exp atts.locality discharge in + let hook = vernac_definition_hook atts.polymorphic kind in + let () = + match id with + | Anonymous -> () + | Name n -> let lid = CAst.make ?loc n in + match local with + | Discharge -> Dumpglob.dump_definition lid true "var" + | Local | Global -> Dumpglob.dump_definition lid false "def" + in + let program_mode = Flags.is_program_mode () in + let name = + match id with + | Anonymous -> fresh_name_for_anonymous_theorem () + | Name n -> n + in + (match def with + | ProveBody (bl,t) -> (* local binders, typ *) + start_proof_and_print (local, atts.polymorphic, DefinitionBody kind) + [(CAst.make ?loc name, pl), (bl, t)] hook + | DefineBody (bl,red_option,c,typ_opt) -> + let red_option = match red_option with + | None -> None + | Some r -> + let sigma, env = Pfedit.get_current_context () in + Some (snd (Hook.get f_interp_redexp env sigma r)) in + ComDefinition.do_definition ~program_mode name + (local, atts.polymorphic, kind) pl bl red_option c typ_opt hook) + +let vernac_start_proof ~atts kind l = + let local = enforce_locality_exp atts.locality NoDischarge in + if Dumpglob.dump () then + List.iter (fun ((id, _), _) -> Dumpglob.dump_definition id false "prf") l; + start_proof_and_print (local, atts.polymorphic, Proof kind) l no_hook + +let vernac_end_proof ?proof = function + | Admitted -> save_proof ?proof Admitted + | Proved (_,_) as e -> save_proof ?proof e + +let vernac_exact_proof c = + (* spiwack: for simplicity I do not enforce that "Proof proof_term" is + called only at the begining of a proof. *) + let status = Pfedit.by (Tactics.exact_proof c) in + save_proof (Vernacexpr.(Proved(Opaque,None))); + if not status then Feedback.feedback Feedback.AddedAxiom + +let vernac_assumption ~atts discharge kind l nl = + let local = enforce_locality_exp atts.locality discharge in + let global = local == Global in + let kind = local, atts.polymorphic, kind in + List.iter (fun (is_coe,(idl,c)) -> + if Dumpglob.dump () then + List.iter (fun (lid, _) -> + if global then Dumpglob.dump_definition lid false "ax" + else Dumpglob.dump_definition lid true "var") idl) l; + let status = ComAssumption.do_assumptions kind nl l in + if not status then Feedback.feedback Feedback.AddedAxiom + +let should_treat_as_cumulative cum poly = + if poly then + match cum with + | GlobalCumulativity | LocalCumulativity -> true + | GlobalNonCumulativity | LocalNonCumulativity -> false + else + match cum with + | GlobalCumulativity | GlobalNonCumulativity -> false + | LocalCumulativity -> + user_err Pp.(str "The Cumulative prefix can only be used in a polymorphic context.") + | LocalNonCumulativity -> + user_err Pp.(str "The NonCumulative prefix can only be used in a polymorphic context.") + +let vernac_record cum k poly finite struc binders sort nameopt cfs = + let is_cumulative = should_treat_as_cumulative cum poly in + let const = match nameopt with + | None -> add_prefix "Build_" (fst (snd struc)).v + | Some ({v=id} as lid) -> + Dumpglob.dump_definition lid false "constr"; id in + if Dumpglob.dump () then ( + Dumpglob.dump_definition (fst (snd struc)) false "rec"; + List.iter (fun (((_, x), _), _) -> + match x with + | Vernacexpr.AssumExpr ({loc;v=Name id}, _) -> Dumpglob.dump_definition (make ?loc id) false "proj" + | _ -> ()) cfs); + ignore(Record.definition_structure (k,is_cumulative,poly,finite,struc,binders,cfs,const,sort)) + +(** When [poly] is true the type is declared polymorphic. When [lo] is true, + then the type is declared private (as per the [Private] keyword). [finite] + indicates whether the type is inductive, co-inductive or + neither. *) +let vernac_inductive ~atts cum lo finite indl = + let is_cumulative = should_treat_as_cumulative cum atts.polymorphic in + if Dumpglob.dump () then + List.iter (fun (((coe,(lid,_)), _, _, _, cstrs), _) -> + match cstrs with + | Constructors cstrs -> + Dumpglob.dump_definition lid false "ind"; + List.iter (fun (_, (lid, _)) -> + Dumpglob.dump_definition lid false "constr") cstrs + | _ -> () (* dumping is done by vernac_record (called below) *) ) + indl; + match indl with + | [ ( _ , _ , _ ,(Record|Structure), Constructors _ ),_ ] -> + user_err Pp.(str "The Record keyword is for types defined using the syntax { ... }.") + | [ (_ , _ , _ ,Variant, RecordDecl _),_ ] -> + user_err Pp.(str "The Variant keyword does not support syntax { ... }.") + | [ ( id , bl , c , b, RecordDecl (oc,fs) ), [] ] -> + vernac_record cum (match b with Class _ -> Class false | _ -> b) + atts.polymorphic finite id bl c oc fs + | [ ( id , bl , c , Class _, Constructors [l]), [] ] -> + let f = + let (coe, ({loc;v=id}, ce)) = l in + let coe' = if coe then Some true else None in + (((coe', AssumExpr ((make ?loc @@ Name id), ce)), None), []) + in vernac_record cum (Class true) atts.polymorphic finite id bl c None [f] + | [ ( _ , _, _, Class _, Constructors _), [] ] -> + user_err Pp.(str "Inductive classes not supported") + | [ ( id , bl , c , Class _, _), _ :: _ ] -> + user_err Pp.(str "where clause not supported for classes") + | [ ( _ , _ , _ , _, RecordDecl _ ) , _ ] -> + user_err Pp.(str "where clause not supported for (co)inductive records") + | _ -> let unpack = function + | ( (false, id) , bl , c , _ , Constructors l ) , ntn -> ( id , bl , c , l ) , ntn + | ( (true,_),_,_,_,Constructors _),_ -> + user_err Pp.(str "Variant types do not handle the \"> Name\" syntax, which is reserved for records. Use the \":>\" syntax on constructors instead.") + | _ -> user_err Pp.(str "Cannot handle mutually (co)inductive records.") + in + let indl = List.map unpack indl in + ComInductive.do_mutual_inductive indl is_cumulative atts.polymorphic lo finite + +let vernac_fixpoint ~atts discharge l = + let local = enforce_locality_exp atts.locality discharge in + if Dumpglob.dump () then + List.iter (fun (((lid,_), _, _, _, _), _) -> Dumpglob.dump_definition lid false "def") l; + (* XXX: Switch to the attribute system and match on ~atts *) + let do_fixpoint = if Flags.is_program_mode () then + ComProgramFixpoint.do_fixpoint + else + ComFixpoint.do_fixpoint + in + do_fixpoint local atts.polymorphic l + +let vernac_cofixpoint ~atts discharge l = + let local = enforce_locality_exp atts.locality discharge in + if Dumpglob.dump () then + List.iter (fun (((lid,_), _, _, _), _) -> Dumpglob.dump_definition lid false "def") l; + let do_cofixpoint = if Flags.is_program_mode () then + ComProgramFixpoint.do_cofixpoint + else + ComFixpoint.do_cofixpoint + in + do_cofixpoint local atts.polymorphic l + +let vernac_scheme l = + if Dumpglob.dump () then + List.iter (fun (lid, s) -> + Option.iter (fun lid -> Dumpglob.dump_definition lid false "def") lid; + match s with + | InductionScheme (_, r, _) + | CaseScheme (_, r, _) + | EqualityScheme r -> dump_global r) l; + Indschemes.do_scheme l + +let vernac_combined_scheme lid l = + if Dumpglob.dump () then + (Dumpglob.dump_definition lid false "def"; + List.iter (fun {loc;v=id} -> dump_global (make ?loc @@ Misctypes.AN (make ?loc @@ Ident id))) l); + Indschemes.do_combined_scheme lid l + +let vernac_universe ~atts l = + if atts.polymorphic && not (Lib.sections_are_opened ()) then + user_err ?loc:atts.loc ~hdr:"vernac_universe" + (str"Polymorphic universes can only be declared inside sections, " ++ + str "use Monomorphic Universe instead"); + Declare.do_universe atts.polymorphic l + +let vernac_constraint ~atts l = + if atts.polymorphic && not (Lib.sections_are_opened ()) then + user_err ?loc:atts.loc ~hdr:"vernac_constraint" + (str"Polymorphic universe constraints can only be declared" + ++ str " inside sections, use Monomorphic Constraint instead"); + Declare.do_constraint atts.polymorphic l + +(**********************) +(* Modules *) + +let vernac_import export refl = + Library.import_module export (List.map qualid_of_reference refl) + +let vernac_declare_module export {loc;v=id} binders_ast mty_ast = + (* We check the state of the system (in section, in module type) + and what module information is supplied *) + if Lib.sections_are_opened () then + user_err Pp.(str "Modules and Module Types are not allowed inside sections."); + let binders_ast = List.map + (fun (export,idl,ty) -> + if not (Option.is_empty export) then + user_err Pp.(str "Arguments of a functor declaration cannot be exported. Remove the \"Export\" and \"Import\" keywords from every functor argument.") + else (idl,ty)) binders_ast in + let mp = + Declaremods.declare_module Modintern.interp_module_ast + id binders_ast (Enforce mty_ast) [] + in + Dumpglob.dump_moddef ?loc mp "mod"; + Flags.if_verbose Feedback.msg_info (str "Module " ++ Id.print id ++ str " is declared"); + Option.iter (fun export -> vernac_import export [make @@ Ident id]) export + +let vernac_define_module export {loc;v=id} (binders_ast : module_binder list) mty_ast_o mexpr_ast_l = + (* We check the state of the system (in section, in module type) + and what module information is supplied *) + if Lib.sections_are_opened () then + user_err Pp.(str "Modules and Module Types are not allowed inside sections."); + match mexpr_ast_l with + | [] -> + Proof_global.check_no_pending_proof (); + let binders_ast,argsexport = + List.fold_right + (fun (export,idl,ty) (args,argsexport) -> + (idl,ty)::args, (List.map (fun {v=i} -> export,i)idl)@argsexport) binders_ast + ([],[]) in + let mp = + Declaremods.start_module Modintern.interp_module_ast + export id binders_ast mty_ast_o + in + Dumpglob.dump_moddef ?loc mp "mod"; + Flags.if_verbose Feedback.msg_info + (str "Interactive Module " ++ Id.print id ++ str " started"); + List.iter + (fun (export,id) -> + Option.iter + (fun export -> vernac_import export [make @@ Ident id]) export + ) argsexport + | _::_ -> + let binders_ast = List.map + (fun (export,idl,ty) -> + if not (Option.is_empty export) then + user_err Pp.(str "Arguments of a functor definition can be imported only if the definition is interactive. Remove the \"Export\" and \"Import\" keywords from every functor argument.") + else (idl,ty)) binders_ast in + let mp = + Declaremods.declare_module Modintern.interp_module_ast + id binders_ast mty_ast_o mexpr_ast_l + in + Dumpglob.dump_moddef ?loc mp "mod"; + Flags.if_verbose Feedback.msg_info + (str "Module " ++ Id.print id ++ str " is defined"); + Option.iter (fun export -> vernac_import export [make @@ Ident id]) + export + +let vernac_end_module export {loc;v=id} = + let mp = Declaremods.end_module () in + Dumpglob.dump_modref ?loc mp "mod"; + Flags.if_verbose Feedback.msg_info (str "Module " ++ Id.print id ++ str " is defined"); + Option.iter (fun export -> vernac_import export [make ?loc @@ Ident id]) export + +let vernac_declare_module_type {loc;v=id} binders_ast mty_sign mty_ast_l = + if Lib.sections_are_opened () then + user_err Pp.(str "Modules and Module Types are not allowed inside sections."); + + match mty_ast_l with + | [] -> + Proof_global.check_no_pending_proof (); + let binders_ast,argsexport = + List.fold_right + (fun (export,idl,ty) (args,argsexport) -> + (idl,ty)::args, (List.map (fun {v=i} -> export,i)idl)@argsexport) binders_ast + ([],[]) in + + let mp = + Declaremods.start_modtype Modintern.interp_module_ast + id binders_ast mty_sign + in + Dumpglob.dump_moddef ?loc mp "modtype"; + Flags.if_verbose Feedback.msg_info + (str "Interactive Module Type " ++ Id.print id ++ str " started"); + List.iter + (fun (export,id) -> + Option.iter + (fun export -> vernac_import export [make ?loc @@ Ident id]) export + ) argsexport + + | _ :: _ -> + let binders_ast = List.map + (fun (export,idl,ty) -> + if not (Option.is_empty export) then + user_err Pp.(str "Arguments of a functor definition can be imported only if the definition is interactive. Remove the \"Export\" and \"Import\" keywords from every functor argument.") + else (idl,ty)) binders_ast in + let mp = + Declaremods.declare_modtype Modintern.interp_module_ast + id binders_ast mty_sign mty_ast_l + in + Dumpglob.dump_moddef ?loc mp "modtype"; + Flags.if_verbose Feedback.msg_info + (str "Module Type " ++ Id.print id ++ str " is defined") + +let vernac_end_modtype {loc;v=id} = + let mp = Declaremods.end_modtype () in + Dumpglob.dump_modref ?loc mp "modtype"; + Flags.if_verbose Feedback.msg_info (str "Module Type " ++ Id.print id ++ str " is defined") + +let vernac_include l = + Declaremods.declare_include Modintern.interp_module_ast l + +(**********************) +(* Gallina extensions *) + +(* Sections *) + +let vernac_begin_section ({v=id} as lid) = + Proof_global.check_no_pending_proof (); + Dumpglob.dump_definition lid true "sec"; + Lib.open_section id + +let vernac_end_section {CAst.loc} = + Dumpglob.dump_reference ?loc + (DirPath.to_string (Lib.current_dirpath true)) "<>" "sec"; + Lib.close_section () + +let vernac_name_sec_hyp {v=id} set = Proof_using.name_set id set + +(* Dispatcher of the "End" command *) + +let vernac_end_segment ({v=id} as lid) = + Proof_global.check_no_pending_proof (); + match Lib.find_opening_node id with + | Lib.OpenedModule (false,export,_,_) -> vernac_end_module export lid + | Lib.OpenedModule (true,_,_,_) -> vernac_end_modtype lid + | Lib.OpenedSection _ -> vernac_end_section lid + | _ -> assert false + +(* Libraries *) + +let warn_require_in_section = + let name = "require-in-section" in + let category = "deprecated" in + CWarnings.create ~name ~category + (fun () -> strbrk "Use of “Require” inside a section is deprecated.") + +let vernac_require from import qidl = + if Lib.sections_are_opened () then warn_require_in_section (); + let qidl = List.map qualid_of_reference qidl in + let root = match from with + | None -> None + | Some from -> + let qid = Libnames.qualid_of_reference from in + let (hd, tl) = Libnames.repr_qualid qid.v in + Some (Libnames.add_dirpath_suffix hd tl) + in + let locate {loc;v=qid} = + try + let warn = not !Flags.quiet in + let (_, dir, f) = Library.locate_qualified_library ?root ~warn qid in + (dir, f) + with + | Library.LibUnmappedDir -> err_unmapped_library ?loc ?from:root qid + | Library.LibNotFound -> err_notfound_library ?loc ?from:root qid + in + let modrefl = List.map locate qidl in + if Dumpglob.dump () then + List.iter2 (fun {CAst.loc} dp -> Dumpglob.dump_libref ?loc dp "lib") qidl (List.map fst modrefl); + Library.require_library_from_dirpath modrefl import + +(* Coercions and canonical structures *) + +let vernac_canonical r = + Recordops.declare_canonical_structure (smart_global r) + +let vernac_coercion ~atts ref qids qidt = + let local = enforce_locality atts.locality in + let target = cl_of_qualid qidt in + let source = cl_of_qualid qids in + let ref' = smart_global ref in + Class.try_add_new_coercion_with_target ref' ~local atts.polymorphic ~source ~target; + Flags.if_verbose Feedback.msg_info (pr_global ref' ++ str " is now a coercion") + +let vernac_identity_coercion ~atts id qids qidt = + let local = enforce_locality atts.locality in + let target = cl_of_qualid qidt in + let source = cl_of_qualid qids in + Class.try_add_new_identity_coercion id ~local atts.polymorphic ~source ~target + +(* Type classes *) + +let vernac_instance ~atts abst sup inst props pri = + let global = not (make_section_locality atts.locality) in + Dumpglob.dump_constraint inst false "inst"; + let program_mode = Flags.is_program_mode () in + ignore(Classes.new_instance ~program_mode ~abstract:abst ~global atts.polymorphic sup inst props pri) + +let vernac_context ~atts l = + if not (Classes.context atts.polymorphic l) then Feedback.feedback Feedback.AddedAxiom + +let vernac_declare_instances ~atts insts = + let glob = not (make_section_locality atts.locality) in + List.iter (fun (id, info) -> Classes.existing_instance glob id (Some info)) insts + +let vernac_declare_class id = + Record.declare_existing_class (Nametab.global id) + +(***********) +(* Solving *) + +let command_focus = Proof.new_focus_kind () +let focus_command_cond = Proof.no_cond command_focus + + (* A command which should be a tactic. It has been + added by Christine to patch an error in the design of the proof + machine, and enables to instantiate existential variables when + there are no more goals to solve. It cannot be a tactic since + all tactics fail if there are no further goals to prove. *) + +let vernac_solve_existential = Pfedit.instantiate_nth_evar_com + +let vernac_set_end_tac tac = + let env = Genintern.empty_glob_sign (Global.env ()) in + let _, tac = Genintern.generic_intern env tac in + if not (Proof_global.there_are_pending_proofs ()) then + user_err Pp.(str "Unknown command of the non proof-editing mode."); + Proof_global.set_endline_tactic tac + (* TO DO verifier s'il faut pas mettre exist s | TacId s ici*) + +let vernac_set_used_variables e = + let env = Global.env () in + let tys = + List.map snd (Proof.initial_goals (Proof_global.give_me_the_proof ())) in + let tys = List.map EConstr.Unsafe.to_constr tys in + let l = Proof_using.process_expr env e tys in + let vars = Environ.named_context env in + List.iter (fun id -> + if not (List.exists (NamedDecl.get_id %> Id.equal id) vars) then + user_err ~hdr:"vernac_set_used_variables" + (str "Unknown variable: " ++ Id.print id)) + l; + let _, to_clear = Proof_global.set_used_variables l in + let to_clear = List.map (fun x -> x.CAst.v) to_clear in + Proof_global.with_current_proof begin fun _ p -> + if List.is_empty to_clear then (p, ()) + else + let tac = Tactics.clear to_clear in + fst (Pfedit.solve SelectAll None tac p), () + end + +(*****************************) +(* Auxiliary file management *) + +let expand filename = + Envars.expand_path_macros ~warn:(fun x -> Feedback.msg_warning (str x)) filename + +let vernac_add_loadpath implicit pdir ldiropt = + let open Mltop in + let pdir = expand pdir in + let alias = Option.default Libnames.default_root_prefix ldiropt in + add_coq_path { recursive = true; + path_spec = VoPath { unix_path = pdir; coq_path = alias; has_ml = AddTopML; implicit } } + +let vernac_remove_loadpath path = + Loadpath.remove_load_path (expand path) + + (* Coq syntax for ML or system commands *) + +let vernac_add_ml_path isrec path = + let open Mltop in + add_coq_path { recursive = isrec; path_spec = MlPath (expand path) } + +let vernac_declare_ml_module ~atts l = + let local = make_locality atts.locality in + Mltop.declare_ml_modules local (List.map expand l) + +let vernac_chdir = function + | None -> Feedback.msg_notice (str (Sys.getcwd())) + | Some path -> + begin + try Sys.chdir (expand path) + with Sys_error err -> + (* Cd is typically used to control the output directory of + extraction. A failed Cd could lead to overwriting .ml files + so we make it an error. *) + user_err Pp.(str ("Cd failed: " ^ err)) + end; + Flags.if_verbose Feedback.msg_info (str (Sys.getcwd())) + +(********************) +(* State management *) + +let vernac_write_state file = + Proof_global.discard_all (); + let file = CUnix.make_suffix file ".coq" in + States.extern_state file + +let vernac_restore_state file = + Proof_global.discard_all (); + let file = Loadpath.locate_file (CUnix.make_suffix file ".coq") in + States.intern_state file + +(************) +(* Commands *) + +let vernac_create_hintdb ~atts id b = + let local = make_module_locality atts.locality in + Hints.create_hint_db local id full_transparent_state b + +let vernac_remove_hints ~atts dbs ids = + let local = make_module_locality atts.locality in + Hints.remove_hints local dbs (List.map Smartlocate.global_with_alias ids) + +let vernac_hints ~atts lb h = + let local = enforce_module_locality atts.locality in + Hints.add_hints local lb (Hints.interp_hints atts.polymorphic h) + +let vernac_syntactic_definition ~atts lid x y = + Dumpglob.dump_definition lid false "syndef"; + let local = enforce_module_locality atts.locality in + Metasyntax.add_syntactic_definition (Global.env()) lid.v x local y + +let vernac_declare_implicits ~atts r l = + let local = make_section_locality atts.locality in + match l with + | [] -> + Impargs.declare_implicits local (smart_global r) + | _::_ as imps -> + Impargs.declare_manual_implicits local (smart_global r) ~enriching:false + (List.map (List.map (fun (ex,b,f) -> ex, (b,true,f))) imps) + +let warn_arguments_assert = + CWarnings.create ~name:"arguments-assert" ~category:"vernacular" + (fun sr -> + strbrk "This command is just asserting the names of arguments of " ++ + pr_global sr ++ strbrk". If this is what you want add " ++ + strbrk "': assert' to silence the warning. If you want " ++ + strbrk "to clear implicit arguments add ': clear implicits'. " ++ + strbrk "If you want to clear notation scopes add ': clear scopes'") + +(* [nargs_for_red] is the number of arguments required to trigger reduction, + [args] is the main list of arguments statuses, + [more_implicits] is a list of extra lists of implicit statuses *) +let vernac_arguments ~atts reference args more_implicits nargs_for_red flags = + let assert_flag = List.mem `Assert flags in + let rename_flag = List.mem `Rename flags in + let clear_scopes_flag = List.mem `ClearScopes flags in + let extra_scopes_flag = List.mem `ExtraScopes flags in + let clear_implicits_flag = List.mem `ClearImplicits flags in + let default_implicits_flag = List.mem `DefaultImplicits flags in + let never_unfold_flag = List.mem `ReductionNeverUnfold flags in + + let err_incompat x y = + user_err Pp.(str ("Options \""^x^"\" and \""^y^"\" are incompatible.")) in + + if assert_flag && rename_flag then + err_incompat "assert" "rename"; + if Option.has_some nargs_for_red && never_unfold_flag then + err_incompat "simpl never" "/"; + if never_unfold_flag && List.mem `ReductionDontExposeCase flags then + err_incompat "simpl never" "simpl nomatch"; + if clear_scopes_flag && extra_scopes_flag then + err_incompat "clear scopes" "extra scopes"; + if clear_implicits_flag && default_implicits_flag then + err_incompat "clear implicits" "default implicits"; + + let sr = smart_global reference in + let inf_names = + let ty, _ = Global.type_of_global_in_context (Global.env ()) sr in + let env = Global.env () in + let sigma = Evd.from_env env in + Impargs.compute_implicits_names env sigma (EConstr.of_constr ty) + in + let prev_names = + try Arguments_renaming.arguments_names sr with Not_found -> inf_names + in + let num_args = List.length inf_names in + assert (Int.equal num_args (List.length prev_names)); + + let names_of args = List.map (fun a -> a.name) args in + + (* Checks *) + + let err_extra_args names = + user_err ~hdr:"vernac_declare_arguments" + (strbrk "Extra arguments: " ++ + prlist_with_sep pr_comma Name.print names ++ str ".") + in + let err_missing_args names = + user_err ~hdr:"vernac_declare_arguments" + (strbrk "The following arguments are not declared: " ++ + prlist_with_sep pr_comma Name.print names ++ str ".") + in + + let rec check_extra_args extra_args = + match extra_args with + | [] -> () + | { notation_scope = None } :: _ -> err_extra_args (names_of extra_args) + | { name = Anonymous; notation_scope = Some _ } :: args -> + check_extra_args args + | _ -> + user_err Pp.(str "Extra notation scopes can be set on anonymous and explicit arguments only.") + in + + let args, scopes = + let scopes = List.map (fun { notation_scope = s } -> s) args in + if List.length args > num_args then + let args, extra_args = List.chop num_args args in + if extra_scopes_flag then + (check_extra_args extra_args; (args, scopes)) + else err_extra_args (names_of extra_args) + else args, scopes + in + + if Option.cata (fun n -> n > num_args) false nargs_for_red then + user_err Pp.(str "The \"/\" modifier should be put before any extra scope."); + + let scopes_specified = List.exists Option.has_some scopes in + + if scopes_specified && clear_scopes_flag then + user_err Pp.(str "The \"clear scopes\" flag is incompatible with scope annotations."); + + let names = List.map (fun { name } -> name) args in + let names = names :: List.map (List.map fst) more_implicits in + + let rename_flag_required = ref false in + let example_renaming = ref None in + let save_example_renaming renaming = + rename_flag_required := !rename_flag_required + || not (Name.equal (fst renaming) Anonymous); + if Option.is_empty !example_renaming then + example_renaming := Some renaming + in + + let rec names_union names1 names2 = + match names1, names2 with + | [], [] -> [] + | _ :: _, [] -> names1 + | [], _ :: _ -> names2 + | (Name _ as name) :: names1, Anonymous :: names2 + | Anonymous :: names1, (Name _ as name) :: names2 -> + name :: names_union names1 names2 + | name1 :: names1, name2 :: names2 -> + if Name.equal name1 name2 then + name1 :: names_union names1 names2 + else user_err Pp.(str "Argument lists should agree on the names they provide.") + in + + let names = List.fold_left names_union [] names in + + let rec rename prev_names names = + match prev_names, names with + | [], [] -> [] + | [], _ :: _ -> err_extra_args names + | _ :: _, [] when assert_flag -> + (* Error messages are expressed in terms of original names, not + renamed ones. *) + err_missing_args (List.lastn (List.length prev_names) inf_names) + | _ :: _, [] -> prev_names + | prev :: prev_names, Anonymous :: names -> + prev :: rename prev_names names + | prev :: prev_names, (Name id as name) :: names -> + if not (Name.equal prev name) then save_example_renaming (prev,name); + name :: rename prev_names names + in + + let names = rename prev_names names in + let renaming_specified = Option.has_some !example_renaming in + + if !rename_flag_required && not rename_flag then + user_err ~hdr:"vernac_declare_arguments" + (strbrk "To rename arguments the \"rename\" flag must be specified." + ++ spc () ++ + match !example_renaming with + | None -> mt () + | Some (o,n) -> + str "Argument " ++ Name.print o ++ + str " renamed to " ++ Name.print n ++ str "."); + + let duplicate_names = + List.duplicates Name.equal (List.filter ((!=) Anonymous) names) + in + if not (List.is_empty duplicate_names) then begin + let duplicates = prlist_with_sep pr_comma Name.print duplicate_names in + user_err (strbrk "Some argument names are duplicated: " ++ duplicates) + end; + + (* Parts of this code are overly complicated because the implicit arguments + API is completely crazy: positions (ExplByPos) are elaborated to + names. This is broken by design, since not all arguments have names. So + even though we eventually want to map only positions to implicit statuses, + we have to check whether the corresponding arguments have names, not to + trigger an error in the impargs code. Even better, the names we have to + check are not the current ones (after previous renamings), but the original + ones (inferred from the type). *) + + let implicits = + List.map (fun { name; implicit_status = i } -> (name,i)) args + in + let implicits = implicits :: more_implicits in + + let open Vernacexpr in + let rec build_implicits inf_names implicits = + match inf_names, implicits with + | _, [] -> [] + | _ :: inf_names, (_, NotImplicit) :: implicits -> + build_implicits inf_names implicits + + (* With the current impargs API, it is impossible to make an originally + anonymous argument implicit *) + | Anonymous :: _, (name, _) :: _ -> + user_err ~hdr:"vernac_declare_arguments" + (strbrk"Argument "++ Name.print name ++ + strbrk " cannot be declared implicit.") + + | Name id :: inf_names, (name, impl) :: implicits -> + let max = impl = MaximallyImplicit in + (ExplByName id,max,false) :: build_implicits inf_names implicits + + | _ -> assert false (* already checked in [names_union] *) + in + + let implicits = List.map (build_implicits inf_names) implicits in + let implicits_specified = match implicits with [[]] -> false | _ -> true in + + if implicits_specified && clear_implicits_flag then + user_err Pp.(str "The \"clear implicits\" flag is incompatible with implicit annotations"); + + if implicits_specified && default_implicits_flag then + user_err Pp.(str "The \"default implicits\" flag is incompatible with implicit annotations"); + + let rargs = + Util.List.map_filter (function (n, true) -> Some n | _ -> None) + (Util.List.map_i (fun i { recarg_like = b } -> i, b) 0 args) + in + + let rec narrow = function + | #Reductionops.ReductionBehaviour.flag as x :: tl -> x :: narrow tl + | [] -> [] | _ :: tl -> narrow tl + in + let red_flags = narrow flags in + let red_modifiers_specified = + not (List.is_empty rargs) || Option.has_some nargs_for_red + || not (List.is_empty red_flags) + in + + if not (List.is_empty rargs) && never_unfold_flag then + err_incompat "simpl never" "!"; + + + (* Actions *) + + if renaming_specified then begin + let local = make_section_locality atts.locality in + Arguments_renaming.rename_arguments local sr names + end; + + if scopes_specified || clear_scopes_flag then begin + let scopes = List.map (Option.map (fun {loc;v=k} -> + try ignore (Notation.find_scope k); k + with UserError _ -> + Notation.find_delimiters_scope ?loc k)) scopes + in + vernac_arguments_scope ~atts reference scopes + end; + + if implicits_specified || clear_implicits_flag then + vernac_declare_implicits ~atts reference implicits; + + if default_implicits_flag then + vernac_declare_implicits ~atts reference []; + + if red_modifiers_specified then begin + match sr with + | ConstRef _ as c -> + Reductionops.ReductionBehaviour.set + (make_section_locality atts.locality) c + (rargs, Option.default ~-1 nargs_for_red, red_flags) + | _ -> user_err + (strbrk "Modifiers of the behavior of the simpl tactic "++ + strbrk "are relevant for constants only.") + end; + + if not (renaming_specified || + implicits_specified || + scopes_specified || + red_modifiers_specified) && (List.is_empty flags) then + warn_arguments_assert sr + +let default_env () = { + Notation_term.ninterp_var_type = Id.Map.empty; + ninterp_rec_vars = Id.Map.empty; +} + +let vernac_reserve bl = + let sb_decl = (fun (idl,c) -> + let env = Global.env() in + let sigma = Evd.from_env env in + let t,ctx = Constrintern.interp_type env sigma c in + let t = Detyping.detype Detyping.Now false Id.Set.empty env (Evd.from_ctx ctx) t in + let t,_ = Notation_ops.notation_constr_of_glob_constr (default_env ()) t in + Reserve.declare_reserved_type idl t) + in List.iter sb_decl bl + +let vernac_generalizable ~atts = + let local = make_non_locality atts.locality in + Implicit_quantifiers.declare_generalizable local + +let _ = + declare_bool_option + { optdepr = false; + optname = "silent"; + optkey = ["Silent"]; + optread = (fun () -> !Flags.quiet); + optwrite = ((:=) Flags.quiet) } + +let _ = + declare_bool_option + { optdepr = false; + optname = "implicit arguments"; + optkey = ["Implicit";"Arguments"]; + optread = Impargs.is_implicit_args; + optwrite = Impargs.make_implicit_args } + +let _ = + declare_bool_option + { optdepr = false; + optname = "strict implicit arguments"; + optkey = ["Strict";"Implicit"]; + optread = Impargs.is_strict_implicit_args; + optwrite = Impargs.make_strict_implicit_args } + +let _ = + declare_bool_option + { optdepr = false; + optname = "strong strict implicit arguments"; + optkey = ["Strongly";"Strict";"Implicit"]; + optread = Impargs.is_strongly_strict_implicit_args; + optwrite = Impargs.make_strongly_strict_implicit_args } + +let _ = + declare_bool_option + { optdepr = false; + optname = "contextual implicit arguments"; + optkey = ["Contextual";"Implicit"]; + optread = Impargs.is_contextual_implicit_args; + optwrite = Impargs.make_contextual_implicit_args } + +let _ = + declare_bool_option + { optdepr = false; + optname = "implicit status of reversible patterns"; + optkey = ["Reversible";"Pattern";"Implicit"]; + optread = Impargs.is_reversible_pattern_implicit_args; + optwrite = Impargs.make_reversible_pattern_implicit_args } + +let _ = + declare_bool_option + { optdepr = false; + optname = "maximal insertion of implicit"; + optkey = ["Maximal";"Implicit";"Insertion"]; + optread = Impargs.is_maximal_implicit_args; + optwrite = Impargs.make_maximal_implicit_args } + +let _ = + declare_bool_option + { optdepr = true; (* remove in 8.8 *) + optname = "automatic introduction of variables"; + optkey = ["Automatic";"Introduction"]; + optread = Flags.is_auto_intros; + optwrite = Flags.make_auto_intros } + +let _ = + declare_bool_option + { optdepr = false; + optname = "coercion printing"; + optkey = ["Printing";"Coercions"]; + optread = (fun () -> !Constrextern.print_coercions); + optwrite = (fun b -> Constrextern.print_coercions := b) } + +let _ = + declare_bool_option + { optdepr = false; + optname = "printing of existential variable instances"; + optkey = ["Printing";"Existential";"Instances"]; + optread = (fun () -> !Detyping.print_evar_arguments); + optwrite = (:=) Detyping.print_evar_arguments } + +let _ = + declare_bool_option + { optdepr = false; + optname = "implicit arguments printing"; + optkey = ["Printing";"Implicit"]; + optread = (fun () -> !Constrextern.print_implicits); + optwrite = (fun b -> Constrextern.print_implicits := b) } + +let _ = + declare_bool_option + { optdepr = false; + optname = "implicit arguments defensive printing"; + optkey = ["Printing";"Implicit";"Defensive"]; + optread = (fun () -> !Constrextern.print_implicits_defensive); + optwrite = (fun b -> Constrextern.print_implicits_defensive := b) } + +let _ = + declare_bool_option + { optdepr = false; + optname = "projection printing using dot notation"; + optkey = ["Printing";"Projections"]; + optread = (fun () -> !Constrextern.print_projections); + optwrite = (fun b -> Constrextern.print_projections := b) } + +let _ = + declare_bool_option + { optdepr = false; + optname = "notations printing"; + optkey = ["Printing";"Notations"]; + optread = (fun () -> not !Constrextern.print_no_symbol); + optwrite = (fun b -> Constrextern.print_no_symbol := not b) } + +let _ = + declare_bool_option + { optdepr = false; + optname = "raw printing"; + optkey = ["Printing";"All"]; + optread = (fun () -> !Flags.raw_print); + optwrite = (fun b -> Flags.raw_print := b) } + +let _ = + declare_bool_option + { optdepr = false; + optname = "use of the program extension"; + optkey = ["Program";"Mode"]; + optread = (fun () -> !Flags.program_mode); + optwrite = (fun b -> Flags.program_mode:=b) } + +let universe_polymorphism_option_name = ["Universe"; "Polymorphism"] + +let _ = + declare_bool_option + { optdepr = false; + optname = "universe polymorphism"; + optkey = universe_polymorphism_option_name; + optread = Flags.is_universe_polymorphism; + optwrite = Flags.make_universe_polymorphism } + +let _ = + declare_bool_option + { optdepr = false; + optname = "Polymorphic inductive cumulativity"; + optkey = ["Polymorphic"; "Inductive"; "Cumulativity"]; + optread = Flags.is_polymorphic_inductive_cumulativity; + optwrite = Flags.make_polymorphic_inductive_cumulativity } + +let _ = + declare_int_option + { optdepr = false; + optname = "the level of inlining during functor application"; + optkey = ["Inline";"Level"]; + optread = (fun () -> Some (Flags.get_inline_level ())); + optwrite = (fun o -> + let lev = Option.default Flags.default_inline_level o in + Flags.set_inline_level lev) } + +let _ = + declare_bool_option + { optdepr = false; + optname = "kernel term sharing"; + optkey = ["Kernel"; "Term"; "Sharing"]; + optread = (fun () -> !CClosure.share); + optwrite = (fun b -> CClosure.share := b) } + +let _ = + declare_bool_option + { optdepr = false; + optname = "display compact goal contexts"; + optkey = ["Printing";"Compact";"Contexts"]; + optread = (fun () -> Printer.get_compact_context()); + optwrite = (fun b -> Printer.set_compact_context b) } + +let _ = + declare_int_option + { optdepr = false; + optname = "the printing depth"; + optkey = ["Printing";"Depth"]; + optread = Topfmt.get_depth_boxes; + optwrite = Topfmt.set_depth_boxes } + +let _ = + declare_int_option + { optdepr = false; + optname = "the printing width"; + optkey = ["Printing";"Width"]; + optread = Topfmt.get_margin; + optwrite = Topfmt.set_margin } + +let _ = + declare_bool_option + { optdepr = false; + optname = "printing of universes"; + optkey = ["Printing";"Universes"]; + optread = (fun () -> !Constrextern.print_universes); + optwrite = (fun b -> Constrextern.print_universes:=b) } + +let _ = + declare_bool_option + { optdepr = false; + optname = "dumping bytecode after compilation"; + optkey = ["Dump";"Bytecode"]; + optread = Flags.get_dump_bytecode; + optwrite = Flags.set_dump_bytecode } + +let _ = + declare_bool_option + { optdepr = false; + optname = "dumping VM lambda code after compilation"; + optkey = ["Dump";"Lambda"]; + optread = Flags.get_dump_lambda; + optwrite = Flags.set_dump_lambda } + +let _ = + declare_bool_option + { optdepr = false; + optname = "explicitly parsing implicit arguments"; + optkey = ["Parsing";"Explicit"]; + optread = (fun () -> !Constrintern.parsing_explicit); + optwrite = (fun b -> Constrintern.parsing_explicit := b) } + +let _ = + declare_string_option ~preprocess:CWarnings.normalize_flags_string + { optdepr = false; + optname = "warnings display"; + optkey = ["Warnings"]; + optread = CWarnings.get_flags; + optwrite = CWarnings.set_flags } + +let _ = + declare_string_option + { optdepr = false; + optname = "native_compute profiler output"; + optkey = ["NativeCompute"; "Profile"; "Filename"]; + optread = Nativenorm.get_profile_filename; + optwrite = Nativenorm.set_profile_filename } + +let _ = + declare_bool_option + { optdepr = false; + optname = "enable native compute profiling"; + optkey = ["NativeCompute"; "Profiling"]; + optread = Nativenorm.get_profiling_enabled; + optwrite = Nativenorm.set_profiling_enabled } + +let vernac_set_strategy ~atts l = + let local = make_locality atts.locality in + let glob_ref r = + match smart_global r with + | ConstRef sp -> EvalConstRef sp + | VarRef id -> EvalVarRef id + | _ -> user_err Pp.(str + "cannot set an inductive type or a constructor as transparent") in + let l = List.map (fun (lev,ql) -> (lev,List.map glob_ref ql)) l in + Redexpr.set_strategy local l + +let vernac_set_opacity ~atts (v,l) = + let local = make_non_locality atts.locality in + let glob_ref r = + match smart_global r with + | ConstRef sp -> EvalConstRef sp + | VarRef id -> EvalVarRef id + | _ -> user_err Pp.(str + "cannot set an inductive type or a constructor as transparent") in + let l = List.map glob_ref l in + Redexpr.set_strategy local [v,l] + +let get_option_locality export local = + if export then + if Option.is_empty local then OptExport + else user_err Pp.(str "Locality modifiers forbidden with Export") + else match local with + | Some true -> OptLocal + | Some false -> OptGlobal + | None -> OptDefault + +let vernac_set_option0 ~atts export key opt = + let locality = get_option_locality export atts.locality in + match opt with + | StringValue s -> set_string_option_value_gen ~locality key s + | StringOptValue (Some s) -> set_string_option_value_gen ~locality key s + | StringOptValue None -> unset_option_value_gen ~locality key + | IntValue n -> set_int_option_value_gen ~locality key n + | BoolValue b -> set_bool_option_value_gen ~locality key b + +let vernac_set_append_option ~atts export key s = + let locality = get_option_locality export atts.locality in + set_string_option_append_value_gen ~locality key s + +let vernac_set_option ~atts export table v = match v with +| StringValue s -> + (* We make a special case for warnings because appending is their + natural semantics *) + if CString.List.equal table ["Warnings"] then + vernac_set_append_option ~atts export table s + else + let (last, prefix) = List.sep_last table in + if String.equal last "Append" && not (List.is_empty prefix) then + vernac_set_append_option ~atts export prefix s + else + vernac_set_option0 ~atts export table v +| _ -> vernac_set_option0 ~atts export table v + +let vernac_unset_option ~atts export key = + let locality = get_option_locality export atts.locality in + unset_option_value_gen ~locality key + +let vernac_add_option key lv = + let f = function + | StringRefValue s -> (get_string_table key)#add s + | QualidRefValue locqid -> (get_ref_table key)#add locqid + in + try List.iter f lv with Not_found -> error_undeclared_key key + +let vernac_remove_option key lv = + let f = function + | StringRefValue s -> (get_string_table key)#remove s + | QualidRefValue locqid -> (get_ref_table key)#remove locqid + in + try List.iter f lv with Not_found -> error_undeclared_key key + +let vernac_mem_option key lv = + let f = function + | StringRefValue s -> (get_string_table key)#mem s + | QualidRefValue locqid -> (get_ref_table key)#mem locqid + in + try List.iter f lv with Not_found -> error_undeclared_key key + +let vernac_print_option key = + try (get_ref_table key)#print + with Not_found -> + try (get_string_table key)#print + with Not_found -> + try print_option_value key + with Not_found -> error_undeclared_key key + +let get_current_context_of_args = function + | Some n -> Pfedit.get_goal_context n + | None -> Pfedit.get_current_context () + +let query_command_selector ?loc = function + | None -> None + | Some (SelectNth n) -> Some n + | _ -> user_err ?loc ~hdr:"query_command_selector" + (str "Query commands only support the single numbered goal selector.") + +let vernac_check_may_eval ~atts redexp glopt rc = + let glopt = query_command_selector ?loc:atts.loc glopt in + let (sigma, env) = get_current_context_of_args glopt in + let sigma, c = interp_open_constr env sigma rc in + let c = EConstr.Unsafe.to_constr c in + let sigma = Evarconv.solve_unif_constraints_with_heuristics env sigma in + Evarconv.check_problems_are_solved env sigma; + let sigma,nf = Evarutil.nf_evars_and_universes sigma in + let uctx = Evd.universe_context_set sigma in + let env = Environ.push_context_set uctx (Evarutil.nf_env_evar sigma env) in + let c = nf c in + let j = + if Evarutil.has_undefined_evars sigma (EConstr.of_constr c) then + Evarutil.j_nf_evar sigma (Retyping.get_judgment_of env sigma (EConstr.of_constr c)) + else + (* OK to call kernel which does not support evars *) + Termops.on_judgment EConstr.of_constr (Arguments_renaming.rename_typing env c) + in + let pp = match redexp with + | None -> + let evars_of_term c = Evarutil.undefined_evars_of_term sigma c in + let l = Evar.Set.union (evars_of_term j.Environ.uj_val) (evars_of_term j.Environ.uj_type) in + let j = { j with Environ.uj_type = Reductionops.nf_betaiota env sigma j.Environ.uj_type } in + print_judgment env sigma j ++ + pr_ne_evar_set (fnl () ++ str "where" ++ fnl ()) (mt ()) sigma l + | Some r -> + let (sigma,r_interp) = Hook.get f_interp_redexp env sigma r in + let redfun env evm c = + let (redfun, _) = reduction_of_red_expr env r_interp in + let (_, c) = redfun env evm c in + c + in + print_eval redfun env sigma rc j + in + pp ++ Printer.pr_universe_ctx_set sigma uctx + +let vernac_declare_reduction ~atts s r = + let local = make_locality atts.locality in + declare_red_expr local s (snd (Hook.get f_interp_redexp (Global.env()) Evd.empty r)) + + (* The same but avoiding the current goal context if any *) +let vernac_global_check c = + let env = Global.env() in + let sigma = Evd.from_env env in + let c,uctx = interp_constr env sigma c in + let senv = Global.safe_env() in + let uctx = UState.context_set uctx in + let senv = Safe_typing.push_context_set false uctx senv in + let c = EConstr.to_constr sigma c in + let j = Safe_typing.typing senv c in + let env = Safe_typing.env_of_safe_env senv in + print_safe_judgment env sigma j ++ + pr_universe_ctx_set sigma uctx + + +let get_nth_goal n = + let pf = Proof_global.give_me_the_proof() in + let gls,_,_,_,sigma = Proof.proof pf in + let gl = {Evd.it=List.nth gls (n-1) ; sigma = sigma; } in + gl + +exception NoHyp +(* Printing "About" information of a hypothesis of the current goal. + We only print the type and a small statement to this comes from the + goal. Precondition: there must be at least one current goal. *) +let print_about_hyp_globs ?loc ref_or_by_not udecl glopt = + let open Context.Named.Declaration in + try + (* FIXME error on non None udecl if we find the hyp. *) + let glnumopt = query_command_selector ?loc glopt in + let gl,id = + match glnumopt, ref_or_by_not.v with + | None,AN {v=Ident id} -> (* goal number not given, catch any failure *) + (try get_nth_goal 1,id with _ -> raise NoHyp) + | Some n,AN {v=Ident id} -> (* goal number given, catch if wong *) + (try get_nth_goal n,id + with + Failure _ -> user_err ?loc ~hdr:"print_about_hyp_globs" + (str "No such goal: " ++ int n ++ str ".")) + | _ , _ -> raise NoHyp in + let hyps = pf_hyps gl in + let decl = Context.Named.lookup id hyps in + let natureofid = match decl with + | LocalAssum _ -> "Hypothesis" + | LocalDef (_,bdy,_) ->"Constant (let in)" in + let sigma, env = Pfedit.get_current_context () in + v 0 (Id.print id ++ str":" ++ pr_econstr_env env sigma (NamedDecl.get_type decl) ++ fnl() ++ fnl() + ++ str natureofid ++ str " of the goal context.") + with (* fallback to globals *) + | NoHyp | Not_found -> + let sigma, env = Pfedit.get_current_context () in + print_about env sigma ref_or_by_not udecl + +let vernac_print ~atts env sigma = + let loc = atts.loc in + function + | PrintTables -> print_tables () + | PrintFullContext-> print_full_context_typ env sigma + | PrintSectionContext qid -> print_sec_context_typ env sigma qid + | PrintInspect n -> inspect env sigma n + | PrintGrammar ent -> Metasyntax.pr_grammar ent + | PrintLoadPath dir -> (* For compatibility ? *) print_loadpath dir + | PrintModules -> print_modules () + | PrintModule qid -> print_module qid + | PrintModuleType qid -> print_modtype qid + | PrintNamespace ns -> print_namespace ns + | PrintMLLoadPath -> Mltop.print_ml_path () + | PrintMLModules -> Mltop.print_ml_modules () + | PrintDebugGC -> Mltop.print_gc () + | PrintName (qid,udecl) -> + dump_global qid; + print_name env sigma qid udecl + | PrintGraph -> Prettyp.print_graph env sigma + | PrintClasses -> Prettyp.print_classes() + | PrintTypeClasses -> Prettyp.print_typeclasses() + | PrintInstances c -> Prettyp.print_instances (smart_global c) + | PrintCoercions -> Prettyp.print_coercions env sigma + | PrintCoercionPaths (cls,clt) -> + Prettyp.print_path_between env sigma (cl_of_qualid cls) (cl_of_qualid clt) + | PrintCanonicalConversions -> Prettyp.print_canonical_projections env sigma + | PrintUniverses (b, dst) -> + let univ = Global.universes () in + let univ = if b then UGraph.sort_universes univ else univ in + let pr_remaining = + if Global.is_joined_environment () then mt () + else str"There may remain asynchronous universe constraints" + in + begin match dst with + | None -> UGraph.pr_universes Universes.pr_with_global_universes univ ++ pr_remaining + | Some s -> dump_universes_gen univ s + end + | PrintHint r -> Hints.pr_hint_ref env sigma (smart_global r) + | PrintHintGoal -> Hints.pr_applicable_hint () + | PrintHintDbName s -> Hints.pr_hint_db_by_name env sigma s + | PrintHintDb -> Hints.pr_searchtable env sigma + | PrintScopes -> + Notation.pr_scopes (Constrextern.without_symbols (pr_lglob_constr_env env)) + | PrintScope s -> + Notation.pr_scope (Constrextern.without_symbols (pr_lglob_constr_env env)) s + | PrintVisibility s -> + Notation.pr_visibility (Constrextern.without_symbols (pr_lglob_constr_env env)) s + | PrintAbout (ref_or_by_not,udecl,glnumopt) -> + print_about_hyp_globs ?loc ref_or_by_not udecl glnumopt + | PrintImplicit qid -> + dump_global qid; + print_impargs qid + | PrintAssumptions (o,t,r) -> + (* Prints all the axioms and section variables used by a term *) + let gr = smart_global r in + let cstr = printable_constr_of_global gr in + let st = Conv_oracle.get_transp_state (Environ.oracle (Global.env())) in + let nassums = + Assumptions.assumptions st ~add_opaque:o ~add_transparent:t gr cstr in + Printer.pr_assumptionset env sigma nassums + | PrintStrategy r -> print_strategy r + +let global_module r = + let {loc;v=qid} = qualid_of_reference r in + try Nametab.full_name_module qid + with Not_found -> + user_err ?loc ~hdr:"global_module" + (str "Module/section " ++ pr_qualid qid ++ str " not found.") + +let interp_search_restriction = function + | SearchOutside l -> (List.map global_module l, true) + | SearchInside l -> (List.map global_module l, false) + +open Search + +let interp_search_about_item env sigma = + function + | SearchSubPattern pat -> + let _,pat = intern_constr_pattern env sigma pat in + GlobSearchSubPattern pat + | SearchString (s,None) when Id.is_valid s -> + GlobSearchString s + | SearchString (s,sc) -> + try + let ref = + Notation.interp_notation_as_global_reference + (fun _ -> true) s sc in + GlobSearchSubPattern (Pattern.PRef ref) + with UserError _ -> + user_err ~hdr:"interp_search_about_item" + (str "Unable to interp \"" ++ str s ++ str "\" either as a reference or as an identifier component") + +(* 05f22a5d6d5b8e3e80f1a37321708ce401834430 introduced the + `search_output_name_only` option to avoid excessive printing when + searching. + + The motivation was to make search usable for IDE completion, + however, it is still too slow due to the non-indexed nature of the + underlying search mechanism. + + In the future we should deprecate the option and provide a fast, + indexed name-searching interface. +*) +let search_output_name_only = ref false + +let _ = + declare_bool_option + { optdepr = false; + optname = "output-name-only search"; + optkey = ["Search";"Output";"Name";"Only"]; + optread = (fun () -> !search_output_name_only); + optwrite = (:=) search_output_name_only } + +let vernac_search ~atts s gopt r = + let gopt = query_command_selector ?loc:atts.loc gopt in + let r = interp_search_restriction r in + let env,gopt = + match gopt with | None -> + (* 1st goal by default if it exists, otherwise no goal at all *) + (try snd (Pfedit.get_goal_context 1) , Some 1 + with _ -> Global.env (),None) + (* if goal selector is given and wrong, then let exceptions be raised. *) + | Some g -> snd (Pfedit.get_goal_context g) , Some g + in + let get_pattern c = snd (intern_constr_pattern env Evd.(from_env env) c) in + let pr_search ref env c = + let pr = pr_global ref in + let pp = if !search_output_name_only + then pr + else begin + let pc = pr_lconstr_env env Evd.(from_env env) c in + hov 2 (pr ++ str":" ++ spc () ++ pc) + end + in Feedback.msg_notice pp + in + match s with + | SearchPattern c -> + (Search.search_pattern gopt (get_pattern c) r |> Search.prioritize_search) pr_search + | SearchRewrite c -> + (Search.search_rewrite gopt (get_pattern c) r |> Search.prioritize_search) pr_search + | SearchHead c -> + (Search.search_by_head gopt (get_pattern c) r |> Search.prioritize_search) pr_search + | SearchAbout sl -> + (Search.search_about gopt (List.map (on_snd (interp_search_about_item env Evd.(from_env env))) sl) r |> + Search.prioritize_search) pr_search + +let vernac_locate = function + | LocateAny {v=AN qid} -> print_located_qualid qid + | LocateTerm {v=AN qid} -> print_located_term qid + | LocateAny {v=ByNotation (ntn, sc)} (** TODO : handle Ltac notations *) + | LocateTerm {v=ByNotation (ntn, sc)} -> + let _, env = Pfedit.get_current_context () in + Notation.locate_notation + (Constrextern.without_symbols (pr_lglob_constr_env env)) ntn sc + | LocateLibrary qid -> print_located_library qid + | LocateModule qid -> print_located_module qid + | LocateOther (s, qid) -> print_located_other s qid + | LocateFile f -> locate_file f + +let vernac_register id r = + if Proof_global.there_are_pending_proofs () then + user_err Pp.(str "Cannot register a primitive while in proof editing mode."); + let kn = Constrintern.global_reference id.v in + if not (isConstRef kn) then + user_err Pp.(str "Register inline: a constant is expected"); + match r with + | RegisterInline -> Global.register_inline (destConstRef kn) + +(********************) +(* Proof management *) + +let vernac_focus gln = + Proof_global.simple_with_current_proof (fun _ p -> + match gln with + | None -> Proof.focus focus_command_cond () 1 p + | Some 0 -> + user_err Pp.(str "Invalid goal number: 0. Goal numbering starts with 1.") + | Some n -> + Proof.focus focus_command_cond () n p) + + (* Unfocuses one step in the focus stack. *) +let vernac_unfocus () = + Proof_global.simple_with_current_proof + (fun _ p -> Proof.unfocus command_focus p ()) + +(* Checks that a proof is fully unfocused. Raises an error if not. *) +let vernac_unfocused () = + let p = Proof_global.give_me_the_proof () in + if Proof.unfocused p then + str"The proof is indeed fully unfocused." + else + user_err Pp.(str "The proof is not fully unfocused.") + + +(* "{" focuses on the first goal, "n: {" focuses on the n-th goal + "}" unfocuses, provided that the proof of the goal has been completed. +*) +let subproof_kind = Proof.new_focus_kind () +let subproof_cond = Proof.done_cond subproof_kind + +let vernac_subproof gln = + Proof_global.simple_with_current_proof (fun _ p -> + match gln with + | None -> Proof.focus subproof_cond () 1 p + | Some (SelectNth n) -> Proof.focus subproof_cond () n p + | _ -> user_err ~hdr:"bracket_selector" + (str "Brackets only support the single numbered goal selector.")) + +let vernac_end_subproof () = + Proof_global.simple_with_current_proof (fun _ p -> + Proof.unfocus subproof_kind p ()) + +let vernac_bullet (bullet : Proof_bullet.t) = + Proof_global.simple_with_current_proof (fun _ p -> + Proof_bullet.put p bullet) + +let vernac_show = function + | ShowScript -> assert false (* Only the stm knows the script *) + | ShowGoal goalref -> + let proof = Proof_global.give_me_the_proof () in + begin match goalref with + | OpenSubgoals -> pr_open_subgoals ~proof + | NthGoal n -> pr_nth_open_subgoal ~proof n + | GoalId id -> pr_goal_by_id ~proof id + end + | ShowProof -> show_proof () + | ShowExistentials -> show_top_evars () + | ShowUniverses -> show_universes () + | ShowProofNames -> + pr_sequence Id.print (Proof_global.get_all_proof_names()) + | ShowIntros all -> show_intro all + | ShowMatch id -> show_match id + +let vernac_check_guard () = + let pts = Proof_global.give_me_the_proof () in + let pfterm = List.hd (Proof.partial_proof pts) in + let message = + try + let { Evd.it=gl ; sigma=sigma } = Proof.V82.top_goal pts in + Inductiveops.control_only_guard (Goal.V82.env sigma gl) sigma pfterm; + (str "The condition holds up to here") + with UserError(_,s) -> + (str ("Condition violated: ") ++s) + in message + +exception End_of_input + +(* XXX: This won't properly set the proof mode, as of today, it is + controlled by the STM. Thus, we would need access information from + the classifier. The proper fix is to move it to the STM, however, + the way the proof mode is set there makes the task non trivial + without a considerable amount of refactoring. + *) +let vernac_load interp fname = + if Proof_global.there_are_pending_proofs () then + CErrors.user_err Pp.(str "Load is not supported inside proofs."); + let interp x = + let proof_mode = Proof_global.get_default_proof_mode_name () [@ocaml.warning "-3"] in + Proof_global.activate_proof_mode proof_mode [@ocaml.warning "-3"]; + interp x in + let parse_sentence = Flags.with_option Flags.we_are_parsing + (fun po -> + match Pcoq.Gram.entry_parse Pcoq.main_entry po with + | Some x -> x + | None -> raise End_of_input) in + let fname = + Envars.expand_path_macros ~warn:(fun x -> Feedback.msg_warning (str x)) fname in + let fname = CUnix.make_suffix fname ".v" in + let input = + let longfname = Loadpath.locate_file fname in + let in_chan = open_utf8_file_in longfname in + Pcoq.Gram.parsable ~file:(Loc.InFile longfname) (Stream.of_channel in_chan) in + begin + try while true do interp (snd (parse_sentence input)) done + with End_of_input -> () + end; + (* If Load left a proof open, we fail too. *) + if Proof_global.there_are_pending_proofs () then + CErrors.user_err Pp.(str "Files processed by Load cannot leave open proofs.") + +(* "locality" is the prefix "Local" attribute, while the "local" component + * is the outdated/deprecated "Local" attribute of some vernacular commands + * still parsed as the obsolete_locality grammar entry for retrocompatibility. + * loc is the Loc.t of the vernacular command being interpreted. *) +let interp ?proof ~atts ~st c = + let open Vernacinterp in + vernac_pperr_endline (fun () -> str "interpreting: " ++ Ppvernac.pr_vernac_expr c); + match c with + + (* Loading a file requires access to the control interpreter *) + | VernacLoad _ -> assert false + + (* The STM should handle that, but LOAD bypasses the STM... *) + | VernacAbortAll -> CErrors.user_err (str "AbortAll cannot be used through the Load command") + | VernacRestart -> CErrors.user_err (str "Restart cannot be used through the Load command") + | VernacUndo _ -> CErrors.user_err (str "Undo cannot be used through the Load command") + | VernacUndoTo _ -> CErrors.user_err (str "UndoTo cannot be used through the Load command") + | VernacBacktrack _ -> CErrors.user_err (str "Backtrack cannot be used through the Load command") + + (* Toplevel control *) + | VernacToplevelControl e -> raise e + + (* Resetting *) + | VernacResetName _ -> anomaly (str "VernacResetName not handled by Stm.") + | VernacResetInitial -> anomaly (str "VernacResetInitial not handled by Stm.") + | VernacBack _ -> anomaly (str "VernacBack not handled by Stm.") + | VernacBackTo _ -> anomaly (str "VernacBackTo not handled by Stm.") + + (* This one is possible to handle here *) + | VernacAbort id -> CErrors.user_err (str "Abort cannot be used through the Load command") + + (* Syntax *) + | VernacSyntaxExtension (infix, sl) -> + vernac_syntax_extension atts infix sl + | VernacDelimiters (sc,lr) -> vernac_delimiters sc lr + | VernacBindScope (sc,rl) -> vernac_bind_scope sc rl + | VernacOpenCloseScope (b, s) -> vernac_open_close_scope ~atts (b,s) + | VernacArgumentsScope (qid,scl) -> vernac_arguments_scope ~atts qid scl + | VernacInfix (mv,qid,sc) -> vernac_infix ~atts mv qid sc + | VernacNotation (c,infpl,sc) -> + vernac_notation ~atts c infpl sc + | VernacNotationAddFormat(n,k,v) -> + Metasyntax.add_notation_extra_printing_rule n k v + + (* Gallina *) + | VernacDefinition ((discharge,kind),lid,d) -> + vernac_definition ~atts discharge kind lid d + | VernacStartTheoremProof (k,l) -> vernac_start_proof ~atts k l + | VernacEndProof e -> vernac_end_proof ?proof e + | VernacExactProof c -> vernac_exact_proof c + | VernacAssumption ((discharge,kind),nl,l) -> + vernac_assumption ~atts discharge kind l nl + | VernacInductive (cum, priv,finite,l) -> vernac_inductive ~atts cum priv finite l + | VernacFixpoint (discharge, l) -> vernac_fixpoint ~atts discharge l + | VernacCoFixpoint (discharge, l) -> vernac_cofixpoint ~atts discharge l + | VernacScheme l -> vernac_scheme l + | VernacCombinedScheme (id, l) -> vernac_combined_scheme id l + | VernacUniverse l -> vernac_universe ~atts l + | VernacConstraint l -> vernac_constraint ~atts l + + (* Modules *) + | VernacDeclareModule (export,lid,bl,mtyo) -> + vernac_declare_module export lid bl mtyo + | VernacDefineModule (export,lid,bl,mtys,mexprl) -> + vernac_define_module export lid bl mtys mexprl + | VernacDeclareModuleType (lid,bl,mtys,mtyo) -> + vernac_declare_module_type lid bl mtys mtyo + | VernacInclude in_asts -> + vernac_include in_asts + (* Gallina extensions *) + | VernacBeginSection lid -> vernac_begin_section lid + + | VernacEndSegment lid -> vernac_end_segment lid + + | VernacNameSectionHypSet (lid, set) -> vernac_name_sec_hyp lid set + + | VernacRequire (from, export, qidl) -> vernac_require from export qidl + | VernacImport (export,qidl) -> vernac_import export qidl + | VernacCanonical qid -> vernac_canonical qid + | VernacCoercion (r,s,t) -> vernac_coercion ~atts r s t + | VernacIdentityCoercion ({v=id},s,t) -> + vernac_identity_coercion ~atts id s t + + (* Type classes *) + | VernacInstance (abst, sup, inst, props, info) -> + vernac_instance ~atts abst sup inst props info + | VernacContext sup -> vernac_context ~atts sup + | VernacDeclareInstances insts -> vernac_declare_instances ~atts insts + | VernacDeclareClass id -> vernac_declare_class id + + (* Solving *) + | VernacSolveExistential (n,c) -> vernac_solve_existential n c + + (* Auxiliary file and library management *) + | VernacAddLoadPath (isrec,s,alias) -> vernac_add_loadpath isrec s alias + | VernacRemoveLoadPath s -> vernac_remove_loadpath s + | VernacAddMLPath (isrec,s) -> vernac_add_ml_path isrec s + | VernacDeclareMLModule l -> vernac_declare_ml_module ~atts l + | VernacChdir s -> vernac_chdir s + + (* State management *) + | VernacWriteState s -> vernac_write_state s + | VernacRestoreState s -> vernac_restore_state s + + (* Commands *) + | VernacCreateHintDb (dbname,b) -> vernac_create_hintdb ~atts dbname b + | VernacRemoveHints (dbnames,ids) -> vernac_remove_hints ~atts dbnames ids + | VernacHints (dbnames,hints) -> + vernac_hints ~atts dbnames hints + | VernacSyntacticDefinition (id,c,b) -> + vernac_syntactic_definition ~atts id c b + | VernacDeclareImplicits (qid,l) -> + vernac_declare_implicits ~atts qid l + | VernacArguments (qid, args, more_implicits, nargs, flags) -> + vernac_arguments ~atts qid args more_implicits nargs flags + | VernacReserve bl -> vernac_reserve bl + | VernacGeneralizable gen -> vernac_generalizable ~atts gen + | VernacSetOpacity qidl -> vernac_set_opacity ~atts qidl + | VernacSetStrategy l -> vernac_set_strategy ~atts l + | VernacSetOption (export, key,v) -> vernac_set_option ~atts export key v + | VernacUnsetOption (export, key) -> vernac_unset_option ~atts export key + | VernacRemoveOption (key,v) -> vernac_remove_option key v + | VernacAddOption (key,v) -> vernac_add_option key v + | VernacMemOption (key,v) -> vernac_mem_option key v + | VernacPrintOption key -> vernac_print_option key + | VernacCheckMayEval (r,g,c) -> + Feedback.msg_notice @@ vernac_check_may_eval ~atts r g c + | VernacDeclareReduction (s,r) -> vernac_declare_reduction ~atts s r + | VernacGlobalCheck c -> + Feedback.msg_notice @@ vernac_global_check c + | VernacPrint p -> + let sigma, env = Pfedit.get_current_context () in + Feedback.msg_notice @@ vernac_print ~atts env sigma p + | VernacSearch (s,g,r) -> vernac_search ~atts s g r + | VernacLocate l -> + Feedback.msg_notice @@ vernac_locate l + | VernacRegister (id, r) -> vernac_register id r + | VernacComments l -> Flags.if_verbose Feedback.msg_info (str "Comments ok\n") + + (* Proof management *) + | VernacFocus n -> vernac_focus n + | VernacUnfocus -> vernac_unfocus () + | VernacUnfocused -> + Feedback.msg_notice @@ vernac_unfocused () + | VernacBullet b -> vernac_bullet b + | VernacSubproof n -> vernac_subproof n + | VernacEndSubproof -> vernac_end_subproof () + | VernacShow s -> + Feedback.msg_notice @@ vernac_show s + | VernacCheckGuard -> + Feedback.msg_notice @@ vernac_check_guard () + | VernacProof (tac, using) -> + let using = Option.append using (Proof_using.get_default_proof_using ()) in + let tacs = if Option.is_empty tac then "tac:no" else "tac:yes" in + let usings = if Option.is_empty using then "using:no" else "using:yes" in + Aux_file.record_in_aux_at ?loc:atts.loc "VernacProof" (tacs^" "^usings); + Option.iter vernac_set_end_tac tac; + Option.iter vernac_set_used_variables using + | VernacProofMode mn -> Proof_global.set_proof_mode mn [@ocaml.warning "-3"] + + (* Extensions *) + | VernacExtend (opn,args) -> + (* XXX: Here we are returning the state! :) *) + let _st : Vernacstate.t = Vernacinterp.call ~atts opn args ~st in + () + +(* Vernaculars that take a locality flag *) +let check_vernac_supports_locality c l = + match l, c with + | None, _ -> () + | Some _, ( + VernacOpenCloseScope _ + | VernacSyntaxExtension _ | VernacInfix _ | VernacNotation _ + | VernacDefinition _ | VernacFixpoint _ | VernacCoFixpoint _ + | VernacAssumption _ | VernacStartTheoremProof _ + | VernacCoercion _ | VernacIdentityCoercion _ + | VernacInstance _ | VernacDeclareInstances _ + | VernacDeclareMLModule _ + | VernacCreateHintDb _ | VernacRemoveHints _ | VernacHints _ + | VernacSyntacticDefinition _ + | VernacArgumentsScope _ | VernacDeclareImplicits _ | VernacArguments _ + | VernacGeneralizable _ + | VernacSetOpacity _ | VernacSetStrategy _ + | VernacSetOption _ | VernacUnsetOption _ + | VernacDeclareReduction _ + | VernacExtend _ + | VernacInductive _) -> () + | Some _, _ -> user_err Pp.(str "This command does not support Locality") + +(* Vernaculars that take a polymorphism flag *) +let check_vernac_supports_polymorphism c p = + match p, c with + | None, _ -> () + | Some _, ( + VernacDefinition _ | VernacFixpoint _ | VernacCoFixpoint _ + | VernacAssumption _ | VernacInductive _ + | VernacStartTheoremProof _ + | VernacCoercion _ | VernacIdentityCoercion _ + | VernacInstance _ | VernacDeclareInstances _ + | VernacHints _ | VernacContext _ + | VernacExtend _ | VernacUniverse _ | VernacConstraint _) -> () + | Some _, _ -> user_err Pp.(str "This command does not support Polymorphism") + +(** A global default timeout, controlled by option "Set Default Timeout n". + Use "Unset Default Timeout" to deactivate it (or set it to 0). *) + +let default_timeout = ref None + +let _ = + Goptions.declare_int_option + { Goptions.optdepr = false; + Goptions.optname = "the default timeout"; + Goptions.optkey = ["Default";"Timeout"]; + Goptions.optread = (fun () -> !default_timeout); + Goptions.optwrite = ((:=) default_timeout) } + +(** When interpreting a command, the current timeout is initially + the default one, but may be modified locally by a Timeout command. *) + +let current_timeout = ref None + +let vernac_timeout f = + match !current_timeout, !default_timeout with + | Some n, _ | None, Some n -> + let f () = f (); current_timeout := None in + Control.timeout n f () Timeout + | None, None -> f () + +let restore_timeout () = current_timeout := None + +let locate_if_not_already ?loc (e, info) = + match Loc.get_loc info with + | None -> (e, Option.cata (Loc.add_loc info) info loc) + | Some l -> (e, info) + +exception HasNotFailed +exception HasFailed of Pp.t + +(* XXX STATE: this type hints that restoring the state should be the + caller's responsibility *) +let with_fail st b f = + if not b + then f () + else begin try + (* If the command actually works, ignore its effects on the state. + * Note that error has to be printed in the right state, hence + * within the purified function *) + try f (); raise HasNotFailed + with + | HasNotFailed as e -> raise e + | e -> + let e = CErrors.push e in + raise (HasFailed (CErrors.iprint + (ExplainErr.process_vernac_interp_error ~allow_uncaught:false e))) + with e when CErrors.noncritical e -> + (* Restore the previous state XXX Careful here with the cache! *) + Vernacstate.invalidate_cache (); + Vernacstate.unfreeze_interp_state st; + let (e, _) = CErrors.push e in + match e with + | HasNotFailed -> + user_err ~hdr:"Fail" (str "The command has not failed!") + | HasFailed msg -> + if not !Flags.quiet || !Flags.test_mode || !Flags.ide_slave then Feedback.msg_info + (str "The command has indeed failed with message:" ++ fnl () ++ msg) + | _ -> assert false + end + +let interp ?(verbosely=true) ?proof ~st {CAst.loc;v=c} = + let orig_univ_poly = Flags.is_universe_polymorphism () in + let orig_program_mode = Flags.is_program_mode () in + let flags f atts = + List.fold_left + (fun (polymorphism, atts) f -> + match f with + | VernacProgram when not atts.program -> + (polymorphism, { atts with program = true }) + | VernacProgram -> + user_err Pp.(str "Program mode specified twice") + | VernacPolymorphic b when polymorphism = None -> + (Some b, atts) + | VernacPolymorphic _ -> + user_err Pp.(str "Polymorphism specified twice") + | VernacLocal b when Option.is_empty atts.locality -> + (polymorphism, { atts with locality = Some b }) + | VernacLocal _ -> + user_err Pp.(str "Locality specified twice") + ) + (None, atts) + f + in + let rec control = function + | VernacExpr (f, v) -> + let (polymorphism, atts) = flags f { loc; locality = None; polymorphic = false; program = orig_program_mode; } in + aux ~polymorphism ~atts v + | VernacFail v -> with_fail st true (fun () -> control v) + | VernacTimeout (n,v) -> + current_timeout := Some n; + control v + | VernacRedirect (s, {v}) -> + Topfmt.with_output_to_file s control v + | VernacTime (batch, {v}) -> + System.with_time ~batch control v; + + and aux ~polymorphism ~atts : _ -> unit = + function + + | VernacLoad (_,fname) -> vernac_load control fname + + | c -> + check_vernac_supports_locality c atts.locality; + check_vernac_supports_polymorphism c polymorphism; + let polymorphic = Option.default (Flags.is_universe_polymorphism ()) polymorphism in + Flags.make_universe_polymorphism polymorphic; + Obligations.set_program_mode atts.program; + try + vernac_timeout begin fun () -> + let atts = { atts with polymorphic } in + if verbosely + then Flags.verbosely (interp ?proof ~atts ~st) c + else Flags.silently (interp ?proof ~atts ~st) c; + (* If the command is `(Un)Set Program Mode` or `(Un)Set Universe Polymorphism`, + we should not restore the previous state of the flag... *) + if orig_program_mode || not !Flags.program_mode || atts.program then + Flags.program_mode := orig_program_mode; + if (Flags.is_universe_polymorphism() = polymorphic) then + Flags.make_universe_polymorphism orig_univ_poly; + end + with + | reraise when + (match reraise with + | Timeout -> true + | e -> CErrors.noncritical e) + -> + let e = CErrors.push reraise in + let e = locate_if_not_already ?loc e in + let () = restore_timeout () in + Flags.make_universe_polymorphism orig_univ_poly; + Flags.program_mode := orig_program_mode; + iraise e + in + if verbosely + then Flags.verbosely control c + else control c + +(* Be careful with the cache here in case of an exception. *) +let interp ?verbosely ?proof ~st cmd = + Vernacstate.unfreeze_interp_state st; + try + interp ?verbosely ?proof ~st cmd; + Vernacstate.freeze_interp_state `No + with exn -> + let exn = CErrors.push exn in + Vernacstate.invalidate_cache (); + iraise exn diff --git a/vernac/vernacentries.mli b/vernac/vernacentries.mli new file mode 100644 index 00000000..f6199e82 --- /dev/null +++ b/vernac/vernacentries.mli @@ -0,0 +1,42 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* unit + +(** Vernacular entries *) +val vernac_require : + Libnames.reference option -> bool option -> Libnames.reference list -> unit + +(** The main interpretation function of vernacular expressions *) +val interp : + ?verbosely:bool -> + ?proof:Proof_global.closed_proof -> + st:Vernacstate.t -> Vernacexpr.vernac_control CAst.t -> Vernacstate.t + +(** Prepare a "match" template for a given inductive type. + For each branch of the match, we list the constructor name + followed by enough pattern variables. + [Not_found] is raised if the given string isn't the qualid of + a known inductive type. *) + +val make_cases : string -> string list list + +(* XXX STATE: this type hints that restoring the state should be the + caller's responsibility *) +val with_fail : Vernacstate.t -> bool -> (unit -> unit) -> unit + +val command_focus : unit Proof.focus_kind + +val interp_redexp_hook : (Environ.env -> Evd.evar_map -> Genredexpr.raw_red_expr -> + Evd.evar_map * Redexpr.red_expr) Hook.t + +val universe_polymorphism_option_name : string list diff --git a/vernac/vernacinterp.ml b/vernac/vernacinterp.ml new file mode 100644 index 00000000..1f2d2e4b --- /dev/null +++ b/vernac/vernacinterp.ml @@ -0,0 +1,87 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(*