Skip to content

Commit

Permalink
Names: shortcuts for building {kn, constant, mind} with empty sections
Browse files Browse the repository at this point in the history
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@16249 85f007b7-540e-0410-9357-904b9bb8a0f7
  • Loading branch information
letouzey committed Feb 26, 2013
1 parent 2b09b02 commit 7dd28b4
Show file tree
Hide file tree
Showing 19 changed files with 61 additions and 55 deletions.
4 changes: 2 additions & 2 deletions checker/mod_checking.ml
Original file line number Diff line number Diff line change
Expand Up @@ -258,10 +258,10 @@ and check_module env mp mb =

and check_structure_field env mp lab res = function
| SFBconst cb ->
let c = make_con mp DirPath.empty lab in
let c = Constant.make2 mp lab in
check_constant_declaration env c cb
| SFBmind mib ->
let kn = make_mind mp DirPath.empty lab in
let kn = MutInd.make2 mp lab in
let kn = mind_of_delta res kn in
Indtypes.check_inductive env kn mib
| SFBmodule msb ->
Expand Down
8 changes: 4 additions & 4 deletions checker/modops.ml
Original file line number Diff line number Diff line change
Expand Up @@ -67,9 +67,9 @@ let module_body_of_type mp mtb =

let rec add_signature mp sign resolver env =
let add_one env (l,elem) =
let kn = make_kn mp DirPath.empty l in
let con = constant_of_kn kn in
let mind = mind_of_delta resolver (mind_of_kn kn) in
let kn = KerName.make2 mp l in
let con = Constant.make1 kn in
let mind = mind_of_delta resolver (MutInd.make1 kn) in
match elem with
| SFBconst cb ->
(* let con = constant_of_delta resolver con in*)
Expand Down Expand Up @@ -97,7 +97,7 @@ let strengthen_const mp_from l cb resolver =
match cb.const_body with
| Def _ -> cb
| _ ->
let con = make_con mp_from DirPath.empty l in
let con = Constant.make2 mp_from l in
(* let con = constant_of_delta resolver con in*)
{ cb with const_body = Def (Declarations.from_val (Const con)) }

Expand Down
4 changes: 2 additions & 2 deletions checker/subtyping.ml
Original file line number Diff line number Diff line change
Expand Up @@ -35,7 +35,7 @@ type namedmodule =
constructors *)

let add_mib_nameobjects mp l mib map =
let ind = make_mind mp DirPath.empty l in
let ind = MutInd.make2 mp l in
let add_mip_nameobjects j oib map =
let ip = (ind,j) in
let map =
Expand Down Expand Up @@ -83,7 +83,7 @@ let check_conv_error error f env a1 a2 =

(* for now we do not allow reorderings *)
let check_inductive env mp1 l info1 mib2 spec2 subst1 subst2=
let kn = make_mind mp1 DirPath.empty l in
let kn = MutInd.make2 mp1 l in
let error () = error_not_match l spec2 in
let check_conv f = check_conv_error error f in
let mib1 =
Expand Down
12 changes: 6 additions & 6 deletions kernel/mod_subst.ml
Original file line number Diff line number Diff line change
Expand Up @@ -190,16 +190,16 @@ let kn_of_deltas resolve1 resolve2 kn =
if kn' == kn then kn_of_delta resolve2 kn else kn'

let constant_of_delta_kn resolve kn =
Constant.make2 kn (kn_of_delta resolve kn)
Constant.make kn (kn_of_delta resolve kn)

let constant_of_deltas_kn resolve1 resolve2 kn =
Constant.make2 kn (kn_of_deltas resolve1 resolve2 kn)
Constant.make kn (kn_of_deltas resolve1 resolve2 kn)

let mind_of_delta_kn resolve kn =
MutInd.make2 kn (kn_of_delta resolve kn)
MutInd.make kn (kn_of_delta resolve kn)

let mind_of_deltas_kn resolve1 resolve2 kn =
MutInd.make2 kn (kn_of_deltas resolve1 resolve2 kn)
MutInd.make kn (kn_of_deltas resolve1 resolve2 kn)

