aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar soubiran <soubiran@85f007b7-540e-0410-9357-904b9bb8a0f7>2009-10-21 15:12:52 +0000
committerGravatar soubiran <soubiran@85f007b7-540e-0410-9357-904b9bb8a0f7>2009-10-21 15:12:52 +0000
commitfe1979bf47951352ce32a6709cb5138fd26f311d (patch)
tree5921dabde1ab3e16da5ae08fe16adf514f1fc07a
parent148a8ee9dec2c04a8d73967b427729c95f039a6a (diff)
This big commit addresses two problems:
1- Management of the name-space in a modular development / sharing of non-logical objects. 2- Performance of atomic module operations (adding a module to the environment, subtyping ...). 1- There are 3 module constructions which derive equalities on fields from a module to another: Let P be a module path and foo a field of P Module M := P. Module M. Include P. ... End M. Declare Module K : S with Module M := P. In this 3 cases we don't want to be bothered by the duplication of names. Of course, M.foo delta reduce to P.foo but many non-logical features of coq do not work modulo conversion (they use eq_constr or constr_pat object). To engender a transparent name-space (ie using P.foo or M.foo is the same thing) we quotient the name-space by the equivalence relation on names induced by the 3 constructions above. To implement this, the types constant and mutual_inductive are now couples of kernel_names. The first projection correspond to the name used by the user and the second projection to the canonical name, for example the internal name of M.foo is (M.foo,P.foo). So: ************************************************************************************* * Use the eq_(con,mind,constructor,gr,egr...) function and not = on names values * ************************************************************************************* Map and Set indexed on names are ordered on user name for the kernel side and on canonical name outside. Thus we have sharing of notation, hints... for free (also for a posteriori declaration of them, ex: a notation on M.foo will be avaible on P.foo). If you want to use this, use the appropriate compare function defined in name.ml or libnames.ml. 2- No more time explosion (i hoppe) when using modules i have re-implemented atomic module operations so that they are all linear in the size of the module. We also have no more unique identifier (internal module names) for modules, it is now based on a section_path like mechanism => we have less substitutions to perform at require, module closing and subtyping but we pre-compute more information hence if we instanciate several functors then we have bigger vo. Last thing, the checker will not work well on vo(s) that contains one of the 3 constructions above, i will work on it soon... git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@12406 85f007b7-540e-0410-9357-904b9bb8a0f7
-rw-r--r--checker/check_stat.ml2
-rw-r--r--checker/closure.ml2
-rw-r--r--checker/declarations.ml715
-rw-r--r--checker/declarations.mli39
-rw-r--r--checker/environ.ml46
-rw-r--r--checker/environ.mli13
-rw-r--r--checker/indtypes.ml6
-rw-r--r--checker/inductive.ml2
-rw-r--r--checker/mod_checking.ml254
-rw-r--r--checker/modops.ml412
-rw-r--r--checker/modops.mli24
-rw-r--r--checker/safe_typing.ml15
-rw-r--r--checker/subtyping.ml188
-rw-r--r--checker/term.ml2
-rw-r--r--checker/typeops.ml4
-rw-r--r--dev/db5
-rw-r--r--dev/printers.mllib2
-rw-r--r--dev/top_printers.ml29
-rw-r--r--dev/vm_printers.ml4
-rw-r--r--interp/coqlib.ml10
-rw-r--r--interp/notation.ml26
-rw-r--r--interp/syntax_def.ml2
-rw-r--r--interp/topconstr.ml8
-rw-r--r--kernel/cbytegen.ml5
-rw-r--r--kernel/cemitcodes.ml4
-rw-r--r--kernel/closure.ml14
-rw-r--r--kernel/closure.mli4
-rw-r--r--kernel/cooking.ml14
-rw-r--r--kernel/cooking.mli2
-rw-r--r--kernel/declarations.ml59
-rw-r--r--kernel/declarations.mli40
-rw-r--r--kernel/entries.ml12
-rw-r--r--kernel/entries.mli11
-rw-r--r--kernel/environ.ml30
-rw-r--r--kernel/environ.mli7
-rw-r--r--kernel/indtypes.ml3
-rw-r--r--kernel/inductive.ml12
-rw-r--r--kernel/mod_subst.ml732
-rw-r--r--kernel/mod_subst.mli87
-rw-r--r--kernel/mod_typing.ml456
-rw-r--r--kernel/mod_typing.mli17
-rw-r--r--kernel/modops.ml762
-rw-r--r--kernel/modops.mli46
-rw-r--r--kernel/names.ml137
-rw-r--r--kernel/names.mli54
-rw-r--r--kernel/pre_env.ml24
-rw-r--r--kernel/pre_env.mli9
-rw-r--r--kernel/reduction.ml8
-rw-r--r--kernel/safe_typing.ml502
-rw-r--r--kernel/safe_typing.mli23
-rw-r--r--kernel/subtyping.ml192
-rw-r--r--kernel/term.ml9
-rw-r--r--kernel/term.mli2
-rw-r--r--kernel/vconv.ml10
-rw-r--r--lib/lib.mllib1
-rw-r--r--lib/tries.ml78
-rw-r--r--lib/tries.mli34
-rw-r--r--library/declare.ml25
-rw-r--r--library/declaremods.ml835
-rw-r--r--library/global.ml47
-rw-r--r--library/global.mli19
-rw-r--r--library/goptions.ml2
-rw-r--r--library/heads.ml16
-rw-r--r--library/impargs.ml4
-rw-r--r--library/impargs.mli2
-rw-r--r--library/lib.ml28
-rw-r--r--library/lib.mli10
-rw-r--r--library/libnames.ml58
-rw-r--r--library/libnames.mli15
-rw-r--r--library/libobject.ml14
-rw-r--r--library/libobject.mli6
-rw-r--r--library/library.ml2
-rw-r--r--library/nametab.ml22
-rw-r--r--parsing/prettyp.ml9
-rw-r--r--parsing/printmod.ml25
-rw-r--r--plugins/extraction/common.ml3
-rw-r--r--plugins/extraction/common.mli2
-rw-r--r--plugins/extraction/extract_env.ml124
-rw-r--r--plugins/extraction/extraction.ml21
-rw-r--r--plugins/extraction/extraction.mli2
-rw-r--r--plugins/extraction/haskell.ml6
-rw-r--r--plugins/extraction/miniml.mli4
-rw-r--r--plugins/extraction/mlutil.ml6
-rw-r--r--plugins/extraction/mlutil.mli2
-rw-r--r--plugins/extraction/modutil.ml21
-rw-r--r--plugins/extraction/ocaml.ml24
-rw-r--r--plugins/extraction/table.ml55
-rw-r--r--plugins/extraction/table.mli9
-rw-r--r--plugins/field/field.ml42
-rw-r--r--plugins/funind/indfun_common.ml4
-rw-r--r--plugins/interface/centaur.ml42
-rw-r--r--plugins/interface/name_to_ast.ml5
-rw-r--r--plugins/interface/name_to_ast.mli2
-rw-r--r--plugins/ring/ring.ml2
-rw-r--r--plugins/setoid_ring/newring.ml44
-rw-r--r--plugins/subtac/subtac_cases.ml2
-rw-r--r--plugins/subtac/subtac_obligations.ml2
-rw-r--r--plugins/syntax/ascii_syntax.ml2
-rw-r--r--plugins/syntax/numbers_syntax.ml13
-rw-r--r--plugins/syntax/z_syntax.ml2
-rw-r--r--plugins/xml/cic2acic.ml4
-rw-r--r--plugins/xml/xmlcommand.ml2
-rw-r--r--pretyping/cases.ml2
-rw-r--r--pretyping/classops.ml4
-rw-r--r--pretyping/detyping.ml6
-rw-r--r--pretyping/evarconv.ml4
-rw-r--r--pretyping/indrec.ml7
-rw-r--r--pretyping/inductiveops.ml2
-rw-r--r--pretyping/pattern.ml9
-rw-r--r--pretyping/recordops.ml8
-rw-r--r--pretyping/termops.ml21
-rw-r--r--pretyping/typeclasses.ml7
-rw-r--r--proofs/redexpr.ml2
-rw-r--r--tactics/auto.ml76
-rw-r--r--tactics/auto.mli2
-rw-r--r--tactics/autorewrite.ml3
-rw-r--r--tactics/btermdn.ml180
-rw-r--r--tactics/btermdn.mli20
-rw-r--r--tactics/decl_interp.mli2
-rw-r--r--tactics/dhyp.ml31
-rw-r--r--tactics/dn.ml183
-rw-r--r--tactics/dn.mli65
-rw-r--r--tactics/extratactics.ml42
-rw-r--r--tactics/nbtermdn.ml92
-rw-r--r--tactics/nbtermdn.mli55
-rw-r--r--tactics/tacinterp.ml4
-rw-r--r--tactics/termdn.ml35
-rw-r--r--tactics/termdn.mli75
-rw-r--r--toplevel/auto_ind_decl.ml14
-rw-r--r--toplevel/auto_ind_decl.mli2
-rw-r--r--toplevel/class.ml2
-rw-r--r--toplevel/command.ml19
-rw-r--r--toplevel/libtypes.ml21
-rw-r--r--toplevel/metasyntax.ml8
-rw-r--r--toplevel/mltop.ml42
-rw-r--r--toplevel/search.ml11
-rw-r--r--toplevel/whelp.ml44
137 files changed, 4081 insertions, 3575 deletions
diff --git a/checker/check_stat.ml b/checker/check_stat.ml
index 290e6ff8e..170ac6384 100644
--- a/checker/check_stat.ml
+++ b/checker/check_stat.ml
@@ -33,7 +33,7 @@ let pr_engt = function
str "Theory: Set is predicative"
let cst_filter f csts =
- Cmap.fold
+ Cmap_env.fold
(fun c ce acc -> if f c ce then c::acc else acc)
csts []
diff --git a/checker/closure.ml b/checker/closure.ml
index b55c5848e..c710df816 100644
--- a/checker/closure.ml
+++ b/checker/closure.ml
@@ -378,7 +378,7 @@ let defined_rels flags env =
(rel_context env) ~init:(0,[])
(* else (0,[])*)
-let mind_equiv_infos info = mind_equiv info.i_env
+let mind_equiv_infos info = eq_ind
let create mk_cl flgs env =
{ i_flags = flgs;
diff --git a/checker/declarations.ml b/checker/declarations.ml
index 0066e7848..3211cc28f 100644
--- a/checker/declarations.ml
+++ b/checker/declarations.ml
@@ -31,60 +31,241 @@ let val_cst_type =
type substitution_domain =
- MSI of mod_self_id
| MBI of mod_bound_id
| MPI of module_path
let val_subst_dom =
- val_sum "substitution_domain" 0 [|[|val_uid|];[|val_uid|];[|val_mp|]|]
+ val_sum "substitution_domain" 0 [|[|val_uid|];[|val_mp|]|]
module Umap = Map.Make(struct
type t = substitution_domain
let compare = Pervasives.compare
end)
-type resolver
-type substitution = (module_path * resolver option) Umap.t
+type delta_hint =
+ Inline of constr option
+ | Equiv of kernel_name
+ | Prefix_equiv of module_path
+
+type delta_key =
+ KN of kernel_name
+ | MP of module_path
+
+module Deltamap = Map.Make(struct
+ type t = delta_key
+ let compare = Pervasives.compare
+ end)
+
+type delta_resolver = delta_hint Deltamap.t
+
+let empty_delta_resolver = Deltamap.empty
+
+type substitution = (module_path * delta_resolver) Umap.t
type 'a subst_fun = substitution -> 'a -> 'a
-let val_res = val_opt no_val
+let val_res_dom = no_val
+ (*val_sum "resolver_domain" 0 [|[|val_kn|];[|val_mp|]|]*)
+
+let val_res = no_val
+ (* val_map ~name:"delta_resolver"
+ val_res_dom
+ (val_sum "resolver_codomain" 0 [|[|val_opt val_constr||];[|val_kn|];[|val_mp|]|])*)
let val_subst =
val_map ~name:"substitution"
val_subst_dom (val_tuple "substition range" [|val_mp;val_res|])
-let fold_subst fs fb fp =
+let fold_subst fb fp =
Umap.fold
(fun k (mp,_) acc ->
match k with
- MSI msid -> fs msid mp acc
| MBI mbid -> fb mbid mp acc
| MPI mp1 -> fp mp1 mp acc)
let empty_subst = Umap.empty
-let add_msid msid mp =
- Umap.add (MSI msid) (mp,None)
let add_mbid mbid mp =
- Umap.add (MBI mbid) (mp,None)
+ Umap.add (MBI mbid) (mp,empty_delta_resolver)
let add_mp mp1 mp2 =
- Umap.add (MPI mp1) (mp2,None)
+ Umap.add (MPI mp1) (mp2,empty_delta_resolver)
-let map_msid msid mp = add_msid msid mp empty_subst
let map_mbid mbid mp = add_mbid mbid mp empty_subst
let map_mp mp1 mp2 = add_mp mp1 mp2 empty_subst
+let add_inline_delta_resolver con =
+ Deltamap.add (KN(user_con con)) (Inline None)
+
+let add_inline_constr_delta_resolver con cstr =
+ Deltamap.add (KN(user_con con)) (Inline (Some cstr))
+
+let add_constant_delta_resolver con =
+ Deltamap.add (KN(user_con con)) (Equiv (canonical_con con))
+
+let add_mind_delta_resolver mind =
+ Deltamap.add (KN(user_mind mind)) (Equiv (canonical_mind mind))
+
+let add_mp_delta_resolver mp1 mp2 =
+ Deltamap.add (MP mp1) (Prefix_equiv mp2)
+
+let mp_in_delta mp =
+ Deltamap.mem (MP mp)
+
+let con_in_delta con resolver =
+try
+ match Deltamap.find (KN(user_con con)) resolver with
+ | Inline _ | Prefix_equiv _ -> false
+ | Equiv _ -> true
+with
+ Not_found -> false
+
+let mind_in_delta mind resolver =
+try
+ match Deltamap.find (KN(user_mind mind)) resolver with
+ | Inline _ | Prefix_equiv _ -> false
+ | Equiv _ -> true
+with
+ Not_found -> false
+
+let delta_of_mp resolve mp =
+ try
+ match Deltamap.find (MP mp) resolve with
+ | Prefix_equiv mp1 -> mp1
+ | _ -> anomaly "mod_subst: bad association in delta_resolver"
+ with
+ Not_found -> mp
+
+let delta_of_kn resolve kn =
+ try
+ match Deltamap.find (KN kn) resolve with
+ | Equiv kn1 -> kn1
+ | Inline _ -> kn
+ | _ -> anomaly
+ "mod_subst: bad association in delta_resolver"
+ with
+ Not_found -> kn
+
+let remove_mp_delta_resolver resolver mp =
+ Deltamap.remove (MP mp) resolver
+
+exception Inline_kn
+
+let rec find_prefix resolve mp =
+ let rec sub_mp = function
+ | MPdot(mp,l) as mp_sup ->
+ (try
+ match Deltamap.find (MP mp_sup) resolve with
+ | Prefix_equiv mp1 -> mp1
+ | _ -> anomaly
+ "mod_subst: bad association in delta_resolver"
+ with
+ Not_found -> MPdot(sub_mp mp,l))
+ | p ->
+ match Deltamap.find (MP p) resolve with
+ | Prefix_equiv mp1 -> mp1
+ | _ -> anomaly
+ "mod_subst: bad association in delta_resolver"
+ in
+ try
+ sub_mp mp
+ with
+ Not_found -> mp
+
+let solve_delta_kn resolve kn =
+ try
+ match Deltamap.find (KN kn) resolve with
+ | Equiv kn1 -> kn1
+ | Inline _ -> raise Inline_kn
+ | _ -> anomaly
+ "mod_subst: bad association in delta_resolver"
+ with
+ Not_found | Inline_kn ->
+ let mp,dir,l = repr_kn kn in
+ let new_mp = find_prefix resolve mp in
+ if mp == new_mp then
+ kn
+ else
+ make_kn new_mp dir l
+
+
+let constant_of_delta resolve con =
+ let kn = user_con con in
+ let new_kn = solve_delta_kn resolve kn in
+ if kn == new_kn then
+ con
+ else
+ constant_of_kn_equiv kn new_kn
+
+let constant_of_delta2 resolve con =
+ let kn = canonical_con con in
+ let kn1 = user_con con in
+ let new_kn = solve_delta_kn resolve kn in
+ if kn == new_kn then
+ con
+ else
+ constant_of_kn_equiv kn1 new_kn
+
+let mind_of_delta resolve mind =
+ let kn = user_mind mind in
+ let new_kn = solve_delta_kn resolve kn in
+ if kn == new_kn then
+ mind
+ else
+ mind_of_kn_equiv kn new_kn
+
+let mind_of_delta2 resolve mind =
+ let kn = canonical_mind mind in
+ let kn1 = user_mind mind in
+ let new_kn = solve_delta_kn resolve kn in
+ if kn == new_kn then
+ mind
+ else
+ mind_of_kn_equiv kn1 new_kn
+
+
+
+let inline_of_delta resolver =
+ let extract key hint l =
+ match key,hint with
+ |KN kn, Inline _ -> kn::l
+ | _,_ -> l
+ in
+ Deltamap.fold extract resolver []
+
+exception Not_inline
+
+let constant_of_delta_with_inline resolve con =
+ let kn1,kn2 = canonical_con con,user_con con in
+ try
+ match Deltamap.find (KN kn2) resolve with
+ | Inline None -> None
+ | Inline (Some const) -> Some const
+ | _ -> raise Not_inline
+ with
+ Not_found | Not_inline ->
+ try match Deltamap.find (KN kn1) resolve with
+ | Inline None -> None
+ | Inline (Some const) -> Some const
+ | _ -> raise Not_inline
+ with
+ Not_found | Not_inline -> None
+
let subst_mp0 sub mp = (* 's like subst *)
let rec aux mp =
match mp with
- | MPself sid ->
- let mp',resolve = Umap.find (MSI sid) sub in
+ | MPfile sid ->
+ let mp',resolve = Umap.find (MPI (MPfile sid)) sub in
mp',resolve
| MPbound bid ->
- let mp',resolve = Umap.find (MBI bid) sub in
- mp',resolve
+ begin
+ try
+ let mp',resolve = Umap.find (MBI bid) sub in
+ mp',resolve
+ with Not_found ->
+ let mp',resolve = Umap.find (MPI mp) sub in
+ mp',resolve
+ end
| MPdot (mp1,l) as mp2 ->
begin
try
@@ -94,59 +275,135 @@ let subst_mp0 sub mp = (* 's like subst *)
let mp1',resolve = aux mp1 in
MPdot (mp1',l),resolve
end
- | _ -> raise Not_found
in
try
Some (aux mp)
with Not_found -> None
-
-
-let subst_mp0 sub mp = (* 's like subst *)
- let rec aux mp =
- match mp with
- | MPself sid -> fst (Umap.find (MSI sid) sub)
- | MPbound bid -> fst (Umap.find (MBI bid) sub)
- | MPdot (mp1,l) as mp2 ->
- begin
- try fst (Umap.find (MPI mp2) sub)
- with Not_found -> MPdot (aux mp1,l)
- end
-
- | _ -> raise Not_found
- in
- try Some (aux mp) with Not_found -> None
-
let subst_mp sub mp =
match subst_mp0 sub mp with
None -> mp
- | Some mp' -> mp'
+ | Some (mp',_) -> mp'
-let subst_kn0 sub kn =
+let subst_kn_delta sub kn =
let mp,dir,l = repr_kn kn in
match subst_mp0 sub mp with
- Some mp' ->
- Some (make_kn mp' dir l)
- | None -> None
+ Some (mp',resolve) ->
+ solve_delta_kn resolve (make_kn mp' dir l)
+ | None -> kn
let subst_kn sub kn =
- match subst_kn0 sub kn with
- None -> kn
- | Some kn' -> kn'
+ let mp,dir,l = repr_kn kn in
+ match subst_mp0 sub mp with
+ Some (mp',_) ->
+ make_kn mp' dir l
+ | None -> kn
+
+exception No_subst
+
+type sideconstantsubst =
+ | User
+ | Canonical
+
+let subst_ind sub mind =
+ let kn1,kn2 = user_mind mind,canonical_mind mind in
+ let mp1,dir,l = repr_kn kn1 in
+ let mp2,_,_ = repr_kn kn2 in
+ try
+ let side,mind',resolve =
+ match subst_mp0 sub mp1,subst_mp0 sub mp2 with
+ None,None ->raise No_subst
+ | Some (mp',resolve),None -> User,(make_mind_equiv mp' mp2 dir l), resolve
+ | None, Some(mp',resolve)-> Canonical,(make_mind_equiv mp1 mp' dir l), resolve
+ | Some(mp1',resolve1),Some(mp2',resolve2)->Canonical,
+ (make_mind_equiv mp1' mp2' dir l), resolve2
+ in
+ match side with
+ |User ->
+ let mind = mind_of_delta resolve mind' in
+ mind
+ |Canonical ->
+ let mind = mind_of_delta2 resolve mind' in
+ mind
+ with
+ No_subst -> mind
+
+let subst_mind0 sub mind =
+ let kn1,kn2 = user_mind mind,canonical_mind mind in
+ let mp1,dir,l = repr_kn kn1 in
+ let mp2,_,_ = repr_kn kn2 in
+ try
+ let side,mind',resolve =
+ match subst_mp0 sub mp1,subst_mp0 sub mp2 with
+ None,None ->raise No_subst
+ | Some (mp',resolve),None -> User,(make_mind_equiv mp' mp2 dir l), resolve
+ | None, Some(mp',resolve)-> Canonical,(make_mind_equiv mp1 mp' dir l), resolve
+ | Some(mp1',resolve1),Some(mp2',resolve2)->Canonical,
+ (make_mind_equiv mp1' mp2' dir l), resolve2
+ in
+ match side with
+ |User ->
+ let mind = mind_of_delta resolve mind' in
+ Some mind
+ |Canonical ->
+ let mind = mind_of_delta2 resolve mind' in
+ Some mind
+ with
+ No_subst -> Some mind
let subst_con sub con =
- let mp,dir,l = repr_con con in
- match subst_mp0 sub mp with
- None -> con
- | Some mp' -> make_con mp' dir l
+ let kn1,kn2 = user_con con,canonical_con con in
+ let mp1,dir,l = repr_kn kn1 in
+ let mp2,_,_ = repr_kn kn2 in
+ try
+ let side,con',resolve =
+ match subst_mp0 sub mp1,subst_mp0 sub mp2 with
+ None,None ->raise No_subst
+ | Some (mp',resolve),None -> User,(make_con_equiv mp' mp2 dir l), resolve
+ | None, Some(mp',resolve)-> Canonical,(make_con_equiv mp1 mp' dir l), resolve
+ | Some(mp1',resolve1),Some(mp2',resolve2)->Canonical,
+ (make_con_equiv mp1' mp2' dir l), resolve2
+ in
+ match constant_of_delta_with_inline resolve con' with
+ None -> begin
+ match side with
+ |User ->
+ let con = constant_of_delta resolve con' in
+ con,Const con
+ |Canonical ->
+ let con = constant_of_delta2 resolve con' in
+ con,Const con
+ end
+ | Some t -> con',t
+ with No_subst -> con , Const con
+
let subst_con0 sub con =
- let mp,dir,l = repr_con con in
- match subst_mp0 sub mp with
- None -> None
- | Some mp' ->
- let con' = make_con mp' dir l in
- Some (Const con')
+ let kn1,kn2 = user_con con,canonical_con con in
+ let mp1,dir,l = repr_kn kn1 in
+ let mp2,_,_ = repr_kn kn2 in
+ try
+ let side,con',resolve =
+ match subst_mp0 sub mp1,subst_mp0 sub mp2 with
+ None,None ->raise No_subst
+ | Some (mp',resolve),None -> User,(make_con_equiv mp' mp2 dir l), resolve
+ | None, Some(mp',resolve)-> Canonical,(make_con_equiv mp1 mp' dir l), resolve
+ | Some(mp1',resolve1),Some(mp2',resolve2)->Canonical,
+ (make_con_equiv mp1' mp2' dir l), resolve2
+ in
+ match constant_of_delta_with_inline resolve con' with
+ None ->begin
+ match side with
+ |User ->
+ let con = constant_of_delta resolve con' in
+ Some (Const con)
+ |Canonical ->
+ let con = constant_of_delta2 resolve con' in
+ Some (Const con)
+ end
+ | t -> t
+ with No_subst -> Some (Const con)
+
let rec map_kn f f' c =
let func = map_kn f f' in
@@ -220,22 +477,8 @@ let rec map_kn f f' c =
| _ -> c
let subst_mps sub =
- map_kn (subst_kn0 sub) (subst_con0 sub)
+ map_kn (subst_mind0 sub) (subst_con0 sub)
-let rec replace_mp_in_mp mpfrom mpto mp =
- match mp with
- | _ when mp = mpfrom -> mpto
- | MPdot (mp1,l) ->
- let mp1' = replace_mp_in_mp mpfrom mpto mp1 in
- if mp1==mp1' then mp
- else MPdot (mp1',l)
- | _ -> mp
-
-let replace_mp_in_con mpfrom mpto kn =
- let mp,dir,l = kn in
- let mp'' = replace_mp_in_mp mpfrom mpto mp in
- if mp==mp'' then kn
- else (mp'', dir, l)
type 'a lazy_subst =
| LSval of 'a
@@ -253,117 +496,115 @@ let force fsubst r =
r := LSval a';
a'
-
-
-let join (subst1 : substitution) (subst2 : substitution) =
- let apply_subst (sub : substitution) key (mp,_) =
- match subst_mp0 sub mp with
- None -> mp,None
- | Some mp' -> mp',None in
- let subst = Umap.mapi (apply_subst subst2) subst1 in
- (Umap.fold Umap.add subst2 subst)
-
-let subst_key subst1 subst2 =
- let replace_in_key key mp sub=
- let newkey =
- match key with
- | MPI mp1 ->
- begin
- match subst_mp0 subst1 mp1 with
- | None -> None
- | Some mp2 -> Some (MPI mp2)
- end
- | _ -> None
- in
- match newkey with
- | None -> Umap.add key mp sub
- | Some mpi -> Umap.add mpi mp sub
+let mp_in_key mp key =
+ let rec mp_rec mp1 =
+ match mp1 with
+ | _ when mp1 = mp -> true
+ | MPdot (mp2,l) -> mp_rec mp2
+ | _ -> false
+ in
+ match key with
+ | MP mp1 ->
+ mp_rec mp1
+ | KN kn ->
+ let mp1,dir,l = repr_kn kn in
+ mp_rec mp1
+
+let subset_prefixed_by mp resolver =
+ let prefixmp key hint resolv =
+ if mp_in_key mp key then
+ Deltamap.add key hint resolv
+ else
+ resolv
in
- Umap.fold replace_in_key subst2 empty_subst
-
-let update_subst_alias subst1 subst2 =
- let subst_inv key (mp,_) sub =
- let newmp =
- match key with
- | MBI msid -> Some (MPbound msid)
- | MSI msid -> Some (MPself msid)
- | _ -> None
- in
- match newmp with
- | None -> sub
- | Some mpi -> match mp with
- | MPbound mbid -> Umap.add (MBI mbid) (mpi,None) sub
- | MPself msid -> Umap.add (MSI msid) (mpi,None) sub
- | _ -> Umap.add (MPI mp) (mpi,None) sub
+ Deltamap.fold prefixmp resolver empty_delta_resolver
+
+let subst_dom_delta_resolver subst resolver =
+ let apply_subst key hint resolver =
+ match key with
+ (MP mp) ->
+ Deltamap.add (MP (subst_mp subst mp)) hint resolver
+ | (KN kn) ->
+ Deltamap.add (KN (subst_kn subst kn)) hint resolver
in
- let subst_mbi = Umap.fold subst_inv subst2 empty_subst in
- let alias_subst key (mp,_) sub=
- let newkey =
- match key with
- | MPI mp1 ->
- begin
- match subst_mp0 subst_mbi mp1 with
- | None -> None
- | Some mp2 -> Some (MPI mp2)
- end
- | _ -> None
- in
- match newkey with
- | None -> Umap.add key (mp,None) sub
- | Some mpi -> Umap.add mpi (mp,None) sub
- in
- Umap.fold alias_subst subst1 empty_subst
+ Deltamap.fold apply_subst resolver empty_delta_resolver
-let join_alias (subst1 : substitution) (subst2 : substitution) =
- let apply_subst (sub : substitution) key (mp,_) =
- match subst_mp0 sub mp with
- None -> mp,None
- | Some mp' -> mp',None in
- Umap.mapi (apply_subst subst2) subst1
-
-
-let update_subst subst1 subst2 =
- let subst_inv key (mp,_) l =
- let newmp =
- match key with
- | MBI msid -> MPbound msid
- | MSI msid -> MPself msid
- | MPI mp -> mp
- in
- match mp with
- | MPbound mbid -> ((MBI mbid),newmp)::l
- | MPself msid -> ((MSI msid),newmp)::l
- | _ -> ((MPI mp),newmp)::l
+let subst_mp_delta sub mp key=
+ match subst_mp0 sub mp with
+ None -> empty_delta_resolver,mp
+ | Some (mp',resolve) ->
+ let mp1 = find_prefix resolve mp' in
+ let resolve1 = subset_prefixed_by mp1 resolve in
+ match key with
+ MP mpk ->
+ (subst_dom_delta_resolver
+ (map_mp mp1 mpk) resolve1),mp1
+ | _ -> anomaly "Mod_subst: Bad association in resolver"
+
+let subst_codom_delta_resolver subst resolver =
+ let apply_subst key hint resolver =
+ match hint with
+ Prefix_equiv mp ->
+ let derived_resolve,mpnew = subst_mp_delta subst mp key in
+ Deltamap.fold Deltamap.add derived_resolve
+ (Deltamap.add key (Prefix_equiv mpnew) resolver)
+ | (Equiv kn) ->
+ Deltamap.add key (Equiv (subst_kn_delta subst kn)) resolver
+ | Inline None ->
+ Deltamap.add key hint resolver
+ | Inline (Some t) ->
+ Deltamap.add key (Inline (Some (subst_mps subst t))) resolver
in
- let subst_mbi = Umap.fold subst_inv subst2 [] in
- let alias_subst key (mp,_) sub=
- let newsetkey =
- match key with
- | MPI mp1 ->
- let compute_set_newkey l (k,mp') =
- let mp_from_key = match k with
- | MBI msid -> MPbound msid
- | MSI msid -> MPself msid
- | MPI mp -> mp
- in
- let new_mp1 = replace_mp_in_mp mp_from_key mp' mp1 in
- if new_mp1 == mp1 then l else (MPI new_mp1)::l
- in
- begin
- match List.fold_left compute_set_newkey [] subst_mbi with
- | [] -> None
- | l -> Some (l)
- end
- | _ -> None
+ Deltamap.fold apply_subst resolver empty_delta_resolver
+
+let subst_dom_codom_delta_resolver subst resolver =
+ subst_dom_delta_resolver subst
+ (subst_codom_delta_resolver subst resolver)
+
+let update_delta_resolver resolver1 resolver2 =
+ let apply_res key hint res =
+ try
+ match hint with
+ Prefix_equiv mp ->
+ let new_hint =
+ Prefix_equiv (find_prefix resolver2 mp)
+ in Deltamap.add key new_hint res
+ | Equiv kn ->
+ let new_hint =
+ Equiv (solve_delta_kn resolver2 kn)
+ in Deltamap.add key new_hint res
+ | _ -> Deltamap.add key hint res
+ with not_found ->
+ Deltamap.add key hint res
in
- match newsetkey with
- | None -> sub
- | Some l ->
- List.fold_left (fun s k -> Umap.add k (mp,None) s)
- sub l
- in
- Umap.fold alias_subst subst1 empty_subst
+ Deltamap.fold apply_res resolver1 empty_delta_resolver
+
+let add_delta_resolver resolver1 resolver2 =
+ if resolver1 == resolver2 then
+ resolver2
+ else
+ Deltamap.fold Deltamap.add (update_delta_resolver resolver1 resolver2)
+ resolver2
+let join (subst1 : substitution) (subst2 : substitution) =
+ let apply_subst (sub : substitution) key (mp,resolve) =
+ let mp',resolve' =
+ match subst_mp0 sub mp with
+ None -> mp, None
+ | Some (mp',resolve') -> mp'
+ ,Some resolve' in
+ let resolve'' : delta_resolver =
+ match resolve' with
+ Some res ->
+ add_delta_resolver
+ (subst_dom_codom_delta_resolver sub resolve)
+ res
+ | None ->
+ subst_codom_delta_resolver sub resolve
+ in
+ mp',resolve'' in
+ let subst = Umap.mapi (apply_subst subst2) subst1 in
+ (Umap.fold Umap.add subst2 subst)
let subst_substituted s r =
match !r with
@@ -414,7 +655,7 @@ let val_recarg = val_sum "recarg" 1 (* Norec *)
let subst_recarg sub r = match r with
| Norec | Mrec _ -> r
- | Imbr (kn,i) -> let kn' = subst_kn sub kn in
+ | Imbr (kn,i) -> let kn' = subst_ind sub kn in
if kn==kn' then r else Imbr (kn',i)
type wf_paths = recarg Rtree.t
@@ -547,8 +788,6 @@ type mutual_inductive_body = {
(* Universes constraints enforced by the inductive declaration *)
mind_constraints : Univ.constraints;
- (* Source of the inductive block when aliased in a module *)
- mind_equiv : kernel_name option
}
let val_ind_pack = val_tuple "mutual_inductive_body"
[|val_array val_one_ind;val_bool;val_bool;val_int;val_nctxt;
@@ -605,8 +844,7 @@ let subst_mind sub mib =
mind_params_ctxt =
map_rel_context (subst_mps sub) mib.mind_params_ctxt;
mind_packets = array_smartmap (subst_mind_packet sub) mib.mind_packets ;
- mind_constraints = mib.mind_constraints ;
- mind_equiv = Option.map (subst_kn sub) mib.mind_equiv }
+ mind_constraints = mib.mind_constraints }
(* Modules *)
@@ -616,7 +854,6 @@ type structure_field_body =
| SFBconst of constant_body
| SFBmind of mutual_inductive_body
| SFBmodule of module_body
- | SFBalias of module_path * struct_expr_body option * Univ.constraints option
| SFBmodtype of module_type_body
and structure_body = (label * structure_field_body) list
@@ -624,116 +861,120 @@ and structure_body = (label * structure_field_body) list
and struct_expr_body =
| SEBident of module_path
| SEBfunctor of mod_bound_id * module_type_body * struct_expr_body
- | SEBstruct of mod_self_id * structure_body
- | SEBapply of struct_expr_body * struct_expr_body
- * Univ.constraints
+ | SEBapply of struct_expr_body * struct_expr_body * Univ.constraints
+ | SEBstruct of structure_body
| SEBwith of struct_expr_body * with_declaration_body
and with_declaration_body =
- With_module_body of identifier list * module_path *
- struct_expr_body option * Univ.constraints
+ With_module_body of identifier list * module_path
| With_definition_body of identifier list * constant_body
and module_body =
- { mod_expr : struct_expr_body option;
- mod_type : struct_expr_body option;
+ { mod_mp : module_path;
+ mod_expr : struct_expr_body option;
+ mod_type : struct_expr_body;
+ mod_type_alg : struct_expr_body option;
mod_constraints : Univ.constraints;
- mod_alias : substitution;
+ mod_delta : delta_resolver;
mod_retroknowledge : action list}
and module_type_body =
- { typ_expr : struct_expr_body;
- typ_strength : module_path option;
- typ_alias : substitution}
+ { typ_mp : module_path;
+ typ_expr : struct_expr_body;
+ typ_expr_alg : struct_expr_body option ;
+ typ_constraints : Univ.constraints;
+ typ_delta :delta_resolver}
(* the validation functions: *)
let rec val_sfb o = val_sum "struct_field_body" 0
[|[|val_cb|]; (* SFBconst *)
[|val_ind_pack|]; (* SFBmind *)
[|val_module|]; (* SFBmodule *)
- [|val_mp;val_opt val_seb;val_opt val_cstrs|]; (* SFBalias *)
[|val_modtype|] (* SFBmodtype *)
|] o
and val_sb o = val_list (val_tuple"label*sfb"[|val_id;val_sfb|]) o
and val_seb o = val_sum "struct_expr_body" 0
[|[|val_mp|]; (* SEBident *)
[|val_uid;val_modtype;val_seb|]; (* SEBfunctor *)
- [|val_uid;val_sb|]; (* SEBstruct *)
+ [|val_sb|]; (* SEBstruct *)
[|val_seb;val_seb;val_cstrs|]; (* SEBapply *)
[|val_seb;val_with|] (* SEBwith *)
|] o
and val_with o = val_sum "with_declaration_body" 0
- [|[|val_list val_id;val_mp;val_cstrs|];
+ [|[|val_list val_id;val_mp|];
[|val_list val_id;val_cb|]|] o
and val_module o = val_tuple "module_body"
- [|val_opt val_seb;val_opt val_seb;val_cstrs;val_subst;no_val|] o
+ [|val_mp;val_opt val_seb;val_opt val_seb;
+ val_seb;val_cstrs;val_res;no_val|] o
and val_modtype o = val_tuple "module_type_body"
- [|val_seb;val_opt val_mp;val_subst|] o
+ [|val_mp;val_seb;val_opt val_seb;val_cstrs;val_res|] o
let rec subst_with_body sub = function
- | With_module_body(id,mp,typ_opt,cst) ->
- With_module_body(id,subst_mp sub mp,
- Option.smartmap (subst_struct_expr sub) typ_opt,cst)
+ | With_module_body(id,mp) ->
+ With_module_body(id,subst_mp sub mp)
| With_definition_body(id,cb) ->
With_definition_body( id,subst_const_body sub cb)
-and subst_modtype sub mtb =
+and subst_modtype sub mtb=
let typ_expr' = subst_struct_expr sub mtb.typ_expr in
- if typ_expr'==mtb.typ_expr then
- mtb
+ let typ_alg' =
+ Option.smartmap
+ (subst_struct_expr sub) mtb.typ_expr_alg in
+ let mp = subst_mp sub mtb.typ_mp
+ in
+ if typ_expr'==mtb.typ_expr &&
+ typ_alg'==mtb.typ_expr_alg && mp==mtb.typ_mp then
+ mtb
else
- { mtb with
- typ_expr = typ_expr'}
-
-and subst_structure sub sign =
- let subst_body = function
- SFBconst cb ->
+ {mtb with
+ typ_mp = mp;
+ typ_expr = typ_expr';
+ typ_expr_alg = typ_alg'}
+
+and subst_structure sub sign =
+ let subst_body = function
+ SFBconst cb ->
SFBconst (subst_const_body sub cb)
- | SFBmind mib ->
+ | SFBmind mib ->
SFBmind (subst_mind sub mib)
- | SFBmodule mb ->
- SFBmodule (subst_module sub mb)
- | SFBmodtype mtb ->
+ | SFBmodule mb ->
+ SFBmodule (subst_module sub mb)
+ | SFBmodtype mtb ->
SFBmodtype (subst_modtype sub mtb)
- | SFBalias (mp,typ_opt ,cst) ->
- SFBalias (subst_mp sub mp,
- Option.smartmap (subst_struct_expr sub) typ_opt,cst)
in
List.map (fun (l,b) -> (l,subst_body b)) sign
-and subst_module sub mb =
- let mtb' = Option.smartmap (subst_struct_expr sub) mb.mod_type in
- (* This is similar to the previous case. In this case we have
- a module M in a signature that is knows to be equivalent to a module M'
- (because the signature is "K with Module M := M'") and we are substituting
- M' with some M''. *)
- let me' = Option.smartmap (subst_struct_expr sub) mb.mod_expr in
- let mb_alias = join_alias mb.mod_alias sub in
- if mtb'==mb.mod_type && mb.mod_expr == me'
- && mb_alias == mb.mod_alias
- then mb else
- { mod_expr = me';
- mod_type=mtb';
- mod_constraints=mb.mod_constraints;
- mod_alias = mb_alias;
- mod_retroknowledge=mb.mod_retroknowledge}
+and subst_module sub mb =
+ let mtb' = subst_struct_expr sub mb.mod_type in
+ let typ_alg' = Option.smartmap
+ (subst_struct_expr sub ) mb.mod_type_alg in
+ let me' = Option.smartmap
+ (subst_struct_expr sub) mb.mod_expr in
+ let mp = subst_mp sub mb.mod_mp in
+ if mtb'==mb.mod_type && mb.mod_expr == me'
+ && mp == mb.mod_mp
+ then mb else
+ { mb with
+ mod_mp = mp;
+ mod_expr = me';
+ mod_type_alg = typ_alg';
+ mod_type=mtb'}
and subst_struct_expr sub = function
- | SEBident mp -> SEBident (subst_mp sub mp)
- | SEBfunctor (msid, mtb, meb') ->
- SEBfunctor(msid,subst_modtype sub mtb,subst_struct_expr sub meb')
- | SEBstruct (msid,str)->
- SEBstruct(msid, subst_structure sub str)
+ | SEBident mp -> SEBident (subst_mp sub mp)
+ | SEBfunctor (mbid, mtb, meb') ->
+ SEBfunctor(mbid,subst_modtype sub mtb
+ ,subst_struct_expr sub meb')
+ | SEBstruct (str)->
+ SEBstruct( subst_structure sub str)
| SEBapply (meb1,meb2,cst)->
SEBapply(subst_struct_expr sub meb1,
subst_struct_expr sub meb2,
cst)
- | SEBwith (meb,wdb)->
+ | SEBwith (meb,wdb)->
SEBwith(subst_struct_expr sub meb,
subst_with_body sub wdb)
-let subst_signature_msid msid mp =
- subst_structure (map_msid msid mp)
diff --git a/checker/declarations.mli b/checker/declarations.mli
index 3d061b4c2..bae7da908 100644
--- a/checker/declarations.mli
+++ b/checker/declarations.mli
@@ -141,19 +141,18 @@ type mutual_inductive_body = {
(* Universes constraints enforced by the inductive declaration *)
mind_constraints : Univ.constraints;
- (* Source of the inductive block when aliased in a module *)
- mind_equiv : kernel_name option
}
(* Modules *)
type substitution
+type delta_resolver
+val empty_delta_resolver : delta_resolver
type structure_field_body =
| SFBconst of constant_body
| SFBmind of mutual_inductive_body
| SFBmodule of module_body
- | SFBalias of module_path * struct_expr_body option * Univ.constraints option
| SFBmodtype of module_type_body
and structure_body = (label * structure_field_body) list
@@ -161,32 +160,33 @@ and structure_body = (label * structure_field_body) list
and struct_expr_body =
| SEBident of module_path
| SEBfunctor of mod_bound_id * module_type_body * struct_expr_body
- | SEBstruct of mod_self_id * structure_body
- | SEBapply of struct_expr_body * struct_expr_body
- * Univ.constraints
+ | SEBapply of struct_expr_body * struct_expr_body * Univ.constraints
+ | SEBstruct of structure_body
| SEBwith of struct_expr_body * with_declaration_body
and with_declaration_body =
- With_module_body of identifier list * module_path *
- struct_expr_body option * Univ.constraints
+ With_module_body of identifier list * module_path
| With_definition_body of identifier list * constant_body
and module_body =
- { mod_expr : struct_expr_body option;
- mod_type : struct_expr_body option;
+ { mod_mp : module_path;
+ mod_expr : struct_expr_body option;
+ mod_type : struct_expr_body;
+ mod_type_alg : struct_expr_body option;
mod_constraints : Univ.constraints;
- mod_alias : substitution;
+ mod_delta : delta_resolver;
mod_retroknowledge : action list}
and module_type_body =
- { typ_expr : struct_expr_body;
- typ_strength : module_path option;
- typ_alias : substitution}
+ { typ_mp : module_path;
+ typ_expr : struct_expr_body;
+ typ_expr_alg : struct_expr_body option ;
+ typ_constraints : Univ.constraints;
+ typ_delta :delta_resolver}
(* Substitutions *)
val fold_subst :
- (mod_self_id -> module_path -> 'a -> 'a) ->
(mod_bound_id -> module_path -> 'a -> 'a) ->
(module_path -> module_path -> 'a -> 'a) ->
substitution -> 'a -> 'a
@@ -194,10 +194,8 @@ val fold_subst :
type 'a subst_fun = substitution -> 'a -> 'a
val empty_subst : substitution
-val add_msid : mod_self_id -> module_path -> substitution -> substitution
val add_mbid : mod_bound_id -> module_path -> substitution -> substitution
val add_mp : module_path -> module_path -> substitution -> substitution
-val map_msid : mod_self_id -> module_path -> substitution
val map_mbid : mod_bound_id -> module_path -> substitution
val map_mp : module_path -> module_path -> substitution
@@ -206,14 +204,9 @@ val subst_mind : mutual_inductive_body subst_fun
val subst_modtype : substitution -> module_type_body -> module_type_body
val subst_struct_expr : substitution -> struct_expr_body -> struct_expr_body
val subst_structure : substitution -> structure_body -> structure_body
-val subst_signature_msid :
- mod_self_id -> module_path -> structure_body -> structure_body
+val subst_module : substitution -> module_body -> module_body
val join : substitution -> substitution -> substitution
-val join_alias : substitution -> substitution -> substitution
-val update_subst_alias : substitution -> substitution -> substitution
-val update_subst : substitution -> substitution -> substitution
-val subst_key : substitution -> substitution -> substitution
(* Validation *)
val val_eng : Obj.t -> unit
diff --git a/checker/environ.ml b/checker/environ.ml
index 2d5ff3e43..3d14b995b 100644
--- a/checker/environ.ml
+++ b/checker/environ.ml
@@ -5,11 +5,10 @@ open Term
open Declarations
type globals = {
- env_constants : constant_body Cmap.t;
- env_inductives : mutual_inductive_body KNmap.t;
+ env_constants : constant_body Cmap_env.t;
+ env_inductives : mutual_inductive_body Mindmap_env.t;
env_modules : module_body MPmap.t;
- env_modtypes : module_type_body MPmap.t;
- env_alias : module_path MPmap.t }
+ env_modtypes : module_type_body MPmap.t}
type stratification = {
env_universes : universes;
@@ -25,11 +24,10 @@ type env = {
let empty_env = {
env_globals =
- { env_constants = Cmap.empty;
- env_inductives = KNmap.empty;
+ { env_constants = Cmap_env.empty;
+ env_inductives = Mindmap_env.empty;
env_modules = MPmap.empty;
- env_modtypes = MPmap.empty;
- env_alias = MPmap.empty };
+ env_modtypes = MPmap.empty};
env_named_context = [];
env_rel_context = [];
env_stratification =
@@ -108,11 +106,11 @@ let add_constraints c env =
(* Global constants *)
let lookup_constant kn env =
- Cmap.find kn env.env_globals.env_constants
+ Cmap_env.find kn env.env_globals.env_constants
let add_constant kn cs env =
let new_constants =
- Cmap.add kn cs env.env_globals.env_constants in
+ Cmap_env.add kn cs env.env_globals.env_constants in
let new_globals =
{ env.env_globals with
env_constants = new_constants } in
@@ -145,26 +143,15 @@ let evaluable_constant cst env =
(* Mutual Inductives *)
let lookup_mind kn env =
- KNmap.find kn env.env_globals.env_inductives
-
-let rec scrape_mind env kn =
- match (lookup_mind kn env).mind_equiv with
- | None -> kn
- | Some kn' -> scrape_mind env kn'
+ Mindmap_env.find kn env.env_globals.env_inductives
let add_mind kn mib env =
- let new_inds = KNmap.add kn mib env.env_globals.env_inductives in
+ let new_inds = Mindmap_env.add kn mib env.env_globals.env_inductives in
let new_globals =
{ env.env_globals with
env_inductives = new_inds } in
{ env with env_globals = new_globals }
-let rec mind_equiv env (kn1,i1) (kn2,i2) =
- let rec equiv kn1 kn2 =
- kn1 = kn2 ||
- scrape_mind env kn1 = scrape_mind env kn2 in
- i1 = i2 && equiv kn1 kn2
-
(* Modules *)
@@ -182,19 +169,6 @@ let shallow_add_module mp mb env =
env_modules = new_mods } in
{ env with env_globals = new_globals }
-let register_alias mp1 mp2 env =
- let new_alias = MPmap.add mp1 mp2 env.env_globals.env_alias in
- let new_globals =
- { env.env_globals with
- env_alias = new_alias } in
- { env with env_globals = new_globals }
-
-let rec scrape_alias mp env =
- try
- let mp1 = MPmap.find mp env.env_globals.env_alias in
- scrape_alias mp1 env
- with
- Not_found -> mp
let lookup_module mp env =
MPmap.find mp env.env_globals.env_modules
diff --git a/checker/environ.mli b/checker/environ.mli
index 1541bf0d8..776698ed8 100644
--- a/checker/environ.mli
+++ b/checker/environ.mli
@@ -4,12 +4,10 @@ open Term
(* Environments *)
type globals = {
- env_constants : Declarations.constant_body Cmap.t;
- env_inductives : Declarations.mutual_inductive_body KNmap.t;
+ env_constants : Declarations.constant_body Cmap_env.t;
+ env_inductives : Declarations.mutual_inductive_body Mindmap_env.t;
env_modules : Declarations.module_body MPmap.t;
- env_modtypes : Declarations.module_type_body MPmap.t;
- env_alias : module_path MPmap.t;
-}
+ env_modtypes : Declarations.module_type_body MPmap.t}
type stratification = {
env_universes : Univ.universes;
env_engagement : Declarations.engagement option;
@@ -61,17 +59,14 @@ val evaluable_constant : constant -> env -> bool
(* Inductives *)
val lookup_mind :
mutual_inductive -> env -> Declarations.mutual_inductive_body
-val scrape_mind : env -> mutual_inductive -> mutual_inductive
+
val add_mind :
mutual_inductive -> Declarations.mutual_inductive_body -> env -> env
-val mind_equiv : env -> inductive -> inductive -> bool
(* Modules *)
val add_modtype :
module_path -> Declarations.module_type_body -> env -> env
val shallow_add_module :
module_path -> Declarations.module_body -> env -> env
-val register_alias : module_path -> module_path -> env -> env
-val scrape_alias : module_path -> env -> module_path
val lookup_module : module_path -> env -> Declarations.module_body
val lookup_modtype : module_path -> env -> Declarations.module_type_body
diff --git a/checker/indtypes.ml b/checker/indtypes.ml
index 3d4f6be79..e0aa9e7c3 100644
--- a/checker/indtypes.ml
+++ b/checker/indtypes.ml
@@ -22,13 +22,11 @@ open Environ
let rec debug_string_of_mp = function
| MPfile sl -> string_of_dirpath sl
| MPbound uid -> "bound("^string_of_mbid uid^")"
- | MPself uid -> "self("^string_of_msid uid^")"
| MPdot (mp,l) -> string_of_mp mp ^ "." ^ string_of_label l
let rec string_of_mp = function
| MPfile sl -> string_of_dirpath sl
| MPbound uid -> string_of_mbid uid
- | MPself uid -> string_of_msid uid
| MPdot (mp,l) -> string_of_mp mp ^ "." ^ string_of_label l
let string_of_mp mp =
@@ -515,7 +513,7 @@ let check_positivity env_ar params nrecp inds =
(************************************************************************)
let check_inductive env kn mib =
- Flags.if_verbose msgnl (str " checking ind: " ++ prkn kn);
+ Flags.if_verbose msgnl (str " checking ind: " ++ pr_mind kn);
(* check mind_constraints: should be consistent with env *)
let env = add_constraints mib.mind_constraints env in
(* check mind_record : TODO ? check #constructor = 1 ? *)
@@ -540,8 +538,6 @@ let check_inductive env kn mib =
(* check mind_nparams_rec: positivity condition *)
check_positivity env_ar params mib.mind_nparams_rec mib.mind_packets;
(* check mind_equiv... *)
- if mib.mind_equiv <> None then
- msg_warning (str"TODO: mind_equiv not checked");
(* Now we can add the inductive *)
add_mind kn mib env
diff --git a/checker/inductive.ml b/checker/inductive.ml
index e08efbe5b..19c7a6cfe 100644
--- a/checker/inductive.ml
+++ b/checker/inductive.ml
@@ -351,7 +351,7 @@ let type_case_branches env (ind,largs) (p,pj) c =
let check_case_info env indsp ci =
let (mib,mip) = lookup_mind_specif env indsp in
if
- not (mind_equiv env indsp ci.ci_ind) or
+ not (eq_ind indsp ci.ci_ind) or
(mib.mind_nparams <> ci.ci_npar) or
(mip.mind_consnrealdecls <> ci.ci_cstr_nargs)
then raise (TypeError(env,WrongCaseInfo(indsp,ci)))
diff --git a/checker/mod_checking.ml b/checker/mod_checking.ml
index 99babe632..e95109fc5 100644
--- a/checker/mod_checking.ml
+++ b/checker/mod_checking.ml
@@ -65,10 +65,6 @@ let rec list_fold_map2 f e = function
let e'',t1',t2' = list_fold_map2 f e' t in
e'',h1'::t1',h2'::t2'
-
-let check_alias (s1:substitution) s2 =
- if s1 <> s2 then failwith "Incorrect alias"
-
let check_definition_sub env cb1 cb2 =
let check_type env t1 t2 =
@@ -135,31 +131,31 @@ let lookup_modtype mp env =
with Not_found ->
failwith ("Unknown module type: "^string_of_mp mp)
-
-let rec check_with env mtb with_decl =
+let rec check_with env mtb with_decl mp=
match with_decl with
| With_definition_body _ ->
- check_with_aux_def env mtb with_decl;
- empty_subst
+ check_with_aux_def env mtb with_decl mp;
+ mtb
| With_module_body _ ->
- check_with_aux_mod env mtb with_decl
+ check_with_aux_mod env mtb with_decl mp;
+ mtb
-and check_with_aux_def env mtb with_decl =
- let msid,sig_b = match (eval_struct env mtb) with
- | SEBstruct(msid,sig_b) ->
- msid,sig_b
+and check_with_aux_def env mtb with_decl mp =
+ let sig_b = match mtb with
+ | SEBstruct(sig_b) ->
+ sig_b
| _ -> error_signature_expected mtb
in
let id,idl = match with_decl with
- | With_definition_body (id::idl,_) | With_module_body (id::idl,_,_,_) ->
+ | With_definition_body (id::idl,_) | With_module_body (id::idl,_) ->
id,idl
- | With_definition_body ([],_) | With_module_body ([],_,_,_) -> assert false
+ | With_definition_body ([],_) | With_module_body ([],_) -> assert false
in
let l = label_of_id id in
try
let rev_before,spec,after = list_split_assoc l [] sig_b in
let before = List.rev rev_before in
- let env' = Modops.add_signature (MPself msid) before env in
+ let env' = Modops.add_signature mp before empty_delta_resolver env in
match with_decl with
| With_definition_body ([],_) -> assert false
| With_definition_body ([id],c) ->
@@ -179,9 +175,9 @@ and check_with_aux_def env mtb with_decl =
let new_with_decl = match with_decl with
With_definition_body (_,c) ->
With_definition_body (idl,c)
- | With_module_body (_,c,t,cst) ->
- With_module_body (idl,c,t,cst) in
- check_with_aux_def env' (type_of_mb env old) new_with_decl
+ | With_module_body (_,c) ->
+ With_module_body (idl,c) in
+ check_with_aux_def env' old.mod_type new_with_decl (MPdot(mp,l))
| Some msb ->
error_a_generative_module_expected l
end
@@ -190,46 +186,35 @@ and check_with_aux_def env mtb with_decl =
Not_found -> error_no_such_label l
| Reduction.NotConvertible -> error_with_incorrect l
-and check_with_aux_mod env mtb with_decl =
- let initmsid,msid,sig_b =
- match eval_struct env mtb with
- | SEBstruct(msid,sig_b) ->
- let msid'=(refresh_msid msid) in
- msid,msid',(subst_signature_msid msid (MPself(msid')) sig_b)
+and check_with_aux_mod env mtb with_decl mp =
+ let sig_b =
+ match mtb with
+ | SEBstruct(sig_b) ->
+ sig_b
| _ -> error_signature_expected mtb in
let id,idl = match with_decl with
- | With_definition_body (id::idl,_) | With_module_body (id::idl,_,_,_) ->
+ | With_definition_body (id::idl,_) | With_module_body (id::idl,_) ->
id,idl
- | With_definition_body ([],_) | With_module_body ([],_,_,_) -> assert false
+ | With_definition_body ([],_) | With_module_body ([],_) -> assert false
in
let l = label_of_id id in
try
let rev_before,spec,after = list_split_assoc l [] sig_b in
let before = List.rev rev_before in
let rec mp_rec = function
- | [] -> MPself initmsid
+ | [] -> mp
| i::r -> MPdot(mp_rec r,label_of_id i)
in
- let env' = Modops.add_signature (MPself msid) before env in
+ let env' = Modops.add_signature mp before empty_delta_resolver env in
match with_decl with
- | With_module_body ([],_,_,_) -> assert false
- | With_module_body ([id], mp,_,_) ->
- let old,alias = match spec with
- SFBmodule msb -> Some msb,None
- | SFBalias (mp',_,_) -> None,Some mp'
+ | With_module_body ([],_) -> assert false
+ | With_module_body ([id], mp1) ->
+ let _ = match spec with
+ SFBmodule msb -> msb
| _ -> error_not_a_module l
in
- let mtb' = lookup_modtype mp env' in
- let _ =
- match old,alias with
- Some msb,None -> ()
- | None,Some mp' ->
- check_modpath_equiv env' mp mp'
- | _,_ ->
- anomaly "Mod_typing:no implementation and no alias"
- in
- join (map_mp (mp_rec [id]) mp) mtb'.typ_alias
- | With_module_body (_::_,mp,_,_) ->
+ let _ = (lookup_module mp1 env) in ()
+ | With_module_body (_::_,mp) ->
let old = match spec with
SFBmodule msb -> msb
| _ -> error_not_a_module l
@@ -240,12 +225,10 @@ and check_with_aux_mod env mtb with_decl =
let new_with_decl = match with_decl with
With_definition_body (_,c) ->
With_definition_body (idl,c)
- | With_module_body (_,c,t,cst) ->
- With_module_body (idl,c,t,cst) in
- let sub =
- check_with_aux_mod env'
- (type_of_mb env old) new_with_decl in
- join (map_mp (mp_rec idl) mp) sub
+ | With_module_body (_,c) ->
+ With_module_body (idl,c) in
+ check_with_aux_mod env'
+ old.mod_type new_with_decl (MPdot(mp,l))
| Some msb ->
error_a_generative_module_expected l
end
@@ -255,110 +238,105 @@ and check_with_aux_mod env mtb with_decl =
| Reduction.NotConvertible -> error_with_incorrect l
and check_module_type env mty =
- if mty.typ_strength <> None then
- failwith "strengthening of module types not supported";
- let sub = check_modexpr env mty.typ_expr in
- check_alias mty.typ_alias sub
-
-and check_module env mb =
- let sub =
- match mb.mod_expr, mb.mod_type with
- | None, None ->
- anomaly "Mod_typing.translate_module: empty type and expr in module entry"
- | None, Some mtb -> check_modexpr env mtb
-
- | Some mexpr, _ ->
- let sub1 = check_modexpr env mexpr in
- (match mb.mod_type with
- | None -> sub1
- | Some mte ->
- let sub2 = check_modexpr env mte in
- check_subtypes env
- {typ_expr = mexpr;
- typ_strength = None;
- typ_alias = sub1;}
- {typ_expr = mte;
- typ_strength = None;
- typ_alias = sub2;};
- sub2) in
- check_alias mb.mod_alias sub
-
-and check_structure_field (s,env) mp lab = function
+ let _ = check_modtype env mty.typ_expr mty.typ_mp in ()
+
+
+and check_module env mp mb =
+ match mb.mod_expr, mb.mod_type with
+ | None,mtb ->
+ let _ = check_modtype env mtb mb.mod_mp in ()
+ | Some mexpr, _ ->
+ let sign = check_modexpr env mexpr mb.mod_mp in
+ let _ = check_modtype env mb.mod_type mb.mod_mp in
+ check_subtypes env
+ {typ_mp=mp;
+ typ_expr=sign;
+ typ_expr_alg=None;
+ typ_constraints=Univ.Constraint.empty;
+ typ_delta = empty_delta_resolver;}
+ {typ_mp=mp;
+ typ_expr=mb.mod_type;
+ typ_expr_alg=None;
+ typ_constraints=Univ.Constraint.empty;
+ typ_delta = empty_delta_resolver;};
+
+and check_structure_field env mp lab = function
| SFBconst cb ->
let c = make_con mp empty_dirpath lab in
- (s,check_constant_declaration env c cb)
+ check_constant_declaration env c cb
| SFBmind mib ->
- let kn = make_kn mp empty_dirpath lab in
- (s,Indtypes.check_inductive env kn mib)
+ let kn = make_mind mp empty_dirpath lab in
+ Indtypes.check_inductive env kn mib
| SFBmodule msb ->
- check_module env msb;
- let mp1 = MPdot(mp,lab) in
- let is_fun, sub = Modops.update_subst env msb mp1 in
- ((if is_fun then s else join s sub),
- Modops.add_module (MPdot(mp,lab)) msb env)
- | SFBalias(mp2,_,cst) ->
- (* cf Safe_typing.add_alias *)
- (try
- let mp' = MPdot(mp,lab) in
- let mp2' = scrape_alias mp2 env in
- let _,sub = Modops.update_subst env (lookup_module mp2' env) mp2' in
- let sub = update_subst sub (map_mp mp' mp2') in
- let sub = join_alias sub (map_mp mp' mp2') in
- let sub = add_mp mp' mp2' sub in
- (join s sub, register_alias mp' mp2 env)
- with Not_found -> failwith "unkown aliased module")
+ let _= check_module env (MPdot(mp,lab)) msb in
+ Modops.add_module msb env
| SFBmodtype mty ->
- let kn = MPdot(mp, lab) in
check_module_type env mty;
- (join s mty.typ_alias, add_modtype kn mty env)
-
-and check_modexpr env mse = match mse with
+ add_modtype (MPdot(mp,lab)) mty env
+
+and check_modexpr env mse mp_mse = match mse with
+ | SEBident mp ->
+ let mb = lookup_module mp env in
+ (subst_and_strengthen mb mp_mse env).mod_type
+ | SEBfunctor (arg_id, mtb, body) ->
+ check_module_type env mtb ;
+ let env' = add_module (module_body_of_type (MPbound arg_id) mtb) env in
+ let sign = check_modexpr env' body mp_mse in
+ SEBfunctor (arg_id, mtb, sign)
+ | SEBapply (f,m,cst) ->
+ let sign = check_modexpr env f mp_mse in
+ let farg_id, farg_b, fbody_b = destr_functor env sign in
+ let mp =
+ try (path_of_mexpr m)
+ with Not_path -> error_application_to_not_path m
+ (* place for nondep_supertype *) in
+ let mtb = module_type_of_module env (Some mp) (lookup_module mp env) in
+ check_subtypes env mtb farg_b;
+ (subst_struct_expr (map_mbid farg_id mp) fbody_b)
+ | SEBwith(mte, with_decl) ->
+ let sign = check_modexpr env mte mp_mse in
+ let sign = check_with env sign with_decl mp_mse in
+ sign
+ | SEBstruct(msb) ->
+ let _ = List.fold_left (fun env (lab,mb) ->
+ check_structure_field env mp_mse lab mb) env msb in
+ SEBstruct(msb)
+
+and check_modtype env mse mp_mse= match mse with
| SEBident mp ->
- let mp = scrape_alias mp env in
let mtb = lookup_modtype mp env in
- mtb.typ_alias
+ mtb.typ_expr
| SEBfunctor (arg_id, mtb, body) ->
check_module_type env mtb;
- let env' = add_module (MPbound arg_id) (module_body_of_type mtb) env in
- let sub = check_modexpr env' body in
- sub
+ let env' = add_module (module_body_of_type (MPbound arg_id) mtb) env in
+ let body = check_modtype env' body mp_mse in
+ SEBfunctor(arg_id,mtb,body)
| SEBapply (f,m,cst) ->
- let sub1 = check_modexpr env f in
- let f'= eval_struct env f in
- let farg_id, farg_b, fbody_b = destr_functor env f' in
+ let sign = check_modtype env f mp_mse in
+ let farg_id, farg_b, fbody_b = destr_functor env sign in
let mp =
- try scrape_alias (path_of_mexpr m) env
+ try (path_of_mexpr m)
with Not_path -> error_application_to_not_path m
- (* place for nondep_supertype *) in
- let mtb = lookup_modtype mp env in
- check_subtypes env mtb farg_b;
- let sub2 = match eval_struct env m with
- | SEBstruct (msid,sign) ->
- join_alias
- (subst_key (map_msid msid mp) mtb.typ_alias)
- (map_msid msid mp)
- | _ -> mtb.typ_alias in
- let sub3 = join_alias sub1 (map_mbid farg_id mp) in
- let sub4 = update_subst sub2 sub3 in
- join sub3 sub4
+ (* place for nondep_supertype *) in
+ let mtb = module_type_of_module env (Some mp) (lookup_module mp env) in
+ check_subtypes env mtb farg_b;
+ subst_struct_expr (map_mbid farg_id mp) fbody_b
| SEBwith(mte, with_decl) ->
- let sub1 = check_modexpr env mte in
- let sub2 = check_with env mte with_decl in
- join sub1 sub2
- | SEBstruct(msid,msb) ->
- let mp = MPself msid in
- let (sub,_) =
- List.fold_left (fun env (lab,mb) ->
- check_structure_field env mp lab mb) (empty_subst,env) msb in
- sub
-
+ let sign = check_modtype env mte mp_mse in
+ let sign = check_with env sign with_decl mp_mse in
+ sign
+ | SEBstruct(msb) ->
+ let _ = List.fold_left (fun env (lab,mb) ->
+ check_structure_field env mp_mse lab mb) env msb in
+ SEBstruct(msb)
+
(*
-let rec add_struct_expr_constraints env = function
+ let rec add_struct_expr_constraints env = function
| SEBident _ -> env
-
+
| SEBfunctor (_,mtb,meb) ->
- add_struct_expr_constraints
- (add_modtype_constraints env mtb) meb
+ add_struct_expr_constraints
+ (add_modtype_constraints env mtb) meb
| SEBstruct (_,structure_body) ->
List.fold_left
diff --git a/checker/modops.ml b/checker/modops.ml
index a986e1898..3d07135d1 100644
--- a/checker/modops.ml
+++ b/checker/modops.ml
@@ -32,11 +32,10 @@ let error_not_match l _ =
let error_no_such_label l = error ("No such label "^string_of_label l)
-let error_no_such_label_sub l l1 l2 =
- let l1 = string_of_msid l1 in
- let l2 = string_of_msid l2 in
- error (l1^" is not a subtype of "^l2^".\nThe field "^
- string_of_label l^" is missing (or invisible) in "^l1^".")
+let error_no_such_label_sub l l1 =
+ let l1 = string_of_mp l1 in
+ error ("The field "^
+ string_of_label l^" is missing in "^l1^".")
let error_not_a_module_loc loc s =
user_err_loc (loc,"",str ("\""^string_of_label s^"\" is not a module"))
@@ -56,38 +55,6 @@ let error_signature_expected mtb =
let error_application_to_not_path _ = error "Application to not path"
-
-let module_body_of_type mtb =
- { mod_type = Some mtb.typ_expr;
- mod_expr = None;
- mod_constraints = Constraint.empty;
- mod_alias = mtb.typ_alias;
- mod_retroknowledge = []}
-
-let module_type_of_module mp mb =
- {typ_expr =
- (match mb.mod_type with
- | Some expr -> expr
- | None -> (match mb.mod_expr with
- | Some expr -> expr
- | None ->
- anomaly "Modops: empty expr and type"));
- typ_alias = mb.mod_alias;
- typ_strength = mp
- }
-
-
-
-let rec list_split_assoc k rev_before = function
- | [] -> raise Not_found
- | (k',b)::after when k=k' -> rev_before,b,after
- | h::tail -> list_split_assoc k (h::rev_before) tail
-
-let path_of_seb = function
- | SEBident mp -> mp
- | _ -> anomaly "Modops: evaluation failed."
-
-
let destr_functor env mtb =
match mtb with
| SEBfunctor (arg_id,arg_t,body_t) ->
@@ -95,254 +62,157 @@ let destr_functor env mtb =
| _ -> error_not_a_functor mtb
-let rec check_modpath_equiv env mp1 mp2 =
- if mp1=mp2 then () else
- let mp1 = scrape_alias mp1 env in
- let mp2 = scrape_alias mp2 env in
- if mp1=mp2 then ()
- else
- error_not_equal mp1 mp2
-
-
-
-let strengthen_const env mp l cb =
- match cb.const_opaque, cb.const_body with
- | false, Some _ -> cb
- | true, Some _
- | _, None ->
- let const = Const (make_con mp empty_dirpath l) in
- let const_subs = Some (Declarations.from_val const) in
- {cb with
- const_body = const_subs;
- const_opaque = false
- }
-
-let strengthen_mind env mp l mib = match mib.mind_equiv with
- | Some _ -> mib
- | None -> {mib with mind_equiv = Some (make_kn mp empty_dirpath l)}
-
+let is_functor = function
+ | SEBfunctor (arg_id,arg_t,body_t) -> true
+ | _ -> false
-let rec eval_struct env = function
- | SEBident mp ->
- begin
- let mp = scrape_alias mp env in
- let mtb =lookup_modtype mp env in
- match mtb.typ_expr,mtb.typ_strength with
- mtb,None -> eval_struct env mtb
- | mtb,Some mp -> strengthen_mtb env mp (eval_struct env mtb)
- end
- | SEBapply (seb1,seb2,_) ->
- let svb1 = eval_struct env seb1 in
- let farg_id, farg_b, fbody_b = destr_functor env svb1 in
- let mp = path_of_seb seb2 in
- let mp = scrape_alias mp env in
- let sub_alias = (lookup_modtype mp env).typ_alias in
- let sub_alias = match eval_struct env (SEBident mp) with
- | SEBstruct (msid,sign) -> subst_key (map_msid msid mp) sub_alias
- | _ -> sub_alias in
- let sub_alias = update_subst_alias sub_alias
- (map_mbid farg_id mp) in
- eval_struct env (subst_struct_expr
- (join sub_alias (map_mbid farg_id mp)) fbody_b)
- | SEBwith (mtb,(With_definition_body _ as wdb)) ->
- merge_with env mtb wdb empty_subst
- | SEBwith (mtb, (With_module_body (_,mp,_,_) as wdb)) ->
- let alias_in_mp =
- (lookup_modtype mp env).typ_alias in
- merge_with env mtb wdb alias_in_mp
-(* | SEBfunctor(mbid,mtb,body) ->
- let env = add_module (MPbound mbid) (module_body_of_type mtb) env in
- SEBfunctor(mbid,mtb,eval_struct env body) *)
- | mtb -> mtb
-
-and type_of_mb env mb =
- match mb.mod_type,mb.mod_expr with
- None,Some b -> eval_struct env b
- | Some t, _ -> eval_struct env t
- | _,_ -> anomaly
- "Modops: empty type and empty expr"
+let module_body_of_type mp mtb =
+ { mod_mp = mp;
+ mod_type = mtb.typ_expr;
+ mod_type_alg = mtb.typ_expr_alg;
+ mod_expr = None;
+ mod_constraints = mtb.typ_constraints;
+ mod_delta = mtb.typ_delta;
+ mod_retroknowledge = []}
-and merge_with env mtb with_decl alias=
- let msid,sig_b = match (eval_struct env mtb) with
- | SEBstruct(msid,sig_b) -> msid,sig_b
- | _ -> error_signature_expected mtb
- in
- let id,idl = match with_decl with
- | With_definition_body (id::idl,_) | With_module_body (id::idl,_,_,_) -> id,idl
- | With_definition_body ([],_) | With_module_body ([],_,_,_) -> assert false
- in
- let l = label_of_id id in
- try
- let rev_before,spec,after = list_split_assoc l [] sig_b in
- let before = List.rev rev_before in
- let rec mp_rec = function
- | [] -> MPself msid
- | i::r -> MPdot(mp_rec r,label_of_id i)
- in
- let new_spec,subst = match with_decl with
- | With_definition_body ([],_)
- | With_module_body ([],_,_,_) -> assert false
- | With_definition_body ([id],c) ->
- SFBconst c,None
- | With_module_body ([id], mp,typ_opt,cst) ->
- let mp' = scrape_alias mp env in
- SFBalias (mp,typ_opt,Some cst),
- Some(join (map_mp (mp_rec [id]) mp') alias)
- | With_definition_body (_::_,_)
- | With_module_body (_::_,_,_,_) ->
- let old = match spec with
- SFBmodule msb -> msb
- | _ -> error_not_a_module l
- in
- let new_with_decl,subst1 =
- match with_decl with
- With_definition_body (_,c) -> With_definition_body (idl,c),None
- | With_module_body (idc,mp,t,cst) ->
- With_module_body (idl,mp,t,cst),
- Some(map_mp (mp_rec idc) mp)
- in
- let subst = Option.fold_right join subst1 alias in
- let modtype =
- merge_with env (type_of_mb env old) new_with_decl alias in
- let msb =
- { mod_expr = None;
- mod_type = Some modtype;
- mod_constraints = old.mod_constraints;
- mod_alias = subst;
- mod_retroknowledge = old.mod_retroknowledge}
- in
- (SFBmodule msb),Some subst
- in
- SEBstruct(msid, before@(l,new_spec)::
- (Option.fold_right subst_structure subst after))
- with
- Not_found -> error_no_such_label l
+let check_modpath_equiv env mp1 mp2 =
+ if mp1=mp2 then () else
+ (* let mb1=lookup_module mp1 env in
+ let mb2=lookup_module mp2 env in
+ if (delta_of_mp mb1.mod_delta mp1)=(delta_of_mp mb2.mod_delta mp2)
+ then ()
+ else*) error_not_equal mp1 mp2
-and add_signature mp sign env =
+let rec add_signature mp sign resolver env =
let add_one env (l,elem) =
let kn = make_kn mp empty_dirpath l in
- let con = make_con mp empty_dirpath l in
+ let con = constant_of_kn kn in
+ let mind = mind_of_kn kn in
match elem with
- | SFBconst cb -> Environ.add_constant con cb env
- | SFBmind mib -> Environ.add_mind kn mib env
- | SFBmodule mb ->
- add_module (MPdot (mp,l)) mb env
+ | SFBconst cb ->
+ (* let con = constant_of_delta resolver con in*)
+ Environ.add_constant con cb env
+ | SFBmind mib ->
+ (* let mind = mind_of_delta resolver mind in*)
+ Environ.add_mind mind mib env
+ | SFBmodule mb -> add_module mb env
(* adds components as well *)
- | SFBalias (mp1,_,cst) ->
- Environ.register_alias (MPdot(mp,l)) mp1 env
- | SFBmodtype mtb -> Environ.add_modtype (MPdot(mp,l))
- mtb env
+ | SFBmodtype mtb -> Environ.add_modtype mtb.typ_mp mtb env
in
List.fold_left add_one env sign
-and add_module mp mb env =
+and add_module mb env =
+ let mp = mb.mod_mp in
let env = Environ.shallow_add_module mp mb env in
- let env =
- Environ.add_modtype mp (module_type_of_module (Some mp) mb) env
- in
- let mod_typ = type_of_mb env mb in
- match mod_typ with
- | SEBstruct (msid,sign) ->
- add_signature mp (subst_signature_msid msid mp sign) env
+ match mb.mod_type with
+ | SEBstruct (sign) ->
+ add_signature mp sign mb.mod_delta env
| SEBfunctor _ -> env
| _ -> anomaly "Modops:the evaluation of the structure failed "
-
-and constants_of_specification env mp sign =
- let aux (env,res) (l,elem) =
- match elem with
- | SFBconst cb -> env,((make_con mp empty_dirpath l),cb)::res
- | SFBmind _ -> env,res
- | SFBmodule mb ->
- let new_env = add_module (MPdot (mp,l)) mb env in
- new_env,(constants_of_modtype env (MPdot (mp,l))
- (type_of_mb env mb)) @ res
- | SFBalias (mp1,_,cst) ->
- let new_env = register_alias (MPdot (mp,l)) mp1 env in
- new_env,(constants_of_modtype env (MPdot (mp,l))
- (eval_struct env (SEBident mp1))) @ res
- | SFBmodtype mtb ->
- (* module type dans un module type.
- Il faut au moins mettre mtb dans l'environnement (avec le bon
- kn pour pouvoir continuer aller deplier les modules utilisant ce
- mtb
- ex:
- Module Type T1.
- Module Type T2.
- ....
- End T2.
- .....
- Declare Module M : T2.
- End T2
- si on ne rajoute pas T2 dans l'environement de typage
- on va exploser au moment du Declare Module
- *)
- let new_env = Environ.add_modtype (MPdot(mp,l)) mtb env in
- new_env, (constants_of_modtype env (MPdot(mp,l)) mtb.typ_expr) @ res
- in
- snd (List.fold_left aux (env,[]) sign)
-
-and constants_of_modtype env mp modtype =
- match (eval_struct env modtype) with
- SEBstruct (msid,sign) ->
- constants_of_specification env mp
- (subst_signature_msid msid mp sign)
- | SEBfunctor _ -> []
- | _ -> anomaly "Modops:the evaluation of the structure failed "
-
-and strengthen_mtb env mp mtb =
- let mtb1 = eval_struct env mtb in
- match mtb1 with
- | SEBfunctor _ -> mtb1
- | SEBstruct (msid,sign) ->
- SEBstruct (msid,strengthen_sig env msid sign mp)
- | _ -> anomaly "Modops:the evaluation of the structure failed "
-
-and strengthen_mod env mp mb =
- let mod_typ = type_of_mb env mb in
- { mod_expr = mb.mod_expr;
- mod_type = Some (strengthen_mtb env mp mod_typ);
- mod_constraints = mb.mod_constraints;
- mod_alias = mb.mod_alias;
- mod_retroknowledge = mb.mod_retroknowledge}
-
-and strengthen_sig env msid sign mp = match sign with
- | [] -> []
- | (l,SFBconst cb) :: rest ->
- let item' = l,SFBconst (strengthen_const env mp l cb) in
- let rest' = strengthen_sig env msid rest mp in
- item'::rest'
- | (l,SFBmind mib) :: rest ->
- let item' = l,SFBmind (strengthen_mind env mp l mib) in
- let rest' = strengthen_sig env msid rest mp in
- item'::rest'
- | (l,SFBmodule mb) :: rest ->
- let mp' = MPdot (mp,l) in
- let item' = l,SFBmodule (strengthen_mod env mp' mb) in
- let env' = add_module
- (MPdot (MPself msid,l)) mb env in
- let rest' = strengthen_sig env' msid rest mp in
+let strengthen_const env mp_from l cb resolver =
+ match cb.const_opaque, cb.const_body with
+ | false, Some _ -> cb
+ | true, Some _
+ | _, None ->
+ let con = make_con mp_from empty_dirpath l in
+ (* let con = constant_of_delta resolver con in*)
+ let const = Const con in
+ let const_subs = Some (Declarations.from_val const) in
+ {cb with
+ const_body = const_subs;
+ const_opaque = false;
+ }
+
+
+let rec strengthen_mod env mp_from mp_to mb =
+ assert(mp_from = mb.mod_mp);
+(* if mp_in_delta mb.mod_mp mb.mod_delta then
+ mb
+ else*)
+ match mb.mod_type with
+ | SEBstruct (sign) ->
+ let resolve_out,sign_out =
+ strengthen_sig env mp_from sign mp_to mb.mod_delta in
+ { mb with
+ mod_expr = Some (SEBident mp_to);
+ mod_type = SEBstruct(sign_out);
+ mod_type_alg = mb.mod_type_alg;
+ mod_constraints = mb.mod_constraints;
+ mod_delta = resolve_out(*add_mp_delta_resolver mp_from mp_to
+ (add_delta_resolver mb.mod_delta resolve_out)*);
+ mod_retroknowledge = mb.mod_retroknowledge}
+ | SEBfunctor _ -> mb
+ | _ -> anomaly "Modops:the evaluation of the structure failed "
+
+and strengthen_sig env mp_from sign mp_to resolver =
+ match sign with
+ | [] -> empty_delta_resolver,[]
+ | (l,SFBconst cb) :: rest ->
+ let item' =
+ l,SFBconst (strengthen_const env mp_from l cb resolver) in
+ let resolve_out,rest' =
+ strengthen_sig env mp_from rest mp_to resolver in
+ resolve_out,item'::rest'
+ | (_,SFBmind _ as item):: rest ->
+ let resolve_out,rest' =
+ strengthen_sig env mp_from rest mp_to resolver in
+ resolve_out,item::rest'
+ | (l,SFBmodule mb) :: rest ->
+ let mp_from' = MPdot (mp_from,l) in
+ let mp_to' = MPdot(mp_to,l) in
+ let mb_out =
+ strengthen_mod env mp_from' mp_to' mb in
+ let item' = l,SFBmodule (mb_out) in
+ let env' = add_module mb_out env in
+ let resolve_out,rest' =
+ strengthen_sig env' mp_from rest mp_to resolver in
+ resolve_out
+ (*add_delta_resolver resolve_out mb.mod_delta*),
item':: rest'
- | ((l,SFBalias (mp1,_,cst)) as item) :: rest ->
- let env' = register_alias (MPdot(MPself msid,l)) mp1 env in
- let rest' = strengthen_sig env' msid rest mp in
- item::rest'
- | (l,SFBmodtype mty as item) :: rest ->
- let env' = add_modtype
- (MPdot((MPself msid),l))
- mty
- env
- in
- let rest' = strengthen_sig env' msid rest mp in
- item::rest'
-
-
-let strengthen env mtb mp = strengthen_mtb env mp mtb
+ | (l,SFBmodtype mty as item) :: rest ->
+ let env' = add_modtype
+ (MPdot(mp_from,l)) mty env
+ in
+ let resolve_out,rest' =
+ strengthen_sig env' mp_from rest mp_to resolver in
+ resolve_out,item::rest'
+
+let strengthen env mtb mp =
+(* if mp_in_delta mtb.typ_mp mtb.typ_delta then
+ (* in this case mtb has already been strengthened*)
+ mtb
+ else*)
+ match mtb.typ_expr with
+ | SEBstruct (sign) ->
+ let resolve_out,sign_out =
+ strengthen_sig env mtb.typ_mp sign mp mtb.typ_delta in
+ {mtb with
+ typ_expr = SEBstruct(sign_out);
+ typ_delta = resolve_out(*add_delta_resolver mtb.typ_delta
+ (add_mp_delta_resolver mtb.typ_mp mp resolve_out)*)}
+ | SEBfunctor _ -> mtb
+ | _ -> anomaly "Modops:the evaluation of the structure failed "
-let update_subst env mb mp =
- match type_of_mb env mb with
- | SEBstruct(msid,str) -> false, join_alias
- (subst_key (map_msid msid mp) mb.mod_alias)
- (map_msid msid mp)
- | _ -> true, mb.mod_alias
+let subst_and_strengthen mb mp env =
+ strengthen_mod env mb.mod_mp mp
+ (subst_module (map_mp mb.mod_mp mp) mb)
+
+
+let module_type_of_module env mp mb =
+ match mp with
+ Some mp ->
+ strengthen env {
+ typ_mp = mp;
+ typ_expr = mb.mod_type;
+ typ_expr_alg = None;
+ typ_constraints = mb.mod_constraints;
+ typ_delta = mb.mod_delta} mp
+
+ | None ->
+ {typ_mp = mb.mod_mp;
+ typ_expr = mb.mod_type;
+ typ_expr_alg = None;
+ typ_constraints = mb.mod_constraints;
+ typ_delta = mb.mod_delta}
diff --git a/checker/modops.mli b/checker/modops.mli
index d5c9f4de6..4476013c6 100644
--- a/checker/modops.mli
+++ b/checker/modops.mli
@@ -19,33 +19,25 @@ open Environ
(* Various operations on modules and module types *)
-(* make the environment entry out of type *)
-val module_body_of_type : module_type_body -> module_body
+(* make the envirconment entry out of type *)
+val module_body_of_type : module_path -> module_type_body -> module_body
-val module_type_of_module : module_path option -> module_body ->
+val module_type_of_module : env -> module_path option -> module_body ->
module_type_body
val destr_functor :
env -> struct_expr_body -> mod_bound_id * module_type_body * struct_expr_body
-(* Evaluation functions *)
-val eval_struct : env -> struct_expr_body -> struct_expr_body
-
-val type_of_mb : env -> module_body -> struct_expr_body
-
-(* [add_signature mp sign env] assumes that the substitution [msid]
- $\mapsto$ [mp] has already been performed (or is not necessary, like
- when [mp = MPself msid]) *)
-val add_signature : module_path -> structure_body -> env -> env
+val add_signature : module_path -> structure_body -> delta_resolver -> env -> env
(* adds a module and its components, but not the constraints *)
-val add_module : module_path -> module_body -> env -> env
+val add_module : module_body -> env -> env
val check_modpath_equiv : env -> module_path -> module_path -> unit
-val strengthen : env -> struct_expr_body -> module_path -> struct_expr_body
+val strengthen : env -> module_type_body -> module_path -> module_type_body
-val update_subst : env -> module_body -> module_path -> bool * substitution
+val subst_and_strengthen : module_body -> module_path -> env -> module_body
val error_incompatible_modtypes :
module_type_body -> module_type_body -> 'a
@@ -57,7 +49,7 @@ val error_with_incorrect : label -> 'a
val error_no_such_label : label -> 'a
val error_no_such_label_sub :
- label -> mod_self_id -> mod_self_id -> 'a
+ label -> module_path -> 'a
val error_signature_expected : struct_expr_body -> 'a
diff --git a/checker/safe_typing.ml b/checker/safe_typing.ml
index b0d683ff3..da0e63c23 100644
--- a/checker/safe_typing.ml
+++ b/checker/safe_typing.ml
@@ -30,9 +30,8 @@ let set_engagement c =
(* full_add_module adds module with universes and constraints *)
let full_add_module dp mb digest =
let env = !genv in
- let mp = MPfile dp in
let env = add_constraints mb.mod_constraints env in
- let env = Modops.add_module mp mb env in
+ let env = Modops.add_module mb env in
genv := add_digest env dp digest
(* Check that the engagement expected by a library matches the initial one *)
@@ -66,16 +65,16 @@ let check_imports f caller env needed =
(* Remove the body of opaque constants in modules *)
-(* also remove mod_expr ? *)
+(* also remove mod_expr ? Good idea!*)
let rec lighten_module mb =
{ mb with
mod_expr = Option.map lighten_modexpr mb.mod_expr;
- mod_type = Option.map lighten_modexpr mb.mod_type }
+ mod_type = lighten_modexpr mb.mod_type }
and lighten_struct struc =
let lighten_body (l,body) = (l,match body with
| SFBconst ({const_opaque=true} as x) -> SFBconst {x with const_body=None}
- | (SFBconst _ | SFBmind _ | SFBalias _) as x -> x
+ | (SFBconst _ | SFBmind _ ) as x -> x
| SFBmodule m -> SFBmodule (lighten_module m)
| SFBmodtype m -> SFBmodtype
({m with
@@ -90,8 +89,8 @@ and lighten_modexpr = function
typ_expr = lighten_modexpr mty.typ_expr}),
lighten_modexpr mexpr)
| SEBident mp as x -> x
- | SEBstruct (msid, struc) ->
- SEBstruct (msid, lighten_struct struc)
+ | SEBstruct ( struc) ->
+ SEBstruct ( lighten_struct struc)
| SEBapply (mexpr,marg,u) ->
SEBapply (lighten_modexpr mexpr,lighten_modexpr marg,u)
| SEBwith (seb,wdcl) ->
@@ -125,7 +124,7 @@ let import file (dp,mb,depends,engmt as vo) digest =
let env = !genv in
check_imports msg_warning dp env depends;
check_engagement env engmt;
- check_module (add_constraints mb.mod_constraints env) mb;
+ check_module (add_constraints mb.mod_constraints env) mb.mod_mp mb;
stamp_library file digest;
(* We drop proofs once checked *)
(* let mb = lighten_module mb in*)
diff --git a/checker/subtyping.ml b/checker/subtyping.ml
index 88989b32e..ec1c908a6 100644
--- a/checker/subtyping.ml
+++ b/checker/subtyping.ml
@@ -32,7 +32,6 @@ type namedobject =
| IndConstr of constructor * mutual_inductive_body
| Module of module_body
| Modtype of module_type_body
- | Alias of module_path * struct_expr_body option
(* adds above information about one mutual inductive: all types and
constructors *)
@@ -60,13 +59,13 @@ let make_label_map mp list =
match e with
| SFBconst cb -> add_map (Constant cb)
| SFBmind mib ->
- add_nameobjects_of_mib (make_kn mp empty_dirpath l) mib map
+ add_nameobjects_of_mib (make_mind mp empty_dirpath l) mib map
| SFBmodule mb -> add_map (Module mb)
| SFBmodtype mtb -> add_map (Modtype mtb)
- | SFBalias (mp,typ_opt,cst) -> add_map (Alias (mp,typ_opt))
in
List.fold_right add_one list Labmap.empty
+
let check_conv_error error f env a1 a2 =
try
f env a1 a2
@@ -74,8 +73,8 @@ let check_conv_error error f env a1 a2 =
NotConvertible -> error ()
(* for now we do not allow reorderings *)
-let check_inductive env msid1 l info1 mib2 spec2 =
- let kn = make_kn (MPself msid1) empty_dirpath l in
+let check_inductive env mp1 l info1 mib2 spec2 subst1 subst2=
+ let kn = make_mind mp1 empty_dirpath l in
let error () = error_not_match l spec2 in
let check_conv f = check_conv_error error f in
let mib1 =
@@ -83,6 +82,7 @@ let check_inductive env msid1 l info1 mib2 spec2 =
| IndType ((_,0), mib) -> mib
| _ -> error ()
in
+ let mib2 = subst_mind subst2 mib2 in
let check_inductive_type env t1 t2 =
(* Due to sort-polymorphism in inductive types, the conclusions of
@@ -155,7 +155,7 @@ let check_inductive env msid1 l info1 mib2 spec2 =
(* the inductive types and constructors types have to be convertible *)
check (fun mib -> mib.mind_nparams);
- begin
+ (*begin
match mib2.mind_equiv with
| None -> ()
| Some kn2' ->
@@ -165,7 +165,7 @@ let check_inductive env msid1 l info1 mib2 spec2 =
| Some kn1' -> scrape_mind env kn1'
in
if kn1 <> kn2 then error ()
- end;
+ end;*)
(* we check that records and their field names are preserved. *)
check (fun mib -> mib.mind_record);
if mib1.mind_record then begin
@@ -187,7 +187,7 @@ let check_inductive env msid1 l info1 mib2 spec2 =
let _ = array_map2_i check_cons_types mib1.mind_packets mib2.mind_packets
in ()
-let check_constant env msid1 l info1 cb2 spec2 =
+let check_constant env mp1 l info1 cb2 spec2 subst1 subst2 =
let error () = error_not_match l spec2 in
let check_conv f = check_conv_error error f in
let check_type env t1 t2 =
@@ -236,30 +236,31 @@ let check_constant env msid1 l info1 cb2 spec2 =
(t1,t2) in
check_conv conv_leq env t1 t2
in
-
- match info1 with
- | Constant cb1 ->
- assert (cb1.const_hyps=[] && cb2.const_hyps=[]) ;
- (*Start by checking types*)
- let typ1 = Typeops.type_of_constant_type env cb1.const_type in
- let typ2 = Typeops.type_of_constant_type env cb2.const_type in
- check_type env typ1 typ2;
- let con = make_con (MPself msid1) empty_dirpath l in
- (match cb2 with
- | {const_body=Some lc2;const_opaque=false} ->
- let c2 = force_constr lc2 in
- let c1 = match cb1.const_body with
- | Some lc1 -> force_constr lc1
- | None -> Const con
- in
- check_conv conv env c1 c2
- | _ -> ())
- | IndType ((kn,i),mind1) ->
- ignore (Util.error (
- "The kernel does not recognize yet that a parameter can be " ^
- "instantiated by an inductive type. Hint: you can rename the " ^
- "inductive type and give a definition to map the old name to the new " ^
- "name."));
+ match info1 with
+ | Constant cb1 ->
+ assert (cb1.const_hyps=[] && cb2.const_hyps=[]) ;
+ (*Start by checking types*)
+ let cb1 = subst_const_body subst1 cb1 in
+ let cb2 = subst_const_body subst2 cb2 in
+ let typ1 = Typeops.type_of_constant_type env cb1.const_type in
+ let typ2 = Typeops.type_of_constant_type env cb2.const_type in
+ check_type env typ1 typ2;
+ let con = make_con mp1 empty_dirpath l in
+ (match cb2 with
+ | {const_body=Some lc2;const_opaque=false} ->
+ let c2 = force_constr lc2 in
+ let c1 = match cb1.const_body with
+ | Some lc1 -> force_constr lc1
+ | None -> Const con
+ in
+ check_conv conv env c1 c2
+ | _ -> ())
+ | IndType ((kn,i),mind1) ->
+ ignore (Util.error (
+ "The kernel does not recognize yet that a parameter can be " ^
+ "instantiated by an inductive type. Hint: you can rename the " ^
+ "inductive type and give a definition to map the old name to the new " ^
+ "name."));
assert (mind1.mind_hyps=[] && cb2.const_hyps=[]) ;
if cb2.const_body <> None then error () ;
let arity1 = type_of_inductive env (mind1,mind1.mind_packets.(i)) in
@@ -278,57 +279,31 @@ let check_constant env msid1 l info1 cb2 spec2 =
check_conv conv env ty1 ty2
| _ -> error ()
-let rec check_modules env msid1 l msb1 msb2 =
- let mp = (MPdot(MPself msid1,l)) in
- let mty1 = module_type_of_module (Some mp) msb1 in
- let mty2 = module_type_of_module None msb2 in
- check_modtypes env mty1 mty2 false
+let rec check_modules env msb1 msb2 subst1 subst2 =
+ let mty1 = module_type_of_module env None msb1 in
+ let mty2 = module_type_of_module env None msb2 in
+ check_modtypes env mty1 mty2 subst1 subst2 false;
+
-
-and check_signatures env (msid1,sig1) alias (msid2,sig2') =
- let mp1 = MPself msid1 in
- let env = add_signature mp1 sig1 env in
- let alias = update_subst_alias alias (map_msid msid2 mp1) in
- let sig2 = subst_structure alias sig2' in
- let sig2 = subst_signature_msid msid2 mp1 sig2 in
+and check_signatures env mp1 sig1 sig2 subst1 subst2 =
let map1 = make_label_map mp1 sig1 in
- let check_one_body (l,spec2) =
+ let check_one_body (l,spec2) =
let info1 =
try
Labmap.find l map1
with
- Not_found -> error_no_such_label_sub l msid1 msid2
+ Not_found -> error_no_such_label_sub l mp1
in
match spec2 with
| SFBconst cb2 ->
- check_constant env msid1 l info1 cb2 spec2
+ check_constant env mp1 l info1 cb2 spec2 subst1 subst2
| SFBmind mib2 ->
- check_inductive env msid1 l info1 mib2 spec2
+ check_inductive env mp1 l info1 mib2 spec2 subst1 subst2
| SFBmodule msb2 ->
begin
match info1 with
- | Module msb -> check_modules env msid1 l msb msb2
- | Alias (mp,typ_opt) ->let msb =
- {mod_expr = Some (SEBident mp);
- mod_type = typ_opt;
- mod_constraints = Constraint.empty;
- mod_alias = (lookup_modtype mp env).typ_alias;
- mod_retroknowledge = []} in
- check_modules env msid1 l msb msb2
- | _ -> error_not_match l spec2
- end
- | SFBalias (mp,typ_opt,_) ->
- begin
- match info1 with
- | Alias (mp1,_) -> check_modpath_equiv env mp mp1
- | Module msb ->
- let msb1 =
- {mod_expr = Some (SEBident mp);
- mod_type = typ_opt;
- mod_constraints = Constraint.empty;
- mod_alias = (lookup_modtype mp env).typ_alias;
- mod_retroknowledge = []} in
- check_modules env msid1 l msb msb1
+ | Module msb -> check_modules env msb msb2
+ subst1 subst2
| _ -> error_not_match l spec2
end
| SFBmodtype mtb2 ->
@@ -337,52 +312,61 @@ and check_signatures env (msid1,sig1) alias (msid2,sig2') =
| Modtype mtb -> mtb
| _ -> error_not_match l spec2
in
- check_modtypes env mtb1 mtb2 true
+ let env = add_module (module_body_of_type mtb2.typ_mp mtb2)
+ (add_module (module_body_of_type mtb1.typ_mp mtb1) env) in
+ check_modtypes env mtb1 mtb2 subst1 subst2 true
in
List.iter check_one_body sig2
-and check_modtypes env mtb1 mtb2 equiv =
- if mtb1==mtb2 then () else (* just in case :) *)
- let mtb1',mtb2'=
- (match mtb1.typ_strength with
- None -> eval_struct env mtb1.typ_expr,
- eval_struct env mtb2.typ_expr
- | Some mp -> strengthen env mtb1.typ_expr mp,
- eval_struct env mtb2.typ_expr) in
- let rec check_structure env str1 str2 equiv =
- match str1, str2 with
- | SEBstruct (msid1,list1),
- SEBstruct (msid2,list2) ->
- check_signatures env
- (msid1,list1) mtb1.typ_alias (msid2,list2);
- if equiv then
- check_signatures env
- (msid2,list2) mtb2.typ_alias (msid1,list1)
+and check_modtypes env mtb1 mtb2 subst1 subst2 equiv =
+ if mtb1==mtb2 then () else
+ let mtb1',mtb2'=mtb1.typ_expr,mtb2.typ_expr in
+ let rec check_structure env str1 str2 equiv subst1 subst2 =
+ match str1,str2 with
+ | SEBstruct (list1),
+ SEBstruct (list2) ->
+ check_signatures env
+ mtb1.typ_mp list1 list2 subst1 subst2;
+ if equiv then
+ check_signatures env
+ mtb2.typ_mp list2 list1 subst1 subst2
+ else
+ ()
| SEBfunctor (arg_id1,arg_t1,body_t1),
SEBfunctor (arg_id2,arg_t2,body_t2) ->
- check_modtypes env arg_t2 arg_t1 equiv;
+ check_modtypes env
+ arg_t2 arg_t1
+ (map_mp arg_t1.typ_mp arg_t2.typ_mp) subst2
+ equiv ;
(* contravariant *)
- let env =
- add_module (MPbound arg_id2) (module_body_of_type arg_t2) env
+ let env = add_module
+ (module_body_of_type (MPbound arg_id2) arg_t2) env
in
- let body_t1' =
- (* since we are just checking well-typedness we do not need
- to expand any constant. Hence the identity resolver. *)
- subst_struct_expr
- (map_mbid arg_id1 (MPbound arg_id2))
- body_t1
+ let env = match body_t1 with
+ SEBstruct str ->
+ add_module {mod_mp = mtb1.typ_mp;
+ mod_expr = None;
+ mod_type = body_t1;
+ mod_type_alg= None;
+ mod_constraints=mtb1.typ_constraints;
+ mod_retroknowledge = [];
+ mod_delta = mtb1.typ_delta} env
+ | _ -> env
in
- check_structure env (eval_struct env body_t1')
- (eval_struct env body_t2) equiv
+ check_structure env body_t1 body_t2 equiv
+ (join (map_mbid arg_id1 (MPbound arg_id2)) subst1)
+ subst2
| _ , _ -> error_incompatible_modtypes mtb1 mtb2
in
if mtb1'== mtb2' then ()
- else check_structure env mtb1' mtb2' equiv
+ else check_structure env mtb1' mtb2' equiv subst1 subst2
let check_subtypes env sup super =
(*if sup<>super then*)
- check_modtypes env sup super false
+ check_modtypes env (strengthen env sup sup.typ_mp) super empty_subst
+ (map_mp super.typ_mp sup.typ_mp) false
let check_equal env sup super =
(*if sup<>super then*)
- check_modtypes env sup super true
+ check_modtypes env sup super empty_subst
+ (map_mp super.typ_mp sup.typ_mp) true
diff --git a/checker/term.ml b/checker/term.ml
index 92d898b31..658eac4f0 100644
--- a/checker/term.ml
+++ b/checker/term.ml
@@ -489,7 +489,7 @@ let compare_constr f t1 t2 =
f h1 h2 & List.for_all2 f l1 l2
else false
| Evar (e1,l1), Evar (e2,l2) -> e1 = e2 & array_for_all2 f l1 l2
- | Const c1, Const c2 -> c1 = c2
+ | Const c1, Const c2 -> eq_constant c1 c2
| Ind c1, Ind c2 -> c1 = c2
| Construct c1, Construct c2 -> c1 = c2
| Case (_,p1,c1,bl1), Case (_,p2,c2,bl2) ->
diff --git a/checker/typeops.ml b/checker/typeops.ml
index 3a4f2f825..e5cf6a6d6 100644
--- a/checker/typeops.ml
+++ b/checker/typeops.ml
@@ -228,7 +228,7 @@ let judge_of_inductive_knowing_parameters env ind (paramstyp:constr array) =
let (mib,mip) =
try lookup_mind_specif env ind
with Not_found ->
- failwith ("Cannot find inductive: "^string_of_kn (fst ind)) in
+ failwith ("Cannot find inductive: "^string_of_mind (fst ind)) in
check_args env c mib.mind_hyps;
type_of_inductive_knowing_parameters env mip paramstyp
@@ -244,7 +244,7 @@ let judge_of_constructor env c =
let mib =
try lookup_mind kn env
with Not_found ->
- failwith ("Cannot find inductive: "^string_of_kn (fst (fst c))) in
+ failwith ("Cannot find inductive: "^string_of_mind (fst (fst c))) in
check_args env constr mib.mind_hyps in
let specif = lookup_mind_specif env (inductive_of_constructor c) in
type_of_constructor c specif
diff --git a/dev/db b/dev/db
index 818785701..a355f62a9 100644
--- a/dev/db
+++ b/dev/db
@@ -11,6 +11,7 @@ install_printer Top_printers.ppdir
install_printer Top_printers.ppmp
install_printer Top_printers.ppkn
install_printer Top_printers.ppcon
+install_printer Top_printers.ppmind
install_printer Top_printers.ppsp
install_printer Top_printers.ppqualid
install_printer Top_printers.ppclindex
@@ -39,5 +40,5 @@ install_printer Top_printers.pptac
install_printer Top_printers.ppobj
install_printer Top_printers.pploc
install_printer Top_printers.prsubst
-
-
+install_printer Top_printers.prdelta
+install_printer Top_printers.print_pure_constr
diff --git a/dev/printers.mllib b/dev/printers.mllib
index 107b2904a..b5a54a189 100644
--- a/dev/printers.mllib
+++ b/dev/printers.mllib
@@ -54,9 +54,9 @@ Mod_typing
Safe_typing
Summary
-Global
Nameops
Libnames
+Global
Nametab
Libobject
Lib
diff --git a/dev/top_printers.ml b/dev/top_printers.ml
index d5ebde7cb..310fb1188 100644
--- a/dev/top_printers.ml
+++ b/dev/top_printers.ml
@@ -28,6 +28,7 @@ open Cerrors
open Evd
open Goptions
open Genarg
+open Mod_subst
let _ = Constrextern.print_evar_arguments := true
@@ -44,8 +45,9 @@ let ppmsid msid = pp (str (debug_string_of_msid msid))
let ppmbid mbid = pp (str (debug_string_of_mbid mbid))
let ppdir dir = pp (pr_dirpath dir)
let ppmp mp = pp(str (string_of_mp mp))
-let ppcon con = pp(pr_con con)
+let ppcon con = pp(debug_pr_con con)
let ppkn kn = pp(pr_kn kn)
+let ppmind kn = pp(debug_pr_mind kn)
let ppsp sp = pp(pr_path sp)
let ppqualid qid = pp(pr_qualid qid)
let ppclindex cl = pp(Classops.pr_cl_index cl)
@@ -72,10 +74,10 @@ let ppidset l = pp (prset pr_id (Idset.elements l))
let pP s = pp (hov 0 s)
let safe_pr_global = function
- | ConstRef kn -> pp (str "CONSTREF(" ++ pr_con kn ++ str ")")
- | IndRef (kn,i) -> pp (str "INDREF(" ++ pr_kn kn ++ str "," ++
+ | ConstRef kn -> pp (str "CONSTREF(" ++ debug_pr_con kn ++ str ")")
+ | IndRef (kn,i) -> pp (str "INDREF(" ++ debug_pr_mind kn ++ str "," ++
int i ++ str ")")
- | ConstructRef ((kn,i),j) -> pp (str "INDREF(" ++ pr_kn kn ++ str "," ++
+ | ConstructRef ((kn,i),j) -> pp (str "INDREF(" ++ debug_pr_mind kn ++ str "," ++
int i ++ str "," ++ int j ++ str ")")
| VarRef id -> pp (str "VARREF(" ++ pr_id id ++ str ")")
@@ -92,6 +94,7 @@ let genppj f j = let (c,t) = f j in (c ++ str " : " ++ t)
let ppj j = pp (genppj pr_ljudge j)
let prsubst s = pp (Mod_subst.debug_pr_subst s)
+let prdelta s = pp (Mod_subst.debug_pr_delta s)
let pp_idpred s = pp (pr_idpred s)
let pp_cpred s = pp (pr_cpred s)
@@ -159,9 +162,9 @@ let constr_display csr =
| Evar (e,l) -> "Evar("^(string_of_int e)^","^(array_display l)^")"
| Const c -> "Const("^(string_of_con c)^")"
| Ind (sp,i) ->
- "MutInd("^(string_of_kn sp)^","^(string_of_int i)^")"
+ "MutInd("^(string_of_mind sp)^","^(string_of_int i)^")"
| Construct ((sp,i),j) ->
- "MutConstruct(("^(string_of_kn sp)^","^(string_of_int i)^"),"
+ "MutConstruct(("^(string_of_mind sp)^","^(string_of_int i)^"),"
^(string_of_int j)^")"
| Case (ci,p,c,bl) ->
"MutCase(<abs>,"^(term_display p)^","^(term_display c)^","
@@ -309,7 +312,7 @@ let print_pure_constr csr =
| ("Coq"::_::l) -> l
| l -> l
in List.iter (fun x -> print_string x; print_string ".") ls;*)
- print_string (string_of_kn sp)
+ print_string (debug_string_of_mind sp)
and sp_con_display sp =
(* let dir,l = decode_kn sp in
let ls =
@@ -318,7 +321,7 @@ let print_pure_constr csr =
| ("Coq"::_::l) -> l
| l -> l
in List.iter (fun x -> print_string x; print_string ".") ls;*)
- print_string (string_of_con sp)
+ print_string (debug_string_of_con sp)
in
try
@@ -442,11 +445,11 @@ let raw_string_of_ref loc = function
let (mp,dir,id) = repr_con cst in
encode_path loc "CST" (Some (mp,dir)) [] (id_of_label id)
| IndRef (kn,i) ->
- let (mp,dir,id) = repr_kn kn in
+ let (mp,dir,id) = repr_mind kn in
encode_path loc "IND" (Some (mp,dir)) [id_of_label id]
(id_of_string ("_"^string_of_int i))
| ConstructRef ((kn,i),j) ->
- let (mp,dir,id) = repr_kn kn in
+ let (mp,dir,id) = repr_mind kn in
encode_path loc "CSTR" (Some (mp,dir))
[id_of_label id;id_of_string ("_"^string_of_int i)]
(id_of_string ("_"^string_of_int j))
@@ -456,13 +459,13 @@ let raw_string_of_ref loc = function
let short_string_of_ref loc = function
| VarRef id -> Ident (loc,id)
| ConstRef cst -> Ident (loc,id_of_label (pi3 (repr_con cst)))
- | IndRef (kn,0) -> Ident (loc,id_of_label (pi3 (repr_kn kn)))
+ | IndRef (kn,0) -> Ident (loc,id_of_label (pi3 (repr_mind kn)))
| IndRef (kn,i) ->
- encode_path loc "IND" None [id_of_label (pi3 (repr_kn kn))]
+ encode_path loc "IND" None [id_of_label (pi3 (repr_mind kn))]
(id_of_string ("_"^string_of_int i))
| ConstructRef ((kn,i),j) ->
encode_path loc "CSTR" None
- [id_of_label (pi3 (repr_kn kn));id_of_string ("_"^string_of_int i)]
+ [id_of_label (pi3 (repr_mind kn));id_of_string ("_"^string_of_int i)]
(id_of_string ("_"^string_of_int j))
let _ = Constrextern.set_debug_global_reference_printer
diff --git a/dev/vm_printers.ml b/dev/vm_printers.ml
index 266bd1043..59545d8aa 100644
--- a/dev/vm_printers.ml
+++ b/dev/vm_printers.ml
@@ -10,7 +10,7 @@ let ppripos (ri,pos) =
| Reloc_annot a ->
let sp,i = a.ci.ci_ind in
print_string
- ("annot : MutInd("^(string_of_kn sp)^","^(string_of_int i)^")\n")
+ ("annot : MutInd("^(string_of_mind sp)^","^(string_of_int i)^")\n")
| Reloc_const _ ->
print_string "structured constant\n"
| Reloc_getglobal kn ->
@@ -62,7 +62,7 @@ and ppatom a =
| Aid idk -> print_idkey idk
| Aiddef(idk,_) -> print_string "&";print_idkey idk
| Aind(sp,i) -> print_string "Ind(";
- print_string (string_of_kn sp);
+ print_string (string_of_mind sp);
print_string ","; print_int i;
print_string ")"
diff --git a/interp/coqlib.ml b/interp/coqlib.ml
index b44cabe8b..90432a4a4 100644
--- a/interp/coqlib.ml
+++ b/interp/coqlib.ml
@@ -66,14 +66,20 @@ let gen_constant_in_modules locstr dirs s =
let check_required_library d =
let d' = List.map id_of_string d in
let dir = make_dirpath (List.rev d') in
+ let mp = (fst(Lib.current_prefix())) in
+ let current_dir = match mp with
+ | MPfile dp -> (dir=dp)
+ | _ -> false
+ in
if not (Library.library_is_loaded dir) then
+ if not current_dir then
(* Loading silently ...
let m, prefix = list_sep_last d' in
read_library
(dummy_loc,make_qualid (make_dirpath (List.rev prefix)) m)
*)
(* or failing ...*)
- error ("Library "^(string_of_dirpath dir)^" has to be required first.")
+ error ("Library "^(string_of_dirpath dir)^" has to be required first.")
(************************************************************************)
(* Specific Coq objects *)
@@ -123,7 +129,7 @@ let jmeq_module_name = ["Coq";"Logic";"JMeq"]
let jmeq_module = make_dir jmeq_module_name
(* TODO: temporary hack *)
-let make_kn dir id = Libnames.encode_kn dir id
+let make_kn dir id = Libnames.encode_mind dir id
let make_con dir id = Libnames.encode_con dir id
(** Identity *)
diff --git a/interp/notation.ml b/interp/notation.ml
index 342cb6c9c..8f9f0959d 100644
--- a/interp/notation.ml
+++ b/interp/notation.ml
@@ -109,7 +109,7 @@ let open_scope i (_,(local,op,sc)) =
let cache_scope o =
open_scope 1 o
-let subst_scope (_,subst,sc) = sc
+let subst_scope (subst,sc) = sc
open Libobject
@@ -180,19 +180,29 @@ type key =
let notations_key_table = ref Gmapl.empty
let prim_token_key_table = Hashtbl.create 7
+
+let make_gr = function
+ | ConstRef con ->
+ ConstRef(constant_of_kn(canonical_con con))
+ | IndRef (kn,i) ->
+ IndRef(mind_of_kn(canonical_mind kn),i)
+ | ConstructRef ((kn,i),j )->
+ ConstructRef((mind_of_kn(canonical_mind kn),i),j)
+ | VarRef id -> VarRef id
+
let rawconstr_key = function
- | RApp (_,RRef (_,ref),_) -> RefKey ref
- | RRef (_,ref) -> RefKey ref
+ | RApp (_,RRef (_,ref),_) -> RefKey (make_gr ref)
+ | RRef (_,ref) -> RefKey (make_gr ref)
| _ -> Oth
let cases_pattern_key = function
- | PatCstr (_,ref,_,_) -> RefKey (ConstructRef ref)
+ | PatCstr (_,ref,_,_) -> RefKey (make_gr (ConstructRef ref))
| _ -> Oth
let aconstr_key = function (* Rem: AApp(ARef ref,[]) stands for @ref *)
- | AApp (ARef ref,args) -> RefKey ref, Some (List.length args)
- | AList (_,_,AApp (ARef ref,args),_,_) -> RefKey ref, Some (List.length args)
- | ARef ref -> RefKey ref, None
+ | AApp (ARef ref,args) -> RefKey(make_gr ref), Some (List.length args)
+ | AList (_,_,AApp (ARef ref,args),_,_) -> RefKey (make_gr ref), Some (List.length args)
+ | ARef ref -> RefKey(make_gr ref), None
| _ -> Oth, None
(**********************************************************************)
@@ -458,7 +468,7 @@ let load_arguments_scope _ (_,(_,r,scl)) =
let cache_arguments_scope o =
load_arguments_scope 1 o
-let subst_arguments_scope (_,subst,(req,r,scl)) =
+let subst_arguments_scope (subst,(req,r,scl)) =
(ArgsScopeNoDischarge,fst (subst_global subst r),scl)
let discharge_arguments_scope (_,(req,r,l)) =
diff --git a/interp/syntax_def.ml b/interp/syntax_def.ml
index 939fe01aa..1330a3689 100644
--- a/interp/syntax_def.ml
+++ b/interp/syntax_def.ml
@@ -51,7 +51,7 @@ let open_syntax_constant i ((sp,kn),(_,pat,onlyparse)) =
let cache_syntax_constant d =
load_syntax_constant 1 d
-let subst_syntax_constant ((sp,kn),subst,(local,pat,onlyparse)) =
+let subst_syntax_constant (subst,(local,pat,onlyparse)) =
(local,subst_interpretation subst pat,onlyparse)
let classify_syntax_constant (local,_,_ as o) =
diff --git a/interp/topconstr.ml b/interp/topconstr.ml
index bea0eae31..ec1b20a12 100644
--- a/interp/topconstr.ml
+++ b/interp/topconstr.ml
@@ -143,7 +143,7 @@ let has_ldots =
(function RApp (_,RVar(_,v),_) when v = ldots_var -> true | _ -> false)
let compare_rawconstr f t1 t2 = match t1,t2 with
- | RRef (_,r1), RRef (_,r2) -> r1 = r2
+ | RRef (_,r1), RRef (_,r2) -> eq_gr r1 r2
| RVar (_,v1), RVar (_,v2) -> v1 = v2
| RApp (_,f1,l1), RApp (_,f2,l2) -> f f1 f2 & List.for_all2 f l1 l2
| RLambda (_,na1,bk1,ty1,c1), RLambda (_,na2,bk2,ty2,c2) when na1 = na2 && bk1 = bk2 ->
@@ -281,7 +281,7 @@ let rec subst_pat subst pat =
match pat with
| PatVar _ -> pat
| PatCstr (loc,((kn,i),j),cpl,n) ->
- let kn' = subst_kn subst kn
+ let kn' = subst_ind subst kn
and cpl' = list_smartmap (subst_pat subst) cpl in
if kn' == kn && cpl' == cpl then pat else
PatCstr (loc,((kn',i),j),cpl',n)
@@ -331,7 +331,7 @@ let rec subst_aconstr subst bound raw =
(fun (a,(n,signopt) as x) ->
let a' = subst_aconstr subst bound a in
let signopt' = Option.map (fun ((indkn,i),n,nal as z) ->
- let indkn' = subst_kn subst indkn in
+ let indkn' = subst_ind subst indkn in
if indkn == indkn' then z else ((indkn',i),n,nal)) signopt in
if a' == a && signopt' == signopt then x else (a',(n,signopt')))
rl
@@ -485,7 +485,7 @@ let adjust_application_n n loc f l =
let rec match_ alp metas sigma a1 a2 = match (a1,a2) with
| r1, AVar id2 when List.mem id2 metas -> bind_env alp sigma id2 r1
| RVar (_,id1), AVar id2 when alpha_var id1 id2 alp -> sigma
- | RRef (_,r1), ARef r2 when r1 = r2 -> sigma
+ | RRef (_,r1), ARef r2 when (eq_gr r1 r2) -> sigma
| RPatVar (_,(_,n1)), APatVar n2 when n1=n2 -> sigma
| RApp (loc,f1,l1), AApp (f2,l2) ->
let n1 = List.length l1 and n2 = List.length l2 in
diff --git a/kernel/cbytegen.ml b/kernel/cbytegen.ml
index a7e8b0b26..e7859962e 100644
--- a/kernel/cbytegen.ml
+++ b/kernel/cbytegen.ml
@@ -723,7 +723,10 @@ let compile_constant_body env body opaque boxed =
BCdefined(true, to_patch)
else
match kind_of_term body with
- | Const kn' -> BCallias (get_allias env kn')
+ | Const kn' ->
+ (* we use the canonical name of the constant*)
+ let con= constant_of_kn (canonical_con kn') in
+ BCallias (get_allias env con)
| _ ->
let res = compile env body in
let to_patch = to_memory res in
diff --git a/kernel/cemitcodes.ml b/kernel/cemitcodes.ml
index 89264e88b..4a9c7da2b 100644
--- a/kernel/cemitcodes.ml
+++ b/kernel/cemitcodes.ml
@@ -304,13 +304,13 @@ let rec subst_strcst s sc =
match sc with
| Const_sorts _ | Const_b0 _ -> sc
| Const_bn(tag,args) -> Const_bn(tag,Array.map (subst_strcst s) args)
- | Const_ind(ind) -> let kn,i = ind in Const_ind((subst_kn s kn, i))
+ | Const_ind(ind) -> let kn,i = ind in Const_ind((subst_ind s kn, i))
let subst_patch s (ri,pos) =
match ri with
| Reloc_annot a ->
let (kn,i) = a.ci.ci_ind in
- let ci = {a.ci with ci_ind = (subst_kn s kn,i)} in
+ let ci = {a.ci with ci_ind = (subst_ind s kn,i)} in
(Reloc_annot {a with ci = ci},pos)
| Reloc_const sc -> (Reloc_const (subst_strcst s sc), pos)
| Reloc_getglobal kn -> (Reloc_getglobal (fst (subst_con s kn)), pos)
diff --git a/kernel/closure.ml b/kernel/closure.ml
index bce564397..93788ed42 100644
--- a/kernel/closure.ml
+++ b/kernel/closure.ml
@@ -196,6 +196,8 @@ let unfold_red kn =
type table_key = id_key
+let eq_table_key = Names.eq_id_key
+
type 'a infos = {
i_flags : reds;
i_repr : 'a infos -> constr -> 'a;
@@ -250,18 +252,6 @@ let defined_rels flags env =
(rel_context env) ~init:(0,[])
(* else (0,[])*)
-let rec mind_equiv env (kn1,i1) (kn2,i2) =
- let rec equiv kn1 kn2 =
- kn1 = kn2 ||
- match (lookup_mind kn1 env).mind_equiv with
- Some kn1' -> equiv kn2 kn1'
- | None -> match (lookup_mind kn2 env).mind_equiv with
- Some kn2' -> equiv kn2' kn1
- | None -> false in
- i1 = i2 && equiv kn1 kn2
-
-let mind_equiv_infos info = mind_equiv info.i_env
-
let create mk_cl flgs env evars =
{ i_flags = flgs;
i_repr = mk_cl;
diff --git a/kernel/closure.mli b/kernel/closure.mli
index b6ff1fa15..5cb6fc97c 100644
--- a/kernel/closure.mli
+++ b/kernel/closure.mli
@@ -179,9 +179,7 @@ val whd_stack :
(* [unfold_reference] unfolds references in a [fconstr] *)
val unfold_reference : clos_infos -> table_key -> fconstr option
-(* [mind_equiv] checks whether two inductive types are intentionally equal *)
-val mind_equiv : env -> inductive -> inductive -> bool
-val mind_equiv_infos : clos_infos -> inductive -> inductive -> bool
+val eq_table_key : table_key -> table_key -> bool
(************************************************************************)
(*i This is for lazy debug *)
diff --git a/kernel/cooking.ml b/kernel/cooking.ml
index e42a732d3..c971ed299 100644
--- a/kernel/cooking.ml
+++ b/kernel/cooking.ml
@@ -19,15 +19,15 @@ open Reduction
(*s Cooking the constants. *)
-type work_list = identifier array Cmap.t * identifier array KNmap.t
+type work_list = identifier array Cmap.t * identifier array Mindmap.t
let pop_dirpath p = match repr_dirpath p with
| [] -> anomaly "dirpath_prefix: empty dirpath"
| _::l -> make_dirpath l
-let pop_kn kn =
- let (mp,dir,l) = Names.repr_kn kn in
- Names.make_kn mp (pop_dirpath dir) l
+let pop_mind kn =
+ let (mp,dir,l) = Names.repr_mind kn in
+ Names.make_mind mp (pop_dirpath dir) l
let pop_con con =
let (mp,dir,l) = Names.repr_con con in
@@ -48,9 +48,9 @@ let share r (cstl,knl) =
let f,l =
match r with
| IndRef (kn,i) ->
- mkInd (pop_kn kn,i), KNmap.find kn knl
+ mkInd (pop_mind kn,i), Mindmap.find kn knl
| ConstructRef ((kn,i),j) ->
- mkConstruct ((pop_kn kn,i),j), KNmap.find kn knl
+ mkConstruct ((pop_mind kn,i),j), Mindmap.find kn knl
| ConstRef cst ->
mkConst (pop_con cst), Cmap.find cst cstl in
let c = mkApp (f, Array.map mkVar l) in
@@ -69,7 +69,7 @@ let update_case_info ci modlist =
with Not_found ->
ci
-let empty_modlist = (Cmap.empty, KNmap.empty)
+let empty_modlist = (Cmap.empty, Mindmap.empty)
let expmod_constr modlist c =
let rec substrec c =
diff --git a/kernel/cooking.mli b/kernel/cooking.mli
index 23b1f2534..db35031d9 100644
--- a/kernel/cooking.mli
+++ b/kernel/cooking.mli
@@ -16,7 +16,7 @@ open Univ
(*s Cooking the constants. *)
-type work_list = identifier array Cmap.t * identifier array KNmap.t
+type work_list = identifier array Cmap.t * identifier array Mindmap.t
type recipe = {
d_from : constant_body;
diff --git a/kernel/declarations.ml b/kernel/declarations.ml
index c48c01d78..515009798 100644
--- a/kernel/declarations.ml
+++ b/kernel/declarations.ml
@@ -69,7 +69,7 @@ type recarg =
let subst_recarg sub r = match r with
| Norec | Mrec _ -> r
- | Imbr (kn,i) -> let kn' = subst_kn sub kn in
+ | Imbr (kn,i) -> let kn' = subst_ind sub kn in
if kn==kn' then r else Imbr (kn',i)
type wf_paths = recarg Rtree.t
@@ -190,25 +190,25 @@ type mutual_inductive_body = {
(* Universes constraints enforced by the inductive declaration *)
mind_constraints : constraints;
- (* Source of the inductive block when aliased in a module *)
- mind_equiv : kernel_name option
}
-let subst_arity sub = function
-| NonPolymorphicType s -> NonPolymorphicType (subst_mps sub s)
-| PolymorphicArity (ctx,s) -> PolymorphicArity (subst_rel_context sub ctx,s)
-
+let subst_arity sub arity =
+ if sub = empty_subst then arity
+ else match arity with
+ | NonPolymorphicType s -> NonPolymorphicType (subst_mps sub s)
+ | PolymorphicArity (ctx,s) -> PolymorphicArity (subst_rel_context sub ctx,s)
+
(* TODO: should be changed to non-coping after Term.subst_mps *)
let subst_const_body sub cb = {
- const_hyps = (assert (cb.const_hyps=[]); []);
- const_body = Option.map (subst_constr_subst sub) cb.const_body;
- const_type = subst_arity sub cb.const_type;
- const_body_code = Cemitcodes.subst_to_patch_subst sub cb.const_body_code;
- (*const_type_code = Cemitcodes.subst_to_patch sub cb.const_type_code;*)
- const_constraints = cb.const_constraints;
- const_opaque = cb.const_opaque;
- const_inline = cb.const_inline}
-
+ const_hyps = (assert (cb.const_hyps=[]); []);
+ const_body = Option.map (subst_constr_subst sub) cb.const_body;
+ const_type = subst_arity sub cb.const_type;
+ const_body_code = Cemitcodes.subst_to_patch_subst sub cb.const_body_code;
+ (*const_type_code = Cemitcodes.subst_to_patch sub cb.const_type_code;*)
+ const_constraints = cb.const_constraints;
+ const_opaque = cb.const_opaque;
+ const_inline = cb.const_inline}
+
let subst_arity sub = function
| Monomorphic s ->
Monomorphic {
@@ -244,8 +244,7 @@ let subst_mind sub mib =
mind_params_ctxt =
map_rel_context (subst_mps sub) mib.mind_params_ctxt;
mind_packets = array_smartmap (subst_mind_packet sub) mib.mind_packets ;
- mind_constraints = mib.mind_constraints ;
- mind_equiv = Option.map (subst_kn sub) mib.mind_equiv }
+ mind_constraints = mib.mind_constraints }
(*s Modules: signature component specifications, module types, and
@@ -255,8 +254,6 @@ type structure_field_body =
| SFBconst of constant_body
| SFBmind of mutual_inductive_body
| SFBmodule of module_body
- | SFBalias of module_path * struct_expr_body option
- * constraints option
| SFBmodtype of module_type_body
and structure_body = (label * structure_field_body) list
@@ -264,24 +261,26 @@ and structure_body = (label * structure_field_body) list
and struct_expr_body =
| SEBident of module_path
| SEBfunctor of mod_bound_id * module_type_body * struct_expr_body
- | SEBstruct of mod_self_id * structure_body
- | SEBapply of struct_expr_body * struct_expr_body
- * constraints
+ | SEBapply of struct_expr_body * struct_expr_body * constraints
+ | SEBstruct of structure_body
| SEBwith of struct_expr_body * with_declaration_body
and with_declaration_body =
With_module_body of identifier list * module_path
- * struct_expr_body option * constraints
| With_definition_body of identifier list * constant_body
and module_body =
- { mod_expr : struct_expr_body option;
- mod_type : struct_expr_body option;
+ { mod_mp : module_path;
+ mod_expr : struct_expr_body option;
+ mod_type : struct_expr_body;
+ mod_type_alg : struct_expr_body option;
mod_constraints : constraints;
- mod_alias : substitution;
+ mod_delta : delta_resolver;
mod_retroknowledge : Retroknowledge.action list}
and module_type_body =
- { typ_expr : struct_expr_body;
- typ_strength : module_path option;
- typ_alias : substitution}
+ { typ_mp : module_path;
+ typ_expr : struct_expr_body;
+ typ_expr_alg : struct_expr_body option ;
+ typ_constraints : constraints;
+ typ_delta :delta_resolver}
diff --git a/kernel/declarations.mli b/kernel/declarations.mli
index c7e27db6b..adf1d14e2 100644
--- a/kernel/declarations.mli
+++ b/kernel/declarations.mli
@@ -171,8 +171,6 @@ type mutual_inductive_body = {
(* Universes constraints enforced by the inductive declaration *)
mind_constraints : constraints;
- (* Source of the inductive block when aliased in a module *)
- mind_equiv : kernel_name option
}
val subst_mind : substitution -> mutual_inductive_body -> mutual_inductive_body
@@ -185,8 +183,6 @@ type structure_field_body =
| SFBconst of constant_body
| SFBmind of mutual_inductive_body
| SFBmodule of module_body
- | SFBalias of module_path * struct_expr_body option
- * constraints option
| SFBmodtype of module_type_body
and structure_body = (label * structure_field_body) list
@@ -194,24 +190,38 @@ and structure_body = (label * structure_field_body) list
and struct_expr_body =
| SEBident of module_path
| SEBfunctor of mod_bound_id * module_type_body * struct_expr_body
- | SEBstruct of mod_self_id * structure_body
- | SEBapply of struct_expr_body * struct_expr_body
- * constraints
+ | SEBapply of struct_expr_body * struct_expr_body * constraints
+ | SEBstruct of structure_body
| SEBwith of struct_expr_body * with_declaration_body
and with_declaration_body =
With_module_body of identifier list * module_path
- * struct_expr_body option * constraints
| With_definition_body of identifier list * constant_body
and module_body =
- { mod_expr : struct_expr_body option;
- mod_type : struct_expr_body option;
+ { (*absolute path of the module*)
+ mod_mp : module_path;
+ (* Implementation *)
+ mod_expr : struct_expr_body option;
+ (* Signature *)
+ mod_type : struct_expr_body;
+ (* algebraic structure expression is kept
+ if it's relevant for extraction *)
+ mod_type_alg : struct_expr_body option;
+ (* set of all constraint in the module *)
mod_constraints : constraints;
- mod_alias : substitution;
+ (* quotiented set of equivalent constant and inductive name *)
+ mod_delta : delta_resolver;
mod_retroknowledge : Retroknowledge.action list}
-
+
and module_type_body =
- { typ_expr : struct_expr_body;
- typ_strength : module_path option;
- typ_alias : substitution}
+ {
+ (*Path of the module type*)
+ typ_mp : module_path;
+ typ_expr : struct_expr_body;
+ (* algebraic structure expression is kept
+ if it's relevant for extraction *)
+ typ_expr_alg : struct_expr_body option ;
+ typ_constraints : constraints;
+ (* quotiented set of equivalent constant and inductive name *)
+ typ_delta :delta_resolver}
diff --git a/kernel/entries.ml b/kernel/entries.ml
index 26e9a6250..938d1c60a 100644
--- a/kernel/entries.ml
+++ b/kernel/entries.ml
@@ -62,7 +62,8 @@ type definition_entry = {
const_entry_opaque : bool;
const_entry_boxed : bool}
-type parameter_entry = types*bool
+(* type and the inlining flag *)
+type parameter_entry = types * bool
type constant_entry =
| DefinitionEntry of definition_entry
@@ -70,7 +71,14 @@ type constant_entry =
(*s Modules *)
-type module_struct_entry =
+
+type specification_entry =
+ SPEconst of constant_entry
+ | SPEmind of mutual_inductive_entry
+ | SPEmodule of module_entry
+ | SPEmodtype of module_struct_entry
+
+and module_struct_entry =
MSEident of module_path
| MSEfunctor of mod_bound_id * module_struct_entry * module_struct_entry
| MSEwith of module_struct_entry * with_declaration
diff --git a/kernel/entries.mli b/kernel/entries.mli
index 291ff0d45..20fbbb8e7 100644
--- a/kernel/entries.mli
+++ b/kernel/entries.mli
@@ -61,7 +61,7 @@ type definition_entry = {
const_entry_opaque : bool;
const_entry_boxed : bool }
-type parameter_entry = types*bool (*inline flag*)
+type parameter_entry = types * bool (*inline flag*)
type constant_entry =
| DefinitionEntry of definition_entry
@@ -69,7 +69,14 @@ type constant_entry =
(*s Modules *)
-type module_struct_entry =
+
+type specification_entry =
+ SPEconst of constant_entry
+ | SPEmind of mutual_inductive_entry
+ | SPEmodule of module_entry
+ | SPEmodtype of module_struct_entry
+
+and module_struct_entry =
MSEident of module_path
| MSEfunctor of mod_bound_id * module_struct_entry * module_struct_entry
| MSEwith of module_struct_entry * with_declaration
diff --git a/kernel/environ.ml b/kernel/environ.ml
index fb51660b3..a233d0be1 100644
--- a/kernel/environ.ml
+++ b/kernel/environ.ml
@@ -141,7 +141,7 @@ let lookup_constant = lookup_constant
let add_constant kn cs env =
let new_constants =
- Cmap.add kn (cs,ref None) env.env_globals.env_constants in
+ Cmap_env.add kn (cs,ref None) env.env_globals.env_constants in
let new_globals =
{ env.env_globals with
env_constants = new_constants } in
@@ -174,11 +174,9 @@ let evaluable_constant cst env =
(* Mutual Inductives *)
let lookup_mind = lookup_mind
-let scrape_mind = scrape_mind
-
-
+
let add_mind kn mib env =
- let new_inds = KNmap.add kn mib env.env_globals.env_inductives in
+ let new_inds = Mindmap_env.add kn mib env.env_globals.env_inductives in
let new_globals =
{ env.env_globals with
env_inductives = new_inds } in
@@ -276,30 +274,12 @@ let shallow_add_module mp mb env =
env_modules = new_mods } in
{ env with env_globals = new_globals }
-let rec scrape_alias mp env =
- try
- let mp1 = MPmap.find mp env.env_globals.env_alias in
- scrape_alias mp1 env
- with
- Not_found -> mp
-
let lookup_module mp env =
- let mp = scrape_alias mp env in
MPmap.find mp env.env_globals.env_modules
-let lookup_modtype ln env =
- let mp = scrape_alias ln env in
- MPmap.find mp env.env_globals.env_modtypes
-
-let register_alias mp1 mp2 env =
- let new_alias = MPmap.add mp1 mp2 env.env_globals.env_alias in
- let new_globals =
- { env.env_globals with
- env_alias = new_alias } in
- { env with env_globals = new_globals }
-let lookup_alias mp env =
- MPmap.find mp env.env_globals.env_alias
+let lookup_modtype mp env =
+ MPmap.find mp env.env_globals.env_modtypes
(*s Judgments. *)
diff --git a/kernel/environ.mli b/kernel/environ.mli
index 0ae285528..9401ba6b0 100644
--- a/kernel/environ.mli
+++ b/kernel/environ.mli
@@ -142,9 +142,6 @@ val add_mind : mutual_inductive -> mutual_inductive_body -> env -> env
(* raises [Not_found] if the required path is not found *)
val lookup_mind : mutual_inductive -> env -> mutual_inductive_body
-(* Find the ultimate inductive in the [mind_equiv] chain *)
-val scrape_mind : env -> mutual_inductive -> mutual_inductive
-
(************************************************************************)
(*s Modules *)
val add_modtype : module_path -> module_type_body -> env -> env
@@ -155,10 +152,6 @@ val shallow_add_module : module_path -> module_body -> env -> env
val lookup_module : module_path -> env -> module_body
val lookup_modtype : module_path -> env -> module_type_body
-val register_alias : module_path -> module_path -> env -> env
-val lookup_alias : module_path -> env -> module_path
-val scrape_alias : module_path -> env -> module_path
-
(************************************************************************)
(*s Universe constraints *)
val set_universes : Univ.universes -> env -> env
diff --git a/kernel/indtypes.ml b/kernel/indtypes.ml
index c202d627d..ffb6be771 100644
--- a/kernel/indtypes.ml
+++ b/kernel/indtypes.ml
@@ -651,8 +651,7 @@ let build_inductive env env_ar params isrecord isfinite inds nmr recargs cst =
mind_nparams_rec = nmr;
mind_params_ctxt = params;
mind_packets = packets;
- mind_constraints = cst;
- mind_equiv = None;
+ mind_constraints = cst
}
(************************************************************************)
diff --git a/kernel/inductive.ml b/kernel/inductive.ml
index 19e4130ff..8252f2d5a 100644
--- a/kernel/inductive.ml
+++ b/kernel/inductive.ml
@@ -362,20 +362,10 @@ let type_case_branches env (ind,largs) pj c =
(************************************************************************)
(* Checking the case annotation is relevent *)
-let rec inductive_kn_equiv env kn1 kn2 =
- match (lookup_mind kn1 env).mind_equiv with
- | Some kn1' -> inductive_kn_equiv env kn2 kn1'
- | None -> match (lookup_mind kn2 env).mind_equiv with
- | Some kn2' -> inductive_kn_equiv env kn2' kn1
- | None -> false
-
-let inductive_equiv env (kn1,i1) (kn2,i2) =
- i1=i2 & inductive_kn_equiv env kn1 kn2
-
let check_case_info env indsp ci =
let (mib,mip) = lookup_mind_specif env indsp in
if
- not (Closure.mind_equiv env indsp ci.ci_ind) or
+ not (eq_ind indsp ci.ci_ind) or
(mib.mind_nparams <> ci.ci_npar) or
(mip.mind_consnrealdecls <> ci.ci_cstr_nargs)
then raise (TypeError(env,WrongCaseInfo(indsp,ci)))
diff --git a/kernel/mod_subst.ml b/kernel/mod_subst.ml
index 238aa3544..a9d4a246d 100644
--- a/kernel/mod_subst.ml
+++ b/kernel/mod_subst.ml
@@ -13,27 +13,30 @@ open Util
open Names
open Term
-(* WARNING: not every constant in the associative list domain used to exist
- in the environment. This allows a simple implementation of the join
- operation. However, iterating over the associative list becomes a non-sense
-*)
-type resolver = (constant * constr option) list
-
-let make_resolver resolve = resolve
-
-let apply_opt_resolver resolve kn =
- match resolve with
- None -> None
- | Some resolve ->
- try List.assoc kn resolve with Not_found -> None
-
-type substitution_domain =
- MSI of mod_self_id
+
+type delta_hint =
+ Inline of constr option
+ | Equiv of kernel_name
+ | Prefix_equiv of module_path
+
+type delta_key =
+ KN of kernel_name
+ | MP of module_path
+
+module Deltamap = Map.Make(struct
+ type t = delta_key
+ let compare = Pervasives.compare
+ end)
+
+type delta_resolver = delta_hint Deltamap.t
+
+let empty_delta_resolver = Deltamap.empty
+
+type substitution_domain =
| MBI of mod_bound_id
| MPI of module_path
let string_of_subst_domain = function
- MSI msid -> debug_string_of_msid msid
| MBI mbid -> debug_string_of_mbid mbid
| MPI mp -> string_of_mp mp
@@ -42,48 +45,233 @@ module Umap = Map.Make(struct
let compare = Pervasives.compare
end)
-type substitution = (module_path * resolver option) Umap.t
-
+type substitution = (module_path * delta_resolver) Umap.t
+
let empty_subst = Umap.empty
-let add_msid msid mp =
- Umap.add (MSI msid) (mp,None)
+
+let string_of_subst_domain = function
+ | MBI mbid -> debug_string_of_mbid mbid
+ | MPI mp -> string_of_mp mp
+
let add_mbid mbid mp resolve =
Umap.add (MBI mbid) (mp,resolve)
-let add_mp mp1 mp2 =
- Umap.add (MPI mp1) (mp2,None)
+let add_mp mp1 mp2 resolve =
+ Umap.add (MPI mp1) (mp2,resolve)
-let map_msid msid mp = add_msid msid mp empty_subst
let map_mbid mbid mp resolve = add_mbid mbid mp resolve empty_subst
-let map_mp mp1 mp2 = add_mp mp1 mp2 empty_subst
+let map_mp mp1 mp2 resolve = add_mp mp1 mp2 resolve empty_subst
+
+let add_inline_delta_resolver con =
+ Deltamap.add (KN(user_con con)) (Inline None)
+
+let add_inline_constr_delta_resolver con cstr =
+ Deltamap.add (KN(user_con con)) (Inline (Some cstr))
+
+let add_constant_delta_resolver con =
+ Deltamap.add (KN(user_con con)) (Equiv (canonical_con con))
+
+let add_mind_delta_resolver mind =
+ Deltamap.add (KN(user_mind mind)) (Equiv (canonical_mind mind))
+
+let add_mp_delta_resolver mp1 mp2 =
+ Deltamap.add (MP mp1) (Prefix_equiv mp2)
+
+let mp_in_delta mp =
+ Deltamap.mem (MP mp)
+
+let con_in_delta con resolver =
+try
+ match Deltamap.find (KN(user_con con)) resolver with
+ | Inline _ | Prefix_equiv _ -> false
+ | Equiv _ -> true
+with
+ Not_found -> false
+
+let mind_in_delta mind resolver =
+try
+ match Deltamap.find (KN(user_mind mind)) resolver with
+ | Inline _ | Prefix_equiv _ -> false
+ | Equiv _ -> true
+with
+ Not_found -> false
+
+let delta_of_mp resolve mp =
+ try
+ match Deltamap.find (MP mp) resolve with
+ | Prefix_equiv mp1 -> mp1
+ | _ -> anomaly "mod_subst: bad association in delta_resolver"
+ with
+ Not_found -> mp
+
+let delta_of_kn resolve kn =
+ try
+ match Deltamap.find (KN kn) resolve with
+ | Equiv kn1 -> kn1
+ | Inline _ -> kn
+ | _ -> anomaly
+ "mod_subst: bad association in delta_resolver"
+ with
+ Not_found -> kn
+
+let remove_mp_delta_resolver resolver mp =
+ Deltamap.remove (MP mp) resolver
+
+exception Inline_kn
+
+let rec find_prefix resolve mp =
+ let rec sub_mp = function
+ | MPdot(mp,l) as mp_sup ->
+ (try
+ match Deltamap.find (MP mp_sup) resolve with
+ | Prefix_equiv mp1 -> mp1
+ | _ -> anomaly
+ "mod_subst: bad association in delta_resolver"
+ with
+ Not_found -> MPdot(sub_mp mp,l))
+ | p ->
+ match Deltamap.find (MP p) resolve with
+ | Prefix_equiv mp1 -> mp1
+ | _ -> anomaly
+ "mod_subst: bad association in delta_resolver"
+ in
+ try
+ sub_mp mp
+ with
+ Not_found -> mp
+
+let solve_delta_kn resolve kn =
+ try
+ match Deltamap.find (KN kn) resolve with
+ | Equiv kn1 -> kn1
+ | Inline _ -> raise Inline_kn
+ | _ -> anomaly
+ "mod_subst: bad association in delta_resolver"
+ with
+ Not_found | Inline_kn ->
+ let mp,dir,l = repr_kn kn in
+ let new_mp = find_prefix resolve mp in
+ if mp == new_mp then
+ kn
+ else
+ make_kn new_mp dir l
+
+
+let constant_of_delta resolve con =
+ let kn = user_con con in
+ let new_kn = solve_delta_kn resolve kn in
+ if kn == new_kn then
+ con
+ else
+ constant_of_kn_equiv kn new_kn
+
+let constant_of_delta2 resolve con =
+ let kn = canonical_con con in
+ let kn1 = user_con con in
+ let new_kn = solve_delta_kn resolve kn in
+ if kn == new_kn then
+ con
+ else
+ constant_of_kn_equiv kn1 new_kn
+
+let mind_of_delta resolve mind =
+ let kn = user_mind mind in
+ let new_kn = solve_delta_kn resolve kn in
+ if kn == new_kn then
+ mind
+ else
+ mind_of_kn_equiv kn new_kn
+
+let mind_of_delta2 resolve mind =
+ let kn = canonical_mind mind in
+ let kn1 = user_mind mind in
+ let new_kn = solve_delta_kn resolve kn in
+ if kn == new_kn then
+ mind
+ else
+ mind_of_kn_equiv kn1 new_kn
+
+
+
+let inline_of_delta resolver =
+ let extract key hint l =
+ match key,hint with
+ |KN kn, Inline _ -> kn::l
+ | _,_ -> l
+ in
+ Deltamap.fold extract resolver []
-let list_contents sub =
- let one_pair uid (mp,_) l =
- (string_of_subst_domain uid, string_of_mp mp)::l
+exception Not_inline
+
+let constant_of_delta_with_inline resolve con =
+ let kn1,kn2 = canonical_con con,user_con con in
+ try
+ match Deltamap.find (KN kn2) resolve with
+ | Inline None -> None
+ | Inline (Some const) -> Some const
+ | _ -> raise Not_inline
+ with
+ Not_found | Not_inline ->
+ try match Deltamap.find (KN kn1) resolve with
+ | Inline None -> None
+ | Inline (Some const) -> Some const
+ | _ -> raise Not_inline
+ with
+ Not_found | Not_inline -> None
+
+let string_of_key = function
+ | KN kn -> string_of_kn kn
+ | MP mp -> string_of_mp mp
+
+let string_of_hint = function
+ | Inline _ -> "inline"
+ | Equiv kn -> string_of_kn kn
+ | Prefix_equiv mp -> string_of_mp mp
+
+let debug_string_of_delta resolve =
+ let to_string key hint s =
+ s^", "^(string_of_key key)^"=>"^(string_of_hint hint)
+ in
+ Deltamap.fold to_string resolve ""
+
+let list_contents sub =
+ let one_pair uid (mp,reso) l =
+ (string_of_subst_domain uid, string_of_mp mp,debug_string_of_delta reso)::l
in
Umap.fold one_pair sub []
-
+
let debug_string_of_subst sub =
- let l = List.map (fun (s1,s2) -> s1^"|->"^s2) (list_contents sub) in
+ let l = List.map (fun (s1,s2,s3) -> s1^"|->"^s2^"["^s3^"]")
+ (list_contents sub) in
"{" ^ String.concat "; " l ^ "}"
+
+let debug_pr_delta resolve =
+ str (debug_string_of_delta resolve)
let debug_pr_subst sub =
let l = list_contents sub in
- let f (s1,s2) = hov 2 (str s1 ++ spc () ++ str "|-> " ++ str s2)
+ let f (s1,s2,s3) = hov 2 (str s1 ++ spc () ++ str "|-> " ++ str s2 ++
+ spc () ++ str "[" ++ str s3 ++ str "]")
in
str "{" ++ hov 2 (prlist_with_sep pr_coma f l) ++ str "}"
-
-
+
+
let subst_mp0 sub mp = (* 's like subst *)
let rec aux mp =
match mp with
- | MPself sid ->
- let mp',resolve = Umap.find (MSI sid) sub in
+ | MPfile sid ->
+ let mp',resolve = Umap.find (MPI (MPfile sid)) sub in
mp',resolve
| MPbound bid ->
- let mp',resolve = Umap.find (MBI bid) sub in
- mp',resolve
+ begin
+ try
+ let mp',resolve = Umap.find (MBI bid) sub in
+ mp',resolve
+ with Not_found ->
+ let mp',resolve = Umap.find (MPI mp) sub in
+ mp',resolve
+ end
| MPdot (mp1,l) as mp2 ->
begin
try
@@ -93,7 +281,6 @@ let subst_mp0 sub mp = (* 's like subst *)
let mp1',resolve = aux mp1 in
MPdot (mp1',l),resolve
end
- | _ -> raise Not_found
in
try
Some (aux mp)
@@ -104,39 +291,126 @@ let subst_mp sub mp =
None -> mp
| Some (mp',_) -> mp'
-
-let subst_kn0 sub kn =
+let subst_kn_delta sub kn =
let mp,dir,l = repr_kn kn in
match subst_mp0 sub mp with
- Some (mp',_) ->
- Some (make_kn mp' dir l)
- | None -> None
+ Some (mp',resolve) ->
+ solve_delta_kn resolve (make_kn mp' dir l)
+ | None -> kn
+
let subst_kn sub kn =
- match subst_kn0 sub kn with
- None -> kn
- | Some kn' -> kn'
+ let mp,dir,l = repr_kn kn in
+ match subst_mp0 sub mp with
+ Some (mp',_) ->
+ (make_kn mp' dir l)
+ | None -> kn
+
+exception No_subst
+
+type sideconstantsubst =
+ | User
+ | Canonical
+
+let subst_ind sub mind =
+ let kn1,kn2 = user_mind mind,canonical_mind mind in
+ let mp1,dir,l = repr_kn kn1 in
+ let mp2,_,_ = repr_kn kn2 in
+ try
+ let side,mind',resolve =
+ match subst_mp0 sub mp1,subst_mp0 sub mp2 with
+ None,None ->raise No_subst
+ | Some (mp',resolve),None -> User,(make_mind_equiv mp' mp2 dir l), resolve
+ | None, Some(mp',resolve)-> Canonical,(make_mind_equiv mp1 mp' dir l), resolve
+ | Some(mp1',resolve1),Some(mp2',resolve2)->Canonical,
+ (make_mind_equiv mp1' mp2' dir l), resolve2
+ in
+ match side with
+ |User ->
+ let mind = mind_of_delta resolve mind' in
+ mind
+ |Canonical ->
+ let mind = mind_of_delta2 resolve mind' in
+ mind
+ with
+ No_subst -> mind
+
+let subst_mind0 sub mind =
+ let kn1,kn2 = user_mind mind,canonical_mind mind in
+ let mp1,dir,l = repr_kn kn1 in
+ let mp2,_,_ = repr_kn kn2 in
+ try
+ let side,mind',resolve =
+ match subst_mp0 sub mp1,subst_mp0 sub mp2 with
+ None,None ->raise No_subst
+ | Some (mp',resolve),None -> User,(make_mind_equiv mp' mp2 dir l), resolve
+ | None, Some(mp',resolve)-> Canonical,(make_mind_equiv mp1 mp' dir l), resolve
+ | Some(mp1',resolve1),Some(mp2',resolve2)->Canonical,
+ (make_mind_equiv mp1' mp2' dir l), resolve2
+ in
+ match side with
+ |User ->
+ let mind = mind_of_delta resolve mind' in
+ Some mind
+ |Canonical ->
+ let mind = mind_of_delta2 resolve mind' in
+ Some mind
+ with
+ No_subst -> Some mind
let subst_con sub con =
- let mp,dir,l = repr_con con in
- match subst_mp0 sub mp with
- None -> con,mkConst con
- | Some (mp',resolve) ->
- let con' = make_con mp' dir l in
- match apply_opt_resolver resolve con with
- None -> con',mkConst con'
- | Some t -> con',t
+ let kn1,kn2 = user_con con,canonical_con con in
+ let mp1,dir,l = repr_kn kn1 in
+ let mp2,_,_ = repr_kn kn2 in
+ try
+ let side,con',resolve =
+ match subst_mp0 sub mp1,subst_mp0 sub mp2 with
+ None,None ->raise No_subst
+ | Some (mp',resolve),None -> User,(make_con_equiv mp' mp2 dir l), resolve
+ | None, Some(mp',resolve)-> Canonical,(make_con_equiv mp1 mp' dir l), resolve
+ | Some(mp1',resolve1),Some(mp2',resolve2)->Canonical,
+ (make_con_equiv mp1' mp2' dir l), resolve2
+ in
+ match constant_of_delta_with_inline resolve con' with
+ None -> begin
+ match side with
+ |User ->
+ let con = constant_of_delta resolve con' in
+ con,mkConst con
+ |Canonical ->
+ let con = constant_of_delta2 resolve con' in
+ con,mkConst con
+ end
+ | Some t -> con',t
+ with No_subst -> con , mkConst con
+
let subst_con0 sub con =
- let mp,dir,l = repr_con con in
- match subst_mp0 sub mp with
- None -> None
- | Some (mp',resolve) ->
- let con' = make_con mp' dir l in
- match apply_opt_resolver resolve con with
- None -> Some (mkConst con')
- | Some t -> Some t
-
+ let kn1,kn2 = user_con con,canonical_con con in
+ let mp1,dir,l = repr_kn kn1 in
+ let mp2,_,_ = repr_kn kn2 in
+ try
+ let side,con',resolve =
+ match subst_mp0 sub mp1,subst_mp0 sub mp2 with
+ None,None ->raise No_subst
+ | Some (mp',resolve),None -> User,(make_con_equiv mp' mp2 dir l), resolve
+ | None, Some(mp',resolve)-> Canonical,(make_con_equiv mp1 mp' dir l), resolve
+ | Some(mp1',resolve1),Some(mp2',resolve2)->Canonical,
+ (make_con_equiv mp1' mp2' dir l), resolve2
+ in
+ match constant_of_delta_with_inline resolve con' with
+ None ->begin
+ match side with
+ |User ->
+ let con = constant_of_delta resolve con' in
+ Some (mkConst con)
+ |Canonical ->
+ let con = constant_of_delta2 resolve con' in
+ Some (mkConst con)
+ end
+ | t -> t
+ with No_subst -> Some (mkConst con)
+
(* Here the semantics is completely unclear.
What does "Hint Unfold t" means when "t" is a parameter?
Does the user mean "Unfold X.t" or does she mean "Unfold y"
@@ -220,7 +494,7 @@ let rec map_kn f f' c =
| _ -> c
let subst_mps sub =
- map_kn (subst_kn0 sub) (subst_con0 sub)
+ map_kn (subst_mind0 sub) (subst_con0 sub)
let rec replace_mp_in_mp mpfrom mpto mp =
match mp with
@@ -231,233 +505,145 @@ let rec replace_mp_in_mp mpfrom mpto mp =
else MPdot (mp1',l)
| _ -> mp
-let replace_mp_in_con mpfrom mpto kn =
- let mp,dir,l = repr_con kn in
+let replace_mp_in_kn mpfrom mpto kn =
+ let mp,dir,l = repr_kn kn in
let mp'' = replace_mp_in_mp mpfrom mpto mp in
if mp==mp'' then kn
- else make_con mp'' dir l
-
-exception BothSubstitutionsAreIdentitySubstitutions
-exception ChangeDomain of resolver
-
-let join (subst1 : substitution) (subst2 : substitution) =
- let apply_subst (sub : substitution) key (mp,resolve) =
- let mp',resolve' =
- match subst_mp0 sub mp with
- None -> mp, None
- | Some (mp',resolve') -> mp',resolve' in
- let resolve'' : resolver option =
- try
- let res =
- match resolve with
- Some res -> res
- | None ->
- match resolve' with
- None -> raise BothSubstitutionsAreIdentitySubstitutions
- | Some res -> raise (ChangeDomain res)
- in
- Some
- (List.map
- (fun (kn,topt) ->
- kn,
- match topt with
- None ->
- (match key with
- MSI msid ->
- let kn' = replace_mp_in_con (MPself msid) mp kn in
- apply_opt_resolver resolve' kn'
- | MBI mbid ->
- let kn' = replace_mp_in_con (MPbound mbid) mp kn in
- apply_opt_resolver resolve' kn'
- | MPI mp1 ->
- let kn' = replace_mp_in_con mp1 mp kn in
- apply_opt_resolver resolve' kn')
- | Some t -> Some (subst_mps sub t)) res)
- with
- BothSubstitutionsAreIdentitySubstitutions -> None
- | ChangeDomain res ->
- let rec changeDom = function
- | [] -> []
- | (kn,topt)::r ->
- let key' =
- match key with
- MSI msid -> MPself msid
- | MBI mbid -> MPbound mbid
- | MPI mp1 -> mp1 in
- let kn' = replace_mp_in_con mp key' kn in
- if kn==kn' then
- (*the key does not appear in kn, we remove it
- from the resolver that we are building*)
- changeDom r
- else
- (kn',topt)::(changeDom r)
- in
- Some (changeDom res)
- in
- mp',resolve'' in
- let subst = Umap.mapi (apply_subst subst2) subst1 in
- (Umap.fold Umap.add subst2 subst)
-
-let subst_key subst1 subst2 =
- let replace_in_key key (mp,resolve) sub=
- let newkey =
- match key with
- | MPI mp1 ->
- begin
- match subst_mp0 subst1 mp1 with
- | None -> None
- | Some (mp2,_) -> Some (MPI mp2)
- end
- | _ -> None
- in
- match newkey with
- | None -> Umap.add key (mp,resolve) sub
- | Some mpi -> Umap.add mpi (mp,resolve) sub
+ else make_kn mp'' dir l
+
+let mp_in_key mp key =
+ let rec mp_rec mp1 =
+ match mp1 with
+ | _ when mp1 = mp -> true
+ | MPdot (mp2,l) -> mp_rec mp2
+ | _ -> false
+ in
+ match key with
+ | MP mp1 ->
+ mp_rec mp1
+ | KN kn ->
+ let mp1,dir,l = repr_kn kn in
+ mp_rec mp1
+
+let subset_prefixed_by mp resolver =
+ let prefixmp key hint resolv =
+ if mp_in_key mp key then
+ Deltamap.add key hint resolv
+ else
+ resolv
in
- Umap.fold replace_in_key subst2 empty_subst
-
-let update_subst_alias subst1 subst2 =
- let subst_inv key (mp,resolve) sub =
- let newmp =
- match key with
- | MBI msid -> MPbound msid
- | MSI msid -> MPself msid
- | MPI mp -> mp
- in
- match mp with
- | MPbound mbid -> Umap.add (MBI mbid) (newmp,None) sub
- | MPself msid -> Umap.add (MSI msid) (newmp,None) sub
- | _ -> Umap.add (MPI mp) (newmp,None) sub
+ Deltamap.fold prefixmp resolver empty_delta_resolver
+
+let subst_dom_delta_resolver subst resolver =
+ let apply_subst key hint resolver =
+ match key with
+ (MP mp) ->
+ Deltamap.add (MP (subst_mp subst mp)) hint resolver
+ | (KN kn) ->
+ Deltamap.add (KN (subst_kn subst kn)) hint resolver
in
- let subst_mbi = Umap.fold subst_inv subst2 empty_subst in
- let alias_subst key (mp,resolve) sub=
- let newkey =
- match key with
- | MPI mp1 ->
- begin
- match subst_mp0 subst_mbi mp1 with
- | None -> None
- | Some (mp2,_) -> Some (MPI mp2)
- end
- | _ -> None
- in
- match newkey with
- | None -> Umap.add key (mp,resolve) sub
- | Some mpi -> Umap.add mpi (mp,resolve) sub
+ Deltamap.fold apply_subst resolver empty_delta_resolver
+
+let subst_mp_delta sub mp key=
+ match subst_mp0 sub mp with
+ None -> empty_delta_resolver,mp
+ | Some (mp',resolve) ->
+ let mp1 = find_prefix resolve mp' in
+ let resolve1 = subset_prefixed_by mp1 resolve in
+ match key with
+ MP mpk ->
+ (subst_dom_delta_resolver
+ (map_mp mp1 mpk empty_delta_resolver) resolve1),mp1
+ | _ -> anomaly "Mod_subst: Bad association in resolver"
+
+let subst_codom_delta_resolver subst resolver =
+ let apply_subst key hint resolver =
+ match hint with
+ Prefix_equiv mp ->
+ let derived_resolve,mpnew = subst_mp_delta subst mp key in
+ Deltamap.fold Deltamap.add derived_resolve
+ (Deltamap.add key (Prefix_equiv mpnew) resolver)
+ | (Equiv kn) ->
+ Deltamap.add key (Equiv (subst_kn_delta subst kn)) resolver
+ | Inline None ->
+ Deltamap.add key hint resolver
+ | Inline (Some t) ->
+ Deltamap.add key (Inline (Some (subst_mps subst t))) resolver
in
- Umap.fold alias_subst subst1 empty_subst
-
-let update_subst subst1 subst2 =
- let subst_inv key (mp,resolve) l =
- let newmp =
- match key with
- | MBI msid -> MPbound msid
- | MSI msid -> MPself msid
- | MPI mp -> mp
- in
- match mp with
- | MPbound mbid -> ((MBI mbid),newmp,resolve)::l
- | MPself msid -> ((MSI msid),newmp,resolve)::l
- | _ -> ((MPI mp),newmp,resolve)::l
+ Deltamap.fold apply_subst resolver empty_delta_resolver
+
+let subst_dom_codom_delta_resolver subst resolver =
+ let apply_subst key hint resolver =
+ match key,hint with
+ (MP mp1),Prefix_equiv mp ->
+ let key = MP (subst_mp subst mp1) in
+ let derived_resolve,mpnew = subst_mp_delta subst mp key in
+ Deltamap.fold Deltamap.add derived_resolve
+ (Deltamap.add key (Prefix_equiv mpnew) resolver)
+ | (KN kn1),(Equiv kn) ->
+ let key = KN (subst_kn subst kn1) in
+ Deltamap.add key (Equiv (subst_kn_delta subst kn)) resolver
+ | (KN kn),Inline None ->
+ let key = KN (subst_kn subst kn) in
+ Deltamap.add key hint resolver
+ | (KN kn),Inline (Some t) ->
+ let key = KN (subst_kn subst kn) in
+ Deltamap.add key (Inline (Some (subst_mps subst t))) resolver
+ | _,_ -> anomaly "Mod_subst: Bad association in resolver"
in
- let subst_mbi = Umap.fold subst_inv subst2 [] in
- let alias_subst key (mp,resolve) sub=
- let newsetkey =
- match key with
- | MPI mp1 ->
- let compute_set_newkey l (k,mp',resolve) =
- let mp_from_key = match k with
- | MBI msid -> MPbound msid
- | MSI msid -> MPself msid
- | MPI mp -> mp
- in
- let new_mp1 = replace_mp_in_mp mp_from_key mp' mp1 in
- if new_mp1 == mp1 then l else (MPI new_mp1,resolve)::l
- in
- begin
- match List.fold_left compute_set_newkey [] subst_mbi with
- | [] -> None
- | l -> Some (l)
- end
- | _ -> None
+ Deltamap.fold apply_subst resolver empty_delta_resolver
+
+let update_delta_resolver resolver1 resolver2 =
+ let apply_res key hint res =
+ try
+ match hint with
+ Prefix_equiv mp ->
+ let new_hint =
+ Prefix_equiv (find_prefix resolver2 mp)
+ in Deltamap.add key new_hint res
+ | Equiv kn ->
+ let new_hint =
+ Equiv (solve_delta_kn resolver2 kn)
+ in Deltamap.add key new_hint res
+ | _ -> Deltamap.add key hint res
+ with not_found ->
+ Deltamap.add key hint res
in
- match newsetkey with
- | None -> sub
- | Some l ->
- List.fold_left (fun s (k,r) -> Umap.add k (mp,r) s)
- sub l
- in
- Umap.fold alias_subst subst1 empty_subst
+ Deltamap.fold apply_res resolver1 empty_delta_resolver
+
+let add_delta_resolver resolver1 resolver2 =
+ if resolver1 == resolver2 then
+ resolver2
+ else if resolver2 = empty_delta_resolver then
+ resolver1
+ else
+ Deltamap.fold Deltamap.add (update_delta_resolver resolver1 resolver2)
+ resolver2
-let join_alias (subst1 : substitution) (subst2 : substitution) =
+let join (subst1 : substitution) (subst2 : substitution) =
let apply_subst (sub : substitution) key (mp,resolve) =
let mp',resolve' =
match subst_mp0 sub mp with
None -> mp, None
- | Some (mp',resolve') -> mp',resolve' in
- let resolve'' : resolver option =
- try
- let res =
- match resolve with
- Some res -> res
- | None ->
- match resolve' with
- None -> raise BothSubstitutionsAreIdentitySubstitutions
- | Some res -> raise (ChangeDomain res)
- in
- Some
- (List.map
- (fun (kn,topt) ->
- kn,
- match topt with
- None ->
- (match key with
- MSI msid ->
- let kn' = replace_mp_in_con (MPself msid) mp kn in
- apply_opt_resolver resolve' kn'
- | MBI mbid ->
- let kn' = replace_mp_in_con (MPbound mbid) mp kn in
- apply_opt_resolver resolve' kn'
- | MPI mp1 ->
- let kn' = replace_mp_in_con mp1 mp kn in
- apply_opt_resolver resolve' kn')
- | Some t -> Some (subst_mps sub t)) res)
- with
- BothSubstitutionsAreIdentitySubstitutions -> None
- | ChangeDomain res ->
- let rec changeDom = function
- | [] -> []
- | (kn,topt)::r ->
- let key' =
- match key with
- MSI msid -> MPself msid
- | MBI mbid -> MPbound mbid
- | MPI mp1 -> mp1 in
- let kn' = replace_mp_in_con mp key' kn in
- if kn==kn' then
- (*the key does not appear in kn, we remove it
- from the resolver that we are building*)
- changeDom r
- else
- (kn',topt)::(changeDom r)
- in
- Some (changeDom res)
+ | Some (mp',resolve') -> mp'
+ ,Some resolve' in
+ let resolve'' : delta_resolver =
+ match resolve' with
+ Some res ->
+ add_delta_resolver
+ (subst_dom_codom_delta_resolver sub resolve)
+ res
+ | None ->
+ subst_codom_delta_resolver sub resolve
in
mp',resolve'' in
- Umap.mapi (apply_subst subst2) subst1
+ let subst = Umap.mapi (apply_subst subst2) subst1 in
+ (Umap.fold Umap.add subst2 subst)
-let remove_alias subst =
- let rec remove key (mp,resolve) sub =
- match key with
- MPI _ -> sub
- | _ -> Umap.add key (mp,resolve) sub
- in
- Umap.fold remove subst empty_subst
let rec occur_in_path uid path =
match uid,path with
- | MSI sid,MPself sid' -> sid = sid'
| MBI bid,MPbound bid' -> bid = bid'
| _,MPdot (mp1,_) -> occur_in_path uid mp1
| _ -> false
@@ -471,7 +657,7 @@ let occur_uid uid sub =
false
with Exit -> true
-let occur_msid uid = occur_uid (MSI uid)
+
let occur_mbid uid = occur_uid (MBI uid)
type 'a lazy_subst =
diff --git a/kernel/mod_subst.mli b/kernel/mod_subst.mli
index d30168a1b..a948d1647 100644
--- a/kernel/mod_subst.mli
+++ b/kernel/mod_subst.mli
@@ -13,26 +13,74 @@
open Names
open Term
-type resolver
+(* A delta_resolver maps name (constant, inductive, module_path)
+ to canonical name *)
+type delta_resolver
+
type substitution
-val make_resolver : (constant * constr option) list -> resolver
+val empty_delta_resolver : delta_resolver
+
+val add_inline_delta_resolver : constant -> delta_resolver -> delta_resolver
+
+val add_inline_constr_delta_resolver : constant -> constr -> delta_resolver
+ -> delta_resolver
+
+val add_constant_delta_resolver : constant -> delta_resolver -> delta_resolver
+
+val add_mind_delta_resolver : mutual_inductive -> delta_resolver -> delta_resolver
+
+val add_mp_delta_resolver : module_path -> module_path -> delta_resolver
+ -> delta_resolver
+
+val add_delta_resolver : delta_resolver -> delta_resolver -> delta_resolver
+
+(* Apply the substitution on the domain of the resolver *)
+val subst_dom_delta_resolver : substitution -> delta_resolver -> delta_resolver
+
+(* Apply the substitution on the codomain of the resolver *)
+val subst_codom_delta_resolver : substitution -> delta_resolver -> delta_resolver
+
+val subst_dom_codom_delta_resolver :
+ substitution -> delta_resolver -> delta_resolver
+
+(* *_of_delta return the associated name of arg2 in arg1 *)
+val constant_of_delta : delta_resolver -> constant -> constant
+
+val mind_of_delta : delta_resolver -> mutual_inductive -> mutual_inductive
+
+val delta_of_mp : delta_resolver -> module_path -> module_path
+
+(* Extract the set of inlined constant in the resolver *)
+val inline_of_delta : delta_resolver -> kernel_name list
+
+(* remove_mp is used for the computation of a resolver induced by Include P *)
+val remove_mp_delta_resolver : delta_resolver -> module_path -> delta_resolver
+
+
+(* mem tests *)
+val mp_in_delta : module_path -> delta_resolver -> bool
+
+val con_in_delta : constant -> delta_resolver -> bool
+
+val mind_in_delta : mutual_inductive -> delta_resolver -> bool
+
+(*substitution*)
val empty_subst : substitution
-val add_msid :
- mod_self_id -> module_path -> substitution -> substitution
+(* add_* add [arg2/arg1]{arg3} to the substitution with no
+ sequential composition *)
val add_mbid :
- mod_bound_id -> module_path -> resolver option -> substitution -> substitution
+ mod_bound_id -> module_path -> delta_resolver -> substitution -> substitution
val add_mp :
- module_path -> module_path -> substitution -> substitution
+ module_path -> module_path -> delta_resolver -> substitution -> substitution
-val map_msid :
- mod_self_id -> module_path -> substitution
+(* map_* create a new substitution [arg2/arg1]{arg3} *)
val map_mbid :
- mod_bound_id -> module_path -> resolver option -> substitution
+ mod_bound_id -> module_path -> delta_resolver -> substitution
val map_mp :
- module_path -> module_path -> substitution
+ module_path -> module_path -> delta_resolver -> substitution
(* sequential composition:
[substitute (join sub1 sub2) t = substitute sub2 (substitute sub1 t)]
@@ -47,6 +95,8 @@ val subst_substituted : substitution -> 'a substituted -> 'a substituted
(*i debugging *)
val debug_string_of_subst : substitution -> string
val debug_pr_subst : substitution -> Pp.std_ppcmds
+val debug_string_of_delta : delta_resolver -> string
+val debug_pr_delta : delta_resolver -> Pp.std_ppcmds
(*i*)
(* [subst_mp sub mp] guarantees that whenever the result of the
@@ -56,7 +106,10 @@ val debug_pr_subst : substitution -> Pp.std_ppcmds
val subst_mp :
substitution -> module_path -> module_path
-val subst_kn :
+val subst_ind :
+ substitution -> mutual_inductive -> mutual_inductive
+
+val subst_kn :
substitution -> kernel_name -> kernel_name
val subst_con :
@@ -71,7 +124,7 @@ val subst_evaluable_reference :
substitution -> evaluable_global_reference -> evaluable_global_reference
(* [replace_mp_in_con mp mp' con] replaces [mp] with [mp'] in [con] *)
-val replace_mp_in_con : module_path -> module_path -> constant -> constant
+val replace_mp_in_kn : module_path -> module_path -> kernel_name -> kernel_name
(* [subst_mps sub c] performs the substitution [sub] on all kernel
names appearing in [c] *)
@@ -80,15 +133,5 @@ val subst_mps : substitution -> constr -> constr
(* [occur_*id id sub] returns true iff [id] occurs in [sub]
on either side *)
-val occur_msid : mod_self_id -> substitution -> bool
val occur_mbid : mod_bound_id -> substitution -> bool
-val update_subst_alias : substitution -> substitution -> substitution
-
-val update_subst : substitution -> substitution -> substitution
-
-val subst_key : substitution -> substitution -> substitution
-
-val join_alias : substitution -> substitution -> substitution
-
-val remove_alias : substitution -> substitution
diff --git a/kernel/mod_typing.ml b/kernel/mod_typing.ml
index 3d55fb69a..48928470a 100644
--- a/kernel/mod_typing.ml
+++ b/kernel/mod_typing.ml
@@ -37,52 +37,32 @@ let rec list_fold_map2 f e = function
let e'',t1',t2' = list_fold_map2 f e' t in
e'',h1'::t1',h2'::t2'
+let discr_resolver env mtb =
+ match mtb.typ_expr with
+ SEBstruct _ ->
+ mtb.typ_delta
+ | _ -> (*case mp is a functor *)
+ empty_delta_resolver
+
let rec rebuild_mp mp l =
match l with
[]-> mp
- | i::r -> rebuild_mp (MPdot(mp,i)) r
-
-let type_of_struct env b meb =
- let rec aux env = function
- | SEBfunctor (mp,mtb,body) ->
- let env = add_module (MPbound mp) (module_body_of_type mtb) env in
- SEBfunctor(mp,mtb, aux env body)
- | SEBident mp ->
- strengthen env (lookup_modtype mp env).typ_expr mp
- | SEBapply _ as mtb -> eval_struct env mtb
- | str -> str
- in
- if b then
- Some (aux env meb)
- else
- None
-
-let rec bounded_str_expr = function
- | SEBfunctor (mp,mtb,body) -> bounded_str_expr body
- | SEBident mp -> (check_bound_mp mp)
- | SEBapply (f,a,_)->(bounded_str_expr f)
- | _ -> false
-
-let return_opt_type mp env mtb =
- if (check_bound_mp mp) then
- Some (strengthen env mtb.typ_expr mp)
- else
- None
-
-let rec check_with env mtb with_decl =
+ | i::r -> rebuild_mp (MPdot(mp,i)) r
+
+let rec check_with env sign with_decl alg_sign mp equiv =
match with_decl with
| With_Definition (id,_) ->
- let cb = check_with_aux_def env mtb with_decl in
- SEBwith(mtb,With_definition_body(id,cb)),empty_subst
- | With_Module (id,mp) ->
- let cst,sub,typ_opt = check_with_aux_mod env mtb with_decl true in
- SEBwith(mtb,With_module_body(id,mp,typ_opt,cst)),sub
-
-and check_with_aux_def env mtb with_decl =
- let msid,sig_b = match (eval_struct env mtb) with
- | SEBstruct(msid,sig_b) ->
- msid,sig_b
- | _ -> error_signature_expected mtb
+ let sign,cb,cst = check_with_aux_def env sign with_decl mp equiv in
+ sign,Some (SEBwith(alg_sign,With_definition_body(id,cb))),equiv,cst
+ | With_Module (id,mp1) ->
+ let sign,equiv,cst =
+ check_with_aux_mod env sign with_decl mp equiv in
+ sign,Some (SEBwith(alg_sign,With_module_body(id,mp1))),equiv,cst
+
+and check_with_aux_def env sign with_decl mp equiv =
+ let sig_b = match sign with
+ | SEBstruct(sig_b) -> sig_b
+ | _ -> error_signature_expected sign
in
let id,idl = match with_decl with
| With_Definition (id::idl,_) | With_Module (id::idl,_) -> id,idl
@@ -92,7 +72,7 @@ and check_with_aux_def env mtb with_decl =
try
let rev_before,spec,after = list_split_assoc l [] sig_b in
let before = List.rev rev_before in
- let env' = Modops.add_signature (MPself msid) before env in
+ let env' = Modops.add_signature mp before equiv env in
match with_decl with
| With_Definition ([],_) -> assert false
| With_Definition ([id],c) ->
@@ -116,7 +96,7 @@ and check_with_aux_def env mtb with_decl =
const_body_code = Cemitcodes.from_val
(compile_constant_body env' body false false);
const_constraints = cst} in
- cb'
+ SEBstruct(before@((l,SFBconst(cb'))::after)),cb',cst
| Some b ->
let cst1 = Reduction.conv env' c (Declarations.force b) in
let cst = Constraint.union cb.const_constraints cst1 in
@@ -126,9 +106,9 @@ and check_with_aux_def env mtb with_decl =
const_body_code = Cemitcodes.from_val
(compile_constant_body env' body false false);
const_constraints = cst} in
- cb'
+ SEBstruct(before@((l,SFBconst(cb'))::after)),cb',cst
end
- | With_Definition (_::_,_) ->
+ | With_Definition (_::_,c) ->
let old = match spec with
SFBmodule msb -> msb
| _ -> error_not_a_module (string_of_label l)
@@ -136,10 +116,14 @@ and check_with_aux_def env mtb with_decl =
begin
match old.mod_expr with
| None ->
- let new_with_decl = match with_decl with
- With_Definition (_,c) -> With_Definition (idl,c)
- | With_Module (_,c) -> With_Module (idl,c) in
- check_with_aux_def env' (type_of_mb env old) new_with_decl
+ let new_with_decl = With_Definition (idl,c) in
+ let sign,cb,cst =
+ check_with_aux_def env' old.mod_type new_with_decl
+ (MPdot(mp,l)) old.mod_delta in
+ let new_spec = SFBmodule({old with
+ mod_type = sign;
+ mod_type_alg = None}) in
+ SEBstruct(before@((l,new_spec)::after)),cb,cst
| Some msb ->
error_a_generative_module_expected l
end
@@ -148,11 +132,10 @@ and check_with_aux_def env mtb with_decl =
Not_found -> error_no_such_label l
| Reduction.NotConvertible -> error_with_incorrect l
-and check_with_aux_mod env mtb with_decl now =
- let initmsid,msid,sig_b = match (eval_struct env mtb) with
- | SEBstruct(msid,sig_b) ->let msid'=(refresh_msid msid) in
- msid,msid',(subst_signature_msid msid (MPself(msid')) sig_b)
- | _ -> error_signature_expected mtb
+and check_with_aux_mod env sign with_decl mp equiv =
+ let sig_b = match sign with
+ | SEBstruct(sig_b) ->sig_b
+ | _ -> error_signature_expected sign
in
let id,idl = match with_decl with
| With_Definition (id::idl,_) | With_Module (id::idl,_) -> id,idl
@@ -163,172 +146,234 @@ and check_with_aux_mod env mtb with_decl now =
let rev_before,spec,after = list_split_assoc l [] sig_b in
let before = List.rev rev_before in
let rec mp_rec = function
- | [] -> MPself initmsid
+ | [] -> mp
| i::r -> MPdot(mp_rec r,label_of_id i)
in
- let env' = Modops.add_signature (MPself msid) before env in
+ let env' = Modops.add_signature mp before equiv env in
match with_decl with
| With_Module ([],_) -> assert false
- | With_Module ([id], mp) ->
- let old,alias = match spec with
- SFBmodule msb -> Some msb,None
- | SFBalias (mp',_,cst) -> None,Some (mp',cst)
+ | With_Module ([id], mp1) ->
+ let old = match spec with
+ SFBmodule msb -> msb
| _ -> error_not_a_module (string_of_label l)
in
- let mtb' = lookup_modtype mp env' in
+ let mb_mp1 = (lookup_module mp1 env) in
+ let mtb_mp1 =
+ module_type_of_module env' None mb_mp1 in
let cst =
- match old,alias with
- Some msb,None ->
+ match old.mod_expr with
+ None ->
begin
try Constraint.union
- (check_subtypes env' mtb' (module_type_of_module None msb))
- msb.mod_constraints
+ (check_subtypes env' mtb_mp1
+ (module_type_of_module env' None old))
+ old.mod_constraints
with Failure _ -> error_with_incorrect (label_of_id id)
end
- | None,Some (mp',None) ->
- check_modpath_equiv env' mp mp';
- Constraint.empty
- | None,Some (mp',Some cst) ->
- check_modpath_equiv env' mp mp';
- cst
- | _,_ ->
- anomaly "Mod_typing:no implementation and no alias"
+ | Some (SEBident(mp')) ->
+ check_modpath_equiv env' mp1 mp';
+ old.mod_constraints
+ | _ -> error_a_generative_module_expected l
+ in
+ let new_mb = strengthen_and_subst_mb mb_mp1
+ (MPdot(mp,l)) env false in
+ let new_spec = SFBmodule {new_mb with
+ mod_mp = MPdot(mp,l);
+ mod_expr = Some (SEBident mp1);
+ mod_constraints = cst} in
+ (* we propagate the new equality in the rest of the signature
+ with the identity substitution accompagned by the new resolver*)
+ let id_subst = map_mp (MPdot(mp,l)) (MPdot(mp,l)) new_mb.mod_delta in
+ SEBstruct(before@(l,new_spec)::subst_signature id_subst after),
+ add_delta_resolver equiv new_mb.mod_delta,cst
+ | With_Module (idc,mp1) ->
+ let old = match spec with
+ SFBmodule msb -> msb
+ | _ -> error_not_a_module (string_of_label l)
in
- if now then
- let mp' = scrape_alias mp env' in
- let _,sub = Modops.update_subst env' (module_body_of_type mtb') mp' in
- let up_subst = update_subst sub (map_mp (mp_rec [id]) mp') in
- cst, (join (map_mp (mp_rec [id]) mp') up_subst),(return_opt_type mp env' mtb')
- else
- cst,empty_subst,(return_opt_type mp env' mtb')
- | With_Module (_::_,mp) ->
- let old,alias = match spec with
- SFBmodule msb -> Some msb, None
- | SFBalias (mpold,typ_opt,cst)->None, Some mpold
- | _ -> error_not_a_module (string_of_label l)
- in
begin
- if alias = None then
- let old = Option.get old in
- match old.mod_expr with
- None ->
- let new_with_decl = match with_decl with
- With_Definition (_,c) ->
- With_Definition (idl,c)
- | With_Module (_,c) -> With_Module (idl,c) in
- let cst,_,typ_opt =
- check_with_aux_mod env'
- (type_of_mb env' old) new_with_decl false in
- if now then
- let mtb' = lookup_modtype mp env' in
- let mp' = scrape_alias mp env' in
- let _,sub = Modops.update_subst env' (module_body_of_type mtb') mp' in
- let up_subst = update_subst
- sub (map_mp (mp_rec (List.rev (id::idl))) mp') in
- cst,
- (join (map_mp (mp_rec (List.rev (id::idl))) mp') up_subst),
- typ_opt
- else
- cst,empty_subst,typ_opt
- | Some msb ->
- error_a_generative_module_expected l
- else
- let mpold = Option.get alias in
- let mpnew = rebuild_mp mpold (List.map label_of_id idl) in
- check_modpath_equiv env' mpnew mp;
- let mtb' = lookup_modtype mp env' in
- Constraint.empty,empty_subst,(return_opt_type mp env' mtb')
+ match old.mod_expr with
+ None ->
+ let new_with_decl = With_Module (idl,mp1) in
+ let sign,equiv',cst =
+ check_with_aux_mod env'
+ old.mod_type new_with_decl (MPdot(mp,l)) old.mod_delta in
+ let new_equiv = add_delta_resolver equiv equiv' in
+ let new_spec = SFBmodule {old with
+ mod_type = sign;
+ mod_type_alg = None;
+ mod_delta = equiv'}
+ in
+ let id_subst = map_mp (MPdot(mp,l)) (MPdot(mp,l)) equiv' in
+ SEBstruct(before@(l,new_spec)::subst_signature id_subst after),
+ new_equiv,cst
+ | Some (SEBident(mp')) ->
+ let mpnew = rebuild_mp mp' (List.map label_of_id idl) in
+ check_modpath_equiv env' mpnew mp;
+ SEBstruct(before@(l,spec)::after)
+ ,equiv,Constraint.empty
+ | _ ->
+ error_a_generative_module_expected l
end
- | _ -> anomaly "Modtyping:incorrect use of with"
+ | _ -> anomaly "Modtyping:incorrect use of with"
with
Not_found -> error_no_such_label l
| Reduction.NotConvertible -> error_with_incorrect l
-and translate_module env me =
+and translate_module env mp me =
match me.mod_entry_expr, me.mod_entry_type with
| None, None ->
anomaly "Mod_typing.translate_module: empty type and expr in module entry"
| None, Some mte ->
- let mtb,sub = translate_struct_entry env mte in
- { mod_expr = None;
- mod_type = Some mtb;
- mod_alias = sub;
- mod_constraints = Constraint.empty;
- mod_retroknowledge = []}
- | Some mexpr, _ ->
- let meb,sub1 = translate_struct_entry env mexpr in
- let mod_typ,sub,cst =
+ let mtb = translate_module_type env mp mte in
+ { mod_mp = mp;
+ mod_expr = None;
+ mod_type = mtb.typ_expr;
+ mod_type_alg = mtb.typ_expr_alg;
+ mod_delta = mtb.typ_delta;
+ mod_constraints = mtb.typ_constraints;
+ mod_retroknowledge = []}
+ | Some mexpr, _ ->
+ let sign,alg_implem,resolver,cst1 =
+ translate_struct_module_entry env mp mexpr in
+ let sign,alg1,resolver,cst2 =
match me.mod_entry_type with
| None ->
- (type_of_struct env (bounded_str_expr meb) meb)
- ,sub1,Constraint.empty
+ sign,None,resolver,Constraint.empty
| Some mte ->
- let mtb2,sub2 = translate_struct_entry env mte in
+ let mtb = translate_module_type env mp mte in
let cst = check_subtypes env
- {typ_expr = meb;
- typ_strength = None;
- typ_alias = sub1;}
- {typ_expr = mtb2;
- typ_strength = None;
- typ_alias = sub2;}
+ {typ_mp = mp;
+ typ_expr = sign;
+ typ_expr_alg = None;
+ typ_constraints = Constraint.empty;
+ typ_delta = resolver;}
+ mtb
in
- Some mtb2,sub2,cst
+ mtb.typ_expr,mtb.typ_expr_alg,mtb.typ_delta,cst
in
- { mod_type = mod_typ;
- mod_expr = Some meb;
- mod_constraints = cst;
- mod_alias = sub;
- mod_retroknowledge = []} (* spiwack: not so sure about that. It may
- cause a bug when closing nested modules.
- If it does, I don't really know how to
- fix the bug.*)
-
-
-and translate_struct_entry env mse = match mse with
- | MSEident mp ->
- let mtb = lookup_modtype mp env in
- SEBident mp,mtb.typ_alias
+ { mod_mp = mp;
+ mod_type = sign;
+ mod_expr = Some alg_implem;
+ mod_type_alg = alg1;
+ mod_constraints = Univ.Constraint.union cst1 cst2;
+ mod_delta = resolver;
+ mod_retroknowledge = []}
+ (* spiwack: not so sure about that. It may
+ cause a bug when closing nested modules.
+ If it does, I don't really know how to
+ fix the bug.*)
+
+
+and translate_struct_module_entry env mp mse = match mse with
+ | MSEident mp1 ->
+ let mb = lookup_module mp1 env in
+ let mb' = strengthen_and_subst_mb mb mp env false in
+ mb'.mod_type, SEBident mp1, mb'.mod_delta,Univ.Constraint.empty
| MSEfunctor (arg_id, arg_e, body_expr) ->
- let arg_b,sub = translate_struct_entry env arg_e in
- let mtb =
- {typ_expr = arg_b;
- typ_strength = None;
- typ_alias = sub} in
- let env' = add_module (MPbound arg_id) (module_body_of_type mtb) env in
- let body_b,sub = translate_struct_entry env' body_expr in
- SEBfunctor (arg_id, mtb, body_b),sub
+ let mtb = translate_module_type env (MPbound arg_id) arg_e in
+ let env' = add_module (module_body_of_type (MPbound arg_id) mtb)
+ env in
+ let sign,alg,resolver,cst = translate_struct_module_entry env'
+ mp body_expr in
+ SEBfunctor (arg_id, mtb, sign),SEBfunctor (arg_id, mtb, alg),resolver,
+ Univ.Constraint.union cst mtb.typ_constraints
| MSEapply (fexpr,mexpr) ->
- let feb,sub1 = translate_struct_entry env fexpr in
- let feb'= eval_struct env feb
- in
- let farg_id, farg_b, fbody_b = destr_functor env feb' in
- let mtb,mp =
+ let sign,alg,resolver,cst1 = translate_struct_module_entry env mp fexpr in
+ let farg_id, farg_b, fbody_b = destr_functor env sign in
+ let mtb,mp1 =
try
- let mp = scrape_alias (path_of_mexpr mexpr) env in
- lookup_modtype mp env,mp
+ let mp1 = path_of_mexpr mexpr in
+ let mtb = module_type_of_module env None (lookup_module mp1 env) in
+ mtb,mp1
with
| Not_path -> error_application_to_not_path mexpr
(* place for nondep_supertype *) in
- let meb,sub2= translate_struct_entry env (MSEident mp) in
- if sub1 = empty_subst then
- let cst = check_subtypes env mtb farg_b in
- SEBapply(feb,meb,cst),sub1
- else
- let sub2 = match eval_struct env (SEBident mp) with
- | SEBstruct (msid,sign) ->
- join_alias
- (subst_key (map_msid msid mp) sub2)
- (map_msid msid mp)
- | _ -> sub2 in
- let sub3 = join_alias sub1 (map_mbid farg_id mp None) in
- let sub4 = update_subst sub2 sub3 in
- let cst = check_subtypes env mtb farg_b in
- SEBapply(feb,meb,cst),(join sub3 sub4)
+ let cst = check_subtypes env mtb farg_b in
+ let mp_delta = discr_resolver env mtb in
+ let mp_delta = complete_inline_delta_resolver env mp1
+ farg_id farg_b mp_delta in
+ let subst = map_mbid farg_id mp1 mp_delta in
+ (subst_struct_expr subst fbody_b),SEBapply(alg,SEBident mp1,cst),
+ (subst_codom_delta_resolver subst resolver),
+ Univ.Constraint.union cst1 cst
| MSEwith(mte, with_decl) ->
- let mtb,sub1 = translate_struct_entry env mte in
- let mtb',sub2 = check_with env mtb with_decl in
- mtb',join sub1 sub2
-
+ let sign,alg,resolve,cst1 = translate_struct_module_entry env mp mte in
+ let sign,alg,resolve,cst2 = check_with env sign with_decl alg mp resolve in
+ sign,Option.get alg,resolve,Univ.Constraint.union cst1 cst2
+
+and translate_struct_type_entry env mse = match mse with
+ | MSEident mp1 ->
+ let mtb = lookup_modtype mp1 env in
+ mtb.typ_expr,
+ Some (SEBident mp1),mtb.typ_delta,mp1,Univ.Constraint.empty
+ | MSEfunctor (arg_id, arg_e, body_expr) ->
+ let mtb = translate_module_type env (MPbound arg_id) arg_e in
+ let env' = add_module (module_body_of_type (MPbound arg_id) mtb) env in
+ let sign,alg,resolve,mp_from,cst = translate_struct_type_entry env'
+ body_expr in
+ SEBfunctor (arg_id, mtb, sign),None,resolve,mp_from,
+ Univ.Constraint.union cst mtb.typ_constraints
+ | MSEapply (fexpr,mexpr) ->
+ let sign,alg,resolve,mp_from,cst1 = translate_struct_type_entry env fexpr in
+ let farg_id, farg_b, fbody_b = destr_functor env sign in
+ let mtb,mp1 =
+ try
+ let mp1 = path_of_mexpr mexpr in
+ let mtb = module_type_of_module env None (lookup_module mp1 env) in
+ mtb,mp1
+ with
+ | Not_path -> error_application_to_not_path mexpr
+ (* place for nondep_supertype *) in
+ let cst2 = check_subtypes env mtb farg_b in
+ let mp_delta = discr_resolver env mtb in
+ let mp_delta = complete_inline_delta_resolver env mp1
+ farg_id farg_b mp_delta in
+ let subst = map_mbid farg_id mp1 mp_delta in
+ (subst_struct_expr subst fbody_b),None,
+ (subst_codom_delta_resolver subst resolve),mp_from,Univ.Constraint.union cst1 cst2
+ | MSEwith(mte, with_decl) ->
+ let sign,alg,resolve,mp_from,cst = translate_struct_type_entry env mte in
+ let sign,alg,resolve,cst1 =
+ check_with env sign with_decl (Option.get alg) mp_from resolve in
+ sign,alg,resolve,mp_from,Univ.Constraint.union cst cst1
+
+and translate_module_type env mp mte =
+ let sign,alg,resolve,mp_from,cst = translate_struct_type_entry env mte in
+ subst_modtype_and_resolver
+ {typ_mp = mp_from;
+ typ_expr = sign;
+ typ_expr_alg = alg;
+ typ_constraints = cst;
+ typ_delta = resolve} mp env
+
+let rec translate_struct_include_module_entry env mp mse = match mse with
+ | MSEident mp1 ->
+ let mb = lookup_module mp1 env in
+ let mb' = strengthen_and_subst_mb mb mp env true in
+ mb'.mod_type, mb'.mod_delta,Univ.Constraint.empty
+ | MSEapply (fexpr,mexpr) ->
+ let sign,resolver,cst1 =
+ translate_struct_include_module_entry env mp fexpr in
+ let farg_id, farg_b, fbody_b = destr_functor env sign in
+ let mtb,mp1 =
+ try
+ let mp1 = path_of_mexpr mexpr in
+ let mtb = module_type_of_module env None (lookup_module mp1 env) in
+ mtb,mp1
+ with
+ | Not_path -> error_application_to_not_path mexpr
+ (* place for nondep_supertype *) in
+ let cst = check_subtypes env mtb farg_b in
+ let mp_delta = discr_resolver env mtb in
+ let mp_delta = complete_inline_delta_resolver env mp1
+ farg_id farg_b mp_delta in
+ let subst = map_mbid farg_id mp1 mp_delta in
+ (subst_struct_expr subst fbody_b),
+ (subst_codom_delta_resolver subst resolver),
+ Univ.Constraint.union cst1 cst
+ | _ -> error ("You cannot Include a high-order structure.")
+
let rec add_struct_expr_constraints env = function
| SEBident _ -> env
@@ -337,7 +382,7 @@ let rec add_struct_expr_constraints env = function
add_struct_expr_constraints
(add_modtype_constraints env mtb) meb
- | SEBstruct (_,structure_body) ->
+ | SEBstruct (structure_body) ->
List.fold_left
(fun env (l,item) -> add_struct_elem_constraints env item)
env
@@ -351,16 +396,13 @@ let rec add_struct_expr_constraints env = function
| SEBwith(meb,With_definition_body(_,cb))->
Environ.add_constraints cb.const_constraints
(add_struct_expr_constraints env meb)
- | SEBwith(meb,With_module_body(_,_,_,cst))->
- Environ.add_constraints cst
- (add_struct_expr_constraints env meb)
-
-and add_struct_elem_constraints env = function
+ | SEBwith(meb,With_module_body(_,_))->
+ add_struct_expr_constraints env meb
+
+and add_struct_elem_constraints env = function
| SFBconst cb -> Environ.add_constraints cb.const_constraints env
| SFBmind mib -> Environ.add_constraints mib.mind_constraints env
| SFBmodule mb -> add_module_constraints env mb
- | SFBalias (mp,_,Some cst) -> Environ.add_constraints cst env
- | SFBalias (mp,_,None) -> env
| SFBmodtype mtb -> add_modtype_constraints env mtb
and add_module_constraints env mb =
@@ -368,15 +410,14 @@ and add_module_constraints env mb =
| None -> env
| Some meb -> add_struct_expr_constraints env meb
in
- let env = match mb.mod_type with
- | None -> env
- | Some mtb ->
- add_struct_expr_constraints env mtb
+ let env =
+ add_struct_expr_constraints env mb.mod_type
in
Environ.add_constraints mb.mod_constraints env
and add_modtype_constraints env mtb =
- add_struct_expr_constraints env mtb.typ_expr
+ Environ.add_constraints mtb.typ_constraints
+ (add_struct_expr_constraints env mtb.typ_expr)
let rec struct_expr_constraints cst = function
@@ -386,7 +427,7 @@ let rec struct_expr_constraints cst = function
struct_expr_constraints
(modtype_constraints cst mtb) meb
- | SEBstruct (_,structure_body) ->
+ | SEBstruct (structure_body) ->
List.fold_left
(fun cst (l,item) -> struct_elem_constraints cst item)
cst
@@ -399,28 +440,25 @@ let rec struct_expr_constraints cst = function
| SEBwith(meb,With_definition_body(_,cb))->
struct_expr_constraints
(Univ.Constraint.union cb.const_constraints cst) meb
- | SEBwith(meb,With_module_body(_,_,_,cst1))->
- struct_expr_constraints (Univ.Constraint.union cst1 cst) meb
-
-and struct_elem_constraints cst = function
+ | SEBwith(meb,With_module_body(_,_))->
+ struct_expr_constraints cst meb
+
+and struct_elem_constraints cst = function
| SFBconst cb -> cst
| SFBmind mib -> cst
| SFBmodule mb -> module_constraints cst mb
- | SFBalias (mp,_,Some cst1) -> Univ.Constraint.union cst1 cst
- | SFBalias (mp,_,None) -> cst
| SFBmodtype mtb -> modtype_constraints cst mtb
and module_constraints cst mb =
let cst = match mb.mod_expr with
| None -> cst
| Some meb -> struct_expr_constraints cst meb in
- let cst = match mb.mod_type with
- | None -> cst
- | Some mtb -> struct_expr_constraints cst mtb in
+ let cst =
+ struct_expr_constraints cst mb.mod_type in
Univ.Constraint.union mb.mod_constraints cst
and modtype_constraints cst mtb =
- struct_expr_constraints cst mtb.typ_expr
+ struct_expr_constraints (Univ.Constraint.union mtb.typ_constraints cst) mtb.typ_expr
let struct_expr_constraints = struct_expr_constraints Univ.Constraint.empty
diff --git a/kernel/mod_typing.mli b/kernel/mod_typing.mli
index 1fadec2ad..746a80e15 100644
--- a/kernel/mod_typing.mli
+++ b/kernel/mod_typing.mli
@@ -13,13 +13,24 @@ open Declarations
open Environ
open Entries
open Mod_subst
+open Names
(*i*)
-val translate_module : env -> module_entry -> module_body
+val translate_module : env -> module_path -> module_entry
+ -> module_body
-val translate_struct_entry : env -> module_struct_entry ->
- struct_expr_body * substitution
+val translate_module_type : env -> module_path -> module_struct_entry ->
+ module_type_body
+
+val translate_struct_module_entry : env -> module_path -> module_struct_entry ->
+ struct_expr_body * struct_expr_body * delta_resolver * Univ.constraints
+
+val translate_struct_type_entry : env -> module_struct_entry ->
+ struct_expr_body * struct_expr_body option * delta_resolver * module_path * Univ.constraints
+
+val translate_struct_include_module_entry : env -> module_path
+ -> module_struct_entry -> struct_expr_body * delta_resolver * Univ.constraints
val add_modtype_constraints : env -> module_type_body -> env
diff --git a/kernel/modops.ml b/kernel/modops.ml
index 3f38cc2f7..8193c02e5 100644
--- a/kernel/modops.ml
+++ b/kernel/modops.ml
@@ -22,7 +22,7 @@ open Mod_subst
-let error_existing_label l =
+let error_existing_label l =
error ("The label "^string_of_label l^" is already declared.")
let error_declaration_not_path _ = error "Declaration is not a path."
@@ -39,31 +39,31 @@ let error_not_match l _ = error ("Signature components for label "^string_of_lab
let error_no_such_label l = error ("No such label "^string_of_label l^".")
-let error_incompatible_labels l l' =
+let error_incompatible_labels l l' =
error ("Opening and closing labels are not the same: "
^string_of_label l^" <> "^string_of_label l'^" !")
-let error_result_must_be_signature () =
+let error_result_must_be_signature () =
error "The result module type must be a signature."
let error_signature_expected mtb =
error "Signature expected."
-let error_no_module_to_end _ =
+let error_no_module_to_end _ =
error "No open module to end."
let error_no_modtype_to_end _ =
error "No open module type to end."
-let error_not_a_modtype_loc loc s =
+let error_not_a_modtype_loc loc s =
user_err_loc (loc,"",str ("\""^s^"\" is not a module type."))
-let error_not_a_module_loc loc s =
+let error_not_a_module_loc loc s =
user_err_loc (loc,"",str ("\""^s^"\" is not a module."))
let error_not_a_module s = error_not_a_module_loc dummy_loc s
-let error_not_a_constant l =
+let error_not_a_constant l =
error ("\""^(string_of_label l)^"\" is not a constant.")
let error_with_incorrect l =
@@ -74,28 +74,17 @@ let error_a_generative_module_expected l =
"component of generative modules can be changed using the \"with\" " ^
"construct.")
-let error_local_context lo =
+let error_local_context lo =
match lo with
- None ->
+ None ->
error ("The local context is not empty.")
| (Some l) ->
error ("The local context of the component "^
(string_of_label l)^" is not empty.")
-let error_no_such_label_sub l l1 l2 =
- error (l1^" is not a subtype of "^l2^".\nThe field "^(string_of_label l)^" is missing in "^l1^".")
-
-
-
-let rec list_split_assoc k rev_before = function
- | [] -> raise Not_found
- | (k',b)::after when k=k' -> rev_before,b,after
- | h::tail -> list_split_assoc k (h::rev_before) tail
-
-let path_of_seb = function
- | SEBident mp -> mp
- | _ -> anomaly "Modops: evaluation failed."
+let error_no_such_label_sub l l1 =
+ error ("The field "^(string_of_label l)^" is missing in "^l1^".")
let destr_functor env mtb =
@@ -104,123 +93,126 @@ let destr_functor env mtb =
(arg_id,arg_t,body_t)
| _ -> error_not_a_functor mtb
-(* the constraints are not important here *)
+let is_functor = function
+ | SEBfunctor (arg_id,arg_t,body_t) -> true
+ | _ -> false
-let module_body_of_type mtb =
- { mod_type = Some mtb.typ_expr;
+let module_body_of_type mp mtb =
+ { mod_mp = mp;
+ mod_type = mtb.typ_expr;
+ mod_type_alg = mtb.typ_expr_alg;
mod_expr = None;
- mod_constraints = Constraint.empty;
- mod_alias = mtb.typ_alias;
+ mod_constraints = mtb.typ_constraints;
+ mod_delta = mtb.typ_delta;
mod_retroknowledge = []}
-let module_type_of_module mp mb =
- let mp1,expr =
- (match mb.mod_type with
- | Some expr -> mp,expr
- | None -> (match mb.mod_expr with
- | Some (SEBident mp') ->(Some mp'),(SEBident mp')
- | Some expr -> mp,expr
- | None ->
- anomaly "Modops: empty expr and type")) in
- {typ_expr = expr;
- typ_alias = mb.mod_alias;
- typ_strength = mp1
- }
-
-let rec check_modpath_equiv env mp1 mp2 =
+let check_modpath_equiv env mp1 mp2 =
if mp1=mp2 then () else
- let mp1 = scrape_alias mp1 env in
- let mp2 = scrape_alias mp2 env in
- if mp1=mp2 then ()
- else
- error_not_equal mp1 mp2
-
+ let mb1=lookup_module mp1 env in
+ let mb2=lookup_module mp2 env in
+ if (delta_of_mp mb1.mod_delta mp1)=(delta_of_mp mb2.mod_delta mp2)
+ then ()
+ else error_not_equal mp1 mp2
+
let rec subst_with_body sub = function
- | With_module_body(id,mp,typ_opt,cst) ->
- With_module_body(id,subst_mp sub mp,Option.smartmap
- (subst_struct_expr sub) typ_opt,cst)
+ | With_module_body(id,mp) ->
+ With_module_body(id,subst_mp sub mp)
| With_definition_body(id,cb) ->
With_definition_body( id,subst_const_body sub cb)
-and subst_modtype sub mtb =
- let typ_expr' = subst_struct_expr sub mtb.typ_expr in
- let sub_mtb = join_alias mtb.typ_alias sub in
- if typ_expr'==mtb.typ_expr && sub_mtb==mtb.typ_alias then
- mtb
+and subst_modtype sub do_delta mtb=
+ let mp = subst_mp sub mtb.typ_mp in
+ let sub = add_mp mtb.typ_mp mp empty_delta_resolver sub in
+ let typ_expr' = subst_struct_expr sub do_delta mtb.typ_expr in
+ let typ_alg' =
+ Option.smartmap
+ (subst_struct_expr sub (fun x -> x)) mtb.typ_expr_alg in
+ let mtb_delta = do_delta mtb.typ_delta in
+ if typ_expr'==mtb.typ_expr &&
+ typ_alg'==mtb.typ_expr_alg && mp==mtb.typ_mp then
+ mtb
else
- { mtb with
- typ_expr = typ_expr';
- typ_alias = sub_mtb}
-
-and subst_structure sub sign =
- let subst_body = function
- SFBconst cb ->
+ {mtb with
+ typ_mp = mp;
+ typ_expr = typ_expr';
+ typ_expr_alg = typ_alg';
+ typ_delta = mtb_delta}
+
+and subst_structure sub do_delta sign =
+ let subst_body = function
+ SFBconst cb ->
SFBconst (subst_const_body sub cb)
- | SFBmind mib ->
+ | SFBmind mib ->
SFBmind (subst_mind sub mib)
- | SFBmodule mb ->
- SFBmodule (subst_module sub mb)
- | SFBmodtype mtb ->
- SFBmodtype (subst_modtype sub mtb)
- | SFBalias (mp,typ_opt,cst) ->
- SFBalias (subst_mp sub mp,Option.smartmap
- (subst_struct_expr sub) typ_opt,cst)
+ | SFBmodule mb ->
+ SFBmodule (subst_module sub do_delta mb)
+ | SFBmodtype mtb ->
+ SFBmodtype (subst_modtype sub do_delta mtb)
in
List.map (fun (l,b) -> (l,subst_body b)) sign
-and subst_module sub mb =
- let mtb' = Option.smartmap (subst_struct_expr sub) mb.mod_type in
- (* This is similar to the previous case. In this case we have
- a module M in a signature that is knows to be equivalent to a module M'
- (because the signature is "K with Module M := M'") and we are substituting
- M' with some M''. *)
- let me' = Option.smartmap (subst_struct_expr sub) mb.mod_expr in
- let mb_alias = update_subst sub mb.mod_alias in
- let mb_alias = if mb_alias = empty_subst then
- join_alias mb.mod_alias sub
- else
- join mb_alias (join_alias mb.mod_alias sub)
- in
- if mtb'==mb.mod_type && mb.mod_expr == me'
- && mb_alias == mb.mod_alias
+and subst_module sub do_delta mb =
+ let mp = subst_mp sub mb.mod_mp in
+ let sub = if is_functor mb.mod_type && not(mp=mb.mod_mp) then
+ add_mp mb.mod_mp mp
+ empty_delta_resolver sub else sub in
+ let id_delta = (fun x -> x) in
+ let mtb',me' =
+ let mtb = subst_struct_expr sub do_delta mb.mod_type in
+ match mb.mod_expr with
+ None -> mtb,None
+ | Some me -> if me==mb.mod_type then
+ mtb,Some mtb
+ else mtb,Option.smartmap
+ (subst_struct_expr sub id_delta) mb.mod_expr
+ in
+ let typ_alg' = Option.smartmap
+ (subst_struct_expr sub id_delta) mb.mod_type_alg in
+ let mb_delta = do_delta mb.mod_delta in
+ if mtb'==mb.mod_type && mb.mod_expr == me'
+ && mb_delta == mb.mod_delta && mp == mb.mod_mp
then mb else
- { mod_expr = me';
- mod_type=mtb';
- mod_constraints=mb.mod_constraints;
- mod_alias = mb_alias;
- mod_retroknowledge=mb.mod_retroknowledge}
-
-
-and subst_struct_expr sub = function
- | SEBident mp -> SEBident (subst_mp sub mp)
- | SEBfunctor (msid, mtb, meb') ->
- SEBfunctor(msid,subst_modtype sub mtb,subst_struct_expr sub meb')
- | SEBstruct (msid,str)->
- SEBstruct(msid, subst_structure sub str)
+ { mb with
+ mod_mp = mp;
+ mod_expr = me';
+ mod_type_alg = typ_alg';
+ mod_type=mtb';
+ mod_delta = mb_delta}
+
+and subst_struct_expr sub do_delta = function
+ | SEBident mp -> SEBident (subst_mp sub mp)
+ | SEBfunctor (mbid, mtb, meb') ->
+ SEBfunctor(mbid,subst_modtype sub do_delta mtb
+ ,subst_struct_expr sub do_delta meb')
+ | SEBstruct (str)->
+ SEBstruct( subst_structure sub do_delta str)
| SEBapply (meb1,meb2,cst)->
- SEBapply(subst_struct_expr sub meb1,
- subst_struct_expr sub meb2,
+ SEBapply(subst_struct_expr sub do_delta meb1,
+ subst_struct_expr sub do_delta meb2,
cst)
- | SEBwith (meb,wdb)->
- SEBwith(subst_struct_expr sub meb,
+ | SEBwith (meb,wdb)->
+ SEBwith(subst_struct_expr sub do_delta meb,
subst_with_body sub wdb)
+let subst_signature subst =
+ subst_structure subst
+ (fun resolver -> subst_codom_delta_resolver subst resolver)
-let subst_signature_msid msid mp =
- subst_structure (map_msid msid mp)
+let subst_struct_expr subst =
+ subst_struct_expr subst
+ (fun resolver -> subst_codom_delta_resolver subst resolver)
-(* spiwack: here comes the function which takes care of importing
+(* spiwack: here comes the function which takes care of importing
the retroknowledge declared in the library *)
(* lclrk : retroknowledge_action list, rkaction : retroknowledge action *)
-let add_retroknowledge msid mp =
- let subst = add_msid msid mp empty_subst in
- let subst_and_perform rkaction env =
+let add_retroknowledge mp =
+ let perform rkaction env =
match rkaction with
| Retroknowledge.RKRegister (f, e) ->
- Environ.register env f
- (match e with
- | Const kn -> kind_of_term (subst_mps subst (mkConst kn))
- | Ind ind -> kind_of_term (subst_mps subst (mkInd ind))
+ Environ.register env f
+ (match e with
+ | Const kn -> kind_of_term (mkConst kn)
+ | Ind ind -> kind_of_term (mkInd ind)
| _ -> anomaly "Modops.add_retroknowledge: had to import an unsupported kind of term")
in
fun lclrk env ->
@@ -229,305 +221,283 @@ let add_retroknowledge msid mp =
int31 type registration absolutely needs int31 bits to be registered.
Since the local_retroknowledge is stored in reverse order (each new
registration is added at the top of the list) we need a fold_right
- for things to go right (the pun is not intented). So we lose
+ for things to go right (the pun is not intented). So we lose
tail recursivity, but the world will have exploded before any module
imports 10 000 retroknowledge registration.*)
- List.fold_right subst_and_perform lclrk env
+ List.fold_right perform lclrk env
+let rec add_signature mp sign resolver env =
+ let add_one env (l,elem) =
+ let kn = make_kn mp empty_dirpath l in
+ let con = constant_of_kn kn in
+ let mind = mind_of_kn kn in
+ match elem with
+ | SFBconst cb ->
+ let con = constant_of_delta resolver con in
+ Environ.add_constant con cb env
+ | SFBmind mib ->
+ let mind = mind_of_delta resolver mind in
+ Environ.add_mind mind mib env
+ | SFBmodule mb -> add_module mb env
+ (* adds components as well *)
+ | SFBmodtype mtb -> Environ.add_modtype mtb.typ_mp mtb env
+ in
+ List.fold_left add_one env sign
+and add_module mb env =
+ let mp = mb.mod_mp in
+ let env = Environ.shallow_add_module mp mb env in
+ match mb.mod_type with
+ | SEBstruct (sign) ->
+ add_retroknowledge mp mb.mod_retroknowledge
+ (add_signature mp sign mb.mod_delta env)
+ | SEBfunctor _ -> env
+ | _ -> anomaly "Modops:the evaluation of the structure failed "
-let strengthen_const env mp l cb =
+let strengthen_const env mp_from l cb resolver =
match cb.const_opaque, cb.const_body with
| false, Some _ -> cb
- | true, Some _
+ | true, Some _
| _, None ->
- let const = mkConst (make_con mp empty_dirpath l) in
+ let con = make_con mp_from empty_dirpath l in
+ let con = constant_of_delta resolver con in
+ let const = mkConst con in
let const_subs = Some (Declarations.from_val const) in
- {cb with
+ {cb with
const_body = const_subs;
const_opaque = false;
const_body_code = Cemitcodes.from_val
(compile_constant_body env const_subs false false)
}
-
-let strengthen_mind env mp l mib = match mib.mind_equiv with
- | Some _ -> mib
- | None -> {mib with mind_equiv = Some (make_kn mp empty_dirpath l)}
-
-
-let rec eval_struct env = function
- | SEBident mp ->
- begin
- let mtb =lookup_modtype mp env in
- match mtb.typ_expr,mtb.typ_strength with
- mtb,None -> eval_struct env mtb
- | mtb,Some mp -> strengthen_mtb env mp (eval_struct env mtb)
- end
- | SEBapply (seb1,seb2,_) ->
- let svb1 = eval_struct env seb1 in
- let farg_id, farg_b, fbody_b = destr_functor env svb1 in
- let mp = path_of_seb seb2 in
- let mp = scrape_alias mp env in
- let sub_alias = (lookup_modtype mp env).typ_alias in
- let sub_alias = match eval_struct env (SEBident mp) with
- | SEBstruct (msid,sign) ->
- join_alias
- (subst_key (map_msid msid mp) sub_alias)
- (map_msid msid mp)
- | _ -> sub_alias in
- let resolve = resolver_of_environment farg_id farg_b mp sub_alias env in
- let sub_alias1 = update_subst sub_alias
- (map_mbid farg_id mp (Some resolve)) in
- eval_struct env (subst_struct_expr
- (join sub_alias1
- (map_mbid farg_id mp (Some resolve))) fbody_b)
- | SEBwith (mtb,(With_definition_body _ as wdb)) ->
- let mtb',_ = merge_with env mtb wdb empty_subst in
- mtb'
- | SEBwith (mtb, (With_module_body (_,mp,_,_) as wdb)) ->
- let alias_in_mp =
- (lookup_modtype mp env).typ_alias in
- let alias_in_mp = match eval_struct env (SEBident mp) with
- | SEBstruct (msid,sign) -> subst_key (map_msid msid mp) alias_in_mp
- | _ -> alias_in_mp in
- let mtb',_ = merge_with env mtb wdb alias_in_mp in
- mtb'
-(* | SEBfunctor(mbid,mtb,body) ->
- let env = add_module (MPbound mbid) (module_body_of_type mtb) env in
- SEBfunctor(mbid,mtb,eval_struct env body) *)
- | mtb -> mtb
-
-and type_of_mb env mb =
- match mb.mod_type,mb.mod_expr with
- None,Some b -> eval_struct env b
- | Some t, _ -> eval_struct env t
- | _,_ -> anomaly
- "Modops: empty type and empty expr"
-
-and merge_with env mtb with_decl alias=
- let msid,sig_b = match (eval_struct env mtb) with
- | SEBstruct(msid,sig_b) -> msid,sig_b
- | _ -> error_signature_expected mtb
- in
- let id,idl = match with_decl with
- | With_definition_body (id::idl,_) | With_module_body (id::idl,_,_,_) -> id,idl
- | With_definition_body ([],_) | With_module_body ([],_,_,_) -> assert false
- in
- let l = label_of_id id in
- try
- let rev_before,spec,after = list_split_assoc l [] sig_b in
- let before = List.rev rev_before in
- let rec mp_rec = function
- | [] -> MPself msid
- | i::r -> MPdot(mp_rec r,label_of_id i)
- in
- let env' = add_signature (MPself msid) before env in
- let new_spec,subst = match with_decl with
- | With_definition_body ([],_)
- | With_module_body ([],_,_,_) -> assert false
- | With_definition_body ([id],c) ->
- SFBconst c,None
- | With_module_body ([id], mp,typ_opt,cst) ->
- let mp' = scrape_alias mp env' in
- let new_alias = update_subst alias (map_mp (mp_rec [id]) mp') in
- SFBalias (mp,typ_opt,Some cst),
- Some(join (map_mp (mp_rec [id]) mp') new_alias)
- | With_definition_body (_::_,_)
- | With_module_body (_::_,_,_,_) ->
- let old,aliasold = match spec with
- SFBmodule msb -> Some msb, None
- | SFBalias (mpold,typ_opt,cst) ->None, Some (mpold,typ_opt,cst)
- | _ -> error_not_a_module (string_of_label l)
- in
- if aliasold = None then
- let old = Option.get old in
- let new_with_decl,subst1 =
- match with_decl with
- With_definition_body (_,c) -> With_definition_body (idl,c),None
- | With_module_body (idc,mp,typ_opt,cst) ->
- let mp' = scrape_alias mp env' in
- With_module_body (idl,mp,typ_opt,cst),
- Some(map_mp (mp_rec (List.rev idc)) mp')
- in
- let subst = match subst1 with
- | None -> None
- | Some s -> Some (join s (update_subst alias s)) in
- let modtype,subst_msb =
- merge_with env' (type_of_mb env' old) new_with_decl alias in
- let msb =
- { mod_expr = None;
- mod_type = Some modtype;
- mod_constraints = old.mod_constraints;
- mod_alias = begin
- match subst_msb with
- |None -> empty_subst
- |Some s -> s
- end;
- mod_retroknowledge = old.mod_retroknowledge}
- in
- (SFBmodule msb),subst
+
+
+let rec strengthen_mod env mp_from mp_to mb =
+ assert(mp_from = mb.mod_mp);
+ if mp_in_delta mb.mod_mp mb.mod_delta then
+ mb
+ else
+ match mb.mod_type with
+ | SEBstruct (sign) ->
+ let resolve_out,sign_out =
+ strengthen_sig env mp_from sign mp_to mb.mod_delta in
+ { mb with
+ mod_expr = Some (SEBident mp_to);
+ mod_type = SEBstruct(sign_out);
+ mod_type_alg = mb.mod_type_alg;
+ mod_constraints = mb.mod_constraints;
+ mod_delta = add_mp_delta_resolver mp_from mp_to
+ (add_delta_resolver mb.mod_delta resolve_out);
+ mod_retroknowledge = mb.mod_retroknowledge}
+ | SEBfunctor _ -> mb
+ | _ -> anomaly "Modops:the evaluation of the structure failed "
+
+and strengthen_sig env mp_from sign mp_to resolver =
+ match sign with
+ | [] -> empty_delta_resolver,[]
+ | (l,SFBconst cb) :: rest ->
+ let item' =
+ l,SFBconst (strengthen_const env mp_from l cb resolver) in
+ let resolve_out,rest' =
+ strengthen_sig env mp_from rest mp_to resolver in
+ resolve_out,item'::rest'
+ | (_,SFBmind _ as item):: rest ->
+ let resolve_out,rest' =
+ strengthen_sig env mp_from rest mp_to resolver in
+ resolve_out,item::rest'
+ | (l,SFBmodule mb) :: rest ->
+ let mp_from' = MPdot (mp_from,l) in
+ let mp_to' = MPdot(mp_to,l) in
+ let mb_out =
+ strengthen_mod env mp_from' mp_to' mb in
+ let item' = l,SFBmodule (mb_out) in
+ let env' = add_module mb_out env in
+ let resolve_out,rest' =
+ strengthen_sig env' mp_from rest mp_to resolver in
+ add_delta_resolver resolve_out mb.mod_delta,
+ item':: rest'
+ | (l,SFBmodtype mty as item) :: rest ->
+ let env' = add_modtype
+ (MPdot(mp_from,l)) mty env
+ in
+ let resolve_out,rest' =
+ strengthen_sig env' mp_from rest mp_to resolver in
+ resolve_out,item::rest'
+
+let strengthen env mtb mp =
+ if mp_in_delta mtb.typ_mp mtb.typ_delta then
+ (* in this case mtb has already been strengthened*)
+ mtb
+ else
+ match mtb.typ_expr with
+ | SEBstruct (sign) ->
+ let resolve_out,sign_out =
+ strengthen_sig env mtb.typ_mp sign mp mtb.typ_delta in
+ {mtb with
+ typ_expr = SEBstruct(sign_out);
+ typ_delta = add_delta_resolver mtb.typ_delta
+ (add_mp_delta_resolver mtb.typ_mp mp resolve_out)}
+ | SEBfunctor _ -> mtb
+ | _ -> anomaly "Modops:the evaluation of the structure failed "
+
+let module_type_of_module env mp mb =
+ match mp with
+ Some mp ->
+ strengthen env {
+ typ_mp = mp;
+ typ_expr = mb.mod_type;
+ typ_expr_alg = None;
+ typ_constraints = mb.mod_constraints;
+ typ_delta = mb.mod_delta} mp
+
+ | None ->
+ {typ_mp = mb.mod_mp;
+ typ_expr = mb.mod_type;
+ typ_expr_alg = None;
+ typ_constraints = mb.mod_constraints;
+ typ_delta = mb.mod_delta}
+
+let complete_inline_delta_resolver env mp mbid mtb delta =
+ let constants = inline_of_delta mtb.typ_delta in
+ let rec make_inline delta = function
+ | [] -> delta
+ | kn::r ->
+ let kn = replace_mp_in_kn (MPbound mbid) mp kn in
+ let con = constant_of_kn kn in
+ let con' = constant_of_delta delta con in
+ let constant = lookup_constant con' env in
+ if (not constant.Declarations.const_opaque) then
+ let constr = Option.map Declarations.force
+ constant.Declarations.const_body in
+ if constr = None then
+ (make_inline delta r)
else
- let mpold,typ_opt,cst = Option.get aliasold in
- SFBalias (mpold,typ_opt,cst),None
- in
- SEBstruct(msid, before@(l,new_spec)::
- (Option.fold_right subst_structure subst after)),subst
- with
- Not_found -> error_no_such_label l
-
-and add_signature mp sign env =
- let add_one env (l,elem) =
- let kn = make_kn mp empty_dirpath l in
- let con = make_con mp empty_dirpath l in
- match elem with
- | SFBconst cb -> Environ.add_constant con cb env
- | SFBmind mib -> Environ.add_mind kn mib env
- | SFBmodule mb ->
- add_module (MPdot (mp,l)) mb env
- (* adds components as well *)
- | SFBalias (mp1,_,cst) ->
- Environ.register_alias (MPdot(mp,l)) mp1 env
- | SFBmodtype mtb -> Environ.add_modtype (MPdot(mp,l))
- mtb env
+ add_inline_constr_delta_resolver con (Option.get constr)
+ (make_inline delta r)
+ else
+ (make_inline delta r)
in
- List.fold_left add_one env sign
-
-and add_module mp mb env =
- let env = Environ.shallow_add_module mp mb env in
- let env =
- Environ.add_modtype mp (module_type_of_module (Some mp) mb) env
- in
- let mod_typ = type_of_mb env mb in
- match mod_typ with
- | SEBstruct (msid,sign) ->
- add_retroknowledge msid mp (mb.mod_retroknowledge)
- (add_signature mp (subst_signature_msid msid mp sign) env)
- | SEBfunctor _ -> env
- | _ -> anomaly "Modops:the evaluation of the structure failed "
-
-
-
-and constants_of_specification env mp sign =
- let aux (env,res) (l,elem) =
- match elem with
- | SFBconst cb -> env,((make_con mp empty_dirpath l),cb)::res
- | SFBmind _ -> env,res
- | SFBmodule mb ->
- let new_env = add_module (MPdot (mp,l)) mb env in
- new_env,(constants_of_modtype env (MPdot (mp,l))
- (type_of_mb env mb)) @ res
- | SFBalias (mp1,typ_opt,cst) ->
- let new_env = register_alias (MPdot (mp,l)) mp1 env in
- new_env,(constants_of_modtype env (MPdot (mp,l))
- (eval_struct env (SEBident mp1))) @ res
- | SFBmodtype mtb ->
- (* module type dans un module type.
- Il faut au moins mettre mtb dans l'environnement (avec le bon
- kn pour pouvoir continuer aller deplier les modules utilisant ce
- mtb
- ex:
- Module Type T1.
- Module Type T2.
- ....
- End T2.
- .....
- Declare Module M : T2.
- End T2
- si on ne rajoute pas T2 dans l'environement de typage
- on va exploser au moment du Declare Module
- *)
- let new_env = Environ.add_modtype (MPdot(mp,l)) mtb env in
- new_env, (constants_of_modtype env (MPdot(mp,l)) mtb.typ_expr) @ res
- in
- snd (List.fold_left aux (env,[]) sign)
-
-and constants_of_modtype env mp modtype =
- match (eval_struct env modtype) with
- SEBstruct (msid,sign) ->
- constants_of_specification env mp
- (subst_signature_msid msid mp sign)
- | SEBfunctor _ -> []
- | _ -> anomaly "Modops:the evaluation of the structure failed "
-
-(* returns a resolver for kn that maps mbid to mp. We only keep
- constants that have the inline flag *)
-and resolver_of_environment mbid modtype mp alias env =
- let constants = constants_of_modtype env (MPbound mbid) modtype.typ_expr in
- let constants = List.map (fun (l,cb) -> (l,subst_const_body alias cb)) constants in
- let rec make_resolve = function
- | [] -> []
- | (con,expecteddef)::r ->
- let con' = replace_mp_in_con (MPbound mbid) mp con in
- let con',_ = subst_con alias con' in
- (* let con' = replace_mp_in_con (MPbound mbid) mp con' in *)
- try
- if expecteddef.Declarations.const_inline then
- let constant = lookup_constant con' env in
- if (not constant.Declarations.const_opaque) then
- let constr = Option.map Declarations.force
- constant.Declarations.const_body in
- (con,constr)::(make_resolve r)
- else make_resolve r
- else make_resolve r
- with Not_found -> error_no_such_label (con_label con')
- in
- let resolve = make_resolve constants in
- Mod_subst.make_resolver resolve
-
-
-and strengthen_mtb env mp mtb =
- let mtb1 = eval_struct env mtb in
- match mtb1 with
- | SEBfunctor _ -> mtb1
- | SEBstruct (msid,sign) ->
- SEBstruct (msid,strengthen_sig env msid sign mp)
- | _ -> anomaly "Modops:the evaluation of the structure failed "
-
-and strengthen_mod env mp mb =
- let mod_typ = type_of_mb env mb in
- { mod_expr = mb.mod_expr;
- mod_type = Some (strengthen_mtb env mp mod_typ);
- mod_constraints = mb.mod_constraints;
- mod_alias = mb.mod_alias;
- mod_retroknowledge = mb.mod_retroknowledge}
-
-and strengthen_sig env msid sign mp = match sign with
- | [] -> []
- | (l,SFBconst cb) :: rest ->
- let item' = l,SFBconst (strengthen_const env mp l cb) in
- let rest' = strengthen_sig env msid rest mp in
+ make_inline delta constants
+
+let rec strengthen_and_subst_mod
+ mb subst env mp_from mp_to env resolver =
+ match mb.mod_type with
+ SEBstruct(str) ->
+ let mb_is_an_alias = mp_in_delta mb.mod_mp mb.mod_delta in
+ if mb_is_an_alias then
+ subst_module subst
+ (fun resolver -> subst_dom_delta_resolver subst resolver) mb
+ else
+ let resolver,new_sig =
+ strengthen_and_subst_struct str subst env
+ mp_from mp_from mp_to false false mb.mod_delta
+ in
+ {mb with
+ mod_mp = mp_to;
+ mod_expr = Some (SEBident mp_from);
+ mod_type = SEBstruct(new_sig);
+ mod_delta = add_mp_delta_resolver mp_to mp_from resolver}
+ | SEBfunctor(arg_id,arg_b,body) ->
+ subst_module subst
+ (fun resolver -> subst_dom_delta_resolver subst resolver) mb
+ | _ -> anomaly "Modops:the evaluation of the structure failed "
+
+and strengthen_and_subst_struct
+ str subst env mp_alias mp_from mp_to alias incl resolver =
+ match str with
+ | [] -> empty_delta_resolver,[]
+ | (l,SFBconst cb) :: rest ->
+ let item' = if alias then
+ l,SFBconst (subst_const_body subst cb)
+ else
+ l,SFBconst (strengthen_const env mp_from l
+ (subst_const_body subst cb) resolver)
+ in
+ let con = make_con mp_from empty_dirpath l in
+ let resolve_out,rest' =
+ strengthen_and_subst_struct rest subst env
+ mp_alias mp_from mp_to alias incl resolver in
+ if incl && not (con_in_delta con resolver) then
+ (add_constant_delta_resolver
+ (make_con_equiv mp_to mp_alias empty_dirpath l) resolve_out),
item'::rest'
- | (l,SFBmind mib) :: rest ->
- let item' = l,SFBmind (strengthen_mind env mp l mib) in
- let rest' = strengthen_sig env msid rest mp in
+ else
+ resolve_out,item'::rest'
+ | (l,SFBmind mib) :: rest ->
+ let item' = l,SFBmind (subst_mind subst mib) in
+ let mind = make_mind mp_from empty_dirpath l in
+ let resolve_out,rest' =
+ strengthen_and_subst_struct rest subst env
+ mp_alias mp_from mp_to alias incl resolver in
+ if incl && not (mind_in_delta mind resolver) then
+ (add_mind_delta_resolver
+ (make_mind_equiv mp_to mp_alias empty_dirpath l) resolve_out),
item'::rest'
- | (l,SFBmodule mb) :: rest ->
- let mp' = MPdot (mp,l) in
- let item' = l,SFBmodule (strengthen_mod env mp' mb) in
- let env' = add_module
- (MPdot (MPself msid,l)) mb env in
- let rest' = strengthen_sig env' msid rest mp in
+ else
+ resolve_out,item'::rest'
+ | (l,SFBmodule mb) :: rest ->
+ let mp_from' = MPdot (mp_from,l) in
+ let mp_to' = MPdot(mp_to,l) in
+ let mb_out = if alias then
+ subst_module subst
+ (fun resolver -> subst_dom_delta_resolver subst resolver) mb
+ else
+ strengthen_and_subst_mod
+ mb subst env mp_from' mp_to' env resolver
+ in
+ let item' = l,SFBmodule (mb_out) in
+ let env' = add_module mb_out env in
+ let resolve_out,rest' =
+ strengthen_and_subst_struct rest subst env'
+ mp_alias mp_from mp_to alias incl resolver in
+ add_delta_resolver resolve_out mb_out.mod_delta,
item':: rest'
- | ((l,SFBalias (mp1,_,cst)) as item) :: rest ->
- let env' = register_alias (MPdot(MPself msid,l)) mp1 env in
- let rest' = strengthen_sig env' msid rest mp in
- item::rest'
- | (l,SFBmodtype mty as item) :: rest ->
- let env' = add_modtype
- (MPdot((MPself msid),l))
- mty
- env
- in
- let rest' = strengthen_sig env' msid rest mp in
- item::rest'
-
-
-let strengthen env mtb mp = strengthen_mtb env mp mtb
-
-
-let update_subst env mb mp =
- match type_of_mb env mb with
- | SEBstruct(msid,str) -> false, join_alias
- (subst_key (map_msid msid mp) mb.mod_alias)
- (map_msid msid mp)
- | _ -> true, mb.mod_alias
+ | (l,SFBmodtype mty) :: rest ->
+ let mp_from' = MPdot (mp_from,l) in
+ let mp_to' = MPdot(mp_to,l) in
+ let subst' = add_mp mp_from' mp_to' empty_delta_resolver subst in
+ let mty = subst_modtype subst'
+ (fun resolver -> subst_dom_codom_delta_resolver subst' resolver) mty in
+ let env' = add_modtype mp_from' mty env in
+ let resolve_out,rest' = strengthen_and_subst_struct rest subst env'
+ mp_alias mp_from mp_to alias incl resolver in
+ (add_mp_delta_resolver
+ mp_to' mp_to' resolve_out),(l,SFBmodtype mty)::rest'
+
+let strengthen_and_subst_mb mb mp env include_b =
+ match mb.mod_type with
+ SEBstruct str ->
+ let mb_is_an_alias = mp_in_delta mb.mod_mp mb.mod_delta in
+ (*if mb is an alias then the strengthening is useless
+ (i.e. it is already done)*)
+ let mp_alias = delta_of_mp mb.mod_delta mb.mod_mp in
+ let subst_resolver = map_mp mb.mod_mp mp empty_delta_resolver in
+ let resolver =
+ add_mp_delta_resolver mp mp_alias
+ (subst_dom_delta_resolver subst_resolver mb.mod_delta) in
+ let subst = map_mp mb.mod_mp mp resolver in
+ let resolver = if mb_is_an_alias && include_b then
+ remove_mp_delta_resolver mb.mod_delta mb.mod_mp
+ else mb.mod_delta in
+ let resolver,new_sig =
+ strengthen_and_subst_struct str subst env
+ mp_alias mb.mod_mp mp mb_is_an_alias include_b resolver
+ in
+ {mb with
+ mod_mp = mp;
+ mod_type = SEBstruct(new_sig);
+ mod_expr = Some (SEBident mb.mod_mp);
+ mod_delta = if include_b then resolver
+ else add_mp_delta_resolver mp mp_alias resolver}
+ | SEBfunctor(arg_id,argb,body) ->
+ let subst = map_mp mb.mod_mp mp empty_delta_resolver in
+ subst_module subst
+ (fun resolver -> subst_dom_codom_delta_resolver subst resolver) mb
+ | _ -> anomaly "Modops:the evaluation of the structure failed "
+
+
+let subst_modtype_and_resolver mtb mp env =
+ let subst = (map_mp mtb.typ_mp mp empty_delta_resolver) in
+ let new_delta = subst_dom_codom_delta_resolver subst mtb.typ_delta in
+ let full_subst = (map_mp mtb.typ_mp mp new_delta) in
+ subst_modtype full_subst
+ (fun resolver -> subst_dom_codom_delta_resolver subst resolver) mtb
diff --git a/kernel/modops.mli b/kernel/modops.mli
index 4cd72a2ef..44d29ee3e 100644
--- a/kernel/modops.mli
+++ b/kernel/modops.mli
@@ -20,46 +20,38 @@ open Mod_subst
(* Various operations on modules and module types *)
-(* make the environment entry out of type *)
-val module_body_of_type : module_type_body -> module_body
-val module_type_of_module : module_path option -> module_body ->
+val module_body_of_type : module_path -> module_type_body -> module_body
+
+val module_type_of_module : env -> module_path option -> module_body ->
module_type_body
val destr_functor :
env -> struct_expr_body -> mod_bound_id * module_type_body * struct_expr_body
-val subst_modtype : substitution -> module_type_body -> module_type_body
-val subst_structure : substitution -> structure_body -> structure_body
-
val subst_struct_expr : substitution -> struct_expr_body -> struct_expr_body
-val subst_signature_msid :
- mod_self_id -> module_path ->
- structure_body -> structure_body
-
-val subst_structure : substitution -> structure_body -> structure_body
-
-(* Evaluation functions *)
-val eval_struct : env -> struct_expr_body -> struct_expr_body
-
-val type_of_mb : env -> module_body -> struct_expr_body
+val subst_signature : substitution -> structure_body -> structure_body
-(* [add_signature mp sign env] assumes that the substitution [msid]
- $\mapsto$ [mp] has already been performed (or is not necessary, like
- when [mp = MPself msid]) *)
val add_signature :
- module_path -> structure_body -> env -> env
+ module_path -> structure_body -> delta_resolver -> env -> env
(* adds a module and its components, but not the constraints *)
-val add_module :
- module_path -> module_body -> env -> env
+val add_module : module_body -> env -> env
val check_modpath_equiv : env -> module_path -> module_path -> unit
-val strengthen : env -> struct_expr_body -> module_path -> struct_expr_body
+val strengthen : env -> module_type_body -> module_path -> module_type_body
-val update_subst : env -> module_body -> module_path -> bool * substitution
+val complete_inline_delta_resolver :
+ env -> module_path -> mod_bound_id -> module_type_body ->
+ delta_resolver -> delta_resolver
+
+val strengthen_and_subst_mb : module_body -> module_path -> env -> bool
+ -> module_body
+
+val subst_modtype_and_resolver : module_type_body -> module_path -> env ->
+ module_type_body
val error_existing_label : label -> 'a
@@ -102,9 +94,7 @@ val error_a_generative_module_expected : label -> 'a
val error_local_context : label option -> 'a
-val error_no_such_label_sub : label->string->string->'a
+val error_no_such_label_sub : label->string->'a
+
-val resolver_of_environment :
- mod_bound_id -> module_type_body -> module_path -> substitution
- -> env -> resolver
diff --git a/kernel/names.ml b/kernel/names.ml
index 0d61a29aa..06ca87b4d 100644
--- a/kernel/names.ml
+++ b/kernel/names.ml
@@ -108,7 +108,7 @@ module Labmap = Idmap
type module_path =
| MPfile of dir_path
| MPbound of mod_bound_id
- | MPself of mod_self_id
+ (* | MPapp of module_path * module_path *)
| MPdot of module_path * label
let rec check_bound_mp = function
@@ -119,7 +119,9 @@ let rec check_bound_mp = function
let rec string_of_mp = function
| MPfile sl -> "MPfile (" ^ string_of_dirpath sl ^ ")"
| MPbound uid -> string_of_uid uid
- | MPself uid -> string_of_uid uid
+ (* | MPapp (mp1,mp2) ->
+ "("^string_of_mp mp ^ " " ^
+ string_of_mp mp^")"*)
| MPdot (mp,l) -> string_of_mp mp ^ "." ^ string_of_label l
(* we compare labels first if both are MPdots *)
@@ -172,6 +174,28 @@ let kn_ord kn1 kn2 =
else
MPord.compare mp1 mp2
+(* a constant name is a kernel name couple (kn1,kn2)
+ where kn1 corresponds to the name used at toplevel
+ (i.e. what the user see)
+ and kn2 corresponds to the canonical kernel name
+ i.e. in the environment we have
+ kn1 \rhd_{\delta}^* kn2 \rhd_{\delta} t *)
+type constant = kernel_name*kernel_name
+
+(* For the environment we distinguish constants by their
+ user part*)
+module User_ord = struct
+ type t = kernel_name*kernel_name
+ let compare x y= kn_ord (fst x) (fst y)
+end
+
+(* For other uses (ex: non-logical things) it is enough
+ to deal with the canonical part *)
+module Canonical_ord = struct
+ type t = kernel_name*kernel_name
+ let compare x y= kn_ord (snd x) (snd y)
+end
+
module KNord = struct
type t = kernel_name
@@ -181,47 +205,89 @@ end
module KNmap = Map.Make(KNord)
module KNpred = Predicate.Make(KNord)
module KNset = Set.Make(KNord)
-module Cmap = KNmap
-module Cpred = KNpred
-module Cset = KNset
+
+module Cmap = Map.Make(Canonical_ord)
+module Cmap_env = Map.Make(User_ord)
+module Cpred = Predicate.Make(Canonical_ord)
+module Cset = Set.Make(Canonical_ord)
+module Cset_env = Set.Make(User_ord)
+
+module Mindmap = Map.Make(Canonical_ord)
+module Mindset = Set.Make(Canonical_ord)
+module Mindmap_env = Map.Make(User_ord)
let default_module_name = "If you see this, it's a bug"
let initial_dir = make_dirpath [default_module_name]
let initial_msid = (make_msid initial_dir "If you see this, it's a bug")
-let initial_path = MPself initial_msid
+let initial_path = MPfile initial_dir
type variable = identifier
-type constant = kernel_name
-type mutual_inductive = kernel_name
+
+(* The same thing is done for mutual inductive names
+ it replaces also the old mind_equiv field of mutual
+ inductive types*)
+type mutual_inductive = kernel_name*kernel_name
type inductive = mutual_inductive * int
type constructor = inductive * int
-let constant_of_kn kn = kn
-let make_con mp dir l = (mp,dir,l)
-let repr_con con = con
-let string_of_con = string_of_kn
-let con_label = label
-let pr_con = pr_kn
-let con_modpath = modpath
-
-let mind_modpath = modpath
+let constant_of_kn kn = (kn,kn)
+let constant_of_kn_equiv kn1 kn2 = (kn1,kn2)
+let make_con mp dir l = ((mp,dir,l),(mp,dir,l))
+let make_con_equiv mp1 mp2 dir l = ((mp1,dir,l),(mp2,dir,l))
+let canonical_con con = snd con
+let user_con con = fst con
+let repr_con con = fst con
+let string_of_con con = string_of_kn (fst con)
+let con_label con = label (fst con)
+let pr_con con = pr_kn (fst con)
+let debug_pr_con con = str "("++ pr_kn (fst con) ++ str ","++ pr_kn (snd con)++ str ")"
+let eq_constant (_,kn1) (_,kn2) = kn1=kn2
+let debug_string_of_con con = string_of_kn (fst con)^"'"^string_of_kn (snd con)
+
+let con_modpath con = modpath (fst con)
+
+let mind_modpath mind = modpath (fst mind)
let ind_modpath ind = mind_modpath (fst ind)
let constr_modpath c = ind_modpath (fst c)
+
+let mind_of_kn kn = (kn,kn)
+let mind_of_kn_equiv kn1 kn2 = (kn1,kn2)
+let make_mind mp dir l = ((mp,dir,l),(mp,dir,l))
+let make_mind_equiv mp1 mp2 dir l = ((mp1,dir,l),(mp2,dir,l))
+let canonical_mind mind = snd mind
+let user_mind mind = fst mind
+let repr_mind mind = fst mind
+let string_of_mind mind = string_of_kn (fst mind)
+let mind_label mind= label (fst mind)
+let pr_mind mind = pr_kn (fst mind)
+let debug_pr_mind mind = str "("++ pr_kn (fst mind) ++ str ","++ pr_kn (snd mind)++ str ")"
+let eq_mind (_,kn1) (_,kn2) = kn1=kn2
+let debug_string_of_mind mind = string_of_kn (fst mind)^"'"^string_of_kn (snd mind)
+
let ith_mutual_inductive (kn,_) i = (kn,i)
let ith_constructor_of_inductive ind i = (ind,i)
let inductive_of_constructor (ind,i) = ind
let index_of_constructor (ind,i) = i
+let eq_ind (kn1,i1) (kn2,i2) = i1=i2&&eq_mind kn1 kn2
+let eq_constructor (kn1,i1) (kn2,i2) = i1=i2&&eq_ind kn1 kn2
module InductiveOrdered = struct
type t = inductive
let compare (spx,ix) (spy,iy) =
- let c = ix - iy in if c = 0 then KNord.compare spx spy else c
+ let c = ix - iy in if c = 0 then Canonical_ord.compare spx spy else c
+end
+
+module InductiveOrdered_env = struct
+ type t = inductive
+ let compare (spx,ix) (spy,iy) =
+ let c = ix - iy in if c = 0 then User_ord.compare spx spy else c
end
module Indmap = Map.Make(InductiveOrdered)
+module Indmap_env = Map.Make(InductiveOrdered_env)
module ConstructorOrdered = struct
type t = constructor
@@ -229,13 +295,24 @@ module ConstructorOrdered = struct
let c = ix - iy in if c = 0 then InductiveOrdered.compare indx indy else c
end
+module ConstructorOrdered_env = struct
+ type t = constructor
+ let compare (indx,ix) (indy,iy) =
+ let c = ix - iy in if c = 0 then InductiveOrdered_env.compare indx indy else c
+end
+
module Constrmap = Map.Make(ConstructorOrdered)
+module Constrmap_env = Map.Make(ConstructorOrdered_env)
(* Better to have it here that in closure, since used in grammar.cma *)
type evaluable_global_reference =
| EvalVarRef of identifier
| EvalConstRef of constant
+let eq_egr e1 e2 = match e1,e2 with
+ EvalConstRef con1, EvalConstRef con2 -> eq_constant con1 con2
+ | _,_ -> e1 = e2
+
(* Hash-consing of name objects *)
module Hname = Hashcons.Make(
struct
@@ -281,24 +358,24 @@ module Hmod = Hashcons.Make(
let rec hash_sub (hdir,huniqid,hstr as hfuns) = function
| MPfile dir -> MPfile (hdir dir)
| MPbound m -> MPbound (huniqid m)
- | MPself m -> MPself (huniqid m)
| MPdot (md,l) -> MPdot (hash_sub hfuns md, hstr l)
let rec equal d1 d2 = match (d1,d2) with
| MPfile dir1, MPfile dir2 -> dir1 == dir2
| MPbound m1, MPbound m2 -> m1 == m2
- | MPself m1, MPself m2 -> m1 == m2
| MPdot (mod1,l1), MPdot (mod2,l2) -> equal mod1 mod2 & l1 = l2
| _ -> false
let hash = Hashtbl.hash
end)
-module Hkn = Hashcons.Make(
- struct
- type t = kernel_name
+
+module Hcn = Hashcons.Make(
+ struct
+ type t = kernel_name*kernel_name
type u = (module_path -> module_path)
* (dir_path -> dir_path) * (string -> string)
- let hash_sub (hmod,hdir,hstr) (md,dir,l) = (hmod md, hdir dir, hstr l)
- let equal (mod1,dir1,l1) (mod2,dir2,l2) =
+ let hash_sub (hmod,hdir,hstr) ((md,dir,l),(mde,dire,le)) =
+ ((hmod md, hdir dir, hstr l),(hmod mde, hdir dire, hstr le))
+ let equal ((mod1,dir1,l1),_) ((mod2,dir2,l2),_) =
mod1 == mod2 && dir1 == dir2 && l1 == l2
let hash = Hashtbl.hash
end)
@@ -310,8 +387,9 @@ let hcons_names () =
let hdir = Hashcons.simple_hcons Hdir.f hident in
let huniqid = Hashcons.simple_hcons Huniqid.f (hstring,hdir) in
let hmod = Hashcons.simple_hcons Hmod.f (hdir,huniqid,hstring) in
- let hkn = Hashcons.simple_hcons Hkn.f (hmod,hdir,hstring) in
- (hkn,hkn,hdir,hname,hident,hstring)
+ let hmind = Hashcons.simple_hcons Hcn.f (hmod,hdir,hstring) in
+ let hcn = Hashcons.simple_hcons Hcn.f (hmod,hdir,hstring) in
+ (hcn,hmind,hdir,hname,hident,hstring)
(*******)
@@ -335,3 +413,8 @@ type inv_rel_key = int (* index in the [rel_context] part of environment
type id_key = inv_rel_key tableKey
+let eq_id_key ik1 ik2 =
+ match ik1,ik2 with
+ ConstKey (_,kn1),
+ ConstKey (_,kn2) -> kn1=kn2
+ | a,b -> a=b
diff --git a/kernel/names.mli b/kernel/names.mli
index fb3b5c81b..e42f8ea71 100644
--- a/kernel/names.mli
+++ b/kernel/names.mli
@@ -80,9 +80,9 @@ module Labmap : Map.S with type key = label
type module_path =
| MPfile of dir_path
| MPbound of mod_bound_id
- | MPself of mod_self_id
+ (* | MPapp of module_path * module_path very soon *)
| MPdot of module_path * label
-(*i | MPapply of module_path * module_path in the future (maybe) i*)
+
val check_bound_mp : module_path -> bool
@@ -122,25 +122,64 @@ module KNmap : Map.S with type key = kernel_name
type variable = identifier
type constant
-type mutual_inductive = kernel_name
+type mutual_inductive
(* Beware: first inductive has index 0 *)
type inductive = mutual_inductive * int
(* Beware: first constructor has index 1 *)
type constructor = inductive * int
+(* *_env modules consider an order on user part of names
+ the others consider an order on canonical part of names*)
module Cmap : Map.S with type key = constant
+module Cmap_env : Map.S with type key = constant
module Cpred : Predicate.S with type elt = constant
module Cset : Set.S with type elt = constant
+module Cset_env : Set.S with type elt = constant
+module Mindmap : Map.S with type key = mutual_inductive
+module Mindmap_env : Map.S with type key = mutual_inductive
+module Mindset : Set.S with type elt = mutual_inductive
module Indmap : Map.S with type key = inductive
module Constrmap : Map.S with type key = constructor
+module Indmap_env : Map.S with type key = inductive
+module Constrmap_env : Map.S with type key = constructor
val constant_of_kn : kernel_name -> constant
+val constant_of_kn_equiv : kernel_name -> kernel_name -> constant
val make_con : module_path -> dir_path -> label -> constant
+val make_con_equiv : module_path -> module_path -> dir_path
+ -> label -> constant
+val user_con : constant -> kernel_name
+val canonical_con : constant -> kernel_name
val repr_con : constant -> module_path * dir_path * label
+val eq_constant : constant -> constant -> bool
+
val string_of_con : constant -> string
val con_label : constant -> label
val con_modpath : constant -> module_path
val pr_con : constant -> Pp.std_ppcmds
+val debug_pr_con : constant -> Pp.std_ppcmds
+val debug_string_of_con : constant -> string
+
+
+
+val mind_of_kn : kernel_name -> mutual_inductive
+val mind_of_kn_equiv : kernel_name -> kernel_name -> mutual_inductive
+val make_mind : module_path -> dir_path -> label -> mutual_inductive
+val make_mind_equiv : module_path -> module_path -> dir_path
+ -> label -> mutual_inductive
+val user_mind : mutual_inductive -> kernel_name
+val canonical_mind : mutual_inductive -> kernel_name
+val repr_mind : mutual_inductive -> module_path * dir_path * label
+val eq_mind : mutual_inductive -> mutual_inductive -> bool
+
+val string_of_mind : mutual_inductive -> string
+val mind_label : mutual_inductive -> label
+val mind_modpath : mutual_inductive -> module_path
+val pr_mind : mutual_inductive -> Pp.std_ppcmds
+val debug_pr_mind : mutual_inductive -> Pp.std_ppcmds
+val debug_string_of_mind : mutual_inductive -> string
+
+
val mind_modpath : mutual_inductive -> module_path
val ind_modpath : inductive -> module_path
@@ -150,16 +189,21 @@ val ith_mutual_inductive : inductive -> int -> inductive
val ith_constructor_of_inductive : inductive -> int -> constructor
val inductive_of_constructor : constructor -> inductive
val index_of_constructor : constructor -> int
+val eq_ind : inductive -> inductive -> bool
+val eq_constructor : constructor -> constructor -> bool
(* Better to have it here that in Closure, since required in grammar.cma *)
type evaluable_global_reference =
| EvalVarRef of identifier
| EvalConstRef of constant
+val eq_egr : evaluable_global_reference -> evaluable_global_reference
+ -> bool
+
(* Hash-consing *)
val hcons_names : unit ->
(constant -> constant) *
- (kernel_name -> kernel_name) * (dir_path -> dir_path) *
+ (mutual_inductive -> mutual_inductive) * (dir_path -> dir_path) *
(name -> name) * (identifier -> identifier) * (string -> string)
@@ -182,3 +226,5 @@ type inv_rel_key = int (* index in the [rel_context] part of environment
of de Bruijn indice *)
type id_key = inv_rel_key tableKey
+
+val eq_id_key : id_key -> id_key -> bool
diff --git a/kernel/pre_env.ml b/kernel/pre_env.ml
index 421672201..b58951e98 100644
--- a/kernel/pre_env.ml
+++ b/kernel/pre_env.ml
@@ -23,11 +23,10 @@ type key = int option ref
type constant_key = constant_body * key
type globals = {
- env_constants : constant_key Cmap.t;
- env_inductives : mutual_inductive_body KNmap.t;
+ env_constants : constant_key Cmap_env.t;
+ env_inductives : mutual_inductive_body Mindmap_env.t;
env_modules : module_body MPmap.t;
- env_modtypes : module_type_body MPmap.t;
- env_alias : module_path MPmap.t }
+ env_modtypes : module_type_body MPmap.t}
type stratification = {
env_universes : universes;
@@ -58,11 +57,10 @@ let empty_named_context_val = [],[]
let empty_env = {
env_globals = {
- env_constants = Cmap.empty;
- env_inductives = KNmap.empty;
+ env_constants = Cmap_env.empty;
+ env_inductives = Mindmap_env.empty;
env_modules = MPmap.empty;
- env_modtypes = MPmap.empty;
- env_alias = MPmap.empty };
+ env_modtypes = MPmap.empty};
env_named_context = empty_named_context;
env_named_vals = [];
env_rel_context = empty_rel_context;
@@ -123,16 +121,12 @@ let env_of_named id env = env
(* Global constants *)
let lookup_constant_key kn env =
- Cmap.find kn env.env_globals.env_constants
+ Cmap_env.find kn env.env_globals.env_constants
let lookup_constant kn env =
- fst (Cmap.find kn env.env_globals.env_constants)
+ fst (Cmap_env.find kn env.env_globals.env_constants)
(* Mutual Inductives *)
let lookup_mind kn env =
- KNmap.find kn env.env_globals.env_inductives
+ Mindmap_env.find kn env.env_globals.env_inductives
-let rec scrape_mind env kn =
- match (lookup_mind kn env).mind_equiv with
- | None -> kn
- | Some kn' -> scrape_mind env kn'
diff --git a/kernel/pre_env.mli b/kernel/pre_env.mli
index abbf9b1b5..718132b32 100644
--- a/kernel/pre_env.mli
+++ b/kernel/pre_env.mli
@@ -23,11 +23,10 @@ type key = int option ref
type constant_key = constant_body * key
type globals = {
- env_constants : constant_key Cmap.t;
- env_inductives : mutual_inductive_body KNmap.t;
+ env_constants : constant_key Cmap_env.t;
+ env_inductives : mutual_inductive_body Mindmap_env.t;
env_modules : module_body MPmap.t;
- env_modtypes : module_type_body MPmap.t;
- env_alias : module_path MPmap.t }
+ env_modtypes : module_type_body MPmap.t}
type stratification = {
env_universes : universes;
@@ -80,5 +79,3 @@ val lookup_constant : constant -> env -> constant_body
(* Mutual Inductives *)
val lookup_mind : mutual_inductive -> env -> mutual_inductive_body
-(* Find the ultimate inductive in the [mind_equiv] chain *)
-val scrape_mind : env -> mutual_inductive -> mutual_inductive
diff --git a/kernel/reduction.ml b/kernel/reduction.ml
index 0a404fff3..92386c4d3 100644
--- a/kernel/reduction.ml
+++ b/kernel/reduction.ml
@@ -274,7 +274,7 @@ and eqappr cv_pb infos (lft1,st1) (lft2,st2) cuniv =
(* 2 constants, 2 local defined vars or 2 defined rels *)
| (FFlex fl1, FFlex fl2) ->
(try (* try first intensional equality *)
- if fl1 = fl2
+ if eq_table_key fl1 fl2
then convert_stacks infos lft1 lft2 v1 v2 cuniv
else raise NotConvertible
with NotConvertible ->
@@ -325,13 +325,13 @@ and eqappr cv_pb infos (lft1,st1) (lft2,st2) cuniv =
(* Inductive types: MutInd MutConstruct Fix Cofix *)
| (FInd ind1, FInd ind2) ->
- if mind_equiv_infos (snd infos) ind1 ind2
+ if eq_ind ind1 ind2
then
convert_stacks infos lft1 lft2 v1 v2 cuniv
else raise NotConvertible
| (FConstruct (ind1,j1), FConstruct (ind2,j2)) ->
- if j1 = j2 && mind_equiv_infos (snd infos) ind1 ind2
+ if j1 = j2 && eq_ind ind1 ind2
then
convert_stacks infos lft1 lft2 v1 v2 cuniv
else raise NotConvertible
@@ -377,7 +377,7 @@ and eqappr cv_pb infos (lft1,st1) (lft2,st2) cuniv =
and convert_stacks infos lft1 lft2 stk1 stk2 cuniv =
compare_stacks
(fun (l1,t1) (l2,t2) c -> ccnv CONV infos l1 l2 t1 t2 c)
- (mind_equiv_infos (snd infos))
+ (eq_ind)
lft1 stk1 lft2 stk2 cuniv
and convert_vect infos lft1 lft2 v1 v2 cuniv =
diff --git a/kernel/safe_typing.ml b/kernel/safe_typing.ml
index e73689bc8..b0dc6dd8a 100644
--- a/kernel/safe_typing.ml
+++ b/kernel/safe_typing.ml
@@ -35,13 +35,12 @@ type modvariant =
| LIBRARY of dir_path
type module_info =
- { msid : mod_self_id;
- modpath : module_path;
- seed : dir_path; (* the "seed" of unique identifier generator *)
- label : label;
- variant : modvariant;
- alias_subst : substitution}
-
+ {modpath : module_path;
+ label : label;
+ variant : modvariant;
+ resolver : delta_resolver;
+ resolver_of_param : delta_resolver;}
+
let check_label l labset =
if Labset.mem l labset then error_existing_label l
@@ -80,12 +79,11 @@ let rec empty_environment =
{ old = empty_environment;
env = empty_env;
modinfo = {
- msid = initial_msid;
modpath = initial_path;
- seed = initial_dir;
label = mk_label "_";
variant = NONE;
- alias_subst = empty_subst};
+ resolver = empty_delta_resolver;
+ resolver_of_param = empty_delta_resolver};
labset = Labset.empty;
revstruct = [];
univ = Univ.Constraint.empty;
@@ -215,9 +213,16 @@ let add_constant dir l decl senv =
in
let senv' = add_constraints cb.const_constraints senv in
let env'' = Environ.add_constant kn cb senv'.env in
+ let resolver =
+ if cb.const_inline then
+ add_inline_delta_resolver kn senv'.modinfo.resolver
+ else
+ senv'.modinfo.resolver
+ in
kn, { old = senv'.old;
env = env'';
- modinfo = senv'.modinfo;
+ modinfo = {senv'.modinfo with
+ resolver = resolver};
labset = Labset.add l senv'.labset;
revstruct = (l,SFBconst cb)::senv'.revstruct;
univ = senv'.univ;
@@ -242,7 +247,7 @@ let add_mind dir l mie senv =
all labels *)
let mib = translate_mind senv.env mie in
let senv' = add_constraints mib.mind_constraints senv in
- let kn = make_kn senv.modinfo.modpath dir l in
+ let kn = make_mind senv.modinfo.modpath dir l in
let env'' = Environ.add_mind kn mib senv'.env in
kn, { old = senv'.old;
env = env'';
@@ -259,118 +264,82 @@ let add_mind dir l mie senv =
let add_modtype l mte senv =
check_label l senv.labset;
- let mtb_expr,sub = translate_struct_entry senv.env mte in
- let mtb = { typ_expr = mtb_expr;
- typ_strength = None;
- typ_alias = sub} in
- let senv' = add_constraints
- (struct_expr_constraints mtb_expr) senv in
let mp = MPdot(senv.modinfo.modpath, l) in
- let env'' = Environ.add_modtype mp mtb senv'.env in
- mp, { old = senv'.old;
- env = env'';
- modinfo = senv'.modinfo;
- labset = Labset.add l senv'.labset;
- revstruct = (l,SFBmodtype mtb)::senv'.revstruct;
- univ = senv'.univ;
- engagement = senv'.engagement;
- imports = senv'.imports;
- loads = senv'.loads;
- local_retroknowledge = senv'.local_retroknowledge }
+ let mtb = translate_module_type senv.env mp mte in
+ let senv' = add_constraints mtb.typ_constraints senv in
+ let env'' = Environ.add_modtype mp mtb senv'.env in
+ mp, { old = senv'.old;
+ env = env'';
+ modinfo = senv'.modinfo;
+ labset = Labset.add l senv'.labset;
+ revstruct = (l,SFBmodtype mtb)::senv'.revstruct;
+ univ = senv'.univ;
+ engagement = senv'.engagement;
+ imports = senv'.imports;
+ loads = senv'.loads;
+ local_retroknowledge = senv'.local_retroknowledge }
(* full_add_module adds module with universes and constraints *)
-let full_add_module mp mb senv =
- let senv = add_constraints (module_constraints mb) senv in
- let env = Modops.add_module mp mb senv.env in
+let full_add_module mb senv =
+ let senv = add_constraints mb.mod_constraints senv in
+ let env = Modops.add_module mb senv.env in
{senv with env = env}
(* Insertion of modules *)
let add_module l me senv =
check_label l senv.labset;
- let mb = translate_module senv.env me in
let mp = MPdot(senv.modinfo.modpath, l) in
- let senv' = full_add_module mp mb senv in
- let is_functor,sub = Modops.update_subst senv'.env mb mp in
- mp, { old = senv'.old;
- env = senv'.env;
- modinfo =
- if is_functor then
- senv'.modinfo
- else
- {senv'.modinfo with
- alias_subst = join senv'.modinfo.alias_subst sub};
- labset = Labset.add l senv'.labset;
- revstruct = (l,SFBmodule mb)::senv'.revstruct;
- univ = senv'.univ;
- engagement = senv'.engagement;
- imports = senv'.imports;
- loads = senv'.loads;
- local_retroknowledge = senv'.local_retroknowledge }
-
-let add_alias l mp senv =
- check_label l senv.labset;
- let mp' = MPdot(senv.modinfo.modpath, l) in
- let mp1 = scrape_alias mp senv.env in
- let typ_opt =
- if check_bound_mp mp then
- Some (strengthen senv.env
- (lookup_modtype mp senv.env).typ_expr mp)
- else
- None
+ let mb = translate_module senv.env mp me in
+ let senv' = full_add_module mb senv in
+ let modinfo = match mb.mod_type with
+ SEBstruct _ ->
+ { senv'.modinfo with
+ resolver =
+ add_delta_resolver mb.mod_delta senv'.modinfo.resolver}
+ | _ -> senv'.modinfo
in
- (* we get all updated alias substitution {mp1.K\M} that comes from mp1 *)
- let _,sub = Modops.update_subst senv.env (lookup_module mp1 senv.env) mp1 in
- (* transformation of {mp1.K\M} to {mp.K\M}*)
- let sub = update_subst sub (map_mp mp' mp1) in
- (* transformation of {mp.K\M} to {mp.K\M'} where M'=M{mp1\mp'}*)
- let sub = join_alias sub (map_mp mp' mp1) in
- (* we add the alias substitution *)
- let sub = add_mp mp' mp1 sub in
- let env' = register_alias mp' mp senv.env in
- mp', { old = senv.old;
- env = env';
- modinfo = { senv.modinfo with
- alias_subst = join
- senv.modinfo.alias_subst sub};
- labset = Labset.add l senv.labset;
- revstruct = (l,SFBalias (mp,typ_opt,None))::senv.revstruct;
- univ = senv.univ;
- engagement = senv.engagement;
- imports = senv.imports;
- loads = senv.loads;
- local_retroknowledge = senv.local_retroknowledge }
-
+ mp,mb.mod_delta,
+ { old = senv'.old;
+ env = senv'.env;
+ modinfo = modinfo;
+ labset = Labset.add l senv'.labset;
+ revstruct = (l,SFBmodule mb)::senv'.revstruct;
+ univ = senv'.univ;
+ engagement = senv'.engagement;
+ imports = senv'.imports;
+ loads = senv'.loads;
+ local_retroknowledge = senv'.local_retroknowledge }
+
(* Interactive modules *)
let start_module l senv =
check_label l senv.labset;
- let msid = make_msid senv.modinfo.seed (string_of_label l) in
- let mp = MPself msid in
- let modinfo = { msid = msid;
- modpath = mp;
- seed = senv.modinfo.seed;
- label = l;
- variant = STRUCT [];
- alias_subst = empty_subst}
- in
- mp, { old = senv;
- env = senv.env;
- modinfo = modinfo;
- labset = Labset.empty;
- revstruct = [];
- univ = Univ.Constraint.empty;
- engagement = None;
- imports = senv.imports;
- loads = [];
- (* spiwack : not sure, but I hope it's correct *)
- local_retroknowledge = [] }
+ let mp = MPdot(senv.modinfo.modpath, l) in
+ let modinfo = { modpath = mp;
+ label = l;
+ variant = STRUCT [];
+ resolver = empty_delta_resolver;
+ resolver_of_param = empty_delta_resolver}
+ in
+ mp, { old = senv;
+ env = senv.env;
+ modinfo = modinfo;
+ labset = Labset.empty;
+ revstruct = [];
+ univ = Univ.Constraint.empty;
+ engagement = None;
+ imports = senv.imports;
+ loads = [];
+ (* spiwack : not sure, but I hope it's correct *)
+ local_retroknowledge = [] }
let end_module l restype senv =
let oldsenv = senv.old in
let modinfo = senv.modinfo in
- let restype = Option.map (translate_struct_entry senv.env) restype in
+ let mp = senv.modinfo.modpath in
+ let restype = Option.map (translate_module_type senv.env mp) restype in
let params,is_functor =
match modinfo.variant with
| NONE | LIBRARY _ | SIG _ -> error_no_module_to_end ()
@@ -386,83 +355,107 @@ let end_module l restype senv =
params
in
let auto_tb =
- SEBstruct (modinfo.msid, List.rev senv.revstruct)
+ SEBstruct (List.rev senv.revstruct)
in
- let mod_typ,subst,cst =
+ let mexpr,mod_typ,mod_typ_alg,resolver,cst =
match restype with
- | None -> None,modinfo.alias_subst,Constraint.empty
- | Some (res_tb,subst) ->
- let cst = check_subtypes senv.env
- {typ_expr = auto_tb;
- typ_strength = None;
- typ_alias = modinfo.alias_subst}
- {typ_expr = res_tb;
- typ_strength = None;
- typ_alias = subst} in
- let mtb = functorize_struct res_tb in
- Some mtb,subst,cst
+ | None -> let mexpr = functorize_struct auto_tb in
+ mexpr,mexpr,None,modinfo.resolver,Constraint.empty
+ | Some mtb ->
+ let auto_mtb = {
+ typ_mp = senv.modinfo.modpath;
+ typ_expr = auto_tb;
+ typ_expr_alg = None;
+ typ_constraints = Constraint.empty;
+ typ_delta = empty_delta_resolver} in
+ let cst = check_subtypes senv.env auto_mtb
+ mtb in
+ let mod_typ = functorize_struct mtb.typ_expr in
+ let mexpr = functorize_struct auto_tb in
+ let typ_alg =
+ Option.map functorize_struct mtb.typ_expr_alg in
+ mexpr,mod_typ,typ_alg,mtb.typ_delta,cst
in
- let mexpr = functorize_struct auto_tb in
let cst = Constraint.union cst senv.univ in
let mb =
- { mod_expr = Some mexpr;
+ { mod_mp = mp;
+ mod_expr = Some mexpr;
mod_type = mod_typ;
+ mod_type_alg = mod_typ_alg;
mod_constraints = cst;
- mod_alias = subst;
+ mod_delta = resolver;
mod_retroknowledge = senv.local_retroknowledge }
in
- let mp = MPdot (oldsenv.modinfo.modpath, l) in
let newenv = oldsenv.env in
let newenv = set_engagement_opt senv.engagement newenv in
let senv'= {senv with env=newenv} in
let senv' =
List.fold_left
- (fun env (mp,mb) -> full_add_module mp mb env)
+ (fun env (_,mb) -> full_add_module mb env)
senv'
(List.rev senv'.loads)
in
let newenv = Environ.add_constraints cst senv'.env in
let newenv =
- Modops.add_module mp mb newenv
- in
- let is_functor,subst = Modops.update_subst newenv mb mp in
- let newmodinfo =
- if is_functor then
- oldsenv.modinfo
- else
- { oldsenv.modinfo with
- alias_subst = join
- oldsenv.modinfo.alias_subst
- subst };
+ Modops.add_module mb newenv in
+ let modinfo = match mb.mod_type with
+ SEBstruct _ ->
+ { oldsenv.modinfo with
+ resolver =
+ add_delta_resolver resolver oldsenv.modinfo.resolver}
+ | _ -> oldsenv.modinfo
in
- mp, { old = oldsenv.old;
- env = newenv;
- modinfo = newmodinfo;
- labset = Labset.add l oldsenv.labset;
- revstruct = (l,SFBmodule mb)::oldsenv.revstruct;
- univ = Univ.Constraint.union senv'.univ oldsenv.univ;
- (* engagement is propagated to the upper level *)
- engagement = senv'.engagement;
- imports = senv'.imports;
- loads = senv'.loads@oldsenv.loads;
- local_retroknowledge = senv'.local_retroknowledge@oldsenv.local_retroknowledge }
+ mp,resolver,{ old = oldsenv.old;
+ env = newenv;
+ modinfo = modinfo;
+ labset = Labset.add l oldsenv.labset;
+ revstruct = (l,SFBmodule mb)::oldsenv.revstruct;
+ univ = Univ.Constraint.union senv'.univ oldsenv.univ;
+ (* engagement is propagated to the upper level *)
+ engagement = senv'.engagement;
+ imports = senv'.imports;
+ loads = senv'.loads@oldsenv.loads;
+ local_retroknowledge =
+ senv'.local_retroknowledge@oldsenv.local_retroknowledge }
(* Include for module and module type*)
- let add_include me senv =
- let struct_expr,_ = translate_struct_entry senv.env me in
- let senv = add_constraints (struct_expr_constraints struct_expr) senv in
- let msid,str = match (eval_struct senv.env struct_expr) with
- | SEBstruct(msid,str_l) -> msid,str_l
- | _ -> error ("You cannot Include a higher-order Module or Module Type.")
+ let add_include me is_module senv =
+ let sign,cst,resolver =
+ if is_module then
+ let sign,resolver,cst =
+ translate_struct_include_module_entry senv.env
+ senv.modinfo.modpath me in
+ sign,cst,resolver
+ else
+ let mtb =
+ translate_module_type senv.env
+ senv.modinfo.modpath me in
+ mtb.typ_expr,mtb.typ_constraints,mtb.typ_delta
in
+ let senv = add_constraints cst senv in
let mp_sup = senv.modinfo.modpath in
- let str1 = subst_signature_msid msid mp_sup str in
+ let str = match sign with
+ | SEBstruct(str_l) ->
+ str_l
+ | _ -> error ("You cannot Include a high-order structure.")
+ in
+ let senv =
+ {senv
+ with modinfo =
+ {senv.modinfo
+ with resolver =
+ add_delta_resolver resolver senv.modinfo.resolver}}
+ in
let add senv (l,elem) =
check_label l senv.labset;
match elem with
| SFBconst cb ->
- let con = make_con mp_sup empty_dirpath l in
+ let kn = make_kn mp_sup empty_dirpath l in
+ let con = constant_of_kn_equiv kn
+ (canonical_con
+ (constant_of_delta resolver (constant_of_kn kn)))
+ in
let senv' = add_constraints cb.const_constraints senv in
let env'' = Environ.add_constant con cb senv'.env in
{ old = senv'.old;
@@ -475,11 +468,14 @@ let end_module l restype senv =
imports = senv'.imports;
loads = senv'.loads ;
local_retroknowledge = senv'.local_retroknowledge }
-
| SFBmind mib ->
let kn = make_kn mp_sup empty_dirpath l in
+ let mind = mind_of_kn_equiv kn
+ (canonical_mind
+ (mind_of_delta resolver (mind_of_kn kn)))
+ in
let senv' = add_constraints mib.mind_constraints senv in
- let env'' = Environ.add_mind kn mib senv'.env in
+ let env'' = Environ.add_mind mind mib senv'.env in
{ old = senv'.old;
env = env'';
modinfo = senv'.modinfo;
@@ -492,17 +488,10 @@ let end_module l restype senv =
local_retroknowledge = senv'.local_retroknowledge }
| SFBmodule mb ->
- let mp = MPdot(senv.modinfo.modpath, l) in
- let is_functor,sub = Modops.update_subst senv.env mb mp in
- let senv' = full_add_module mp mb senv in
+ let senv' = full_add_module mb senv in
{ old = senv'.old;
env = senv'.env;
- modinfo =
- if is_functor then
- senv'.modinfo
- else
- {senv'.modinfo with
- alias_subst = join senv'.modinfo.alias_subst sub};
+ modinfo = senv'.modinfo;
labset = Labset.add l senv'.labset;
revstruct = (l,SFBmodule mb)::senv'.revstruct;
univ = senv'.univ;
@@ -510,87 +499,69 @@ let end_module l restype senv =
imports = senv'.imports;
loads = senv'.loads;
local_retroknowledge = senv'.local_retroknowledge }
- | SFBalias (mp',typ_opt,cst) ->
- let env' = Option.fold_right
- Environ.add_constraints cst senv.env in
- let mp = MPdot(senv.modinfo.modpath, l) in
- let mp1 = scrape_alias mp' senv.env in
- let _,sub = Modops.update_subst senv.env (lookup_module mp1 senv.env) mp1 in
- let sub = update_subst sub (map_mp mp mp1) in
- let sub = join_alias sub (map_mp mp mp1) in
- let sub = add_mp mp mp1 sub in
- let env' = register_alias mp mp' env' in
- { old = senv.old;
- env = env';
- modinfo = { senv.modinfo with
- alias_subst = join
- senv.modinfo.alias_subst sub};
- labset = Labset.add l senv.labset;
- revstruct = (l,SFBalias (mp',typ_opt,cst))::senv.revstruct;
- univ = senv.univ;
- engagement = senv.engagement;
- imports = senv.imports;
- loads = senv.loads;
- local_retroknowledge = senv.local_retroknowledge }
| SFBmodtype mtb ->
- let env' = add_modtype_constraints senv.env mtb in
+ let senv' = add_constraints mtb.typ_constraints senv in
let mp = MPdot(senv.modinfo.modpath, l) in
- let env'' = Environ.add_modtype mp mtb env' in
+ let env' = Environ.add_modtype mp mtb senv'.env in
{ old = senv.old;
- env = env'';
- modinfo = senv.modinfo;
+ env = env';
+ modinfo = senv'.modinfo;
labset = Labset.add l senv.labset;
- revstruct = (l,SFBmodtype mtb)::senv.revstruct;
- univ = senv.univ;
- engagement = senv.engagement;
- imports = senv.imports;
- loads = senv.loads;
- local_retroknowledge = senv.local_retroknowledge }
+ revstruct = (l,SFBmodtype mtb)::senv'.revstruct;
+ univ = senv'.univ;
+ engagement = senv'.engagement;
+ imports = senv'.imports;
+ loads = senv'.loads;
+ local_retroknowledge = senv'.local_retroknowledge }
in
- List.fold_left add senv str1
+ resolver,(List.fold_left add senv str)
(* Adding parameters to modules or module types *)
let add_module_parameter mbid mte senv =
if senv.revstruct <> [] or senv.loads <> [] then
anomaly "Cannot add a module parameter to a non empty module";
- let mtb_expr,sub = translate_struct_entry senv.env mte in
- let mtb = {typ_expr = mtb_expr;
- typ_strength = None;
- typ_alias = sub} in
- let senv = full_add_module (MPbound mbid) (module_body_of_type mtb) senv
+ let mtb = translate_module_type senv.env (MPbound mbid) mte in
+ let senv =
+ full_add_module (module_body_of_type (MPbound mbid) mtb) senv
in
let new_variant = match senv.modinfo.variant with
| STRUCT params -> STRUCT ((mbid,mtb) :: params)
| SIG params -> SIG ((mbid,mtb) :: params)
| _ ->
- anomaly "Module parameters can only be added to modules or signatures"
+ anomaly "Module parameters can only be added to modules or signatures"
in
- { old = senv.old;
- env = senv.env;
- modinfo = { senv.modinfo with variant = new_variant };
- labset = senv.labset;
- revstruct = [];
- univ = senv.univ;
- engagement = senv.engagement;
- imports = senv.imports;
- loads = [];
- local_retroknowledge = senv.local_retroknowledge }
-
+
+ let resolver_of_param = match mtb.typ_expr with
+ SEBstruct _ -> mtb.typ_delta
+ | _ -> empty_delta_resolver
+ in
+ mtb.typ_delta, { old = senv.old;
+ env = senv.env;
+ modinfo = { senv.modinfo with
+ variant = new_variant;
+ resolver_of_param = add_delta_resolver
+ resolver_of_param senv.modinfo.resolver_of_param};
+ labset = senv.labset;
+ revstruct = [];
+ univ = senv.univ;
+ engagement = senv.engagement;
+ imports = senv.imports;
+ loads = [];
+ local_retroknowledge = senv.local_retroknowledge }
+
(* Interactive module types *)
let start_modtype l senv =
check_label l senv.labset;
- let msid = make_msid senv.modinfo.seed (string_of_label l) in
- let mp = MPself msid in
- let modinfo = { msid = msid;
- modpath = mp;
- seed = senv.modinfo.seed;
- label = l;
- variant = SIG [];
- alias_subst = empty_subst }
- in
+ let mp = MPdot(senv.modinfo.modpath, l) in
+ let modinfo = { modpath = mp;
+ label = l;
+ variant = SIG [];
+ resolver = empty_delta_resolver;
+ resolver_of_param = empty_delta_resolver}
+ in
mp, { old = senv;
env = senv.env;
modinfo = modinfo;
@@ -614,7 +585,7 @@ let end_modtype l senv =
if l <> modinfo.label then error_incompatible_labels l modinfo.label;
if not (empty_context senv.env) then error_local_context None;
let auto_tb =
- SEBstruct (modinfo.msid, List.rev senv.revstruct)
+ SEBstruct (List.rev senv.revstruct)
in
let mtb_expr =
List.fold_left
@@ -625,42 +596,39 @@ let end_modtype l senv =
in
let mp = MPdot (oldsenv.modinfo.modpath, l) in
let newenv = oldsenv.env in
- (* since universes constraints cannot be stored in the modtype,
- they are propagated to the upper level *)
let newenv = Environ.add_constraints senv.univ newenv in
let newenv = set_engagement_opt senv.engagement newenv in
let senv = {senv with env=newenv} in
let senv =
List.fold_left
- (fun env (mp,mb) -> full_add_module mp mb env)
+ (fun env (mp,mb) -> full_add_module mb env)
senv
(List.rev senv.loads)
in
- let subst = senv.modinfo.alias_subst in
- let mtb = {typ_expr = mtb_expr;
- typ_strength = None;
- typ_alias = subst} in
+ let mtb = {typ_mp = mp;
+ typ_expr = mtb_expr;
+ typ_expr_alg = None;
+ typ_constraints = senv.univ;
+ typ_delta = senv.modinfo.resolver} in
let newenv =
Environ.add_modtype mp mtb senv.env
in
mp, { old = oldsenv.old;
env = newenv;
- modinfo = oldsenv.modinfo;
- labset = Labset.add l oldsenv.labset;
- revstruct = (l,SFBmodtype mtb)::oldsenv.revstruct;
- univ = Univ.Constraint.union senv.univ oldsenv.univ;
- engagement = senv.engagement;
- imports = senv.imports;
- loads = senv.loads@oldsenv.loads;
- (* spiwack : if there is a bug with retroknowledge in nested modules
- it's likely to come from here *)
- local_retroknowledge =
- senv.local_retroknowledge@oldsenv.local_retroknowledge}
-
+ modinfo = oldsenv.modinfo;
+ labset = Labset.add l oldsenv.labset;
+ revstruct = (l,SFBmodtype mtb)::oldsenv.revstruct;
+ univ = Univ.Constraint.union senv.univ oldsenv.univ;
+ engagement = senv.engagement;
+ imports = senv.imports;
+ loads = senv.loads@oldsenv.loads;
+ (* spiwack : if there is a bug with retroknowledge in nested modules
+ it's likely to come from here *)
+ local_retroknowledge =
+ senv.local_retroknowledge@oldsenv.local_retroknowledge}
let current_modpath senv = senv.modinfo.modpath
-let current_msid senv = senv.modinfo.msid
-
+let delta_of_senv senv = senv.modinfo.resolver,senv.modinfo.resolver_of_param
(* Check that the engagement expected by a library matches the initial one *)
let check_engagement env c =
@@ -685,7 +653,7 @@ type compiled_library =
let is_empty senv =
senv.revstruct = [] &&
- senv.modinfo.msid = initial_msid &&
+ senv.modinfo.modpath = initial_path &&
senv.modinfo.variant = NONE
let start_library dir senv =
@@ -697,14 +665,12 @@ let start_library dir senv =
| hd::tl ->
make_dirpath tl, label_of_id hd
in
- let msid = make_msid dir_path (string_of_label l) in
- let mp = MPself msid in
- let modinfo = { msid = msid;
- modpath = mp;
- seed = dir;
- label = l;
- variant = LIBRARY dir;
- alias_subst = empty_subst }
+ let mp = MPfile dir in
+ let modinfo = {modpath = mp;
+ label = l;
+ variant = LIBRARY dir;
+ resolver = empty_delta_resolver;
+ resolver_of_param = empty_delta_resolver}
in
mp, { old = senv;
env = senv.env;
@@ -731,14 +697,18 @@ let export senv dir =
end;
(*if senv.modinfo.params <> [] || senv.modinfo.restype <> None then
(* error_export_simple *) (); *)
- let mb =
- { mod_expr = Some (SEBstruct (modinfo.msid, List.rev senv.revstruct));
- mod_type = None;
+ let str = SEBstruct (List.rev senv.revstruct) in
+ let mp = senv.modinfo.modpath in
+ let mb =
+ { mod_mp = mp;
+ mod_expr = Some str;
+ mod_type = str;
+ mod_type_alg = None;
mod_constraints = senv.univ;
- mod_alias = senv.modinfo.alias_subst;
+ mod_delta = senv.modinfo.resolver;
mod_retroknowledge = senv.local_retroknowledge}
in
- modinfo.msid, (dir,mb,senv.imports,engagement senv.env)
+ mp, (dir,mb,senv.imports,engagement senv.env)
let check_imports senv needed =
@@ -775,9 +745,13 @@ let import (dp,mb,depends,engmt) digest senv =
let mp = MPfile dp in
let env = senv.env in
let env = Environ.add_constraints mb.mod_constraints env in
- let env = Modops.add_module mp mb env in
+ let env = Modops.add_module mb env in
mp, { senv with
env = env;
+ modinfo =
+ {senv.modinfo with
+ resolver =
+ add_delta_resolver mb.mod_delta senv.modinfo.resolver};
imports = (dp,digest)::senv.imports;
loads = (mp,mb)::senv.loads }
@@ -785,14 +759,14 @@ let import (dp,mb,depends,engmt) digest senv =
(* Remove the body of opaque constants in modules *)
let rec lighten_module mb =
{ mb with
- mod_expr = Option.map lighten_modexpr mb.mod_expr;
- mod_type = Option.map lighten_modexpr mb.mod_type;
+ mod_expr = None;
+ mod_type = lighten_modexpr mb.mod_type;
}
and lighten_struct struc =
let lighten_body (l,body) = (l,match body with
| SFBconst ({const_opaque=true} as x) -> SFBconst {x with const_body=None}
- | (SFBconst _ | SFBmind _ | SFBalias _) as x -> x
+ | (SFBconst _ | SFBmind _ ) as x -> x
| SFBmodule m -> SFBmodule (lighten_module m)
| SFBmodtype m -> SFBmodtype
({m with
@@ -807,8 +781,8 @@ and lighten_modexpr = function
typ_expr = lighten_modexpr mty.typ_expr}),
lighten_modexpr mexpr)
| SEBident mp as x -> x
- | SEBstruct (msid, struc) ->
- SEBstruct (msid, lighten_struct struc)
+ | SEBstruct (struc) ->
+ SEBstruct (lighten_struct struc)
| SEBapply (mexpr,marg,u) ->
SEBapply (lighten_modexpr mexpr,lighten_modexpr marg,u)
| SEBwith (seb,wdcl) ->
diff --git a/kernel/safe_typing.mli b/kernel/safe_typing.mli
index ac1e3863a..f011d42b7 100644
--- a/kernel/safe_typing.mli
+++ b/kernel/safe_typing.mli
@@ -13,6 +13,7 @@ open Names
open Term
open Declarations
open Entries
+open Mod_subst
(*i*)
(*s Safe environments. Since we are now able to type terms, we can
@@ -55,12 +56,8 @@ val add_mind :
(* Adding a module *)
val add_module :
label -> module_entry -> safe_environment
- -> module_path * safe_environment
+ -> module_path * delta_resolver * safe_environment
-(* Adding a module alias*)
-val add_alias :
- label -> module_path -> safe_environment
- -> module_path * safe_environment
(* Adding a module type *)
val add_modtype :
label -> module_struct_entry -> safe_environment
@@ -77,12 +74,13 @@ val set_engagement : engagement -> safe_environment -> safe_environment
(*s Interactive module functions *)
val start_module :
label -> safe_environment -> module_path * safe_environment
+
val end_module :
- label -> module_struct_entry option
- -> safe_environment -> module_path * safe_environment
+ label -> module_struct_entry option
+ -> safe_environment -> module_path * delta_resolver * safe_environment
val add_module_parameter :
- mod_bound_id -> module_struct_entry -> safe_environment -> safe_environment
+ mod_bound_id -> module_struct_entry -> safe_environment -> delta_resolver * safe_environment
val start_modtype :
label -> safe_environment -> module_path * safe_environment
@@ -91,12 +89,11 @@ val end_modtype :
label -> safe_environment -> module_path * safe_environment
val add_include :
- module_struct_entry -> safe_environment -> safe_environment
+ module_struct_entry -> bool -> safe_environment -> delta_resolver * safe_environment
val current_modpath : safe_environment -> module_path
-val current_msid : safe_environment -> mod_self_id
-
-
+val delta_of_senv : safe_environment -> delta_resolver*delta_resolver
+
(* Loading and saving compilation units *)
(* exporting and importing modules *)
@@ -106,7 +103,7 @@ val start_library : dir_path -> safe_environment
-> module_path * safe_environment
val export : safe_environment -> dir_path
- -> mod_self_id * compiled_library
+ -> module_path * compiled_library
val import : compiled_library -> Digest.t -> safe_environment
-> module_path * safe_environment
diff --git a/kernel/subtyping.ml b/kernel/subtyping.ml
index 861dc9a3f..41f4c640c 100644
--- a/kernel/subtyping.ml
+++ b/kernel/subtyping.ml
@@ -33,7 +33,6 @@ type namedobject =
| IndConstr of constructor * mutual_inductive_body
| Module of module_body
| Modtype of module_type_body
- | Alias of module_path * struct_expr_body option
(* adds above information about one mutual inductive: all types and
constructors *)
@@ -61,10 +60,9 @@ let make_label_map mp list =
match e with
| SFBconst cb -> add_map (Constant cb)
| SFBmind mib ->
- add_nameobjects_of_mib (make_kn mp empty_dirpath l) mib map
+ add_nameobjects_of_mib (make_mind mp empty_dirpath l) mib map
| SFBmodule mb -> add_map (Module mb)
| SFBmodtype mtb -> add_map (Modtype mtb)
- | SFBalias (mp,typ_opt,cst) -> add_map (Alias (mp,typ_opt))
in
List.fold_right add_one list Labmap.empty
@@ -75,15 +73,18 @@ let check_conv_error error cst f env a1 a2 =
NotConvertible -> error ()
(* for now we do not allow reorderings *)
-let check_inductive cst env msid1 l info1 mib2 spec2 =
- let kn = make_kn (MPself msid1) empty_dirpath l in
+
+let check_inductive cst env mp1 l info1 mp2 mib2 spec2 subst1 subst2 reso1 reso2=
+ let kn1 = make_mind mp1 empty_dirpath l in
+ let kn2 = make_mind mp2 empty_dirpath l in
let error () = error_not_match l spec2 in
let check_conv cst f = check_conv_error error cst f in
let mib1 =
match info1 with
- | IndType ((_,0), mib) -> mib
+ | IndType ((_,0), mib) -> subst_mind subst1 mib
| _ -> error ()
in
+ let mib2 = subst_mind subst2 mib2 in
let check_inductive_type cst env t1 t2 =
(* Due to sort-polymorphism in inductive types, the conclusions of
@@ -141,8 +142,8 @@ let check_inductive cst env msid1 l info1 mib2 spec2 =
array_fold_left2
(fun cst t1 t2 -> check_conv cst conv env t1 t2)
cst
- (arities_of_specif kn (mib1,p1))
- (arities_of_specif kn (mib2,p2))
+ (arities_of_specif kn1 (mib1,p1))
+ (arities_of_specif kn1 (mib2,p2))
in
let check f = if f mib1 <> f mib2 then error () in
check (fun mib -> mib.mind_finite);
@@ -158,16 +159,12 @@ let check_inductive cst env msid1 l info1 mib2 spec2 =
(* the inductive types and constructors types have to be convertible *)
check (fun mib -> mib.mind_nparams);
- begin
- match mib2.mind_equiv with
- | None -> ()
- | Some kn2' ->
- let kn2 = scrape_mind env kn2' in
- let kn1 = match mib1.mind_equiv with
- None -> kn
- | Some kn1' -> scrape_mind env kn1'
- in
- if kn1 <> kn2 then error ()
+ begin
+ match mind_of_delta reso2 kn2 with
+ | kn2' when kn2=kn2' -> ()
+ | kn2' ->
+ if not (eq_mind (mind_of_delta reso1 kn1) kn2') then
+ error ()
end;
(* we check that records and their field names are preserved. *)
check (fun mib -> mib.mind_record);
@@ -194,7 +191,8 @@ let check_inductive cst env msid1 l info1 mib2 spec2 =
in
cst
-let check_constant cst env msid1 l info1 cb2 spec2 =
+
+let check_constant cst env mp1 l info1 cb2 spec2 subst1 subst2 =
let error () = error_not_match l spec2 in
let check_conv cst f = check_conv_error error cst f in
let check_type cst env t1 t2 =
@@ -245,13 +243,15 @@ let check_constant cst env msid1 l info1 cb2 spec2 =
in
match info1 with
- | Constant cb1 ->
- assert (cb1.const_hyps=[] && cb2.const_hyps=[]) ;
- (*Start by checking types*)
+ | Constant cb1 ->
+ assert (cb1.const_hyps=[] && cb2.const_hyps=[]) ;
+ (*Start by checking types*)
+ let cb1 = subst_const_body subst1 cb1 in
+ let cb2 = subst_const_body subst2 cb2 in
let typ1 = Typeops.type_of_constant_type env cb1.const_type in
let typ2 = Typeops.type_of_constant_type env cb2.const_type in
let cst = check_type cst env typ1 typ2 in
- let con = make_con (MPself msid1) empty_dirpath l in
+ let con = make_con mp1 empty_dirpath l in
let cst =
if cb2.const_opaque then
match cb2.const_body with
@@ -264,7 +264,8 @@ let check_constant cst env msid1 l info1 cb2 spec2 =
begin
match (kind_of_term c) with
Const n ->
- let cb = lookup_constant n env in
+ let cb = subst_const_body subst1
+ (lookup_constant n env) in
(match cb.const_opaque,
cb.const_body with
| true, Some lc1 ->
@@ -311,28 +312,13 @@ let check_constant cst env msid1 l info1 cb2 spec2 =
check_conv cst conv env ty1 ty2
| _ -> error ()
-let rec check_modules cst env msid1 l msb1 msb2 alias =
- let mp = (MPdot(MPself msid1,l)) in
- let mty1 = module_type_of_module (Some mp) msb1 in
- let alias1,struct_expr = match eval_struct env mty1.typ_expr with
- | SEBstruct (msid,sign) as str ->
- update_subst alias (map_msid msid mp),str
- | _ as str -> empty_subst,str in
- let mty1 = {mty1 with
- typ_expr = struct_expr;
- typ_alias = join alias1 mty1.typ_alias } in
- let mty2 = module_type_of_module None msb2 in
- let cst = check_modtypes cst env mty1 mty2 false in
+let rec check_modules cst env msb1 msb2 subst1 subst2 =
+ let mty1 = module_type_of_module env None msb1 in
+ let mty2 = module_type_of_module env None msb2 in
+ let cst = check_modtypes cst env mty1 mty2 subst1 subst2 false in
cst
-
-and check_signatures cst env (msid1,sig1) alias (msid2,sig2') =
- let mp1 = MPself msid1 in
- let env = add_signature mp1 sig1 env in
- let sig1 = subst_structure alias sig1 in
- let alias1 = update_subst alias (map_msid msid2 mp1) in
- let sig2 = subst_structure alias1 sig2' in
- let sig2 = subst_signature_msid msid2 mp1 sig2 in
+and check_signatures cst env mp1 sig1 mp2 sig2 subst1 subst2 reso1 reso2=
let map1 = make_label_map mp1 sig1 in
let check_one_body cst (l,spec2) =
let info1 =
@@ -340,38 +326,19 @@ and check_signatures cst env (msid1,sig1) alias (msid2,sig2') =
Labmap.find l map1
with
Not_found -> error_no_such_label_sub l
- (string_of_msid msid1) (string_of_msid msid2)
+ (string_of_mp mp1)
in
match spec2 with
| SFBconst cb2 ->
- check_constant cst env msid1 l info1 cb2 spec2
+ check_constant cst env mp1 l info1 cb2 spec2 subst1 subst2
| SFBmind mib2 ->
- check_inductive cst env msid1 l info1 mib2 spec2
+ check_inductive cst env
+ mp1 l info1 mp2 mib2 spec2 subst1 subst2 reso1 reso2
| SFBmodule msb2 ->
begin
match info1 with
- | Module msb -> check_modules cst env msid1 l msb msb2 alias
- | Alias (mp,typ_opt) ->let msb =
- {mod_expr = Some (SEBident mp);
- mod_type = typ_opt;
- mod_constraints = Constraint.empty;
- mod_alias = (lookup_modtype mp env).typ_alias;
- mod_retroknowledge = []} in
- check_modules cst env msid1 l msb msb2 alias
- | _ -> error_not_match l spec2
- end
- | SFBalias (mp,typ_opt,_) ->
- begin
- match info1 with
- | Alias (mp1,_) -> check_modpath_equiv env mp mp1; cst
- | Module msb ->
- let msb1 =
- {mod_expr = Some (SEBident mp);
- mod_type = typ_opt;
- mod_constraints = Constraint.empty;
- mod_alias = (lookup_modtype mp env).typ_alias;
- mod_retroknowledge = []} in
- check_modules cst env msid1 l msb msb1 alias
+ | Module msb -> check_modules cst env msb msb2
+ subst1 subst2
| _ -> error_not_match l spec2
end
| SFBmodtype mtb2 ->
@@ -380,50 +347,69 @@ and check_signatures cst env (msid1,sig1) alias (msid2,sig2') =
| Modtype mtb -> mtb
| _ -> error_not_match l spec2
in
- check_modtypes cst env mtb1 mtb2 true
+ let env = add_module (module_body_of_type mtb2.typ_mp mtb2)
+ (add_module (module_body_of_type mtb1.typ_mp mtb1) env) in
+ check_modtypes cst env mtb1 mtb2 subst1 subst2 true
in
List.fold_left check_one_body cst sig2
-
-and check_modtypes cst env mtb1 mtb2 equiv =
- if mtb1==mtb2 then cst else (* just in case :) *)
- let mtb1',mtb2'=
- (match mtb1.typ_strength with
- None -> eval_struct env mtb1.typ_expr,
- eval_struct env mtb2.typ_expr
- | Some mp -> strengthen env mtb1.typ_expr mp,
- eval_struct env mtb2.typ_expr) in
- let rec check_structure cst env str1 str2 equiv =
- match str1, str2 with
- | SEBstruct (msid1,list1),
- SEBstruct (msid2,list2) ->
- let cst = check_signatures cst env
- (msid1,list1) mtb1.typ_alias (msid2,list2) in
- if equiv then
- check_signatures cst env
- (msid2,list2) mtb2.typ_alias (msid1,list1)
- else
- cst
+and check_modtypes cst env mtb1 mtb2 subst1 subst2 equiv =
+ if mtb1==mtb2 then cst else
+ let mtb1',mtb2'=mtb1.typ_expr,mtb2.typ_expr in
+ let rec check_structure cst env str1 str2 equiv subst1 subst2 =
+ match str1,str2 with
+ | SEBstruct (list1),
+ SEBstruct (list2) ->
+ if equiv then
+ let subst2 =
+ add_mp mtb2.typ_mp mtb1.typ_mp mtb1.typ_delta subst2 in
+ Univ.Constraint.union
+ (check_signatures cst env
+ mtb1.typ_mp list1 mtb2.typ_mp list2 subst1 subst2
+ mtb1.typ_delta mtb2.typ_delta)
+ (check_signatures cst env
+ mtb2.typ_mp list2 mtb1.typ_mp list1 subst2 subst1
+ mtb2.typ_delta mtb1.typ_delta)
+ else
+ check_signatures cst env
+ mtb1.typ_mp list1 mtb2.typ_mp list2 subst1 subst2
+ mtb1.typ_delta mtb2.typ_delta
| SEBfunctor (arg_id1,arg_t1,body_t1),
SEBfunctor (arg_id2,arg_t2,body_t2) ->
- let cst = check_modtypes cst env arg_t2 arg_t1 equiv in
+ let subst1 =
+ (join (map_mbid arg_id1 (MPbound arg_id2) arg_t2.typ_delta) subst1) in
+ let cst = check_modtypes cst env
+ arg_t2 arg_t1 subst2 subst1
+ equiv in
(* contravariant *)
- let env =
- add_module (MPbound arg_id2) (module_body_of_type arg_t2) env
+ let env = add_module
+ (module_body_of_type (MPbound arg_id2) arg_t2) env
in
- let body_t1' =
- (* since we are just checking well-typedness we do not need
- to expand any constant. Hence the identity resolver. *)
- subst_struct_expr
- (map_mbid arg_id1 (MPbound arg_id2) None)
- body_t1
+ let subst1 =
+ (join (map_mbid arg_id1 (MPbound arg_id2) arg_t2.typ_delta) subst1) in
+ let env = match body_t1 with
+ SEBstruct str ->
+ add_module {mod_mp = mtb1.typ_mp;
+ mod_expr = None;
+ mod_type = subst_struct_expr subst1 body_t1;
+ mod_type_alg= None;
+ mod_constraints=mtb1.typ_constraints;
+ mod_retroknowledge = [];
+ mod_delta = mtb1.typ_delta} env
+ | _ -> env
in
- check_structure cst env (eval_struct env body_t1')
- (eval_struct env body_t2) equiv
+ check_structure cst env body_t1 body_t2 equiv
+ subst1
+ subst2
| _ , _ -> error_incompatible_modtypes mtb1 mtb2
in
if mtb1'== mtb2' then cst
- else check_structure cst env mtb1' mtb2' equiv
+ else check_structure cst env mtb1' mtb2' equiv subst1 subst2
let check_subtypes env sup super =
- check_modtypes Constraint.empty env sup super false
+ let env = add_module
+ (module_body_of_type sup.typ_mp sup) env in
+ check_modtypes Constraint.empty env
+ (strengthen env sup sup.typ_mp) super empty_subst
+ (map_mp super.typ_mp sup.typ_mp sup.typ_delta) false
+
diff --git a/kernel/term.ml b/kernel/term.ml
index 68ea2ed3f..685656592 100644
--- a/kernel/term.ml
+++ b/kernel/term.ml
@@ -184,7 +184,7 @@ module Hconstr =
type t = constr
type u = (constr -> constr) *
((sorts -> sorts) * (constant -> constant) *
- (kernel_name -> kernel_name) * (name -> name) *
+ (mutual_inductive -> mutual_inductive) * (name -> name) *
(identifier -> identifier))
let hash_sub = hash_term
let equal = comp_term
@@ -609,6 +609,7 @@ let map_constr_with_binders g f l c = match kind_of_term c with
application associativity, binders name and Cases annotations are
not taken into account *)
+
let compare_constr f t1 t2 =
match kind_of_term t1, kind_of_term t2 with
| Rel n1, Rel n2 -> n1 = n2
@@ -630,9 +631,9 @@ let compare_constr f t1 t2 =
f h1 h2 & List.for_all2 f l1 l2
else false
| Evar (e1,l1), Evar (e2,l2) -> e1 = e2 & array_for_all2 f l1 l2
- | Const c1, Const c2 -> c1 = c2
- | Ind c1, Ind c2 -> c1 = c2
- | Construct c1, Construct c2 -> c1 = c2
+ | Const c1, Const c2 -> eq_constant c1 c2
+ | Ind c1, Ind c2 -> eq_ind c1 c2
+ | Construct c1, Construct c2 -> eq_constructor c1 c2
| Case (_,p1,c1,bl1), Case (_,p2,c2,bl2) ->
f p1 p2 & f c1 c2 & array_for_all2 f bl1 bl2
| Fix (ln1,(_,tl1,bl1)), Fix (ln2,(_,tl2,bl2)) ->
diff --git a/kernel/term.mli b/kernel/term.mli
index 5929250db..0de831668 100644
--- a/kernel/term.mli
+++ b/kernel/term.mli
@@ -595,7 +595,7 @@ val compare_constr : (constr -> constr -> bool) -> constr -> constr -> bool
val hcons_constr:
(constant -> constant) *
- (kernel_name -> kernel_name) *
+ (mutual_inductive -> mutual_inductive) *
(dir_path -> dir_path) *
(name -> name) *
(identifier -> identifier) *
diff --git a/kernel/vconv.ml b/kernel/vconv.ml
index 0dd119f7b..a35d1d88e 100644
--- a/kernel/vconv.ml
+++ b/kernel/vconv.ml
@@ -83,7 +83,7 @@ and conv_whd pb k whd1 whd2 cu =
and conv_atom pb k a1 stk1 a2 stk2 cu =
match a1, a2 with
| Aind (kn1,i1), Aind(kn2,i2) ->
- if mind_equiv_infos !infos (kn1,i1) (kn2,i2) && compare_stack stk1 stk2
+ if eq_ind (kn1,i1) (kn2,i2) && compare_stack stk1 stk2
then
conv_stack k stk1 stk2 cu
else raise NotConvertible
@@ -94,7 +94,7 @@ and conv_atom pb k a1 stk1 a2 stk2 cu =
| Aiddef(ik1,v1), Aiddef(ik2,v2) ->
begin
try
- if ik1 = ik2 && compare_stack stk1 stk2 then
+ if eq_table_key ik1 ik2 && compare_stack stk1 stk2 then
conv_stack k stk1 stk2 cu
else raise NotConvertible
with NotConvertible ->
@@ -191,11 +191,11 @@ let rec conv_eq pb t1 t2 cu =
if e1 = e2 then conv_eq_vect l1 l2 cu
else raise NotConvertible
| Const c1, Const c2 ->
- if c1 = c2 then cu else raise NotConvertible
+ if eq_constant c1 c2 then cu else raise NotConvertible
| Ind c1, Ind c2 ->
- if c1 = c2 then cu else raise NotConvertible
+ if eq_ind c1 c2 then cu else raise NotConvertible
| Construct c1, Construct c2 ->
- if c1 = c2 then cu else raise NotConvertible
+ if eq_constructor c1 c2 then cu else raise NotConvertible
| Case (_,p1,c1,bl1), Case (_,p2,c2,bl2) ->
let pcu = conv_eq CONV p1 p2 cu in
let ccu = conv_eq CONV c1 c2 pcu in
diff --git a/lib/lib.mllib b/lib/lib.mllib
index 2321abd1b..56fe32cbd 100644
--- a/lib/lib.mllib
+++ b/lib/lib.mllib
@@ -13,6 +13,7 @@ Edit
Gset
Gmap
Tlm
+tries
Gmapl
Profile
Explore
diff --git a/lib/tries.ml b/lib/tries.ml
new file mode 100644
index 000000000..1159ea233
--- /dev/null
+++ b/lib/tries.ml
@@ -0,0 +1,78 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+
+
+module Make =
+ functor (X : Set.OrderedType) ->
+ functor (Y : Map.OrderedType) ->
+struct
+ module T_dom = Set.Make(X)
+ module T_codom = Map.Make(Y)
+
+ type t = Node of T_dom.t * t T_codom.t
+
+ let codom_to_list m = T_codom.fold (fun x y l -> (x,y)::l) m []
+
+ let codom_rng m = T_codom.fold (fun _ y acc -> y::acc) m []
+
+ let codom_dom m = T_codom.fold (fun x _ acc -> x::acc) m []
+
+ let empty = Node (T_dom.empty, T_codom.empty)
+
+ let map (Node (_,m)) lbl = T_codom.find lbl m
+
+ let xtract (Node (hereset,_)) = T_dom.elements hereset
+
+ let dom (Node (_,m)) = codom_dom m
+
+ let in_dom (Node (_,m)) lbl = T_codom.mem lbl m
+
+ let is_empty_node (Node(a,b)) = (T_dom.elements a = []) & (codom_to_list b = [])
+
+let assure_arc m lbl =
+ if T_codom.mem lbl m then
+ m
+ else
+ T_codom.add lbl (Node (T_dom.empty,T_codom.empty)) m
+
+let cleanse_arcs (Node (hereset,m)) =
+ let l = codom_rng m in
+ Node(hereset, if List.for_all is_empty_node l then T_codom.empty else m)
+
+let rec at_path f (Node (hereset,m)) = function
+ | [] ->
+ cleanse_arcs (Node(f hereset,m))
+ | h::t ->
+ let m = assure_arc m h in
+ cleanse_arcs (Node(hereset,
+ T_codom.add h (at_path f (T_codom.find h m) t) m))
+
+let add tm (path,v) =
+ at_path (fun hereset -> T_dom.add v hereset) tm path
+
+let rmv tm (path,v) =
+ at_path (fun hereset -> T_dom.remove v hereset) tm path
+
+let app f tlm =
+ let rec apprec pfx (Node(hereset,m)) =
+ let path = List.rev pfx in
+ T_dom.iter (fun v -> f(path,v)) hereset;
+ T_codom.iter (fun l tm -> apprec (l::pfx) tm) m
+ in
+ apprec [] tlm
+
+let to_list tlm =
+ let rec torec pfx (Node(hereset,m)) =
+ let path = List.rev pfx in
+ List.flatten((List.map (fun v -> (path,v)) (T_dom.elements hereset))::
+ (List.map (fun (l,tm) -> torec (l::pfx) tm) (codom_to_list m)))
+ in
+ torec [] tlm
+
+end
diff --git a/lib/tries.mli b/lib/tries.mli
new file mode 100644
index 000000000..342c81ecd
--- /dev/null
+++ b/lib/tries.mli
@@ -0,0 +1,34 @@
+
+
+
+
+
+module Make :
+ functor (X : Set.OrderedType) ->
+ functor (Y : Map.OrderedType) ->
+sig
+
+ type t
+
+ val empty : t
+
+ (* Work on labels, not on paths. *)
+
+ val map : t -> Y.t -> t
+
+ val xtract : t -> X.t list
+
+ val dom : t -> Y.t list
+
+ val in_dom : t -> Y.t -> bool
+
+ (* Work on paths, not on labels. *)
+
+ val add : t -> Y.t list * X.t -> t
+
+ val rmv : t -> Y.t list * X.t -> t
+
+ val app : ((Y.t list * X.t) -> unit) -> t -> unit
+
+ val to_list : t -> (Y.t list * X.t) list
+end
diff --git a/library/declare.ml b/library/declare.ml
index da723aa4b..1084aa07d 100644
--- a/library/declare.ml
+++ b/library/declare.ml
@@ -100,12 +100,14 @@ let load_constant i ((sp,kn),(_,_,kind)) =
if Nametab.exists_cci sp then
errorlabstrm "cache_constant"
(pr_id (basename sp) ++ str " already exists");
- Nametab.push (Nametab.Until i) sp (ConstRef (constant_of_kn kn));
- add_constant_kind (constant_of_kn kn) kind
-
+ let con = Global.constant_of_delta (constant_of_kn kn) in
+ Nametab.push (Nametab.Until i) sp (ConstRef con);
+ add_constant_kind con kind
+
(* Opening means making the name without its module qualification available *)
let open_constant i ((sp,kn),_) =
- Nametab.push (Nametab.Exactly i) sp (ConstRef (constant_of_kn kn))
+ let con = constant_of_kn kn in
+ Nametab.push (Nametab.Exactly i) sp (ConstRef con)
let cache_constant ((sp,kn),(cdt,dhyps,kind)) =
let id = basename sp in
@@ -187,6 +189,7 @@ let declare_inductive_argument_scopes kn mie =
let inductive_names sp kn mie =
let (dp,_) = repr_path sp in
+ let kn = Global.mind_of_delta (mind_of_kn kn) in
let names, _ =
List.fold_left
(fun (names, n) ind ->
@@ -228,17 +231,18 @@ let cache_inductive ((sp,kn),(dhyps,mie)) =
let id = basename sp in
let _,dir,_ = repr_kn kn in
let kn' = Global.add_mind dir id mie in
- assert (kn'=kn);
- add_section_kn kn (Global.lookup_mind kn').mind_hyps;
+ assert (kn'= mind_of_kn kn);
+ add_section_kn kn' (Global.lookup_mind kn').mind_hyps;
Dischargedhypsmap.set_discharged_hyps sp dhyps;
List.iter (fun (sp, ref) -> Nametab.push (Nametab.Until 1) sp ref) names;
List.iter (fun (sp,_) -> !cache_hook sp) (inductive_names sp kn mie)
let discharge_inductive ((sp,kn),(dhyps,mie)) =
- let mie = Global.lookup_mind kn in
+ let mind = (Global.mind_of_delta (mind_of_kn kn)) in
+ let mie = Global.lookup_mind mind in
let repl = replacement_context () in
- let sechyps = section_segment_of_mutual_inductive kn in
+ let sechyps = section_segment_of_mutual_inductive mind in
Some (discharged_hyps kn sechyps,
Discharge.process_inductive (named_of_variable_context sechyps) repl mie)
@@ -271,8 +275,9 @@ let declare_mind isrecord mie =
| ind::_ -> ind.mind_entry_typename
| [] -> anomaly "cannot declare an empty list of inductives" in
let (sp,kn as oname) = add_leaf id (inInductive ([],mie)) in
- declare_mib_implicits kn;
- declare_inductive_argument_scopes kn mie;
+ let mind = (Global.mind_of_delta (mind_of_kn kn)) in
+ declare_mib_implicits mind;
+ declare_inductive_argument_scopes mind mie;
!xml_declare_inductive (isrecord,oname);
oname
diff --git a/library/declaremods.ml b/library/declaremods.ml
index d796a2906..548a9b0f3 100644
--- a/library/declaremods.ml
+++ b/library/declaremods.ml
@@ -22,7 +22,7 @@ open Mod_subst
(* modules and components *)
-(* This type is a functional closure of substitutive lib_objects.
+(* OBSOLETE This type is a functional closure of substitutive lib_objects.
The first part is a partial substitution (which will be later
applied to lib_objects when completed).
@@ -41,7 +41,7 @@ open Mod_subst
*)
type substitutive_objects =
- substitution * mod_bound_id list * mod_self_id * lib_objects
+ mod_bound_id list * module_path * lib_objects
(* For each module, we store the following things:
@@ -77,9 +77,10 @@ let modtab_objects =
(* currently started interactive module (if any) - its arguments (if it
is a functor) and declared output type *)
-let openmod_info =
- ref (([],None,None) : mod_bound_id list * module_struct_entry option
- * struct_expr_body option)
+let openmod_info =
+ ref ((MPfile(initial_dir),[],None,None)
+ : module_path * mod_bound_id list * module_struct_entry option
+ * module_type_body option)
(* The library_cache here is needed to avoid recalculations of
substituted modules object during "reloading" of libraries *)
@@ -99,7 +100,8 @@ let _ = Summary.declare_summary "MODULE-INFO"
Summary.init_function = (fun () ->
modtab_substobjs := MPmap.empty;
modtab_objects := MPmap.empty;
- openmod_info := ([],None,None);
+ openmod_info := ((MPfile(initial_dir),
+ [],None,None));
library_cache := Dirmap.empty) }
(* auxiliary functions to transform full_path and kernel_name given
@@ -116,50 +118,18 @@ let dir_of_sp sp =
let dir,id = repr_path sp in
add_dirpath_suffix dir id
-let msid_of_mp = function
- MPself msid -> msid
- | _ -> anomaly "'Self' module path expected!"
-
-let msid_of_prefix (_,(mp,sec)) =
- if sec=empty_dirpath then
- msid_of_mp mp
- else
- anomaly ("Non-empty section in module name!" ^
- string_of_mp mp ^ "." ^ string_of_dirpath sec)
-
-let scrape_alias mp =
- Environ.scrape_alias mp (Global.env())
-
(* This function checks if the type calculated for the module [mp] is
a subtype of [sub_mtb]. Uses only the global environment. *)
let check_subtypes mp sub_mtb =
let env = Global.env () in
- let mtb = Environ.lookup_modtype mp env in
- let sub_mtb =
- {typ_expr = sub_mtb;
- typ_strength = None;
- typ_alias = empty_subst} in
- let _ = Environ.add_constraints
- (Subtyping.check_subtypes env mtb sub_mtb)
+ let mb = Environ.lookup_module mp env in
+ let mtb = Modops.module_type_of_module env None mb in
+ let _ = Environ.add_constraints
+ (Subtyping.check_subtypes env mtb sub_mtb)
in
() (* The constraints are checked and forgot immediately! *)
-let compute_subst_objects mp (subst,mbids,msid,objs) =
- match mbids with
- | [] ->
- let subst' = join_alias (map_msid msid mp) subst in
- Some (join (map_msid msid mp) (join subst' subst), objs)
- | _ ->
- None
-
-let subst_substobjs dir mp substobjs =
- match compute_subst_objects mp substobjs with
- | Some (subst, objs) ->
- let prefix = dir,(mp,empty_dirpath) in
- Some (subst_objects prefix subst objs)
- | None -> None
-
(* These functions register the visibility of the module and iterates
through its components. They are called by plenty module functions *)
@@ -178,12 +148,12 @@ let compute_visibility exists what i dir dirinfo =
errorlabstrm (what^"_module") (pr_dirpath dir ++ str " already exists")
else
Nametab.Until i
-
+(*
let do_load_and_subst_module i dir mp substobjs keep =
let prefix = (dir,(mp,empty_dirpath)) in
let dirinfo = DirModule (dir,(mp,empty_dirpath)) in
let vis = compute_visibility false "load_and_subst" i dir dirinfo in
- let objects = compute_subst_objects mp substobjs in
+ let objects = compute_subst_objects mp substobjs resolver in
Nametab.push_dir vis dir dirinfo;
modtab_substobjs := MPmap.add mp substobjs !modtab_substobjs;
match objects with
@@ -194,55 +164,33 @@ let do_load_and_subst_module i dir mp substobjs keep =
Some (seg@keep)
| None ->
None
+*)
-let do_module exists what iter_objects i dir mp substobjs objects =
+let do_module exists what iter_objects i dir mp substobjs keep=
let prefix = (dir,(mp,empty_dirpath)) in
let dirinfo = DirModule (dir,(mp,empty_dirpath)) in
let vis = compute_visibility exists what i dir dirinfo in
Nametab.push_dir vis dir dirinfo;
modtab_substobjs := MPmap.add mp substobjs !modtab_substobjs;
- match objects with
- Some seg ->
- modtab_objects := MPmap.add mp (prefix,seg) !modtab_objects;
- iter_objects (i+1) prefix seg
- | None -> ()
+ match substobjs with
+ ([],mp1,objs) ->
+ modtab_objects := MPmap.add mp (prefix,objs@keep) !modtab_objects;
+ iter_objects (i+1) prefix (objs@keep)
+ | (mbids,_,_) -> ()
let conv_names_do_module exists what iter_objects i
- (sp,kn) substobjs substituted =
+ (sp,kn) substobjs =
let dir,mp = dir_of_sp sp, mp_of_kn kn in
- do_module exists what iter_objects i dir mp substobjs substituted
+ do_module exists what iter_objects i dir mp substobjs []
(* Interactive modules and module types cannot be recached! cache_mod*
functions can be called only once (and "end_mod*" set the flag to
false then)
*)
-
-let cache_module ((sp,kn as oname),(entry,substobjs,substituted)) =
- let _ = match entry with
- | None ->
- anomaly "You must not recache interactive modules!"
- | Some (me,sub_mte_o) ->
- let sub_mtb_o = match sub_mte_o with
- None -> None
- | Some mte -> Some (Mod_typing.translate_struct_entry
- (Global.env()) mte)
- in
-
- let mp = Global.add_module (basename sp) me in
- if mp <> mp_of_kn kn then
- anomaly "Kernel and Library names do not match";
-
- match sub_mtb_o with
- None -> ()
- | Some (sub_mtb,sub) ->
- check_subtypes mp sub_mtb
-
- in
- conv_names_do_module false "cache" load_objects 1 oname substobjs substituted
-
-
-
-
+let cache_module ((sp,kn),(entry,substobjs)) =
+ let dir,mp = dir_of_sp sp, mp_of_kn kn in
+ do_module false "cache" load_objects 1 dir mp substobjs []
+
(* TODO: This check is not essential *)
let check_empty s = function
| None -> ()
@@ -253,42 +201,26 @@ let check_empty s = function
(* When this function is called the module itself is already in the
environment. This function loads its objects only *)
-let load_module i (oname,(entry,substobjs,substituted)) =
+let load_module i (oname,(entry,substobjs)) =
(* TODO: This check is not essential *)
check_empty "load_module" entry;
- conv_names_do_module false "load" load_objects i oname substobjs substituted
+ conv_names_do_module false "load" load_objects i oname substobjs
-let open_module i (oname,(entry,substobjs,substituted)) =
+let open_module i (oname,(entry,substobjs)) =
(* TODO: This check is not essential *)
check_empty "open_module" entry;
- conv_names_do_module true "open" open_objects i oname substobjs substituted
+ conv_names_do_module true "open" open_objects i oname substobjs
-let subst_module ((sp,kn),subst,(entry,substobjs,_)) =
+let subst_module (subst,(entry,(mbids,mp,objs))) =
check_empty "subst_module" entry;
- let dir,mp = dir_of_sp sp, mp_of_kn kn in
- let (sub,mbids,msid,objs) = substobjs in
- let sub = subst_key subst sub in
- let sub' = update_subst_alias subst sub in
- let sub' = update_subst_alias sub' (map_msid msid mp) in
- (* let sub = join_alias sub sub' in*)
- let sub = join sub' sub in
- let subst' = join sub subst in
- (* substitutive_objects get the new substitution *)
- let substobjs = (subst',mbids,msid,objs) in
- (* if we are not a functor - calculate substitued.
- We add "msid |-> mp" to the substitution *)
- let substituted = subst_substobjs dir mp substobjs
- in
- (None,substobjs,substituted)
-
-
-let classify_module (_,substobjs,_) =
- Substitute (None,substobjs,None)
+ (None,(mbids,subst_mp subst mp, subst_objects subst objs))
+let classify_module (_,substobjs) =
+ Substitute (None,substobjs)
let (in_module,out_module) =
declare_object {(default_object "MODULE") with
@@ -298,168 +230,6 @@ let (in_module,out_module) =
subst_function = subst_module;
classify_function = classify_module }
-
-let rec replace_alias modalias_obj obj =
- let rec put_alias (id_alias,obj_alias) l =
- match l with
- [] -> []
- | (id,o)::r
- when ( object_tag o = "MODULE") ->
- if id = id_alias then
-(* let (entry,subst_o,substed_o) = out_module_alias obj_alias in
- let (entry',subst_o',substed_o') = out_module o in
- begin
- match substed_o,substed_o' with
- Some a,Some b ->
- (id,in_module_alias
- (entry,subst_o',Some (dump_alias_object a b)))::r*)
- (id_alias,obj_alias)::r
- (* | _,_ -> (id,o)::r
- end*)
- else (id,o)::(put_alias (id_alias,obj_alias) r)
- | e::r -> e::(put_alias (id_alias,obj_alias) r) in
- let rec choose_obj_alias list_alias list_obj =
- match list_alias with
- | [] -> list_obj
- | o::r ->choose_obj_alias r (put_alias o list_obj) in
- choose_obj_alias modalias_obj obj
-
-and dump_alias_object alias_obj obj =
- let rec alias_in_obj seg =
- match seg with
- | [] -> []
- | (id,o)::r when (object_tag o = "MODULE ALIAS") ->
- (id,o)::(alias_in_obj r)
- | e::r -> (alias_in_obj r) in
- let modalias_obj = alias_in_obj alias_obj in
- replace_alias modalias_obj obj
-
-and do_module_alias exists what iter_objects i dir mp alias substobjs objects =
- let prefix = (dir,(alias,empty_dirpath)) in
- let alias_objects =
- try Some (MPmap.find alias !modtab_objects) with
- Not_found -> None in
- let dirinfo = DirModule (dir,(mp,empty_dirpath)) in
- let vis = compute_visibility exists what i dir dirinfo in
- Nametab.push_dir vis dir dirinfo;
- modtab_substobjs := MPmap.add mp substobjs !modtab_substobjs;
- match alias_objects,objects with
- Some (_,seg), Some seg' ->
- let new_seg = dump_alias_object seg seg' in
- modtab_objects := MPmap.add mp (prefix,new_seg) !modtab_objects;
- iter_objects (i+1) prefix new_seg
- | _,_-> ()
-
-and cache_module_alias ((sp,kn),(entry,substobjs,substituted)) =
- let dir,mp,alias = match entry with
- | None ->
- anomaly "You must not recache interactive modules!"
- | Some (me,sub_mte_o) ->
- let sub_mtb_o = match sub_mte_o with
- None -> None
- | Some mte -> Some (Mod_typing.translate_struct_entry
- (Global.env()) mte)
- in
-
- let mp' = match me with
- | {mod_entry_type = None;
- mod_entry_expr = Some (MSEident mp)} ->
- Global.add_alias (basename sp) mp
- | _ -> anomaly "cache module alias"
- in
- if mp' <> mp_of_kn kn then
- anomaly "Kernel and Library names do not match";
-
- let _ = match sub_mtb_o with
- None -> ()
- | Some (sub_mtb,sub) ->
- check_subtypes mp' sub_mtb in
- match me with
- | {mod_entry_type = None;
- mod_entry_expr = Some (MSEident mp)} ->
- dir_of_sp sp,mp_of_kn kn,scrape_alias mp
- | _ -> anomaly "cache module alias"
- in
- do_module_alias false "cache" load_objects 1 dir mp alias substobjs substituted
-
-and load_module_alias i ((sp,kn),(entry,substobjs,substituted)) =
- let dir,mp,alias=
- match entry with
- | Some (me,_)->
- begin
- match me with
- |{mod_entry_type = None;
- mod_entry_expr = Some (MSEident mp)} ->
- dir_of_sp sp,mp_of_kn kn,scrape_alias mp
- | _ -> anomaly "Modops: Not an alias"
- end
- | None -> anomaly "Modops: Empty info"
- in
- do_module_alias false "load" load_objects i dir mp alias substobjs substituted
-
-and open_module_alias i ((sp,kn),(entry,substobjs,substituted)) =
- let dir,mp,alias=
- match entry with
- | Some (me,_)->
- begin
- match me with
- |{mod_entry_type = None;
- mod_entry_expr = Some (MSEident mp)} ->
- dir_of_sp sp,mp_of_kn kn,scrape_alias mp
- | _ -> anomaly "Modops: Not an alias"
- end
- | None -> anomaly "Modops: Empty info"
- in
- do_module_alias true "open" open_objects i dir mp alias substobjs substituted
-
-and subst_module_alias ((sp,kn),subst,(entry,substobjs,_)) =
- let dir,mp = dir_of_sp sp, mp_of_kn kn in
- let (sub,mbids,msid,objs) = substobjs in
- let sub' = update_subst_alias subst (map_msid msid mp) in
- let subst' = join sub' subst in
- let subst' = join sub subst' in
- (* substitutive_objects get the new substitution *)
- let substobjs = (subst',mbids,msid,objs) in
- (* if we are not a functor - calculate substitued.
- We add "msid |-> mp" to the substitution *)
- match entry with
- | Some (me,sub)->
- begin
- match me with
- |{mod_entry_type = None;
- mod_entry_expr = Some (MSEident mp')} ->
- let mp' = subst_mp subst' mp' in
- let mp' = scrape_alias mp' in
- (Some ({mod_entry_type = None;
- mod_entry_expr =
- Some (MSEident mp')},sub),
- substobjs, match mbids with
- | [] -> let subst = update_subst subst' (map_mp mp' mp) in
- Some (subst_objects (dir,(mp',empty_dirpath))
- (join (join subst' subst) (join (map_msid msid mp')
- (map_mp mp mp')))
- objs)
-
- | _ -> None)
-
- | _ -> anomaly "Modops: Not an alias"
- end
- | None -> anomaly "Modops: Empty info"
-
-and classify_module_alias (entry,substobjs,_) =
- Substitute (entry,substobjs,None)
-
-let (in_module_alias,out_module_alias) =
- declare_object {(default_object "MODULE ALIAS") with
- cache_function = cache_module_alias;
- open_function = open_module_alias;
- classify_function = classify_module_alias;
- subst_function = subst_module_alias;
- load_function = load_module_alias }
-
-
-
-
let cache_keep _ = anomaly "This module should not be cached!"
let load_keep i ((sp,kn),seg) =
@@ -553,9 +323,9 @@ let open_modtype i ((sp,kn),(entry,_)) =
Nametab.push_modtype (Nametab.Exactly i) sp (mp_of_kn kn)
-let subst_modtype (_,subst,(entry,(subs,mbids,msid,objs))) =
+let subst_modtype (subst,(entry,(mbids,mp,objs))) =
check_empty "subst_modtype" entry;
- (entry,(join subs subst,mbids,msid,objs))
+ (entry,(mbids,subst_mp subst mp,subst_objects subst objs))
let classify_modtype (_,substobjs) =
@@ -571,147 +341,143 @@ let (in_modtype,_) =
classify_function = classify_modtype }
-
-let rec replace_module_object idl (subst, mbids, msid, lib_stack) modobjs mp =
- let rec mp_rec = function
- | [] -> MPself msid
- | i::r -> MPdot(mp_rec r,label_of_id i)
- in
- if mbids<>[] then
+let rec replace_module_object idl ( mbids, mp, lib_stack) (mbids2,mp2,objs) mp1=
+ if mbids<>[] then
error "Unexpected functor objects"
- else
- let rec replace_idl = function
- | _,[] -> []
- | id::idl,(id',obj)::tail when id = id' ->
- let tag = object_tag obj in
- if tag = "MODULE" or tag ="MODULE ALIAS" then
- (match idl with
- [] -> (id, in_module_alias (Some
- ({mod_entry_type = None;
- mod_entry_expr = Some (MSEident mp)},None)
- ,modobjs,None))::tail
- | _ ->
- let (a,substobjs,_) = if tag = "MODULE ALIAS" then
- out_module_alias obj else out_module obj in
- let substobjs' = replace_module_object idl substobjs modobjs mp in
- if tag = "MODULE ALIAS" then
- (id, in_module_alias (a,substobjs',None))::tail
- else
- (id, in_module (None,substobjs',None))::tail
- )
- else error "MODULE expected!"
- | idl,lobj::tail -> lobj::replace_idl (idl,tail)
- in
- (join (map_mp (mp_rec (List.rev idl)) mp) subst, mbids, msid, replace_idl (idl,lib_stack))
-
-let abstract_substobjs mbids1 (subst, mbids2, msid, lib_stack) =
- (subst, mbids1@mbids2, msid, lib_stack)
-
-let rec get_modtype_substobjs env = function
- MSEident ln -> MPmap.find ln !modtypetab
+ else
+ let rec replace_idl = function
+ | _,[] -> []
+ | id::idl,(id',obj)::tail when id = id' ->
+ if object_tag obj = "MODULE" then
+ (match idl with
+ [] -> (id, in_module
+ (None,(mbids,(MPdot(mp,label_of_id id)),subst_objects
+ (map_mp mp1 (MPdot(mp,label_of_id id)) empty_delta_resolver) objs)))::tail
+ | _ ->
+ let (_,substobjs) = out_module obj in
+ let substobjs' = replace_module_object idl substobjs
+ (mbids2,mp2,objs) mp in
+ (id, in_module (None,substobjs'))::tail
+ )
+ else error "MODULE expected!"
+ | idl,lobj::tail -> lobj::replace_idl (idl,tail)
+ in
+ (mbids, mp, replace_idl (idl,lib_stack))
+
+let discr_resolver mb =
+ match mb.mod_type with
+ SEBstruct _ ->
+ Some mb.mod_delta
+ | _ -> (*case mp is a functor *)
+ None
+
+(* Small function to avoid module typing during substobjs retrivial *)
+let rec get_objs_modtype_application env = function
+| MSEident mp ->
+ MPmap.find mp !modtypetab,Environ.lookup_modtype mp env,[]
+| MSEapply (fexpr, MSEident mp) ->
+ let objs,mtb,mp_l= get_objs_modtype_application env fexpr in
+ objs,mtb,mp::mp_l
+| MSEapply (_,mexpr) ->
+ Modops.error_application_to_not_path mexpr
+| _ -> error "Application of a non-functor."
+
+let rec get_modtype_substobjs env mp_from= function
+ MSEident ln ->
+ MPmap.find ln !modtypetab
| MSEfunctor (mbid,_,mte) ->
- let (subst, mbids, msid, objs) = get_modtype_substobjs env mte in
- (subst, mbid::mbids, msid, objs)
- | MSEwith (mty, With_Definition _) -> get_modtype_substobjs env mty
- | MSEwith (mty, With_Module (idl,mp)) ->
- let substobjs = get_modtype_substobjs env mty in
- let mp = Environ.scrape_alias mp env in
- let modobjs = MPmap.find mp !modtab_substobjs in
- replace_module_object idl substobjs modobjs mp
- | MSEapply (mexpr, MSEident mp) ->
- let ftb,sub1 = Mod_typing.translate_struct_entry env mexpr in
- let farg_id, farg_b, fbody_b = Modops.destr_functor env
- (Modops.eval_struct env ftb) in
- let mp = Environ.scrape_alias mp env in
- let sub_alias = (Environ.lookup_modtype mp env).typ_alias in
- let sub_alias = match Modops.eval_struct env (SEBident mp) with
- | SEBstruct (msid,sign) -> join_alias
- (subst_key (map_msid msid mp) sub_alias)
- (map_msid msid mp)
- | _ -> sub_alias in
- let (subst, mbids, msid, objs) = get_modtype_substobjs env mexpr in
- (match mbids with
- | mbid::mbids ->
- let resolve =
- Modops.resolver_of_environment farg_id farg_b mp sub_alias env in
- let sub3=
- if sub1 = empty_subst then
- update_subst sub_alias (map_mbid farg_id mp None)
- else
- let sub1' = join_alias sub1 (map_mbid farg_id mp None) in
- let sub_alias' = update_subst sub_alias sub1' in
- join sub1' sub_alias'
- in
- let sub3 = join sub3 (update_subst sub_alias (map_mbid farg_id mp None)) in
- (* application outside the kernel, only for substitutive
- objects (that are all non-logical objects) *)
- ((join
- (join subst sub3)
- (map_mbid mbid mp (Some resolve)))
- , mbids, msid, objs)
- | [] -> match mexpr with
- | MSEident _ -> error "Application of a non-functor"
- | _ -> error "Application of a functor with too few arguments")
+ let (mbids, mp, objs) = get_modtype_substobjs env mp_from mte in
+ (mbid::mbids, mp, objs)
+ | MSEwith (mty, With_Definition _) -> get_modtype_substobjs env mp_from mty
+ | MSEwith (mty, With_Module (idl,mp1)) ->
+ let substobjs = get_modtype_substobjs env mp_from mty in
+ let modobjs = MPmap.find mp1 !modtab_substobjs in
+ replace_module_object idl substobjs modobjs mp1
+ | MSEapply (fexpr, MSEident mp) ->
+ let (mbids, mp1, objs),mtb_mp1,mp_l =
+ get_objs_modtype_application env (MSEapply(fexpr, MSEident mp)) in
+ let rec compute_subst mbids sign mp_l =
+ match mbids,mp_l with
+ [],[] -> [],empty_subst
+ | mbid,[] -> mbid,empty_subst
+ | [],r -> error ("Application of a functor with too few arguments.")
+ | mbid::mbids,mp::mp_l ->
+ let farg_id, farg_b, fbody_b = Modops.destr_functor env sign in
+ let mb = Environ.lookup_module mp env in
+ let mp_delta = discr_resolver mb in
+ let mbid_left,subst=compute_subst mbids fbody_b mp_l in
+ if mp_delta = None then
+ mbid_left,join (map_mbid mbid mp empty_delta_resolver) subst
+ else
+ let mp_delta = Modops.complete_inline_delta_resolver env mp
+ farg_id farg_b (Option.get mp_delta) in
+ mbid_left,join (map_mbid mbid mp mp_delta) subst
+ in
+ let mbids_left,subst = compute_subst mbids mtb_mp1.typ_expr (List.rev mp_l) in
+ (mbids_left, mp1,subst_objects subst objs)
| MSEapply (_,mexpr) ->
Modops.error_application_to_not_path mexpr
-
(* push names of bound modules (and their components) to Nametab *)
(* add objects associated to them *)
let process_module_bindings argids args =
let process_arg id (mbid,mty) =
let dir = make_dirpath [id] in
let mp = MPbound mbid in
- let substobjs = get_modtype_substobjs (Global.env()) mty in
- ignore (do_load_and_subst_module 1 dir mp substobjs [])
- in
- List.iter2 process_arg argids args
-
+ let (mbids,mp_from,objs) = get_modtype_substobjs (Global.env()) mp mty in
+ let substobjs = (mbids,mp,subst_objects
+ (map_mp mp_from mp empty_delta_resolver) objs)in
+ do_module false "start" load_objects 1 dir mp substobjs []
+ in
+ List.iter2 process_arg argids args
+
let intern_args interp_modtype (idl,arg) =
let lib_dir = Lib.library_dp() in
let mbids = List.map (fun (_,id) -> make_mbid lib_dir (string_of_id id)) idl in
let mty = interp_modtype (Global.env()) arg in
let dirs = List.map (fun (_,id) -> make_dirpath [id]) idl in
- let substobjs = get_modtype_substobjs (Global.env()) mty in
+ let (mbi,mp_from,objs) = get_modtype_substobjs (Global.env())
+ (MPbound (List.hd mbids)) mty in
List.map2
(fun dir mbid ->
- Global.add_module_parameter mbid mty;
- let mp = MPbound mbid in
- ignore (do_load_and_subst_module 1 dir mp substobjs []);
- (mbid,mty))
+ let resolver = Global.add_module_parameter mbid mty in
+ let mp = MPbound mbid in
+ let substobjs = (mbi,mp,subst_objects
+ (map_mp mp_from mp resolver) objs) in
+ do_module false "interp" load_objects 1 dir mp substobjs [];
+ (mbid,mty))
dirs mbids
let start_module interp_modtype export id args res_o =
let fs = Summary.freeze_summaries () in
-
let mp = Global.start_module id in
let arg_entries = List.concat (List.map (intern_args interp_modtype) args) in
-
let res_entry_o, sub_body_o = match res_o with
None -> None, None
| Some (res, restricted) ->
(* we translate the module here to catch errors as early as possible *)
let mte = interp_modtype (Global.env()) res in
- if restricted then
- Some mte, None
- else
- let mtb,_ = Mod_typing.translate_struct_entry (Global.env()) mte in
- let sub_mtb =
- List.fold_right
- (fun (arg_id,arg_t) mte ->
- let arg_t,sub = Mod_typing.translate_struct_entry (Global.env()) arg_t
+ if restricted then
+ let _ = Mod_typing.translate_struct_type_entry (Global.env()) mte in
+ Some mte, None
+ else
+ let mtb = Mod_typing.translate_module_type (Global.env())
+ mp mte in
+ let funct_mtb =
+ List.fold_right
+ (fun (arg_id,arg_t) mte ->
+ let arg_t = Mod_typing.translate_module_type (Global.env())
+ (MPbound arg_id) arg_t
in
- let arg_t = {typ_expr = arg_t;
- typ_strength = None;
- typ_alias = sub} in
SEBfunctor(arg_id,arg_t,mte))
- arg_entries mtb
- in
- None, Some sub_mtb
+ arg_entries mtb.typ_expr
+ in
+ None, Some {mtb with
+ typ_expr = funct_mtb}
in
let mbids = List.map fst arg_entries in
- openmod_info:=(mbids,res_entry_o,sub_body_o);
+ openmod_info:=(mp,mbids,res_entry_o,sub_body_o);
let prefix = Lib.start_module export id mp fs in
Nametab.push_dir (Nametab.Until 1) (fst prefix) (DirOpenModule prefix);
Lib.add_frozen_state (); mp
@@ -720,24 +486,25 @@ let start_module interp_modtype export id args res_o =
let end_module () =
let oldoname,oldprefix,fs,lib_stack = Lib.end_module () in
- let mbids, res_o, sub_o = !openmod_info in
+ let mp,mbids, res_o, sub_o = !openmod_info in
let substitute, keep, special = Lib.classify_segment lib_stack in
- let dir = fst oldprefix in
- let msid = msid_of_prefix oldprefix in
-
- let substobjs, keep, special = try
+ let mp_from,substobjs, keep, special = try
match res_o with
| None ->
- (empty_subst, mbids, msid, substitute), keep, special
- | Some (MSEident ln) ->
- abstract_substobjs mbids (MPmap.find ln (!modtypetab)), [], []
+ (* the module is not sealed *)
+ None,( mbids, mp, substitute), keep, special
+ | Some (MSEident ln as mty) ->
+ let (mbids1,mp1,objs) = get_modtype_substobjs (Global.env()) mp mty in
+ Some mp1,(mbids@mbids1,mp1,objs), [], []
| Some (MSEwith _ as mty) ->
- abstract_substobjs mbids (get_modtype_substobjs (Global.env()) mty), [], []
+ let (mbids1,mp1,objs) = get_modtype_substobjs (Global.env()) mp mty in
+ Some mp1,(mbids@mbids1,mp1,objs), [], []
| Some (MSEfunctor _) ->
anomaly "Funsig cannot be here..."
| Some (MSEapply _ as mty) ->
- abstract_substobjs mbids (get_modtype_substobjs (Global.env()) mty), [], []
+ let (mbids1,mp1,objs) = get_modtype_substobjs (Global.env()) mp mty in
+ Some mp1,(mbids@mbids1,mp1,objs), [], []
with
Not_found -> anomaly "Module objects not found..."
in
@@ -745,15 +512,21 @@ let end_module () =
dependencies on functor arguments *)
let id = basename (fst oldoname) in
- let mp = Global.end_module fs id res_o in
+ let mp,resolver = Global.end_module fs id res_o in
begin match sub_o with
None -> ()
| Some sub_mtb -> check_subtypes mp sub_mtb
end;
- let substituted = subst_substobjs dir mp substobjs in
- let node = in_module (None,substobjs,substituted) in
+(* we substitute objects if the module is
+ sealed by a signature (ie. mp_from != None *)
+ let substobjs = match mp_from,substobjs with
+ None,_ -> substobjs
+ | Some mp_from,(mbids,_,objs) ->
+ (mbids,mp,subst_objects (map_mp mp_from mp resolver) objs)
+ in
+ let node = in_module (None,substobjs) in
let objects =
if keep = [] || mbids <> [] then
special@[node] (* no keep objects or we are defining a functor *)
@@ -785,28 +558,30 @@ type library_name = dir_path
(* The first two will form substitutive_objects, the last one is keep *)
type library_objects =
- mod_self_id * lib_objects * lib_objects
+ module_path * lib_objects * lib_objects
let register_library dir cenv objs digest =
let mp = MPfile dir in
- try
+ let substobjs, keep =
+ try
ignore(Global.lookup_module mp);
(* if it's in the environment, the cached objects should be correct *)
- let substobjs, objects = Dirmap.find dir !library_cache in
- do_module false "register_library" load_objects 1 dir mp substobjs objects
+ Dirmap.find dir !library_cache
with Not_found ->
if mp <> Global.import cenv digest then
anomaly "Unexpected disk module name";
- let msid,substitute,keep = objs in
- let substobjs = empty_subst, [], msid, substitute in
- let objects = do_load_and_subst_module 1 dir mp substobjs keep in
- let modobjs = substobjs, objects in
- library_cache := Dirmap.add dir modobjs !library_cache
+ let mp,substitute,keep = objs in
+ let substobjs = [], mp, substitute in
+ let modobjs = substobjs, keep in
+ library_cache := Dirmap.add dir modobjs !library_cache;
+ modobjs
+ in
+ do_module false "register_library" load_objects 1 dir mp substobjs keep
let start_library dir =
let mp = Global.start_library dir in
- openmod_info:=[],None,None;
+ openmod_info:=mp,[],None,None;
Lib.start_compilation dir mp;
Lib.add_frozen_state ()
@@ -816,10 +591,9 @@ let set_end_library_hook f = end_library_hook := f
let end_library dir =
!end_library_hook();
let prefix, lib_stack = Lib.end_compilation dir in
- let cenv = Global.export dir in
- let msid = msid_of_prefix prefix in
+ let mp,cenv = Global.export dir in
let substitute, keep, _ = Lib.classify_segment lib_stack in
- cenv,(msid,substitute,keep)
+ cenv,(mp,substitute,keep)
(* implementation of Export M and Import M *)
@@ -838,8 +612,8 @@ let cache_import (_,(_,mp)) =
let classify_import (export,_ as obj) =
if export then Substitute obj else Dispose
-let subst_import (_,subst,(export,mp as obj)) =
- let mp' = subst_mp subst mp in
+let subst_import (subst,(export,mp as obj)) =
+ let mp' = subst_mp subst mp in
if mp'==mp then obj else
(export,mp')
@@ -870,27 +644,23 @@ let start_modtype interp_modtype id args =
let end_modtype () =
-
let oldoname,prefix,fs,lib_stack = Lib.end_modtype () in
let id = basename (fst oldoname) in
let substitute, _, special = Lib.classify_segment lib_stack in
-
- let msid = msid_of_prefix prefix in
let mbids = !openmodtype_info in
- let modtypeobjs = empty_subst, mbids, msid, substitute in
-
- let ln = Global.end_modtype fs id in
+ let mp = Global.end_modtype fs id in
+ let modtypeobjs = mbids, mp, substitute in
let oname = Lib.add_leaves id (special@[in_modtype (None, modtypeobjs)]) in
if fst oname <> fst oldoname then
anomaly
"Section paths generated on start_ and end_modtype do not match";
- if (mp_of_kn (snd oname)) <> ln then
+ if (mp_of_kn (snd oname)) <> mp then
anomaly
"Kernel and Library names do not match";
Lib.add_frozen_state ()(* to prevent recaching *);
- ln
+ mp
let declare_modtype interp_modtype id args mty =
@@ -907,64 +677,79 @@ let declare_modtype interp_modtype id args mty =
arg_entries
base_mty
in
- let substobjs = get_modtype_substobjs (Global.env()) entry in
- (* Undo the simulated interactive building of the module type *)
- (* and declare the module type as a whole *)
- Summary.unfreeze_summaries fs;
+ let (mbids,mp_from,objs) = get_modtype_substobjs (Global.env()) mmp entry in
+ (* Undo the simulated interactive building of the module type *)
+ (* and declare the module type as a whole *)
- ignore (add_leaf id (in_modtype (Some entry, substobjs)));
- mmp
+ let substobjs = (mbids,mmp,
+ subst_objects (map_mp mp_from mmp empty_delta_resolver) objs) in
+ Summary.unfreeze_summaries fs;
+ ignore (add_leaf id (in_modtype (Some entry, substobjs)));
+ mmp
with e ->
(* Something wrong: undo the whole process *)
Summary.unfreeze_summaries fs; raise e
+(* Small function to avoid module typing during substobjs retrivial *)
+let rec get_objs_module_application env = function
+| MSEident mp ->
+ MPmap.find mp !modtab_substobjs,Environ.lookup_module mp env,[]
+| MSEapply (fexpr, MSEident mp) ->
+ let objs,mtb,mp_l= get_objs_module_application env fexpr in
+ objs,mtb,mp::mp_l
+| MSEapply (_,mexpr) ->
+ Modops.error_application_to_not_path mexpr
+| _ -> error "Application of a non-functor."
-let rec get_module_substobjs env = function
+
+let rec get_module_substobjs env mp_from = function
| MSEident mp -> MPmap.find mp !modtab_substobjs
| MSEfunctor (mbid,mty,mexpr) ->
- let (subst, mbids, msid, objs) = get_module_substobjs env mexpr in
- (subst, mbid::mbids, msid, objs)
- | MSEapply (mexpr, MSEident mp) ->
- let ftb,sub1 = Mod_typing.translate_struct_entry env mexpr in
- let farg_id, farg_b, fbody_b = Modops.destr_functor env
- (Modops.eval_struct env ftb) in
- let mp = Environ.scrape_alias mp env in
- let sub_alias = (Environ.lookup_modtype mp env).typ_alias in
- let sub_alias = match Modops.eval_struct env (SEBident mp) with
- | SEBstruct (msid,sign) -> join_alias
- (subst_key (map_msid msid mp) sub_alias)
- (map_msid msid mp)
- | _ -> sub_alias in
- let (subst, mbids, msid, objs) = get_module_substobjs env mexpr in
+ let (mbids, mp, objs) = get_module_substobjs env mp_from mexpr in
+ (mbid::mbids, mp, objs)
+ | MSEapply (fexpr, MSEident mp) ->
+ let (mbids, mp1, objs),mb_mp1,mp_l =
+ get_objs_module_application env (MSEapply(fexpr, MSEident mp)) in
+ let rec compute_subst mbids sign mp_l =
+ match mbids,mp_l with
+ [],[] -> [],empty_subst
+ | mbid,[] -> mbid,empty_subst
+ | [],r -> error ("Application of a functor with too few arguments.")
+ | mbid::mbids,mp::mp_l ->
+ let farg_id, farg_b, fbody_b = Modops.destr_functor env sign in
+ let mb = Environ.lookup_module mp env in
+ let mp_delta = discr_resolver mb in
+ let mbid_left,subst=compute_subst mbids fbody_b mp_l in
+ if mp_delta = None then
+ mbid_left,join (map_mbid mbid mp empty_delta_resolver) subst
+ else
+ let mp_delta = Modops.complete_inline_delta_resolver env mp
+ farg_id farg_b (Option.get mp_delta) in
+ mbid_left,join (map_mbid mbid mp mp_delta) subst
+ in
+ let mbids_left,subst = compute_subst mbids mb_mp1.mod_type (List.rev mp_l) in
+ (mbids_left, mp1,subst_objects subst objs)
+ (* let sign,alg,equiv,_ = Mod_typing.translate_struct_module_entry env mp_from fexpr in
+ let farg_id, farg_b, fbody_b = Modops.destr_functor env sign in
+ let mb = Environ.lookup_module mp env in
+ let mp_delta = discr_resolver mb in
+ let (mbids, mp1, objs) = get_module_substobjs env mp_from fexpr in
(match mbids with
| mbid::mbids ->
- let resolve =
- Modops.resolver_of_environment farg_id farg_b mp sub_alias env in
- let sub3=
- if sub1 = empty_subst then
- update_subst sub_alias (map_mbid farg_id mp None)
- else
- let sub1' = join_alias sub1 (map_mbid farg_id mp None) in
- let sub_alias' = update_subst sub_alias sub1' in
- join sub1' sub_alias'
- in
- let sub3 = join sub3 (update_subst sub_alias (map_mbid farg_id mp None)) in
- (* application outside the kernel, only for substitutive
- objects (that are all non-logical objects) *)
- ((join
- (join subst sub3)
- (map_mbid mbid mp (Some resolve)))
- , mbids, msid, objs)
- | [] -> match mexpr with
+ if mp_delta = None then
+ (mbids, mp1,subst_objects (map_mbid mbid mp empty_delta_resolver) objs)
+ else
+ let mp_delta = Modops.complete_inline_delta_resolver env mp
+ farg_id farg_b (Option.get mp_delta) in
+ (mbids, mp1,subst_objects (map_mbid mbid mp mp_delta) objs)
+ | [] -> match fexpr with
| MSEident _ -> error "Application of a non-functor."
- | _ -> error "Application of a functor with too few arguments.")
+ | _ -> error "Application of a functor with too few arguments.")*)
| MSEapply (_,mexpr) ->
Modops.error_application_to_not_path mexpr
- | MSEwith (mty, With_Definition _) -> get_module_substobjs env mty
- | MSEwith (mty, With_Module (idl,mp)) ->
- let substobjs = get_module_substobjs env mty in
- let modobjs = MPmap.find mp !modtab_substobjs in
- replace_module_object idl substobjs modobjs mp
+ | MSEwith (mty, With_Definition _) -> get_module_substobjs env mp_from mty
+ | MSEwith (mty, With_Module (idl,mp)) -> assert false
+
(* Include *)
@@ -991,46 +776,32 @@ let lift_oname (sp,kn) =
let dir,_ = Libnames.repr_path sp in
(dir,mp)
-let cache_include (oname,((me,is_mod),substobjs,substituted)) =
+let cache_include (oname,((me,is_mod),(mbis,mp1,objs))) =
let dir,mp1 = lift_oname oname in
let prefix = (dir,(mp1,empty_dirpath)) in
- Global.add_include me;
- match substituted with
- Some seg ->
- load_objects 1 prefix seg;
- open_objects 1 prefix seg;
- | None -> ()
-
-let load_include i (oname,((me,is_mod),substobjs,substituted)) =
+ load_objects 1 prefix objs;
+ open_objects 1 prefix objs
+
+let load_include i (oname,((me,is_mod),(mbis,mp1,objs))) =
let dir,mp1 = lift_oname oname in
let prefix = (dir,(mp1,empty_dirpath)) in
- match substituted with
- Some seg ->
- load_objects i prefix seg
- | None -> ()
-
-let open_include i (oname,((me,is_mod),substobjs,substituted)) =
+ load_objects i prefix objs
+
+
+let open_include i (oname,((me,is_mod),(mbis,mp1,objs))) =
let dir,mp1 = lift_oname oname in
let prefix = (dir,(mp1,empty_dirpath)) in
- match substituted with
- Some seg ->
- if is_mod then
- open_objects i prefix seg
- else
- if i = 1 then
- open_objects i prefix seg
- | None -> ()
-
-let subst_include (oname,subst,((me,is_mod),substobj,_)) =
- let dir,mp1 = lift_oname oname in
- let (sub,mbids,msid,objs) = substobj in
- let subst' = join sub subst in
- let substobjs = (subst',mbids,msid,objs) in
- let substituted = subst_substobjs dir mp1 substobjs in
- ((subst_inc_expr subst' me,is_mod),substobjs,substituted)
-
-let classify_include ((me,is_mod),substobjs,_) =
- Substitute ((me,is_mod),substobjs,None)
+ if is_mod || i = 1 then
+ open_objects i prefix objs
+ else ()
+
+let subst_include (subst,((me,is_mod),substobj)) =
+ let (mbids,mp,objs) = substobj in
+ let substobjs = (mbids,subst_mp subst mp,subst_objects subst objs) in
+ ((subst_inc_expr subst me,is_mod),substobjs)
+
+let classify_include ((me,is_mod),substobjs) =
+ Substitute ((me,is_mod),substobjs)
let (in_include,out_include) =
declare_object {(default_object "INCLUDE") with
@@ -1040,20 +811,19 @@ let (in_include,out_include) =
subst_function = subst_include;
classify_function = classify_include }
-let rec update_include (sub,mbids,msid,objs) =
+let rec update_include (mbids,mp,objs) =
let rec replace_include = function
| [] -> []
| (id,obj)::tail ->
if object_tag obj = "INCLUDE" then
- let ((me,is_mod),substobjs,substituted) = out_include obj in
+ let ((me,is_mod),substobjs) = out_include obj in
let substobjs' = update_include substobjs in
- (id, in_include ((me,true),substobjs',substituted))::
+ (id, in_include ((me,true),substobjs'))::
(replace_include tail)
else
(id,obj)::(replace_include tail)
in
- (sub,mbids,msid,replace_include objs)
-
+ (mbids,mp,replace_include objs)
let declare_module interp_modtype interp_modexpr id args mty_o mexpr_o =
@@ -1084,8 +854,7 @@ let declare_module interp_modtype interp_modexpr id args mty_o mexpr_o =
Some (List.fold_right
(fun (mbid,mte) me -> MSEfunctor(mbid,mte,me))
arg_entries
- (interp_modexpr (Global.env()) mexpr))
- in
+ (interp_modexpr (Global.env()) mexpr)) in
let entry =
{mod_entry_type = mty_entry_o;
mod_entry_expr = mexpr_entry_o }
@@ -1093,45 +862,33 @@ let declare_module interp_modtype interp_modexpr id args mty_o mexpr_o =
let env = Global.env() in
let substobjs =
match entry with
- | {mod_entry_type = Some mte} -> get_modtype_substobjs env mte
- | {mod_entry_expr = Some mexpr} -> get_module_substobjs env mexpr
+ | {mod_entry_type = Some mte} -> get_modtype_substobjs env mmp mte
+ | {mod_entry_expr = Some mexpr} -> get_module_substobjs env mmp mexpr
| _ -> anomaly "declare_module: No type, no body ..."
in
- let substobjs = update_include substobjs in
+ let (mbids,mp_from,objs) = update_include substobjs in
(* Undo the simulated interactive building of the module *)
(* and declare the module as a whole *)
Summary.unfreeze_summaries fs;
- match entry with
- |{mod_entry_type = None;
- mod_entry_expr = Some (MSEident mp) } ->
- let dir,mp' = dir_of_sp (Lib.make_path id), mp_of_kn (Lib.make_kn id) in
- let (sub,mbids,msid,objs) = substobjs in
- let mp1 = Environ.scrape_alias mp env in
- let prefix = dir,(mp1,empty_dirpath) in
- let substituted =
- match mbids with
- | [] ->
- Some (subst_objects prefix
- (join sub (join (map_msid msid mp1) (map_mp mp' mp1))) objs)
- | _ -> None in
- ignore (add_leaf
- id
- (in_module_alias (Some ({mod_entry_type = None;
- mod_entry_expr = Some (MSEident mp1) }, mty_sub_o),
- substobjs, substituted)));
- mmp
- | _ ->
- let dir,mp = dir_of_sp (Lib.make_path id), mp_of_kn (Lib.make_kn id) in
- let (sub,mbids,msid,objs) = substobjs in
- let sub' = join_alias (subst_key (map_msid msid mp) sub) (map_msid msid mp) in
- let substobjs = ( sub',mbids,msid,objs) in
- let substituted = subst_substobjs dir mp substobjs in
- ignore (add_leaf
- id
- (in_module (Some (entry, mty_sub_o), substobjs, substituted)));
- mmp
-
- with e ->
+ let dir,mp = dir_of_sp (Lib.make_path id), mp_of_kn (Lib.make_kn id) in
+ let mp_env,resolver = Global.add_module id entry in
+ let _ = if mp_env <> mp then
+ anomaly "Kernel and Library names do not match";
+ match mty_sub_o with
+ | None -> ()
+ | Some sub_mte ->
+ let sub_mtb = Mod_typing.translate_module_type
+ env mp sub_mte in
+ check_subtypes mp_env sub_mtb
+ in
+ let substobjs = (mbids,mp_env,
+ subst_objects(map_mp mp_from mp_env resolver) objs) in
+ ignore (add_leaf
+ id
+ (in_module (Some (entry, mty_sub_o), substobjs)));
+ mmp
+
+ with e ->
(* Something wrong: undo the whole process *)
Summary.unfreeze_summaries fs; raise e
@@ -1142,19 +899,19 @@ let declare_include interp_struct me_ast is_mod =
try
let env = Global.env() in
- let me = interp_struct env me_ast in
- let substobjs =
+ let me = interp_struct env me_ast in
+ let mp1,_ = current_prefix () in
+ let (mbids,mp,objs)=
if is_mod then
- get_module_substobjs env me
+ get_module_substobjs env mp1 me
else
- get_modtype_substobjs env me in
- let mp1,_ = current_prefix () in
- let dir = dir_of_sp (Lib.path_of_include()) in
- let substituted = subst_substobjs dir mp1 substobjs in
+ get_modtype_substobjs env mp1 me in
let id = current_mod_id() in
-
+ let resolver = Global.add_include me is_mod in
+ let substobjs = (mbids,mp1,
+ subst_objects (map_mp mp mp1 resolver) objs) in
ignore (add_leaf id
- (in_include ((me,is_mod), substobjs, substituted)))
+ (in_include ((me,is_mod), substobjs)))
with e ->
(* Something wrong: undo the whole process *)
Summary.unfreeze_summaries fs; raise e
diff --git a/library/global.ml b/library/global.ml
index e228de23a..6d7942ec0 100644
--- a/library/global.ml
+++ b/library/global.ml
@@ -48,14 +48,6 @@ let push_named_def d =
global_env := env;
cst
-(*let add_thing add kn thing =
- let _,dir,l = repr_kn kn in
- let kn',newenv = add dir l thing !global_env in
- if kn = kn' then
- global_env := newenv
- else
- anomaly "Kernel names do not match."
-*)
let add_thing add dir id thing =
let kn, newenv = add dir (label_of_id id) thing !global_env in
@@ -65,14 +57,23 @@ let add_thing add dir id thing =
let add_constant = add_thing add_constant
let add_mind = add_thing add_mind
let add_modtype = add_thing (fun _ -> add_modtype) ()
-let add_module = add_thing (fun _ -> add_module) ()
-let add_alias = add_thing (fun _ -> add_alias) ()
+
+
+let add_module id me =
+ let l = label_of_id id in
+ let mp,resolve,new_env = add_module l me !global_env in
+ global_env := new_env;
+ mp,resolve
+
let add_constraints c = global_env := add_constraints c !global_env
let set_engagement c = global_env := set_engagement c !global_env
-let add_include me = global_env := add_include me !global_env
+let add_include me is_module =
+ let resolve,newenv = add_include me is_module !global_env in
+ global_env := newenv;
+ resolve
let start_module id =
let l = label_of_id id in
@@ -82,15 +83,16 @@ let start_module id =
let end_module fs id mtyo =
let l = label_of_id id in
- let mp,newenv = end_module l mtyo !global_env in
+ let mp,resolve,newenv = end_module l mtyo !global_env in
Summary.unfreeze_summaries fs;
global_env := newenv;
- mp
+ mp,resolve
let add_module_parameter mbid mte =
- let newenv = add_module_parameter mbid mte !global_env in
- global_env := newenv
+ let resolve,newenv = add_module_parameter mbid mte !global_env in
+ global_env := newenv;
+ resolve
let start_modtype id =
@@ -117,15 +119,22 @@ let lookup_mind kn = lookup_mind kn (env())
let lookup_module mp = lookup_module mp (env())
let lookup_modtype kn = lookup_modtype kn (env())
+let constant_of_delta con =
+ let resolver,resolver_param = (delta_of_senv !global_env) in
+ Mod_subst.constant_of_delta resolver_param
+ (Mod_subst.constant_of_delta resolver con)
-
-
+let mind_of_delta mind =
+ let resolver,resolver_param = (delta_of_senv !global_env) in
+ Mod_subst.mind_of_delta resolver_param
+ (Mod_subst.mind_of_delta resolver mind)
+
let start_library dir =
let mp,newenv = start_library dir !global_env in
global_env := newenv;
mp
-let export s = snd (export !global_env s)
+let export s = export !global_env s
let import cenv digest =
let mp,newenv = import cenv digest !global_env in
@@ -161,3 +170,5 @@ let register field value by_clause =
let entry = kind_of_term value in
let senv = Safe_typing.register !global_env field entry by_clause in
global_env := senv
+
+
diff --git a/library/global.mli b/library/global.mli
index 3c2317122..30bd04150 100644
--- a/library/global.mli
+++ b/library/global.mli
@@ -15,6 +15,7 @@ open Term
open Declarations
open Entries
open Indtypes
+open Mod_subst
open Safe_typing
(*i*)
@@ -47,12 +48,11 @@ val push_named_def : (identifier * constr * types option) -> Univ.constraints
val add_constant :
dir_path -> identifier -> global_declaration -> constant
val add_mind :
- dir_path -> identifier -> mutual_inductive_entry -> kernel_name
+ dir_path -> identifier -> mutual_inductive_entry -> mutual_inductive
-val add_module : identifier -> module_entry -> module_path
+val add_module : identifier -> module_entry -> module_path * delta_resolver
val add_modtype : identifier -> module_struct_entry -> module_path
-val add_include : module_struct_entry -> unit
-val add_alias : identifier -> module_path -> module_path
+val add_include : module_struct_entry -> bool -> delta_resolver
val add_constraints : constraints -> unit
@@ -66,10 +66,11 @@ val set_engagement : engagement -> unit
of the started module / module type *)
val start_module : identifier -> module_path
-val end_module :
- Summary.frozen -> identifier -> module_struct_entry option -> module_path
-val add_module_parameter : mod_bound_id -> module_struct_entry -> unit
+val end_module : Summary.frozen ->identifier -> module_struct_entry option ->
+ module_path * delta_resolver
+
+val add_module_parameter : mod_bound_id -> module_struct_entry -> delta_resolver
val start_modtype : identifier -> module_path
val end_modtype : Summary.frozen -> identifier -> module_path
@@ -83,10 +84,12 @@ val lookup_inductive : inductive -> mutual_inductive_body * one_inductive_body
val lookup_mind : mutual_inductive -> mutual_inductive_body
val lookup_module : module_path -> module_body
val lookup_modtype : module_path -> module_type_body
+val constant_of_delta : constant -> constant
+val mind_of_delta : mutual_inductive -> mutual_inductive
(* Compiled modules *)
val start_library : dir_path -> module_path
-val export : dir_path -> compiled_library
+val export : dir_path -> module_path * compiled_library
val import : compiled_library -> Digest.t -> module_path
(*s Function to get an environment from the constants part of the global
diff --git a/library/goptions.ml b/library/goptions.ml
index 032016c3d..06d4b618e 100644
--- a/library/goptions.ml
+++ b/library/goptions.ml
@@ -91,7 +91,7 @@ module MakeTable =
| GOadd -> t := MySet.add p !t
| GOrmv -> t := MySet.remove p !t in
let load_options i o = if i=1 then cache_options o in
- let subst_options (_,subst,(f,p as obj)) =
+ let subst_options (subst,(f,p as obj)) =
let p' = A.subst subst p in
if p' == p then obj else
(f,p')
diff --git a/library/heads.ml b/library/heads.ml
index 150ba8942..056f78a5f 100644
--- a/library/heads.ml
+++ b/library/heads.ml
@@ -39,8 +39,20 @@ type head_approximation =
(** Registration as global tables and rollback. *)
+module Evalreford = struct
+ type t = evaluable_global_reference
+ let compare x y =
+ let make_name = function
+ | EvalConstRef con ->
+ EvalConstRef(constant_of_kn(canonical_con con))
+ | k -> k
+ in
+ Pervasives.compare (make_name x) (make_name y)
+end
+
module Evalrefmap =
- Map.Make (struct type t = evaluable_global_reference let compare = compare end)
+ Map.Make (Evalreford)
+
let head_map = ref Evalrefmap.empty
@@ -144,7 +156,7 @@ let subst_head_approximation subst = function
kind_of_head (Global.env()) c
| x -> x
-let subst_head (_,subst,(ref,k)) =
+let subst_head (subst,(ref,k)) =
(subst_evaluable_reference subst ref, subst_head_approximation subst k)
let discharge_head (_,(ref,k)) =
diff --git a/library/impargs.ml b/library/impargs.ml
index d27ced220..1edac69aa 100644
--- a/library/impargs.ml
+++ b/library/impargs.ml
@@ -435,7 +435,7 @@ type implicit_interactive_request =
type implicit_discharge_request =
| ImplLocal
| ImplConstant of constant * implicits_flags
- | ImplMutualInductive of kernel_name * implicits_flags
+ | ImplMutualInductive of mutual_inductive * implicits_flags
| ImplInteractive of global_reference * implicits_flags *
implicit_interactive_request
@@ -455,7 +455,7 @@ let cache_implicits o =
let subst_implicits_decl subst (r,imps as o) =
let r' = fst (subst_global subst r) in if r==r' then o else (r',imps)
-let subst_implicits (_,subst,(req,l)) =
+let subst_implicits (subst,(req,l)) =
(ImplLocal,list_smartmap (subst_implicits_decl subst) l)
let impls_of_context ctx =
diff --git a/library/impargs.mli b/library/impargs.mli
index 6d2b01e8f..e8191e863 100644
--- a/library/impargs.mli
+++ b/library/impargs.mli
@@ -108,7 +108,7 @@ type implicit_interactive_request =
type implicit_discharge_request =
| ImplLocal
| ImplConstant of constant * implicits_flags
- | ImplMutualInductive of kernel_name * implicits_flags
+ | ImplMutualInductive of mutual_inductive * implicits_flags
| ImplInteractive of global_reference * implicits_flags *
implicit_interactive_request
diff --git a/library/lib.ml b/library/lib.ml
index 961a4ebb9..b8dcee9d2 100644
--- a/library/lib.ml
+++ b/library/lib.ml
@@ -37,21 +37,21 @@ let iter_objects f i prefix =
let load_objects = iter_objects load_object
let open_objects = iter_objects open_object
-let subst_objects prefix subst seg =
+let subst_objects subst seg =
let subst_one = fun (id,obj as node) ->
- let obj' = subst_object (make_oname prefix id, subst, obj) in
+ let obj' = subst_object (subst,obj) in
if obj' == obj then node else
(id, obj')
in
list_smartmap subst_one seg
-let load_and_subst_objects i prefix subst seg =
+(*let load_and_subst_objects i prefix subst seg =
List.rev (List.fold_left (fun seg (id,obj as node) ->
let obj' = subst_object (make_oname prefix id, subst, obj) in
let node = if obj == obj' then node else (id, obj') in
load_object i (make_oname prefix id, obj');
node :: seg) [] seg)
-
+*)
let classify_segment seg =
let rec clean ((substl,keepl,anticipl) as acc) = function
| (_,CompilingLibrary _) :: _ | [] -> acc
@@ -435,13 +435,13 @@ type binding_kind = Explicit | Implicit
type variable_info = Names.identifier * binding_kind * Term.constr option * Term.types
type variable_context = variable_info list
-type abstr_list = variable_context Names.Cmap.t * variable_context Names.KNmap.t
+type abstr_list = variable_context Names.Cmap.t * variable_context Names.Mindmap.t
let sectab =
ref ([] : ((Names.identifier * binding_kind) list * Cooking.work_list * abstr_list) list)
let add_section () =
- sectab := ([],(Names.Cmap.empty,Names.KNmap.empty),(Names.Cmap.empty,Names.KNmap.empty)) :: !sectab
+ sectab := ([],(Names.Cmap.empty,Names.Mindmap.empty),(Names.Cmap.empty,Names.Mindmap.empty)) :: !sectab
let add_section_variable id impl =
match !sectab with
@@ -474,7 +474,7 @@ let add_section_replacement f g hyps =
sectab := (vars,f args exps,g sechyps abs)::sl
let add_section_kn kn =
- let f x (l1,l2) = (l1,Names.KNmap.add kn x l2) in
+ let f x (l1,l2) = (l1,Names.Mindmap.add kn x l2) in
add_section_replacement f f
let add_section_constant kn =
@@ -489,7 +489,7 @@ let section_segment_of_constant con =
Names.Cmap.find con (fst (pi3 (List.hd !sectab)))
let section_segment_of_mutual_inductive kn =
- Names.KNmap.find kn (snd (pi3 (List.hd !sectab)))
+ Names.Mindmap.find kn (snd (pi3 (List.hd !sectab)))
let rec list_mem_assoc x = function
| [] -> raise Not_found
@@ -502,7 +502,7 @@ let section_instance = function
| ConstRef con ->
Names.Cmap.find con (fst (pi2 (List.hd !sectab)))
| IndRef (kn,_) | ConstructRef ((kn,_),_) ->
- Names.KNmap.find kn (snd (pi2 (List.hd !sectab)))
+ Names.Mindmap.find kn (snd (pi2 (List.hd !sectab)))
let is_in_section ref =
try ignore (section_instance ref); true with Not_found -> false
@@ -772,7 +772,7 @@ let mp_of_global ref =
let rec dp_of_mp modp =
match modp with
| Names.MPfile dp -> dp
- | Names.MPbound _ | Names.MPself _ -> library_dp ()
+ | Names.MPbound _ -> library_dp ()
| Names.MPdot (mp,_) -> dp_of_mp mp
let rec split_mp mp =
@@ -781,7 +781,6 @@ let rec split_mp mp =
| Names.MPdot (prfx, lbl) ->
let mprec, dprec = split_mp prfx in
mprec, Names.make_dirpath (Names.id_of_string (Names.string_of_label lbl) :: (Names.repr_dirpath dprec))
- | Names.MPself msid -> let (_, id, dp) = Names.repr_msid msid in library_dp(), Names.make_dirpath [Names.id_of_string id]
| Names.MPbound mbid -> let (_, id, dp) = Names.repr_mbid mbid in library_dp(), Names.make_dirpath [Names.id_of_string id]
let split_modpath mp =
@@ -789,7 +788,6 @@ let split_modpath mp =
| Names.MPfile dp -> dp, []
| Names.MPbound mbid ->
library_dp (), [Names.id_of_mbid mbid]
- | Names.MPself msid -> library_dp (), [Names.id_of_msid msid]
| Names.MPdot (mp,l) -> let (mp', lab) = aux mp in
(mp', Names.id_of_label l :: lab)
in
@@ -819,8 +817,8 @@ let remove_section_part ref =
(* Discharging names *)
let pop_kn kn =
- let (mp,dir,l) = Names.repr_kn kn in
- Names.make_kn mp (pop_dirpath dir) l
+ let (mp,dir,l) = Names.repr_mind kn in
+ Names.make_mind mp (pop_dirpath dir) l
let pop_con con =
let (mp,dir,l) = Names.repr_con con in
@@ -831,7 +829,7 @@ let con_defined_in_sec kn =
dir <> Names.empty_dirpath && fst (split_dirpath dir) = snd (current_prefix ())
let defined_in_sec kn =
- let _,dir,_ = Names.repr_kn kn in
+ let _,dir,_ = Names.repr_mind kn in
dir <> Names.empty_dirpath && fst (split_dirpath dir) = snd (current_prefix ())
let discharge_global = function
diff --git a/library/lib.mli b/library/lib.mli
index 0e2e304cd..32d1c0009 100644
--- a/library/lib.mli
+++ b/library/lib.mli
@@ -32,8 +32,8 @@ type lib_objects = (Names.identifier * Libobject.obj) list
val open_objects : int -> Libnames.object_prefix -> lib_objects -> unit
val load_objects : int -> Libnames.object_prefix -> lib_objects -> unit
-val subst_objects : Libnames.object_prefix -> Mod_subst.substitution -> lib_objects -> lib_objects
-val load_and_subst_objects : int -> Libnames.object_prefix -> Mod_subst.substitution -> lib_objects -> lib_objects
+val subst_objects : Mod_subst.substitution -> lib_objects -> lib_objects
+(*val load_and_subst_objects : int -> Libnames.object_prefix -> Mod_subst.substitution -> lib_objects -> lib_objects*)
(* [classify_segment seg] verifies that there are no OpenedThings,
clears ClosedSections and FrozenStates and divides Leafs according
@@ -183,13 +183,13 @@ val is_in_section : Libnames.global_reference -> bool
val add_section_variable : Names.identifier -> binding_kind -> unit
val add_section_constant : Names.constant -> Sign.named_context -> unit
-val add_section_kn : Names.kernel_name -> Sign.named_context -> unit
+val add_section_kn : Names.mutual_inductive -> Sign.named_context -> unit
val replacement_context : unit ->
- (Names.identifier array Names.Cmap.t * Names.identifier array Names.KNmap.t)
+ (Names.identifier array Names.Cmap.t * Names.identifier array Names.Mindmap.t)
(*s Discharge: decrease the section level if in the current section *)
-val discharge_kn : Names.kernel_name -> Names.kernel_name
+val discharge_kn : Names.mutual_inductive -> Names.mutual_inductive
val discharge_con : Names.constant -> Names.constant
val discharge_global : Libnames.global_reference -> Libnames.global_reference
val discharge_inductive : Names.inductive -> Names.inductive
diff --git a/library/libnames.ml b/library/libnames.ml
index 2b335ea6c..fad0336fc 100644
--- a/library/libnames.ml
+++ b/library/libnames.ml
@@ -27,13 +27,21 @@ let isConstRef = function ConstRef _ -> true | _ -> false
let isIndRef = function IndRef _ -> true | _ -> false
let isConstructRef = function ConstructRef _ -> true | _ -> false
+let eq_gr gr1 gr2 =
+ match gr1,gr2 with
+ ConstRef con1, ConstRef con2 ->
+ eq_constant con1 con2
+ | IndRef kn1,IndRef kn2 -> eq_ind kn1 kn2
+ | ConstructRef kn1,ConstructRef kn2 -> eq_constructor kn1 kn2
+ | _,_ -> gr1=gr2
+
let destVarRef = function VarRef ind -> ind | _ -> failwith "destVarRef"
let destConstRef = function ConstRef ind -> ind | _ -> failwith "destConstRef"
let destIndRef = function IndRef ind -> ind | _ -> failwith "destIndRef"
let destConstructRef = function ConstructRef ind -> ind | _ -> failwith "destConstructRef"
let subst_constructor subst ((kn,i),j as ref) =
- let kn' = subst_kn subst kn in
+ let kn' = subst_ind subst kn in
if kn==kn' then ref, mkConstruct ref
else ((kn',i),j), mkConstruct ((kn',i),j)
@@ -43,7 +51,7 @@ let subst_global subst ref = match ref with
let kn',t = subst_con subst kn in
if kn==kn' then ref, mkConst kn else ConstRef kn', t
| IndRef (kn,i) ->
- let kn' = subst_kn subst kn in
+ let kn' = subst_ind subst kn in
if kn==kn' then ref, mkInd (kn,i) else IndRef(kn',i), mkInd (kn',i)
| ConstructRef ((kn,i),j as c) ->
let c',t = subst_constructor subst c in
@@ -65,15 +73,25 @@ let constr_of_global = function
let constr_of_reference = constr_of_global
let reference_of_constr = global_of_constr
-module RefOrdered =
- struct
- type t = global_reference
- let compare = Pervasives.compare
- end
-
+(* outside of the kernel, names are ordered on their canonical part *)
+module RefOrdered = struct
+ type t = global_reference
+ let compare x y =
+ let make_name = function
+ | ConstRef con ->
+ ConstRef(constant_of_kn(canonical_con con))
+ | IndRef (kn,i) ->
+ IndRef(mind_of_kn(canonical_mind kn),i)
+ | ConstructRef ((kn,i),j )->
+ ConstructRef((mind_of_kn(canonical_mind kn),i),j)
+ | VarRef id -> VarRef id
+ in
+ Pervasives.compare (make_name x) (make_name y)
+end
+
module Refset = Set.Make(RefOrdered)
module Refmap = Map.Make(RefOrdered)
-
+
(* Extended global references *)
type syndef_name = kernel_name
@@ -191,26 +209,22 @@ let restrict_path n sp =
let dir' = list_firstn n (repr_dirpath dir) in
make_path (make_dirpath dir') s
-let encode_kn dir id = make_kn (MPfile dir) empty_dirpath (label_of_id id)
+let encode_mind dir id = make_mind (MPfile dir) empty_dirpath (label_of_id id)
let encode_con dir id = make_con (MPfile dir) empty_dirpath (label_of_id id)
-let decode_kn kn =
- let rec dirpath_of_module = function
+let decode_mind kn =
+ let rec dir_of_mp = function
| MPfile dir -> repr_dirpath dir
- | MPbound mbid ->
+ | MPbound mbid ->
let _,_,dp = repr_mbid mbid in
let id = id_of_mbid mbid in
id::(repr_dirpath dp)
- | MPself msid ->
- let _,_,dp = repr_msid msid in
- let id = id_of_msid msid in
- id::(repr_dirpath dp)
- | MPdot(mp,l) -> (id_of_label l)::(dirpath_of_module mp)
+ | MPdot(mp,l) -> (id_of_label l)::(dir_of_mp mp)
in
- let mp,sec_dir,l = repr_kn kn in
+ let mp,sec_dir,l = repr_mind kn in
if (repr_dirpath sec_dir) = [] then
- (make_dirpath (dirpath_of_module mp)),id_of_label l
+ (make_dirpath (dir_of_mp mp)),id_of_label l
else
anomaly "Section part should be empty!"
@@ -289,8 +303,8 @@ let pop_con con =
Names.make_con mp (pop_dirpath dir) l
let pop_kn kn =
- let (mp,dir,l) = repr_kn kn in
- Names.make_kn mp (pop_dirpath dir) l
+ let (mp,dir,l) = repr_mind kn in
+ Names.make_mind mp (pop_dirpath dir) l
let pop_global_reference = function
| ConstRef con -> ConstRef (pop_con con)
diff --git a/library/libnames.mli b/library/libnames.mli
index 43ca252c1..fd2ca37ae 100644
--- a/library/libnames.mli
+++ b/library/libnames.mli
@@ -28,11 +28,14 @@ val isConstRef : global_reference -> bool
val isIndRef : global_reference -> bool
val isConstructRef : global_reference -> bool
+val eq_gr : global_reference -> global_reference -> bool
+
val destVarRef : global_reference -> variable
val destConstRef : global_reference -> constant
val destIndRef : global_reference -> inductive
val destConstructRef : global_reference -> constructor
+
val subst_constructor : substitution -> constructor -> constructor * constr
val subst_global : substitution -> global_reference -> global_reference * constr
@@ -47,6 +50,12 @@ val global_of_constr : constr -> global_reference
val constr_of_reference : global_reference -> constr
val reference_of_constr : constr -> global_reference
+module RefOrdered : sig
+ type t = global_reference
+ val compare : global_reference -> global_reference -> int
+end
+
+
module Refset : Set.S with type elt = global_reference
module Refmap : Map.S with type key = global_reference
@@ -108,8 +117,8 @@ val restrict_path : int -> full_path -> full_path
(*s Temporary function to brutally form kernel names from section paths *)
-val encode_kn : dir_path -> identifier -> kernel_name
-val decode_kn : kernel_name -> dir_path * identifier
+val encode_mind : dir_path -> identifier -> mutual_inductive
+val decode_mind : mutual_inductive -> dir_path * identifier
val encode_con : dir_path -> identifier -> constant
val decode_con : constant -> dir_path * identifier
@@ -170,7 +179,7 @@ val loc_of_reference : reference -> loc
(*s Popping one level of section in global names *)
val pop_con : constant -> constant
-val pop_kn : kernel_name -> kernel_name
+val pop_kn : mutual_inductive-> mutual_inductive
val pop_global_reference : global_reference -> global_reference
(* Deprecated synonyms *)
diff --git a/library/libobject.ml b/library/libobject.ml
index 4bd701e13..ecdcacf1d 100644
--- a/library/libobject.ml
+++ b/library/libobject.ml
@@ -34,7 +34,7 @@ type 'a object_declaration = {
load_function : int -> object_name * 'a -> unit;
open_function : int -> object_name * 'a -> unit;
classify_function : 'a -> 'a substitutivity;
- subst_function : object_name * substitution * 'a -> 'a;
+ subst_function : substitution * 'a -> 'a;
discharge_function : object_name * 'a -> 'a option;
rebuild_function : 'a -> 'a }
@@ -63,7 +63,7 @@ let default_object s = {
This helps introducing new functions in objects.
*)
-let ident_subst_function (_,_,a) = a
+let ident_subst_function (_,a) = a
type obj = Dyn.t (* persistent dynamic objects *)
@@ -71,7 +71,7 @@ type dynamic_object_declaration = {
dyn_cache_function : object_name * obj -> unit;
dyn_load_function : int -> object_name * obj -> unit;
dyn_open_function : int -> object_name * obj -> unit;
- dyn_subst_function : object_name * substitution * obj -> obj;
+ dyn_subst_function : substitution * obj -> obj;
dyn_classify_function : obj -> obj substitutivity;
dyn_discharge_function : object_name * obj -> obj option;
dyn_rebuild_function : obj -> obj }
@@ -93,9 +93,9 @@ let declare_object odecl =
and opener i (oname,lobj) =
if Dyn.tag lobj = na then odecl.open_function i (oname,outfun lobj)
else anomaly "somehow we got the wrong dynamic object in the openfun"
- and substituter (oname,sub,lobj) =
- if Dyn.tag lobj = na then
- infun (odecl.subst_function (oname,sub,outfun lobj))
+ and substituter (sub,lobj) =
+ if Dyn.tag lobj = na then
+ infun (odecl.subst_function (sub,outfun lobj))
else anomaly "somehow we got the wrong dynamic object in the substfun"
and classifier lobj =
if Dyn.tag lobj = na then
@@ -158,7 +158,7 @@ let load_object i ((_,lobj) as node) =
let open_object i ((_,lobj) as node) =
apply_dyn_fun () (fun d -> d.dyn_open_function i node) lobj
-let subst_object ((_,_,lobj) as node) =
+let subst_object ((_,lobj) as node) =
apply_dyn_fun lobj (fun d -> d.dyn_subst_function node) lobj
let classify_object lobj =
diff --git a/library/libobject.mli b/library/libobject.mli
index 834a70c64..9c0abafde 100644
--- a/library/libobject.mli
+++ b/library/libobject.mli
@@ -70,7 +70,7 @@ type 'a object_declaration = {
load_function : int -> object_name * 'a -> unit;
open_function : int -> object_name * 'a -> unit;
classify_function : 'a -> 'a substitutivity;
- subst_function : object_name * substitution * 'a -> 'a;
+ subst_function : substitution * 'a -> 'a;
discharge_function : object_name * 'a -> 'a option;
rebuild_function : 'a -> 'a }
@@ -86,7 +86,7 @@ type 'a object_declaration = {
val default_object : string -> 'a object_declaration
(* the identity substitution function *)
-val ident_subst_function : object_name * substitution * 'a -> 'a
+val ident_subst_function : substitution * 'a -> 'a
(*s Given an object declaration, the function [declare_object]
will hand back two functions, the "injection" and "projection"
@@ -102,7 +102,7 @@ val object_tag : obj -> string
val cache_object : object_name * obj -> unit
val load_object : int -> object_name * obj -> unit
val open_object : int -> object_name * obj -> unit
-val subst_object : object_name * substitution * obj -> obj
+val subst_object : substitution * obj -> obj
val classify_object : obj -> obj substitutivity
val discharge_object : object_name * obj -> obj option
val rebuild_object : obj -> obj
diff --git a/library/library.ml b/library/library.ml
index d129a24db..fbe437fc4 100644
--- a/library/library.ml
+++ b/library/library.ml
@@ -302,7 +302,7 @@ let open_import i (_,(dir,export)) =
let cache_import obj =
open_import 1 obj
-let subst_import (_,_,o) = o
+let subst_import (_,o) = o
let classify_import (_,export as obj) =
if export then Substitute obj else Dispose
diff --git a/library/nametab.ml b/library/nametab.ml
index 31915c95a..c14b6cfc0 100644
--- a/library/nametab.ml
+++ b/library/nametab.ml
@@ -314,18 +314,26 @@ let the_tacticrevtab = ref (KNmap.empty : knrevtab)
Parameter but also Remark and Fact) *)
let push_xref visibility sp xref =
- the_ccitab := SpTab.push visibility sp xref !the_ccitab;
match visibility with
| Until _ ->
- if Globrevtab.mem xref !the_globrevtab then
- ()
- else
- the_globrevtab := Globrevtab.add xref sp !the_globrevtab
- | _ -> ()
+ the_ccitab := SpTab.push visibility sp xref !the_ccitab;
+ the_globrevtab := Globrevtab.add xref sp !the_globrevtab
+ | _ ->
+ begin
+ if SpTab.exists sp !the_ccitab then
+ match SpTab.find sp !the_ccitab with
+ | TrueGlobal( ConstRef _) | TrueGlobal( IndRef _) |
+ TrueGlobal( ConstructRef _) as xref ->
+ the_ccitab := SpTab.push visibility sp xref !the_ccitab;
+ | _ ->
+ the_ccitab := SpTab.push visibility sp xref !the_ccitab;
+ else
+ the_ccitab := SpTab.push visibility sp xref !the_ccitab;
+ end
let push_cci visibility sp ref =
push_xref visibility sp (TrueGlobal ref)
-
+
(* This is for Syntactic Definitions *)
let push_syndef visibility sp kn =
push_xref visibility sp (SynDef kn)
diff --git a/parsing/prettyp.ml b/parsing/prettyp.ml
index 12a3bb572..f96a94fb4 100644
--- a/parsing/prettyp.ml
+++ b/parsing/prettyp.ml
@@ -425,7 +425,7 @@ let gallina_print_leaf_entry with_values ((sp,kn as oname),lobj) =
| (_,"CONSTANT") ->
Some (print_constant with_values sep (constant_of_kn kn))
| (_,"INDUCTIVE") ->
- Some (gallina_print_inductive kn)
+ Some (gallina_print_inductive (mind_of_kn kn))
| (_,"MODULE") ->
let (mp,_,l) = repr_kn kn in
Some (print_module with_values (MPdot (mp,l)))
@@ -540,7 +540,7 @@ let print_full_pure_context () =
| ((_,kn),Lib.Leaf lobj)::rest ->
let pp = match object_tag lobj with
| "CONSTANT" ->
- let con = constant_of_kn kn in
+ let con = Global.constant_of_delta (constant_of_kn kn) in
let cb = Global.lookup_constant con in
let val_0 = cb.const_body in
let typ = ungeneralized_type_of_constant_type cb.const_type in
@@ -559,9 +559,10 @@ let print_full_pure_context () =
str " : " ++ pr_ltype typ ++ cut () ++ str " := " ++
print_body val_0) ++ str "." ++ fnl () ++ fnl ()
| "INDUCTIVE" ->
- let (mib,mip) = Global.lookup_inductive (kn,0) in
+ let mind = Global.mind_of_delta (mind_of_kn kn) in
+ let (mib,mip) = Global.lookup_inductive (mind,0) in
let mipv = mib.mind_packets in
- let names = list_tabulate (fun x -> (kn,x)) (Array.length mipv) in
+ let names = list_tabulate (fun x -> (mind,x)) (Array.length mipv) in
pr_mutual_inductive mib.mind_finite names ++ str "." ++
fnl () ++ fnl ()
| "MODULE" ->
diff --git a/parsing/printmod.ml b/parsing/printmod.ml
index a5470a892..80f05bf92 100644
--- a/parsing/printmod.ml
+++ b/parsing/printmod.ml
@@ -27,7 +27,7 @@ let rec print_local_modpath locals = function
| MPbound mbid -> pr_id (List.assoc mbid locals)
| MPdot(mp,l) ->
print_local_modpath locals mp ++ str "." ++ pr_lab l
- | MPself _ | MPfile _ -> raise Not_found
+ | MPfile _ -> raise Not_found
let print_modpath locals mp =
try (* must be with let because streams are lazy! *)
@@ -58,9 +58,10 @@ let rec print_module name locals with_body mb =
| true, Some mexpr ->
spc () ++ str ":= " ++ print_modexpr locals mexpr
in
- let modtype = match mb.mod_type with
- None -> str ""
- | Some t -> spc () ++ str": " ++
+
+ let modtype =
+ match mb.mod_type with
+ | t -> spc () ++ str": " ++
print_modtype locals t
in
hv 2 (str "Module " ++ name ++ modtype ++ body)
@@ -78,8 +79,8 @@ and print_modtype locals mty =
pr_id (id_of_mbid mbid) ++ str " : " ++
print_modtype locals mtb1.typ_expr ++
str ")" ++ spc() ++ print_modtype locals' mtb2)
- | SEBstruct (msid,sign) ->
- hv 2 (str "Sig" ++ spc () ++ print_sig locals msid sign ++ brk (1,-2) ++ str "End")
+ | SEBstruct (sign) ->
+ hv 2 (str "Sig" ++ spc () ++ print_sig locals sign ++ brk (1,-2) ++ str "End")
| SEBapply (mexpr,marg,_) ->
let lapp = flatten_app mexpr [marg] in
let fapp = List.hd lapp in
@@ -90,31 +91,29 @@ and print_modtype locals mty =
let s = (String.concat "." (List.map string_of_id idl)) in
hov 2 (print_modtype locals seb ++ spc() ++ str "with" ++ spc() ++
str "Definition"++ spc() ++ str s ++ spc() ++ str ":="++ spc())
- | SEBwith(seb,With_module_body(idl,mp,_,_))->
+ | SEBwith(seb,With_module_body(idl,mp))->
let s =(String.concat "." (List.map string_of_id idl)) in
hov 2 (print_modtype locals seb ++ spc() ++ str "with" ++ spc() ++
str "Module"++ spc() ++ str s ++ spc() ++ str ":="++ spc())
-and print_sig locals msid sign =
+and print_sig locals sign =
let print_spec (l,spec) = (match spec with
| SFBconst {const_body=Some _; const_opaque=false} -> str "Definition "
| SFBconst {const_body=None}
| SFBconst {const_opaque=true} -> str "Parameter "
| SFBmind _ -> str "Inductive "
| SFBmodule _ -> str "Module "
- | SFBalias (mp,_,_) -> str "Module "
| SFBmodtype _ -> str "Module Type ") ++ str (string_of_label l)
in
prlist_with_sep spc print_spec sign
-and print_struct locals msid struc =
+and print_struct locals struc =
let print_body (l,body) = (match body with
| SFBconst {const_body=Some _; const_opaque=false} -> str "Definition "
| SFBconst {const_body=Some _; const_opaque=true} -> str "Theorem "
| SFBconst {const_body=None} -> str "Parameter "
| SFBmind _ -> str "Inductive "
| SFBmodule _ -> str "Module "
- | SFBalias (mp,_,_) -> str "Module "
| SFBmodtype _ -> str "Module Type ") ++ str (string_of_label l)
in
prlist_with_sep spc print_body struc
@@ -128,8 +127,8 @@ and print_modexpr locals mexpr = match mexpr with
hov 2 (str "Functor" ++ spc() ++ str"(" ++ pr_id(id_of_mbid mbid) ++
str ":" ++ print_modtype locals mty.typ_expr ++
str ")" ++ spc () ++ print_modexpr locals' mexpr)
- | SEBstruct (msid, struc) ->
- hv 2 (str "Struct" ++ spc () ++ print_struct locals msid struc ++ brk (1,-2) ++ str "End")
+ | SEBstruct ( struc) ->
+ hv 2 (str "Struct" ++ spc () ++ print_struct locals struc ++ brk (1,-2) ++ str "End")
| SEBapply (mexpr,marg,_) ->
let lapp = flatten_app mexpr [marg] in
hov 3 (str"(" ++ prlist_with_sep spc (print_modexpr locals) lapp ++ str")")
diff --git a/plugins/extraction/common.ml b/plugins/extraction/common.ml
index e1edeec37..47381f6d8 100644
--- a/plugins/extraction/common.ml
+++ b/plugins/extraction/common.ml
@@ -203,7 +203,7 @@ let pop_visible, push_visible, get_visible, subst_mp =
and push mp o =
vis := { mp = mp; content = Hashtbl.create 97 } :: !vis;
let s = List.hd !sub in
- let s = match o with None -> s | Some msid -> add_msid msid mp s in
+ let s = match o with None -> s | Some msid -> add_mp msid mp empty_delta_resolver s in
sub := s :: !sub
and get () = !vis
and subst mp = subst_mp (List.hd !sub) mp
@@ -278,7 +278,6 @@ let rec mp_renaming_fun mp = match mp with
let lmp = mp_renaming mp in
if lmp = [""] then (modfstlev_rename l)::lmp
else (modular_rename Mod (id_of_label l))::lmp
- | MPself msid -> [modular_rename Mod (id_of_msid msid)]
| MPbound mbid -> [modular_rename Mod (id_of_mbid mbid)]
| MPfile _ when not (modular ()) -> assert false (* see [at_toplevel] above *)
| MPfile _ ->
diff --git a/plugins/extraction/common.mli b/plugins/extraction/common.mli
index c401bd059..a68e72d26 100644
--- a/plugins/extraction/common.mli
+++ b/plugins/extraction/common.mli
@@ -45,7 +45,7 @@ val pp_global : kind -> global_reference -> string
val pp_module : module_path -> string
val top_visible_mp : unit -> module_path
-val push_visible : module_path -> mod_self_id option -> unit
+val push_visible : module_path -> module_path option -> unit
val pop_visible : unit -> unit
val check_duplicate : module_path -> label -> string
diff --git a/plugins/extraction/extract_env.ml b/plugins/extraction/extract_env.ml
index ba4786d37..37721c9b9 100644
--- a/plugins/extraction/extract_env.ml
+++ b/plugins/extraction/extract_env.ml
@@ -32,7 +32,7 @@ let toplevel_env () =
let mp,_,l = repr_kn kn in
let seb = match Libobject.object_tag o with
| "CONSTANT" -> SFBconst (Global.lookup_constant (constant_of_kn kn))
- | "INDUCTIVE" -> SFBmind (Global.lookup_mind kn)
+ | "INDUCTIVE" -> SFBmind (Global.lookup_mind (mind_of_kn kn))
| "MODULE" -> SFBmodule (Global.lookup_module (MPdot (mp,l)))
| "MODULE TYPE" ->
SFBmodtype (Global.lookup_modtype (MPdot (mp,l)))
@@ -41,8 +41,8 @@ let toplevel_env () =
| _ -> failwith "caught"
in
match current_toplevel () with
- | MPself msid -> SEBstruct (msid, List.rev (map_succeed get_reference seg))
- | _ -> assert false
+ | _ -> SEBstruct (List.rev (map_succeed get_reference seg))
+
let environment_until dir_opt =
let rec parse = function
@@ -69,7 +69,7 @@ module type VISIT = sig
(* Add kernel_name / constant / reference / ... in the visit lists.
These functions silently add the mp of their arg in the mp list *)
- val add_kn : kernel_name -> unit
+ val add_kn : mutual_inductive -> unit
val add_con : constant -> unit
val add_ref : global_reference -> unit
val add_decl_deps : ml_decl -> unit
@@ -77,7 +77,7 @@ module type VISIT = sig
(* Test functions:
is a particular object a needed dependency for the current extraction ? *)
- val needed_kn : kernel_name -> bool
+ val needed_kn : mutual_inductive -> bool
val needed_con : constant -> bool
val needed_mp : module_path -> bool
end
@@ -87,17 +87,17 @@ module Visit : VISIT = struct
(for inductives and modules names) and a Cset for constants
(and still the remaining MPset) *)
type must_visit =
- { mutable kn : KNset.t; mutable con : Cset.t; mutable mp : MPset.t }
+ { mutable kn : Mindset.t; mutable con : Cset.t; mutable mp : MPset.t }
(* the imperative internal visit lists *)
- let v = { kn = KNset.empty ; con = Cset.empty ; mp = MPset.empty }
+ let v = { kn = Mindset.empty ; con = Cset.empty ; mp = MPset.empty }
(* the accessor functions *)
- let reset () = v.kn <- KNset.empty; v.con <- Cset.empty; v.mp <- MPset.empty
- let needed_kn kn = KNset.mem kn v.kn
+ let reset () = v.kn <- Mindset.empty; v.con <- Cset.empty; v.mp <- MPset.empty
+ let needed_kn kn = Mindset.mem kn v.kn
let needed_con c = Cset.mem c v.con
let needed_mp mp = MPset.mem mp v.mp
let add_mp mp =
check_loaded_modfile mp; v.mp <- MPset.union (prefixes_mp mp) v.mp
- let add_kn kn = v.kn <- KNset.add kn v.kn; add_mp (modpath kn)
+ let add_kn kn = v.kn <- Mindset.add kn v.kn; add_mp (mind_modpath kn)
let add_con c = v.con <- Cset.add c v.con; add_mp (con_modpath c)
let add_ref = function
| ConstRef c -> add_con c
@@ -140,29 +140,29 @@ let factor_fix env l cb msb =
labels, recd, msb''
end
-let build_mb expr typ_opt =
- { mod_expr = Some expr;
+let build_mb mp expr typ_opt =
+ { mod_mp = mp;
+ mod_expr = Some expr;
mod_type = typ_opt;
+ mod_type_alg = None;
mod_constraints = Univ.Constraint.empty;
- mod_alias = Mod_subst.empty_subst;
+ mod_delta = Mod_subst.empty_delta_resolver;
mod_retroknowledge = [] }
let my_type_of_mb env mb =
- match mb.mod_type with
- | Some mtb -> mtb
- | None -> Modops.eval_struct env (Option.get mb.mod_expr)
+ mb.mod_type
(** Ad-hoc update of environment, inspired by [Mod_type.check_with_aux_def].
To check with Elie. *)
-let env_for_mtb_with env mtb idl =
- let msid,sig_b = match Modops.eval_struct env mtb with
- | SEBstruct(msid,sig_b) -> msid,sig_b
+let env_for_mtb_with env mp mtb idl =
+ let sig_b = match mtb with
+ | SEBstruct(sig_b) -> sig_b
| _ -> assert false
in
let l = label_of_id (List.hd idl) in
let before = fst (list_split_when (fun (l',_) -> l=l') sig_b) in
- Modops.add_signature (MPself msid) before env
+ Modops.add_signature mp before empty_delta_resolver env
(* From a [structure_body] (i.e. a list of [structure_field_body])
to specifications. *)
@@ -177,20 +177,18 @@ let rec extract_sfb_spec env mp = function
else begin Visit.add_spec_deps s; (l,Spec s) :: specs end
| (l,SFBmind _) :: msig ->
let kn = make_kn mp empty_dirpath l in
- let s = Sind (kn, extract_inductive env kn) in
+ let mind = mind_of_kn kn in
+ let s = Sind (kn, extract_inductive env mind) 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,SFBmodule mb) :: msig ->
let specs = extract_sfb_spec env mp msig in
- let spec = extract_seb_spec env (my_type_of_mb env mb) in
+ let spec = extract_seb_spec env mb.mod_mp (my_type_of_mb env mb) in
(l,Smodule spec) :: specs
| (l,SFBmodtype mtb) :: msig ->
let specs = extract_sfb_spec env mp msig in
- (l,Smodtype (extract_seb_spec env mtb.typ_expr)) :: specs
- | (l,SFBalias (mp1,typ_opt,_)) :: msig ->
- let mb = build_mb (SEBident mp1) typ_opt in
- extract_sfb_spec env mp ((l,SFBmodule mb) :: msig)
+ (l,Smodtype (extract_seb_spec env mtb.typ_mp mtb.typ_expr)) :: specs
(* From [struct_expr_body] to specifications *)
@@ -201,34 +199,35 @@ let rec extract_sfb_spec env mp = function
For instance, [my_type_of_mb] ensures this invariant.
*)
-and extract_seb_spec env = function
+and extract_seb_spec env mp1 = function
| SEBident mp -> Visit.add_mp mp; MTident mp
| SEBwith(mtb',With_definition_body(idl,cb))->
- let env' = env_for_mtb_with env mtb' idl in
- let mtb''= extract_seb_spec env mtb' in
+ let env' = env_for_mtb_with env mp1 mtb' idl in
+ let mtb''= extract_seb_spec env mp1 mtb' in
(match extract_with_type env' cb with (* cb peut contenir des kn *)
| None -> mtb''
| Some (vl,typ) -> MTwith(mtb'',ML_With_type(idl,vl,typ)))
- | SEBwith(mtb',With_module_body(idl,mp,_,_))->
+ | SEBwith(mtb',With_module_body(idl,mp))->
Visit.add_mp mp;
- MTwith(extract_seb_spec env mtb',
+ MTwith(extract_seb_spec env mp1 mtb',
ML_With_module(idl,mp))
(* TODO: On pourrait peut-etre oter certaines eta-expansion, du genre:
| SEBfunctor(mbid,_,SEBapply(m,SEBident (MPbound mbid2),_))
when mbid = mbid2 -> extract_seb_spec env m
(* faudrait alors ajouter un test de non-apparition de mbid dans mb *)
*)
- | SEBfunctor (mbid, mtb, mtb') ->
+ | SEBfunctor (mbid, mtb, mtb') ->
let mp = MPbound mbid in
- let env' = Modops.add_module mp (Modops.module_body_of_type mtb) env in
- MTfunsig (mbid, extract_seb_spec env mtb.typ_expr,
- extract_seb_spec env' mtb')
- | SEBstruct (msid, msig) ->
- let mp = MPself msid in
- let env' = Modops.add_signature mp msig env in
- MTsig (msid, extract_sfb_spec env' mp msig)
- | SEBapply _ as mtb ->
- extract_seb_spec env (Modops.eval_struct env mtb)
+ let env' = Modops.add_module (Modops.module_body_of_type mp mtb)
+ env in
+ MTfunsig (mbid, extract_seb_spec env mp mtb.typ_expr,
+ extract_seb_spec env' mp1 mtb')
+ | SEBstruct (msig) ->
+ let env' = Modops.add_signature mp1 msig empty_delta_resolver env in
+ MTsig (mp1, extract_sfb_spec env' mp1 msig)
+ | SEBapply _ ->
+ assert false
+
(* From a [structure_body] (i.e. a list of [structure_field_body])
@@ -263,9 +262,10 @@ let rec extract_sfb env mp all = function
| (l,SFBmind mib) :: msb ->
let ms = extract_sfb env mp all msb in
let kn = make_kn mp empty_dirpath l in
- let b = Visit.needed_kn kn in
+ let mind = mind_of_kn kn in
+ let b = Visit.needed_kn mind in
if all || b then
- let d = Dind (kn, extract_inductive env kn) in
+ let d = Dind (kn, extract_inductive env mind) in
if (not b) && (logical_decl d) then ms
else begin Visit.add_decl_deps d; (l,SEdecl d) :: ms end
else ms
@@ -279,40 +279,34 @@ let rec extract_sfb env mp all = function
let ms = extract_sfb env mp all msb in
let mp = MPdot (mp,l) in
if all || Visit.needed_mp mp then
- (l,SEmodtype (extract_seb_spec env mtb.typ_expr)) :: ms
+ (l,SEmodtype (extract_seb_spec env mp mtb.typ_expr)) :: ms
else ms
- | (l,SFBalias (mp1,typ_opt,_)) :: msb ->
- let mb = build_mb (SEBident mp1) typ_opt in
- extract_sfb env mp all ((l,SFBmodule mb) :: msb)
(* From [struct_expr_body] to implementations *)
-and extract_seb env mpo all = function
+and extract_seb env mp all = function
| SEBident mp ->
if is_modfile mp && not (modular ()) then error_MPfile_as_mod mp false;
Visit.add_mp mp; MEident mp
| SEBapply (meb, meb',_) ->
- MEapply (extract_seb env None true meb,
- extract_seb env None true meb')
+ MEapply (extract_seb env mp true meb,
+ extract_seb env mp true meb')
| SEBfunctor (mbid, mtb, meb) ->
- let mp = MPbound mbid in
- let env' = Modops.add_module mp (Modops.module_body_of_type mtb) env in
- MEfunctor (mbid, extract_seb_spec env mtb.typ_expr,
- extract_seb env' None true meb)
- | SEBstruct (msid, msb) ->
- let mp,msb = match mpo with
- | None -> MPself msid, msb
- | Some mp -> mp, Modops.subst_structure (map_msid msid mp) msb
- in
- let env' = Modops.add_signature mp msb env in
- MEstruct (msid, extract_sfb env' mp all msb)
+ let mp1 = MPbound mbid in
+ let env' = Modops.add_module (Modops.module_body_of_type mp1 mtb)
+ env in
+ MEfunctor (mbid, extract_seb_spec env mp1 mtb.typ_expr,
+ extract_seb env' mp true meb)
+ | SEBstruct (msb) ->
+ let env' = Modops.add_signature mp msb empty_delta_resolver env in
+ MEstruct (mp,extract_sfb env' mp all msb)
| SEBwith (_,_) -> anomaly "Not available yet"
and extract_module env mp all mb =
(* [mb.mod_expr <> None ], since we look at modules from outside. *)
(* Example of module with empty [mod_expr] is X inside a Module F [X:SIG]. *)
- { ml_mod_expr = extract_seb env (Some mp) all (Option.get mb.mod_expr);
- ml_mod_type = extract_seb_spec env (my_type_of_mb env mb) }
+ { ml_mod_expr = extract_seb env mp all (Option.get mb.mod_expr);
+ ml_mod_type = extract_seb_spec env mp (my_type_of_mb env mb) }
let unpack = function MEstruct (_,sel) -> sel | _ -> assert false
@@ -324,7 +318,7 @@ let mono_environment refs mpl =
let env = Global.env () in
let l = List.rev (environment_until None) in
List.rev_map
- (fun (mp,m) -> mp, unpack (extract_seb env (Some mp) false m)) l
+ (fun (mp,m) -> mp, unpack (extract_seb env mp false m)) l
(**************************************)
(*S Part II : Input/Output primitives *)
@@ -514,7 +508,7 @@ let extraction_library is_rec m =
let l = List.rev (environment_until (Some dir_m)) in
let select l (mp,meb) =
if Visit.needed_mp mp
- then (mp, unpack (extract_seb env (Some mp) true meb)) :: l
+ then (mp, unpack (extract_seb env mp true meb)) :: l
else l
in
let struc = List.fold_left select [] l in
diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml
index 3468e8a36..d119dbe8e 100644
--- a/plugins/extraction/extraction.ml
+++ b/plugins/extraction/extraction.ml
@@ -310,7 +310,15 @@ and extract_ind env kn = (* kn is supposed to be in long form *)
with Not_found ->
(* First, if this inductive is aliased via a Module, *)
(* we process the original inductive. *)
- Option.iter (fun kn -> ignore (extract_ind env kn)) mib.mind_equiv;
+ let equiv =
+ if (canonical_mind kn) = (user_mind kn) then
+ NoEquiv
+ else
+ begin
+ ignore (extract_ind env (mind_of_kn (canonical_mind kn)));
+ Equiv (canonical_mind kn)
+ end
+ in
(* Everything concerning parameters. *)
(* We do that first, since they are common to all the [mib]. *)
let mip0 = mib.mind_packets.(0) in
@@ -333,13 +341,12 @@ and extract_ind env kn = (* kn is supposed to be in long form *)
ip_types = t })
mib.mind_packets
in
+
add_ind kn mib
{ind_info = Standard;
ind_nparams = npar;
ind_packets = packets;
- ind_equiv = match mib.mind_equiv with
- | None -> NoEquiv
- | Some kn -> Equiv kn
+ ind_equiv = equiv
};
(* Second pass: we extract constructors *)
for i = 0 to mib.mind_ntypes - 1 do
@@ -388,7 +395,7 @@ 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_kn kn in
+ let mp,d,_ = repr_mind kn in
let rec select_fields l typs = match l,typs with
| [],[] -> []
| (Name id)::l, typ::typs ->
@@ -424,9 +431,7 @@ and extract_ind env kn = (* kn is supposed to be in long form *)
let i = {ind_info = ind_info;
ind_nparams = npar;
ind_packets = packets;
- ind_equiv = match mib.mind_equiv with
- | None -> NoEquiv
- | Some kn -> Equiv kn }
+ ind_equiv = equiv }
in
add_ind kn mib i;
i
diff --git a/plugins/extraction/extraction.mli b/plugins/extraction/extraction.mli
index 2c534ea7b..6bcd24763 100644
--- a/plugins/extraction/extraction.mli
+++ b/plugins/extraction/extraction.mli
@@ -26,7 +26,7 @@ val extract_with_type : env -> constant_body -> ( identifier list * ml_type ) op
val extract_fixpoint :
env -> constant array -> (constr, types) prec_declaration -> ml_decl
-val extract_inductive : env -> kernel_name -> ml_ind
+val extract_inductive : env -> mutual_inductive -> ml_ind
(*s Is a [ml_decl] or a [ml_spec] logical ? *)
diff --git a/plugins/extraction/haskell.ml b/plugins/extraction/haskell.ml
index 9d45c08b7..6973bb454 100644
--- a/plugins/extraction/haskell.ml
+++ b/plugins/extraction/haskell.ml
@@ -84,7 +84,7 @@ let rec pp_type par vl t =
| Tvar i -> (try pr_id (List.nth vl (pred i)) with _ -> (str "a" ++ int i))
| Tglob (r,[]) -> pp_global Type r
| Tglob (r,l) ->
- if r = IndRef (kn_sig,0) then
+ if r = IndRef (mind_of_kn kn_sig,0) then
pp_type true vl (List.hd l)
else
pp_par par
@@ -269,8 +269,8 @@ let pp_string_parameters ids = prlist (fun id -> str id ++ str " ")
let pp_decl = function
| Dind (kn,i) when i.ind_info = Singleton ->
- pp_singleton kn i.ind_packets.(0) ++ fnl ()
- | Dind (kn,i) -> hov 0 (pp_ind true kn 0 i)
+ pp_singleton (mind_of_kn kn) i.ind_packets.(0) ++ fnl ()
+ | Dind (kn,i) -> hov 0 (pp_ind true (mind_of_kn kn) 0 i)
| Dtype (r, l, t) ->
if is_inline_custom r then mt ()
else
diff --git a/plugins/extraction/miniml.mli b/plugins/extraction/miniml.mli
index 55231d766..91c60d205 100644
--- a/plugins/extraction/miniml.mli
+++ b/plugins/extraction/miniml.mli
@@ -128,7 +128,7 @@ type ml_specif =
and ml_module_type =
| MTident of module_path
| MTfunsig of mod_bound_id * ml_module_type * ml_module_type
- | MTsig of mod_self_id * ml_module_sig
+ | MTsig of module_path * ml_module_sig
| MTwith of ml_module_type * ml_with_declaration
and ml_with_declaration =
@@ -145,7 +145,7 @@ type ml_structure_elem =
and ml_module_expr =
| MEident of module_path
| MEfunctor of mod_bound_id * ml_module_type * ml_module_expr
- | MEstruct of mod_self_id * ml_module_structure
+ | MEstruct of module_path * ml_module_structure
| MEapply of ml_module_expr * ml_module_expr
and ml_module_structure = (label * ml_structure_elem) list
diff --git a/plugins/extraction/mlutil.ml b/plugins/extraction/mlutil.ml
index 213df31e6..2e9fb8bf2 100644
--- a/plugins/extraction/mlutil.ml
+++ b/plugins/extraction/mlutil.ml
@@ -649,9 +649,9 @@ let generalizable_list =
let datatypes = MPfile (dirpath_of_string "Coq.Init.Datatypes")
and specif = MPfile (dirpath_of_string "Coq.Init.Specif")
in
- [ make_kn datatypes empty_dirpath (mk_label "bool");
- make_kn specif empty_dirpath (mk_label "sumbool");
- make_kn specif empty_dirpath (mk_label "sumor") ]
+ [ make_mind datatypes empty_dirpath (mk_label "bool");
+ make_mind specif empty_dirpath (mk_label "sumbool");
+ make_mind specif empty_dirpath (mk_label "sumor") ]
let check_generalizable_case unsafe br =
if not unsafe then
diff --git a/plugins/extraction/mlutil.mli b/plugins/extraction/mlutil.mli
index 4b6b47ab9..9c5feac82 100644
--- a/plugins/extraction/mlutil.mli
+++ b/plugins/extraction/mlutil.mli
@@ -50,7 +50,7 @@ end
(*s Utility functions over ML types without meta *)
-val type_mem_kn : kernel_name -> ml_type -> bool
+val type_mem_kn : mutual_inductive -> ml_type -> bool
val type_maxvar : ml_type -> int
diff --git a/plugins/extraction/modutil.ml b/plugins/extraction/modutil.ml
index 1b1a39770..9195b747e 100644
--- a/plugins/extraction/modutil.ml
+++ b/plugins/extraction/modutil.ml
@@ -21,11 +21,7 @@ open Mod_subst
(*S Functions upon ML modules. *)
let rec msid_of_mt = function
- | MTident mp -> begin
- match Modops.eval_struct (Global.env()) (SEBident mp) with
- | SEBstruct(msid,_) -> MPself msid
- | _ -> anomaly "Extraction:the With can't be applied to a funsig"
- end
+ | MTident mp -> mp
| MTwith(mt,_)-> msid_of_mt mt
| _ -> anomaly "Extraction:the With operator isn't applied to a name"
@@ -101,25 +97,26 @@ let ind_iter_references do_term do_cons do_type kn ind =
do_type (IndRef ip);
if lang () = Ocaml then
(match ind.ind_equiv with
- | Equiv kne -> do_type (IndRef (kne, snd ip));
+ | Miniml.Equiv kne -> do_type (IndRef (mind_of_kn kne, snd ip));
| _ -> ());
Array.iteri (fun j -> cons_iter (ip,j+1)) p.ip_types
in
if lang () = Ocaml then record_iter_references do_term ind.ind_info;
- Array.iteri (fun i -> packet_iter (kn,i)) ind.ind_packets
+ Array.iteri (fun i -> packet_iter (kn,i)) ind.ind_packets
let decl_iter_references do_term do_cons do_type =
let type_iter = type_iter_references do_type
and ast_iter = ast_iter_references do_term do_cons do_type in
function
- | Dind (kn,ind) -> ind_iter_references do_term do_cons do_type kn ind
+ | Dind (kn,ind) -> ind_iter_references do_term do_cons do_type
+ (mind_of_kn kn) ind
| Dtype (r,_,t) -> do_type r; type_iter t
| Dterm (r,a,t) -> do_term r; ast_iter a; type_iter t
| Dfix(rv,c,t) ->
Array.iter do_term rv; Array.iter ast_iter c; Array.iter type_iter t
let spec_iter_references do_term do_cons do_type = function
- | Sind (kn,ind) -> ind_iter_references do_term do_cons do_type kn ind
+ | Sind (kn,ind) -> ind_iter_references do_term do_cons do_type (mind_of_kn kn) ind
| Stype (r,_,ot) -> do_type r; Option.iter (type_iter_references do_type) ot
| Sval (r,t) -> do_term r; type_iter_references do_type t
@@ -190,7 +187,7 @@ let signature_of_structure s =
(*s Searching one [ml_decl] in a [ml_structure] by its [global_reference] *)
let get_decl_in_structure r struc =
- try
+ try
let base_mp,ll = labels_of_ref r in
if not (at_toplevel base_mp) then error_not_visible r;
let sel = List.assoc base_mp struc in
@@ -311,7 +308,7 @@ let reset_needed, add_needed, found_needed, is_needed =
(fun r -> Refset.mem (base_r r) !needed))
let declared_refs = function
- | Dind (kn,_) -> [|IndRef (kn,0)|]
+ | Dind (kn,_) -> [|IndRef (mind_of_kn kn,0)|]
| Dtype (r,_,_) -> [|r|]
| Dterm (r,_,_) -> [|r|]
| Dfix (rv,_,_) -> rv
@@ -322,7 +319,7 @@ let declared_refs = function
let compute_deps_decl = function
| Dind (kn,ind) ->
(* Todo Later : avoid dependencies when Extract Inductive *)
- ind_iter_references add_needed add_needed add_needed kn ind
+ ind_iter_references add_needed add_needed add_needed (mind_of_kn kn) ind
| Dtype (r,ids,t) ->
if not (is_custom r) then type_iter_references add_needed t
| Dterm (r,u,t) ->
diff --git a/plugins/extraction/ocaml.ml b/plugins/extraction/ocaml.ml
index 93fc2c2dc..107b5368f 100644
--- a/plugins/extraction/ocaml.ml
+++ b/plugins/extraction/ocaml.ml
@@ -122,7 +122,7 @@ let find_projections = function Record l -> l | _ -> raise NoRecord
let kn_sig =
let specif = MPfile (dirpath_of_string "Coq.Init.Specif") in
- make_kn specif empty_dirpath (mk_label "sig")
+ make_mind specif empty_dirpath (mk_label "sig")
let rec pp_type par vl t =
let rec pp_rec par = function
@@ -380,7 +380,7 @@ let pp_Dfix (rv,c,t) =
let pp_equiv param_list name = function
| NoEquiv, _ -> mt ()
| Equiv kn, i ->
- str " = " ++ pp_parameters param_list ++ pp_global Type (IndRef (kn,i))
+ str " = " ++ pp_parameters param_list ++ pp_global Type (IndRef (mind_of_kn kn,i))
| RenEquiv ren, _ ->
str " = " ++ pp_parameters param_list ++ str (ren^".") ++ name
@@ -408,7 +408,7 @@ let pp_logical_ind packet =
fnl ()
let pp_singleton kn packet =
- let name = pp_global Type (IndRef (kn,0)) in
+ let name = pp_global Type (IndRef (mind_of_kn kn,0)) in
let l = rename_tvars keywords packet.ip_vars in
hov 2 (str "type " ++ pp_parameters l ++ name ++ str " =" ++ spc () ++
pp_type false l (List.hd packet.ip_types.(0)) ++ fnl () ++
@@ -416,7 +416,7 @@ let pp_singleton kn packet =
pr_id packet.ip_consnames.(0)))
let pp_record kn projs ip_equiv packet =
- let name = pp_global Type (IndRef (kn,0)) in
+ let name = pp_global Type (IndRef (mind_of_kn kn,0)) in
let projnames = List.map (pp_global Term) projs in
let l = List.combine projnames packet.ip_types.(0) in
let pl = rename_tvars keywords packet.ip_vars in
@@ -438,20 +438,20 @@ let pp_ind co kn ind =
let init= ref (str "type ") in
let names =
Array.mapi (fun i p -> if p.ip_logical then mt () else
- pp_global Type (IndRef (kn,i)))
+ pp_global Type (IndRef (mind_of_kn kn,i)))
ind.ind_packets
in
let cnames =
Array.mapi
(fun i p -> if p.ip_logical then [||] else
- Array.mapi (fun j _ -> pp_global Cons (ConstructRef ((kn,i),j+1)))
+ Array.mapi (fun j _ -> pp_global Cons (ConstructRef ((mind_of_kn kn,i),j+1)))
p.ip_types)
ind.ind_packets
in
let rec pp i =
if i >= Array.length ind.ind_packets then mt ()
else
- let ip = (kn,i) in
+ let ip = (mind_of_kn kn,i) in
let ip_equiv = ind.ind_equiv, i in
let p = ind.ind_packets.(i) in
if is_custom (IndRef ip) then pp (i+1)
@@ -601,12 +601,12 @@ and pp_module_type ol = function
let name = pp_modname (MPbound mbid) in
let def = pp_module_type None mt' in
str "functor (" ++ name ++ str ":" ++ typ ++ str ") ->" ++ fnl () ++ def
- | MTsig (msid, sign) ->
+ | MTsig (mp1, sign) ->
let tvm = top_visible_mp () in
- let mp = match ol with None -> MPself msid | Some l -> MPdot (tvm,l) in
+ let mp = match ol with None -> mp1 | Some l -> MPdot (tvm,l) in
(* References in [sign] are in short form (relative to [msid]).
In push_visible, [msid-->mp] is added to the current subst. *)
- push_visible mp (Some msid);
+ push_visible mp (Some mp1);
let l = map_succeed pp_specif sign in
pop_visible ();
str "sig " ++ fnl () ++
@@ -684,9 +684,9 @@ and pp_module_expr ol = function
str "functor (" ++ name ++ str ":" ++ typ ++ str ") ->" ++ fnl () ++ def
| MEapply (me, me') ->
pp_module_expr None me ++ str "(" ++ pp_module_expr None me' ++ str ")"
- | MEstruct (msid, sel) ->
+ | MEstruct (mp, sel) ->
let tvm = top_visible_mp () in
- let mp = match ol with None -> MPself msid | Some l -> MPdot (tvm,l) in
+ let mp = match ol with None -> mp | Some l -> MPdot (tvm,l) in
(* No need to update the subst with [Some msid] below : names are
already in long form (see [subst_structure] in [Extract_env]). *)
push_visible mp None;
diff --git a/plugins/extraction/table.ml b/plugins/extraction/table.ml
index 9451188aa..043bf0fe7 100644
--- a/plugins/extraction/table.ml
+++ b/plugins/extraction/table.ml
@@ -31,13 +31,13 @@ let occur_kn_in_ref kn = function
let modpath_of_r = function
| ConstRef kn -> con_modpath kn
| IndRef (kn,_)
- | ConstructRef ((kn,_),_) -> modpath kn
+ | ConstructRef ((kn,_),_) -> mind_modpath kn
| VarRef _ -> assert false
let label_of_r = function
| ConstRef kn -> con_label kn
| IndRef (kn,_)
- | ConstructRef ((kn,_),_) -> label kn
+ | ConstructRef ((kn,_),_) -> mind_label kn
| VarRef _ -> assert false
let rec base_mp = function
@@ -93,15 +93,34 @@ let rec parse_labels ll = function
let labels_of_mp mp = parse_labels [] mp
+let rec parse_labels2 ll mp1 = function
+ | mp when mp1=mp -> mp,ll
+ | MPdot (mp,l) -> parse_labels (l::ll) mp
+ | mp -> mp,ll
+
let labels_of_ref r =
- let mp,_,l =
+ let mp_top = fst(Lib.current_prefix()) in
+ let mp,_,l =
+ match r with
+ ConstRef con -> repr_con con
+ | IndRef (kn,_)
+ | ConstructRef ((kn,_),_) -> repr_mind kn
+ | VarRef _ -> assert false
+ in
+ parse_labels2 [l] mp_top mp
+
+
+
+
+let labels_of_ref2 r =
+ let mp1,_,l =
match r with
ConstRef con -> repr_con con
| IndRef (kn,_)
- | ConstructRef ((kn,_),_) -> repr_kn kn
+ | ConstructRef ((kn,_),_) -> repr_mind kn
| VarRef _ -> assert false
- in
- parse_labels [l] mp
+ in mp1,l
+
let rec add_labels_mp mp = function
| [] -> mp
@@ -127,10 +146,10 @@ let lookup_type kn = Cmap.find kn !types
(*s Inductives table. *)
-let inductives = ref (KNmap.empty : (mutual_inductive_body * ml_ind) KNmap.t)
-let init_inductives () = inductives := KNmap.empty
-let add_ind kn mib ml_ind = inductives := KNmap.add kn (mib,ml_ind) !inductives
-let lookup_ind kn = KNmap.find kn !inductives
+let inductives = ref (Mindmap.empty : (mutual_inductive_body * ml_ind) Mindmap.t)
+let init_inductives () = inductives := Mindmap.empty
+let add_ind kn mib ml_ind = inductives := Mindmap.add kn (mib,ml_ind) !inductives
+let lookup_ind kn = Mindmap.find kn !inductives
(*s Recursors table. *)
@@ -138,7 +157,7 @@ let recursors = ref Cset.empty
let init_recursors () = recursors := Cset.empty
let add_recursors env kn =
- let make_kn id = make_con (modpath kn) empty_dirpath (label_of_id id) in
+ let make_kn id = make_con (mind_modpath kn) empty_dirpath (label_of_id id) in
let mib = Environ.lookup_mind kn env in
Array.iter
(fun mip ->
@@ -305,7 +324,13 @@ let error_record r =
let check_loaded_modfile mp = match base_mp mp with
| MPfile dp -> if not (Library.library_is_loaded dp) then
- err (str ("Please load library "^(string_of_dirpath dp^" first.")))
+ let mp1 = (fst(Lib.current_prefix())) in
+ let rec find_prefix = function
+ |MPfile dp1 when dp=dp1 -> ()
+ |MPdot(mp,_) -> find_prefix mp
+ |MPbound(_) -> ()
+ | _ -> err (str ("Please load library "^(string_of_dirpath dp^" first.")))
+ in find_prefix mp1
| _ -> ()
let info_file f =
@@ -450,7 +475,7 @@ let (inline_extraction,_) =
load_function = (fun _ (_,(b,l)) -> add_inline_entries b l);
classify_function = (fun o -> Substitute o);
subst_function =
- (fun (_,s,(b,l)) -> (b,(List.map (fun x -> fst (subst_global s x)) l)))
+ (fun (s,(b,l)) -> (b,(List.map (fun x -> fst (subst_global s x)) l)))
}
let _ = declare_summary "Extraction Inline"
@@ -528,7 +553,7 @@ let (blacklist_extraction,_) =
cache_function = (fun (_,l) -> add_blacklist_entries l);
load_function = (fun _ (_,l) -> add_blacklist_entries l);
classify_function = (fun o -> Libobject.Keep o);
- subst_function = (fun (_,_,x) -> x)
+ subst_function = (fun (_,x) -> x)
}
let _ = declare_summary "Extraction Blacklist"
@@ -585,7 +610,7 @@ let (in_customs,_) =
load_function = (fun _ (_,(r,ids,s)) -> add_custom r ids s);
classify_function = (fun o -> Substitute o);
subst_function =
- (fun (_,s,(r,ids,str)) -> (fst (subst_global s r), ids, str))
+ (fun (s,(r,ids,str)) -> (fst (subst_global s r), ids, str))
}
let _ = declare_summary "ML extractions"
diff --git a/plugins/extraction/table.mli b/plugins/extraction/table.mli
index 42ed6eef0..512ecca74 100644
--- a/plugins/extraction/table.mli
+++ b/plugins/extraction/table.mli
@@ -38,7 +38,7 @@ val info_file : string -> unit
(*s utilities about [module_path] and [kernel_names] and [global_reference] *)
-val occur_kn_in_ref : kernel_name -> global_reference -> bool
+val occur_kn_in_ref : mutual_inductive -> global_reference -> bool
val modpath_of_r : global_reference -> module_path
val label_of_r : global_reference -> label
val current_toplevel : unit -> module_path
@@ -56,6 +56,7 @@ val common_prefix_from_list : module_path -> module_path list -> module_path
val add_labels_mp : module_path -> label list -> module_path
val get_nth_label_mp : int -> module_path -> label
val labels_of_ref : global_reference -> module_path * label list
+val labels_of_ref2 : global_reference -> module_path * label
(*s Some table-related operations *)
@@ -65,10 +66,10 @@ val lookup_term : constant -> ml_decl
val add_type : constant -> ml_schema -> unit
val lookup_type : constant -> ml_schema
-val add_ind : kernel_name -> mutual_inductive_body -> ml_ind -> unit
-val lookup_ind : kernel_name -> mutual_inductive_body * ml_ind
+val add_ind : mutual_inductive -> mutual_inductive_body -> ml_ind -> unit
+val lookup_ind : mutual_inductive -> mutual_inductive_body * ml_ind
-val add_recursors : Environ.env -> kernel_name -> unit
+val add_recursors : Environ.env -> mutual_inductive -> unit
val is_recursor : global_reference -> bool
val add_projection : int -> constant -> unit
diff --git a/plugins/field/field.ml4 b/plugins/field/field.ml4
index 7ac5f4e89..91c0ca21e 100644
--- a/plugins/field/field.ml4
+++ b/plugins/field/field.ml4
@@ -60,7 +60,7 @@ let _ =
let load_addfield _ = ()
let cache_addfield (_,(typ,th)) = th_tab := Gmap.add typ th !th_tab
-let subst_addfield (_,subst,(typ,th as obj)) =
+let subst_addfield (subst,(typ,th as obj)) =
let typ' = subst_mps subst typ in
let th' = subst_mps subst th in
if typ' == typ && th' == th then obj else
diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml
index 600a9123b..19a884323 100644
--- a/plugins/funind/indfun_common.ml
+++ b/plugins/funind/indfun_common.ml
@@ -317,9 +317,9 @@ let cache_Function (_,finfos) =
let load_Function _ = cache_Function
let open_Function _ = cache_Function
-let subst_Function (_,subst,finfos) =
+let subst_Function (subst,finfos) =
let do_subst_con c = fst (Mod_subst.subst_con subst c)
- and do_subst_ind (kn,i) = (Mod_subst.subst_kn subst kn,i)
+ and do_subst_ind (kn,i) = (Mod_subst.subst_ind subst kn,i)
in
let function_constant' = do_subst_con finfos.function_constant in
let graph_ind' = do_subst_ind finfos.graph_ind in
diff --git a/plugins/interface/centaur.ml4 b/plugins/interface/centaur.ml4
index e7084fbb0..07f32b6d4 100644
--- a/plugins/interface/centaur.ml4
+++ b/plugins/interface/centaur.ml4
@@ -416,7 +416,7 @@ let inspect n =
| (sp,kn), "MUTUALINDUCTIVE" ->
add_search2 (Nametab.locate (qualid_of_path sp))
(Pretyping.Default.understand Evd.empty (Global.env())
- (RRef(dummy_loc, IndRef(kn,0))))
+ (RRef(dummy_loc, IndRef((mind_of_kn kn),0))))
| _ -> failwith ("unexpected value 1 for "^
(string_of_id (basename (fst oname)))))
| _ -> failwith "unexpected value")
diff --git a/plugins/interface/name_to_ast.ml b/plugins/interface/name_to_ast.ml
index ef61a8202..142116ade 100644
--- a/plugins/interface/name_to_ast.ml
+++ b/plugins/interface/name_to_ast.ml
@@ -176,10 +176,11 @@ let inductive_to_ast_list sp =
let leaf_entry_to_ast_list ((sp,kn),lobj) =
let tag = object_tag lobj in
+ let con = constant_of_kn kn in
match tag with
| "VARIABLE" -> variable_to_ast_list (basename sp)
- | "CONSTANT" -> constant_to_ast_list (constant_of_kn kn)
- | "INDUCTIVE" -> inductive_to_ast_list kn
+ | "CONSTANT" -> constant_to_ast_list con
+ | "INDUCTIVE" -> inductive_to_ast_list (mind_of_kn kn)
| s ->
errorlabstrm
"print" (str ("printing of unrecognized object " ^
diff --git a/plugins/interface/name_to_ast.mli b/plugins/interface/name_to_ast.mli
index a532604f5..db2555550 100644
--- a/plugins/interface/name_to_ast.mli
+++ b/plugins/interface/name_to_ast.mli
@@ -2,4 +2,4 @@ val name_to_ast : Libnames.reference -> Vernacexpr.vernac_expr;;
val inductive_to_ast_list : Names.mutual_inductive -> Vernacexpr.vernac_expr list;;
val constant_to_ast_list : Names.constant -> Vernacexpr.vernac_expr list;;
val variable_to_ast_list : Names.variable -> Vernacexpr.vernac_expr list;;
-val leaf_entry_to_ast_list : (Libnames.full_path * Names.mutual_inductive) * Libobject.obj -> Vernacexpr.vernac_expr list;;
+val leaf_entry_to_ast_list : (Libnames.full_path * Names.kernel_name) * Libobject.obj -> Vernacexpr.vernac_expr list;;
diff --git a/plugins/ring/ring.ml b/plugins/ring/ring.ml
index 6ee12ebcb..c0b5542ee 100644
--- a/plugins/ring/ring.ml
+++ b/plugins/ring/ring.ml
@@ -258,7 +258,7 @@ let subst_theory subst th =
}
-let subst_th (_,subst,(c,th as obj)) =
+let subst_th (subst,(c,th as obj)) =
let c' = subst_mps subst c in
let th' = subst_theory subst th in
if c' == c && th' == th then obj else
diff --git a/plugins/setoid_ring/newring.ml4 b/plugins/setoid_ring/newring.ml4
index a930fee17..e85c12c59 100644
--- a/plugins/setoid_ring/newring.ml4
+++ b/plugins/setoid_ring/newring.ml4
@@ -400,7 +400,7 @@ let add_entry (sp,_kn) e =
from_name := Spmap.add sp e !from_name
-let subst_th (_,subst,th) =
+let subst_th (subst,th) =
let c' = subst_mps subst th.ring_carrier in
let eq' = subst_mps subst th.ring_req in
let set' = subst_mps subst th.ring_setoid in
@@ -980,7 +980,7 @@ let add_field_entry (sp,_kn) e =
field_from_relation := Cmap.add e.field_req e !field_from_relation;
field_from_name := Spmap.add sp e !field_from_name
-let subst_th (_,subst,th) =
+let subst_th (subst,th) =
let c' = subst_mps subst th.field_carrier in
let eq' = subst_mps subst th.field_req in
let thm1' = subst_mps subst th.field_ok in
diff --git a/plugins/subtac/subtac_cases.ml b/plugins/subtac/subtac_cases.ml
index d54bbee4e..4eb05df74 100644
--- a/plugins/subtac/subtac_cases.ml
+++ b/plugins/subtac/subtac_cases.ml
@@ -484,7 +484,7 @@ let check_and_adjust_constructor env ind cstrs = function
| PatCstr (loc,((_,i) as cstr),args,alias) as pat ->
(* Check it is constructor of the right type *)
let ind' = inductive_of_constructor cstr in
- if Closure.mind_equiv env ind' ind then
+ if Names.eq_ind ind' ind then
(* Check the constructor has the right number of args *)
let ci = cstrs.(i-1) in
let nb_args_constr = ci.cs_nargs in
diff --git a/plugins/subtac/subtac_obligations.ml b/plugins/subtac/subtac_obligations.ml
index 76a8a0322..f617c1008 100644
--- a/plugins/subtac/subtac_obligations.ml
+++ b/plugins/subtac/subtac_obligations.ml
@@ -155,7 +155,7 @@ let cache (_, (infos, tac)) =
let load (_, (_, tac)) =
set_default_tactic tac
-let subst (_, s, (infos, tac)) =
+let subst (s, (infos, tac)) =
(infos, Tacinterp.subst_tactic s tac)
let (input,output) =
diff --git a/plugins/syntax/ascii_syntax.ml b/plugins/syntax/ascii_syntax.ml
index f60abaf85..19473dfa6 100644
--- a/plugins/syntax/ascii_syntax.ml
+++ b/plugins/syntax/ascii_syntax.ml
@@ -21,7 +21,7 @@ open Bigint
exception Non_closed_ascii
let make_dir l = make_dirpath (List.map id_of_string (List.rev l))
-let make_kn dir id = Libnames.encode_kn (make_dir dir) (id_of_string id)
+let make_kn dir id = Libnames.encode_mind (make_dir dir) (id_of_string id)
let make_path dir id = Libnames.make_path (make_dir dir) (id_of_string id)
let ascii_module = ["Coq";"Strings";"Ascii"]
diff --git a/plugins/syntax/numbers_syntax.ml b/plugins/syntax/numbers_syntax.ml
index e58618219..21b7115b6 100644
--- a/plugins/syntax/numbers_syntax.ml
+++ b/plugins/syntax/numbers_syntax.ml
@@ -24,7 +24,7 @@ let make_path dir id = Libnames.make_path (make_dir dir) (Names.id_of_string id)
(* copied on g_zsyntax.ml, where it is said to be a temporary hack*)
(* takes a path an identifier in the form of a string list and a string,
returns a kernel_name *)
-let make_kn dir id = Libnames.encode_kn (make_dir dir) (Names.id_of_string id)
+let make_kn dir id = Libnames.encode_mind (make_dir dir) (Names.id_of_string id)
(* int31 stuff *)
@@ -50,9 +50,10 @@ let zn2z_WW = ConstructRef ((zn2z_id "zn2z",0),2)
let bigN_module = ["Coq"; "Numbers"; "Natural"; "BigN"; "BigN" ]
let bigN_path = make_path (bigN_module@["BigN"]) "t"
(* big ugly hack *)
-let bigN_id id = (Obj.magic ((Names.MPdot ((Names.MPfile (make_dir bigN_module)),
- Names.mk_label "BigN")),
- [], Names.id_of_string id) : Names.kernel_name)
+let bigN_id id = let kn =
+ ((Names.MPdot ((Names.MPfile (make_dir bigN_module)),
+ Names.mk_label "BigN")),
+ [], Names.id_of_string id) in (Obj.magic (kn,kn) : Names.mutual_inductive)
let bigN_scope = "bigN_scope"
(* number of inlined level of bigN (actually the level 0 to n_inlined-1 are inlined) *)
@@ -81,9 +82,9 @@ let bigN_constructor =
let bigZ_module = ["Coq"; "Numbers"; "Integer"; "BigZ"; "BigZ" ]
let bigZ_path = make_path (bigZ_module@["BigZ"]) "t"
(* big ugly hack bis *)
-let bigZ_id id = (Obj.magic ((Names.MPdot ((Names.MPfile (make_dir bigZ_module)),
+let bigZ_id id = let kn = ((Names.MPdot ((Names.MPfile (make_dir bigZ_module)),
Names.mk_label "BigZ")),
- [], Names.id_of_string id) : Names.kernel_name)
+ [], Names.id_of_string id) in (Obj.magic (kn,kn): Names.mutual_inductive)
let bigZ_scope = "bigZ_scope"
let bigZ_pos = ConstructRef ((bigZ_id "t_",0),1)
diff --git a/plugins/syntax/z_syntax.ml b/plugins/syntax/z_syntax.ml
index a10c76013..f6afd080f 100644
--- a/plugins/syntax/z_syntax.ml
+++ b/plugins/syntax/z_syntax.ml
@@ -31,7 +31,7 @@ let make_path dir id = Libnames.make_path (make_dir dir) (id_of_string id)
let positive_path = make_path positive_module "positive"
(* TODO: temporary hack *)
-let make_kn dir id = Libnames.encode_kn dir id
+let make_kn dir id = Libnames.encode_mind dir id
let positive_kn =
make_kn (make_dir positive_module) (id_of_string "positive")
diff --git a/plugins/xml/cic2acic.ml b/plugins/xml/cic2acic.ml
index 5bb7635b9..7706a3eb5 100644
--- a/plugins/xml/cic2acic.ml
+++ b/plugins/xml/cic2acic.ml
@@ -80,7 +80,7 @@ let get_uri_of_var v pvars =
type tag =
Constant of Names.constant
- | Inductive of Names.kernel_name
+ | Inductive of Names.mutual_inductive
| Variable of Names.kernel_name
;;
@@ -165,7 +165,7 @@ let token_list_of_kernel_name tag =
N.id_of_label (N.con_label con),
Lib.remove_section_part (LN.ConstRef con)
| Inductive kn ->
- N.id_of_label (N.label kn),
+ N.id_of_label (N.mind_label kn),
Lib.remove_section_part (LN.IndRef (kn,0))
in
token_list_of_path dir id (etag_of_tag tag)
diff --git a/plugins/xml/xmlcommand.ml b/plugins/xml/xmlcommand.ml
index a46500b89..294f5154d 100644
--- a/plugins/xml/xmlcommand.ml
+++ b/plugins/xml/xmlcommand.ml
@@ -629,7 +629,7 @@ let _ =
let _ =
Declare.set_xml_declare_inductive
(function (isrecord,(sp,kn)) ->
- print false (Libnames.IndRef (kn,0)) (kind_of_inductive isrecord kn)
+ print false (Libnames.IndRef (Names.mind_of_kn kn,0)) (kind_of_inductive isrecord (Names.mind_of_kn kn))
xml_library_root)
;;
diff --git a/pretyping/cases.ml b/pretyping/cases.ml
index 899fb64e1..9406a57d7 100644
--- a/pretyping/cases.ml
+++ b/pretyping/cases.ml
@@ -488,7 +488,7 @@ let check_and_adjust_constructor env ind cstrs = function
| PatCstr (loc,((_,i) as cstr),args,alias) as pat ->
(* Check it is constructor of the right type *)
let ind' = inductive_of_constructor cstr in
- if Closure.mind_equiv env ind' ind then
+ if eq_ind ind' ind then
(* Check the constructor has the right number of args *)
let ci = cstrs.(i-1) in
let nb_args_constr = ci.cs_nargs in
diff --git a/pretyping/classops.ml b/pretyping/classops.ml
index 35ffda0a1..7ec77f7dc 100644
--- a/pretyping/classops.ml
+++ b/pretyping/classops.ml
@@ -156,7 +156,7 @@ let subst_cl_typ subst ct = match ct with
if kn' == kn then ct else
fst (find_class_type (Global.env()) Evd.empty t)
| CL_IND (kn,i) ->
- let kn' = subst_kn subst kn in
+ let kn' = subst_ind subst kn in
if kn' == kn then ct else
CL_IND (kn',i)
@@ -355,7 +355,7 @@ let load_coercion i (_,(coe,stre,isid,cls,clt,ps)) =
let cache_coercion o =
load_coercion 1 o
-let subst_coercion (_,subst,(coe,stre,isid,cls,clt,ps as obj)) =
+let subst_coercion (subst,(coe,stre,isid,cls,clt,ps as obj)) =
let coe' = subst_coe_typ subst coe in
let cls' = subst_cl_typ subst cls in
let clt' = subst_cl_typ subst clt in
diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml
index f9c872f9e..16520916c 100644
--- a/pretyping/detyping.ml
+++ b/pretyping/detyping.ml
@@ -70,7 +70,7 @@ module PrintingCasesMake =
type t = inductive * int array
let encode = Test.encode
let subst subst ((kn,i), ints as obj) =
- let kn' = subst_kn subst kn in
+ let kn' = subst_ind subst kn in
if kn' == kn then obj else
(kn',i), ints
let printer (ind,_) = pr_global_env Idset.empty (IndRef ind)
@@ -566,7 +566,7 @@ let rec subst_cases_pattern subst pat =
match pat with
| PatVar _ -> pat
| PatCstr (loc,((kn,i),j),cpl,n) ->
- let kn' = subst_kn subst kn
+ let kn' = subst_ind subst kn
and cpl' = list_smartmap (subst_cases_pattern subst) cpl in
if kn' == kn && cpl' == cpl then pat else
PatCstr (loc,((kn',i),j),cpl',n)
@@ -610,7 +610,7 @@ let rec subst_rawconstr subst raw =
let (n,topt) = x in
let topt' = Option.smartmap
(fun (loc,(sp,i),x,y as t) ->
- let sp' = subst_kn subst sp in
+ let sp' = subst_ind subst sp in
if sp == sp' then t else (loc,(sp',i),x,y)) topt in
if a == a' && topt == topt' then y else (a',(n,topt'))) rl
and branches' = list_smartmap
diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml
index b6e697e4d..10f4db77a 100644
--- a/pretyping/evarconv.ml
+++ b/pretyping/evarconv.ml
@@ -428,12 +428,12 @@ and evar_eqappr_x env evd pbty (term1,l1 as appr1) (term2,l2 as appr2) =
evar_conv_x (push_rel (n,None,c) env) i pbty c'1 c'2)]
| Ind sp1, Ind sp2 ->
- if sp1=sp2 then
+ if eq_ind sp1 sp2 then
ise_list2 evd (fun i -> evar_conv_x env i CONV) l1 l2
else (evd, false)
| Construct sp1, Construct sp2 ->
- if sp1=sp2 then
+ if eq_constructor sp1 sp2 then
ise_list2 evd (fun i -> evar_conv_x env i CONV) l1 l2
else (evd, false)
diff --git a/pretyping/indrec.ml b/pretyping/indrec.ml
index eed795cdc..f134614cb 100644
--- a/pretyping/indrec.ml
+++ b/pretyping/indrec.ml
@@ -574,15 +574,16 @@ let make_elimination_ident id s = add_suffix id (elimination_suffix s)
let lookup_eliminator ind_sp s =
let kn,i = ind_sp in
- let mp,dp,l = repr_kn kn in
+ let mp,dp,l = repr_mind kn in
let ind_id = (Global.lookup_mind kn).mind_packets.(i).mind_typename in
let id = add_suffix ind_id (elimination_suffix s) in
(* Try first to get an eliminator defined in the same section as the *)
(* inductive type *)
try
- let cst = make_con mp dp (label_of_id id) in
+ let cst =Global.constant_of_delta
+ (make_con mp dp (label_of_id id)) in
let _ = Global.lookup_constant cst in
- mkConst cst
+ mkConst cst
with Not_found ->
(* Then try to get a user-defined eliminator in some other places *)
(* using short name (e.g. for "eq_rec") *)
diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml
index bfe1522f9..1f196f43e 100644
--- a/pretyping/inductiveops.ml
+++ b/pretyping/inductiveops.ml
@@ -471,5 +471,5 @@ let control_only_guard env c =
iter env c
let subst_inductive subst (kn,i as ind) =
- let kn' = Mod_subst.subst_kn subst kn in
+ let kn' = Mod_subst.subst_ind subst kn in
if kn == kn' then ind else (kn',i)
diff --git a/pretyping/pattern.ml b/pretyping/pattern.ml
index be37e6531..ec4a1260c 100644
--- a/pretyping/pattern.ml
+++ b/pretyping/pattern.ml
@@ -101,7 +101,7 @@ let rec pattern_of_constr t =
| Prod (na,c,b) -> PProd (na,pattern_of_constr c,pattern_of_constr b)
| Lambda (na,c,b) -> PLambda (na,pattern_of_constr c,pattern_of_constr b)
| App (f,a) -> PApp (pattern_of_constr f,Array.map pattern_of_constr a)
- | Const sp -> PRef (ConstRef sp)
+ | Const sp -> PRef (ConstRef (constant_of_kn(canonical_con sp)))
| Ind sp -> PRef (IndRef sp)
| Construct sp -> PRef (ConstructRef sp)
| Evar (n,ctxt) -> PEvar (n,Array.map pattern_of_constr ctxt)
@@ -212,8 +212,9 @@ let rec pat_of_raw metas vars = function
with Not_found -> PVar id)
| RPatVar (_,(false,n)) ->
metas := n::!metas; PMeta (Some n)
- | RRef (_,r) ->
- PRef r
+ | RRef (_,ConstRef con) ->
+ PRef (ConstRef (constant_of_kn(canonical_con con)))
+ | RRef (_,r) -> PRef r
(* Hack pour ne pas réécrire une interprétation complčte des patterns*)
| RApp (_, RPatVar (_,(true,n)), cl) ->
metas := n::!metas; PSoApp (n, List.map (pat_of_raw metas vars) cl)
@@ -295,5 +296,5 @@ and pat_of_raw_branch loc metas vars ind brs i =
let pattern_of_rawconstr c =
let metas = ref [] in
- let p = pat_of_raw metas [] c in
+ let p = pat_of_raw metas [] c in
(!metas,p)
diff --git a/pretyping/recordops.ml b/pretyping/recordops.ml
index 057b26394..13f2aef26 100644
--- a/pretyping/recordops.ml
+++ b/pretyping/recordops.ml
@@ -52,8 +52,8 @@ let load_structure i (_,(ind,id,kl,projs)) =
let cache_structure o =
load_structure 1 o
-let subst_structure (_,subst,((kn,i),id,kl,projs as obj)) =
- let kn' = subst_kn subst kn in
+let subst_structure (subst,((kn,i),id,kl,projs as obj)) =
+ let kn' = subst_ind subst kn in
let projs' =
(* invariant: struc.s_PROJ is an evaluable reference. Thus we can take *)
(* the first component of subst_con. *)
@@ -136,7 +136,7 @@ let (in_method,out_method) =
{ (default_object "RECMETHODS") with
load_function = (fun _ -> load_method);
cache_function = load_method;
- subst_function = (fun (_,s,(ty,id)) -> Mod_subst.subst_mps s ty,subst_id s id);
+ subst_function = (fun (s,(ty,id)) -> Mod_subst.subst_mps s ty,subst_id s id);
classify_function = (fun x -> Substitute x)
}
@@ -252,7 +252,7 @@ let open_canonical_structure i (_,o) =
let cache_canonical_structure o =
open_canonical_structure 1 o
-let subst_canonical_structure (_,subst,(cst,ind as obj)) =
+let subst_canonical_structure (subst,(cst,ind as obj)) =
(* invariant: cst is an evaluable reference. Thus we can take *)
(* the first component of subst_con. *)
let cst' = fst (subst_con subst cst) in
diff --git a/pretyping/termops.ml b/pretyping/termops.ml
index f0a7ce4c8..3b6d13d47 100644
--- a/pretyping/termops.ml
+++ b/pretyping/termops.ml
@@ -65,9 +65,9 @@ let rec pr_constr c = match kind_of_term c with
(str"Evar#" ++ int e ++ str"{" ++
prlist_with_sep spc pr_constr (Array.to_list l) ++str"}")
| Const c -> str"Cst(" ++ pr_con c ++ str")"
- | Ind (sp,i) -> str"Ind(" ++ pr_path sp ++ str"," ++ int i ++ str")"
+ | Ind (sp,i) -> str"Ind(" ++ pr_mind sp ++ str"," ++ int i ++ str")"
| Construct ((sp,i),j) ->
- str"Constr(" ++ pr_path sp ++ str"," ++ int i ++ str"," ++ int j ++ str")"
+ str"Constr(" ++ pr_mind sp ++ str"," ++ int i ++ str"," ++ int j ++ str")"
| Case (ci,p,c,bl) -> v 0
(hv 0 (str"<"++pr_constr p++str">"++ cut() ++ str"Case " ++
pr_constr c ++ str"of") ++ cut() ++
@@ -729,7 +729,7 @@ let hdchar env c =
lowercase_first_char (id_of_label (con_label kn))
| Ind ((kn,i) as x) ->
if i=0 then
- lowercase_first_char (id_of_label (label kn))
+ lowercase_first_char (id_of_label (mind_label kn))
else
lowercase_first_char (basename_of_global (IndRef x))
| Construct ((sp,i) as x) ->
@@ -825,15 +825,22 @@ let names_of_rel_context env =
(**** Globality of identifiers *)
-let rec is_imported_modpath = function
- | MPfile dp -> true
- | p -> false
+let rec is_imported_modpath mp =
+ let current_mp,_ = Lib.current_prefix() in
+ match mp with
+ | MPfile dp ->
+ let rec find_prefix = function
+ |MPfile dp1 -> not (dp1=dp)
+ |MPdot(mp,_) -> find_prefix mp
+ |MPbound(_) -> false
+ in find_prefix current_mp
+ | p -> false
let is_imported_ref = function
| VarRef _ -> false
| IndRef (kn,_)
| ConstructRef ((kn,_),_) ->
- let (mp,_,_) = repr_kn kn in is_imported_modpath mp
+ let (mp,_,_) = repr_mind kn in is_imported_modpath mp
| ConstRef kn ->
let (mp,_,_) = repr_con kn in is_imported_modpath mp
diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml
index 9d3522573..539821403 100644
--- a/pretyping/typeclasses.ml
+++ b/pretyping/typeclasses.ml
@@ -111,7 +111,7 @@ let load_class (_, cl) =
let cache_class = load_class
-let subst_class (_,subst,cl) =
+let subst_class (subst,cl) =
let do_subst_con c = fst (Mod_subst.subst_con subst c)
and do_subst c = Mod_subst.subst_mps subst c
and do_subst_gr gr = fst (subst_global subst gr) in
@@ -187,8 +187,9 @@ let load_instance (_,inst) =
instances := Gmap.add inst.is_class insts !instances
let cache_instance = load_instance
-let subst_instance (_,subst,inst) =
- { inst with
+
+let subst_instance (subst,inst) =
+ { inst with
is_class = fst (subst_global subst inst.is_class);
is_impl = fst (Mod_subst.subst_con subst inst.is_impl) }
diff --git a/proofs/redexpr.ml b/proofs/redexpr.ml
index d12b5f8e5..de1e3d2bd 100644
--- a/proofs/redexpr.ml
+++ b/proofs/redexpr.ml
@@ -51,7 +51,7 @@ let cache_strategy (_,str) =
(fun (lev,ql) -> List.iter (fun q -> set_strategy_one q lev) ql)
str
-let subst_strategy (_,subs,(local,obj)) =
+let subst_strategy (subs,(local,obj)) =
local,
list_smartmap
(fun (k,ql as entry) ->
diff --git a/tactics/auto.ml b/tactics/auto.ml
index 9cdbc6627..fafc0b592 100644
--- a/tactics/auto.ml
+++ b/tactics/auto.ml
@@ -89,28 +89,30 @@ let insert v l =
type stored_data = pri_auto_tactic
-type search_entry = stored_data list * stored_data list * stored_data Btermdn.t
+module Bounded_net = Btermdn.Make(struct
+ type t = stored_data
+ let compare = Pervasives.compare
+ end)
-let empty_se = ([],[],Btermdn.create ())
+type search_entry = stored_data list * stored_data list * Bounded_net.t
+
+let empty_se = ([],[],Bounded_net.create ())
let add_tac pat t st (l,l',dn) =
match pat with
| None -> if not (List.mem t l) then (insert t l, l', dn) else (l, l', dn)
- | Some pat -> if not (List.mem t l') then (l, insert t l', Btermdn.add st dn (pat,t)) else (l, l', dn)
+ | Some pat -> if not (List.mem t l') then (l, insert t l', Bounded_net.add st dn (pat,t)) else (l, l', dn)
let rebuild_dn st (l,l',dn) =
- (l, l', List.fold_left (fun dn t -> Btermdn.add (Some st) dn (Option.get t.pat, t))
- (Btermdn.create ()) l')
+ (l, l', List.fold_left (fun dn t -> Bounded_net.add (Some st) dn (Option.get t.pat, t))
+ (Bounded_net.create ()) l')
let lookup_tacs (hdc,c) st (l,l',dn) =
- let l' = List.map snd (Btermdn.lookup st dn c) in
+ let l' = List.map snd (Bounded_net.lookup st dn c) in
let sl' = Sort.list pri_order l' in
Sort.merge pri_order l sl'
-module Constr_map = Map.Make(struct
- type t = global_reference
- let compare = Pervasives.compare
- end)
+module Constr_map = Map.Make(RefOrdered)
let is_transparent_gr (ids, csts) = function
| VarRef id -> Idpred.mem id ids
@@ -389,14 +391,8 @@ let forward_subst_tactic =
let set_extern_subst_tactic f = forward_subst_tactic := f
-let subst_autohint (_,subst,(local,name,hintlist as obj)) =
+let subst_autohint (subst,(local,name,hintlist as obj)) =
let trans_clenv clenv = Clenv.subst_clenv subst clenv in
- let trans_data data code =
- { data with
- pat = Option.smartmap (subst_pattern subst) data.pat ;
- code = code ;
- }
- in
let subst_key gr =
let (lab'', elab') = subst_global subst gr in
let gr' =
@@ -409,29 +405,49 @@ let subst_autohint (_,subst,(local,name,hintlist as obj)) =
let data' = match data.code with
| Res_pf (c, clenv) ->
let c' = subst_mps subst c in
- if c==c' then data else
- trans_data data (Res_pf (c', trans_clenv clenv))
+ let clenv' = trans_clenv clenv in
+ let pat' = Option.smartmap (subst_pattern subst) data.pat in
+ if c==c' && clenv'==clenv && pat'==data.pat then data else
+ {data with
+ pat=pat';
+ code=Res_pf (c', clenv')}
| ERes_pf (c, clenv) ->
let c' = subst_mps subst c in
- if c==c' then data else
- trans_data data (ERes_pf (c', trans_clenv clenv))
+ let clenv' = trans_clenv clenv in
+ let pat' = Option.smartmap (subst_pattern subst) data.pat in
+ if c==c' && clenv'==clenv && pat'==data.pat then data else
+ {data with
+ pat=pat';
+ code=ERes_pf (c', clenv')}
| Give_exact c ->
let c' = subst_mps subst c in
- if c==c' then data else
- trans_data data (Give_exact c')
+ let pat' = Option.smartmap (subst_pattern subst) data.pat in
+ if c==c' && pat'==data.pat then data else
+ {data with
+ pat=pat';
+ code=(Give_exact c')}
| Res_pf_THEN_trivial_fail (c, clenv) ->
let c' = subst_mps subst c in
- if c==c' then data else
- let code' = Res_pf_THEN_trivial_fail (c', trans_clenv clenv) in
- trans_data data code'
+ let clenv' = trans_clenv clenv in
+ let pat' = Option.smartmap (subst_pattern subst) data.pat in
+ if c==c' && clenv'==clenv && pat'==data.pat then data else
+ {data with
+ pat=pat';
+ code=Res_pf_THEN_trivial_fail (c',clenv')}
| Unfold_nth ref ->
let ref' = subst_evaluable_reference subst ref in
- if ref==ref' then data else
- trans_data data (Unfold_nth ref')
+ let pat' = Option.smartmap (subst_pattern subst) data.pat in
+ if ref==ref' && pat'==data.pat then data else
+ {data with
+ pat=pat';
+ code=(Unfold_nth ref')}
| Extern tac ->
let tac' = !forward_subst_tactic subst tac in
- if tac==tac' then data else
- trans_data data (Extern tac')
+ let pat' = Option.smartmap (subst_pattern subst) data.pat in
+ if tac==tac' && pat'==data.pat then data else
+ {data with
+ pat=pat';
+ code=(Extern tac')}
in
if k' == k && data' == data then hint else
(k',data')
diff --git a/tactics/auto.mli b/tactics/auto.mli
index 1910bd139..83ad60bc3 100644
--- a/tactics/auto.mli
+++ b/tactics/auto.mli
@@ -42,7 +42,7 @@ type pri_auto_tactic = {
type stored_data = pri_auto_tactic
-type search_entry = stored_data list * stored_data list * stored_data Btermdn.t
+type search_entry
(* The head may not be bound. *)
diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml
index 6abe8938c..db0bbd867 100644
--- a/tactics/autorewrite.ml
+++ b/tactics/autorewrite.ml
@@ -215,7 +215,8 @@ let cache_hintrewrite (_,(rbase,lrl)) =
let lrl = HintDN.map (fun (i,h) -> (i + max, h)) lrl in
rewtab:=Stringmap.add rbase (HintDN.union lrl base) !rewtab
-let subst_hintrewrite (_,subst,(rbase,list as node)) =
+
+let subst_hintrewrite (subst,(rbase,list as node)) =
let list' = HintDN.subst subst list in
if list' == list then node else
(rbase,list')
diff --git a/tactics/btermdn.ml b/tactics/btermdn.ml
index e28a9e3c3..198ee7859 100644
--- a/tactics/btermdn.ml
+++ b/tactics/btermdn.ml
@@ -20,71 +20,133 @@ open Libnames
let dnet_depth = ref 8
-let bounded_constr_pat_discr_st st (t,depth) =
- if depth = 0 then
- None
- else
- match constr_pat_discr_st st t with
- | None -> None
- | Some (c,l) -> Some(c,List.map (fun c -> (c,depth-1)) l)
-let bounded_constr_val_discr_st st (t,depth) =
- if depth = 0 then
- Dn.Nothing
- else
- match constr_val_discr_st st t with
- | Dn.Label (c,l) -> Dn.Label(c,List.map (fun c -> (c,depth-1)) l)
- | Dn.Nothing -> Dn.Nothing
- | Dn.Everything -> Dn.Everything
+module Make =
+ functor (Z : Map.OrderedType) ->
+struct
+ module Term_dn = Termdn.Make(Z)
-let bounded_constr_pat_discr (t,depth) =
- if depth = 0 then
- None
- else
- match constr_pat_discr t with
- | None -> None
- | Some (c,l) -> Some(c,List.map (fun c -> (c,depth-1)) l)
+ module X = struct
+ type t = constr_pattern*int
+ let compare = Pervasives.compare
+ end
-let bounded_constr_val_discr (t,depth) =
- if depth = 0 then
- Dn.Nothing
- else
- match constr_val_discr t with
- | Dn.Label (c,l) -> Dn.Label(c,List.map (fun c -> (c,depth-1)) l)
- | Dn.Nothing -> Dn.Nothing
- | Dn.Everything -> Dn.Everything
+ module Y = struct
+ type t = Term_dn.term_label
+ let compare x y =
+ let make_name n =
+ match n with
+ | Term_dn.GRLabel(ConstRef con) ->
+ Term_dn.GRLabel(ConstRef(constant_of_kn(canonical_con con)))
+ | Term_dn.GRLabel(IndRef (kn,i)) ->
+ Term_dn.GRLabel(IndRef(mind_of_kn(canonical_mind kn),i))
+ | Term_dn.GRLabel(ConstructRef ((kn,i),j ))->
+ Term_dn.GRLabel(ConstructRef((mind_of_kn(canonical_mind kn),i),j))
+ | k -> k
+ in
+ Pervasives.compare (make_name x) (make_name y)
+ end
+
+ module Dn = Dn.Make(X)(Y)(Z)
+
+ type t = Dn.t
-type 'a t = (term_label,constr_pattern * int,'a) Dn.t
+ let create = Dn.create
-let create = Dn.create
+ let decomp =
+ let rec decrec acc c = match kind_of_term c with
+ | App (f,l) -> decrec (Array.fold_right (fun a l -> a::l) l acc) f
+ | Cast (c1,_,_) -> decrec acc c1
+ | _ -> (c,acc)
+ in
+ decrec []
-let add = function
- | None ->
- (fun dn (c,v) ->
- Dn.add dn bounded_constr_pat_discr ((c,!dnet_depth),v))
- | Some st ->
- (fun dn (c,v) ->
- Dn.add dn (bounded_constr_pat_discr_st st) ((c,!dnet_depth),v))
+ let constr_val_discr t =
+ let c, l = decomp t in
+ match kind_of_term c with
+ | Ind ind_sp -> Dn.Label(Term_dn.GRLabel (IndRef ind_sp),l)
+ | Construct cstr_sp -> Dn.Label(Term_dn.GRLabel (ConstructRef cstr_sp),l)
+ | Var id -> Dn.Label(Term_dn.GRLabel (VarRef id),l)
+ | Const _ -> Dn.Everything
+ | _ -> Dn.Nothing
-let rmv = function
- | None ->
- (fun dn (c,v) ->
- Dn.rmv dn bounded_constr_pat_discr ((c,!dnet_depth),v))
- | Some st ->
- (fun dn (c,v) ->
- Dn.rmv dn (bounded_constr_pat_discr_st st) ((c,!dnet_depth),v))
+let constr_val_discr_st (idpred,cpred) t =
+ let c, l = decomp t in
+ match kind_of_term c with
+ | Const c -> if Cpred.mem c cpred then Dn.Everything else Dn.Label(Term_dn.GRLabel (ConstRef c),l)
+ | Ind ind_sp -> Dn.Label(Term_dn.GRLabel (IndRef ind_sp),l)
+ | Construct cstr_sp -> Dn.Label(Term_dn.GRLabel (ConstructRef cstr_sp),l)
+ | Var id when not (Idpred.mem id idpred) -> Dn.Label(Term_dn.GRLabel (VarRef id),l)
+ | Prod (n, d, c) -> Dn.Label(Term_dn.ProdLabel, [d; c])
+ | Lambda (n, d, c) -> Dn.Label(Term_dn.LambdaLabel, [d; c] @ l)
+ | Sort s -> Dn.Label(Term_dn.SortLabel (Some s), [])
+ | Evar _ -> Dn.Everything
+ | _ -> Dn.Nothing
-let lookup = function
- | None ->
- (fun dn t ->
- List.map
- (fun ((c,_),v) -> (c,v))
- (Dn.lookup dn bounded_constr_val_discr (t,!dnet_depth)))
- | Some st ->
- (fun dn t ->
- List.map
- (fun ((c,_),v) -> (c,v))
- (Dn.lookup dn (bounded_constr_val_discr_st st) (t,!dnet_depth)))
-
-let app f dn = Dn.app (fun ((c,_),v) -> f(c,v)) dn
+ let bounded_constr_pat_discr_st st (t,depth) =
+ if depth = 0 then
+ None
+ else
+ match Term_dn.constr_pat_discr_st st t with
+ | None -> None
+ | Some (c,l) -> Some(c,List.map (fun c -> (c,depth-1)) l)
+
+ let bounded_constr_val_discr_st st (t,depth) =
+ if depth = 0 then
+ Dn.Nothing
+ else
+ match constr_val_discr_st st t with
+ | Dn.Label (c,l) -> Dn.Label(c,List.map (fun c -> (c,depth-1)) l)
+ | Dn.Nothing -> Dn.Nothing
+ | Dn.Everything -> Dn.Everything
+ let bounded_constr_pat_discr (t,depth) =
+ if depth = 0 then
+ None
+ else
+ match Term_dn.constr_pat_discr t with
+ | None -> None
+ | Some (c,l) -> Some(c,List.map (fun c -> (c,depth-1)) l)
+
+ let bounded_constr_val_discr (t,depth) =
+ if depth = 0 then
+ Dn.Nothing
+ else
+ match constr_val_discr t with
+ | Dn.Label (c,l) -> Dn.Label(c,List.map (fun c -> (c,depth-1)) l)
+ | Dn.Nothing -> Dn.Nothing
+ | Dn.Everything -> Dn.Everything
+
+
+ let add = function
+ | None ->
+ (fun dn (c,v) ->
+ Dn.add dn bounded_constr_pat_discr ((c,!dnet_depth),v))
+ | Some st ->
+ (fun dn (c,v) ->
+ Dn.add dn (bounded_constr_pat_discr_st st) ((c,!dnet_depth),v))
+
+ let rmv = function
+ | None ->
+ (fun dn (c,v) ->
+ Dn.rmv dn bounded_constr_pat_discr ((c,!dnet_depth),v))
+ | Some st ->
+ (fun dn (c,v) ->
+ Dn.rmv dn (bounded_constr_pat_discr_st st) ((c,!dnet_depth),v))
+
+ let lookup = function
+ | None ->
+ (fun dn t ->
+ List.map
+ (fun ((c,_),v) -> (c,v))
+ (Dn.lookup dn bounded_constr_val_discr (t,!dnet_depth)))
+ | Some st ->
+ (fun dn t ->
+ List.map
+ (fun ((c,_),v) -> (c,v))
+ (Dn.lookup dn (bounded_constr_val_discr_st st) (t,!dnet_depth)))
+
+ let app f dn = Dn.app (fun ((c,_),v) -> f(c,v)) dn
+
+end
+
diff --git a/tactics/btermdn.mli b/tactics/btermdn.mli
index b41ecbf77..ebded23ac 100644
--- a/tactics/btermdn.mli
+++ b/tactics/btermdn.mli
@@ -15,15 +15,19 @@ open Names
(*i*)
(* Discrimination nets with bounded depth. *)
+module Make :
+ functor (Z : Map.OrderedType) ->
+sig
+ type t
-type 'a t
+ val create : unit -> t
-val create : unit -> 'a t
-
-val add : transparent_state option -> 'a t -> (constr_pattern * 'a) -> 'a t
-val rmv : transparent_state option -> 'a t -> (constr_pattern * 'a) -> 'a t
-
-val lookup : transparent_state option -> 'a t -> constr -> (constr_pattern * 'a) list
-val app : ((constr_pattern * 'a) -> unit) -> 'a t -> unit
+ val add : transparent_state option -> t -> (constr_pattern * Z.t) -> t
+ val rmv : transparent_state option -> t -> (constr_pattern * Z.t) -> t
+ val lookup : transparent_state option -> t -> constr -> (constr_pattern * Z.t) list
+ val app : ((constr_pattern * Z.t) -> unit) -> t -> unit
+end
+
val dnet_depth : int ref
+
diff --git a/tactics/decl_interp.mli b/tactics/decl_interp.mli
index 08d97646e..bd0859382 100644
--- a/tactics/decl_interp.mli
+++ b/tactics/decl_interp.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id:$ *)
+(* $Id$ *)
open Tacinterp
open Decl_expr
diff --git a/tactics/dhyp.ml b/tactics/dhyp.ml
index 07086e05a..96d83b972 100644
--- a/tactics/dhyp.ml
+++ b/tactics/dhyp.ml
@@ -157,35 +157,44 @@ let subst_located_destructor_pattern subst = function
| ConclLocation d ->
ConclLocation (subst_destructor_pattern subst d)
+
type destructor_data = {
d_pat : located_destructor_pattern;
d_pri : int;
d_code : identifier option * glob_tactic_expr (* should be of phylum tactic *)
}
-type t = (identifier,destructor_data) Nbtermdn.t
-type frozen_t = (identifier,destructor_data) Nbtermdn.frozen_t
+module Dest_data = struct
+ type t = destructor_data
+ let compare = Pervasives.compare
+ end
+
+module Nbterm_net = Nbtermdn.Make(Dest_data)
+
+type t = identifier Nbterm_net.t
+type frozen_t = identifier Nbterm_net.frozen_t
+
+let tactab = (Nbterm_net.create () : t)
-let tactab = (Nbtermdn.create () : t)
+let lookup pat = Nbterm_net.lookup tactab pat
-let lookup pat = Nbtermdn.lookup tactab pat
-let init () = Nbtermdn.empty tactab
+let init () = Nbterm_net.empty tactab
-let freeze () = Nbtermdn.freeze tactab
-let unfreeze fs = Nbtermdn.unfreeze fs tactab
+let freeze () = Nbterm_net.freeze tactab
+let unfreeze fs = Nbterm_net.unfreeze fs tactab
let add (na,dd) =
let pat = match dd.d_pat with
| HypLocation(_,p,_) -> p.d_typ
| ConclLocation p -> p.d_typ
in
- if Nbtermdn.in_dn tactab na then begin
+ if Nbterm_net.in_dn tactab na then begin
msgnl (str "Warning [Overriding Destructor Entry " ++
str (string_of_id na) ++ str"]");
- Nbtermdn.remap tactab na (pat,dd)
+ Nbterm_net.remap tactab na (pat,dd)
end else
- Nbtermdn.add tactab (na,(pat,dd))
+ Nbterm_net.add tactab (na,(pat,dd))
let _ =
Summary.declare_summary "destruct-hyp-concl"
@@ -207,7 +216,7 @@ let cache_dd (_,(_,na,dd)) =
let classify_dd (local,_,_ as o) =
if local then Dispose else Substitute o
-let subst_dd (_,subst,(local,na,dd)) =
+let subst_dd (subst,(local,na,dd)) =
(local,na,
{ d_pat = subst_located_destructor_pattern subst dd.d_pat;
d_pri = dd.d_pri;
diff --git a/tactics/dn.ml b/tactics/dn.ml
index 359e3fe7f..8076e252c 100644
--- a/tactics/dn.ml
+++ b/tactics/dn.ml
@@ -1,100 +1,99 @@
-(* -*- compile-command: "make -C .. bin/coqtop.byte" -*- *)
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-(* $Id$ *)
-(* This file implements the basic structure of what Chet called
- ``discrimination nets''. If my understanding is right, it serves
- to associate actions (for example, tactics) with a priority to term
- patterns, so that if a hypothesis matches a pattern in the net,
- then the associated tactic is applied. Discrimination nets are used
- (only) to implement the tactics Auto, DHyp and Point.
- A discrimination net is a tries structure, that is, a tree structure
- specially conceived for searching patterns, like for example strings
- --see the file Tlm.ml in the directory lib/util--. Here the tries
- structure are used for looking for term patterns.
- This module is then used in :
- - termdn.ml (discrimination nets of terms);
- - btermdn.ml (discrimination nets of terms with bounded depth,
- used in the tactic auto);
- - nbtermdn.ml (named discrimination nets with bounded depth, used
- in the tactics Dhyp and Point).
- Eduardo (4/8/97) *)
-(* Definition of the basic structure *)
-
-type ('lbl,'pat) decompose_fun = 'pat -> ('lbl * 'pat list) option
-
-type 'res lookup_res = Label of 'res | Nothing | Everything
-
-type ('lbl,'tree) lookup_fun = 'tree -> ('lbl * 'tree list) lookup_res
-
-type ('lbl,'pat,'inf) t = (('lbl * int) option,'pat * 'inf) Tlm.t
-
-let create () = Tlm.empty
-
-(* [path_of dna pat] returns the list of nodes of the pattern [pat] read in
+module Make =
+ functor (X : Set.OrderedType) ->
+ functor (Y : Map.OrderedType) ->
+ functor (Z : Map.OrderedType) ->
+struct
+
+ module Y_tries = struct
+ type t = (Y.t * int) option
+ let compare x y =
+ match x,y with
+ None,None -> 0
+ | Some (l,n),Some (l',n') ->
+ if (Y.compare l l') = 0 && n = n' then 0
+ else Pervasives.compare x y
+ | a,b -> Pervasives.compare a b
+ end
+ module X_tries = struct
+ type t = X.t * Z.t
+ let compare (x1,x2) (y1,y2) =
+ if (X.compare x1 y1) = 0 && (Z.compare x2 y2)=0 then 0
+ else Pervasives.compare (x1,x2) (y1,y2)
+ end
+
+ module T = Tries.Make(X_tries)(Y_tries)
+
+ type decompose_fun = X.t -> (Y.t * X.t list) option
+
+ type 'res lookup_res = Label of 'res | Nothing | Everything
+
+ type 'tree lookup_fun = 'tree -> (Y.t * 'tree list) lookup_res
+
+ type t = T.t
+
+ let create () = T.empty
+
+(* [path_of dna pat] returns the list of nodes of the pattern [pat] read in
prefix ordering, [dna] is the function returning the main node of a pattern *)
-let path_of dna =
- let rec path_of_deferred = function
- | [] -> []
- | h::tl -> pathrec tl h
-
- and pathrec deferred t =
- match dna t with
- | None ->
- None :: (path_of_deferred deferred)
- | Some (lbl,[]) ->
- (Some (lbl,0))::(path_of_deferred deferred)
- | Some (lbl,(h::def_subl as v)) ->
- (Some (lbl,List.length v))::(pathrec (def_subl@deferred) h)
- in
- pathrec []
-
-let tm_of tm lbl =
- try [Tlm.map tm lbl, true] with Not_found -> []
-
-let rec skip_arg n tm =
- if n = 0 then [tm,true]
- else
- List.flatten
- (List.map
- (fun a -> match a with
- | None -> skip_arg (pred n) (Tlm.map tm a)
- | Some (lbl,m) ->
- skip_arg (pred n + m) (Tlm.map tm a))
- (Tlm.dom tm))
-
-let lookup tm dna t =
- let rec lookrec t tm =
- match dna t with
- | Nothing -> tm_of tm None
- | Label(lbl,v) ->
- tm_of tm None@
- (List.fold_left
- (fun l c ->
+ let path_of dna =
+ let rec path_of_deferred = function
+ | [] -> []
+ | h::tl -> pathrec tl h
+
+ and pathrec deferred t =
+ match dna t with
+ | None ->
+ None :: (path_of_deferred deferred)
+ | Some (lbl,[]) ->
+ (Some (lbl,0))::(path_of_deferred deferred)
+ | Some (lbl,(h::def_subl as v)) ->
+ (Some (lbl,List.length v))::(pathrec (def_subl@deferred) h)
+ in
+ pathrec []
+
+ let tm_of tm lbl =
+ try [T.map tm lbl, true] with Not_found -> []
+
+ let rec skip_arg n tm =
+ if n = 0 then [tm,true]
+ else
+ List.flatten
+ (List.map
+ (fun a -> match a with
+ | None -> skip_arg (pred n) (T.map tm a)
+ | Some (lbl,m) ->
+ skip_arg (pred n + m) (T.map tm a))
+ (T.dom tm))
+
+ let lookup tm dna t =
+ let rec lookrec t tm =
+ match dna t with
+ | Nothing -> tm_of tm None
+ | Label(lbl,v) ->
+ tm_of tm None@
+ (List.fold_left
+ (fun l c ->
List.flatten(List.map (fun (tm, b) ->
- if b then lookrec c tm
- else [tm,b]) l))
- (tm_of tm (Some(lbl,List.length v))) v)
- | Everything -> skip_arg 1 tm
- in
- List.flatten (List.map (fun (tm,b) -> Tlm.xtract tm) (lookrec t tm))
-
-let add tm dna (pat,inf) =
- let p = path_of dna pat in Tlm.add tm (p,(pat,inf))
-
-let rmv tm dna (pat,inf) =
- let p = path_of dna pat in Tlm.rmv tm (p,(pat,inf))
-
-let app f tm = Tlm.app (fun (_,p) -> f p) tm
-
+ if b then lookrec c tm
+ else [tm,b]) l))
+ (tm_of tm (Some(lbl,List.length v))) v)
+ | Everything -> skip_arg 1 tm
+ in
+ List.flatten (List.map (fun (tm,b) -> T.xtract tm) (lookrec t tm))
+
+ let add tm dna (pat,inf) =
+ let p = path_of dna pat in T.add tm (p,(pat,inf))
+
+ let rmv tm dna (pat,inf) =
+ let p = path_of dna pat in T.rmv tm (p,(pat,inf))
+
+ let app f tm = T.app (fun (_,p) -> f p) tm
+
+end
+
diff --git a/tactics/dn.mli b/tactics/dn.mli
index b4b2e6c89..3cb52a565 100644
--- a/tactics/dn.mli
+++ b/tactics/dn.mli
@@ -1,46 +1,47 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-(*i $Id$ i*)
-(* Discrimination nets. *)
-type ('lbl,'tree) decompose_fun = 'tree -> ('lbl * 'tree list) option
-type ('lbl,'pat,'inf) t (* = (('lbl * int) option,'pat * 'inf) Tlm.t *)
-val create : unit -> ('lbl,'pat,'inf) t
-(* [add t f (tree,inf)] adds a structured object [tree] together with
- the associated information [inf] to the table [t]; the function
- [f] is used to translated [tree] into its prefix decomposition: [f]
- must decompose any tree into a label characterizing its root node and
- the list of its subtree *)
-val add : ('lbl,'pat,'inf) t -> ('lbl,'pat) decompose_fun -> 'pat * 'inf
- -> ('lbl,'pat,'inf) t
+module Make :
+ functor (X : Set.OrderedType) ->
+ functor (Y : Map.OrderedType) ->
+ functor (Z : Map.OrderedType) ->
+sig
-val rmv : ('lbl,'pat,'inf) t -> ('lbl,'pat) decompose_fun -> 'pat * 'inf
- -> ('lbl,'pat,'inf) t
-
-type 'res lookup_res = Label of 'res | Nothing | Everything
-
-type ('lbl,'tree) lookup_fun = 'tree -> ('lbl * 'tree list) lookup_res
+ type decompose_fun = X.t -> (Y.t * X.t list) option
+
+ type t
+ val create : unit -> t
+
+ (* [add t f (tree,inf)] adds a structured object [tree] together with
+ the associated information [inf] to the table [t]; the function
+ [f] is used to translated [tree] into its prefix decomposition: [f]
+ must decompose any tree into a label characterizing its root node and
+ the list of its subtree *)
+
+ val add : t -> decompose_fun -> X.t * Z.t -> t
+
+ val rmv : t -> decompose_fun -> X.t * Z.t -> t
+
+ type 'res lookup_res = Label of 'res | Nothing | Everything
+
+ type 'tree lookup_fun = 'tree -> (Y.t * 'tree list) lookup_res
+
(* [lookup t f tree] looks for trees (and their associated
information) in table [t] such that the structured object [tree]
matches against them; [f] is used to translated [tree] into its
prefix decomposition: [f] must decompose any tree into a label
characterizing its root node and the list of its subtree *)
-
-val lookup : ('lbl,'pat,'inf) t -> ('lbl,'term) lookup_fun -> 'term
- -> ('pat * 'inf) list
-
-val app : (('pat * 'inf) -> unit) -> ('lbl,'pat,'inf) t -> unit
-
-val skip_arg : int -> ('lbl,'pat,'inf) t -> (('lbl,'pat,'inf) t * bool) list
+
+ val lookup : t -> 'term lookup_fun -> 'term
+ -> (X.t * Z.t) list
+
+ val app : ((X.t * Z.t) -> unit) -> t -> unit
+
+ val skip_arg : int -> t -> (t * bool) list
+
+end
diff --git a/tactics/extratactics.ml4 b/tactics/extratactics.ml4
index 6ac35317c..58ab6dfa0 100644
--- a/tactics/extratactics.ml4
+++ b/tactics/extratactics.ml4
@@ -401,7 +401,7 @@ let cache_transitivity_lemma (_,(left,lem)) =
else
transitivity_right_table := lem :: !transitivity_right_table
-let subst_transitivity_lemma (_,subst,(b,ref)) = (b,subst_mps subst ref)
+let subst_transitivity_lemma (subst,(b,ref)) = (b,subst_mps subst ref)
let (inTransitivity,_) =
declare_object {(default_object "TRANSITIVITY-STEPS") with
diff --git a/tactics/nbtermdn.ml b/tactics/nbtermdn.ml
index 3fd5236a8..7d6e1c4c9 100644
--- a/tactics/nbtermdn.ml
+++ b/tactics/nbtermdn.ml
@@ -26,34 +26,64 @@ open Libnames
(* The former comments are from Chet.
See the module dn.ml for further explanations.
Eduardo (5/8/97) *)
-
-type ('na,'a) t = {
- mutable table : ('na,constr_pattern * 'a) Gmap.t;
- mutable patterns : (Termdn.term_label option,'a Btermdn.t) Gmap.t }
-
-type ('na,'a) frozen_t =
- ('na,constr_pattern * 'a) Gmap.t
- * (Termdn.term_label option,'a Btermdn.t) Gmap.t
+module Make =
+ functor (Y:Map.OrderedType) ->
+struct
+ module X = struct
+ type t = constr_pattern*int
+ let compare = Pervasives.compare
+ end
+
+ module Term_dn = Termdn.Make(Y)
+ open Term_dn
+ module Z = struct
+ type t = Term_dn.term_label
+ let compare x y =
+ let make_name n =
+ match n with
+ | GRLabel(ConstRef con) ->
+ GRLabel(ConstRef(constant_of_kn(canonical_con con)))
+ | GRLabel(IndRef (kn,i)) ->
+ GRLabel(IndRef(mind_of_kn(canonical_mind kn),i))
+ | GRLabel(ConstructRef ((kn,i),j ))->
+ GRLabel(ConstructRef((mind_of_kn(canonical_mind kn),i),j))
+ | k -> k
+ in
+ Pervasives.compare (make_name x) (make_name y)
+ end
+
+ module Dn = Dn.Make(X)(Z)(Y)
+ module Bounded_net = Btermdn.Make(Y)
+
+
+type 'na t = {
+ mutable table : ('na,constr_pattern * Y.t) Gmap.t;
+ mutable patterns : (Term_dn.term_label option,Bounded_net.t) Gmap.t }
+
+
+type 'na frozen_t =
+ ('na,constr_pattern * Y.t) Gmap.t
+ * (Term_dn.term_label option, Bounded_net.t) Gmap.t
let create () =
{ table = Gmap.empty;
patterns = Gmap.empty }
let get_dn dnm hkey =
- try Gmap.find hkey dnm with Not_found -> Btermdn.create ()
+ try Gmap.find hkey dnm with Not_found -> Bounded_net.create ()
let add dn (na,(pat,valu)) =
- let hkey = Option.map fst (Termdn.constr_pat_discr pat) in
+ let hkey = Option.map fst (Term_dn.constr_pat_discr pat) in
dn.table <- Gmap.add na (pat,valu) dn.table;
let dnm = dn.patterns in
- dn.patterns <- Gmap.add hkey (Btermdn.add None (get_dn dnm hkey) (pat,valu)) dnm
+ dn.patterns <- Gmap.add hkey (Bounded_net.add None (get_dn dnm hkey) (pat,valu)) dnm
let rmv dn na =
let (pat,valu) = Gmap.find na dn.table in
- let hkey = Option.map fst (Termdn.constr_pat_discr pat) in
+ let hkey = Option.map fst (Term_dn.constr_pat_discr pat) in
dn.table <- Gmap.remove na dn.table;
let dnm = dn.patterns in
- dn.patterns <- Gmap.add hkey (Btermdn.rmv None (get_dn dnm hkey) (pat,valu)) dnm
+ dn.patterns <- Gmap.add hkey (Bounded_net.rmv None (get_dn dnm hkey) (pat,valu)) dnm
let in_dn dn na = Gmap.mem na dn.table
@@ -61,13 +91,43 @@ let remap ndn na (pat,valu) =
rmv ndn na;
add ndn (na,(pat,valu))
+let decomp =
+ let rec decrec acc c = match kind_of_term c with
+ | App (f,l) -> decrec (Array.fold_right (fun a l -> a::l) l acc) f
+ | Cast (c1,_,_) -> decrec acc c1
+ | _ -> (c,acc)
+ in
+ decrec []
+
+ let constr_val_discr t =
+ let c, l = decomp t in
+ match kind_of_term c with
+ | Ind ind_sp -> Dn.Label(Term_dn.GRLabel (IndRef ind_sp),l)
+ | Construct cstr_sp -> Dn.Label(Term_dn.GRLabel (ConstructRef cstr_sp),l)
+ | Var id -> Dn.Label(Term_dn.GRLabel (VarRef id),l)
+ | Const _ -> Dn.Everything
+ | _ -> Dn.Nothing
+
+let constr_val_discr_st (idpred,cpred) t =
+ let c, l = decomp t in
+ match kind_of_term c with
+ | Const c -> if Cpred.mem c cpred then Dn.Everything else Dn.Label(Term_dn.GRLabel (ConstRef c),l)
+ | Ind ind_sp -> Dn.Label(Term_dn.GRLabel (IndRef ind_sp),l)
+ | Construct cstr_sp -> Dn.Label(Term_dn.GRLabel (ConstructRef cstr_sp),l)
+ | Var id when not (Idpred.mem id idpred) -> Dn.Label(Term_dn.GRLabel (VarRef id),l)
+ | Prod (n, d, c) -> Dn.Label(Term_dn.ProdLabel, [d; c])
+ | Lambda (n, d, c) -> Dn.Label(Term_dn.LambdaLabel, [d; c] @ l)
+ | Sort s -> Dn.Label(Term_dn.SortLabel (Some s), [])
+ | Evar _ -> Dn.Everything
+ | _ -> Dn.Nothing
+
let lookup dn valu =
let hkey =
- match (Termdn.constr_val_discr valu) with
+ match (constr_val_discr valu) with
| Dn.Label(l,_) -> Some l
| _ -> None
in
- try Btermdn.lookup None (Gmap.find hkey dn.patterns) valu with Not_found -> []
+ try Bounded_net.lookup None (Gmap.find hkey dn.patterns) valu with Not_found -> []
let app f dn = Gmap.iter f dn.table
@@ -85,4 +145,4 @@ let empty dn =
let to2lists dn =
(Gmap.to_list dn.table, Gmap.to_list dn.patterns)
-
+end
diff --git a/tactics/nbtermdn.mli b/tactics/nbtermdn.mli
index e8279b76f..027ea5734 100644
--- a/tactics/nbtermdn.mli
+++ b/tactics/nbtermdn.mli
@@ -15,24 +15,37 @@ open Libnames
(*i*)
(* Named, bounded-depth, term-discrimination nets. *)
-
-type ('na,'a) t
-type ('na,'a) frozen_t
-
-val create : unit -> ('na,'a) t
-
-val add : ('na,'a) t -> ('na * (constr_pattern * 'a)) -> unit
-val rmv : ('na,'a) t -> 'na -> unit
-val in_dn : ('na,'a) t -> 'na -> bool
-val remap : ('na,'a) t -> 'na -> (constr_pattern * 'a) -> unit
-
-val lookup : ('na,'a) t -> constr -> (constr_pattern * 'a) list
-val app : ('na -> (constr_pattern * 'a) -> unit) -> ('na,'a) t -> unit
-
-val dnet_depth : int ref
-
-val freeze : ('na,'a) t -> ('na,'a) frozen_t
-val unfreeze : ('na,'a) frozen_t -> ('na,'a) t -> unit
-val empty : ('na,'a) t -> unit
-val to2lists : ('na,'a) t -> ('na * (constr_pattern * 'a)) list *
- (Termdn.term_label option * 'a Btermdn.t) list
+module Make :
+ functor (Y:Map.OrderedType) ->
+sig
+
+ module Term_dn : sig
+ type term_label =
+ | GRLabel of global_reference
+ | ProdLabel
+ | LambdaLabel
+ | SortLabel of sorts option
+ end
+
+ type 'na t
+ type 'na frozen_t
+
+ val create : unit -> 'na t
+
+ val add : 'na t -> ('na * (constr_pattern * Y.t)) -> unit
+ val rmv : 'na t -> 'na -> unit
+ val in_dn : 'na t -> 'na -> bool
+ val remap : 'na t -> 'na -> (constr_pattern * Y.t) -> unit
+
+ val lookup : 'na t -> constr -> (constr_pattern * Y.t) list
+ val app : ('na -> (constr_pattern * Y.t) -> unit) -> 'na t -> unit
+
+ val dnet_depth : int ref
+
+
+ val freeze : 'na t -> 'na frozen_t
+ val unfreeze : 'na frozen_t -> 'na t -> unit
+ val empty : 'na t -> unit
+ val to2lists : 'na t -> ('na * (constr_pattern * Y.t)) list *
+ (Term_dn.term_label option * Btermdn.Make(Y).t) list
+end
diff --git a/tactics/tacinterp.ml b/tactics/tacinterp.ml
index a7d041323..10d2a026e 100644
--- a/tactics/tacinterp.ml
+++ b/tactics/tacinterp.ml
@@ -2837,9 +2837,9 @@ let cache_md x = load_md 1 x
let subst_kind subst id =
match id with
| NewTac _ -> id
- | UpdateTac kn -> UpdateTac (Mod_subst.subst_kn subst kn)
+ | UpdateTac kn -> UpdateTac (subst_kn subst kn)
-let subst_md (_,subst,defs) =
+let subst_md (subst,defs) =
List.map (fun (id,t) -> (subst_kind subst id,subst_tactic subst t)) defs
let (inMD,outMD) =
diff --git a/tactics/termdn.ml b/tactics/termdn.ml
index 5084635e8..a2bc95044 100644
--- a/tactics/termdn.ml
+++ b/tactics/termdn.ml
@@ -20,14 +20,43 @@ open Nametab
(* Discrimination nets of terms.
See the module dn.ml for further explanations.
Eduardo (5/8/97) *)
+module Make =
+ functor (Z : Map.OrderedType) ->
+struct
-type term_label =
+ module X = struct
+ type t = constr_pattern
+ let compare = Pervasives.compare
+ end
+
+type term_label =
| GRLabel of global_reference
| ProdLabel
| LambdaLabel
| SortLabel of sorts option
-type 'a t = (term_label,constr_pattern,'a) Dn.t
+ module Y = struct
+ type t = term_label
+ let compare x y =
+ let make_name n =
+ match n with
+ | GRLabel(ConstRef con) ->
+ GRLabel(ConstRef(constant_of_kn(canonical_con con)))
+ | GRLabel(IndRef (kn,i)) ->
+ GRLabel(IndRef(mind_of_kn(canonical_mind kn),i))
+ | GRLabel(ConstructRef ((kn,i),j ))->
+ GRLabel(ConstructRef((mind_of_kn(canonical_mind kn),i),j))
+ | k -> k
+ in
+ Pervasives.compare (make_name x) (make_name y)
+ end
+
+
+ module Dn = Dn.Make(X)(Y)(Z)
+
+ type t = Dn.t
+
+ type 'a lookup_res = 'a Dn.lookup_res
(*If we have: f a b c ..., decomp gives: (f,[a;b;c;...])*)
@@ -109,3 +138,5 @@ let rmv dn st = Dn.rmv dn (constr_pat_discr_st st)
let lookup dn st t = Dn.lookup dn (constr_val_discr_st st) t
let app f dn = Dn.app f dn
+
+end
diff --git a/tactics/termdn.mli b/tactics/termdn.mli
index e60aea6b4..aea49b073 100644
--- a/tactics/termdn.mli
+++ b/tactics/termdn.mli
@@ -28,38 +28,45 @@ indicates which constants and variables can be considered as rigid.
These dnets are able to cope with existential variables as well, which match
[Everything]. *)
-type 'a t
-
-val create : unit -> 'a t
-
-(* [add t (c,a)] adds to table [t] pattern [c] associated to action [act] *)
-
-val add : 'a t -> transparent_state -> (constr_pattern * 'a) -> 'a t
-
-val rmv : 'a t -> transparent_state -> (constr_pattern * 'a) -> 'a t
-
-(* [lookup t c] looks for patterns (with their action) matching term [c] *)
-
-val lookup : 'a t -> transparent_state -> constr -> (constr_pattern * 'a) list
-
-val app : ((constr_pattern * 'a) -> unit) -> 'a t -> unit
-
-
-(*i*)
-(* These are for Nbtermdn *)
-
-type term_label =
- | GRLabel of global_reference
- | ProdLabel
- | LambdaLabel
- | SortLabel of sorts option
-
-val constr_pat_discr_st : transparent_state ->
- constr_pattern -> (term_label * constr_pattern list) option
-val constr_val_discr_st : transparent_state ->
- constr -> (term_label * constr list) Dn.lookup_res
-
-val constr_pat_discr : constr_pattern -> (term_label * constr_pattern list) option
-val constr_val_discr : constr -> (term_label * constr list) Dn.lookup_res
-
+module Make :
+ functor (Z : Map.OrderedType) ->
+sig
+
+ type t
+
+ type 'a lookup_res
+
+ val create : unit -> t
+
+ (* [add t (c,a)] adds to table [t] pattern [c] associated to action [act] *)
+
+ val add : t -> transparent_state -> (constr_pattern * Z.t) -> t
+
+ val rmv : t -> transparent_state -> (constr_pattern * Z.t) -> t
+
+ (* [lookup t c] looks for patterns (with their action) matching term [c] *)
+
+ val lookup : t -> transparent_state -> constr -> (constr_pattern * Z.t) list
+
+ val app : ((constr_pattern * Z.t) -> unit) -> t -> unit
+
+
+ (*i*)
+ (* These are for Nbtermdn *)
+
+ type term_label =
+ | GRLabel of global_reference
+ | ProdLabel
+ | LambdaLabel
+ | SortLabel of sorts option
+
+ val constr_pat_discr_st : transparent_state ->
+ constr_pattern -> (term_label * constr_pattern list) option
+ val constr_val_discr_st : transparent_state ->
+ constr -> (term_label * constr list) lookup_res
+
+ val constr_pat_discr : constr_pattern -> (term_label * constr_pattern list) option
+ val constr_val_discr : constr -> (term_label * constr list) lookup_res
+
(*i*)
+end
diff --git a/toplevel/auto_ind_decl.ml b/toplevel/auto_ind_decl.ml
index 3e025b032..5a21b9090 100644
--- a/toplevel/auto_ind_decl.ml
+++ b/toplevel/auto_ind_decl.ml
@@ -47,8 +47,8 @@ let rec deconstruct_type t =
let l,r = decompose_prod t in
(List.map (fun (_,b) -> b) (List.rev l))@[r]
-let subst_in_constr (_,subst,(ind,const)) =
- let ind' = (subst_kn subst (fst ind)),(snd ind)
+let subst_in_constr (subst,(ind,const)) =
+ let ind' = (subst_ind subst (fst ind)),(snd ind)
and const' = subst_mps subst const in
ind',const'
@@ -156,7 +156,7 @@ let make_eq_scheme sp =
| Var x -> mkVar (id_of_string ("eq_"^(string_of_id x)))
| Cast (x,_,_) -> aux (kind_of_term x) a
| App (x,newa) -> aux (kind_of_term x) newa
- | Ind (sp',i) -> if sp=sp' then mkRel(eqA-nlist-i+nb_ind-1)
+ | Ind (sp',i) -> if eq_mind sp sp' then mkRel(eqA-nlist-i+nb_ind-1)
else ( try
let eq = find_eq_scheme (sp',i)
and eqa = Array.map
@@ -167,7 +167,7 @@ let make_eq_scheme sp =
in if args = [||] then eq
else mkApp (eq,Array.append
(Array.map (fun x->lift lifti x) a) eqa)
- with Not_found -> raise(EqNotFound (string_of_kn sp'))
+ with Not_found -> raise(EqNotFound (string_of_mind sp'))
)
| Sort _ -> raise (EqUnknown "Sort" )
| Prod _ -> raise (EqUnknown "Prod" )
@@ -257,7 +257,7 @@ let make_eq_scheme sp =
(mkArrow (mkFullInd (sp,i) 1) bb);
cores.(i) <- make_one_eq i
done;
- if (string_of_mp (modpath sp ))="Coq.Init.Logic"
+ if (string_of_mp (mind_modpath sp ))="Coq.Init.Logic"
then print_string "Logic time, do nothing.\n"
else (
for i=0 to (nb_ind-1) do
@@ -451,7 +451,7 @@ let eqI ind l =
(List.map (fun (_,seq,_,_)-> mkVar seq) list_id ))
and e = try find_eq_scheme ind with
Not_found -> error
- ("The boolean equality on "^(string_of_kn (fst ind))^" is needed.");
+ ("The boolean equality on "^(string_of_mind (fst ind))^" is needed.");
in (if eA = [||] then e else mkApp(e,eA))
let compute_bl_goal ind lnamesparrec nparrec =
@@ -556,7 +556,7 @@ repeat ( apply andb_prop in z;let z1:= fresh "Z" in destruct z as [z1 z]).
| App (c,ca) -> (
match (kind_of_term c) with
| Ind (i1,i2) ->
- if(string_of_label (label i1) = "eq")
+ if(string_of_label (mind_label i1) = "eq")
then (
tclTHENSEQ ((do_replace_bl ind gls (!avoid)
nparrec (ca.(2))
diff --git a/toplevel/auto_ind_decl.mli b/toplevel/auto_ind_decl.mli
index 291ce7bb1..5e8c1a07f 100644
--- a/toplevel/auto_ind_decl.mli
+++ b/toplevel/auto_ind_decl.mli
@@ -13,7 +13,7 @@ open Mod_subst
open Sign
-val subst_in_constr : (object_name*substitution*(inductive*constr))
+val subst_in_constr : (substitution*(inductive*constr))
-> (inductive*constr)
val compute_bl_goal : inductive -> rel_context -> int -> types
diff --git a/toplevel/class.ml b/toplevel/class.ml
index 3a3588743..0166b5213 100644
--- a/toplevel/class.ml
+++ b/toplevel/class.ml
@@ -169,7 +169,7 @@ let ident_key_of_class = function
| CL_FUN -> "Funclass"
| CL_SORT -> "Sortclass"
| CL_CONST sp -> string_of_label (con_label sp)
- | CL_IND (sp,_) -> string_of_label (label sp)
+ | CL_IND (sp,_) -> string_of_label (mind_label sp)
| CL_SECVAR id -> string_of_id id
(* coercion identité *)
diff --git a/toplevel/command.ml b/toplevel/command.ml
index b50bcca77..34afbbcd5 100644
--- a/toplevel/command.ml
+++ b/toplevel/command.ml
@@ -644,19 +644,20 @@ let _ =
let declare_mutual_with_eliminations isrecord mie impls =
let names = List.map (fun e -> e.mind_entry_typename) mie.mind_entry_inds in
let (_,kn) = declare_mind isrecord mie in
+ let mind =Global.mind_of_delta (mind_of_kn kn) in
list_iter_i (fun i (indimpls, constrimpls) ->
- let ind = (kn,i) in
- Autoinstance.search_declaration (IndRef ind);
- maybe_declare_manual_implicits false (IndRef ind) indimpls;
- list_iter_i
- (fun j impls ->
+ let ind = (mind,i) in
+ Autoinstance.search_declaration (IndRef ind);
+ maybe_declare_manual_implicits false (IndRef ind) indimpls;
+ list_iter_i
+ (fun j impls ->
(* Autoinstance.search_declaration (ConstructRef (ind,j));*)
- maybe_declare_manual_implicits false (ConstructRef (ind, succ j)) impls)
- constrimpls)
+ maybe_declare_manual_implicits false (ConstructRef (ind, succ j)) impls)
+ constrimpls)
impls;
if_verbose ppnl (minductive_message names);
- if !elim_flag then declare_eliminations kn;
- kn
+ if !elim_flag then declare_eliminations mind;
+ mind
let build_mutual l finite =
let indl,ntnl = List.split l in
diff --git a/toplevel/libtypes.ml b/toplevel/libtypes.ml
index e04480700..7a065ae02 100644
--- a/toplevel/libtypes.ml
+++ b/toplevel/libtypes.ml
@@ -9,7 +9,8 @@
open Term
open Summary
open Libobject
-
+open Libnames
+open Names
(*
* Module construction
*)
@@ -18,14 +19,16 @@ let reduce c = Reductionops.head_unfold_under_prod
(Auto.Hint_db.transparent_state (Auto.searchtable_map "typeclass_instances"))
(Global.env()) Evd.empty c
-module TypeDnet = Term_dnet.Make(struct
- type t = Libnames.global_reference
- let compare = Pervasives.compare
- let subst s gr = fst (Libnames.subst_global s gr)
- let constr_of = Global.type_of_global
- end)
+module TypeDnet = Term_dnet.Make
+ (struct
+ type t = Libnames.global_reference
+ let compare = RefOrdered.compare
+ let subst s gr = fst (Libnames.subst_global s gr)
+ let constr_of = Global.type_of_global
+ end)
(struct let reduce = reduce
- let direction = false end)
+ let direction = false
+ end)
type result = Libnames.global_reference * (constr*existential_key) * Termops.subst
@@ -70,7 +73,7 @@ let (input,output) =
declare_object
{ (default_object "LIBTYPES") with
load_function = (fun _ -> load);
- subst_function = (fun (_,s,t) -> subst s t);
+ subst_function = (fun (s,t) -> subst s t);
classify_function = (fun x -> Substitute x)
}
diff --git a/toplevel/metasyntax.ml b/toplevel/metasyntax.ml
index d19a43c8d..8bb53b378 100644
--- a/toplevel/metasyntax.ml
+++ b/toplevel/metasyntax.ml
@@ -65,7 +65,7 @@ let cache_tactic_notation (_,(pa,pp)) =
let subst_tactic_parule subst (key,n,p,(d,tac)) =
(key,n,p,(d,Tacinterp.subst_tactic subst tac))
-let subst_tactic_notation (_,subst,(pa,pp)) =
+let subst_tactic_notation (subst,(pa,pp)) =
(subst_tactic_parule subst pa,pp)
let (inTacticGrammar, outTacticGrammar) =
@@ -676,7 +676,7 @@ let subst_parsing_rule subst x = x
let subst_printing_rule subst x = x
-let subst_syntax_extension (_,subst,(local,sy)) =
+let subst_syntax_extension (subst,(local,sy)) =
(local, List.map (fun (prec,ntn,gr,pp) ->
(prec,ntn, subst_parsing_rule subst gr, subst_printing_rule subst pp)) sy)
@@ -865,7 +865,7 @@ let cache_notation o =
load_notation 1 o;
open_notation 1 o
-let subst_notation (_,subst,(lc,scope,pat,b,ndf)) =
+let subst_notation (subst,(lc,scope,pat,b,ndf)) =
(lc,scope,subst_interpretation subst pat,b,ndf)
let classify_notation (local,_,_,_,_ as o) =
@@ -1024,7 +1024,7 @@ 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
+let subst_scope_command (subst,(scope,o as x)) = match o with
| ScopeClasses cl ->
let cl' = Classops.subst_cl_typ subst cl in if cl'==cl then x else
scope, ScopeClasses cl'
diff --git a/toplevel/mltop.ml4 b/toplevel/mltop.ml4
index 367ae1567..ee4370305 100644
--- a/toplevel/mltop.ml4
+++ b/toplevel/mltop.ml4
@@ -309,7 +309,7 @@ let (inMLModule,outMLModule) =
declare_object {(default_object "ML-MODULE") with
load_function = (fun _ -> cache_ml_module_object);
cache_function = cache_ml_module_object;
- subst_function = (fun (_,_,o) -> o);
+ subst_function = (fun (_,o) -> o);
classify_function = classify_ml_module_object }
let declare_ml_modules local l =
diff --git a/toplevel/search.ml b/toplevel/search.ml
index 8457ef020..8f31134c7 100644
--- a/toplevel/search.ml
+++ b/toplevel/search.ml
@@ -57,22 +57,23 @@ let gen_crible refopt (fn : global_reference -> env -> constr -> unit) =
fn (VarRef id) env typ
with Not_found -> (* we are in a section *) ())
| "CONSTANT" ->
- let cst = constant_of_kn kn in
+ let cst = Global.constant_of_delta(constant_of_kn kn) in
let typ = Typeops.type_of_constant env cst in
if refopt = None
|| head_const typ = constr_of_global (Option.get refopt)
then
fn (ConstRef cst) env typ
| "INDUCTIVE" ->
- let mib = Global.lookup_mind kn in
+ let mind = Global.mind_of_delta(mind_of_kn kn) in
+ let mib = Global.lookup_mind mind in
(match refopt with
- | Some (IndRef ((kn',tyi) as ind)) when kn=kn' ->
+ | Some (IndRef ((kn',tyi) as ind)) when eq_mind mind kn' ->
print_constructors ind fn env
(Array.length (mib.mind_packets.(tyi).mind_user_lc))
| Some _ -> ()
| _ ->
Array.iteri
- (fun i mip -> print_constructors (kn,i) fn env
+ (fun i mip -> print_constructors (mind,i) fn env
(Array.length mip.mind_user_lc)) mib.mind_packets)
| _ -> ()
in
@@ -118,7 +119,7 @@ let filter_by_module (module_list:dir_path list) (accept:bool)
with No_full_path ->
false
-let ref_eq = Libnames.encode_kn Coqlib.logic_module (id_of_string "eq"), 0
+let ref_eq = Libnames.encode_mind Coqlib.logic_module (id_of_string "eq"), 0
let c_eq = mkInd ref_eq
let gref_eq = IndRef ref_eq
diff --git a/toplevel/whelp.ml4 b/toplevel/whelp.ml4
index dac56e7d6..0d4168ec6 100644
--- a/toplevel/whelp.ml4
+++ b/toplevel/whelp.ml4
@@ -110,9 +110,9 @@ let uri_of_global ref =
| ConstRef cst ->
uri_of_repr_kn ref (repr_con cst); url_string ".con"
| IndRef (kn,i) ->
- uri_of_repr_kn ref (repr_kn kn); uri_of_ind_pointer [1;i+1]
+ uri_of_repr_kn ref (repr_mind kn); uri_of_ind_pointer [1;i+1]
| ConstructRef ((kn,i),j) ->
- uri_of_repr_kn ref (repr_kn kn); uri_of_ind_pointer [1;i+1;j]
+ uri_of_repr_kn ref (repr_mind kn); uri_of_ind_pointer [1;i+1;j]
let whelm_special = id_of_string "WHELM_ANON_VAR"