diff options
| -rw-r--r-- | interp/constrintern.ml | 9 | ||||
| -rw-r--r-- | interp/constrintern.mli | 4 | ||||
| -rw-r--r-- | interp/implicit_quantifiers.ml | 86 | ||||
| -rw-r--r-- | interp/implicit_quantifiers.mli | 21 | ||||
| -rw-r--r-- | pretyping/typeclasses_errors.ml | 4 | ||||
| -rw-r--r-- | pretyping/typeclasses_errors.mli | 2 | ||||
| -rw-r--r-- | test-suite/bugs/closed/bug_10225.v | 7 | ||||
| -rw-r--r-- | vernac/classes.ml | 13 | ||||
| -rw-r--r-- | vernac/classes.mli | 6 | ||||
| -rw-r--r-- | vernac/explainErr.ml | 2 | ||||
| -rw-r--r-- | vernac/himsg.ml | 13 | ||||
| -rw-r--r-- | vernac/himsg.mli | 2 |
12 files changed, 36 insertions, 133 deletions
diff --git a/interp/constrintern.ml b/interp/constrintern.ml index e55f66e856..ff0c06e705 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -392,9 +392,8 @@ let intern_generalized_binder ?(global_level=false) intern_type ntnvars env {loc;v=na} b' t ty = let ids = (match na with Anonymous -> fun x -> x | Name na -> Id.Set.add na) env.ids in let ty, ids' = - if t then ty, ids else - Implicit_quantifiers.implicit_application ids - Implicit_quantifiers.combine_params_freevar ty + if t then ty, ids + else Implicit_quantifiers.implicit_application ids ty in let ty' = intern_type {env with ids = ids; unb = true} ty in let fvs = Implicit_quantifiers.generalizable_vars_of_glob_constr ~bound:ids ~allowed:ids' ty' in @@ -1300,7 +1299,7 @@ let find_pattern_variable qid = if qualid_is_ident qid then qualid_basename qid else raise (InternalizationError(qid.CAst.loc,NotAConstructor qid)) -let check_duplicate loc fields = +let check_duplicate ?loc fields = let eq (ref1, _) (ref2, _) = qualid_eq ref1 ref2 in let dups = List.duplicates eq fields in match dups with @@ -1345,7 +1344,7 @@ let sort_fields ~complete loc fields completer = try Nametab.shortest_qualid_of_global ?loc Id.Set.empty global_record_id with Not_found -> anomaly (str "Environment corruption for records.") in - let () = check_duplicate loc fields in + let () = check_duplicate ?loc fields in let (end_index, (* one past the last field index *) first_field_index, (* index of the first field of the record *) proj_list) (* list of projections *) diff --git a/interp/constrintern.mli b/interp/constrintern.mli index 4bf8ee9429..450daea75c 100644 --- a/interp/constrintern.mli +++ b/interp/constrintern.mli @@ -189,3 +189,7 @@ val for_grammar : ('a -> 'b) -> 'a -> 'b (** Placeholder for global option, should be moved to a parameter *) val get_asymmetric_patterns : unit -> bool + +val check_duplicate : ?loc:Loc.t -> (qualid * constr_expr) list -> unit +(** Check that a list of record field definitions doesn't contain + duplicates. *) diff --git a/interp/implicit_quantifiers.ml b/interp/implicit_quantifiers.ml index bab9024415..d7bae6b3fd 100644 --- a/interp/implicit_quantifiers.ml +++ b/interp/implicit_quantifiers.ml @@ -23,9 +23,6 @@ open Libobject open Nameops open Context.Rel.Declaration -exception MismatchedContextInstance of Environ.env * Typeclasses_errors.contexts * constr_expr list * Constr.rel_context (* found, expected *) -let mismatched_ctx_inst_err env c n m = raise (MismatchedContextInstance (env, c, n, m)) - module RelDecl = Context.Rel.Declaration (*i*) @@ -66,9 +63,6 @@ let declare_generalizable ~local gen = let find_generalizable_ident id = Id.Pred.mem (root_of_id id) !generalizable_table -let ids_of_list l = - List.fold_right Id.Set.add l Id.Set.empty - let is_global id = try ignore (Nametab.locate_extended (qualid_of_ident id)); true with Not_found -> false @@ -105,26 +99,6 @@ let free_vars_of_constr_expr c ?(bound=Id.Set.empty) l = | _ -> Constrexpr_ops.fold_constr_expr_with_binders (fun a l -> Id.Set.add a l) aux bdvars l c in aux bound l c -let ids_of_names l = - List.fold_left (fun acc x -> match x.CAst.v with Name na -> na :: acc | Anonymous -> acc) [] l - -let free_vars_of_binders ?(bound=Id.Set.empty) l (binders : local_binder_expr list) = - let rec aux bdvars l c = match c with - ((CLocalAssum (n, _, c)) :: tl) -> - let bound = ids_of_names n in - let l' = free_vars_of_constr_expr c ~bound:bdvars l in - aux (Id.Set.union (ids_of_list bound) bdvars) l' tl - - | ((CLocalDef (n, c, t)) :: tl) -> - let bound = match n.CAst.v with Anonymous -> [] | Name n -> [n] in - let l' = free_vars_of_constr_expr c ~bound:bdvars l in - let l'' = Option.fold_left (fun l t -> free_vars_of_constr_expr t ~bound:bdvars l) l' t in - aux (Id.Set.union (ids_of_list bound) bdvars) l'' tl - - | CLocalPattern _ :: tl -> assert false - | [] -> bdvars, l - in aux bound l binders - let generalizable_vars_of_glob_constr ?(bound=Id.Set.empty) ?(allowed=Id.Set.empty) = let rec vars bound vs c = match DAst.get c with | GVar id -> @@ -149,7 +123,7 @@ let next_name_away_from na avoid = | Anonymous -> make_fresh avoid (Global.env ()) (Id.of_string "anon") | Name id -> make_fresh avoid (Global.env ()) id -let combine_params avoid fn applied needed = +let combine_params avoid applied needed = let named, applied = List.partition (function @@ -167,47 +141,30 @@ let combine_params avoid fn applied needed = (fun x -> match x with (t, Some {CAst.loc;v=ExplByName id}) -> id, t | _ -> assert false) named in - let is_unset (_, decl) = match decl with - | LocalAssum _ -> true - | LocalDef _ -> false - in - let needed = List.filter is_unset needed in let rec aux ids avoid app need = match app, need with - [], [] -> List.rev ids, avoid - | app, (_, (LocalAssum ({binder_name=Name id}, _) | LocalDef ({binder_name=Name id}, _, _))) :: need when Id.List.mem_assoc id named -> + | _, (_, LocalDef _) :: need -> aux ids avoid app need + + | [], [] -> List.rev ids, avoid + + | app, (_, (LocalAssum ({binder_name=Name id}, _))) :: need when Id.List.mem_assoc id named -> aux (Id.List.assoc id named :: ids) avoid app need - | (x, None) :: app, (None, (LocalAssum ({binder_name=Name id}, _) | LocalDef ({binder_name=Name id}, _, _))) :: need -> + | (x, None) :: app, (None, (LocalAssum ({binder_name=Name id}, _))) :: need -> aux (x :: ids) avoid app need - | _, (Some cl, _ as d) :: need -> - let t', avoid' = fn avoid d in - aux (t' :: ids) avoid' app need - | x :: app, (None, _) :: need -> aux (fst x :: ids) avoid app need - | [], (None, _ as decl) :: need -> - let t', avoid' = fn avoid decl in - aux (t' :: ids) avoid' app need + | _, (Some _, decl) :: need | [], (None, decl) :: need -> + let id' = next_name_away_from (RelDecl.get_name decl) avoid in + let t' = CAst.make @@ CRef (qualid_of_ident id',None) in + aux (t' :: ids) (Id.Set.add id' avoid) app need | (x,_) :: _, [] -> user_err ?loc:(Constrexpr_ops.constr_loc x) (str "Typeclass does not expect more arguments") - in aux [] avoid applied needed - -let combine_params_freevar avoid (_, decl) = - let id' = next_name_away_from (RelDecl.get_name decl) avoid in - (CAst.make @@ CRef (qualid_of_ident id',None), Id.Set.add id' avoid) - -let destClassApp cl = - let open CAst in - let loc = cl.loc in - match cl.v with - | CApp ((None, { v = CRef (ref, inst) }), l) -> CAst.make ?loc (ref, List.map fst l, inst) - | CAppExpl ((None, ref, inst), l) -> CAst.make ?loc (ref, l, inst) - | CRef (ref, inst) -> CAst.make ?loc:cl.loc (ref, [], inst) - | _ -> raise Not_found + in + aux [] avoid applied needed let destClassAppExpl cl = let open CAst in @@ -217,7 +174,7 @@ let destClassAppExpl cl = | CRef (ref, inst) -> CAst.make ?loc:cl.loc (ref, [], inst) | _ -> raise Not_found -let implicit_application env ?(allow_partial=true) f ty = +let implicit_application env ty = let is_class = try let ({CAst.v=(qid, _, _)} as clapp) = destClassAppExpl ty in @@ -230,24 +187,13 @@ let implicit_application env ?(allow_partial=true) f ty = match is_class with | None -> ty, env | Some ({CAst.loc;v=(id, par, inst)}, gr) -> - let avoid = Id.Set.union env (ids_of_list (free_vars_of_constr_expr ty ~bound:env [])) in + let avoid = Id.Set.union env (Id.Set.of_list (free_vars_of_constr_expr ty ~bound:env [])) in let env = Global.env () in let sigma = Evd.from_env env in let c = class_info env sigma gr in let (ci, rd) = c.cl_context in - if not allow_partial then - begin - let opt_succ x n = match x with - | None -> succ n - | Some _ -> n - in - let applen = List.fold_left (fun acc (x, y) -> opt_succ y acc) 0 par in - let needlen = List.fold_left (fun acc x -> opt_succ x acc) 0 ci in - if not (Int.equal needlen applen) then - mismatched_ctx_inst_err (Global.env ()) Typeclasses_errors.Parameters (List.map fst par) rd - end; let pars = List.rev (List.combine ci rd) in - let args, avoid = combine_params avoid f par pars in + let args, avoid = combine_params avoid par pars in CAst.make ?loc @@ CAppExpl ((None, id, inst), args), avoid let warn_ignoring_implicit_status = diff --git a/interp/implicit_quantifiers.mli b/interp/implicit_quantifiers.mli index 437fef1753..8536a2dee0 100644 --- a/interp/implicit_quantifiers.mli +++ b/interp/implicit_quantifiers.mli @@ -11,22 +11,14 @@ open Names open Glob_term open Constrexpr -open Libnames val declare_generalizable : local:bool -> lident list option -> unit -val ids_of_list : Id.t list -> Id.Set.t -val destClassApp : constr_expr -> (qualid * constr_expr list * instance_expr option) CAst.t -val destClassAppExpl : constr_expr -> (qualid * (constr_expr * explicitation CAst.t option) list * instance_expr option) CAst.t - (** Fragile, should be used only for construction a set of identifiers to avoid *) val free_vars_of_constr_expr : constr_expr -> ?bound:Id.Set.t -> Id.t list -> Id.t list -val free_vars_of_binders : - ?bound:Id.Set.t -> Id.t list -> local_binder_expr list -> Id.Set.t * Id.t list - (** Returns the generalizable free ids in left-to-right order with the location of their first occurrence *) @@ -37,15 +29,4 @@ val make_fresh : Id.Set.t -> Environ.env -> Id.t -> Id.t val implicits_of_glob_constr : ?with_products:bool -> Glob_term.glob_constr -> Impargs.manual_implicits -val combine_params_freevar : - Id.Set.t -> GlobRef.t option * Constr.rel_declaration -> - Constrexpr.constr_expr * Id.Set.t - -val implicit_application : Id.Set.t -> ?allow_partial:bool -> - (Id.Set.t -> GlobRef.t option * Constr.rel_declaration -> - Constrexpr.constr_expr * Id.Set.t) -> - constr_expr -> constr_expr * Id.Set.t - -(* Should be likely located elsewhere *) -exception MismatchedContextInstance of Environ.env * Typeclasses_errors.contexts * constr_expr list * Constr.rel_context (* found, expected *) -val mismatched_ctx_inst_err : Environ.env -> Typeclasses_errors.contexts -> constr_expr list -> Constr.rel_context -> 'a +val implicit_application : Id.Set.t -> constr_expr -> constr_expr * Id.Set.t diff --git a/pretyping/typeclasses_errors.ml b/pretyping/typeclasses_errors.ml index af5b3016c9..90ecab1608 100644 --- a/pretyping/typeclasses_errors.ml +++ b/pretyping/typeclasses_errors.ml @@ -8,13 +8,9 @@ (* * (see LICENSE file for the text of the license) *) (************************************************************************) -(*i*) open Names open EConstr open Environ -(*i*) - -type contexts = Parameters | Properties type typeclass_error = | NotAClass of constr diff --git a/pretyping/typeclasses_errors.mli b/pretyping/typeclasses_errors.mli index fd75781ed5..21bb5f8443 100644 --- a/pretyping/typeclasses_errors.mli +++ b/pretyping/typeclasses_errors.mli @@ -12,8 +12,6 @@ open Names open EConstr open Environ -type contexts = Parameters | Properties - type typeclass_error = | NotAClass of constr | UnboundMethod of GlobRef.t * lident (** Class name, method *) diff --git a/test-suite/bugs/closed/bug_10225.v b/test-suite/bugs/closed/bug_10225.v new file mode 100644 index 0000000000..6d6bb39a65 --- /dev/null +++ b/test-suite/bugs/closed/bug_10225.v @@ -0,0 +1,7 @@ + +Class Bar := {}. +Instance bb : Bar := {}. + +Class Foo := { xx : Bar; foo : nat }. + +Fail Instance bar : Foo := { foo := 1 + 1; foo := 2 + 2 }. diff --git a/vernac/classes.ml b/vernac/classes.ml index bd5a211f1d..178387dd17 100644 --- a/vernac/classes.ml +++ b/vernac/classes.ml @@ -281,9 +281,6 @@ let existing_instance glob g info = ~hdr:"declare_instance" (Pp.str "Constant does not build instances of a declared type class.") -let mismatched_params env n m = Implicit_quantifiers.mismatched_ctx_inst_err env Parameters n m -let mismatched_props env n m = Implicit_quantifiers.mismatched_ctx_inst_err env Properties n m - (* Declare everything in the parameters as implicit, and the class instance as well *) let type_ctx_instance ~program_mode env sigma ctx inst subst = @@ -479,9 +476,8 @@ let do_instance_interactive env sigma ?hook ~tac ~global ~poly cty k u ctx ctx' let do_instance env env' sigma ?hook ~global ~poly cty k u ctx ctx' pri decl imps subst id props = let term, termtype, sigma = match props with - | (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; + | (true, { CAst.v = CRecord fs; loc }) -> + check_duplicate ?loc fs; let subst, sigma = do_instance_type_ctx_instance fs k env' ctx' sigma ~program_mode:false subst in let term, termtype = do_instance_subst_constructor_and_ty subst k u (ctx' @ ctx) in @@ -502,9 +498,8 @@ let do_instance env env' sigma ?hook ~global ~poly cty k u ctx ctx' pri decl imp let do_instance_program env env' sigma ?hook ~global ~poly cty k u ctx ctx' pri decl imps subst id opt_props = let term, termtype, sigma = match opt_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 (true, { CAst.v = CRecord fs; loc }) -> + check_duplicate ?loc fs; let subst, sigma = do_instance_type_ctx_instance fs k env' ctx' sigma ~program_mode:true subst in let term, termtype = diff --git a/vernac/classes.mli b/vernac/classes.mli index ace9096469..36c2d22c1d 100644 --- a/vernac/classes.mli +++ b/vernac/classes.mli @@ -14,12 +14,6 @@ open Constrexpr open Typeclasses open Libnames -(** Errors *) - -val mismatched_params : env -> constr_expr list -> Constr.rel_context -> 'a - -val mismatched_props : env -> constr_expr list -> Constr.rel_context -> 'a - (** Instance declaration *) val declare_instance : ?warn:bool -> env -> Evd.evar_map -> diff --git a/vernac/explainErr.ml b/vernac/explainErr.ml index 2bc95dbfcd..9733a751a8 100644 --- a/vernac/explainErr.ml +++ b/vernac/explainErr.ml @@ -66,8 +66,6 @@ let process_vernac_interp_error exn = match fst exn with wrap_vernac_error exn (Himsg.explain_prim_token_notation_error kind ctx sigma te) | Typeclasses_errors.TypeClassError(env, sigma, te) -> wrap_vernac_error exn (Himsg.explain_typeclass_error env sigma te) - | Implicit_quantifiers.MismatchedContextInstance(e,c,l,x) -> - wrap_vernac_error exn (Himsg.explain_mismatched_contexts e c l x) | InductiveError e -> wrap_vernac_error exn (Himsg.explain_inductive_error e) | Modops.ModuleTypingError e -> diff --git a/vernac/himsg.ml b/vernac/himsg.ml index a1f7835cbe..d10b704a48 100644 --- a/vernac/himsg.ml +++ b/vernac/himsg.ml @@ -1095,19 +1095,6 @@ let explain_unbound_method env sigma cid { CAst.v = id } = str "Unbound method name " ++ Id.print (id) ++ spc () ++ str"of class" ++ spc () ++ pr_global cid ++ str "." -let pr_constr_exprs env sigma exprs = - hv 0 (List.fold_right - (fun d pps -> ws 2 ++ Ppconstr.pr_constr_expr env sigma d ++ pps) - exprs (mt ())) - -let explain_mismatched_contexts env c i j = - let sigma = Evd.from_env env in - let pm, pn = with_diffs (pr_rel_context env sigma j) (pr_constr_exprs env sigma i) in - str"Mismatched contexts while declaring instance: " ++ brk (1,1) ++ - hov 1 (str"Expected:" ++ brk (1, 1) ++ pm) ++ - fnl () ++ brk (1,1) ++ - hov 1 (str"Found:" ++ brk (1, 1) ++ pn) - let explain_typeclass_error env sigma = function | NotAClass c -> explain_not_a_class env sigma c | UnboundMethod (cid, id) -> explain_unbound_method env sigma cid id diff --git a/vernac/himsg.mli b/vernac/himsg.mli index d1c1c092e3..09384cb859 100644 --- a/vernac/himsg.mli +++ b/vernac/himsg.mli @@ -24,8 +24,6 @@ val explain_pretype_error : env -> Evd.evar_map -> pretype_error -> Pp.t val explain_inductive_error : inductive_error -> Pp.t -val explain_mismatched_contexts : env -> contexts -> Constrexpr.constr_expr list -> Constr.rel_context -> Pp.t - val explain_typeclass_error : env -> Evd.evar_map -> typeclass_error -> Pp.t val explain_recursion_scheme_error : env -> recursion_scheme_error -> Pp.t |