let inline_of_delta inline resolver =
match inline with
Expand Down Expand Up @@ -288,7 +288,7 @@ let subst_ind sub mind =
let knc' =
progress (kn_of_delta resolve) (if user then knu else knc) ~orelse:knc
in
MutInd.make2 knu knc'
MutInd.make knu knc'
with No_subst -> mind

let subst_con0 sub cst =
Expand All @@ -305,7 +305,7 @@ let subst_con0 sub cst =
let knc' =
progress (kn_of_delta resolve) (if user then knu else knc) ~orelse:knc
in
let cst' = Constant.make2 knu knc' in
let cst' = Constant.make knu knc' in
cst', mkConst cst'

let subst_con sub cst =
Expand Down
12 changes: 6 additions & 6 deletions kernel/modops.ml
Original file line number Diff line number Diff line change
Expand Up @@ -259,7 +259,7 @@ let add_retroknowledge mp =

let rec add_signature mp sign resolver env =
let add_one env (l,elem) =
let kn = make_kn mp DirPath.empty l in
let kn = KerName.make2 mp l in
match elem with
| SFBconst cb ->
Environ.add_constant (constant_of_delta_kn resolver kn) cb env
Expand All @@ -284,7 +284,7 @@ let strengthen_const mp_from l cb resolver =
match cb.const_body with
| Def _ -> cb
| _ ->
let kn = make_kn mp_from DirPath.empty l in
let kn = KerName.make2 mp_from l in
let con = constant_of_delta_kn resolver kn in
{ cb with
const_body = Def (Declarations.from_val (mkConst con));
Expand Down Expand Up @@ -429,8 +429,8 @@ and strengthen_and_subst_struct
(* If we are performing an inclusion we need to add
the fact that the constant mp_to.l is \Delta-equivalent
to resolver(mp_from.l) *)
let kn_from = make_kn mp_from DirPath.empty l in
let kn_to = make_kn mp_to DirPath.empty l in
let kn_from = KerName.make2 mp_from l in
let kn_to = KerName.make2 mp_to l in
let old_name = kn_of_delta resolver kn_from in
(add_kn_delta_resolver kn_to old_name resolve_out),
item'::rest'
Expand All @@ -446,8 +446,8 @@ and strengthen_and_subst_struct
strengthen_and_subst_struct rest subst
mp_alias mp_from mp_to alias incl resolver in
if incl then
let kn_from = make_kn mp_from DirPath.empty l in
let kn_to = make_kn mp_to DirPath.empty l in
let kn_from = KerName.make2 mp_from l in
let kn_to = KerName.make2 mp_to l in
let old_name = kn_of_delta resolver kn_from in
(add_kn_delta_resolver kn_to old_name resolve_out),
item'::rest'
Expand Down
8 changes: 5 additions & 3 deletions kernel/names.ml
Original file line number Diff line number Diff line change
Expand Up @@ -323,6 +323,8 @@ module KerName = struct
let make mp dir l = (mp,dir,l)
let repr kn = kn

let make2 mp l = (mp,DirPath.empty,l)

let modpath (mp,_,_) = mp
let label (_,_,l) = l

Expand Down Expand Up @@ -409,7 +411,7 @@ module KerPair = struct
let make knu knc = if knu == knc then Same knc else Dual (knu,knc)

let make1 = same
let make2 = make
let make2 mp l = same (KerName.make2 mp l)
let make3 mp dir l = same (KerName.make mp dir l)
let repr3 kp = KerName.repr (user kp)
let label kp = KerName.label (user kp)
Expand Down Expand Up @@ -701,7 +703,7 @@ let kn_equal = KerName.equal
type constant = Constant.t

let constant_of_kn = Constant.make1
let constant_of_kn_equiv = Constant.make2
let constant_of_kn_equiv = Constant.make
let make_con = Constant.make3
let repr_con = Constant.repr3
let canonical_con = Constant.canonical
Expand All @@ -721,7 +723,7 @@ let con_with_label = Constant.change_label

type mutual_inductive = MutInd.t
let mind_of_kn = MutInd.make1
let mind_of_kn_equiv = MutInd.make2
let mind_of_kn_equiv = MutInd.make
let make_mind = MutInd.make3
let canonical_mind = MutInd.canonical
let user_mind = MutInd.user
Expand Down
17 changes: 12 additions & 5 deletions kernel/names.mli
Original file line number Diff line number Diff line change
Expand Up @@ -207,6 +207,7 @@ sig

(** Constructor and destructor *)
val make : ModPath.t -> DirPath.t -> Label.t -> t
val make2 : ModPath.t -> Label.t -> t
val repr : t -> ModPath.t * DirPath.t * Label.t

(** Projections *)
Expand Down Expand Up @@ -235,11 +236,14 @@ sig

(** Constructors *)

val make2 : KerName.t -> KerName.t -> t
val make : KerName.t -> KerName.t -> t
(** Builds a constant name from a user and a canonical kernel name. *)

val make1 : KerName.t -> t
(** Special case of [make2] where the user name is canonical. *)
(** Special case of [make] where the user name is canonical. *)

val make2 : ModPath.t -> Label.t -> t
(** Shortcut for [(make1 (KerName.make2 ...))] *)

val make3 : ModPath.t -> DirPath.t -> Label.t -> t
(** Shortcut for [(make1 (KerName.make ...))] *)
Expand Down Expand Up @@ -301,11 +305,14 @@ sig

(** Constructors *)

val make2 : KerName.t -> KerName.t -> t
val make : KerName.t -> KerName.t -> t
(** Builds a mutual inductive name from a user and a canonical kernel name. *)

val make1 : KerName.t -> t
(** Special case of [make2] where the user name is canonical. *)
(** Special case of [make] where the user name is canonical. *)

val make2 : ModPath.t -> Label.t -> t
(** Shortcut for [(make1 (KerName.make2 ...))] *)

val make3 : ModPath.t -> DirPath.t -> Label.t -> t
(** Shortcut for [(make1 (KerName.make ...))] *)
Expand Down Expand Up @@ -595,7 +602,7 @@ type constant = Constant.t
(** @deprecated Alias type *)

val constant_of_kn_equiv : KerName.t -> KerName.t -> constant
(** @deprecated Same as [Constant.make2] *)
(** @deprecated Same as [Constant.make] *)

val constant_of_kn : KerName.t -> constant
(** @deprecated Same as [Constant.make1] *)
Expand Down
2 changes: 1 addition & 1 deletion kernel/nativecode.ml
Original file line number Diff line number Diff line change
Expand Up @@ -1440,7 +1440,7 @@ let compile_constant_field env prefix con (code, symb, (mupds, cupds)) cb =
gl@code, !symbols_list, (mupds, cupds)

let compile_mind_field prefix mp l (code, symb, (mupds, cupds)) mb =
let mind = make_mind mp empty_dirpath l in
let mind = MutInd.make2 mp l in
reset_symbols_list symb;
let code, upd = compile_mind prefix mb mind code in
let mupds = Mindmap_env.add mind upd mupds in
Expand Down
6 changes: 2 additions & 4 deletions kernel/safe_typing.ml
Original file line number Diff line number Diff line change
Expand Up @@ -482,11 +482,9 @@ let end_module l restype senv =
let add senv ((l,elem) as field) =
let new_name = match elem with
| SFBconst _ ->
let kn = make_kn mp_sup DirPath.empty l in
C (constant_of_delta_kn resolver kn)
C (constant_of_delta_kn resolver (KerName.make2 mp_sup l))
| SFBmind _ ->
let kn = make_kn mp_sup DirPath.empty l in
I (mind_of_delta_kn resolver kn)
I (mind_of_delta_kn resolver (KerName.make2 mp_sup l))
| SFBmodule _ -> M
| SFBmodtype _ -> MT (MPdot(senv.modinfo.modpath, l))
in
Expand Down
8 changes: 4 additions & 4 deletions kernel/subtyping.ml
Original file line number Diff line number Diff line change
Expand Up @@ -39,7 +39,7 @@ type namedmodule =
constructors *)

let add_mib_nameobjects mp l mib map =
let ind = make_mind mp DirPath.empty l in
let ind = MutInd.make2 mp l in
let add_mip_nameobjects j oib map =
let ip = (ind,j) in
let map =
Expand Down Expand Up @@ -88,8 +88,8 @@ let check_conv_error error why cst f env a1 a2 =
(* for now we do not allow reorderings *)

let check_inductive cst env mp1 l info1 mp2 mib2 spec2 subst1 subst2 reso1 reso2=
let kn1 = make_kn mp1 DirPath.empty l in
let kn2 = make_kn mp2 DirPath.empty l in
let kn1 = KerName.make2 mp1 l in
let kn2 = KerName.make2 mp2 l in
let error why = error_signature_mismatch l spec2 why in
let check_conv why cst f = check_conv_error error why cst f in
let mib1 =
Expand Down Expand Up @@ -180,7 +180,7 @@ let check_inductive cst env mp1 l info1 mp2 mib2 spec2 subst1 subst2 reso1 reso2
let kn2' = kn_of_delta reso2 kn2 in
if KerName.equal kn2 kn2' ||
MutInd.equal (mind_of_delta_kn reso1 kn1)
(subst_ind subst2 (MutInd.make2 kn2 kn2'))
(subst_ind subst2 (MutInd.make kn2 kn2'))
then ()
else error NotEqualInductiveAliases
end;
Expand Down
4 changes: 2 additions & 2 deletions library/globnames.ml
Original file line number Diff line number Diff line change
Expand Up @@ -138,9 +138,9 @@ let constr_of_global_or_constr = function

(** {6 Temporary function to brutally form kernel names from section paths } *)

let encode_mind dir id = make_mind (MPfile dir) DirPath.empty (Label.of_id id)
let encode_mind dir id = MutInd.make2 (MPfile dir) (Label.of_id id)

let encode_con dir id = make_con (MPfile dir) DirPath.empty (Label.of_id id)
let encode_con dir id = Constant.make2 (MPfile dir) (Label.of_id id)

let check_empty_section dp =
if not (DirPath.is_empty dp) then
Expand Down
2 changes: 1 addition & 1 deletion plugins/extraction/common.ml
Original file line number Diff line number Diff line change
Expand Up @@ -584,7 +584,7 @@ let pp_module mp =
the constants are directly turned into chars *)

let mk_ind path s =
make_mind (MPfile (dirpath_of_string path)) DirPath.empty (Label.make s)
MutInd.make2 (MPfile (dirpath_of_string path)) (Label.make s)

let ind_ascii = mk_ind "Coq.Strings.Ascii" "ascii"

Expand Down
10 changes: 5 additions & 5 deletions plugins/extraction/extract_env.ml
Original file line number Diff line number Diff line change
Expand Up @@ -219,13 +219,13 @@ let env_for_mtb_with_def env mp seb idl =
let rec extract_sfb_spec env mp = function
| [] -> []
| (l,SFBconst cb) :: msig ->
let kn = make_con mp DirPath.empty l in
let kn = Constant.make2 mp l in
let s = extract_constant_spec env kn cb in
let specs = extract_sfb_spec env mp msig in
if logical_spec s then specs
else begin Visit.add_spec_deps s; (l,Spec s) :: specs end
| (l,SFBmind _) :: msig ->
let mind = make_mind mp DirPath.empty l in
let mind = MutInd.make2 mp l in
let s = Sind (mind, extract_inductive env mind) in
let specs = extract_sfb_spec env mp msig in
if logical_spec s then specs
Expand Down Expand Up @@ -288,7 +288,7 @@ let rec extract_sfb env mp all = function
| (l,SFBconst cb) :: msb ->
(try
let vl,recd,msb = factor_fix env l cb msb in
let vc = Array.map (make_con mp DirPath.empty) vl in
let vc = Array.map (Constant.make2 mp) vl in
let ms = extract_sfb env mp all msb in
let b = Array.exists Visit.needed_con vc in
if all || b then
Expand All @@ -298,7 +298,7 @@ let rec extract_sfb env mp all = function
else ms
with Impossible ->
let ms = extract_sfb env mp all msb in
let c = make_con mp DirPath.empty l in
let c = Constant.make2 mp l in
let b = Visit.needed_con c in
if all || b then
let d = extract_constant env c cb in
Expand All @@ -307,7 +307,7 @@ let rec extract_sfb env mp all = function
else ms)
| (l,SFBmind mib) :: msb ->
let ms = extract_sfb env mp all msb in
let mind = make_mind mp DirPath.empty l in
let mind = MutInd.make2 mp l in
let b = Visit.needed_ind mind in
if all || b then
let d = Dind (mind, extract_inductive env mind) in
Expand Down
4 changes: 2 additions & 2 deletions plugins/extraction/extraction.ml
Original file line number Diff line number Diff line change
Expand Up @@ -442,15 +442,15 @@ and extract_ind env kn = (* kn is supposed to be in long form *)
List.skipn mib.mind_nparams (names_prod mip0.mind_user_lc.(0)) in
assert (List.length field_names = List.length typ);
let projs = ref Cset.empty in
let mp,d,_ = repr_mind kn in
let mp = MutInd.modpath kn in
let rec select_fields l typs = match l,typs with
| [],[] -> []
| _::l, typ::typs when isDummy (expand env typ) ->
select_fields l typs
| Anonymous::l, typ::typs ->
None :: (select_fields l typs)
| Name id::l, typ::typs ->
let knp = make_con mp d (Label.of_id id) in
let knp = Constant.make2 mp (Label.of_id id) in
(* Is it safe to use [id] for projections [foo.id] ? *)
if List.for_all ((=) Keep) (type2signature env typ)
then projs := Cset.add knp !projs;
Expand Down
3 changes: 1 addition & 2 deletions plugins/extraction/mlutil.ml
Original file line number Diff line number Diff line change
Expand Up @@ -1304,9 +1304,8 @@ let inline_test r t =
not (is_fix t2) && ml_size t < 12 && is_not_strict t)

let con_of_string s =
let null = DirPath.empty in
match DirPath.repr (dirpath_of_string s) with
| id :: d -> make_con (MPfile (DirPath.make d)) null (Label.of_id id)
| id :: d -> Constant.make2 (MPfile (DirPath.make d)) (Label.of_id id)
| [] -> assert false

let manual_inline_set =
Expand Down
2 changes: 1 addition & 1 deletion plugins/extraction/modutil.ml
Original file line number Diff line number Diff line change
Expand Up @@ -34,7 +34,7 @@ let se_iter do_decl do_spec do_mp =
let mp_w =
List.fold_left (fun mp l -> MPdot(mp,Label.of_id l)) mp_mt idl'
in
let r = ConstRef (make_con mp_w DirPath.empty (Label.of_id l')) in
let r = ConstRef (Constant.make2 mp_w (Label.of_id l')) in
mt_iter mt; do_decl (Dtype(r,l,t))
| MTwith (mt,ML_With_module(idl,mp))->
let mp_mt = msid_of_mt mt in
Expand Down
2 changes: 1 addition & 1 deletion plugins/extraction/ocaml.ml
Original file line number Diff line number Diff line change
Expand Up @@ -650,7 +650,7 @@ and pp_module_type params = function
let mp_w =
List.fold_left (fun mp l -> MPdot(mp,Label.of_id l)) mp_mt idl'
in
let r = ConstRef (make_con mp_w DirPath.empty (Label.of_id l)) in
let r = ConstRef (Constant.make2 mp_w (Label.of_id l)) in
push_visible mp_mt [];
let pp_w = str " with type " ++ ids ++ pp_global Type r in
pop_visible();
Expand Down
2 changes: 1 addition & 1 deletion plugins/syntax/numbers_syntax.ml
Original file line number Diff line number Diff line change
Expand Up @@ -18,7 +18,7 @@ open Glob_term
let make_dir l = DirPath.make (List.rev_map Id.of_string l)
let make_path dir id = Libnames.make_path (make_dir dir) (Id.of_string id)

let make_mind mp id = Names.make_mind mp DirPath.empty (Label.make id)
let make_mind mp id = Names.MutInd.make2 mp (Label.make id)
let make_mind_mpfile dir id = make_mind (MPfile (make_dir dir)) id
let make_mind_mpdot dir modname id =
let mp = MPdot (MPfile (make_dir dir), Label.make modname)
Expand Down
Loading

0 comments on commit 7dd28b4

Please sign in to comment.