From 34fb7c5558f9f951e35eaf7a90dd88df768b1b34 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ga=C3=ABtan=20Gilbert?= Date: Wed, 14 Jun 2023 15:47:05 +0200 Subject: [PATCH] sort poly (not for inductive output type) --- checker/checker.ml | 2 + checker/values.ml | 14 +- dev/top_printers.ml | 37 +- engine/eConstr.ml | 100 +-- engine/eConstr.mli | 4 +- engine/evarutil.ml | 35 +- engine/evarutil.mli | 7 +- engine/evd.ml | 91 +-- engine/evd.mli | 15 +- engine/termops.ml | 3 + engine/termops.mli | 1 + engine/uState.ml | 567 ++++++++++++------ engine/uState.mli | 45 +- engine/univGen.ml | 72 ++- engine/univGen.mli | 37 +- engine/univNames.ml | 24 +- engine/univNames.mli | 15 +- engine/univProblem.ml | 72 ++- engine/univProblem.mli | 10 +- engine/univSubst.ml | 16 +- engine/univSubst.mli | 17 +- interp/constrexpr.mli | 19 +- interp/constrexpr_ops.ml | 9 +- interp/constrextern.ml | 21 +- interp/constrintern.ml | 62 +- interp/constrintern.mli | 2 +- interp/notation_ops.ml | 17 +- interp/notation_term.mli | 4 +- kernel/cClosure.ml | 11 +- kernel/cPrimitives.ml | 4 +- kernel/constant_typing.ml | 10 +- kernel/constant_typing.mli | 4 +- kernel/constr.ml | 4 +- kernel/conversion.ml | 17 +- kernel/conversion.mli | 2 +- kernel/cooking.ml | 9 +- kernel/cooking.mli | 2 +- kernel/declareops.ml | 2 +- kernel/declareops.mli | 2 +- kernel/discharge.ml | 14 +- kernel/environ.ml | 23 +- kernel/environ.mli | 1 + kernel/indTyping.ml | 16 +- kernel/indTyping.mli | 2 +- kernel/indtypes.ml | 9 +- kernel/indtypes.mli | 2 +- kernel/inductive.ml | 2 +- kernel/inferCumulativity.ml | 16 +- kernel/nativecode.ml | 39 +- kernel/nativevalues.ml | 36 +- kernel/safe_typing.ml | 12 +- kernel/section.ml | 15 +- kernel/section.mli | 2 +- kernel/sorts.ml | 217 ++++++- kernel/sorts.mli | 72 ++- kernel/type_errors.ml | 11 +- kernel/type_errors.mli | 6 + kernel/typeops.ml | 52 +- kernel/uGraph.ml | 16 +- kernel/uVars.ml | 300 +++++---- kernel/uVars.mli | 73 ++- kernel/univ.ml | 3 - kernel/univ.mli | 3 - kernel/vars.ml | 141 ++++- kernel/vars.mli | 7 +- kernel/vconv.ml | 12 +- kernel/vmbytecodes.ml | 2 + kernel/vmbytecodes.mli | 1 + kernel/vmbytegen.ml | 151 +++-- kernel/vmbytegen.mli | 2 +- kernel/vmemitcodes.ml | 2 +- kernel/vmsymtable.ml | 2 +- kernel/vmvalues.ml | 58 +- kernel/vmvalues.mli | 8 +- library/nametab.ml | 43 +- library/nametab.mli | 10 +- parsing/g_constr.mlg | 51 +- parsing/pcoq.ml | 5 +- parsing/pcoq.mli | 3 +- plugins/ltac/extratactics.mlg | 2 +- plugins/ltac2/tac2core.ml | 8 +- plugins/ltac2/tac2ffi.ml | 1 + plugins/ltac2/tac2ffi.mli | 1 + plugins/ring/ring.ml | 4 +- pretyping/detyping.ml | 30 +- pretyping/evarconv.ml | 13 +- pretyping/evardefine.ml | 2 +- pretyping/glob_ops.ml | 23 +- pretyping/glob_ops.mli | 2 + pretyping/glob_term.mli | 20 +- pretyping/pretyping.ml | 103 +++- pretyping/pretyping.mli | 6 +- pretyping/reductionops.ml | 27 +- pretyping/structures.ml | 2 +- pretyping/typeclasses.ml | 2 +- pretyping/typing.ml | 2 +- pretyping/vnorm.ml | 14 +- printing/ppconstr.ml | 15 +- printing/printer.ml | 69 ++- printing/printer.mli | 2 +- proofs/clenv.ml | 4 +- proofs/clenv.mli | 2 +- tactics/autorewrite.ml | 1 + tactics/cbn.ml | 23 +- tactics/eqschemes.ml | 17 +- tactics/hints.ml | 16 +- tactics/hints.mli | 4 +- tactics/tactics.ml | 3 +- .../poly-capture-global-univs/src/evilImpl.ml | 3 +- test-suite/output/PrintGrammar.out | 3 +- test-suite/output/SortQuality.out | 2 +- test-suite/success/sort_poly.v | 36 ++ vernac/auto_ind_decl.ml | 8 +- vernac/comAssumption.ml | 2 +- vernac/comHints.ml | 5 +- vernac/comInductive.ml | 12 +- vernac/comPrimitive.ml | 4 +- vernac/declare.ml | 44 +- vernac/declareUctx.ml | 33 - vernac/declareUctx.mli | 11 - vernac/declareUniv.ml | 37 +- vernac/g_vernac.mlg | 55 +- vernac/himsg.ml | 25 +- vernac/ppvernac.ml | 16 +- vernac/prettyp.ml | 2 +- vernac/prettyp.mli | 4 +- vernac/printmod.mli | 3 +- vernac/record.ml | 2 +- vernac/vernacentries.ml | 2 +- vernac/vernacexpr.mli | 4 +- 130 files changed, 2448 insertions(+), 1108 deletions(-) create mode 100644 test-suite/success/sort_poly.v delete mode 100644 vernac/declareUctx.ml delete mode 100644 vernac/declareUctx.mli diff --git a/checker/checker.ml b/checker/checker.ml index a3b943c5cac31..96e236db129e7 100644 --- a/checker/checker.ml +++ b/checker/checker.ml @@ -287,11 +287,13 @@ let explain_exn = function | CantApplyNonFunctional _ -> str"CantApplyNonFunctional" | IllFormedRecBody _ -> str"IllFormedRecBody" | IllTypedRecBody _ -> str"IllTypedRecBody" + | UnsatisfiedQConstraints _ -> str"UnsatisfiedQConstraints" | UnsatisfiedConstraints _ -> str"UnsatisfiedConstraints" | DisallowedSProp -> str"DisallowedSProp" | BadBinderRelevance _ -> str"BadBinderRelevance" | BadCaseRelevance _ -> str"BadCaseRelevance" | BadInvert -> str"BadInvert" + | UndeclaredQualities _ -> str"UndeclaredQualities" | UndeclaredUniverse _ -> str"UndeclaredUniverse" | BadVariance _ -> str "BadVariance" )) diff --git a/checker/values.ml b/checker/values.ml index c997c576b9ff9..036c1160b1a83 100644 --- a/checker/values.ml +++ b/checker/values.ml @@ -94,6 +94,12 @@ let v_level = v_tuple "level" [|Int;v_raw_level|] let v_expr = v_tuple "levelexpr" [|v_level;Int|] let v_univ = List v_expr +let v_qvar = v_sum "qvar" 0 [|[|Int|];[|String;Int|]|] + +let v_constant_quality = v_enum "constant_quality" 3 + +let v_quality = v_sum "quality" 0 [|[|v_qvar|];[|v_constant_quality|]|] + let v_cstrs = Annot ("Univ.constraints", @@ -103,16 +109,16 @@ let v_cstrs = let v_variance = v_enum "variance" 3 -let v_instance = Annot ("instance", Array v_level) -let v_abs_context = v_tuple "abstract_universe_context" [|Array v_name; v_cstrs|] +let v_instance = Annot ("instance", v_pair (Array v_quality) (Array v_level)) +let v_abs_context = v_tuple "abstract_universe_context" [|v_pair (Array v_name) (Array v_name); v_cstrs|] let v_context_set = v_tuple "universe_context_set" [|v_hset v_level;v_cstrs|] (** kernel/term *) -let v_sort = v_sum "sort" 3 (*SProp, Prop, Set*) [|[|v_univ(*Type*)|]|] +let v_sort = v_sum "sort" 3 (*SProp, Prop, Set*) [|[|v_univ(*Type*)|];[|v_qvar;v_univ(*QSort*)|]|] let v_sortfam = v_enum "sorts_family" 4 -let v_relevance = v_sum "relevance" 2 [||] +let v_relevance = v_sum "relevance" 2 [|[|v_qvar|]|] let v_binder_annot x = v_tuple "binder_annot" [|x;v_relevance|] let v_puniverses v = v_tuple "punivs" [|v;v_instance|] diff --git a/dev/top_printers.ml b/dev/top_printers.ml index 2e14a296517e2..c57171e74b834 100644 --- a/dev/top_printers.ml +++ b/dev/top_printers.ml @@ -260,10 +260,11 @@ let ppuni_level u = pp (Level.raw_pr u) let ppqvar q = pp (QVar.raw_pr q) let ppesorts s = pp (Sorts.debug_print (Evd.MiniEConstr.ESorts.unsafe_to_sorts s)) -let prlev l = UnivNames.pr_with_global_universes l +let prlev l = UnivNames.pr_level_with_global_universes l +let prqvar q = Sorts.QVar.raw_pr q let ppuniverse_set l = pp (Level.Set.pr prlev l) -let ppuniverse_instance l = pp (Instance.pr prlev l) -let ppuniverse_context l = pp (pr_universe_context prlev l) +let ppuniverse_instance l = pp (Instance.pr prqvar prlev l) +let ppuniverse_context l = pp (pr_universe_context prqvar prlev l) let ppuniverse_context_set l = pp (pr_universe_context_set prlev l) let ppuniverse_subst l = pp (UnivSubst.pr_universe_subst Level.raw_pr l) let ppuniverse_opt_subst l = pp (UState.pr_universe_opt_subst Level.raw_pr l) @@ -281,14 +282,16 @@ let ppnamedcontextval e = pp (pr_named_context env sigma (named_context_of_val e)) let ppaucontext auctx = - let nas = AbstractContext.names auctx in - let prlev l = match Level.var_index l with - | Some n -> (match nas.(n) with - | Anonymous -> prlev l + let qnas, unas = AbstractContext.names auctx in + let prgen pr var_index nas l = match var_index l with + | Some n -> (match unas.(n) with + | Anonymous -> pr l | Name id -> Id.print id) - | None -> prlev l + | None -> pr l in - pp (pr_universe_context prlev (AbstractContext.repr auctx)) + let prqvar l = prgen prqvar Sorts.QVar.var_index qnas l in + let prlev l = prgen prlev Level.var_index unas l in + pp (pr_universe_context prqvar prlev (AbstractContext.repr auctx)) let ppenv e = pp @@ -375,6 +378,9 @@ let constr_display csr = and univ_display u = incr cnt; pp (str "with " ++ int !cnt ++ str" " ++ Universe.raw_pr u ++ fnl ()) + and quality_display q = + incr cnt; pp (str "with " ++ int !cnt ++ str" " ++ Sorts.Quality.raw_pr q ++ fnl ()) + and level_display u = incr cnt; pp (str "with " ++ int !cnt ++ str" " ++ Level.raw_pr u ++ fnl ()) @@ -387,8 +393,15 @@ let constr_display csr = | QSort (q, u) -> univ_display u; Printf.sprintf "QSort(%s, %i)" (Sorts.QVar.to_string q) !cnt and universes_display l = + let qs, us = Instance.to_array l in + let qs = Array.fold_right (fun x i -> + quality_display x; + (string_of_int !cnt)^ + (if not(i="") then (" "^i) else "")) + qs "" + in Array.fold_right (fun x i -> level_display x; (string_of_int !cnt)^(if not(i="") - then (" "^i) else "")) (Instance.to_array l) "" + then (" "^i) else "")) us (if qs = "" then "" else (qs^" | ")) and name_display x = match x.binder_name with | Name id -> "Name("^(Id.to_string id)^")" @@ -527,7 +540,9 @@ let print_pure_constr csr = and box_display c = open_hovbox 1; term_display c; close_box() and universes_display u = - Array.iter (fun u -> print_space (); pp (Level.raw_pr u)) (Instance.to_array u) + let qs, us = Instance.to_array u in + Array.iter (fun u -> print_space (); pp (Sorts.Quality.raw_pr u)) qs; + Array.iter (fun u -> print_space (); pp (Level.raw_pr u)) us and sort_display = function | SProp -> print_string "SProp" diff --git a/engine/eConstr.ml b/engine/eConstr.ml index 1ad83f9c4e4e2..89b6eaedc2f44 100644 --- a/engine/eConstr.ml +++ b/engine/eConstr.ml @@ -572,31 +572,14 @@ let compare_constr sigma cmp c1 c2 = let cmp nargs c1 c2 = cmp c1 c2 in compare_gen kind eq_inst eq_sorts (eq_existential cmp) cmp 0 c1 c2 -let compare_cumulative_instances cv_pb nargs_ok variances u u' cstrs = - let open UnivProblem in - if not nargs_ok then enforce_eq_instances_univs false u u' cstrs - else - let make u = Sorts.sort_of_univ @@ Univ.Universe.make u in - CArray.fold_left3 - (fun cstrs v u u' -> - let open UVars.Variance in - match v with - | Irrelevant -> Set.add (UWeak (u,u')) cstrs - | Covariant -> - (match cv_pb with - | Conversion.CONV -> Set.add (UEq (make u, make u')) cstrs - | Conversion.CUMUL -> Set.add (ULe (make u, make u')) cstrs) - | Invariant -> - Set.add (UEq (make u, make u')) cstrs) - cstrs variances (UVars.Instance.to_array u) (UVars.Instance.to_array u') - let cmp_inductives cv_pb (mind,ind as spec) nargs u1 u2 cstrs = let open UnivProblem in match mind.Declarations.mind_variance with | None -> enforce_eq_instances_univs false u1 u2 cstrs | Some variances -> let num_param_arity = Conversion.inductive_cumulativity_arguments spec in - compare_cumulative_instances cv_pb (Int.equal num_param_arity nargs) variances u1 u2 cstrs + if not (Int.equal num_param_arity nargs) then enforce_eq_instances_univs false u1 u2 cstrs + else compare_cumulative_instances cv_pb variances u1 u2 cstrs let cmp_constructors (mind, ind, cns as spec) nargs u1 u2 cstrs = let open UnivProblem in @@ -607,8 +590,11 @@ let cmp_constructors (mind, ind, cns as spec) nargs u1 u2 cstrs = if not (Int.equal num_cnstr_args nargs) then enforce_eq_instances_univs false u1 u2 cstrs else + let qs1, us1 = UVars.Instance.to_array u1 + and qs2, us2 = UVars.Instance.to_array u2 in + let cstrs = enforce_eq_qualities qs1 qs2 cstrs in Array.fold_left2 (fun cstrs u1 u2 -> UnivProblem.(Set.add (UWeak (u1,u2)) cstrs)) - cstrs (UVars.Instance.to_array u1) (UVars.Instance.to_array u2) + cstrs us1 us2 let eq_universes env sigma cstrs cv_pb refargs l l' = if EInstance.is_empty l then (assert (EInstance.is_empty l'); true) @@ -619,7 +605,7 @@ let eq_universes env sigma cstrs cv_pb refargs l l' = let open UnivProblem in match refargs with | Some (ConstRef c, 1) when Environ.is_array_type env c -> - cstrs := compare_cumulative_instances cv_pb true [|UVars.Variance.Irrelevant|] l l' !cstrs; + cstrs := compare_cumulative_instances cv_pb [|UVars.Variance.Irrelevant|] l l' !cstrs; true | None | Some (ConstRef _, _) -> cstrs := enforce_eq_instances_univs true l l' !cstrs; true @@ -714,31 +700,77 @@ let eq_constr_universes_proj env sigma m n = let res = eq_constr' 0 m n in if res then Some !cstrs else None +let add_universes_of_instance sigma (qs,us) u = + let u = EInstance.kind sigma u in + let qs', us' = UVars.Instance.levels u in + let qs = Sorts.Quality.(Set.fold (fun q qs -> match q with + | QVar q -> Sorts.QVar.Set.add q qs + | QConstant _ -> qs) + qs' qs) + in + qs, Univ.Level.Set.union us us' + +let fold_annot_relevance f acc na = + f acc na.Context.binder_relevance + +let fold_case_under_context_relevance f acc (nas,_) = + Array.fold_left (fold_annot_relevance f) acc nas + +let fold_rec_declaration_relevance f acc (nas,_,_) = + Array.fold_left (fold_annot_relevance f) acc nas + +let fold_constr_relevance sigma f acc c = + match kind sigma c with + | Rel _ | Var _ | Meta _ | Evar _ + | Sort _ | Cast _ | App _ + | Const _ | Ind _ | Construct _ | Proj _ + | Int _ | Float _ | Array _ -> acc + + | Prod (na,_,_) | Lambda (na,_,_) | LetIn (na,_,_,_) -> + fold_annot_relevance f acc na + + | Case (ci,_u,_params,ret,_iv,_v,brs) -> + let acc = f acc ci.ci_relevance in + let acc = fold_case_under_context_relevance f acc ret in + let acc = CArray.fold_left (fold_case_under_context_relevance f) acc brs in + acc + + | Fix (_,data) + | CoFix (_,data) -> + fold_rec_declaration_relevance f acc data + +let add_relevance sigma (qs,us as v) r = + let open Sorts in + (* NB this normalizes above_prop to Relevant which makes it disappear *) + match UState.nf_relevance (Evd.evar_universe_context sigma) r with + | Irrelevant | Relevant -> v + | RelevanceVar q -> QVar.Set.add q qs, us + let universes_of_constr sigma c = let open Univ in - let open UVars in let rec aux s c = + let s = fold_constr_relevance sigma (add_relevance sigma) s c in match kind sigma c with - | Const (c, u) -> - Level.Set.fold Level.Set.add (Instance.levels (EInstance.kind sigma u)) s - | Ind ((mind,_), u) | Construct (((mind,_),_), u) -> - Level.Set.fold Level.Set.add (Instance.levels (EInstance.kind sigma u)) s - | Sort u -> - let sort = ESorts.kind sigma u in - if Sorts.is_small sort then s - else - Level.Set.fold Level.Set.add (Sorts.levels sort) s + | Const (_, u) | Ind (_, u) | Construct (_,u) -> add_universes_of_instance sigma s u + | Sort u -> begin match ESorts.kind sigma u with + | Type u -> + Util.on_snd (Level.Set.fold Level.Set.add (Universe.levels u)) s + | QSort (q, u) -> + let qs, us = s in + Sorts.QVar.Set.add q qs, Level.Set.union us (Universe.levels u) + | SProp | Prop | Set -> s + end | Evar (k, args) -> let concl = Evd.evar_concl (Evd.find_undefined sigma k) in fold sigma aux (aux s concl) c | Array (u,_,_,_) -> - let s = Level.Set.fold Level.Set.add (Instance.levels (EInstance.kind sigma u)) s in + let s = add_universes_of_instance sigma s u in fold sigma aux s c | Case (_,u,_,_,_,_,_) -> - let s = Level.Set.fold Level.Set.add (Instance.levels (EInstance.kind sigma u)) s in + let s = add_universes_of_instance sigma s u in fold sigma aux s c | _ -> fold sigma aux s c - in aux Level.Set.empty c + in aux (Sorts.QVar.Set.empty,Level.Set.empty) c open Context open Environ diff --git a/engine/eConstr.mli b/engine/eConstr.mli index 4834bda1614cd..ba6ea34f15602 100644 --- a/engine/eConstr.mli +++ b/engine/eConstr.mli @@ -333,7 +333,7 @@ val fold_with_binders : Evd.evar_map -> ('a -> 'a) -> ('a -> 'b -> t -> 'b) -> ' (** Gather the universes transitively used in the term, including in the type of evars appearing in it. *) -val universes_of_constr : Evd.evar_map -> t -> Univ.Level.Set.t +val universes_of_constr : Evd.evar_map -> t -> Sorts.QVar.Set.t * Univ.Level.Set.t (** {6 Substitutions} *) @@ -367,7 +367,7 @@ val noccur_between : Evd.evar_map -> int -> int -> t -> bool val closedn : Evd.evar_map -> int -> t -> bool val closed0 : Evd.evar_map -> t -> bool -val subst_univs_level_constr : Univ.universe_level_subst -> t -> t +val subst_univs_level_constr : UVars.full_level_subst -> t -> t val subst_instance_context : UVars.Instance.t -> rel_context -> rel_context val subst_instance_constr : UVars.Instance.t -> t -> t diff --git a/engine/evarutil.ml b/engine/evarutil.ml index d0f305ce3279a..238f6be979c18 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -35,7 +35,7 @@ let finalize ?abort_on_undefined_evars sigma f = let sigma = minimize_universes sigma in let uvars = ref Univ.Level.Set.empty in let v = f (fun c -> - let varsc = EConstr.universes_of_constr sigma c in + let _, varsc = EConstr.universes_of_constr sigma c in let c = EConstr.to_constr ?abort_on_undefined_evars sigma c in uvars := Univ.Level.Set.union !uvars varsc; c) @@ -67,12 +67,9 @@ let whd_evar = EConstr.whd_evar let nf_evar sigma c = let lsubst = Evd.universe_subst sigma in let evar_value ev = Evd.existential_opt_value0 sigma ev in - let level_value l = - UnivSubst.level_subst_of (fun l -> UnivSubst.normalize_univ_variable_opt_subst lsubst l) l - in - let sort_value s = UState.nf_sort (Evd.evar_universe_context sigma) s in - let rel_value r = UState.nf_relevance (Evd.evar_universe_context sigma) r in - EConstr.of_constr @@ UnivSubst.nf_evars_and_universes_opt_subst evar_value level_value sort_value rel_value (EConstr.Unsafe.to_constr c) + let univ_value l = UnivSubst.normalize_univ_variable_opt_subst lsubst l in + let qvar_value q = UState.nf_qvar (Evd.evar_universe_context sigma) q in + EConstr.of_constr @@ UnivSubst.nf_evars_and_universes_opt_subst evar_value qvar_value univ_value (EConstr.Unsafe.to_constr c) let j_nf_evar sigma j = { uj_val = nf_evar sigma j.uj_val; @@ -763,6 +760,12 @@ let compare_cumulative_instances cv_pb variances u u' sigma = let open UnivProblem in let cstrs = Univ.Constraints.empty in let soft = Set.empty in + let qs, us = UVars.Instance.to_array u + and qs', us' = UVars.Instance.to_array u' in + let qcstrs = enforce_eq_qualities qs qs' Set.empty in + match Evd.add_universe_constraints sigma qcstrs with + | exception UGraph.UniverseInconsistency p -> Inr p + | sigma -> let cstrs, soft = Array.fold_left3 (fun (cstrs, soft) v u u' -> let open UVars.Variance in match v with @@ -770,7 +773,7 @@ let compare_cumulative_instances cv_pb variances u u' sigma = | Covariant when cv_pb == Conversion.CUMUL -> Univ.Constraints.add (u,Univ.Le,u') cstrs, soft | Covariant | Invariant -> Univ.Constraints.add (u,Univ.Eq,u') cstrs, soft) - (cstrs,soft) variances (UVars.Instance.to_array u) (UVars.Instance.to_array u') + (cstrs,soft) variances us us' in match Evd.add_constraints sigma cstrs with | sigma -> @@ -779,11 +782,17 @@ let compare_cumulative_instances cv_pb variances u u' sigma = let compare_constructor_instances evd u u' = let open UnivProblem in - let soft = - Array.fold_left2 (fun cs u u' -> Set.add (UWeak (u,u')) cs) - Set.empty (UVars.Instance.to_array u) (UVars.Instance.to_array u') - in - Evd.add_universe_constraints evd soft + let qs, us = UVars.Instance.to_array u + and qs', us' = UVars.Instance.to_array u' in + let qcstrs = enforce_eq_qualities qs qs' Set.empty in + match Evd.add_universe_constraints evd qcstrs with + | exception UGraph.UniverseInconsistency p -> Inr p + | evd -> + let soft = + Array.fold_left2 (fun cs u u' -> Set.add (UWeak (u,u')) cs) + Set.empty us us' + in + Inl (Evd.add_universe_constraints evd soft) (** [eq_constr_univs_test ~evd ~extended_evd t u] tests equality of [t] and [u] up to existential variable instantiation and diff --git a/engine/evarutil.mli b/engine/evarutil.mli index cb8710490e1c3..eb5b3e80bcaf1 100644 --- a/engine/evarutil.mli +++ b/engine/evarutil.mli @@ -194,9 +194,12 @@ val compare_cumulative_instances : Conversion.conv_pb -> UVars.Variance.t array (evar_map, UGraph.univ_inconsistency) Util.union (** We should only compare constructors at convertible types, so this - is only an opportunity to unify universes. *) + is only an opportunity to unify universes. + + But what about qualities? +*) val compare_constructor_instances : evar_map -> - UVars.Instance.t -> UVars.Instance.t -> evar_map + UVars.Instance.t -> UVars.Instance.t -> (evar_map, UGraph.univ_inconsistency) Util.union (** {6 Unification problems} *) type unification_pb = conv_pb * env * constr * constr diff --git a/engine/evd.ml b/engine/evd.ml index 0ca8266deb025..3931defe5d7e2 100644 --- a/engine/evd.ml +++ b/engine/evd.ml @@ -10,7 +10,6 @@ open Pp open CErrors -open Sorts open Util open Names open Nameops @@ -904,8 +903,8 @@ let mkLEvar sigma (evk, args) = let evar_handler sigma = let evar_expand ev = existential_expand_value0 sigma ev in let qvar_relevant q = match UState.nf_qvar sigma.universes q with - | QSProp -> false - | QProp | QType | QVar _ -> true + | QConstant QSProp -> false + | QConstant QProp | QConstant QType | QVar _ -> true in let evar_relevant (evk, _) = match find sigma evk with | EvarInfo evi -> @@ -934,6 +933,9 @@ let existential_type0 = existential_type let add_constraints d c = { d with universes = UState.add_constraints d.universes c } +let add_quconstraints d c = + { d with universes = UState.add_quconstraints d.universes c } + let add_universe_constraints d c = { d with universes = UState.add_universe_constraints d.universes c } @@ -1071,6 +1073,8 @@ let evar_universe_context d = d.universes let universe_context_set d = UState.context_set d.universes +let full_universe_context_set d = UState.full_context_set d.universes + let to_universe_context evd = UState.context evd.universes let univ_entry ~poly evd = UState.univ_entry ~poly evd.universes @@ -1086,9 +1090,18 @@ let universe_subst evd = let merge_context_set ?loc ?(sideff=false) rigid evd uctx' = {evd with universes = UState.merge ?loc ~sideff rigid evd.universes uctx'} +let merge_full_context_set ?loc ?(sideff=false) rigid evd ctx' = + {evd with universes = UState.merge_full ?loc ~sideff rigid evd.universes ctx'} + +let merge_sort_variables ?loc ?(sideff=false) evd qs = + { evd with universes = UState.merge_sort_variables ?loc ~sideff evd.universes qs } + let with_context_set ?loc rigid evd (a, uctx) = (merge_context_set ?loc rigid evd uctx, a) +let with_full_context_set ?loc rigid d (a, ctx) = + (merge_full_context_set ?loc rigid d ctx, a) + let new_univ_level_variable ?loc ?name rigid evd = let uctx', u = UState.new_univ_variable ?loc rigid name evd.universes in ({evd with universes = uctx'}, u) @@ -1097,8 +1110,12 @@ let new_univ_variable ?loc ?name rigid evd = let uctx', u = UState.new_univ_variable ?loc rigid name evd.universes in ({evd with universes = uctx'}, Univ.Universe.make u) -let new_sort_variable ?loc ?name rigid sigma = - let (sigma, u) = new_univ_variable ?loc rigid ?name sigma in +let new_quality_variable ?loc ?name evd = + let uctx, q = UState.new_sort_variable ?loc ?name evd.universes in + {evd with universes = uctx}, q + +let new_sort_variable ?loc rigid sigma = + let (sigma, u) = new_univ_variable ?loc rigid sigma in let uctx, q = UState.new_sort_variable sigma.universes in ({ sigma with universes = uctx }, Sorts.qsort q u) @@ -1117,22 +1134,22 @@ let make_nonalgebraic_variable evd u = (****************************************) let fresh_sort_in_family ?loc ?(rigid=univ_flexible) evd s = - with_context_set ?loc rigid evd (UnivGen.fresh_sort_in_family s) + with_full_context_set ?loc rigid evd (UnivGen.fresh_sort_in_family s) let fresh_constant_instance ?loc ?(rigid=univ_flexible) env evd c = - with_context_set ?loc rigid evd (UnivGen.fresh_constant_instance env c) + with_full_context_set ?loc rigid evd (UnivGen.fresh_constant_instance env c) let fresh_inductive_instance ?loc ?(rigid=univ_flexible) env evd i = - with_context_set ?loc rigid evd (UnivGen.fresh_inductive_instance env i) + with_full_context_set ?loc rigid evd (UnivGen.fresh_inductive_instance env i) let fresh_constructor_instance ?loc ?(rigid=univ_flexible) env evd c = - with_context_set ?loc rigid evd (UnivGen.fresh_constructor_instance env c) + with_full_context_set ?loc rigid evd (UnivGen.fresh_constructor_instance env c) let fresh_array_instance ?loc ?(rigid=univ_flexible) env evd = - with_context_set ?loc rigid evd (UnivGen.fresh_array_instance env) + with_full_context_set ?loc rigid evd (UnivGen.fresh_array_instance env) let fresh_global ?loc ?(rigid=univ_flexible) ?names env evd gr = - with_context_set ?loc rigid evd (UnivGen.fresh_global_instance ?loc ?names env gr) + with_full_context_set ?loc rigid evd (UnivGen.fresh_global_instance ?loc ?names env gr) let is_sort_variable evd s = UState.is_sort_variable evd.universes s @@ -1151,33 +1168,11 @@ let universe_rigidity evd l = UnivFlexible (Univ.Level.Set.mem l (UState.algebraics uctx)) else UnivRigid -let normalize_universe evd = - let vars = UState.subst evd.universes in - let normalize = UnivSubst.normalize_universe_opt_subst vars in - normalize - let normalize_universe_instance evd l = - let vars = UState.subst evd.universes in - let normalize = UnivSubst.level_subst_of (UnivSubst.normalize_univ_variable_opt_subst vars) in - UnivSubst.subst_instance normalize l + UState.nf_instance evd.universes l let normalize_sort evars s = - match s with - | SProp | Prop | Set -> s - | Type u -> - let u' = normalize_universe evars u in - if u' == u then s else Sorts.sort_of_univ u' - | QSort (q, u) -> - begin match UState.nf_qvar evars.universes q with - | QProp -> Sorts.prop - | QSProp -> Sorts.sprop - | QType -> - let u' = normalize_universe evars u in - Sorts.sort_of_univ u' - | QVar q' -> - let u' = normalize_universe evars u in - if q' == q && u' == u then s else Sorts.qsort q' u' - end + UState.nf_sort evars.universes s (* FIXME inefficient *) let set_eq_sort env d s1 s2 = @@ -1222,6 +1217,12 @@ let check_leq evd s s' = let check_constraints evd csts = UGraph.check_constraints csts (UState.ugraph evd.universes) +let check_qconstraints evd csts = + UState.check_qconstraints evd.universes csts + +let check_quconstraints evd (qcsts,ucsts) = + check_qconstraints evd qcsts && check_constraints evd ucsts + let fix_undefined_variables evd = { evd with universes = UState.fix_undefined_variables evd.universes } @@ -1241,6 +1242,8 @@ let minimize_universes evd = let universe_of_name evd s = UState.universe_of_name evd.universes s +let quality_of_name evd s = UState.quality_of_name evd.universes s + let universe_binders evd = UState.universe_binders evd.universes let universes evd = UState.ugraph evd.universes @@ -1688,12 +1691,11 @@ module MiniEConstr = struct Some (instantiate_evar_array sigma info c args) in let lsubst = universe_subst sigma in - let level_value l = - UnivSubst.level_subst_of (fun l -> UnivSubst.normalize_univ_variable_opt_subst lsubst l) l + let univ_value l = + UnivSubst.normalize_univ_variable_opt_subst lsubst l in - let sort_value s = UState.nf_sort (evar_universe_context sigma) s in - let rel_value r = UState.nf_relevance (evar_universe_context sigma) r in - UnivSubst.nf_evars_and_universes_opt_subst evar_value level_value sort_value rel_value c + let qvar_value q = UState.nf_qvar sigma.universes q in + UnivSubst.nf_evars_and_universes_opt_subst evar_value qvar_value univ_value c let to_constr_gen sigma c = let saw_evar = ref false in @@ -1703,12 +1705,11 @@ module MiniEConstr = struct v in let lsubst = universe_subst sigma in - let level_value l = - UnivSubst.level_subst_of (fun l -> UnivSubst.normalize_univ_variable_opt_subst lsubst l) l + let univ_value l = + UnivSubst.normalize_univ_variable_opt_subst lsubst l in - let sort_value s = UState.nf_sort (evar_universe_context sigma) s in - let rel_value r = UState.nf_relevance (evar_universe_context sigma) r in - let c = UnivSubst.nf_evars_and_universes_opt_subst evar_value level_value sort_value rel_value c in + let qvar_value q = UState.nf_qvar sigma.universes q in + let c = UnivSubst.nf_evars_and_universes_opt_subst evar_value qvar_value univ_value c in let saw_evar = if not !saw_evar then false else let exception SawEvar in diff --git a/engine/evd.mli b/engine/evd.mli index 0d0576414d16e..32c56674ca05d 100644 --- a/engine/evd.mli +++ b/engine/evd.mli @@ -269,6 +269,8 @@ val is_undefined : evar_map -> Evar.t-> bool val add_constraints : evar_map -> Univ.Constraints.t -> evar_map (** Add universe constraints in an evar map. *) +val add_quconstraints : evar_map -> Sorts.QUConstraints.t -> evar_map + val undefined_map : evar_map -> undefined evar_info Evar.Map.t (** Access the undefined evar mapping directly. *) @@ -631,11 +633,13 @@ val restrict_universe_context : evar_map -> Univ.Level.Set.t -> evar_map (** Raises Not_found if not a name for a universe in this map. *) val universe_of_name : evar_map -> Id.t -> Univ.Level.t +val quality_of_name : evar_map -> Id.t -> Sorts.QVar.t val universe_binders : evar_map -> UnivNames.universe_binders val new_univ_level_variable : ?loc:Loc.t -> ?name:Id.t -> rigid -> evar_map -> evar_map * Univ.Level.t -val new_sort_variable : ?loc:Loc.t -> ?name:Id.t -> rigid -> evar_map -> evar_map * esorts +val new_quality_variable : ?loc:Loc.t -> ?name:Id.t -> evar_map -> evar_map * Sorts.QVar.t +val new_sort_variable : ?loc:Loc.t -> rigid -> evar_map -> evar_map * esorts val add_global_univ : evar_map -> Univ.Level.t -> evar_map @@ -665,9 +669,12 @@ val check_eq : evar_map -> esorts -> esorts -> bool val check_leq : evar_map -> esorts -> esorts -> bool val check_constraints : evar_map -> Univ.Constraints.t -> bool +val check_qconstraints : evar_map -> Sorts.QConstraints.t -> bool +val check_quconstraints : evar_map -> Sorts.QUConstraints.t -> bool val evar_universe_context : evar_map -> UState.t val universe_context_set : evar_map -> Univ.ContextSet.t +val full_universe_context_set : evar_map -> UnivGen.full_context_set val universe_subst : evar_map -> UnivSubst.universe_opt_subst val universes : evar_map -> UGraph.t @@ -685,8 +692,14 @@ val set_universe_context : evar_map -> UState.t -> evar_map val merge_context_set : ?loc:Loc.t -> ?sideff:bool -> rigid -> evar_map -> Univ.ContextSet.t -> evar_map +val merge_full_context_set : ?loc:Loc.t -> ?sideff:bool -> rigid -> evar_map -> UnivGen.full_context_set -> evar_map + +val merge_sort_variables : ?loc:Loc.t -> ?sideff:bool -> evar_map -> Sorts.QVar.Set.t -> evar_map + val with_context_set : ?loc:Loc.t -> rigid -> evar_map -> 'a Univ.in_universe_context_set -> evar_map * 'a +val with_full_context_set : ?loc:Loc.t -> rigid -> evar_map -> 'a UnivGen.in_full_context_set -> evar_map * 'a + val nf_univ_variables : evar_map -> evar_map val collapse_sort_variables : evar_map -> evar_map diff --git a/engine/termops.ml b/engine/termops.ml index d5b50cb4c7f84..043c082d5adf1 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -296,6 +296,9 @@ let has_no_evar sigma = with Exit -> false let pr_evd_level sigma = UState.pr_uctx_level (Evd.evar_universe_context sigma) + +let pr_evd_qvar sigma = UState.pr_uctx_qvar (Evd.evar_universe_context sigma) + let reference_of_level sigma l = UState.qualid_of_level (Evd.evar_universe_context sigma) l let pr_evar_universe_context ctx = diff --git a/engine/termops.mli b/engine/termops.mli index da70e9b621b09..5acf326945f63 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -292,6 +292,7 @@ val pr_evar_map_filter : ?with_univs:bool -> (Evar.t -> any_evar_info -> bool) - val pr_metaset : Metaset.t -> Pp.t val pr_evar_universe_context : UState.t -> Pp.t val pr_evd_level : evar_map -> Univ.Level.t -> Pp.t +val pr_evd_qvar : evar_map -> Sorts.QVar.t -> Pp.t module Internal : sig diff --git a/engine/uState.ml b/engine/uState.ml index 08bd7f2d1abfb..8d0dc0bca1095 100644 --- a/engine/uState.ml +++ b/engine/uState.ml @@ -12,20 +12,21 @@ open CErrors open Util open Names open Univ +open Sorts open UVars type universes_entry = | Monomorphic_entry of Univ.ContextSet.t | Polymorphic_entry of UVars.UContext.t -module UNameMap = Names.Id.Map +module UNameMap = Id.Map type uinfo = { uname : Id.t option; uloc : Loc.t option; } -open Sorts.Quality +open Quality let sort_inconsistency ?explain cst l r = let explain = Option.map (fun p -> UGraph.Other p) explain in @@ -33,42 +34,41 @@ let sort_inconsistency ?explain cst l r = module QState : sig type t - type elt = Sorts.QVar.t + type elt = QVar.t val empty : t - val union : fail:(t -> Sorts.Quality.t -> Sorts.Quality.t -> t) -> t -> t -> t - val add : elt -> t -> t - val repr : elt -> t -> Sorts.Quality.t - val unify_quality : fail:(unit -> t) -> Conversion.conv_pb -> Sorts.Quality.t -> Sorts.Quality.t -> t -> t + val union : fail:(t -> Quality.t -> Quality.t -> t) -> t -> t -> t + val add : check_fresh:bool -> named:bool -> elt -> t -> t + val repr : elt -> t -> Quality.t + val unify_quality : fail:(unit -> t) -> Conversion.conv_pb -> Quality.t -> Quality.t -> t -> t val is_above_prop : elt -> t -> bool + val undefined : t -> QVar.Set.t val collapse : t -> t val pr : t -> Pp.t + val of_set : QVar.Set.t -> t end = struct -module QSet = Sorts.QVar.Set -module QMap = Sorts.QVar.Map +module QSet = QVar.Set +module QMap = QVar.Map type t = { - qmap : Sorts.Quality.t option QMap.t; + named : QSet.t; + (** Named variables, may not be set to another *) + qmap : Quality.t option QMap.t; (* TODO: use a persistent union-find structure *) above : QSet.t; (** Set of quality variables known to be either in Prop or Type. If q ∈ above then it must map to None in qmap. *) } -type elt = Sorts.QVar.t +type elt = QVar.t -let empty = { qmap = QMap.empty; above = QSet.empty } - -let quality_eq a b = match a, b with - | QProp, QProp | QSProp, QSProp | QType, QType -> true - | QVar q1, QVar q2 -> Sorts.QVar.equal q1 q2 - | (QVar _ | QProp | QSProp | QType), _ -> false +let empty = { named = QSet.empty; qmap = QMap.empty; above = QSet.empty } let rec repr q m = match QMap.find q m.qmap with | None -> QVar q | Some (QVar q) -> repr q m -| Some (QProp | QSProp | QType as q) -> q +| Some (QConstant _ as q) -> q | exception Not_found -> (* let () = assert !Flags.in_debugger in *) (* FIXME *) QVar q @@ -77,47 +77,54 @@ let is_above_prop q m = QSet.mem q m.above let set q qv m = let q = repr q m in - let q = match q with QVar q -> q | QProp | QSProp | QType -> assert false in - let qv = match qv with QVar qv -> repr qv m | (QSProp | QProp | QType as qv) -> qv in + let q = match q with QVar q -> q | QConstant _ -> assert false in + let qv = match qv with QVar qv -> repr qv m | (QConstant _ as qv) -> qv in match q, qv with | q, QVar qv -> - if Sorts.QVar.equal q qv then Some m + if QVar.equal q qv then Some m + else + if QSet.mem q m.named then None else let above = if QSet.mem q m.above then QSet.add qv (QSet.remove q m.above) else m.above in - Some { qmap = QMap.add q (Some (QVar qv)) m.qmap; above } - | q, (QProp | QSProp | QType as qv) -> - if qv == QSProp && QSet.mem q m.above then None - else Some { qmap = QMap.add q (Some qv) m.qmap; above = QSet.remove q m.above } + Some { named = m.named; qmap = QMap.add q (Some (QVar qv)) m.qmap; above } + | q, (QConstant qc as qv) -> + if qc == QSProp && QSet.mem q m.above then None + else if QSet.mem q m.named then None + else + Some { named = m.named; qmap = QMap.add q (Some qv) m.qmap; above = QSet.remove q m.above } let set_above_prop q m = + (* XXX fail if named? *) let q = repr q m in - let q = match q with QVar q -> q | QProp | QSProp | QType -> assert false in - { qmap = m.qmap; above = QSet.add q m.above } + let q = match q with QVar q -> q | QConstant _ -> assert false in + { named = m.named; qmap = m.qmap; above = QSet.add q m.above } let unify_quality ~fail c q1 q2 local = match q1, q2 with -| QType, QType | QProp, QProp | QSProp, QSProp -> local -| QProp, QVar q when c == Conversion.CUMUL -> +| QConstant QType, QConstant QType +| QConstant QProp, QConstant QProp +| QConstant QSProp, QConstant QSProp -> local +| QConstant QProp, QVar q when c == Conversion.CUMUL -> set_above_prop q local -| QVar q, (QType | QProp | QSProp | QVar _ as qv) -| (QType | QProp | QSProp as qv), QVar q -> +| QVar q, (QConstant (QType | QProp | QSProp) | QVar _ as qv) +| (QConstant (QType | QProp | QSProp) as qv), QVar q -> begin match set q qv local with | Some local -> local | None -> fail () end -| (QType, (QProp | QSProp)) -> fail () -| (QProp, QType) -> +| (QConstant QType, QConstant (QProp | QSProp)) -> fail () +| (QConstant QProp, QConstant QType) -> begin match c with | CONV -> fail () | CUMUL -> local end -| (QSProp, (QType | QProp)) -> fail () -| (QProp, QSProp) -> fail () +| (QConstant QSProp, QConstant (QType | QProp)) -> fail () +| (QConstant QProp, QConstant QSProp) -> fail () let nf_quality m = function - | QSProp | QProp | QType as q -> q + | QConstant _ as q -> q | QVar q -> repr q m let union ~fail s1 s2 = @@ -127,7 +134,7 @@ let union ~fail s1 s2 = | Some q, None | None, Some q -> Some (Some q) | None, None -> Some None | Some q1, Some q2 -> - let () = if not (quality_eq q1 q2) then extra := (q1,q2) :: !extra in + let () = if not (Quality.equal q1 q2) then extra := (q1,q2) :: !extra in Some (Some q1)) s1.qmap s2.qmap in @@ -138,21 +145,33 @@ let union ~fail s1 s2 = | exception Not_found -> false in let above = QSet.filter filter @@ QSet.union s1.above s2.above in - let s = { qmap; above } in + let s = { named = QSet.union s1.named s2.named; qmap; above } in List.fold_left (fun s (q1,q2) -> let q1 = nf_quality s q1 and q2 = nf_quality s q2 in unify_quality ~fail:(fun () -> fail s q1 q2) CONV q1 q2 s) s extra -let add q m = { qmap = QMap.add q None m.qmap; above = m.above } +let add ~check_fresh ~named q m = + if check_fresh then assert (not (QMap.mem q m.qmap)); + { named = if named then QSet.add q m.named else m.named; + qmap = QMap.add q None m.qmap; + above = m.above } + +let of_set qs = + { named = QSet.empty; qmap = QMap.bind (fun _ -> None) qs; above = QSet.empty } + +(* XXX what about [above]? *) +let undefined m = + let m = QMap.filter (fun _ v -> Option.is_empty v) m.qmap in + QMap.domain m let collapse m = let map q v = match v with - | None -> Some QType + | None -> if QSet.mem q m.named then None else Some (QConstant QType) | Some _ -> v in - { qmap = QMap.mapi map m.qmap; above = QSet.empty } + { named = m.named; qmap = QMap.mapi map m.qmap; above = QSet.empty } let pr { qmap; above } = (* TODO names *) @@ -162,16 +181,16 @@ let pr { qmap; above } = if QSet.mem u above then str " >= Prop" else mt () | Some q -> - let q = Sorts.Quality.raw_pr q in + let q = Quality.raw_pr q in str " := " ++ q in - h (prlist_with_sep fnl (fun (u, v) -> Sorts.QVar.raw_pr u ++ prbody u v) (QMap.bindings qmap)) + h (prlist_with_sep fnl (fun (u, v) -> QVar.raw_pr u ++ prbody u v) (QMap.bindings qmap)) end module UPairSet = UnivMinim.UPairSet -type univ_names = UnivNames.universe_binders * uinfo Level.Map.t +type univ_names = UnivNames.universe_binders * (uinfo QVar.Map.t * uinfo Level.Map.t) (* 2nd part used to check consistency on the fly. *) type t = @@ -192,7 +211,7 @@ type t = } let empty = - { names = UNameMap.empty, Level.Map.empty; + { names = UnivNames.empty_binders, (QVar.Map.empty, Level.Map.empty); local = ContextSet.empty; seff_univs = Level.Set.empty; univ_variables = Level.Map.empty; @@ -221,14 +240,20 @@ let uname_union s t = | Some _, _ -> l | _, _ -> r) s t +let names_union ((qbind,ubind),(qrev,urev)) ((qbind',ubind'),(qrev',urev')) = + let qbind = uname_union qbind qbind' + and ubind = uname_union ubind ubind' + and qrev = QVar.Map.union (fun _ l _ -> Some l) qrev qrev' + and urev = Level.Map.lunion urev urev' in + ((qbind,ubind),(qrev,urev)) + let union uctx uctx' = if uctx == uctx' then uctx else if is_empty uctx' then uctx else let local = ContextSet.union uctx.local uctx'.local in let seff = Level.Set.union uctx.seff_univs uctx'.seff_univs in - let names = uname_union (fst uctx.names) (fst uctx'.names) in - let names_rev = Level.Map.lunion (snd uctx.names) (snd uctx'.names) in + let names = names_union uctx.names uctx'.names in let newus = Level.Set.diff (ContextSet.levels uctx'.local) (ContextSet.levels uctx.local) in let newus = Level.Set.diff newus (Level.Map.domain uctx.univ_variables) in @@ -240,9 +265,9 @@ let union uctx uctx' = if UGraph.type_in_type uctx.universes then s else CErrors.user_err Pp.(str "Could not merge universe contexts: could not unify" ++ spc() ++ - Sorts.Quality.raw_pr q1 ++ strbrk " and " ++ Sorts.Quality.raw_pr q2 ++ str ".") + Quality.raw_pr q1 ++ strbrk " and " ++ Quality.raw_pr q2 ++ str ".") in - { names = (names, names_rev); + { names; local = local; seff_univs = seff; univ_variables = @@ -261,18 +286,31 @@ let union uctx uctx' = let context_set uctx = uctx.local +let full_context_set uctx = + let us, csts = uctx.local in + (QState.undefined uctx.sort_variables, us), csts + let constraints uctx = snd uctx.local -let compute_instance_binders rbinders inst = - let map lvl = - try Name (Option.get (Level.Map.find lvl rbinders).uname) +let compute_instance_binders (qrev,urev) inst = + let qinst, uinst = Instance.to_array inst in + let qmap = function + | QVar q -> + begin try Name (Option.get (QVar.Map.find q qrev).uname) + with Option.IsNone | Not_found -> Anonymous + end + | QConstant _ -> assert false + in + let umap lvl = + try Name (Option.get (Level.Map.find lvl urev).uname) with Option.IsNone | Not_found -> Anonymous in - Array.map map (Instance.to_array inst) + Array.map qmap qinst, Array.map umap uinst let context uctx = let (_, rbinders) = uctx.names in - UContext.of_context_set (compute_instance_binders rbinders) uctx.local + let qvars = QState.undefined uctx.sort_variables in + UContext.of_context_set (compute_instance_binders rbinders) qvars uctx.local type named_universes_entry = universes_entry * UnivNames.universe_binders @@ -283,7 +321,11 @@ let univ_entry ~poly uctx = else Monomorphic_entry (context_set uctx) in entry, binders -let of_context_set local = { empty with local } +let of_context_set ((qs,us),csts) = + let sort_variables = QState.of_set qs in + { empty with + local = (us,csts); + sort_variables;} type universe_opt_subst = UnivSubst.universe_opt_subst @@ -295,27 +337,23 @@ let initial_graph uctx = uctx.initial_universes let algebraics uctx = uctx.univ_algebraic -let add_names ?loc s l (names, names_rev) = - if UNameMap.mem s names - then user_err ?loc - Pp.(str "Universe " ++ Names.Id.print s ++ str" already bound."); - (UNameMap.add s l names, Level.Map.add l { uname = Some s; uloc = loc } names_rev) - -let add_loc l loc (names, names_rev) = - match loc with - | None -> (names, names_rev) - | Some _ -> (names, Level.Map.add l { uname = None; uloc = loc } names_rev) - -let of_names (ubind,revubind) = +let of_names (ubind,(revqbind,revubind)) = + let revqbind = QVar.Map.map (fun id -> { uname = Some id; uloc = None }) revqbind in let revubind = Level.Map.map (fun id -> { uname = Some id; uloc = None }) revubind in - {empty with names = (ubind,revubind)} + {empty with names = (ubind,(revqbind,revubind))} let universe_of_name uctx s = - UNameMap.find s (fst uctx.names) + UNameMap.find s (snd (fst uctx.names)) + +let quality_of_name uctx s = + Id.Map.find s (fst (fst uctx.names)) let name_level level id uctx = - assert(not(Names.Id.Map.mem id (fst uctx.names))); - { uctx with names = (Names.Id.Map.add id level (fst uctx.names), Univ.Level.Map.add level { uname = Some id; uloc = None } (snd uctx.names)) } + let ((qbind,ubind),(qrev,urev)) = uctx.names in + assert(not(Id.Map.mem id ubind)); + let ubind = Id.Map.add id level ubind in + let urev = Level.Map.add level { uname = Some id; uloc = None } urev in + { uctx with names = ((qbind,ubind),(qrev,urev)) } let universe_binders uctx = @@ -340,40 +378,38 @@ let level_inconsistency cst l r = let mk u = Sorts.sort_of_univ @@ Universe.make u in raise (UGraph.UniverseInconsistency (cst, mk l, mk r, None)) -let subst_univs_sort normalize qnormalize s = match s with -| Sorts.Set | Sorts.Prop | Sorts.SProp -> s -| Sorts.Type u -> Sorts.sort_of_univ (UnivSubst.subst_univs_universe normalize u) -| Sorts.QSort (q, u) -> - match qnormalize q with - | QSProp -> Sorts.sprop - | QProp -> Sorts.prop - | QType -> Sorts.sort_of_univ (UnivSubst.subst_univs_universe normalize u) - | QVar q -> Sorts.qsort q (UnivSubst.subst_univs_universe normalize u) +let nf_universe uctx u = + UnivSubst.(subst_univs_universe (normalize_univ_variable_opt_subst uctx.univ_variables)) u + +let nf_level uctx u = + UnivSubst.(level_subst_of (normalize_univ_variable_opt_subst uctx.univ_variables)) u + +let nf_instance uctx u = Instance.subst_fn (nf_qvar uctx, nf_level uctx) u + +let nf_quality uctx q = Quality.subst (nf_qvar uctx) q let nf_sort uctx s = - let normalize u = UnivSubst.normalize_univ_variable_opt_subst uctx.univ_variables u in + let normalize u = nf_universe uctx u in let qnormalize q = QState.repr q uctx.sort_variables in - subst_univs_sort normalize qnormalize s + Sorts.subst_fn (qnormalize, normalize) s let nf_relevance uctx r = match r with -| Sorts.Relevant | Sorts.Irrelevant -> r -| Sorts.RelevanceVar q -> +| Relevant | Irrelevant -> r +| RelevanceVar q -> match nf_qvar uctx q with - | QSProp -> Sorts.Irrelevant - | QProp | QType -> Sorts.Relevant + | QConstant QSProp -> Sorts.Irrelevant + | QConstant QProp | QConstant QType -> Sorts.Relevant | QVar q' -> + (* XXX currently not used in nf_evars_and_universes_opt_subst + does it matter? *) if QState.is_above_prop q' uctx.sort_variables then Relevant - else if Sorts.QVar.equal q q' then r + else if QVar.equal q q' then r else Sorts.RelevanceVar q' let nf_universes uctx c = let lsubst = uctx.univ_variables in - let level_value l = - UnivSubst.level_subst_of (fun l -> UnivSubst.normalize_univ_variable_opt_subst lsubst l) l - in - let sort_value s = nf_sort uctx s in - let rel_value r = nf_relevance uctx r in - UnivSubst.nf_evars_and_universes_opt_subst (fun _ -> None) level_value sort_value rel_value c + let nf_univ u = UnivSubst.normalize_univ_variable_opt_subst lsubst u in + UnivSubst.nf_evars_and_universes_opt_subst (fun _ -> None) (nf_qvar uctx) nf_univ c type small_universe = USet | UProp | USProp @@ -386,10 +422,10 @@ type sort_classification = | UAlgebraic of Universe.t (* Arbitrary algebraic expression *) let classify s = match s with -| Sorts.Prop -> USmall UProp -| Sorts.SProp -> USmall USProp -| Sorts.Set -> USmall USet -| Sorts.Type u | Sorts.QSort (_, u) -> +| Prop -> USmall UProp +| SProp -> USmall USProp +| Set -> USmall USet +| Type u | QSort (_, u) -> if Universe.is_levels u then match Universe.level u with | None -> UMax (u, Universe.levels u) | Some u -> ULevel u @@ -410,10 +446,10 @@ let enforce_leq_up u v local = { local with local_cst = UnivSubst.enforce_leq u (Universe.make v) local.local_cst } let quality_of_sort = function -| Sorts.Set | Sorts.Type _ -> QType -| Sorts.Prop -> QProp -| Sorts.SProp -> QSProp -| Sorts.QSort (q, _) -> QVar q +| Set | Type _ -> QConstant QType +| Prop -> QConstant QProp +| SProp -> QConstant QSProp +| QSort (q, _) -> QVar q let get_constraint = function | Conversion.CONV -> Eq @@ -434,11 +470,13 @@ let process_universe_constraints uctx cstrs = let univs = uctx.universes in let vars = ref uctx.univ_variables in let normalize u = normalize_univ_variable_opt_subst !vars u in + let qnormalize sorts q = QState.repr q sorts in let normalize_sort sorts s = - let qnormalize q = QState.repr q sorts in - subst_univs_sort normalize qnormalize s + Sorts.subst_fn ((qnormalize sorts), subst_univs_universe normalize) s in let nf_constraint sorts = function + | QLeq (a, b) -> QLeq (Quality.subst (qnormalize sorts) a, Quality.subst (qnormalize sorts) b) + | QEq (a, b) -> QEq (Quality.subst (qnormalize sorts) a, Quality.subst (qnormalize sorts) b) | ULub (u, v) -> ULub (level_subst_of normalize u, level_subst_of normalize v) | UWeak (u, v) -> UWeak (level_subst_of normalize u, level_subst_of normalize v) | UEq (u, v) -> UEq (normalize_sort sorts u, normalize_sort sorts v) @@ -447,19 +485,19 @@ let process_universe_constraints uctx cstrs = let is_local l = Level.Map.mem l !vars in let equalize_small l s local = let ls = match l with - | USProp -> Sorts.sprop - | UProp -> Sorts.prop - | USet -> Sorts.set + | USProp -> sprop + | UProp -> prop + | USet -> set in if UGraph.check_eq_sort univs ls s then local else if is_uset l then match classify s with - | USmall _ -> sort_inconsistency Eq Sorts.set s + | USmall _ -> sort_inconsistency Eq set s | ULevel r -> if is_local r then let () = instantiate_variable r Universe.type0 vars in add_local (Level.set, Eq, r) local else - sort_inconsistency Eq Sorts.set s + sort_inconsistency Eq set s | UMax (u, _)| UAlgebraic u -> if univ_level_mem Level.set u then let inst = univ_level_rem Level.set u u in @@ -495,7 +533,7 @@ let process_universe_constraints uctx cstrs = else if univ_level_mem l ru then enforce_leq_up inst l local - else sort_inconsistency Eq (Sorts.sort_of_univ (Universe.make l)) (Sorts.sort_of_univ ru) + else sort_inconsistency Eq (sort_of_univ (Universe.make l)) (sort_of_univ ru) in let equalize_universes l r local = match classify l, classify r with | USmall l', (USmall _ | ULevel _ | UMax _ | UAlgebraic _) -> @@ -515,6 +553,26 @@ let process_universe_constraints uctx cstrs = let cst = nf_constraint local.local_sorts cst in if UnivProblem.is_trivial cst then local else match cst with + | QEq (a, b) -> + (* TODO sort_inconsistency should be able to handle raw + qualities instead of having to make a dummy sort *) + let mk = function + | QVar q -> qsort q Universe.type0 + | QConstant QSProp -> sprop + | QConstant QProp -> prop + | QConstant QType -> set + in + unify_quality univs CONV (mk a) (mk b) local + | QLeq (a, b) -> + (* TODO sort_inconsistency should be able to handle raw + qualities instead of having to make a dummy sort *) + let mk = function + | QVar q -> qsort q Universe.type0 + | QConstant QSProp -> sprop + | QConstant QProp -> prop + | QConstant QType -> set + in + unify_quality univs CUMUL (mk a) (mk b) local | ULe (l, r) -> let local = unify_quality univs CUMUL l r local in let l = normalize_sort local.local_sorts l in @@ -547,7 +605,7 @@ let process_universe_constraints uctx cstrs = | UMax (_, levels) -> if is_uset r' then let fold l' local = - let l = Sorts.sort_of_univ @@ Universe.make l' in + let l = sort_of_univ @@ Universe.make l' in if Level.is_set l' || is_local l' then equalize_variables false l' Level.set local else sort_inconsistency Le l r @@ -606,27 +664,6 @@ let process_universe_constraints uctx cstrs = let extra = { UnivMinim.above_prop = local.local_above_prop; UnivMinim.weak_constraints = local.local_weak } in !vars, extra, local.local_cst, local.local_sorts -let add_constraints uctx cstrs = - let univs, old_cstrs = uctx.local in - let cstrs' = Constraints.fold (fun (l,d,r) acc -> - let l = Universe.make l and r = Sorts.sort_of_univ @@ Universe.make r in - let cstr' = let open UnivProblem in - match d with - | Lt -> - ULe (Sorts.sort_of_univ @@ Universe.super l, r) - | Le -> ULe (Sorts.sort_of_univ l, r) - | Eq -> UEq (Sorts.sort_of_univ l, r) - in UnivProblem.Set.add cstr' acc) - cstrs UnivProblem.Set.empty - in - let vars, extra, cstrs', sorts = process_universe_constraints uctx cstrs' in - { uctx with - local = (univs, Constraints.union old_cstrs cstrs'); - univ_variables = vars; - universes = UGraph.merge_constraints cstrs' uctx.universes; - sort_variables = sorts; - minim_extra = extra; } - let add_universe_constraints uctx cstrs = let univs, local = uctx.local in let vars, extra, local', sorts = process_universe_constraints uctx cstrs in @@ -637,6 +674,67 @@ let add_universe_constraints uctx cstrs = sort_variables = sorts; minim_extra = extra; } +let problem_of_constraints cstrs = + Constraints.fold (fun (l,d,r) acc -> + let l = Universe.make l and r = sort_of_univ @@ Universe.make r in + let cstr' = let open UnivProblem in + match d with + | Lt -> + ULe (sort_of_univ @@ Universe.super l, r) + | Le -> ULe (sort_of_univ l, r) + | Eq -> UEq (sort_of_univ l, r) + in UnivProblem.Set.add cstr' acc) + cstrs UnivProblem.Set.empty + +let add_constraints uctx cstrs = + let cstrs = problem_of_constraints cstrs in + add_universe_constraints uctx cstrs + +let add_quconstraints uctx (qcstrs,ucstrs) = + let cstrs = problem_of_constraints ucstrs in + let cstrs = QConstraints.fold (fun (l,d,r) cstrs -> + match d with + | Equal -> UnivProblem.Set.add (QEq (l,r)) cstrs + | Leq -> UnivProblem.Set.add (QLeq (l,r)) cstrs) + qcstrs cstrs + in + add_universe_constraints uctx cstrs + +let check_qconstraints uctx csts = + Sorts.QConstraints.for_all (fun (l,k,r) -> + let l = nf_quality uctx l in + let r = nf_quality uctx r in + if Quality.equal l r then true + else match l,k,r with + | _, Equal, _ -> false + | QConstant QProp, Leq, QConstant QType -> true + | QConstant QProp, Leq, QVar q -> QState.is_above_prop q uctx.sort_variables + | _, Leq, _ -> false) + csts + +let check_universe_constraint uctx (c:UnivProblem.t) = + match c with + | QEq (a,b) -> + let a = nf_quality uctx a in + let b = nf_quality uctx b in + Quality.equal a b + | QLeq (a,b) -> + let a = nf_quality uctx a in + let b = nf_quality uctx b in + if Quality.equal a b then true + else begin match a, b with + | QConstant QProp, QConstant QType -> true + | QConstant QProp, QVar q -> QState.is_above_prop q uctx.sort_variables + | _ -> false + end + | ULe (u,v) -> UGraph.check_leq_sort uctx.universes u v + | UEq (u,v) -> UGraph.check_eq_sort uctx.universes u v + | ULub (u,v) -> UGraph.check_eq_level uctx.universes u v + | UWeak _ -> true + +let check_universe_constraints uctx csts = + UnivProblem.Set.for_all (check_universe_constraint uctx) csts + let constrain_variables diff uctx = let univs, local = uctx.local in let univs, vars, local = @@ -655,17 +753,30 @@ let constrain_variables diff uctx = { uctx with local = (univs, local); univ_variables = vars } let id_of_level uctx l = - try Some (Option.get (Level.Map.find l (snd uctx.names)).uname) + try (Level.Map.find l (snd (snd uctx.names))).uname + with Not_found -> None + +let id_of_qvar uctx l = + try (QVar.Map.find l (fst (snd uctx.names))).uname + with Not_found -> None + +let qualid_of_qvar_names (bind, (qrev,_)) l = + try Some (Libnames.qualid_of_ident (Option.get (QVar.Map.find l qrev).uname)) with Not_found | Option.IsNone -> - None + UnivNames.qualid_of_qvar bind l -let qualid_of_level_names (map, map_rev) l = - try Some (Libnames.qualid_of_ident (Option.get (Level.Map.find l map_rev).uname)) +let qualid_of_level_names (bind, (_,urev)) l = + try Some (Libnames.qualid_of_ident (Option.get (Level.Map.find l urev).uname)) with Not_found | Option.IsNone -> - UnivNames.qualid_of_level map l + UnivNames.qualid_of_level bind l let qualid_of_level uctx l = qualid_of_level_names uctx.names l +let pr_uctx_qvar_names names l = + match qualid_of_qvar_names names l with + | Some qid -> Libnames.pr_qualid qid + | None -> QVar.raw_pr l + let pr_uctx_level_names names l = match qualid_of_level_names names l with | Some qid -> Libnames.pr_qualid qid @@ -673,54 +784,88 @@ let pr_uctx_level_names names l = let pr_uctx_level uctx l = pr_uctx_level_names uctx.names l -type ('a, 'b) gen_universe_decl = { - univdecl_instance : 'a; (* Declared universes *) +let pr_uctx_qvar uctx l = pr_uctx_qvar_names uctx.names l + +type ('a, 'b, 'c) gen_universe_decl = { + univdecl_qualities : 'a; + univdecl_instance : 'b; (* Declared universes *) univdecl_extensible_instance : bool; (* Can new universes be added *) - univdecl_constraints : 'b; (* Declared constraints *) + univdecl_constraints : 'c; (* Declared constraints *) univdecl_extensible_constraints : bool (* Can new constraints be added *) } type universe_decl = - (Level.t list, Constraints.t) gen_universe_decl + (QVar.t list, Level.t list, Constraints.t) gen_universe_decl let default_univ_decl = - { univdecl_instance = []; + { univdecl_qualities = []; + univdecl_instance = []; univdecl_extensible_instance = true; univdecl_constraints = Constraints.empty; univdecl_extensible_constraints = true } -let pr_error_unbound_universes left names = +let pr_error_unbound_universes quals univs names = let open Pp in - let n = Level.Set.cardinal left in + let nqs = QVar.Set.cardinal quals in + let prqvar q = + let info = QVar.Map.find_opt q (fst (snd names)) in + h (pr_uctx_qvar_names names q ++ (match info with + | None | Some {uloc=None} -> mt () + | Some {uloc=Some loc} -> spc() ++ str"(" ++ Loc.pr loc ++ str")")) + in + let nus = Level.Set.cardinal univs in let prlev u = - let info = Level.Map.find_opt u (snd names) in + let info = Level.Map.find_opt u (snd (snd names)) in h (pr_uctx_level_names names u ++ (match info with | None | Some {uloc=None} -> mt () | Some {uloc=Some loc} -> spc() ++ str"(" ++ Loc.pr loc ++ str")")) in + let ppqs = if nqs > 0 then + str (if nqs = 1 then "Quality" else "Qualities") ++ spc () ++ + prlist_with_sep spc prqvar (QVar.Set.elements quals) + else mt() + in + let ppus = if nus > 0 then + let universe_s = CString.plural nus "universe" in + let universe_s = if nqs = 0 then CString.capitalize_ascii universe_s else universe_s in + str universe_s ++ spc () ++ + prlist_with_sep spc prlev (Level.Set.elements univs) + else mt() + in (hv 0 - (str (CString.plural n "Universe") ++ spc () ++ - (prlist_with_sep spc prlev (Level.Set.elements left)) ++ - spc () ++ str (CString.conjugate_verb_to_be n) ++ str" unbound.")) + (ppqs ++ + (if nqs > 0 && nus > 0 then strbrk " and " else mt()) ++ + ppus ++ + spc () ++ str (CString.conjugate_verb_to_be (nus + nqs)) ++ str" unbound.")) -exception UnboundUnivs of Level.Set.t * univ_names +exception UnboundUnivs of QVar.Set.t * Level.Set.t * univ_names (* Deliberately using no location as the location of the univs doesn't correspond to the failing command. *) -let error_unbound_universes left uctx = raise (UnboundUnivs (left,uctx)) +let error_unbound_universes qs us uctx = raise (UnboundUnivs (qs,us,uctx)) let _ = CErrors.register_handler (function - | UnboundUnivs (left,uctx) -> Some (pr_error_unbound_universes left uctx) + | UnboundUnivs (qs,us,uctx) -> Some (pr_error_unbound_universes qs us uctx) | _ -> None) -let universe_context_inst ~prefix ~extensible levels names = - let left = List.fold_left (fun acc l -> Level.Set.remove l acc) levels prefix in - if not extensible && not (Level.Set.is_empty left) - then error_unbound_universes left names - else - let left = UContext.sort_levels (Array.of_list (Level.Set.elements left)) in - let inst = Array.append (Array.of_list prefix) left in - let inst = Instance.of_array inst in - inst +let universe_context_inst decl qvars levels names = + let leftqs = List.fold_left (fun acc l -> QVar.Set.remove l acc) qvars decl.univdecl_qualities in + let leftus = List.fold_left (fun acc l -> Level.Set.remove l acc) levels decl.univdecl_instance in + let () = + let unboundus = if decl.univdecl_extensible_instance then Level.Set.empty else leftus in + if not (QVar.Set.is_empty leftqs && Level.Set.is_empty unboundus) + then error_unbound_universes leftqs unboundus names + in + let leftqs = UContext.sort_qualities + (Array.map_of_list (fun q -> Quality.QVar q) (QVar.Set.elements leftqs)) + in + let leftus = UContext.sort_levels (Array.of_list (Level.Set.elements leftus)) in + let instq = Array.append + (Array.map_of_list (fun q -> Quality.QVar q) decl.univdecl_qualities) + leftqs + in + let instu = Array.append (Array.of_list decl.univdecl_instance) leftus in + let inst = Instance.of_array (instq,instu) in + inst let check_universe_context_set ~prefix levels names = let left = @@ -728,7 +873,7 @@ let check_universe_context_set ~prefix levels names = levels prefix in if not (Level.Set.is_empty left) - then error_unbound_universes left names + then error_unbound_universes QVar.Set.empty left names let check_implication uctx cstrs cstrs' = let gr = initial_graph uctx in @@ -740,6 +885,11 @@ let check_implication uctx cstrs cstrs' = Constraints.pr (pr_uctx_level uctx) cstrs') let check_mono_univ_decl uctx decl = + let () = + if not (List.is_empty decl.univdecl_qualities) + || not (QVar.Set.is_empty (QState.undefined uctx.sort_variables)) + then CErrors.user_err Pp.(str "Monomorphic declarations may not have sort variables.") + in let levels, csts = uctx.local in let () = let prefix = decl.univdecl_instance in @@ -755,10 +905,9 @@ let check_mono_univ_decl uctx decl = end let check_poly_univ_decl uctx decl = - let prefix = decl.univdecl_instance in - let extensible = decl.univdecl_extensible_instance in let levels, csts = uctx.local in - let inst = universe_context_inst ~prefix ~extensible levels uctx.names in + let qvars = QState.undefined uctx.sort_variables in + let inst = universe_context_inst decl qvars levels uctx.names in let nas = compute_instance_binders (snd uctx.names) inst in let csts = if decl.univdecl_extensible_constraints then csts else begin @@ -802,8 +951,8 @@ let restrict_universe_context ~lbound (univs, csts) keep = let restrict uctx vars = let vars = Level.Set.union vars uctx.seff_univs in - let vars = Names.Id.Map.fold (fun na l vars -> Level.Set.add l vars) - (fst uctx.names) vars + let vars = Id.Map.fold (fun na l vars -> Level.Set.add l vars) + (snd (fst uctx.names)) vars in let uctx' = restrict_universe_context ~lbound:uctx.universes_lbound uctx.local vars in { uctx with local = uctx' } @@ -822,11 +971,8 @@ let univ_flexible = UnivFlexible false let univ_flexible_alg = UnivFlexible true (** ~sideff indicates that it is ok to redeclare a universe. - ~extend also merges the universe context in the local constraint structures - and not only in the graph. This depends if the - context we merge comes from a side effect that is already inlined - or defined separately. In the later case, there is no extension, - see [emit_side_effects] for example. *) + Also merges the universe context in the local constraint structures + and not only in the graph. *) let merge ?loc ~sideff rigid uctx uctx' = let levels = ContextSet.levels uctx' in let uctx = @@ -852,14 +998,15 @@ let merge ?loc ~sideff rigid uctx uctx' = in let names = let fold u accu = - let modify _ info = match info.uloc with - | None -> { info with uloc = loc } - | Some _ -> info + let update = function + | None -> Some { uname = None; uloc = loc } + | Some info -> match info.uloc with + | None -> Some { info with uloc = loc } + | Some _ -> Some info in - try Level.Map.modify u modify accu - with Not_found -> Level.Map.add u { uname = None; uloc = loc } accu + Level.Map.update u update accu in - (fst uctx.names, Level.Set.fold fold levels (snd uctx.names)) + (fst uctx.names, (fst (snd uctx.names), Level.Set.fold fold levels (snd (snd uctx.names)))) in let initial = declare uctx.initial_universes in let univs = declare uctx.universes in @@ -867,6 +1014,31 @@ let merge ?loc ~sideff rigid uctx uctx' = { uctx with names; local; universes; initial_universes = initial } +let merge_sort_variables ?loc ~sideff uctx qvars = + let sort_variables = + QVar.Set.fold (fun qv qstate -> QState.add ~check_fresh:(not sideff) ~named:false qv qstate) + qvars + uctx.sort_variables + in + let names = + let fold u accu = + let update = function + | None -> Some { uname = None; uloc = loc } + | Some info -> match info.uloc with + | None -> Some { info with uloc = loc } + | Some _ -> Some info + in + QVar.Map.update u update accu + in + let qrev = QVar.Set.fold fold qvars (fst (snd uctx.names)) in + (fst uctx.names, (qrev, snd (snd uctx.names))) + in + { uctx with sort_variables; names } + +let merge_full ?loc ~sideff rigid uctx ((qvars,levels),csts) = + let uctx = merge_sort_variables ?loc ~sideff uctx qvars in + merge ?loc ~sideff rigid uctx (levels,csts) + (* Check bug_4363 and bug_6323 when changing this code *) let demote_seff_univs univs uctx = let seff = Level.Set.union uctx.seff_univs univs in @@ -906,6 +1078,30 @@ let update_sigma_univs uctx univs = in merge_seff eunivs eunivs.local +let add_qnames ?loc s l ((qnames,unames), (qnames_rev,unames_rev)) = + if Id.Map.mem s qnames + then user_err ?loc + Pp.(str "Quality " ++ Id.print s ++ str" already bound."); + ((Id.Map.add s l qnames, unames), + (QVar.Map.add l { uname = Some s; uloc = loc } qnames_rev, unames_rev)) + +let add_names ?loc s l ((qnames,unames), (qnames_rev,unames_rev)) = + if UNameMap.mem s unames + then user_err ?loc + Pp.(str "Universe " ++ Id.print s ++ str" already bound."); + ((qnames,UNameMap.add s l unames), + (qnames_rev, Level.Map.add l { uname = Some s; uloc = loc } unames_rev)) + +let add_qloc l loc (names, (qnames_rev,unames_rev) as orig) = + match loc with + | None -> orig + | Some _ -> (names, (QVar.Map.add l { uname = None; uloc = loc } qnames_rev, unames_rev)) + +let add_loc l loc (names, (qnames_rev,unames_rev) as orig) = + match loc with + | None -> orig + | Some _ -> (names, (qnames_rev, Level.Map.add l { uname = None; uloc = loc } unames_rev)) + let add_universe ?loc name strict lbound uctx u = let initial_universes = UGraph.add_universe ~lbound ~strict u uctx.initial_universes in let universes = UGraph.add_universe ~lbound ~strict u uctx.universes in @@ -917,10 +1113,17 @@ let add_universe ?loc name strict lbound uctx u = in { uctx with names; local; initial_universes; universes } -let new_sort_variable uctx = +let new_sort_variable ?loc ?name uctx = let q = UnivGen.new_sort_global () in - let sort_variables = QState.add q uctx.sort_variables in - { uctx with sort_variables }, q + (* don't need to check_fresh as it's guaranteed new *) + let sort_variables = QState.add ~check_fresh:false ~named:(Option.has_some name) + q uctx.sort_variables + in + let names = match name with + | Some n -> add_qnames ?loc n q uctx.names + | None -> add_qloc q loc uctx.names + in + { uctx with sort_variables; names }, q let new_univ_variable ?loc rigid name uctx = let u = UnivGen.fresh_level () in @@ -986,16 +1189,16 @@ let make_flexible_nonalgebraic uctx = let is_sort_variable uctx s = match s with (* FIXME: normalize here *) - | Sorts.Type u -> + | Type u -> (match Universe.level u with | Some l as x -> if Level.Set.mem l (ContextSet.levels uctx.local) then x else None | None -> None) - | Sorts.QSort (q, u) -> + | QSort (q, u) -> let q = nf_qvar uctx q in (match q, Universe.level u with - | QType, Some l -> + | QConstant QType, Some l -> if Level.Set.mem l (ContextSet.levels uctx.local) then Some l else None | (_, Some _ | _, None) -> None) diff --git a/engine/uState.mli b/engine/uState.mli index b2f890fa8d980..db4c17528967c 100644 --- a/engine/uState.mli +++ b/engine/uState.mli @@ -14,6 +14,7 @@ open Names open Univ +open Sorts type universes_entry = | Monomorphic_entry of Univ.ContextSet.t @@ -41,10 +42,10 @@ val from_env : ?binders:lident list -> Environ.env -> t (** Main entry point at the beginning of a declaration declaring the binding names as rigid universes. *) -val of_names : (UnivNames.universe_binders * Id.t Level.Map.t) -> t +val of_names : (UnivNames.universe_binders * UnivNames.rev_binders) -> t (** Main entry point when only names matter, e.g. for printing. *) -val of_context_set : Univ.ContextSet.t -> t +val of_context_set : UnivGen.full_context_set -> t (** Main entry point when starting from the instance of a global reference, e.g. when building a scheme. *) @@ -60,6 +61,8 @@ val context_set : t -> Univ.ContextSet.t (** The local context of the state, i.e. a set of bound variables together with their associated constraints. *) +val full_context_set : t -> UnivGen.full_context_set + type universe_opt_subst = UnivSubst.universe_opt_subst (* Reexport because UnivSubst is private *) @@ -93,13 +96,22 @@ val univ_entry : poly:bool -> t -> named_universes_entry val universe_binders : t -> UnivNames.universe_binders (** Return local names of universes. *) -val nf_qvar : t -> Sorts.QVar.t -> Sorts.Quality.t +val nf_qvar : t -> QVar.t -> Quality.t (** Returns the normal form of the sort variable. *) +val nf_quality : t -> Quality.t -> Quality.t + +val nf_instance : t -> UVars.Instance.t -> UVars.Instance.t + +val nf_level : t -> Level.t -> Level.t +(** Must not be allowed to be algebraic *) + +val nf_universe : t -> Universe.t -> Universe.t + val nf_sort : t -> Sorts.t -> Sorts.t (** Returns the normal form of the sort. *) -val nf_relevance : t -> Sorts.relevance -> Sorts.relevance +val nf_relevance : t -> relevance -> relevance (** Returns the normal form of the relevance. *) (** {5 Constraints handling} *) @@ -114,8 +126,16 @@ val add_universe_constraints : t -> UnivProblem.Set.t -> t @raise UniversesDiffer when universes differ *) +val check_qconstraints : t -> QConstraints.t -> bool + +val check_universe_constraints : t -> UnivProblem.Set.t -> bool + +val add_quconstraints : t -> QUConstraints.t -> t + (** {5 Names} *) +val quality_of_name : t -> Id.t -> Sorts.QVar.t + val universe_of_name : t -> Id.t -> Univ.Level.t (** Retrieve the universe associated to the name. *) @@ -154,6 +174,9 @@ val univ_flexible : rigid val univ_flexible_alg : rigid val merge : ?loc:Loc.t -> sideff:bool -> rigid -> t -> Univ.ContextSet.t -> t +val merge_sort_variables : ?loc:Loc.t -> sideff:bool -> t -> QVar.Set.t -> t +val merge_full : ?loc:Loc.t -> sideff:bool -> rigid -> t -> UnivGen.full_context_set -> t + val emit_side_effects : Safe_typing.private_constants -> t -> t val demote_global_univs : Environ.env -> t -> t @@ -166,7 +189,7 @@ val demote_seff_univs : Univ.Level.Set.t -> t -> t globally declared by some side effect. You should be using emit_side_effects instead. *) -val new_sort_variable : t -> t * Sorts.QVar.t +val new_sort_variable : ?loc:Loc.t -> ?name:Id.t -> t -> t * QVar.t (** Declare a new local sort. *) val new_univ_variable : ?loc:Loc.t -> rigid -> Id.t option -> t -> t * Univ.Level.t @@ -211,14 +234,15 @@ val minimize : t -> t val collapse_sort_variables : t -> t -type ('a, 'b) gen_universe_decl = { - univdecl_instance : 'a; (* Declared universes *) +type ('a, 'b, 'c) gen_universe_decl = { + univdecl_qualities : 'a; + univdecl_instance : 'b; (* Declared universes *) univdecl_extensible_instance : bool; (* Can new universes be added *) - univdecl_constraints : 'b; (* Declared constraints *) + univdecl_constraints : 'c; (* Declared constraints *) univdecl_extensible_constraints : bool (* Can new constraints be added *) } type universe_decl = - (Level.t list, Univ.Constraints.t) gen_universe_decl + (QVar.t list, Level.t list, Univ.Constraints.t) gen_universe_decl val default_univ_decl : universe_decl @@ -244,11 +268,14 @@ val update_sigma_univs : t -> UGraph.t -> t (** {5 Pretty-printing} *) val pr_uctx_level : t -> Univ.Level.t -> Pp.t +val pr_uctx_qvar : t -> Sorts.QVar.t -> Pp.t val qualid_of_level : t -> Univ.Level.t -> Libnames.qualid option (** Only looks in the local names, not in the nametab. *) val id_of_level : t -> Univ.Level.t -> Id.t option +val id_of_qvar : t -> Sorts.QVar.t -> Id.t option + val pr_weak : (Univ.Level.t -> Pp.t) -> t -> Pp.t val pr_universe_opt_subst : (Level.t -> Pp.t) -> universe_opt_subst -> Pp.t diff --git a/engine/univGen.ml b/engine/univGen.ml index 8d212ae8b83a6..7d239ea7d83e9 100644 --- a/engine/univGen.ml +++ b/engine/univGen.ml @@ -14,18 +14,38 @@ open Constr open Univ open UVars +type full_context_set = (Sorts.QVar.Set.t * Univ.Level.Set.t) * Univ.Constraints.t + +type 'a in_full_context_set = 'a * full_context_set + +let empty_full_context = (QVar.Set.empty, Level.Set.empty), Constraints.empty + +let is_empty_full_context ((qs,us),csts) = + QVar.Set.is_empty qs && Level.Set.is_empty us && Constraints.is_empty csts + +let full_context_union ((qs,us),csts) ((qs',us'),csts') = + ((QVar.Set.union qs qs', Level.Set.union us us'),Constraints.union csts csts') + +let diff_full_context ((qs,us),csts) ((qs',us'),csts') = + (QVar.Set.diff qs qs', Level.Set.diff us us'), Constraints.diff csts csts' + type univ_length_mismatch = { - actual : int ; - expect : int ; + actual : int * int; + expect : int * int; } (* Due to an OCaml bug ocaml/ocaml#10027 inlining this record will cause compliation with -rectypes to crash. *) exception UniverseLengthMismatch of univ_length_mismatch let () = CErrors.register_handler (function - | UniverseLengthMismatch { actual; expect } -> - Some Pp.(str "Universe instance length is " ++ int actual - ++ str " but should be " ++ int expect ++ str ".") + | UniverseLengthMismatch { actual=(aq,au); expect=(eq,eu) } -> + let ppreal, ppexpected = + if aq = 0 && eq = 0 then Pp.(int au, int eu) + else Pp.(str "(" ++ int aq ++ str " | " ++ int au ++ str ")" + , str "(" ++ int eq ++ str " | " ++ int eu ++ str ")") + in + Some Pp.(str "Universe instance length is " ++ ppreal + ++ str " but should be " ++ ppexpected ++ str".") | _ -> None) (* Generator of levels *) @@ -48,21 +68,29 @@ let new_sort_global () = let s = if Flags.async_proofs_is_worker() then !Flags.async_proofs_worker_id else "" in Sorts.QVar.make_unif s (new_sort_id ()) -let fresh_instance auctx = - let inst = Array.init (AbstractContext.size auctx) (fun _ -> fresh_level()) in - let ctx = Array.fold_right Level.Set.add inst Level.Set.empty in - let inst = Instance.of_array inst in - inst, (ctx, AbstractContext.instantiate inst auctx) +let fresh_instance auctx : _ in_full_context_set = + let qlen, ulen = AbstractContext.size auctx in + let qinst = Array.init qlen (fun _ -> Sorts.Quality.QVar (new_sort_global())) in + let uinst = Array.init ulen (fun _ -> fresh_level()) in + let qctx = Array.fold_left (fun qctx q -> match q with + | Sorts.Quality.QVar q -> Sorts.QVar.Set.add q qctx + | _ -> assert false) + Sorts.QVar.Set.empty + qinst + in + let uctx = Array.fold_right Level.Set.add uinst Level.Set.empty in + let inst = Instance.of_array (qinst,uinst) in + inst, ((qctx,uctx), AbstractContext.instantiate inst auctx) let existing_instance ?loc auctx inst = let () = - let actual = Array.length (Instance.to_array inst) + let actual = Instance.length inst and expect = AbstractContext.size auctx in - if not (Int.equal actual expect) then + if not (UVars.eq_sizes actual expect) then Loc.raise ?loc (UniverseLengthMismatch { actual; expect }) else () in - inst, (Level.Set.empty, AbstractContext.instantiate inst auctx) + inst, ((Sorts.QVar.Set.empty,Level.Set.empty), AbstractContext.instantiate inst auctx) let fresh_instance_from ?loc ctx = function | Some inst -> existing_instance ?loc ctx inst @@ -104,12 +132,12 @@ let constr_of_monomorphic_global env gr = str " would forget universes.") let fresh_sort_in_family = function - | InSProp -> Sorts.sprop, ContextSet.empty - | InProp -> Sorts.prop, ContextSet.empty - | InSet -> Sorts.set, ContextSet.empty + | InSProp -> Sorts.sprop, empty_full_context + | InProp -> Sorts.prop, empty_full_context + | InSet -> Sorts.set, empty_full_context | InType | InQSort (* Treat as Type *) -> let u = fresh_level () in - sort_of_univ (Univ.Universe.make u), ContextSet.singleton u + sort_of_univ (Univ.Universe.make u), ((QVar.Set.empty,Level.Set.singleton u),Constraints.empty) let new_global_univ () = let u = fresh_level () in @@ -127,3 +155,13 @@ let fresh_universe_context_set_instance ctx = in let cst' = subst_univs_level_constraints subst cst in subst, (univs', cst') + +let fresh_full_context_instance ((qs,us),csts) = + let usubst, (us, csts) = fresh_universe_context_set_instance (us,csts) in + let qsubst, qs = QVar.Set.fold (fun q (qsubst,qs) -> + let q' = new_sort_global () in + QVar.Map.add q (Sorts.Quality.QVar q') qsubst, QVar.Set.add q' qs) + qs + (QVar.Map.empty, QVar.Set.empty) + in + (qsubst, usubst), ((qs, us), csts) diff --git a/engine/univGen.mli b/engine/univGen.mli index ffe13acb29239..cb80c56ca6143 100644 --- a/engine/univGen.mli +++ b/engine/univGen.mli @@ -15,8 +15,8 @@ open Univ open UVars type univ_length_mismatch = { - actual : int ; - expect : int ; + actual : int * int; + expect : int * int; } (* Due to an OCaml bug ocaml/ocaml#10027 inlining this record will cause compliation with -rectypes to crash. *) @@ -33,30 +33,47 @@ val new_global_univ : unit -> Universe.t in_universe_context_set (** Build a fresh instance for a given context, its associated substitution and the instantiated constraints. *) -val fresh_instance : AbstractContext.t -> Instance.t in_universe_context_set +type full_context_set = (Sorts.QVar.Set.t * Univ.Level.Set.t) * Univ.Constraints.t + +type 'a in_full_context_set = 'a * full_context_set + +val full_context_union : full_context_set -> full_context_set -> full_context_set + +val empty_full_context : full_context_set + +val is_empty_full_context : full_context_set -> bool + +val diff_full_context : full_context_set -> full_context_set -> full_context_set + +val fresh_instance : AbstractContext.t -> Instance.t in_full_context_set val fresh_instance_from : ?loc:Loc.t -> AbstractContext.t -> Instance.t option -> - Instance.t in_universe_context_set + Instance.t in_full_context_set val fresh_sort_in_family : Sorts.family -> - Sorts.t in_universe_context_set + Sorts.t in_full_context_set +(** NB: InQSort is treated as InType *) + val fresh_constant_instance : env -> Constant.t -> - pconstant in_universe_context_set + pconstant in_full_context_set val fresh_inductive_instance : env -> inductive -> - pinductive in_universe_context_set + pinductive in_full_context_set val fresh_constructor_instance : env -> constructor -> - pconstructor in_universe_context_set + pconstructor in_full_context_set val fresh_array_instance : env -> - Instance.t in_universe_context_set + Instance.t in_full_context_set val fresh_global_instance : ?loc:Loc.t -> ?names:UVars.Instance.t -> env -> GlobRef.t -> - constr in_universe_context_set + constr in_full_context_set (** Get fresh variables for the universe context. Useful to make tactics that manipulate constrs in universe contexts polymorphic. *) val fresh_universe_context_set_instance : ContextSet.t -> universe_level_subst * ContextSet.t +val fresh_full_context_instance : full_context_set -> + full_level_subst * full_context_set + (** Create a fresh global in the environment argument, without side effects. BEWARE: this raises an error on polymorphic constants/inductives: the constraints should be properly added to an evd. diff --git a/engine/univNames.ml b/engine/univNames.ml index 16fd451fe1071..ece6b25d99550 100644 --- a/engine/univNames.ml +++ b/engine/univNames.ml @@ -10,21 +10,37 @@ open Names open Univ +open Sorts -type universe_binders = Level.t Names.Id.Map.t +type universe_binders = QVar.t Id.Map.t * Level.t Id.Map.t -let empty_binders = Id.Map.empty +type rev_binders = Id.t QVar.Map.t * Id.t Level.Map.t + +let empty_binders = Id.Map.empty, Id.Map.empty + +let empty_rev_binders = QVar.Map.empty, Level.Map.empty type univ_name_list = Names.lname list -let qualid_of_level ctx l = +type full_name_list = lname list * lname list + +let qualid_of_level (_,ctx) l = match Level.name l with | Some qid -> (try Some (Nametab.shortest_qualid_of_universe ctx qid) with Not_found -> None) | None -> None -let pr_with_global_universes ?(binders=empty_binders) l = +let pr_level_with_global_universes ?(binders=empty_binders) l = match qualid_of_level binders l with | Some qid -> Libnames.pr_qualid qid | None -> Level.raw_pr l + +let qualid_of_qvar (ctx,_) l = + try Some (Nametab.shortest_qualid_of_qvar ctx l) + with Not_found -> None + +let pr_qvar_with_global_universes ?(binders=empty_binders) l = + match qualid_of_qvar binders l with + | Some qid -> Libnames.pr_qualid qid + | None -> QVar.raw_pr l diff --git a/engine/univNames.mli b/engine/univNames.mli index 85a3ea097ecb2..b15a713d7adda 100644 --- a/engine/univNames.mli +++ b/engine/univNames.mli @@ -8,15 +8,26 @@ (* * (see LICENSE file for the text of the license) *) (************************************************************************) +open Names open Univ +open Sorts (** Local universe name <-> level mapping *) -type universe_binders = Univ.Level.t Names.Id.Map.t +type universe_binders = QVar.t Id.Map.t * Level.t Id.Map.t + +type rev_binders = Id.t QVar.Map.t * Id.t Level.Map.t val empty_binders : universe_binders +val empty_rev_binders : rev_binders + type univ_name_list = Names.lname list -val pr_with_global_universes : ?binders:universe_binders -> Level.t -> Pp.t +type full_name_list = lname list * lname list + +val pr_level_with_global_universes : ?binders:universe_binders -> Level.t -> Pp.t val qualid_of_level : universe_binders -> Level.t -> Libnames.qualid option + +val pr_qvar_with_global_universes : ?binders:universe_binders -> QVar.t -> Pp.t +val qualid_of_qvar : universe_binders -> QVar.t -> Libnames.qualid option diff --git a/engine/univProblem.ml b/engine/univProblem.ml index 3699af9071676..55eb359c7ac66 100644 --- a/engine/univProblem.ml +++ b/engine/univProblem.ml @@ -11,6 +11,8 @@ open Univ type t = + | QEq of Sorts.Quality.t * Sorts.Quality.t + | QLeq of Sorts.Quality.t * Sorts.Quality.t | ULe of Sorts.t * Sorts.t | UEq of Sorts.t * Sorts.t | ULub of Level.t * Level.t @@ -18,21 +20,17 @@ type t = let is_trivial = function + | QLeq (QConstant QProp, QConstant QType) -> true + | QLeq (a,b) | QEq (a, b) -> Sorts.Quality.equal a b | ULe (u, v) | UEq (u, v) -> Sorts.equal u v | ULub (u, v) | UWeak (u, v) -> Level.equal u v let force = function - | ULe _ | UEq _ | UWeak _ as cst -> cst + | QEq _ | QLeq _ | ULe _ | UEq _ | UWeak _ as cst -> cst | ULub (u,v) -> UEq (Sorts.sort_of_univ @@ Universe.make u, Sorts.sort_of_univ @@ Universe.make v) let check_eq_level g u v = UGraph.check_eq_level g u v -let check g = function - | ULe (u,v) -> UGraph.check_leq_sort g u v - | UEq (u,v) -> UGraph.check_eq_sort g u v - | ULub (u,v) -> check_eq_level g u v - | UWeak _ -> true - module Set = struct module S = Set.Make( struct @@ -40,6 +38,14 @@ module Set = struct let compare x y = match x, y with + | QEq (a, b), QEq (a', b') -> + let i = Sorts.Quality.compare a a' in + if i <> 0 then i + else Sorts.Quality.compare b b' + | QLeq (a, b), QLeq (a', b') -> + let i = Sorts.Quality.compare a a' in + if i <> 0 then i + else Sorts.Quality.compare b b' | ULe (u, v), ULe (u', v') -> let i = Sorts.compare u u' in if Int.equal i 0 then Sorts.compare v v' @@ -54,6 +60,10 @@ module Set = struct if Int.equal i 0 then Level.compare v v' else if Level.equal u v' && Level.equal v u' then 0 else i + | QEq _, _ -> -1 + | _, QEq _ -> 1 + | QLeq _, _ -> -1 + | _, QLeq _ -> 1 | ULe _, _ -> -1 | _, ULe _ -> 1 | UEq _, _ -> -1 @@ -69,6 +79,8 @@ module Set = struct else add cst s let pr_one = let open Pp in function + | QEq (a, b) -> Sorts.Quality.raw_pr a ++ str " = " ++ Sorts.Quality.raw_pr b + | QLeq (a, b) -> Sorts.Quality.raw_pr a ++ str " <= " ++ Sorts.Quality.raw_pr b | ULe (u, v) -> Sorts.debug_print u ++ str " <= " ++ Sorts.debug_print v | UEq (u, v) -> Sorts.debug_print u ++ str " = " ++ Sorts.debug_print v | ULub (u, v) -> Level.raw_pr u ++ str " /\\ " ++ Level.raw_pr v @@ -83,8 +95,6 @@ module Set = struct x == y || equal x y let force s = map force s - - let check g s = for_all (check g) s end type 'a constraint_function = 'a -> 'a -> Set.t -> Set.t @@ -92,10 +102,40 @@ type 'a constraint_function = 'a -> 'a -> Set.t -> Set.t let enforce_eq_instances_univs strict x y c = let mkU u = Sorts.sort_of_univ @@ Universe.make u in let mk u v = if strict then ULub (u, v) else UEq (mkU u, mkU v) in - let ax = UVars.Instance.to_array x and ay = UVars.Instance.to_array y in - if Array.length ax != Array.length ay then - CErrors.anomaly Pp.(str "Invalid argument: enforce_eq_instances_univs called with" ++ - str " instances of different lengths."); - CArray.fold_right2 - (fun x y -> Set.add (mk x y)) - ax ay c + if not (UVars.eq_sizes (UVars.Instance.length x) (UVars.Instance.length y)) then + CErrors.anomaly Pp.(str "Invalid argument: enforce_eq_instances_univs called with" ++ + str " instances of different lengths."); + let xq, xu = UVars.Instance.to_array x and yq, yu = UVars.Instance.to_array y in + let c = CArray.fold_left2 + (* TODO strict? *) + (fun c x y -> if Sorts.Quality.equal x y then c else Set.add (QEq (x,y)) c) + c xq yq + in + let c = CArray.fold_left2 + (fun c x y -> Set.add (mk x y) c) + c xu yu + in + c + +let enforce_eq_qualities qs qs' cstrs = + CArray.fold_left2 (fun c a b -> + if Sorts.Quality.equal a b then c else Set.add (QEq (a, b)) c) + cstrs qs qs' + +let compare_cumulative_instances cv_pb variances u u' cstrs = + let make u = Sorts.sort_of_univ @@ Univ.Universe.make u in + let qs, us = UVars.Instance.to_array u + and qs', us' = UVars.Instance.to_array u' in + let cstrs = enforce_eq_qualities qs qs' cstrs in + CArray.fold_left3 + (fun cstrs v u u' -> + let open UVars.Variance in + match v with + | Irrelevant -> Set.add (UWeak (u,u')) cstrs + | Covariant -> + (match cv_pb with + | Conversion.CONV -> Set.add (UEq (make u, make u')) cstrs + | Conversion.CUMUL -> Set.add (ULe (make u, make u')) cstrs) + | Invariant -> + Set.add (UEq (make u, make u')) cstrs) + cstrs variances us us' diff --git a/engine/univProblem.mli b/engine/univProblem.mli index 75be2de1a2b35..5875d43ad7b67 100644 --- a/engine/univProblem.mli +++ b/engine/univProblem.mli @@ -21,6 +21,8 @@ open UVars *) type t = + | QEq of Sorts.Quality.t * Sorts.Quality.t + | QLeq of Sorts.Quality.t * Sorts.Quality.t | ULe of Sorts.t * Sorts.t | UEq of Sorts.t * Sorts.t | ULub of Level.t * Level.t @@ -28,8 +30,6 @@ type t = val is_trivial : t -> bool -val check : UGraph.t -> t -> bool - (** Wrapper around the UGraph function to handle Prop *) val check_eq_level : UGraph.t -> Level.t -> Level.t -> bool @@ -40,10 +40,12 @@ module Set : sig (** Replace ULub constraints by UEq *) val force : t -> t - - val check : UGraph.t -> t -> bool end type 'a constraint_function = 'a -> 'a -> Set.t -> Set.t val enforce_eq_instances_univs : bool -> Instance.t constraint_function + +val enforce_eq_qualities : Sorts.Quality.t array constraint_function + +val compare_cumulative_instances : Conversion.conv_pb -> Variance.t array -> Instance.t constraint_function diff --git a/engine/univSubst.ml b/engine/univSubst.ml index 054c5afa9c074..310278157d6e9 100644 --- a/engine/univSubst.ml +++ b/engine/univSubst.ml @@ -19,8 +19,8 @@ type universe_subst = Universe.t universe_map type universe_subst_fn = Level.t -> Universe.t option type universe_level_subst_fn = Level.t -> Level.t -let subst_instance fn i = - Instance.of_array (Array.Smart.map fn (Instance.to_array i)) +type quality_subst = Quality.t QVar.Map.t +type quality_subst_fn = QVar.t -> Quality.t let subst_univs_universe fn ul = let addn n u = iterate Universe.super n u in @@ -194,7 +194,7 @@ let normalize_univ_variables ctx = in ctx, def, subst let subst_univs_fn_puniverses f (c, u as cu) = - let u' = subst_instance f u in + let u' = Instance.subst_fn f u in if u' == u then cu else (c, u') let nf_binder_annot frel na = @@ -203,7 +203,9 @@ let nf_binder_annot frel na = if rel' == na.binder_relevance then na else { binder_name = na.binder_name; binder_relevance = rel' } -let nf_evars_and_universes_opt_subst fevar flevel fsort frel c = +let nf_evars_and_universes_opt_subst fevar fqual funiv c = + let frel = Sorts.relevance_subst_fn fqual in + let flevel = fqual, level_subst_of funiv in let rec aux c = match kind c with | Evar (evk, args) -> @@ -221,10 +223,10 @@ let nf_evars_and_universes_opt_subst fevar flevel fsort frel c = let pu' = subst_univs_fn_puniverses flevel pu in if pu' == pu then c else mkConstructU pu' | Sort s -> - let s' = fsort s in + let s' = Sorts.subst_fn (fqual, subst_univs_universe funiv) s in if s' == s then c else mkSort s' | Case (ci,u,pms,p,iv,t,br) -> - let u' = subst_instance flevel u in + let u' = Instance.subst_fn flevel u in let ci' = let rel' = frel ci.ci_relevance in if rel' == ci.ci_relevance then ci else { ci with ci_relevance = rel' } @@ -237,7 +239,7 @@ let nf_evars_and_universes_opt_subst fevar flevel fsort frel c = if ci' == ci && u' == u && pms' == pms && p' == p && iv' == iv && t' == t && br' == br then c else mkCase (ci', u', pms', p', iv', t', br') | Array (u,elems,def,ty) -> - let u' = subst_instance flevel u in + let u' = Instance.subst_fn flevel u in let elems' = CArray.Smart.map aux elems in let def' = aux def in let ty' = aux ty in diff --git a/engine/univSubst.mli b/engine/univSubst.mli index 6b1228366636f..ac15d5e8ea34d 100644 --- a/engine/univSubst.mli +++ b/engine/univSubst.mli @@ -10,18 +10,20 @@ open Constr open Univ -open UVars +open Sorts type 'a universe_map = 'a Level.Map.t type universe_subst = Universe.t universe_map type universe_subst_fn = Level.t -> Universe.t option type universe_level_subst_fn = Level.t -> Level.t +type quality_subst = Quality.t QVar.Map.t +type quality_subst_fn = QVar.t -> Quality.t + val level_subst_of : universe_subst_fn -> universe_level_subst_fn (** The resulting function must never be called on a level which would produce an algebraic. *) val subst_univs_constraints : universe_subst_fn -> Constraints.t -> Constraints.t -val subst_instance : universe_level_subst_fn -> Instance.t -> Instance.t type universe_opt_subst = Universe.t option universe_map @@ -41,12 +43,11 @@ val nf_binder_annot : (Sorts.relevance -> Sorts.relevance) -> (** Full universes substitutions into terms *) -val nf_evars_and_universes_opt_subst : - (existential -> constr option) -> - (Level.t -> Level.t) -> - (Sorts.t -> Sorts.t) -> - (Sorts.relevance -> Sorts.relevance) -> - constr -> constr +val nf_evars_and_universes_opt_subst + : (existential -> constr option) + -> quality_subst_fn + -> universe_subst_fn + -> constr -> constr val subst_univs_universe : (Level.t -> Universe.t option) -> Universe.t -> Universe.t diff --git a/interp/constrexpr.mli b/interp/constrexpr.mli index 1db7a0e02fc4e..2edd20cd09549 100644 --- a/interp/constrexpr.mli +++ b/interp/constrexpr.mli @@ -20,16 +20,27 @@ type sort_name_expr = | CRawType of Univ.Level.t (** Universes like "foo.1" have no qualid form *) type univ_level_expr = sort_name_expr Glob_term.glob_sort_gen -type sort_expr = (Sorts.QVar.t option * (sort_name_expr * int) list) Glob_term.glob_sort_gen -type instance_expr = univ_level_expr list +type _ quality_expr_g = + | CQConstant : Sorts.Quality.constant -> [`with_constant] quality_expr_g + | CQVar of qualid + | CQAnon of Loc.t option + | CRawQVar of Sorts.QVar.t + +type quality_expr = [`with_constant] quality_expr_g + +type qvar_expr = [`no_constant] quality_expr_g + +type sort_expr = (qvar_expr option * (sort_name_expr * int) list) Glob_term.glob_sort_gen + +type instance_expr = quality_expr list * univ_level_expr list (** Constraints don't have anonymous universes *) type univ_constraint_expr = sort_name_expr * Univ.constraint_type * sort_name_expr -type universe_decl_expr = (lident list, univ_constraint_expr list) UState.gen_universe_decl +type universe_decl_expr = (lident list, lident list, univ_constraint_expr list) UState.gen_universe_decl type cumul_univ_decl_expr = - ((lident * UVars.Variance.t option) list, univ_constraint_expr list) UState.gen_universe_decl + (lident list, (lident * UVars.Variance.t option) list, univ_constraint_expr list) UState.gen_universe_decl type ident_decl = lident * universe_decl_expr option type cumul_ident_decl = lident * cumul_univ_decl_expr option diff --git a/interp/constrexpr_ops.ml b/interp/constrexpr_ops.ml index 9ecdffb79eb8c..f3379d438882e 100644 --- a/interp/constrexpr_ops.ml +++ b/interp/constrexpr_ops.ml @@ -29,12 +29,19 @@ let sort_name_expr_eq c1 c2 = match c1, c2 with | CRawType u1, CRawType u2 -> Univ.Level.equal u1 u2 | (CSProp|CProp|CSet|CType _|CRawType _), _ -> false +let quality_expr_eq (type a) (c1:a quality_expr_g) (c2:a quality_expr_g) = match c1, c2 with + | CQConstant q1, CQConstant q2 -> Sorts.Quality.Constants.equal q1 q2 + | CQVar q1, CQVar q2 -> Libnames.qualid_eq q1 q2 + | CQAnon _, CQAnon _ -> true + | CRawQVar q1, CRawQVar q2 -> Sorts.QVar.equal q1 q2 + | (CQConstant _ | CQVar _ | CQAnon _ | CRawQVar _), _ -> false + let univ_level_expr_eq u1 u2 = Glob_ops.glob_sort_gen_eq sort_name_expr_eq u1 u2 let sort_expr_eq u1 u2 = let eq (q1, l1) (q2, l2) = - Option.equal Sorts.QVar.equal q1 q2 && + Option.equal quality_expr_eq q1 q2 && List.equal (fun (x,m) (y,n) -> sort_name_expr_eq x y && Int.equal m n) l1 l2 in Glob_ops.glob_sort_gen_eq eq u1 u2 diff --git a/interp/constrextern.ml b/interp/constrextern.ml index 3d33cbbe1ae0a..2ccd7d1c61979 100644 --- a/interp/constrextern.ml +++ b/interp/constrextern.ml @@ -815,7 +815,7 @@ let extern_float f scopes = type extern_env = Id.Set.t * UnivNames.universe_binders let extern_env env sigma = vars_of_env env, Evd.universe_binders sigma -let empty_extern_env = Id.Set.empty, Id.Map.empty +let empty_extern_env = Id.Set.empty, UnivNames.empty_binders let extern_glob_sort_name uvars = function | GSProp -> CSProp @@ -828,9 +828,20 @@ let extern_glob_sort_name uvars = function | None -> CRawType u end +let extern_glob_quality (type a) uvars (q: a glob_quality_g) : a quality_expr_g = + match q with + | GQConstant q -> CQConstant q + | GLocalQVar {v=Anonymous;loc} -> CQAnon loc + | GLocalQVar {v=Name id; loc} -> CQVar (qualid_of_ident ?loc id) + | GRawQVar q -> CRawQVar q + | GQVar q -> begin match UnivNames.qualid_of_qvar uvars q with + | Some qid -> CQVar qid + | None -> CRawQVar q + end + let extern_glob_sort uvars u = let map (q, l) = - q, List.map (on_fst (extern_glob_sort_name uvars)) l + Option.map (extern_glob_quality uvars) q, List.map (on_fst (extern_glob_sort_name uvars)) l in map_glob_sort_gen map u @@ -842,8 +853,10 @@ let extern_glob_sort uvars = function | UNamed _ | UAnonymous _ as u -> extern_glob_sort uvars u let extern_instance uvars = function - | Some l when !print_universes -> - Some (List.map (map_glob_sort_gen (extern_glob_sort_name uvars)) l) + | Some (ql,ul) when !print_universes -> + let ql = List.map (extern_glob_quality uvars) ql in + let ul = List.map (map_glob_sort_gen (extern_glob_sort_name uvars)) ul in + Some (ql,ul) | _ -> None let extern_ref (vars,uvars) ref us = diff --git a/interp/constrintern.ml b/interp/constrintern.ml index dfe9cc13f3111..2fdf3ea50352e 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -234,9 +234,9 @@ let contract_curly_brackets_pat ntn (l,ll) = (* side effect; don't inline *) (InConstrEntry,!ntn'),(l,ll) -type local_univs = { bound : Univ.Level.t Id.Map.t; unb_univs : bool } +type local_univs = { bound : UnivNames.universe_binders; unb_univs : bool } -let empty_local_univs = { bound = Id.Map.empty; unb_univs = false } +let empty_local_univs = { bound = UnivNames.empty_binders; unb_univs = false } type abstraction_kind = AbsLambda | AbsPi @@ -1083,7 +1083,7 @@ let string_of_ty = function | Variable -> "var" let gvar (loc, id) us = match us with - | None | Some [] -> DAst.make ?loc @@ GVar id + | None | Some ([],[]) -> DAst.make ?loc @@ GVar id | Some _ -> user_err ?loc (str "Variable " ++ Id.print id ++ str " cannot have a universe instance") @@ -1166,7 +1166,7 @@ let intern_sort_name ~local_univs = function | CType qid -> let is_id = qualid_is_ident qid in let local = if not is_id then None - else Id.Map.find_opt (qualid_basename qid) local_univs.bound + else Id.Map.find_opt (qualid_basename qid) (snd local_univs.bound) in match local with | Some u -> GUniv u @@ -1176,18 +1176,40 @@ let intern_sort_name ~local_univs = function if is_id && local_univs.unb_univs then GLocalUniv (CAst.make ?loc:qid.loc (qualid_basename qid)) else - CErrors.user_err Pp.(str "Undeclared universe " ++ pr_qualid qid ++ str".") + CErrors.user_err ?loc:qid.loc Pp.(str "Undeclared universe " ++ pr_qualid qid ++ str".") + +let intern_quality (type a) ~local_univs (q:a quality_expr_g) : a glob_quality_g = + match q with + | CQConstant q -> GQConstant q + | CQAnon loc -> GLocalQVar (CAst.make ?loc Anonymous) + | CRawQVar q -> GRawQVar q + | CQVar qid -> + let is_id = qualid_is_ident qid in + let local = if not is_id then None + else Id.Map.find_opt (qualid_basename qid) (fst local_univs.bound) + in + match local with + | Some u -> GQVar u + | None -> + try GQVar (Nametab.locate_qvar qid) + with Not_found -> + if is_id && local_univs.unb_univs + then GLocalQVar (CAst.make ?loc:qid.loc (Name (qualid_basename qid))) + else + CErrors.user_err ?loc:qid.loc Pp.(str "Undeclared quality " ++ pr_qualid qid ++ str".") let intern_sort ~local_univs s = let map (q, l) = - (* No user-facing syntax for qualities *) - let () = assert (Option.is_empty q) in - None, List.map (on_fst (intern_sort_name ~local_univs)) l + Option.map (intern_quality ~local_univs) q, List.map (on_fst (intern_sort_name ~local_univs)) l in map_glob_sort_gen map s -let intern_instance ~local_univs us = - Option.map (List.map (map_glob_sort_gen (intern_sort_name ~local_univs))) us +let intern_instance ~local_univs = function + | None -> None + | Some (qs, us) -> + let qs = List.map (intern_quality ~local_univs) qs in + let us = List.map (map_glob_sort_gen (intern_sort_name ~local_univs)) us in + Some (qs, us) let intern_name_alias = function | { CAst.v = CRef(qid,u) } -> @@ -1305,8 +1327,9 @@ let intern_qualid ?(no_secvar=false) qid intern env ntnvars us args = DAst.make ?loc @@ GApp (DAst.make ?loc:loc' @@ GRef (ref, us), arg) | _ -> err () end - | Some [s], GSort (UAnonymous {rigid=true}) -> DAst.make ?loc @@ GSort (glob_sort_of_level s) - | Some [_old_level], GSort _new_sort -> + | Some ([],[s]), GSort (UAnonymous {rigid=true}) -> + DAst.make ?loc @@ GSort (glob_sort_of_level s) + | Some ([],[_old_level]), GSort _new_sort -> (* TODO: add old_level and new_sort to the error message *) user_err ?loc (str "Cannot change universe level of notation " ++ pr_qualid qid) | Some _, _ -> err () @@ -2859,15 +2882,20 @@ let interp_univ_constraints env evd cstrs = let interp_univ_decl env decl = let open UState in - let binders : lident list = decl.univdecl_instance in let evd = Evd.from_env env in + let evd, qualities = List.fold_left_map (fun evd lid -> + Evd.new_quality_variable ?loc:lid.loc ~name:lid.v evd) + evd + decl.univdecl_qualities + in let evd, instance = List.fold_left_map (fun evd lid -> Evd.new_univ_level_variable ?loc:lid.loc univ_rigid ~name:lid.v evd) evd - binders + decl.univdecl_instance in let evd, cstrs = interp_univ_constraints env evd decl.univdecl_constraints in let decl = { + univdecl_qualities = qualities; univdecl_instance = instance; univdecl_extensible_instance = decl.univdecl_extensible_instance; univdecl_constraints = cstrs; @@ -2880,6 +2908,11 @@ let interp_cumul_univ_decl env decl = let binders = List.map fst decl.univdecl_instance in let variances = Array.map_of_list snd decl.univdecl_instance in let evd = Evd.from_env env in + let evd, qualities = List.fold_left_map (fun evd lid -> + Evd.new_quality_variable ?loc:lid.loc ~name:lid.v evd) + evd + decl.univdecl_qualities + in let evd, instance = List.fold_left_map (fun evd lid -> Evd.new_univ_level_variable ?loc:lid.loc univ_rigid ~name:lid.v evd) evd @@ -2887,6 +2920,7 @@ let interp_cumul_univ_decl env decl = in let evd, cstrs = interp_univ_constraints env evd decl.univdecl_constraints in let decl = { + univdecl_qualities = qualities; univdecl_instance = instance; univdecl_extensible_instance = decl.univdecl_extensible_instance; univdecl_constraints = cstrs; diff --git a/interp/constrintern.mli b/interp/constrintern.mli index 91347d876adcb..c8864b7a3c014 100644 --- a/interp/constrintern.mli +++ b/interp/constrintern.mli @@ -150,7 +150,7 @@ val intern_reference : qualid -> GlobRef.t option (** Returns None if not a reference or a abbrev not bound to a name *) val intern_name_alias : - constr_expr -> (GlobRef.t * Glob_term.glob_level list option) option + constr_expr -> (GlobRef.t * Glob_term.glob_instance option) option (** Expands abbreviations; raise an error if not existing *) val interp_reference : ltac_sign -> qualid -> glob_constr diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml index b28a3e68322a2..29335c7b58712 100644 --- a/interp/notation_ops.ml +++ b/interp/notation_ops.ml @@ -80,15 +80,26 @@ let compare_glob_universe_instances lt strictly_lt us1 us2 = | None, None -> true | Some _, None -> strictly_lt := true; lt | None, Some _ -> false - | Some l1, Some l2 -> - CList.for_all2eq (fun u1 u2 -> + | Some (ql1,ul1), Some (ql2,ul2) -> + let is_anon = function + | GLocalQVar {v=Anonymous} -> true + | _ -> false + in + CList.for_all2eq (fun q1 q2 -> + match is_anon q1, is_anon q2 with + | true, true -> true + | false, true -> strictly_lt := true; lt + | true, false -> false + | false, false -> glob_quality_eq q1 q2) + ql1 ql2 + && CList.for_all2eq (fun u1 u2 -> match u1, u2 with | UAnonymous {rigid=true}, UAnonymous {rigid=true} -> true | UAnonymous {rigid=false}, UAnonymous {rigid=false} -> true | UAnonymous _, UAnonymous _ -> false | UNamed _, UAnonymous _ -> strictly_lt := true; lt | UAnonymous _, UNamed _ -> false - | UNamed _, UNamed _ -> glob_level_eq u1 u2) l1 l2 + | UNamed _, UNamed _ -> glob_level_eq u1 u2) ul1 ul2 (* Compute us1 <= us2, as a boolean *) let compare_glob_universe_instances_le us1 us2 = diff --git a/interp/notation_term.mli b/interp/notation_term.mli index 55bf27a261997..e8f75698d6bb1 100644 --- a/interp/notation_term.mli +++ b/interp/notation_term.mli @@ -21,10 +21,10 @@ open Glob_term type notation_constr = (* Part common to [glob_constr] and [cases_pattern] *) - | NRef of GlobRef.t * glob_level list option + | NRef of GlobRef.t * glob_instance option | NVar of Id.t | NApp of notation_constr * notation_constr list - | NProj of (Constant.t * glob_level list option) * notation_constr list * notation_constr + | NProj of (Constant.t * glob_instance option) * notation_constr list * notation_constr | NHole of Evar_kinds.t * Namegen.intro_pattern_naming_expr | NGenarg of Genarg.glob_generic_argument | NList of Id.t * Id.t * notation_constr * notation_constr * (* associativity: *) bool diff --git a/kernel/cClosure.ml b/kernel/cClosure.ml index d0d74ca943d19..46c6acc4a332c 100644 --- a/kernel/cClosure.ml +++ b/kernel/cClosure.ml @@ -417,14 +417,9 @@ let usubst_punivs (_,u) (v,u' as orig) = if UVars.Instance.is_empty u then orig else v, UVars.subst_instance_instance u u' -let usubst_sort (_,u) s = match s with - | Sorts.Type su -> - if UVars.Instance.is_empty u then s - else Sorts.(sort_of_univ (UVars.subst_instance_universe u su)) - | Sorts.QSort (q, v) -> - if UVars.Instance.is_empty u then s - else Sorts.qsort q (UVars.subst_instance_universe u v) - | Sorts.(SProp | Prop | Set) -> s +let usubst_sort (_,u) s = + if UVars.Instance.is_empty u then s + else UVars.subst_instance_sort u s (* Optimization: do not enclose variables in a closure. Makes variable access much faster *) diff --git a/kernel/cPrimitives.ml b/kernel/cPrimitives.ml index f738f1d447a21..6a9fcf679e234 100644 --- a/kernel/cPrimitives.ml +++ b/kernel/cPrimitives.ml @@ -279,7 +279,7 @@ and ind_or_type = | PITT_param : int -> ind_or_type (* DeBruijn index referring to prenex type quantifiers *) let one_univ = - AbstractContext.make Names.[|Name (Id.of_string "u")|] Constraints.empty + AbstractContext.make ([||],Names.[|Name (Id.of_string "u")|]) Constraints.empty let typ_univs (type a) (t : a prim_type) = match t with | PT_int63 -> AbstractContext.empty @@ -296,7 +296,7 @@ let types = let array_ty = PITT_type (PT_array, - (Instance.of_array [|Level.var 0|], + (Instance.of_array ([||],[|Level.var 0|]), PITT_param 1)) in function diff --git a/kernel/constant_typing.ml b/kernel/constant_typing.ml index 9c7c92562a4a0..57f072dd2cfc7 100644 --- a/kernel/constant_typing.ml +++ b/kernel/constant_typing.ml @@ -74,11 +74,11 @@ let rec unzip ctx j = unzip ctx { j with uj_val = mkApp (mkLambda (n,ty,j.uj_val),arg) } type typing_context = - TyCtx of Environ.env * unsafe_type_judgment * Id.Set.t * universe_level_subst * universes + TyCtx of Environ.env * unsafe_type_judgment * Id.Set.t * UVars.full_level_subst * universes let process_universes env = function | Entries.Monomorphic_entry -> - env, Univ.empty_level_subst, UVars.Instance.empty, Monomorphic + env, UVars.empty_full_subst, UVars.Instance.empty, Monomorphic | Entries.Polymorphic_entry uctx -> (** [ctx] must contain local universes, such that it has no impact on the rest of the graph (up to transitivity). *) @@ -148,7 +148,7 @@ let infer_primitive env { prim_entry_type = utyp; prim_entry_content = p; } = let make_univ_hyps = function | None -> Instance.empty - | Some us -> Instance.of_array us + | Some us -> us let infer_parameter ~sec_univs env entry = let env, usubst, _, univs = process_universes env entry.parameter_entry_universes in @@ -297,12 +297,12 @@ let check_delayed (type a) (handle : a effect_handler) tyenv (body : a proof_out let uctx = ContextSet.union uctx uctx' in let env, univs = match univs with | Monomorphic -> - assert (is_empty_level_subst usubst); + assert (UVars.is_empty_full_subst usubst); push_context_set uctx env, Opaqueproof.PrivateMonomorphic uctx | Polymorphic _ -> assert (Int.equal valid_signatures 0); push_subgraph uctx env, - let private_univs = on_snd (subst_univs_level_constraints usubst) uctx in + let private_univs = on_snd (subst_univs_level_constraints (snd usubst)) uctx in Opaqueproof.PrivatePolymorphic private_univs in (* Note: non-trivial trusted side-effects only in monomorphic case *) diff --git a/kernel/constant_typing.mli b/kernel/constant_typing.mli index 07e230be80a0f..a6260f58d6960 100644 --- a/kernel/constant_typing.mli +++ b/kernel/constant_typing.mli @@ -30,11 +30,11 @@ val translate_local_def : env -> Id.t -> section_def_entry -> val translate_local_assum : env -> types -> types * Sorts.relevance val translate_constant : - sec_univs:Univ.Level.t array option -> env -> Constant.t -> constant_entry -> + sec_univs:UVars.Instance.t option -> env -> Constant.t -> constant_entry -> 'a pconstant_body val translate_opaque : - sec_univs:Univ.Level.t array option -> env -> Constant.t -> 'a opaque_entry -> + sec_univs:UVars.Instance.t option -> env -> Constant.t -> 'a opaque_entry -> unit pconstant_body * typing_context val check_delayed : 'a effect_handler -> typing_context -> 'a proof_output -> (Constr.t * Univ.ContextSet.t Opaqueproof.delayed_universes) diff --git a/kernel/constr.ml b/kernel/constr.ml index 329f80520fe34..4b1cfd2d4dc09 100644 --- a/kernel/constr.ml +++ b/kernel/constr.ml @@ -1502,7 +1502,7 @@ let debug_print_fix pr_constr ((t,i),(lna,tl,bl)) = let pr_puniverses p u = if UVars.Instance.is_empty u then p - else Pp.(p ++ str"(*" ++ UVars.Instance.pr Univ.Level.raw_pr u ++ str"*)") + else Pp.(p ++ str"(*" ++ UVars.Instance.pr Sorts.QVar.raw_pr Univ.Level.raw_pr u ++ str"*)") let rec debug_print c = let open Pp in @@ -1563,7 +1563,7 @@ let rec debug_print c = | Float i -> str"Float("++str (Float64.to_string i) ++ str")" | Array(u,t,def,ty) -> str"Array(" ++ prlist_with_sep pr_comma debug_print (Array.to_list t) ++ str" | " ++ debug_print def ++ str " : " ++ debug_print ty - ++ str")@{" ++ UVars.Instance.pr Univ.Level.raw_pr u ++ str"}" + ++ str")@{" ++ UVars.Instance.pr Sorts.QVar.raw_pr Univ.Level.raw_pr u ++ str"}" and debug_invert = let open Pp in function | NoInvert -> mt() diff --git a/kernel/conversion.ml b/kernel/conversion.ml index 03341f79634a9..78c186625a2b2 100644 --- a/kernel/conversion.ml +++ b/kernel/conversion.ml @@ -175,9 +175,9 @@ let convert_instances_cumul pb var u u' (s, check) = let get_cumulativity_constraints cv_pb variance u u' = match cv_pb with | CONV -> - UVars.enforce_eq_variance_instances variance u u' Univ.Constraints.empty + UVars.enforce_eq_variance_instances variance u u' Sorts.QUConstraints.empty | CUMUL -> - UVars.enforce_leq_variance_instances variance u u' Univ.Constraints.empty + UVars.enforce_leq_variance_instances variance u u' Sorts.QUConstraints.empty let inductive_cumulativity_arguments (mind,ind) = mind.Declarations.mind_nparams + @@ -212,7 +212,8 @@ let convert_constructors_gen cmp_instances cmp_cumul (mind, ind, cns) nargs u1 u else (** By invariant, both constructors have a common supertype, so they are convertible _at that type_. *) - let variance = Array.make (UVars.Instance.length u1) UVars.Variance.Irrelevant in + (* NB: no variance for qualities *) + let variance = Array.make (snd (UVars.Instance.length u1)) UVars.Variance.Irrelevant in cmp_cumul CONV variance u1 u2 s let convert_constructors ctor nargs u1 u2 (s, check) = @@ -593,7 +594,7 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = (* Inductive types: MutInd MutConstruct Fix Cofix *) | (FInd (ind1,u1 as pind1), FInd (ind2,u2 as pind2)) -> if Ind.CanOrd.equal ind1 ind2 then - if UVars.Instance.length u1 = 0 || UVars.Instance.length u2 = 0 then + if UVars.Instance.is_empty u1 || UVars.Instance.is_empty u2 then let cuniv = convert_instances ~flex:false u1 u2 cuniv in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv else @@ -610,7 +611,7 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = | (FConstruct ((ind1,j1),u1 as pctor1), FConstruct ((ind2,j2),u2 as pctor2)) -> if Int.equal j1 j2 && Ind.CanOrd.equal ind1 ind2 then - if UVars.Instance.length u1 = 0 || UVars.Instance.length u2 = 0 then + if UVars.Instance.is_empty u1 || UVars.Instance.is_empty u2 then let cuniv = convert_instances ~flex:false u1 u2 cuniv in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv else @@ -767,7 +768,7 @@ and convert_stacks l2r infos lft1 lft2 stk1 stk2 cuniv = let mind = Environ.lookup_mind (fst ci1.ci_ind) (info_env infos.cnv_inf) in let mip = mind.Declarations.mind_packets.(snd ci1.ci_ind) in let cu = - if UVars.Instance.length u1 = 0 || UVars.Instance.length u2 = 0 then + if UVars.Instance.is_empty u1 || UVars.Instance.is_empty u2 then convert_instances ~flex:false u1 u2 cu else let u1 = CClosure.usubst_instance e1 u1 in @@ -894,8 +895,8 @@ let check_convert_instances ~flex:_ u u' univs = (* general conversion and inference functions *) let check_inductive_instances cv_pb variance u1 u2 univs = - let csts = get_cumulativity_constraints cv_pb variance u1 u2 in - if (UGraph.check_constraints csts univs) then univs + let qcsts, ucsts = get_cumulativity_constraints cv_pb variance u1 u2 in + if Sorts.QConstraints.trivial qcsts && (UGraph.check_constraints ucsts univs) then univs else raise NotConvertible let checked_universes = diff --git a/kernel/conversion.mli b/kernel/conversion.mli index 563a0a6aeeae2..95b28ea61439e 100644 --- a/kernel/conversion.mli +++ b/kernel/conversion.mli @@ -39,7 +39,7 @@ type ('a,'b) generic_conversion_function = env -> 'b universe_state -> 'a -> 'a type 'a infer_conversion_function = env -> 'a -> 'a -> Univ.Constraints.t val get_cumulativity_constraints : conv_pb -> UVars.Variance.t array -> - UVars.Instance.t -> UVars.Instance.t -> Univ.Constraints.t + UVars.Instance.t -> UVars.Instance.t -> Sorts.QUConstraints.t val inductive_cumulativity_arguments : (Declarations.mutual_inductive_body * int) -> int val constructor_cumulativity_arguments : (Declarations.mutual_inductive_body * int * int) -> int diff --git a/kernel/cooking.ml b/kernel/cooking.ml index d9dba971affaa..323128224dd57 100644 --- a/kernel/cooking.ml +++ b/kernel/cooking.ml @@ -348,17 +348,16 @@ let discharge_abstract_universe_context abstr auctx = together with the substitution [u₀ ↦ Var(0), ... ,uₙ₋₁ ↦ Var(n - 1), Var(0) ↦ Var(n), ..., Var(m - 1) ↦ Var (n + m - 1)]. *) - let open Univ in let n = AbstractContext.size abstr.abstr_auctx in - if (Int.equal n 0) then + if (UVars.eq_sizes n (0,0)) then (** Optimization: still need to take the union for the constraints between globals *) abstr, n, AbstractContext.union abstr.abstr_auctx auctx else let subst = abstr.abstr_ausubst in - let suff = Instance.of_array @@ Array.init (AbstractContext.size auctx) (fun i -> Level.var i) in + let suff = UVars.make_abstract_instance auctx in let ainst = Instance.append subst suff in let substf = make_instance_subst ainst in - let auctx = UVars.subst_univs_level_abstract_universe_context substf auctx in + let auctx = UVars.subst_univs_level_abstract_universe_context (snd substf) auctx in let auctx' = AbstractContext.union abstr.abstr_auctx auctx in { abstr with abstr_ausubst = ainst }, n, auctx' @@ -387,5 +386,5 @@ let lift_private_mono_univs info a = a let lift_private_poly_univs info (inst, cstrs) = - let cstrs = Univ.subst_univs_level_constraints (make_instance_subst info.abstr_info.abstr_ausubst) cstrs in + let cstrs = Univ.subst_univs_level_constraints (snd (make_instance_subst info.abstr_info.abstr_ausubst)) cstrs in (inst, cstrs) diff --git a/kernel/cooking.mli b/kernel/cooking.mli index f6810491999d7..9c80d203150f7 100644 --- a/kernel/cooking.mli +++ b/kernel/cooking.mli @@ -68,7 +68,7 @@ val lift_mono_univs : cooking_info -> Univ.ContextSet.t -> cooking_info * Univ.C (** The [int] is how many universes got discharged, ie size of returned context - size of input context. *) -val lift_poly_univs : cooking_info -> UVars.AbstractContext.t -> cooking_info * int * UVars.AbstractContext.t +val lift_poly_univs : cooking_info -> UVars.AbstractContext.t -> cooking_info * (int * int) * UVars.AbstractContext.t val lift_private_mono_univs : cooking_info -> 'a -> 'a diff --git a/kernel/declareops.ml b/kernel/declareops.ml index fba7ef5a6adb9..c10855e96bbf2 100644 --- a/kernel/declareops.ml +++ b/kernel/declareops.ml @@ -61,7 +61,7 @@ let universes_context = function let abstract_universes = function | Entries.Monomorphic_entry -> - Univ.empty_level_subst, Monomorphic + UVars.empty_full_subst, Monomorphic | Entries.Polymorphic_entry uctx -> let (inst, auctx) = UVars.abstract_universes uctx in let inst = UVars.make_instance_subst inst in diff --git a/kernel/declareops.mli b/kernel/declareops.mli index 2eeeb331c11c6..75e8fb0498368 100644 --- a/kernel/declareops.mli +++ b/kernel/declareops.mli @@ -17,7 +17,7 @@ open UVars val universes_context : universes -> AbstractContext.t -val abstract_universes : Entries.universes_entry -> Univ.universe_level_subst * universes +val abstract_universes : Entries.universes_entry -> UVars.full_level_subst * universes (** {6 Arities} *) diff --git a/kernel/discharge.ml b/kernel/discharge.ml index ef1f9109c1280..9d4758a1e9b30 100644 --- a/kernel/discharge.ml +++ b/kernel/discharge.ml @@ -24,12 +24,13 @@ let lift_univs info univ_hyps = function assert (UVars.Instance.is_empty univ_hyps); info, univ_hyps, Monomorphic | Polymorphic auctx -> - let info, n, auctx = lift_poly_univs info auctx in + let info, (qn, un), auctx = lift_poly_univs info auctx in let univ_hyps = let open UVars.Instance in - let us = to_array univ_hyps in - let us = Array.sub us 0 (Array.length us - n) in - of_array us + let qs, us = to_array univ_hyps in + let qs = Array.sub qs 0 (Array.length qs - qn) in + let us = Array.sub us 0 (Array.length us - un) in + of_array (qs,us) in info, univ_hyps, Polymorphic auctx @@ -174,9 +175,10 @@ let cook_inductive info mib = | None, None -> None, None | None, Some _ | Some _, None -> assert false | Some variance, Some sec_variance -> + (* no variance for qualities *) + let ulen = snd (AbstractContext.size (universe_context_of_cooking_info info)) in let sec_variance, newvariance = - Array.chop (Array.length sec_variance - AbstractContext.size (universe_context_of_cooking_info info)) - sec_variance + Array.chop (Array.length sec_variance - ulen) sec_variance in Some (Array.append newvariance variance), Some sec_variance in diff --git a/kernel/environ.ml b/kernel/environ.ml index 0927f2678b97c..2f189550acf10 100644 --- a/kernel/environ.ml +++ b/kernel/environ.ml @@ -81,6 +81,7 @@ type env = { env_nb_rel : int; env_universes : UGraph.t; env_universes_lbound : UGraph.Bound.t; + env_qualities : Sorts.QVar.Set.t; irr_constants : Cset_env.t; irr_inds : Indset_env.t; env_typing_flags : typing_flags; @@ -110,6 +111,7 @@ let empty_env = { env_nb_rel = 0; env_universes = UGraph.initial_universes; env_universes_lbound = UGraph.Bound.Set; + env_qualities = Sorts.QVar.Set.empty; irr_constants = Cset_env.empty; irr_inds = Indset_env.empty; env_typing_flags = Declareops.safe_flags Conv_oracle.empty; @@ -375,13 +377,28 @@ let check_constraints c env = UGraph.check_constraints c env.env_universes let add_universes ~lbound ~strict ctx g = + let _qs, us = UVars.Instance.to_array (UVars.UContext.instance ctx) in let g = Array.fold_left - (fun g v -> UGraph.add_universe ~lbound ~strict v g) - g (UVars.Instance.to_array (UVars.UContext.instance ctx)) + (fun g v -> UGraph.add_universe ~lbound ~strict v g) + g us in - UGraph.merge_constraints (UVars.UContext.constraints ctx) g + UGraph.merge_constraints (UVars.UContext.constraints ctx) g + +let add_qualities qs known = + let open Sorts.Quality in + Array.fold_left (fun known q -> + match q with + | QVar q -> + let known' = Sorts.QVar.Set.add q known in + let () = if known == known' then CErrors.anomaly Pp.(str"multiply bound sort quality") in + known' + | QConstant _ -> CErrors.anomaly Pp.(str "constant quality in ucontext")) + known + qs let push_context ?(strict=false) ctx env = + let qs, _us = UVars.Instance.to_array (UVars.UContext.instance ctx) in + let env = { env with env_qualities = add_qualities qs env.env_qualities } in map_universes (add_universes ~lbound:(universes_lbound env) ~strict ctx) env let add_universes_set ~lbound ~strict ctx g = diff --git a/kernel/environ.mli b/kernel/environ.mli index 177b1a20613f9..8ecdf2f0db188 100644 --- a/kernel/environ.mli +++ b/kernel/environ.mli @@ -73,6 +73,7 @@ type env = private { env_nb_rel : int; env_universes : UGraph.t; env_universes_lbound : UGraph.Bound.t; + env_qualities : Sorts.QVar.Set.t; irr_constants : Cset_env.t; irr_inds : Indset_env.t; env_typing_flags : typing_flags; diff --git a/kernel/indTyping.ml b/kernel/indTyping.ml index 8ddf7e74f5d71..c31e950968525 100644 --- a/kernel/indTyping.ml +++ b/kernel/indTyping.ml @@ -72,7 +72,7 @@ type univ_info = { ind_squashed : bool; ind_has_relevant_arg : bool; } let no_sort_variable () = - CErrors.anomaly (Pp.str "A sort variable was sent to the kernel") + CErrors.user_err (Pp.str "Sort variables not yet supported for the inductive's sort.") let check_univ_leq ?(is_real_arg=false) env u info = let ind_univ = info.ind_univ in @@ -81,6 +81,9 @@ let check_univ_leq ?(is_real_arg=false) env u info = else info in (* Inductive types provide explicit lifting from SProp to other universes, so allow SProp <= any. *) + (* qualities: for now ind_univ will be checked to be non-qsort in abstract_packets + so if u is also qsort it will be in "missing" + (or type in type is on) *) if Sorts.is_sprop u || UGraph.check_leq_sort (universes env) u ind_univ then info else if is_impredicative_sort env ind_univ then { info with ind_squashed = true } else {info with missing = u :: info.missing} @@ -261,8 +264,8 @@ let abstract_packets usubst ((arity,lc),(indices,splayed_lc),univ_info) = in let ind_univ = match univ_info.ind_univ with | Prop | SProp | Set -> univ_info.ind_univ - | Type u -> Sorts.sort_of_univ (Univ.subst_univs_level_universe usubst u) - | QSort _ -> no_sort_variable () + | Type u -> Sorts.sort_of_univ (Univ.subst_univs_level_universe (snd usubst) u) + | QSort _ -> no_sort_variable () (* TODO *) in let arity = @@ -341,11 +344,14 @@ let typecheck_inductive env ~sec_univs (mie:mutual_inductive_entry) = | Monomorphic_ind_entry | Template_ind_entry _ -> CErrors.user_err Pp.(str "Inductive cannot be both monomorphic and universe cumulative.") | Polymorphic_ind_entry uctx -> - let univs = Instance.to_array @@ UContext.instance uctx in + (* no variance for qualities *) + let _qualities, univs = Instance.to_array @@ UContext.instance uctx in let univs = Array.map2 (fun a b -> a,b) univs variances in let univs = match sec_univs with | None -> univs | Some sec_univs -> + (* no variance for qualities *) + let _, sec_univs = UVars.Instance.to_array sec_univs in let sec_univs = Array.map (fun u -> u, None) sec_univs in Array.append sec_univs univs in @@ -362,7 +368,7 @@ let typecheck_inductive env ~sec_univs (mie:mutual_inductive_entry) = (* Abstract universes *) let usubst, univs = match mie.mind_entry_universes with | Monomorphic_ind_entry | Template_ind_entry _ -> - Univ.empty_level_subst, Monomorphic + UVars.empty_full_subst, Monomorphic | Polymorphic_ind_entry uctx -> let (inst, auctx) = UVars.abstract_universes uctx in let inst = UVars.make_instance_subst inst in diff --git a/kernel/indTyping.mli b/kernel/indTyping.mli index 0d5e774f36fd8..426d8c76ec948 100644 --- a/kernel/indTyping.mli +++ b/kernel/indTyping.mli @@ -25,7 +25,7 @@ open Declarations * (indices * splayed constructor types) (both without params) * top allowed elimination *) -val typecheck_inductive : env -> sec_univs:Univ.Level.t array option +val typecheck_inductive : env -> sec_univs:UVars.Instance.t option -> mutual_inductive_entry -> env * universes diff --git a/kernel/indtypes.ml b/kernel/indtypes.ml index 8006f45da33f6..04f33375dccf0 100644 --- a/kernel/indtypes.ml +++ b/kernel/indtypes.ml @@ -542,13 +542,14 @@ let build_inductive env ~sec_univs names prv univs template variance | Some variance -> match sec_univs with | None -> Some variance, None | Some sec_univs -> - let nsec = Array.length sec_univs in - Some (Array.sub variance nsec (Array.length variance - nsec)), - Some (Array.sub variance 0 nsec) + (* no variance for qualities *) + let _nsecq, nsecu = UVars.Instance.length sec_univs in + Some (Array.sub variance nsecu (Array.length variance - nsecu)), + Some (Array.sub variance 0 nsecu) in let univ_hyps = match sec_univs with | None -> UVars.Instance.empty - | Some univs -> UVars.Instance.of_array univs + | Some univs -> univs in let mib = (* Build the mutual inductive *) diff --git a/kernel/indtypes.mli b/kernel/indtypes.mli index f194ab28837e3..08d18538fae68 100644 --- a/kernel/indtypes.mli +++ b/kernel/indtypes.mli @@ -14,5 +14,5 @@ open Environ open Entries (** Check an inductive. *) -val check_inductive : env -> sec_univs:Univ.Level.t array option +val check_inductive : env -> sec_univs:UVars.Instance.t option -> MutInd.t -> mutual_inductive_entry -> mutual_inductive_body diff --git a/kernel/inductive.ml b/kernel/inductive.ml index bb84c1d718885..0b8e00581d8d2 100644 --- a/kernel/inductive.ml +++ b/kernel/inductive.ml @@ -348,7 +348,7 @@ let expand_arity (mib, mip) (ind, u) params nas = let params = Vars.subst_of_rel_context_instance paramdecl params in let realdecls, _ = List.chop mip.mind_nrealdecls mip.mind_arity_ctxt in let self = - let u = Univ.(Instance.of_array (Array.init (Instance.length u) Level.var)) in + let u = Instance.abstract_instance (Instance.length u) in let args = Context.Rel.instance mkRel 0 mip.mind_arity_ctxt in mkApp (mkIndU (ind, u), args) in diff --git a/kernel/inferCumulativity.ml b/kernel/inferCumulativity.ml index e8a1357fff680..6835f87e48eba 100644 --- a/kernel/inferCumulativity.ml +++ b/kernel/inferCumulativity.ml @@ -108,13 +108,17 @@ open Inf let infer_generic_instance_eq variances u = Array.fold_left (fun variances u -> infer_level_eq u variances) - variances (Instance.to_array u) + variances + u + +(* no variance for qualities *) +let instance_univs u = snd (Instance.to_array u) let extend_con_instance cb u = - Instance.(of_array (Array.append (to_array cb.const_univ_hyps) (to_array u))) + (Array.append (instance_univs cb.const_univ_hyps) (instance_univs u)) let extend_ind_instance mib u = - Instance.(of_array (Array.append (to_array mib.mind_univ_hyps) (to_array u))) + (Array.append (instance_univs mib.mind_univ_hyps) (instance_univs u)) let extended_mind_variance mind = match mind.mind_variance, mind.mind_sec_variance with @@ -129,7 +133,9 @@ let infer_cumulative_ind_instance cv_pb mind_variance variances u = | _, Irrelevant -> variances | _, Invariant | CONV, Covariant -> infer_level_eq u variances | CUMUL, Covariant -> infer_level_leq u variances) - variances mind_variance (Instance.to_array u) + variances + mind_variance + u let infer_inductive_instance cv_pb env variances ind nargs u = let mind = Environ.lookup_mind (fst ind) env in @@ -235,7 +241,7 @@ let rec infer_fterm cv_pb infos variances hd stk = in infer_stack infos variances stk | FArray (u,elemsdef,ty) -> - let variances = infer_generic_instance_eq variances u in + let variances = infer_generic_instance_eq variances (instance_univs u) in let variances = infer_fterm CONV infos variances ty [] in let elems, def = Parray.to_array elemsdef in let variances = infer_fterm CONV infos variances def [] in diff --git a/kernel/nativecode.ml b/kernel/nativecode.ml index 0e17c8b56f3d2..8ae960e66ff2b 100644 --- a/kernel/nativecode.ml +++ b/kernel/nativecode.ml @@ -959,9 +959,9 @@ let get_evar_code i = MLprimitive (Get_evar, [|MLglobal symbols_tbl_name; MLint i|]) -let get_level_code i = - MLprimitive (Get_level, - [|MLglobal symbols_tbl_name; MLint i|]) +(* let get_level_code i = *) +(* MLprimitive (Get_level, *) +(* [|MLglobal symbols_tbl_name; MLint i|]) *) let get_proj_code i = MLprimitive (Get_proj, @@ -1054,20 +1054,21 @@ let cast_to_int v = | MLint _ -> v | _ -> MLprimitive (Val_to_int, [|v|]) -let ml_of_instance instance u = - let ml_of_level l = - match Univ.Level.var_index l with - | Some i -> - (* FIXME: use a proper cast function *) - let univ = MLprimitive (MLmagic, [|MLlocal (Option.get instance)|]) in - MLprimitive (MLarrayget, [|univ; MLint i|]) - | None -> let i = push_symbol (SymbLevel l) in get_level_code i - in - let u = UVars.Instance.to_array u in - if Array.is_empty u then [||] - else let u = Array.map ml_of_level u in - (* FIXME: use a proper cast function *) - [|MLprimitive (MLmagic, [|MLarray u|])|] +let ml_of_instance _instance u = + (* let ml_of_level l = *) + (* match Univ.Level.var_index l with *) + (* | Some i -> *) + (* (\* FIXME: use a proper cast function *\) *) + (* let univ = MLprimitive (MLmagic, [|MLlocal (Option.get instance)|]) in *) + (* MLprimitive (MLarrayget, [|univ; MLint i|]) *) + (* | None -> let i = push_symbol (SymbLevel l) in get_level_code i *) + (* in *) + if UVars.Instance.is_empty u then [||] + else assert false (* TODO handle sort poly *) + (* let u = UVars.Instance.to_array u in *) + (* let u = Array.map ml_of_level u in *) + (* (\* FIXME: use a proper cast function *\) *) + (* [|MLprimitive (MLmagic, [|MLarray u|])|] *) let compile_prim env decl cond paux = @@ -2013,7 +2014,7 @@ and compile_named env sigma univ auxdefs id = Glet(Gnamed id, MLprimitive (Mk_var id, [||]))::auxdefs let compile_constant env sigma con cb = - let no_univs = 0 = UVars.AbstractContext.size (Declareops.constant_polymorphic_context cb) in + let no_univs = (0,0) = UVars.AbstractContext.size (Declareops.constant_polymorphic_context cb) in begin match cb.const_body with | Def t -> let code = lambda_of_constr env sigma t in @@ -2074,7 +2075,7 @@ let compile_mind mb mind stack = let name = Gind ("", ind) in let accu = let args = - if Int.equal (UVars.AbstractContext.size u) 0 then + if (UVars.AbstractContext.size u) = (0,0) then [|get_ind_code j; MLarray [||]|] else [|get_ind_code j|] in diff --git a/kernel/nativevalues.ml b/kernel/nativevalues.ml index 575e422d03572..d3a3ac56809f4 100644 --- a/kernel/nativevalues.ml +++ b/kernel/nativevalues.ml @@ -170,24 +170,24 @@ let mk_rels_accu lvl len = let napply (f:t) (args: t array) = Array.fold_left (fun f a -> apply f a) f args -let mk_constant_accu kn u = - mk_accu (Aconstant (kn,UVars.Instance.of_array u)) - -let mk_ind_accu ind u = - mk_accu (Aind (ind,UVars.Instance.of_array u)) - -let mk_sort_accu s u = - let open Sorts in - match s with - | SProp | Prop | Set -> mk_accu (Asort s) - | Type s -> - let u = UVars.Instance.of_array u in - let s = Sorts.sort_of_univ (UVars.subst_instance_universe u s) in - mk_accu (Asort s) - | QSort (q, s) -> - let u = UVars.Instance.of_array u in - let s = Sorts.qsort q (UVars.subst_instance_universe u s) in - mk_accu (Asort s) +let mk_constant_accu _kn _u = assert false (* TODO handle sort poly *) + (* mk_accu (Aconstant (kn,UVars.Instance.of_array u)) *) + +let mk_ind_accu _ind _u = assert false (* TODO handle sort poly *) + (* mk_accu (Aind (ind,UVars.Instance.of_array u)) *) + +let mk_sort_accu _s _u = assert false (* TODO handle sort poly *) + (* let open Sorts in *) + (* match s with *) + (* | SProp | Prop | Set -> mk_accu (Asort s) *) + (* | Type s -> *) + (* let u = UVars.Instance.of_array u in *) + (* let s = Sorts.sort_of_univ (UVars.subst_instance_universe u s) in *) + (* mk_accu (Asort s) *) + (* | QSort (q, s) -> *) + (* let u = UVars.Instance.of_array u in *) + (* let s = Sorts.qsort q (UVars.subst_instance_universe u s) in *) + (* mk_accu (Asort s) *) let mk_var_accu id = mk_accu (Avar id) diff --git a/kernel/safe_typing.ml b/kernel/safe_typing.ml index 56813ae061667..bc1a9cb962024 100644 --- a/kernel/safe_typing.ml +++ b/kernel/safe_typing.ml @@ -521,12 +521,10 @@ let push_section_context uctx senv = let sections = get_section senv.sections in let sections = Section.push_local_universe_context uctx sections in let senv = { senv with sections=Some sections } in - let ctx = UVars.UContext.to_context_set uctx in - (* We check that the universes are fresh. FIXME: This should be done - implicitly, but we have to work around the API. *) - let () = assert (Univ.Level.Set.for_all (fun u -> not (Univ.Level.Set.mem u (fst senv.univ))) (fst ctx)) in + let _qualities, ctx = UVars.UContext.to_context_set uctx in + (* push_context checks freshness *) { senv with - env = Environ.push_context_set ~strict:false ctx senv.env; + env = Environ.push_context ~strict:false uctx senv.env; univ = Univ.ContextSet.union ctx senv.univ } (** {6 Insertion of new declarations to current environment } *) @@ -606,7 +604,7 @@ type global_declaration = type exported_opaque = { exp_handle : Opaqueproof.opaque_handle; exp_body : Constr.t; - exp_univs : int option; + exp_univs : (int * int) option; (* Minimal amount of data needed to rebuild the private universes. We enforce in the API that private constants have no internal constraints. *) } @@ -1293,6 +1291,8 @@ let start_library dir senv = let export ~output_native_objects senv dir = let () = check_current_library dir senv in + (* qualities are in the senv only during sections *) + let () = assert (Sorts.QVar.Set.is_empty senv.env.Environ.env_qualities) in let mp = senv.modpath in let str = NoFunctor (List.rev senv.revstruct) in let mb = diff --git a/kernel/section.ml b/kernel/section.ml index 9c06ddfdb56dd..926528e27fa80 100644 --- a/kernel/section.ml +++ b/kernel/section.ml @@ -32,7 +32,7 @@ type 'a t = { (** Global universes introduced in the section *) poly_universes : UContext.t; (** Universes local to the section *) - all_poly_univs : Univ.Level.t array; + all_poly_univs : Instance.t; (** All polymorphic universes, including from previous sections. *) has_poly_univs : bool; (** Are there polymorphic universes or constraints, including in previous sections. *) @@ -60,15 +60,12 @@ let push_local_universe_context ctx sec = else let sctx = sec.poly_universes in let poly_universes = UContext.union sctx ctx in - let all_poly_univs = - Array.append sec.all_poly_univs (Instance.to_array @@ UContext.instance ctx) - in + let all_poly_univs = Instance.append sec.all_poly_univs (UContext.instance ctx) in { sec with poly_universes; all_poly_univs; has_poly_univs = true } -let rec is_polymorphic_univ u sec = - let uctx = sec.poly_universes in - let here = Array.exists (fun u' -> Level.equal u u') (Instance.to_array (UContext.instance uctx)) in - here || Option.cata (is_polymorphic_univ u) false sec.prev +let is_polymorphic_univ u sec = + let _, us = Instance.to_array sec.all_poly_univs in + Array.exists (fun u' -> Level.equal u u') us let push_constraints uctx sec = if sec.has_poly_univs && @@ -87,7 +84,7 @@ let open_section ~custom prev = context = []; mono_universes = ContextSet.empty; poly_universes = UContext.empty; - all_poly_univs = Option.cata (fun sec -> sec.all_poly_univs) [| |] prev; + all_poly_univs = Option.cata (fun sec -> sec.all_poly_univs) Instance.empty prev; has_poly_univs = Option.cata has_poly_univs false prev; entries = []; expand_info_map = (Cmap.empty, Mindmap.empty); diff --git a/kernel/section.mli b/kernel/section.mli index 3bfc30cd9bc2d..0ba1d219e7e2c 100644 --- a/kernel/section.mli +++ b/kernel/section.mli @@ -59,7 +59,7 @@ val push_global : Environ.env -> poly:bool -> section_entry -> 'a t -> 'a t (** {6 Retrieving section data} *) -val all_poly_univs : 'a t -> Univ.Level.t array +val all_poly_univs : 'a t -> Instance.t (** Returns all polymorphic universes, including those from previous sections. Earlier sections are earlier in the array. diff --git a/kernel/sorts.ml b/kernel/sorts.ml index e471f6c9a31cb..b39a129fc26d8 100644 --- a/kernel/sorts.ml +++ b/kernel/sorts.ml @@ -32,6 +32,29 @@ struct | Var q -> Hashset.Combine.combinesmall 1 q | Unif (s,q) -> Hashset.Combine.(combinesmall 2 (combine (CString.hash s) q)) + module Hstruct = struct + type nonrec t = t + type u = unit + + let hashcons () = function + | Var _ as q -> q + | Unif (s,i) as q -> + let s' = CString.hcons s in + if s == s' then q else Unif (s',i) + + let eq a b = + match a, b with + | Var a, Var b -> Int.equal a b + | Unif (sa, ia), Unif (sb, ib) -> sa == sb && Int.equal ia ib + | (Var _ | Unif _), _ -> false + + let hash = hash + end + + module Hasher = Hashcons.Make(Hstruct) + + let hcons = Hashcons.simple_hcons Hasher.generate Hasher.hcons () + let compare a b = match a, b with | Var a, Var b -> Int.compare a b | Unif (s1,i1), Unif (s2,i2) -> @@ -61,39 +84,168 @@ struct end module Quality = struct - type t = QVar of QVar.t | QProp | QSProp | QType + type constant = QProp | QSProp | QType + type t = QVar of QVar.t | QConstant of constant + + let var i = QVar (QVar.make_var i) + + let var_index = function + | QVar q -> QVar.var_index q + | QConstant _ -> None + + module Constants = struct + let equal a b = match a, b with + | QProp, QProp | QSProp, QSProp | QType, QType -> true + | (QProp | QSProp | QType), _ -> false + + let compare a b = match a, b with + | QProp, QProp -> 0 + | QProp, _ -> -1 + | _, QProp -> 1 + | QSProp, QSProp -> 0 + | QSProp, _ -> -1 + | _, QSProp -> 1 + | QType, QType -> 0 + + let pr = function + | QProp -> Pp.str "Prop" + | QSProp -> Pp.str "SProp" + | QType -> Pp.str "Type" + + let hash = function + | QSProp -> 0 + | QProp -> 1 + | QType -> 2 + + end let equal a b = match a, b with | QVar a, QVar b -> QVar.equal a b - | QProp, QProp | QSProp, QSProp | QType, QType -> true - | (QVar _ | QProp | QSProp | QType), _ -> false + | QConstant a, QConstant b -> Constants.equal a b + | (QVar _ | QConstant _), _ -> false let compare a b = match a, b with | QVar a, QVar b -> QVar.compare a b | QVar _, _ -> -1 | _, QVar _ -> 1 - | QProp, QProp -> 0 - | QProp, _ -> -1 - | _, QProp -> 1 - | QSProp, QSProp -> 0 - | QSProp, _ -> -1 - | _, QSProp -> 1 - | QType, QType -> 0 + | QConstant a, QConstant b -> Constants.compare a b let pr prv = function | QVar v -> prv v - | QProp -> Pp.str "Prop" - | QSProp -> Pp.str "SProp" - | QType -> Pp.str "Type" + | QConstant q -> Constants.pr q let raw_pr q = pr QVar.raw_pr q + let hash = let open Hashset.Combine in function + | QConstant q -> Constants.hash q + | QVar q -> combinesmall 3 (QVar.hash q) + + let subst f = function + | QConstant _ as q -> q + | QVar qv as q -> + match f qv with + | QConstant _ as q -> q + | QVar qv' as q' -> + if qv == qv' then q else q' + + let subst_fn m v = + match QVar.Map.find_opt v m with + | Some v -> v + | None -> QVar v + + module Hstruct = struct + type nonrec t = t + type u = QVar.t -> QVar.t + + let hashcons hv = function + | QConstant _ as q -> q + | QVar qv as q -> + let qv' = hv qv in + if qv == qv' then q else QVar qv' + + let eq a b = + match a, b with + | QVar a, QVar b -> a == b + | QVar _, _ -> false + | (QConstant _), _ -> equal a b + + let hash = hash + end + + module Hasher = Hashcons.Make(Hstruct) + + let hcons = Hashcons.simple_hcons Hasher.generate Hasher.hcons QVar.hcons + module Self = struct type nonrec t = t let compare = compare end module Set = CSet.Make(Self) module Map = CMap.Make(Self) end +module QConstraint = struct + type kind = Equal | Leq + + let eq_kind : kind -> kind -> bool = (=) + let compare_kind : kind -> kind -> int = compare + + let pr_kind = function + | Equal -> Pp.str "=" + | Leq -> Pp.str "<=" + + type t = Quality.t * kind * Quality.t + + let equal (a,k,b) (a',k',b') = + eq_kind k k' && Quality.equal a a' && Quality.equal b b' + + let compare (a,k,b) (a',k',b') = + let c = compare_kind k k' in + if c <> 0 then c + else + let c = Quality.compare a a' in + if c <> 0 then c + else Quality.compare b b' + + let trivial (a,(Equal|Leq),b) = Quality.equal a b + + let pr prq (a,k,b) = + let open Pp in + hov 1 (Quality.pr prq a ++ spc() ++ pr_kind k ++ spc() ++ Quality.pr prq b) + + let raw_pr x = pr QVar.raw_pr x + +end + +module QConstraints = struct include CSet.Make(QConstraint) + let trivial = for_all QConstraint.trivial + + let pr prq c = + let open Pp in + v 0 (prlist_with_sep spc (fun (u1,op,u2) -> + hov 0 (Quality.pr prq u1 ++ QConstraint.pr_kind op ++ Quality.pr prq u2)) + (elements c)) + +end + +let enforce_eq_quality a b csts = + if Quality.equal a b then csts + else QConstraints.add (a,QConstraint.Equal,b) csts + +let enforce_leq_quality a b csts = + if Quality.equal a b then csts + else match a, b with + | Quality.(QConstant QProp), Quality.(QConstant QType) -> csts + | _ -> QConstraints.add (a,QConstraint.Leq,b) csts + +module QUConstraints = struct + + type t = QConstraints.t * Univ.Constraints.t + + let empty = QConstraints.empty, Univ.Constraints.empty + + let union (qcsts,ucsts) (qcsts',ucsts') = + QConstraints.union qcsts qcsts', Constraints.union ucsts ucsts' +end + type t = | SProp | Prop @@ -156,6 +308,22 @@ let levels s = match s with | Set -> Level.Set.singleton Level.set | Type u | QSort (_, u) -> Universe.levels u +let subst_fn (fq,fu) = function + | SProp | Prop | Set as s -> s + | Type v as s -> + let v' = fu v in + if v' == v then s else sort_of_univ v' + | QSort (q, v) as s -> + let open Quality in + match fq q with + | QVar q' -> + let v' = fu v in + if q' == q && v' == v then s + else qsort q' v' + | QConstant QSProp -> sprop + | QConstant QProp -> prop + | QConstant QType -> sort_of_univ (fu v) + let family = function | SProp -> InSProp | Prop -> InProp @@ -178,9 +346,18 @@ let family_compare a b = match a,b with | _, InType -> 1 | InQSort, InQSort -> 0 -let family_equal = (==) +let family_equal a b = match a, b with + | InSProp, InSProp | InProp, InProp | InSet, InSet | InType, InType -> true + | InQSort, InQSort -> true + | (InSProp | InProp | InSet | InType | InQSort), _ -> false -let family_leq a b = family_compare a b <= 0 +let family_leq a b = + family_equal a b + || match a, b with + | InSProp, _ -> true + | InProp, InSet -> true + | _, InType -> true + | _ -> false open Hashset.Combine @@ -239,6 +416,16 @@ let relevance_hash = function | Irrelevant -> 1 | RelevanceVar q -> Hashset.Combine.combinesmall 2 (QVar.hash q) +let relevance_subst_fn f = function + | Relevant | Irrelevant as r -> r + | RelevanceVar qv as r -> + let open Quality in + match f qv with + | QConstant QSProp -> Irrelevant + | QConstant (QProp | QType) -> Relevant + | QVar qv' -> + if qv' == qv then r else RelevanceVar qv' + let relevance_of_sort = function | SProp -> Irrelevant | Prop | Set | Type _ -> Relevant diff --git a/kernel/sorts.mli b/kernel/sorts.mli index 2d88d97b2c720..6a4584a748e51 100644 --- a/kernel/sorts.mli +++ b/kernel/sorts.mli @@ -26,6 +26,8 @@ sig val equal : t -> t -> bool val compare : t -> t -> int + val hash : t -> int + val raw_pr : t -> Pp.t val to_string : t -> string (** Debug printing *) @@ -36,9 +38,19 @@ sig end module Quality : sig - type t = - | QVar of QVar.t - | QProp | QSProp | QType + type constant = QProp | QSProp | QType + type t = QVar of QVar.t | QConstant of constant + + module Constants : sig + val equal : constant -> constant -> bool + val compare : constant -> constant -> int + val pr : constant -> Pp.t + end + + val var : int -> t + (** [var i] is [QVar (QVar.make_var i)] *) + + val var_index : t -> int option val equal : t -> t -> bool @@ -48,9 +60,57 @@ module Quality : sig val raw_pr : t -> Pp.t + val hash : t -> int + + val hcons : t -> t + + (* XXX Inconsistent naming: this one should be subst_fn *) + val subst : (QVar.t -> t) -> t -> t + + val subst_fn : t QVar.Map.t -> QVar.t -> t + module Set : CSig.SetS with type elt = t module Map : CMap.ExtS with type key = t and module Set := Set + +end + +module QConstraint : sig + type kind = Equal | Leq + + val pr_kind : kind -> Pp.t + + type t = Quality.t * kind * Quality.t + + val equal : t -> t -> bool + + val compare : t -> t -> int + + val trivial : t -> bool + + val pr : (QVar.t -> Pp.t) -> t -> Pp.t + + val raw_pr : t -> Pp.t +end + +module QConstraints : sig include CSig.SetS with type elt = QConstraint.t + + val trivial : t -> bool + + val pr : (QVar.t -> Pp.t) -> t -> Pp.t +end + +val enforce_eq_quality : Quality.t -> Quality.t -> QConstraints.t -> QConstraints.t + +val enforce_leq_quality : Quality.t -> Quality.t -> QConstraints.t -> QConstraints.t + +module QUConstraints : sig + + type t = QConstraints.t * Univ.Constraints.t + + val union : t -> t -> t + + val empty : t end type t = private @@ -88,13 +148,19 @@ val levels : t -> Univ.Level.Set.t val super : t -> t +val subst_fn : (QVar.t -> Quality.t) * (Univ.Universe.t -> Univ.Universe.t) + -> t -> t + (** On binders: is this variable proof relevant *) +(* TODO put in submodule or new file *) type relevance = Relevant | Irrelevant | RelevanceVar of QVar.t val relevance_hash : relevance -> int val relevance_equal : relevance -> relevance -> bool +val relevance_subst_fn : (QVar.t -> Quality.t) -> relevance -> relevance + val relevance_of_sort : t -> relevance val relevance_of_sort_family : family -> relevance diff --git a/kernel/type_errors.ml b/kernel/type_errors.ml index 7b062d7c750e3..f6c4fdfc2562f 100644 --- a/kernel/type_errors.ml +++ b/kernel/type_errors.ml @@ -70,7 +70,9 @@ type ('constr, 'types) ptype_error = | IllFormedRecBody of 'constr pguard_error * Name.t Context.binder_annot array * int * env * ('constr, 'types) punsafe_judgment array | IllTypedRecBody of int * Name.t Context.binder_annot array * ('constr, 'types) punsafe_judgment array * 'types array + | UnsatisfiedQConstraints of Sorts.QConstraints.t | UnsatisfiedConstraints of Constraints.t + | UndeclaredQualities of Sorts.QVar.Set.t | UndeclaredUniverse of Level.t | DisallowedSProp | BadBinderRelevance of Sorts.relevance * ('constr, 'types) Context.Rel.Declaration.pt @@ -148,9 +150,15 @@ let error_ill_formed_rec_body env why lna i fixenv vdefj = let error_ill_typed_rec_body env i lna vdefj vargs = raise (TypeError (env, IllTypedRecBody (i,lna,vdefj,vargs))) +let error_unsatisfied_qconstraints env c = + raise (TypeError (env, UnsatisfiedQConstraints c)) + let error_unsatisfied_constraints env c = raise (TypeError (env, UnsatisfiedConstraints c)) +let error_undeclared_qualities env l = + raise (TypeError (env, UndeclaredQualities l)) + let error_undeclared_universe env l = raise (TypeError (env, UndeclaredUniverse l)) @@ -195,7 +203,8 @@ let map_pguard_error f = function let map_ptype_error f = function | UnboundRel _ | UnboundVar _ | CaseOnPrivateInd _ -| UndeclaredUniverse _ | DisallowedSProp | UnsatisfiedConstraints _ +| UndeclaredQualities _ | UndeclaredUniverse _ | DisallowedSProp +| UnsatisfiedQConstraints _ | UnsatisfiedConstraints _ | ReferenceVariables _ | BadInvert | BadVariance _ as e -> e | NotAType j -> NotAType (on_judgment f j) | BadAssumption j -> BadAssumption (on_judgment f j) diff --git a/kernel/type_errors.mli b/kernel/type_errors.mli index 1de91ed2cc810..b3c59f281e8de 100644 --- a/kernel/type_errors.mli +++ b/kernel/type_errors.mli @@ -72,7 +72,9 @@ type ('constr, 'types) ptype_error = | IllFormedRecBody of 'constr pguard_error * Name.t Context.binder_annot array * int * env * ('constr, 'types) punsafe_judgment array | IllTypedRecBody of int * Name.t Context.binder_annot array * ('constr, 'types) punsafe_judgment array * 'types array + | UnsatisfiedQConstraints of Sorts.QConstraints.t | UnsatisfiedConstraints of Constraints.t + | UndeclaredQualities of Sorts.QVar.Set.t | UndeclaredUniverse of Level.t | DisallowedSProp | BadBinderRelevance of Sorts.relevance * ('constr, 'types) Context.Rel.Declaration.pt @@ -145,8 +147,12 @@ val error_ill_formed_rec_body : val error_ill_typed_rec_body : env -> int -> Name.t Context.binder_annot array -> unsafe_judgment array -> types array -> 'a +val error_unsatisfied_qconstraints : env -> Sorts.QConstraints.t -> 'a + val error_unsatisfied_constraints : env -> Constraints.t -> 'a +val error_undeclared_qualities : env -> Sorts.QVar.Set.t -> 'a + val error_undeclared_universe : env -> Level.t -> 'a val error_disallowed_sprop : env -> 'a diff --git a/kernel/typeops.ml b/kernel/typeops.ml index 658984fdcc31a..4408c10525663 100644 --- a/kernel/typeops.ml +++ b/kernel/typeops.ml @@ -42,6 +42,10 @@ let check_constraints cst env = if Environ.check_constraints cst env then () else error_unsatisfied_constraints env cst +let check_qconstraints qcst env = + if Sorts.QConstraints.trivial qcst then () + else error_unsatisfied_qconstraints env qcst + (* This should be a type (a priori without intention to be an assumption) *) let check_type env c t = match kind(Reduction.whd_all env t) with @@ -91,15 +95,8 @@ let warn_bad_relevance_binder ?loc env rlv bnd = | CWarnings.AsError -> error_bad_binder_relevance env rlv bnd -let anomaly_sort_variable q = - anomaly Pp.(str "The kernel received a unification sort variable " ++ Sorts.QVar.raw_pr q) - let check_assumption env x t ty = let r = x.binder_relevance in - let () = match r with - | Sorts.Relevant | Sorts.Irrelevant -> () - | Sorts.RelevanceVar q -> (* TODO check *) anomaly_sort_variable q - in let r' = infer_assumption env t ty in let x = if Sorts.relevance_equal r r' then x @@ -306,7 +303,7 @@ let type_of_prim_type _env u (type a) (prim : a CPrimitives.prim_type) = match p Constr.mkSet | CPrimitives.PT_array -> begin match UVars.Instance.to_array u with - | [|u|] -> + | [||], [|u|] -> let ty = Constr.mkType (Univ.Universe.make u) in Constr.mkProd(Context.anonR, ty , ty) | _ -> anomaly Pp.(str"universe instance for array type should have length 1") @@ -323,7 +320,7 @@ let type_of_float env = | None -> CErrors.user_err Pp.(str"The type float must be registered before this construction can be typechecked.") let type_of_array env u = - assert (UVars.Instance.length u = 1); + assert (UVars.Instance.length u = (0,1)); match env.retroknowledge.Retroknowledge.retro_array with | Some c -> mkConstU (c,u) | None -> CErrors.user_err Pp.(str"The type array must be registered before this construction can be typechecked.") @@ -510,11 +507,12 @@ let type_case_scrutinee env (mib, _mip) (u', largs) u pms (pctx, p) c = let () = Array.iter2 (fun p1 p2 -> Conversion.conv ~l2r:true env p1 p2) (Array.of_list params) pms in (* We use l2r:true for compat with old versions which used CONV with arguments flipped. It is relevant for performance eg in bedrock / Kami. *) - let cst = match mib.mind_variance with - | None -> UVars.enforce_eq_instances u u' Univ.Constraints.empty - | Some variance -> UVars.enforce_leq_variance_instances variance u' u Univ.Constraints.empty + let qcst, ucst = match mib.mind_variance with + | None -> UVars.enforce_eq_instances u u' Sorts.QUConstraints.empty + | Some variance -> UVars.enforce_leq_variance_instances variance u' u Sorts.QUConstraints.empty in - let () = check_constraints cst env in + let () = check_qconstraints qcst env in + let () = check_constraints ucst env in let subst = Vars.subst_of_rel_context_instance_list pctx (realargs @ [c]) in Vars.substl subst p @@ -530,10 +528,6 @@ let type_of_case env (mib, mip as specif) ci u pms (pctx, pnas, p, pt) iv c ct l error_elim_arity env (ind, u') c None in let rp = Sorts.relevance_of_sort sp in - let () = match ci.ci_relevance with - | Sorts.Relevant | Sorts.Irrelevant -> () - | Sorts.RelevanceVar q -> anomaly_sort_variable q - in let ci = if Sorts.relevance_equal ci.ci_relevance rp then ci else @@ -616,10 +610,6 @@ let type_of_global_in_context env r = let check_assum_annot env s x t = let r = x.binder_relevance in - let () = match r with - | Sorts.Relevant | Sorts.Irrelevant -> () - | Sorts.RelevanceVar q -> anomaly_sort_variable q - in let r' = Sorts.relevance_of_sort s in if Sorts.relevance_equal r' r then x @@ -629,10 +619,6 @@ let check_assum_annot env s x t = let check_let_annot env s x c t = let r = x.binder_relevance in - let () = match r with - | Sorts.Relevant | Sorts.Irrelevant -> () - | Sorts.RelevanceVar q -> anomaly_sort_variable q - in let r' = Sorts.relevance_of_sort s in if Sorts.relevance_equal r' r then x @@ -651,8 +637,7 @@ let rec execute env cstr = | Sort s -> let () = match s with | SProp -> if not (Environ.sprop_allowed env) then error_disallowed_sprop env - | QSort (q, _) -> anomaly_sort_variable q - | Prop | Set | Type _ -> () + | QSort _ | Prop | Set | Type _ -> () in cstr, type_of_sort s @@ -752,7 +737,7 @@ let rec execute env cstr = let realdecls, _ = List.chop mip.mind_nrealdecls mip.mind_arity_ctxt in let self = let args = Context.Rel.instance mkRel 0 mip.mind_arity_ctxt in - let inst = UVars.Instance.of_array (Array.init (UVars.Instance.length u) Level.var) in + let inst = UVars.Instance.(abstract_instance (length u)) in mkApp (mkIndU (ci.ci_ind, inst), args) in let realdecls = LocalAssum (Context.make_annot Anonymous mip.mind_relevance, self) :: realdecls in @@ -813,7 +798,7 @@ let rec execute env cstr = | Array(u,t,def,ty) -> (* ty : Type@{u} and all of t,def : ty *) let ulev = match UVars.Instance.to_array u with - | [|u|] -> u + | [||], [|u|] -> u | _ -> assert false in let ty',tyty = execute env ty in @@ -855,8 +840,15 @@ and execute_array env cs = (* Derived functions *) +let check_declared_qualities env qualities = + let module S = Sorts.QVar.Set in + let unknown = S.diff qualities env.env_qualities in + if S.is_empty unknown then () + else error_undeclared_qualities env unknown + let check_wellformed_universes env c = - let univs = universes_of_constr c in + let qualities, univs = universes_of_constr c in + check_declared_qualities env qualities; try UGraph.check_declared_universes (universes env) univs with UGraph.UndeclaredLevel u -> error_undeclared_universe env u diff --git a/kernel/uGraph.ml b/kernel/uGraph.ml index cfabfc707e2d5..2715adac1df9d 100644 --- a/kernel/uGraph.ml +++ b/kernel/uGraph.ml @@ -184,13 +184,14 @@ let constraints_for ~kept g = (** Subtyping of polymorphic contexts *) let check_subtype univs ctxT ctx = - if AbstractContext.size ctxT == AbstractContext.size ctx then + (* NB: size check is the only constraint on qualities *) + if eq_sizes (AbstractContext.size ctxT) (AbstractContext.size ctx) then let uctx = AbstractContext.repr ctx in let inst = UContext.instance uctx in let cst = UContext.constraints uctx in let cstT = UContext.constraints (AbstractContext.repr ctxT) in let push accu v = add_universe v ~lbound:Bound.Set ~strict:false accu in - let univs = Array.fold_left push univs (Instance.to_array inst) in + let univs = Array.fold_left push univs (snd (Instance.to_array inst)) in let univs = merge_constraints cstT univs in check_constraints cst univs else false @@ -198,13 +199,10 @@ let check_subtype univs ctxT ctx = (** Instances *) let check_eq_instances g t1 t2 = - let t1 = Instance.to_array t1 in - let t2 = Instance.to_array t2 in - t1 == t2 || - (Int.equal (Array.length t1) (Array.length t2) && - let rec aux i = - (Int.equal i (Array.length t1)) || (check_eq_level g t1.(i) t2.(i) && aux (i + 1)) - in aux 0) + let qt1, ut1 = Instance.to_array t1 in + let qt2, ut2 = Instance.to_array t2 in + CArray.equal Sorts.Quality.equal qt1 qt2 + && CArray.equal (check_eq_level g) ut1 ut2 let domain g = G.domain g.graph let choose p g u = G.choose p g.graph u diff --git a/kernel/uVars.ml b/kernel/uVars.ml index 9844abb2b41f1..1443a96869c2b 100644 --- a/kernel/uVars.ml +++ b/kernel/uVars.ml @@ -12,6 +12,8 @@ open Pp open Util open Univ +module Quality = Sorts.Quality + module Variance = struct (** A universe position in the instance given to a cumulative @@ -65,62 +67,73 @@ struct end module Instance : sig - type t = Level.t array + type t val empty : t val is_empty : t -> bool - val of_array : Level.t array -> t - val to_array : t -> Level.t array + val of_array : Quality.t array * Level.t array -> t + val to_array : t -> Quality.t array * Level.t array + + val abstract_instance : int * int -> t val append : t -> t -> t val equal : t -> t -> bool - val length : t -> int + val length : t -> int * int val hcons : t -> t val hash : t -> int val share : t -> t * int - val subst_fn : (Level.t -> Level.t) -> t -> t + val subst_fn : (Sorts.QVar.t -> Quality.t) * (Level.t -> Level.t) -> t -> t - val pr : (Level.t -> Pp.t) -> ?variance:Variance.t array -> t -> Pp.t - val levels : t -> Level.Set.t + val pr : (Sorts.QVar.t -> Pp.t) -> (Level.t -> Pp.t) -> ?variance:Variance.t array -> t -> Pp.t + val levels : t -> Quality.Set.t * Level.Set.t end = struct - type t = Level.t array + type t = Quality.t array * Level.t array - let empty : t = [||] + let empty : t = [||], [||] module HInstancestruct = struct type nonrec t = t - type u = Level.t -> Level.t + type u = (Quality.t -> Quality.t) * (Level.t -> Level.t) - let hashcons huniv a = - let len = Array.length a in - if Int.equal len 0 then empty + let hashcons (hqual, huniv) (aq, au as a) = + let qlen = Array.length aq in + let ulen = Array.length au in + if Int.equal qlen 0 && Int.equal ulen 0 then empty else begin - for i = 0 to len - 1 do - let x = Array.unsafe_get a i in + for i = 0 to qlen - 1 do + let x = Array.unsafe_get aq i in + let x' = hqual x in + if x == x' then () + else Array.unsafe_set aq i x' + done; + for i = 0 to ulen - 1 do + let x = Array.unsafe_get au i in let x' = huniv x in if x == x' then () - else Array.unsafe_set a i x' + else Array.unsafe_set au i x' done; a end let eq t1 t2 = - t1 == t2 || - (Int.equal (Array.length t1) (Array.length t2) && - let rec aux i = - (Int.equal i (Array.length t1)) || (t1.(i) == t2.(i) && aux (i + 1)) - in aux 0) + CArray.equal (==) (fst t1) (fst t2) + && CArray.equal (==) (snd t1) (snd t2) - let hash a = + let hash (aq,au) = let accu = ref 0 in - for i = 0 to Array.length a - 1 do - let l = Array.unsafe_get a i in + for i = 0 to Array.length aq - 1 do + let l = Array.unsafe_get aq i in + let h = Quality.hash l in + accu := Hashset.Combine.combine !accu h; + done; + for i = 0 to Array.length au - 1 do + let l = Array.unsafe_get au i in let h = Level.hash l in accu := Hashset.Combine.combine !accu h; done; @@ -131,67 +144,106 @@ struct module HInstance = Hashcons.Make(HInstancestruct) - let hcons = Hashcons.simple_hcons HInstance.generate HInstance.hcons Level.hcons + let hcons = Hashcons.simple_hcons HInstance.generate HInstance.hcons (Quality.hcons,Level.hcons) let hash = HInstancestruct.hash let share a = (hcons a, hash a) - let empty = hcons [||] + let empty = hcons empty + + let is_empty (x,y) = CArray.is_empty x && CArray.is_empty y + - let is_empty x = Int.equal (Array.length x) 0 + let append (xq,xu as x) (yq,yu as y) = + if is_empty x then y + else if is_empty y then x + else Array.append xq yq, Array.append xu yu - let append x y = - if Array.length x = 0 then y - else if Array.length y = 0 then x - else Array.append x y + let of_array a : t = a - let of_array a = - a + let to_array (a:t) = a - let to_array a = a + let abstract_instance (qlen,ulen) = + let qs = Array.init qlen Quality.var in + let us = Array.init ulen Level.var in + of_array (qs,us) - let length a = Array.length a + let length (aq,au) = Array.length aq, Array.length au - let subst_fn fn t = - let t' = CArray.Smart.map fn t in - if t' == t then t else of_array t' + let subst_fn (fq, fn) (q,u as orig) : t = + let q' = CArray.Smart.map (Quality.subst fq) q in + let u' = CArray.Smart.map fn u in + if q' == q && u' == u then orig else q', u' - let levels x = Array.fold_left (fun acc x -> Level.Set.add x acc) Level.Set.empty x + let levels (xq,xu) = + let q = Array.fold_left (fun acc x -> Quality.Set.add x acc) Quality.Set.empty xq in + let u = Array.fold_left (fun acc x -> Level.Set.add x acc) Level.Set.empty xu in + q, u - let pr prl ?variance = + let pr prq prl ?variance (q,u) = let ppu i u = let v = Option.map (fun v -> v.(i)) variance in pr_opt_no_spc Variance.pr v ++ prl u in - prvecti_with_sep spc ppu + (if Array.is_empty q then mt() else prvect_with_sep spc (Quality.pr prq) q ++ strbrk " | ") + ++ prvecti_with_sep spc ppu u - let equal t u = - t == u || - (Array.is_empty t && Array.is_empty u) || - (CArray.for_all2 Level.equal t u - (* Necessary as universe instances might come from different modules and - unmarshalling doesn't preserve sharing *)) + let equal (xq,xu) (yq,yu) = + CArray.equal Quality.equal xq yq + && CArray.equal Level.equal xu yu end -let enforce_eq_instances x y = - let ax = Instance.to_array x and ay = Instance.to_array y in - if Array.length ax != Array.length ay then - CErrors.anomaly (Pp.(++) (Pp.str "Invalid argument: enforce_eq_instances called with") - (Pp.str " instances of different lengths.")); - CArray.fold_right2 enforce_eq_level ax ay +let eq_sizes (a,b) (a',b') = Int.equal a a' && Int.equal b b' -let enforce_eq_variance_instances = Variance.eq_constraints -let enforce_leq_variance_instances = Variance.leq_constraints +type 'a quconstraint_function = 'a -> 'a -> Sorts.QUConstraints.t -> Sorts.QUConstraints.t + +let enforce_eq_instances x y (qcs, ucs as orig) = + let xq, xu = Instance.to_array x and yq, yu = Instance.to_array y in + if Array.length xq != Array.length yq || Array.length xu != Array.length yu then + CErrors.anomaly (Pp.(++) (Pp.str "Invalid argument: enforce_eq_instances called with") + (Pp.str " instances of different lengths.")); + let qcs' = CArray.fold_right2 Sorts.enforce_eq_quality xq yq qcs in + let ucs' = CArray.fold_right2 enforce_eq_level xu yu ucs in + if qcs' == qcs && ucs' == ucs then orig else qcs', ucs' + +let enforce_eq_variance_instances variances x y (qcs,ucs as orig) = + let xq, xu = Instance.to_array x and yq, yu = Instance.to_array y in + let qcs' = CArray.fold_right2 Sorts.enforce_eq_quality xq yq qcs in + let ucs' = Variance.eq_constraints variances xu yu ucs in + if qcs' == qcs && ucs' == ucs then orig else qcs', ucs' + +let enforce_leq_variance_instances variances x y (qcs,ucs as orig) = + let xq, xu = Instance.to_array x and yq, yu = Instance.to_array y in + (* no variance for quality variables -> enforce_eq *) + let qcs' = CArray.fold_right2 Sorts.enforce_eq_quality xq yq qcs in + let ucs' = Variance.leq_constraints variances xu yu ucs in + if qcs' == qcs && ucs' == ucs then orig else qcs', ucs' let subst_instance_level s l = match Level.var_index l with - | Some n -> s.(n) + | Some n -> (snd (Instance.to_array s)).(n) | None -> l +let subst_instance_qvar s v = + match Sorts.QVar.var_index v with + | Some n -> (fst (Instance.to_array s)).(n) + | None -> Quality.QVar v + +let subst_instance_quality s l = + match l with + | Quality.QVar v -> begin match Sorts.QVar.var_index v with + | Some n -> (fst (Instance.to_array s)).(n) + | None -> l + end + | _ -> l + let subst_instance_instance s i = - Array.Smart.map (fun l -> subst_instance_level s l) i + let qs, us = Instance.to_array i in + let qs' = Array.Smart.map (fun l -> subst_instance_quality s l) qs in + let us' = Array.Smart.map (fun l -> subst_instance_level s l) us in + if qs' == qs && us' == us then i else Instance.of_array (qs', us') let subst_instance_universe s univ = let f (v,n as vn) = @@ -204,10 +256,11 @@ let subst_instance_universe s univ = if u == u' then univ else Universe.unrepr u' -let subst_instance_sort u = let open Sorts in function - | SProp | Prop | Set as s -> s - | Type v -> sort_of_univ (subst_instance_universe u v) - | QSort (q, v) -> qsort q (subst_instance_universe u v) +let subst_instance_sort u s = + Sorts.subst_fn ((subst_instance_qvar u), (subst_instance_universe u)) s + +let subst_instance_relevance u r = + Sorts.relevance_subst_fn (subst_instance_qvar u) r let subst_instance_constraint s (u,d,v as c) = let u' = subst_instance_level s u in @@ -226,50 +279,62 @@ let in_punivs x = (x, Instance.empty) let eq_puniverses f (x, u) (y, u') = f x y && Instance.equal u u' +type bound_names = Names.Name.t array * Names.Name.t array + (** A context of universe levels with universe constraints, representing local universe variables and constraints *) module UContext = struct - type t = Names.Name.t array * Instance.t constrained + type t = bound_names * Instance.t constrained - let make names (univs, _ as x) = - assert (Array.length names = Array.length univs); + let make names (univs, _ as x) : t = + let qs, us = Instance.to_array univs in + assert (Array.length (fst names) = Array.length qs && Array.length(snd names) = Array.length us); (names, x) (** Universe contexts (variables as a list) *) - let empty = ([||], (Instance.empty, Constraints.empty)) + let empty = (([||], [||]), (Instance.empty, Constraints.empty)) let is_empty (_, (univs, cst)) = Instance.is_empty univs && Constraints.is_empty cst - let pr prl ?variance (_, (univs, cst) as ctx) = + let pr prq prl ?variance (_, (univs, cst) as ctx) = if is_empty ctx then mt() else - h (Instance.pr prl ?variance univs ++ str " |= ") ++ h (v 0 (Constraints.pr prl cst)) + h (Instance.pr prq prl ?variance univs ++ str " |= ") ++ h (v 0 (Constraints.pr prl cst)) - let hcons (names, (univs, cst)) = - (Array.map Names.Name.hcons names, (Instance.hcons univs, hcons_constraints cst)) + let hcons ((qnames, unames), (univs, cst)) = + ((Array.map Names.Name.hcons qnames, Array.map Names.Name.hcons unames), (Instance.hcons univs, hcons_constraints cst)) - let names (names, _) = names + let names ((names, _) : t) = names let instance (_, (univs, _cst)) = univs let constraints (_, (_univs, cst)) = cst - let union (na, (univs, cst)) (na', (univs', cst')) = - Array.append na na', (Instance.append univs univs', Constraints.union cst cst') + let union ((qna, una), (univs, cst)) ((qna', una'), (univs', cst')) = + (Array.append qna qna', Array.append una una'), (Instance.append univs univs', Constraints.union cst cst') let size (_,(x,_)) = Instance.length x - let refine_names names' (names, x) = + let refine_names (qnames',unames') ((qnames, unames), x) = let merge_names = Array.map2 Names.(fun old refined -> match refined with Anonymous -> old | Name _ -> refined) in - (merge_names names names', x) + ((merge_names qnames qnames', merge_names unames unames'), x) let sort_levels a = Array.sort Level.compare a; a - let of_context_set f (ctx, cst) = - let inst = Instance.of_array (sort_levels (Array.of_list (Level.Set.elements ctx))) in + let sort_qualities a = + Array.sort Quality.compare a; a + + let of_context_set f qctx (uctx, cst) = + let qctx = sort_qualities + (Array.map_of_list (fun q -> Quality.QVar q) + (Sorts.QVar.Set.elements qctx)) + in + let uctx = sort_levels (Array.of_list (Level.Set.elements uctx)) in + let inst = Instance.of_array (qctx, uctx) in (f inst, (inst, cst)) let to_context_set (_, (ctx, cst)) = - (Instance.levels ctx, cst) + let qctx, uctx = Instance.levels ctx in + qctx, (uctx, cst) end @@ -280,31 +345,37 @@ let hcons_universe_context = UContext.hcons module AbstractContext = struct - type t = Names.Name.t array constrained + type t = bound_names constrained let make names csts : t = names, csts - let repr (inst, cst) = - (inst, (Array.init (Array.length inst) (fun i -> Level.var i), cst)) - - let pr f ?variance ctx = UContext.pr f ?variance (repr ctx) - - let instantiate inst (u, cst) = - assert (Array.length u = Array.length inst); + let instantiate inst ((qnames,unames), cst) = + let q, u = Instance.to_array inst in + assert (Array.length q == Array.length qnames && Array.length u = Array.length unames); subst_instance_constraints inst cst let names (nas, _) = nas - let hcons (univs, cst) = - (Array.map Names.Name.hcons univs, hcons_constraints cst) + let hcons ((qnames,unames), cst) = + let qnames = Array.map Names.Name.hcons qnames in + let unames = Array.map Names.Name.hcons unames in + ((qnames, unames), hcons_constraints cst) - let empty = ([||], Constraints.empty) + let empty = (([||],[||]), Constraints.empty) - let is_empty (nas, cst) = Array.is_empty nas && Constraints.is_empty cst + let is_empty ((qnas,unas), cst) = + Array.is_empty qnas && Array.is_empty unas && Constraints.is_empty cst - let union (nas, cst) (nas', cst') = (Array.append nas nas', Constraints.union cst cst') + let union ((qnas,unas), cst) ((qnas',unas'), cst') = + ((Array.append qnas qnas', Array.append unas unas'), Constraints.union cst cst') - let size (nas, _) = Array.length nas + let size ((qnas,unas), _) = Array.length qnas, Array.length unas + + let repr (names, cst as self) : UContext.t = + let inst = Instance.abstract_instance (size self) in + (names, (inst, cst)) + + let pr prq pru ?variance ctx = UContext.pr prq pru ?variance (repr ctx) end @@ -319,29 +390,56 @@ let map_univ_abstracted f {univ_abstracted_value;univ_abstracted_binder} = let hcons_abstract_universe_context = AbstractContext.hcons -let subst_univs_level_instance subst i = - let i' = Instance.subst_fn (subst_univs_level_level subst) i in - if i == i' then i - else i' +type full_level_subst = Quality.t Sorts.QVar.Map.t * universe_level_subst + +let is_empty_full_subst (qsubst,usubst) = Sorts.QVar.Map.is_empty qsubst && is_empty_level_subst usubst + +let empty_full_subst = Sorts.QVar.Map.empty, empty_level_subst + +let subst_full_level_instance (qsubst,usubst) i = + let i' = Instance.subst_fn (Quality.subst_fn qsubst, subst_univs_level_level usubst) i in + if i == i' then i + else i' let subst_univs_level_abstract_universe_context subst (inst, csts) = inst, subst_univs_level_constraints subst csts +let subst_fn_of_qsubst qsubst qv = + match Sorts.QVar.Map.find_opt qv qsubst with + | None -> Quality.QVar qv + | Some q -> q + +let subst_full_level_sort (qsubst,usubst) s = + let fq qv = subst_fn_of_qsubst qsubst qv in + let fu u = subst_univs_level_universe usubst u in + Sorts.subst_fn (fq,fu) s + +let subst_full_level_relevance (qsubst,_) r = + Sorts.relevance_subst_fn (subst_fn_of_qsubst qsubst) r + let make_instance_subst i = - let arr = Instance.to_array i in + let qarr, uarr = Instance.to_array i in + let qsubst = + Array.fold_left_i (fun i acc l -> + let l = match l with Quality.QVar l -> l | _ -> assert false in + Sorts.QVar.Map.add l (Quality.var i) acc) + Sorts.QVar.Map.empty qarr + in + let usubst = Array.fold_left_i (fun i acc l -> Level.Map.add l (Level.var i) acc) - Level.Map.empty arr + Level.Map.empty uarr + in + qsubst, usubst -let make_abstract_instance (ctx, _) = - Array.init (Array.length ctx) (fun i -> Level.var i) +let make_abstract_instance ctx = + UContext.instance (AbstractContext.repr ctx) let abstract_universes uctx = let nas = UContext.names uctx in let instance = UContext.instance uctx in - let () = assert (Int.equal (Array.length nas) (Instance.length instance)) in let subst = make_instance_subst instance in - let cstrs = subst_univs_level_constraints subst + let cstrs = subst_univs_level_constraints (snd subst) (UContext.constraints uctx) in let ctx = (nas, cstrs) in diff --git a/kernel/uVars.mli b/kernel/uVars.mli index 4efae21d52b3d..cc83dc0c90913 100644 --- a/kernel/uVars.mli +++ b/kernel/uVars.mli @@ -9,6 +9,7 @@ (************************************************************************) open Univ +open Sorts (** {6 Support for universe polymorphism } *) @@ -43,8 +44,11 @@ sig val empty : t val is_empty : t -> bool - val of_array : Level.t array -> t - val to_array : t -> Level.t array + val of_array : Quality.t array * Level.t array -> t + val to_array : t -> Quality.t array * Level.t array + + val abstract_instance : int * int -> t + (** Instance of the given size made of QVar/Level.var *) val append : t -> t -> t (** To concatenate two instances, used for discharge *) @@ -52,7 +56,7 @@ sig val equal : t -> t -> bool (** Equality *) - val length : t -> int + val length : t -> int * int (** Instance length *) val hcons : t -> t @@ -64,18 +68,27 @@ sig val share : t -> t * int (** Simultaneous hash-consing and hash-value computation *) - val pr : (Level.t -> Pp.t) -> ?variance:Variance.t array -> t -> Pp.t + val pr : (QVar.t -> Pp.t) -> (Level.t -> Pp.t) -> ?variance:Variance.t array -> t -> Pp.t (** Pretty-printing, no comments *) - val levels : t -> Level.Set.t + val levels : t -> Quality.Set.t * Level.Set.t (** The set of levels in the instance *) + val subst_fn + : (QVar.t -> Quality.t) * (Level.t -> Level.t) + -> t -> t + end -val enforce_eq_instances : Instance.t constraint_function +val eq_sizes : int * int -> int * int -> bool +(** Convenient function to compare the result of Instance.length, UContext.size etc *) + +type 'a quconstraint_function = 'a -> 'a -> Sorts.QUConstraints.t -> Sorts.QUConstraints.t + +val enforce_eq_instances : Instance.t quconstraint_function -val enforce_eq_variance_instances : Variance.t array -> Instance.t constraint_function -val enforce_leq_variance_instances : Variance.t array -> Instance.t constraint_function +val enforce_eq_variance_instances : Variance.t array -> Instance.t quconstraint_function +val enforce_leq_variance_instances : Variance.t array -> Instance.t quconstraint_function type 'a puniverses = 'a * Instance.t val out_punivs : 'a puniverses -> 'a @@ -83,6 +96,8 @@ val in_punivs : 'a -> 'a puniverses val eq_puniverses : ('a -> 'a -> bool) -> 'a puniverses -> 'a puniverses -> bool +type bound_names = Names.Name.t array * Names.Name.t array + (** A vector of universe levels with universe Constraints.t, representing local universe variables and associated Constraints.t; the names are user-facing names for printing *) @@ -91,7 +106,7 @@ module UContext : sig type t - val make : Names.Name.t array -> Instance.t constrained -> t + val make : bound_names -> Instance.t constrained -> t val empty : t val is_empty : t -> bool @@ -102,22 +117,25 @@ sig val union : t -> t -> t (** Keeps the order of the instances *) - val size : t -> int + val size : t -> int * int (** The number of universes in the context *) - val names : t -> Names.Name.t array + val names : t -> bound_names (** Return the user names of the universes *) - val refine_names : Names.Name.t array -> t -> t + val refine_names : bound_names -> t -> t (** Use names to name the possibly yet unnamed universes *) val sort_levels : Level.t array -> Level.t array (** Arbitrary choice of linear order of the variables *) - val of_context_set : (Instance.t -> Names.Name.t array) -> ContextSet.t -> t + val sort_qualities : Quality.t array -> Quality.t array + (** Arbitrary choice of linear order of the variables *) + + val of_context_set : (Instance.t -> bound_names) -> QVar.Set.t -> ContextSet.t -> t (** Build a vector of universe levels assuming a function generating names *) - val to_context_set : t -> ContextSet.t + val to_context_set : t -> Quality.Set.t * ContextSet.t (** Discard the names and order of the universes *) end @@ -136,7 +154,7 @@ sig use de Bruijn indices *) - val make : Names.Name.t array -> Constraints.t -> t + val make : bound_names -> Constraints.t -> t (** Build an abstract context. Constraints may be between universe variables. *) @@ -147,7 +165,7 @@ sig val empty : t val is_empty : t -> bool - val size : t -> int + val size : t -> int * int val union : t -> t -> t (** The constraints are expected to be relative to the concatenated set of universes *) @@ -155,7 +173,7 @@ sig val instantiate : Instance.t -> t -> Constraints.t (** Generate the set of instantiated Constraints.t **) - val names : t -> Names.Name.t array + val names : t -> bound_names (** Return the names of the bound universe variables *) end @@ -170,17 +188,30 @@ val map_univ_abstracted : ('a -> 'b) -> 'a univ_abstracted -> 'b univ_abstracted (** {6 Substitution} *) +type full_level_subst = Quality.t QVar.Map.t * universe_level_subst + +val empty_full_subst : full_level_subst + +val is_empty_full_subst : full_level_subst -> bool + val subst_univs_level_abstract_universe_context : universe_level_subst -> AbstractContext.t -> AbstractContext.t -val subst_univs_level_instance : universe_level_subst -> Instance.t -> Instance.t +(** There are no constraints on qualities, so this only needs a subst for univs *) + +val subst_full_level_instance : full_level_subst -> Instance.t -> Instance.t (** Level to universe substitutions. *) +val subst_full_level_sort : full_level_subst -> Sorts.t -> Sorts.t + +val subst_full_level_relevance : full_level_subst -> Sorts.relevance -> Sorts.relevance + (** Substitution of instances *) val subst_instance_instance : Instance.t -> Instance.t -> Instance.t val subst_instance_universe : Instance.t -> Universe.t -> Universe.t val subst_instance_sort : Instance.t -> Sorts.t -> Sorts.t +val subst_instance_relevance : Instance.t -> Sorts.relevance -> Sorts.relevance -val make_instance_subst : Instance.t -> universe_level_subst +val make_instance_subst : Instance.t -> full_level_subst (** Creates [u(0) ↦ 0; ...; u(n-1) ↦ n - 1] out of [u(0); ...; u(n - 1)] *) val abstract_universes : UContext.t -> Instance.t * AbstractContext.t @@ -190,9 +221,9 @@ val make_abstract_instance : AbstractContext.t -> Instance.t (** {6 Pretty-printing of universes. } *) -val pr_universe_context : (Level.t -> Pp.t) -> ?variance:Variance.t array -> +val pr_universe_context : (QVar.t -> Pp.t) -> (Level.t -> Pp.t) -> ?variance:Variance.t array -> UContext.t -> Pp.t -val pr_abstract_universe_context : (Level.t -> Pp.t) -> ?variance:Variance.t array -> +val pr_abstract_universe_context : (QVar.t -> Pp.t) -> (Level.t -> Pp.t) -> ?variance:Variance.t array -> AbstractContext.t -> Pp.t (** {6 Hash-consing } *) diff --git a/kernel/univ.ml b/kernel/univ.ml index 6fc31ecd911d9..17294b3b82fbb 100644 --- a/kernel/univ.ml +++ b/kernel/univ.ml @@ -633,9 +633,6 @@ type universe_context_set = ContextSet.t (** A value in a universe context (resp. context set). *) type 'a in_universe_context_set = 'a * universe_context_set -let extend_in_context_set (a, ctx) ctx' = - (a, ContextSet.union ctx ctx') - (** Substitutions. *) let empty_level_subst = Level.Map.empty diff --git a/kernel/univ.mli b/kernel/univ.mli index 0a401c7774606..48d34d939d00c 100644 --- a/kernel/univ.mli +++ b/kernel/univ.mli @@ -232,9 +232,6 @@ end (** A value in a universe context set. *) type 'a in_universe_context_set = 'a * ContextSet.t -val extend_in_context_set : 'a in_universe_context_set -> ContextSet.t -> - 'a in_universe_context_set - (** {6 Substitution} *) type universe_level_subst = Level.t Level.Map.t diff --git a/kernel/vars.ml b/kernel/vars.ml index d4f5a3015bb95..2a34c8c6ab292 100644 --- a/kernel/vars.ml +++ b/kernel/vars.ml @@ -275,12 +275,97 @@ let smash_rel_context sign = (** Universe substitutions *) open Constr +let map_annot_relevance f na = + let open Context in + let r = na.binder_relevance in + let r' = f r in + if r' == r then na else { na with binder_relevance = r' } + +let map_case_info_relevance f ci = + let r = ci.ci_relevance in + let r' = f r in + if r' == r then ci else { ci with ci_relevance = r' } + +let map_case_under_context_relevance f (nas,x as v) = + let nas' = CArray.Smart.map (map_annot_relevance f) nas in + if nas' == nas then v else (nas',x) + +let map_rec_declaration_relevance f (i,(nas,x,y) as v) = +let nas' = CArray.Smart.map (map_annot_relevance f) nas in + if nas' == nas then v else (i,(nas',x,y)) + +let map_constr_relevance f c = + match kind c with + | Rel _ | Var _ | Meta _ | Evar _ + | Sort _ | Cast _ | App _ + | Const _ | Ind _ | Construct _ | Proj _ + | Int _ | Float _ | Array _ -> c + + | Prod (na,x,y) -> + let na' = map_annot_relevance f na in + if na' == na then c else mkProd (na',x,y) + + | Lambda (na,x,y) -> + let na' = map_annot_relevance f na in + if na' == na then c else mkLambda (na',x,y) + + | LetIn (na,x,y,z) -> + let na' = map_annot_relevance f na in + if na' == na then c else mkLetIn (na',x,y,z) + + | Case (ci,u,params,ret,iv,v,brs) -> + let ci' = map_case_info_relevance f ci in + let ret' = map_case_under_context_relevance f ret in + let brs' = CArray.Smart.map (map_case_under_context_relevance f) brs in + if ci' == ci && ret' == ret && brs' == brs then c + else mkCase (ci',u,params,ret',iv,v,brs') + + | Fix data -> + let data' = map_rec_declaration_relevance f data in + if data' == data then c else mkFix data' + + | CoFix data -> + let data' = map_rec_declaration_relevance f data in + if data' == data then c else mkCoFix data' + +let fold_annot_relevance f acc na = + f acc na.Context.binder_relevance + +let fold_case_under_context_relevance f acc (nas,_) = + Array.fold_left (fold_annot_relevance f) acc nas + +let fold_rec_declaration_relevance f acc (nas,_,_) = + Array.fold_left (fold_annot_relevance f) acc nas + +let fold_constr_relevance f acc c = + match kind c with + | Rel _ | Var _ | Meta _ | Evar _ + | Sort _ | Cast _ | App _ + | Const _ | Ind _ | Construct _ | Proj _ + | Int _ | Float _ | Array _ -> acc + + | Prod (na,_,_) | Lambda (na,_,_) | LetIn (na,_,_,_) -> + fold_annot_relevance f acc na + + | Case (ci,_u,_params,ret,_iv,_v,brs) -> + let acc = f acc ci.ci_relevance in + let acc = fold_case_under_context_relevance f acc ret in + let acc = CArray.fold_left (fold_case_under_context_relevance f) acc brs in + acc + + | Fix (_,data) + | CoFix (_,data) -> + fold_rec_declaration_relevance f acc data + let subst_univs_level_constr subst c = - if Univ.is_empty_level_subst subst then c + if UVars.is_empty_full_subst subst then c else - let f = UVars.subst_univs_level_instance subst in + let f = UVars.subst_full_level_instance subst in + (* XXX shouldn't Constr.map return the pointer equal term when unchanged instead? *) let changed = ref false in let rec aux t = + let t' = map_constr_relevance (UVars.subst_full_level_relevance subst) t in + let t = if t' == t then t else (changed := true; t') in match kind t with | Const (c, u) -> if UVars.Instance.is_empty u then t @@ -300,10 +385,11 @@ let subst_univs_level_constr subst c = let u' = f u in if u' == u then t else (changed := true; mkConstructU (c, u')) - | Sort (Sorts.Type u) -> - let u' = Univ.subst_univs_level_universe subst u in - if u' == u then t else - (changed := true; mkSort (Sorts.sort_of_univ u')) + | Sort s -> + let s' = UVars.subst_full_level_sort subst s in + if s' == s then t + else + (changed := true; mkSort s') | Case (ci, u, pms, p, CaseInvert {indices}, c, br) -> if UVars.Instance.is_empty u then Constr.map aux t @@ -341,6 +427,7 @@ let subst_instance_constr subst c = else let f u = UVars.subst_instance_instance subst u in let rec aux t = + let t = map_constr_relevance (UVars.subst_instance_relevance subst) t in match kind t with | Const (c, u) -> if UVars.Instance.is_empty u then t @@ -360,10 +447,9 @@ let subst_instance_constr subst c = let u' = f u in if u' == u then t else (mkConstructU (c, u')) - | Sort (Sorts.Type u) -> - let u' = UVars.subst_instance_universe subst u in - if u' == u then t else - (mkSort (Sorts.sort_of_univ u')) + | Sort s -> + let s' = UVars.subst_instance_sort subst s in + if s' == s then t else mkSort s' | Case (ci, u, pms, p, iv, c, br) -> let u' = f u in @@ -384,29 +470,42 @@ let subst_instance_constr subst c = let univ_instantiate_constr u c = let open UVars in - assert (Int.equal (Instance.length u) (AbstractContext.size c.univ_abstracted_binder)); + assert (UVars.eq_sizes (Instance.length u) (AbstractContext.size c.univ_abstracted_binder)); subst_instance_constr u c.univ_abstracted_value let subst_instance_context s ctx = if UVars.Instance.is_empty s then ctx else Context.Rel.map (fun x -> subst_instance_constr s x) ctx +let add_universes_of_instance (qs,us) u = + let qs', us' = UVars.Instance.levels u in + let qs = Sorts.Quality.(Set.fold (fun q qs -> match q with + | QVar q -> Sorts.QVar.Set.add q qs + | QConstant _ -> qs) + qs' qs) + in + qs, Univ.Level.Set.union us us' + +let add_relevance (qs,us as v) = let open Sorts in function + | Irrelevant | Relevant -> v + | RelevanceVar q -> QVar.Set.add q qs, us + let universes_of_constr c = let open Univ in - let open UVars in let rec aux s c = + let s = fold_constr_relevance add_relevance s c in match kind c with - | Const (_c, u) -> - Level.Set.fold Level.Set.add (Instance.levels u) s - | Ind ((_mind,_), u) | Construct (((_mind,_),_), u) -> - Level.Set.fold Level.Set.add (Instance.levels u) s - | Sort u when not (Sorts.is_small u) -> - Level.Set.fold Level.Set.add (Sorts.levels u) s + | Const (_, u) | Ind (_, u) | Construct (_,u) -> add_universes_of_instance s u + | Sort (Sorts.Type u) -> + Util.on_snd (Level.Set.union (Universe.levels u)) s + | Sort (Sorts.QSort (q,u)) -> + let qs, us = s in + Sorts.QVar.Set.add q qs, Level.Set.union us (Universe.levels u) | Array (u,_,_,_) -> - let s = Level.Set.fold Level.Set.add (Instance.levels u) s in + let s = add_universes_of_instance s u in Constr.fold aux s c | Case (_, u, _, _, _,_ ,_) -> - let s = Level.Set.fold Level.Set.add (Instance.levels u) s in + let s = add_universes_of_instance s u in Constr.fold aux s c | _ -> Constr.fold aux s c - in aux Level.Set.empty c + in aux (Sorts.QVar.Set.empty,Level.Set.empty) c diff --git a/kernel/vars.mli b/kernel/vars.mli index 5ca022e87f26a..a0917ab30bd37 100644 --- a/kernel/vars.mli +++ b/kernel/vars.mli @@ -186,13 +186,12 @@ val smash_rel_context : rel_context -> rel_context (** {3 Substitution of universes} *) -open Univ open UVars (** Level substitutions for polymorphism. *) -val subst_univs_level_constr : universe_level_subst -> constr -> constr -val subst_univs_level_context : Univ.universe_level_subst -> Constr.rel_context -> Constr.rel_context +val subst_univs_level_constr : full_level_subst -> constr -> constr +val subst_univs_level_context : full_level_subst -> Constr.rel_context -> Constr.rel_context (** Instance substitution for polymorphism. *) val subst_instance_constr : Instance.t -> constr -> constr @@ -201,7 +200,7 @@ val subst_instance_context : Instance.t -> Constr.rel_context -> Constr.rel_cont val univ_instantiate_constr : Instance.t -> constr univ_abstracted -> constr (** Ignores the constraints carried by [univ_abstracted]. *) -val universes_of_constr : constr -> Univ.Level.Set.t +val universes_of_constr : constr -> Sorts.QVar.Set.t * Univ.Level.Set.t (** {3 Low-level cached lift type} *) diff --git a/kernel/vconv.ml b/kernel/vconv.ml index d0a29aa905914..608edc227ff9c 100644 --- a/kernel/vconv.ml +++ b/kernel/vconv.ml @@ -91,22 +91,22 @@ and conv_atom env pb k a1 stk1 a2 stk2 cu = match a1, a2 with | Aind ((mi,_i) as ind1) , Aind ind2 -> if Ind.CanOrd.equal ind1 ind2 && compare_stack stk1 stk2 then - let ulen = UVars.AbstractContext.size (Environ.mind_context env mi) in + let qlen, ulen = UVars.AbstractContext.size (Environ.mind_context env mi) in + let ulen = qlen + ulen in if ulen = 0 then conv_stack env k stk1 stk2 cu - else + else begin match stk1 , stk2 with | Zapp args1 :: stk1' , Zapp args2 :: stk2' -> assert (ulen <= nargs args1); assert (ulen <= nargs args2); - let u1 = Array.init ulen (fun i -> uni_lvl_val (arg args1 i)) in - let u2 = Array.init ulen (fun i -> uni_lvl_val (arg args2 i)) in - let u1 = UVars.Instance.of_array u1 in - let u2 = UVars.Instance.of_array u2 in + let u1 = instance_of_args ulen args1 in + let u2 = instance_of_args ulen args2 in let cu = convert_instances ~flex:false u1 u2 cu in conv_arguments env ~from:ulen k args1 args2 (conv_stack env k stk1' stk2' cu) | _, _ -> assert false (* Should not happen if problem is well typed *) + end else raise NotConvertible | Aid ik1, Aid ik2 -> if Vmvalues.eq_id_key ik1 ik2 && compare_stack stk1 stk2 then diff --git a/kernel/vmbytecodes.ml b/kernel/vmbytecodes.ml index 09473023bbfce..e69b760364287 100644 --- a/kernel/vmbytecodes.ml +++ b/kernel/vmbytecodes.ml @@ -77,6 +77,7 @@ type fv_elem = | FVnamed of Id.t | FVrel of int | FVuniv_var of int + | FVqual_var of int | FVevar of Evar.t type fv = fv_elem array @@ -99,6 +100,7 @@ let pp_fv_elem = function | FVnamed id -> str "FVnamed(" ++ Id.print id ++ str ")" | FVrel i -> str "Rel(" ++ int i ++ str ")" | FVuniv_var v -> str "FVuniv(" ++ int v ++ str ")" + | FVqual_var v -> str "FVqual(" ++ int v ++ str ")" | FVevar e -> str "FVevar(" ++ int (Evar.repr e) ++ str ")" let rec pp_instr i = diff --git a/kernel/vmbytecodes.mli b/kernel/vmbytecodes.mli index ae8393147aaf6..ef6d9b85d07ff 100644 --- a/kernel/vmbytecodes.mli +++ b/kernel/vmbytecodes.mli @@ -78,6 +78,7 @@ type fv_elem = FVnamed of Id.t | FVrel of int | FVuniv_var of int +| FVqual_var of int | FVevar of Evar.t type fv = fv_elem array diff --git a/kernel/vmbytegen.ml b/kernel/vmbytegen.ml index 27ae2f1496d09..2360984622e23 100644 --- a/kernel/vmbytegen.ml +++ b/kernel/vmbytegen.ml @@ -103,14 +103,17 @@ type t = fv_elem let compare e1 e2 = match e1, e2 with | FVnamed id1, FVnamed id2 -> Id.compare id1 id2 -| FVnamed _, (FVrel _ | FVuniv_var _ | FVevar _) -> -1 -| FVrel _, FVnamed _ -> 1 +| FVnamed _, _ -> -1 +| _, FVnamed _ -> 1 | FVrel r1, FVrel r2 -> Int.compare r1 r2 -| FVrel _, (FVuniv_var _ | FVevar _) -> -1 +| FVrel _, _ -> -1 +| _, FVrel _ -> 1 | FVuniv_var i1, FVuniv_var i2 -> Int.compare i1 i2 -| FVuniv_var _, (FVnamed _ | FVrel _) -> 1 -| FVuniv_var _, FVevar _ -> -1 -| FVevar _, (FVnamed _ | FVrel _ | FVuniv_var _) -> 1 +| FVuniv_var _, _ -> -1 +| _, FVuniv_var _ -> 1 +| FVqual_var i1, FVqual_var i2 -> Int.compare i1 i2 +| FVqual_var _, _ -> -1 +| _, FVqual_var _ -> 1 | FVevar e1, FVevar e2 -> Evar.compare e1 e2 end @@ -128,7 +131,7 @@ type vm_env = { type comp_env = { arity : int; (* arity of the current function, 0 if none *) - nb_uni_stack : int ; (* number of universes on the stack, *) + nb_uni_stack : int * int; (* number of qualities and universes on the stack, *) (* universes are always at the bottom. *) nb_stack : int; (* number of variables on the stack *) in_stack : int Range.t; (* position in the stack *) @@ -153,7 +156,7 @@ module Config = struct let stack_safety_margin = 15 end -type argument = ArgLambda of lambda | ArgUniv of Univ.Level.t +type _ argument = ArgLambda : lambda -> glob_env argument | ArgUniv of Univ.Level.t | ArgQuality of Sorts.Quality.t let empty_fv = { size= 0; fv_rev = []; fv_fwd = FvMap.empty } let push_fv d e = { @@ -166,7 +169,7 @@ let fv r = !(r.in_env) let empty_comp_env ()= { arity = 0; - nb_uni_stack = 0; + nb_uni_stack = (0,0); nb_stack = 0; in_stack = Range.empty; pos_rec = [||]; @@ -192,7 +195,7 @@ let ensure_stack_capacity (cenv : comp_env) code = let rec add_param n sz l = if Int.equal n 0 then l else add_param (n - 1) sz (Range.cons (n+sz) l) -let comp_env_fun ?(univs=0) arity = +let comp_env_fun ?(univs=(0,0)) arity = { arity; nb_uni_stack = univs ; nb_stack = arity; @@ -206,7 +209,7 @@ let comp_env_fun ?(univs=0) arity = let comp_env_fix_type rfv = { arity = 0; - nb_uni_stack = 0; + nb_uni_stack = (0,0); nb_stack = 0; in_stack = Range.empty; pos_rec = [||]; @@ -217,7 +220,7 @@ let comp_env_fix_type rfv = let comp_env_fix ndef arity rfv = { arity; - nb_uni_stack = 0; + nb_uni_stack = (0,0); nb_stack = arity; in_stack = add_param arity 0 Range.empty; pos_rec = Array.init ndef (fun i -> Koffsetclosure i); @@ -228,7 +231,7 @@ let comp_env_fix ndef arity rfv = let comp_env_cofix_type ndef rfv = { arity = 0; - nb_uni_stack = 0; + nb_uni_stack = (0,0); nb_stack = 0; in_stack = Range.empty; pos_rec = [||]; @@ -239,7 +242,7 @@ let comp_env_cofix_type ndef rfv = let comp_env_cofix ndef arity rfv = { arity; - nb_uni_stack = 0; + nb_uni_stack = (0,0); nb_stack = arity; in_stack = add_param arity 0 Range.empty; pos_rec = Array.init ndef (fun i -> Kenvacc (ndef - 1 - i)); @@ -294,12 +297,12 @@ let pos_universe_var i r sz = (* Compilation of a universe variable can happen either at toplevel (the current closure correspond to a constant and has local universes) or in a local closure (which has no local universes). *) - if r.nb_uni_stack != 0 then + if r.nb_uni_stack <> (0,0) then (* Universe variables are represented by De Bruijn levels (not indices), - starting at 0. The shape of the stack will be [v1|..|vn|u1..up|arg1..argq] - with size = n + p + q, and q = r.arity. So Kacc (sz - r.arity - 1) will access + starting at 0. The shape of the stack will be [v1|..|vn|qual1..qualx|u1..up|arg1..argq] + with size = n + x + p + q, and q = r.arity. So Kacc (sz - r.arity - 1) will access the last universe. *) - Kacc (sz - r.arity - (r.nb_uni_stack - i)) + Kacc (sz - r.arity - fst r.nb_uni_stack - (snd r.nb_uni_stack - i)) else let env = !(r.in_env) in let db = FVuniv_var i in @@ -309,6 +312,24 @@ let pos_universe_var i r sz = r.in_env := push_fv db env; Kenvacc(r.offset + pos) +let pos_quality_var i r sz = + (* Compilation of a universe variable can happen either at toplevel (the + current closure correspond to a constant and has local universes) or in a + local closure (which has no local universes). *) + if r.nb_uni_stack <> (0,0) then + (* Universe variables are represented by De Bruijn levels (not indices), + starting at 0. The shape of the stack will be [v1|..|vn|qual1..qualx|u1..up|arg1..argq] + with size = n + x + p + q, and q = r.arity. XXX TODO check this I can't understand debruijn *) + Kacc (sz - r.arity - (fst r.nb_uni_stack - i)) + else + let env = !(r.in_env) in + let db = FVqual_var i in + try Kenvacc (r.offset + find_at db env) + with Not_found -> + let pos = env.size in + r.in_env := push_fv db env; + Kenvacc(r.offset + pos) + let pos_evar evk r = let env = !(r.in_env) in let cid = FVevar evk in @@ -479,6 +500,7 @@ let compile_fv_elem cenv fv sz cont = | FVrel i -> pos_rel i cenv sz :: cont | FVnamed id -> pos_named id cenv :: cont | FVuniv_var i -> pos_universe_var i cenv sz :: cont + | FVqual_var i -> pos_quality_var i cenv sz :: cont | FVevar evk -> pos_evar evk cenv :: cont let rec compile_fv cenv l sz cont = @@ -548,31 +570,43 @@ let rec compile_lam env cenv lam sz cont = | Lind (ind,u) -> if UVars.Instance.is_empty u then compile_structured_constant cenv (Const_ind ind) sz cont - else comp_app compile_structured_constant compile_universe cenv - (Const_ind ind) (UVars.Instance.to_array u) sz cont + else + let qs, us = UVars.Instance.to_array u in + let ulen = Array.length us in + let args = Array.init (ulen + Array.length qs) (fun i -> + if i < ulen then ArgUniv us.(i) + else ArgQuality qs.(i - ulen)) + in + comp_app compile_structured_constant (compile_arg ()) cenv + (Const_ind ind) args sz cont | Lsort s -> (* We represent universes as a global constant with local universes "compacted", i.e. as [u arg0 ... argn] where we will substitute (after evaluation) [Var 0,...,Var n] with values of [arg0,...,argn] *) let s, subs = match s with - | Sorts.Set | Sorts.Prop | Sorts.SProp as s -> s, [] + | Sorts.Set | Sorts.Prop | Sorts.SProp as s -> s, [||] | Sorts.Type u -> let u, s = Univ.compact_univ u in - Sorts.sort_of_univ u, s + Sorts.sort_of_univ u, Array.map_of_list (fun i -> FVuniv_var i) s | Sorts.QSort (q, u) -> let u, s = Univ.compact_univ u in + let s = Array.map_of_list (fun i -> FVuniv_var i) s in + let q, s = match Sorts.QVar.var_index q with + | Some i -> Sorts.QVar.make_var 0, Array.append [|FVqual_var i|] s + | None -> q, s + in Sorts.qsort q u, s in let compile_get_univ cenv idx sz cont = let () = set_max_stack_size cenv sz in - compile_fv_elem cenv (FVuniv_var idx) sz cont + compile_fv_elem cenv idx sz cont in - if List.is_empty subs then + if Array.is_empty subs then compile_structured_constant cenv (Const_sort s) sz cont else comp_app compile_structured_constant compile_get_univ cenv - (Const_sort s) (Array.of_list subs) sz cont + (Const_sort s) subs sz cont | Llet (_id,def,body) -> compile_lam env cenv def sz @@ -806,15 +840,33 @@ and compile_get_global cenv (kn,u) sz cont = if UVars.Instance.is_empty u then Kgetglobal kn :: cont else + let qs, us = UVars.Instance.to_array u in + let args = Array.append + (Array.map (fun u -> ArgUniv u) us) + (Array.map (fun q -> ArgQuality q) qs) + in comp_app (fun _ _ _ cont -> Kgetglobal kn :: cont) - compile_universe cenv () (UVars.Instance.to_array u) sz cont + (compile_arg ()) cenv () args sz cont and compile_universe cenv uni sz cont = let () = set_max_stack_size cenv sz in match Univ.Level.var_index uni with - | None -> compile_structured_constant cenv (Const_univ_level uni) sz cont + | None -> compile_structured_constant cenv (Const_quality_or_univ (Const_univ uni)) sz cont | Some idx -> pos_universe_var idx cenv sz :: cont +and compile_quality cenv q sz cont = + let () = set_max_stack_size cenv sz in + match Sorts.Quality.var_index q with + | None -> compile_structured_constant cenv (Const_quality_or_univ (Const_quality q)) sz cont + | Some idx -> pos_quality_var idx cenv sz :: cont + +and compile_arg : 'env. 'env -> _ -> 'env argument -> _ -> _ -> _ = + fun (type env) (env:env) cenv (constr_or_uni:env argument) sz cont -> + match constr_or_uni with + | ArgLambda t -> compile_lam env cenv t sz cont + | ArgUniv uni -> compile_universe cenv uni sz cont + | ArgQuality q -> compile_quality cenv q sz cont + and compile_constant env cenv kn u args sz cont = let () = set_max_stack_size cenv sz in if UVars.Instance.is_empty u then @@ -823,31 +875,26 @@ and compile_constant env cenv kn u args sz cont = compile_get_global cenv (kn,u) sz cont) (compile_lam env) cenv () args sz cont else - let compile_arg cenv constr_or_uni sz cont = - match constr_or_uni with - | ArgLambda t -> compile_lam env cenv t sz cont - | ArgUniv uni -> compile_universe cenv uni sz cont - in - let u = UVars.Instance.to_array u in - let lu = Array.length u in + let qs, us = UVars.Instance.to_array u in + let lq = Array.length qs in + let lu = Array.length us in let all = - Array.init (lu + Array.length args) - (fun i -> if i < lu then ArgUniv u.(i) else ArgLambda args.(i-lu)) + Array.init (lq + lu + Array.length args) + (fun i -> if i < lu then ArgUniv us.(i) else + if i - lu < lq then ArgQuality qs.(i - lu) + else ArgLambda args.(i-(lu+lq))) in comp_app (fun _ _ _ cont -> Kgetglobal kn :: cont) - compile_arg cenv () all sz cont - -let is_univ_copy max u = - let u = UVars.Instance.to_array u in - if Array.length u = max then - Array.fold_left_i (fun i acc u -> - if acc then - match Univ.Level.var_index u with - | None -> false - | Some l -> l = i - else false) true u - else - false + (compile_arg env) cenv () all sz cont + +let is_univ_copy (maxq,maxu) u = + let qs,us = UVars.Instance.to_array u in + let check_array max var_index a = + Array.length a = max + && Array.for_all_i (fun i x -> Option.equal Int.equal (var_index x) (Some i)) 0 a + in + check_array maxq Sorts.Quality.var_index qs + && check_array maxu Univ.Level.var_index us let dump_bytecode = ref false @@ -860,12 +907,12 @@ let dump_bytecodes init code fvs = prlist_with_sep (fun () -> str "; ") pp_fv_elem fvs ++ fnl ()) -let compile ~fail_on_error ?universes:(universes=0) env sigma c = +let compile ~fail_on_error ?universes:(universes=(0,0)) env sigma c = Label.reset_label_counter (); let cont = [Kstop] in try let cenv, init_code, fun_code = - if Int.equal universes 0 then + if UVars.eq_sizes universes (0,0) then let lam = lambda_of_constr ~optimize:true env sigma c in let cenv = empty_comp_env () in let env = { env; fun_code = [] } in @@ -880,7 +927,7 @@ let compile ~fail_on_error ?universes:(universes=0) env sigma c = let params, body = decompose_Llam lam in let arity = Array.length params in let cenv = empty_comp_env () in - let full_arity = arity + universes in + let full_arity = arity + fst universes + snd universes in let r_fun = comp_env_fun ~univs:universes arity in let lbl_fun = Label.create () in let env = { env; fun_code = [] } in diff --git a/kernel/vmbytegen.mli b/kernel/vmbytegen.mli index 01361a0497144..f34c593c990ad 100644 --- a/kernel/vmbytegen.mli +++ b/kernel/vmbytegen.mli @@ -16,7 +16,7 @@ open Environ (** Should only be used for monomorphic terms *) val compile : - fail_on_error:bool -> ?universes:int -> + fail_on_error:bool -> ?universes:int*int -> env -> Genlambda.evars -> constr -> (to_patch * fv) option (** init, fun, fv *) diff --git a/kernel/vmemitcodes.ml b/kernel/vmemitcodes.ml index 296e1f697b845..d89b1c461aef3 100644 --- a/kernel/vmemitcodes.ml +++ b/kernel/vmemitcodes.ml @@ -475,7 +475,7 @@ type to_patch = emitcodes * patches (* Substitution *) let subst_strcst s sc = match sc with - | Const_sort _ | Const_b0 _ | Const_univ_level _ | Const_val _ | Const_uint _ + | Const_sort _ | Const_b0 _ | Const_quality_or_univ _ | Const_val _ | Const_uint _ | Const_float _ -> sc | Const_ind ind -> let kn,i = ind in Const_ind (subst_mind s kn, i) diff --git a/kernel/vmsymtable.ml b/kernel/vmsymtable.ml index eb9aede569e3b..6983f940222c0 100644 --- a/kernel/vmsymtable.ml +++ b/kernel/vmsymtable.ml @@ -299,7 +299,7 @@ and slot_for_fv env sigma fv envcache table = | Some v -> v end | FVevar evk -> val_of_evar evk - | FVuniv_var _idu -> + | FVuniv_var _ | FVqual_var _ -> assert false and eval_to_patch env sigma (code, fv) envcache table = diff --git a/kernel/vmvalues.ml b/kernel/vmvalues.ml index 03805947fc3f9..9fe869e7d4974 100644 --- a/kernel/vmvalues.ml +++ b/kernel/vmvalues.ml @@ -8,7 +8,6 @@ (* * (see LICENSE file for the text of the license) *) (************************************************************************) open Names -open UVars open Values (********************************************) @@ -43,6 +42,10 @@ let switch_tag = 5 let cofix_tag = 6 let cofix_evaluated_tag = 7 +type quality_or_univ = + | Const_quality of Sorts.Quality.t + | Const_univ of Univ.Level.t + (** Structured constants are constants whose construction is done once. Their occurrences share the same value modulo kernel name substitutions (for functor application). Structured values have the additional property that no @@ -52,7 +55,7 @@ type structured_constant = | Const_sort of Sorts.t | Const_ind of inductive | Const_b0 of tag - | Const_univ_level of Univ.Level.t + | Const_quality_or_univ of quality_or_univ | Const_val of structured_values | Const_uint of Uint63.t | Const_float of Float64.t @@ -94,6 +97,12 @@ let hash_structured_values (v : structured_values) = (* We may want a better hash function here *) Hashtbl.hash v +let eq_quality_or_univ a b = + match a, b with + | Const_quality a, Const_quality b -> Sorts.Quality.equal a b + | Const_univ a, Const_univ b -> Univ.Level.equal a b + | (Const_quality _ | Const_univ _), _ -> false + let eq_structured_constant c1 c2 = match c1, c2 with | Const_sort s1, Const_sort s2 -> Sorts.equal s1 s2 | Const_sort _, _ -> false @@ -101,8 +110,8 @@ let eq_structured_constant c1 c2 = match c1, c2 with | Const_ind _, _ -> false | Const_b0 t1, Const_b0 t2 -> Int.equal t1 t2 | Const_b0 _, _ -> false -| Const_univ_level l1 , Const_univ_level l2 -> Univ.Level.equal l1 l2 -| Const_univ_level _ , _ -> false +| Const_quality_or_univ l1 , Const_quality_or_univ l2 -> eq_quality_or_univ l1 l2 +| Const_quality_or_univ _ , _ -> false | Const_val v1, Const_val v2 -> eq_structured_values v1 v2 | Const_val _, _ -> false | Const_uint i1, Const_uint i2 -> Uint63.equal i1 i2 @@ -116,10 +125,11 @@ let hash_structured_constant c = | Const_sort s -> combinesmall 1 (Sorts.hash s) | Const_ind i -> combinesmall 2 (Ind.CanOrd.hash i) | Const_b0 t -> combinesmall 3 (Int.hash t) - | Const_univ_level l -> combinesmall 4 (Univ.Level.hash l) - | Const_val v -> combinesmall 5 (hash_structured_values v) - | Const_uint i -> combinesmall 6 (Uint63.hash i) - | Const_float f -> combinesmall 7 (Float64.hash f) + | Const_quality_or_univ (Const_univ l) -> combinesmall 4 (Univ.Level.hash l) + | Const_quality_or_univ (Const_quality l) -> combinesmall 5 (Sorts.Quality.hash l) + | Const_val v -> combinesmall 6 (hash_structured_values v) + | Const_uint i -> combinesmall 7 (Uint63.hash i) + | Const_float f -> combinesmall 8 (Float64.hash f) let eq_annot_switch asw1 asw2 = let eq_rlc (i1, j1) (i2, j2) = Int.equal i1 i2 && Int.equal j1 j2 in @@ -147,7 +157,8 @@ let pp_struct_const = function | Const_sort s -> pp_sort s | Const_ind (mind, i) -> Pp.(MutInd.print mind ++ str"#" ++ int i) | Const_b0 i -> Pp.int i - | Const_univ_level l -> Univ.Level.raw_pr l + | Const_quality_or_univ (Const_univ l) -> Univ.Level.raw_pr l + | Const_quality_or_univ (Const_quality l) -> Sorts.Quality.raw_pr l | Const_val _ -> Pp.str "(value)" | Const_uint i -> Pp.str (Uint63.to_string i) | Const_float f -> Pp.str (Float64.to_string f) @@ -290,7 +301,21 @@ let arg args i = (* Destructors ***********************************) (*************************************************) -let uni_lvl_val (v : values) : Univ.Level.t = Obj.magic v +let uni_lvl_val (v : values) : quality_or_univ = Obj.magic v + +let instance_of_args n args = + let args = Array.init n (fun i -> uni_lvl_val (arg args i)) in + let first_univ = + (* TODO check are univs first or qualities *) + CArray.findi (fun _ arg -> match arg with Const_univ _ -> true | _ -> false) args + in + let qs, us = match first_univ with + | None -> args, [||] + | Some first_univ -> CArray.chop first_univ args + in + let qs = Array.map (function Const_quality q -> q | _ -> assert false) qs in + let us = Array.map (function Const_univ u -> u | _ -> assert false) us in + UVars.Instance.of_array (qs,us) let rec whd_accu a stk = let stk = @@ -302,15 +327,10 @@ let rec whd_accu a stk = begin match stk with | [] -> Vaccu (Obj.magic at, stk) | [Zapp args] -> - let args = Array.init (nargs args) (arg args) in let s = Obj.obj (Obj.field at 0) in - begin match s with - | Sorts.Type u -> - let inst = Instance.of_array (Array.map uni_lvl_val args) in - let u = UVars.subst_instance_universe inst u in - Vaccu (Asort (Sorts.sort_of_univ u), []) - | _ -> assert false - end + let inst = instance_of_args (nargs args) args in + let s = UVars.subst_instance_sort inst s in + Vaccu (Asort s, []) | _ -> assert false end | i when i <= max_atom_tag -> @@ -401,7 +421,7 @@ let obj_of_str_const str = | Const_sort s -> obj_of_atom (Asort s) | Const_ind ind -> obj_of_atom (Aind ind) | Const_b0 tag -> Obj.repr tag - | Const_univ_level l -> Obj.repr l + | Const_quality_or_univ l -> Obj.repr l | Const_val v -> Obj.repr v | Const_uint i -> Obj.repr i | Const_float f -> Obj.repr f diff --git a/kernel/vmvalues.mli b/kernel/vmvalues.mli index b713d5f4a3536..c97f64349efb7 100644 --- a/kernel/vmvalues.mli +++ b/kernel/vmvalues.mli @@ -34,11 +34,15 @@ val switch_tag : tag val cofix_tag : tag val cofix_evaluated_tag : tag +type quality_or_univ = + | Const_quality of Sorts.Quality.t + | Const_univ of Univ.Level.t + type structured_constant = | Const_sort of Sorts.t | Const_ind of inductive | Const_b0 of tag - | Const_univ_level of Univ.Level.t + | Const_quality_or_univ of quality_or_univ | Const_val of structured_values | Const_uint of Uint63.t | Const_float of Float64.t @@ -143,7 +147,7 @@ external val_of_annot_switch : annot_switch -> values = "%identity" (** Destructors *) val whd_val : values -> kind -val uni_lvl_val : values -> Univ.Level.t +val instance_of_args : int -> arguments -> UVars.Instance.t (** Arguments *) diff --git a/library/nametab.ml b/library/nametab.ml index a8f13d5bd2576..598b6dccfd1e4 100644 --- a/library/nametab.ml +++ b/library/nametab.ml @@ -382,15 +382,14 @@ end module MPDTab = Make(DirPath')(MPEqual) module DirTab = Make(DirPath')(GlobDirRef) -module UnivIdEqual = -struct - type t = Univ.UGlobal.t - let equal = Univ.UGlobal.equal -end -module UnivTab = Make(FullPath)(UnivIdEqual) +module UnivTab = Make(FullPath)(Univ.UGlobal) type univtab = UnivTab.t let the_univtab = Summary.ref ~name:"univtab" (UnivTab.empty : univtab) +module QVarTab = Make(FullPath)(Sorts.QVar) +type qvartab = QVarTab.t +let the_qvartab = Summary.ref ~name:"qvartab" (QVarTab.empty : qvartab) + (* Reversed name tables ***************************************************) (* This table translates extended_global_references back to section paths *) @@ -404,18 +403,16 @@ type mprevtab = DirPath.t MPmap.t type mptrevtab = full_path MPmap.t -module UnivIdOrdered = -struct - type t = Univ.UGlobal.t - let hash = Univ.UGlobal.hash - let compare = Univ.UGlobal.compare -end - -module UnivIdMap = HMap.Make(UnivIdOrdered) +module UnivIdMap = HMap.Make(Univ.UGlobal) type univrevtab = full_path UnivIdMap.t let the_univrevtab = Summary.ref ~name:"univrevtab" (UnivIdMap.empty : univrevtab) +module QVarIdMap = HMap.Make(Sorts.QVar) + +type qvarrevtab = full_path QVarIdMap.t +let the_qvarrevtab = Summary.ref ~name:"qvarrevtab" (QVarIdMap.empty : qvarrevtab) + (** Module-related nametab *) module Modules = struct @@ -506,6 +503,11 @@ let push_universe vis sp univ = the_univtab := UnivTab.push vis sp univ !the_univtab; the_univrevtab := UnivIdMap.add univ sp !the_univrevtab +let push_qvar vis sp q = + assert (Option.is_empty (Sorts.QVar.var_index q)); + the_qvartab := QVarTab.push vis sp q !the_qvartab; + the_qvarrevtab := QVarIdMap.add q sp !the_qvarrevtab + (* Reverse locate functions ***********************************************) let path_of_global ref = @@ -532,6 +534,9 @@ let path_of_modtype mp = let path_of_universe mp = UnivIdMap.find mp !the_univrevtab +let path_of_qvar mp = + QVarIdMap.find mp !the_qvarrevtab + (* Shortest qualid functions **********************************************) let shortest_qualid_of_global ?loc ctx ref = @@ -556,7 +561,11 @@ let shortest_qualid_of_modtype ?loc kn = let shortest_qualid_of_universe ?loc ctx kn = let sp = UnivIdMap.find kn !the_univrevtab in - UnivTab.shortest_qualid_gen ?loc (fun id -> Id.Map.mem id ctx) sp !the_univtab + UnivTab.shortest_qualid_gen ?loc (fun id -> Id.Map.mem id ctx) sp !the_univtab + +let shortest_qualid_of_qvar ?loc ctx kn = + let sp = QVarIdMap.find kn !the_qvarrevtab in + QVarTab.shortest_qualid_gen ?loc (fun id -> Id.Map.mem id ctx) sp !the_qvartab let pr_global_env env ref = try pr_qualid (shortest_qualid_of_global env ref) @@ -621,6 +630,8 @@ let full_name_modtype qid = MPTab.user_name qid Modules.(!nametab.modtypetab) let locate_universe qid = UnivTab.locate qid !the_univtab +let locate_qvar qid = QVarTab.locate qid !the_qvartab + let locate_dir qid = DirTab.locate qid Modules.(!nametab.dirtab) let locate_module qid = MPDTab.locate qid Modules.(!nametab.modtab) @@ -697,3 +708,5 @@ let exists_module dir = MPDTab.exists dir Modules.(!nametab.modtab) let exists_modtype sp = MPTab.exists sp Modules.(!nametab.modtypetab) let exists_universe kn = UnivTab.exists kn !the_univtab + +let exists_qvar kn = QVarTab.exists kn !the_qvartab diff --git a/library/nametab.mli b/library/nametab.mli index 18d9cd1ffee1c..62279d947253f 100644 --- a/library/nametab.mli +++ b/library/nametab.mli @@ -110,10 +110,10 @@ val push_module : visibility -> DirPath.t -> ModPath.t -> unit val push_dir : visibility -> DirPath.t -> GlobDirRef.t -> unit val push_abbreviation : ?deprecated:Deprecation.t -> visibility -> full_path -> Globnames.abbreviation -> unit -module UnivIdMap : CMap.ExtS with type key = Univ.UGlobal.t - val push_universe : visibility -> full_path -> Univ.UGlobal.t -> unit +val push_qvar : visibility -> full_path -> Sorts.QVar.t -> unit + (** Deprecation info *) val is_deprecated_xref : Globnames.extended_global_reference -> Deprecation.t option @@ -134,6 +134,7 @@ val locate_dir : qualid -> GlobDirRef.t val locate_module : qualid -> ModPath.t val locate_section : qualid -> DirPath.t val locate_universe : qualid -> Univ.UGlobal.t +val locate_qvar : qualid -> Sorts.QVar.t val locate_extended_nowarn : qualid -> Globnames.extended_global_reference * Deprecation.t option @@ -175,6 +176,7 @@ val exists_modtype : full_path -> bool val exists_module : DirPath.t -> bool val exists_dir : DirPath.t -> bool val exists_universe : full_path -> bool +val exists_qvar : full_path -> bool (** {6 These functions locate qualids into full user names } *) @@ -197,6 +199,8 @@ val path_of_modtype : ModPath.t -> full_path @raise Not_found if the universe was not introduced by the user. *) val path_of_universe : Univ.UGlobal.t -> full_path +val path_of_qvar : Sorts.QVar.t -> full_path + (** Returns in particular the dirpath or the basename of the full path associated to global reference *) @@ -221,6 +225,8 @@ val shortest_qualid_of_module : ?loc:Loc.t -> ModPath.t -> qualid (** In general we have a [UnivNames.universe_binders] around rather than a [Id.Set.t] *) val shortest_qualid_of_universe : ?loc:Loc.t -> 'u Id.Map.t -> Univ.UGlobal.t -> qualid +val shortest_qualid_of_qvar : ?loc:Loc.t -> 'u Id.Map.t -> Sorts.QVar.t -> qualid + (** {5 Generic name handling} *) (** NOT FOR PUBLIC USE YET. Plugin writers, please do not rely on this API. *) diff --git a/parsing/g_constr.mlg b/parsing/g_constr.mlg index 278172ff5440d..e144f74940964 100644 --- a/parsing/g_constr.mlg +++ b/parsing/g_constr.mlg @@ -81,11 +81,36 @@ let test_array_closing = lk_kw "|" >> lk_kw "]" >> check_no_space end +let test_sort_qvar = + let open Pcoq.Lookahead in + to_entry "test_sort_qvar" begin + lk_ident >> lk_list lk_field >> lk_kw "|" + end + +type univ_level_or_quality = + | SProp | Prop | Set | Type | Anon of Loc.t | Global of Libnames.qualid + +let force_univ_level ?loc = function + | SProp -> UNamed CSProp + | Prop -> UNamed CProp + | Set -> UNamed CSet + | Type -> UAnonymous {rigid=true} + | Anon _ -> UAnonymous {rigid=false} + | Global qid -> UNamed (CType qid) + +let force_quality ?loc = function + | SProp -> CQConstant QSProp + | Prop -> CQConstant QProp + | Set -> CErrors.user_err ?loc Pp.(str "Universe levels cannot be Set.") + | Type -> CQConstant QType + | Anon loc -> CQAnon (Some loc) + | Global qid -> CQVar qid + } GRAMMAR EXTEND Gram GLOBAL: binder_constr lconstr constr term - universe_level universe_name sort sort_family + universe_name sort sort_family global constr_pattern cpattern Constr.ident closed_binder open_binders binder binders binders_fixannot record_declaration typeclass_constraint pattern arg type_cstr @@ -111,6 +136,7 @@ GRAMMAR EXTEND Gram | "SProp" -> { UNamed (None, [CSProp, 0]) } | "Type" -> { UAnonymous {rigid=true} } | "Type"; "@{"; "_"; "}" -> { UAnonymous {rigid=false} } + | "Type"; "@{"; test_sort_qvar; q = reference; "|"; u = universe; "}" -> { UNamed (Some (CQVar q), u) } | "Type"; "@{"; u = universe; "}" -> { UNamed (None, u) } ] ] ; sort_family: @@ -282,16 +308,23 @@ GRAMMAR EXTEND Gram | -> { [] } ] ] ; univ_annot: - [ [ "@{"; l = LIST0 universe_level; "}" -> { Some l } + [ [ "@{"; l = LIST0 univ_level_or_quality; l' = OPT [ "|"; l' = LIST0 univ_level_or_quality -> { l' } ]; "}" -> + { + let l, l' = match l' with + | None -> [], List.map (force_univ_level ~loc) l + | Some l' -> List.map (force_quality ~loc) l, List.map (force_univ_level ~loc) l' + in + Some (l,l') + } | -> { None } ] ] ; - universe_level: - [ [ "Set" -> { UNamed CSet } - (* no parsing SProp as a level *) - | "Prop" -> { UNamed CProp } - | "Type" -> { UAnonymous {rigid=true} } - | "_" -> { UAnonymous {rigid=false} } - | id = global -> { UNamed (CType id) } ] ] + univ_level_or_quality: + [ [ "Set" -> { Set } + | "SProp" -> { SProp } + | "Prop" -> { Prop } + | "Type" -> { Type } + | "_" -> { Anon loc } + | id = global -> { Global id } ] ] ; fix_decls: [ [ dcl = fix_decl -> { let (id,_,_,_,_) = dcl.CAst.v in (id,[dcl.CAst.v]) } diff --git a/parsing/pcoq.ml b/parsing/pcoq.ml index ecc43659254ea..d364abaedebde 100644 --- a/parsing/pcoq.ml +++ b/parsing/pcoq.ml @@ -90,6 +90,10 @@ struct | Tok.IDENT _ -> Some (n + 1) | _ -> None + let lk_field n kwstate strm = match LStream.peek_nth kwstate n strm with + | Tok.FIELD _ -> Some (n + 1) + | _ -> None + let lk_name = lk_ident <+> lk_kw "_" let lk_ident_except idents n kwstate strm = match LStream.peek_nth kwstate n strm with @@ -337,7 +341,6 @@ module Constr = let ident = Entry.make "ident" let global = Entry.make "global" let universe_name = Entry.make "universe_name" - let universe_level = Entry.make "universe_level" let sort = Entry.make "sort" let sort_family = Entry.make "sort_family" let pattern = Entry.make "pattern" diff --git a/parsing/pcoq.mli b/parsing/pcoq.mli index 4bcec9670e6b9..c8df72b07b3f4 100644 --- a/parsing/pcoq.mli +++ b/parsing/pcoq.mli @@ -29,12 +29,14 @@ module Lookahead : sig val to_entry : string -> t -> unit Entry.t val (>>) : t -> t -> t val (<+>) : t -> t -> t + val lk_empty : t val lk_list : t -> t val check_no_space : t val lk_kw : string -> t val lk_kws : string list -> t val lk_nat : t val lk_ident : t + val lk_field : t val lk_name : t val lk_ident_except : string list -> t val lk_ident_list : t @@ -182,7 +184,6 @@ module Constr : val ident : Id.t Entry.t val global : qualid Entry.t val universe_name : sort_name_expr Entry.t - val universe_level : univ_level_expr Entry.t val sort : sort_expr Entry.t val sort_family : Sorts.family Entry.t val pattern : cases_pattern_expr Entry.t diff --git a/plugins/ltac/extratactics.mlg b/plugins/ltac/extratactics.mlg index 9d86a0d1ea8e7..8b5b28ae2a334 100644 --- a/plugins/ltac/extratactics.mlg +++ b/plugins/ltac/extratactics.mlg @@ -228,7 +228,7 @@ let add_rewrite_hint ~locality ~poly bases ort t lcsr = if poly then ctx else (* This is a global universe context that shouldn't be refreshed at every use of the hint, declare it globally. *) - (DeclareUctx.declare_universe_context ~poly:false ctx; + (Global.push_context_set ~strict:true ctx; Univ.ContextSet.empty) in CAst.make ?loc:(Constrexpr_ops.constr_loc ce) ((c, ctx), ort, Option.map (in_gen (rawwit wit_ltac)) t) in diff --git a/plugins/ltac2/tac2core.ml b/plugins/ltac2/tac2core.ml index 4248190f3ceb0..950f5470fe43d 100644 --- a/plugins/ltac2/tac2core.ml +++ b/plugins/ltac2/tac2core.ml @@ -122,10 +122,14 @@ let to_binder b = let of_instance u = let u = UVars.Instance.to_array (EConstr.Unsafe.to_instance u) in - Value.of_array (fun v -> Value.of_ext Value.val_univ v) u + let toqs = Value.of_array (fun v -> Value.of_ext Value.val_quality v) in + let tous = Value.of_array (fun v -> Value.of_ext Value.val_univ v) in + Value.of_pair toqs tous u let to_instance u = - let u = Value.to_array (fun v -> Value.to_ext Value.val_univ v) u in + let toqs = Value.to_array (fun v -> Value.to_ext Value.val_quality v) in + let tous = Value.to_array (fun v -> Value.to_ext Value.val_univ v) in + let u = Value.to_pair toqs tous u in EConstr.EInstance.make (UVars.Instance.of_array u) let of_rec_declaration (nas, ts, cs) = diff --git a/plugins/ltac2/tac2ffi.ml b/plugins/ltac2/tac2ffi.ml index 7619925b6bc54..8be42b9ff75c9 100644 --- a/plugins/ltac2/tac2ffi.ml +++ b/plugins/ltac2/tac2ffi.ml @@ -109,6 +109,7 @@ let val_projection = Val.create "projection" let val_case = Val.create "case" let val_binder = Val.create "binder" let val_univ = Val.create "universe" +let val_quality = Val.create "quality" let val_free : Names.Id.Set.t Val.tag = Val.create "free" let val_ltac1 : Geninterp.Val.t Val.tag = Val.create "ltac1" let val_uint63 = Val.create "uint63" diff --git a/plugins/ltac2/tac2ffi.mli b/plugins/ltac2/tac2ffi.mli index 7b19b644bb680..b2371e858c412 100644 --- a/plugins/ltac2/tac2ffi.mli +++ b/plugins/ltac2/tac2ffi.mli @@ -213,6 +213,7 @@ val val_projection : Projection.t Val.tag val val_case : Constr.case_info Val.tag val val_binder : (Name.t Context.binder_annot * types) Val.tag val val_univ : Univ.Level.t Val.tag +val val_quality : Sorts.Quality.t Val.tag val val_free : Id.Set.t Val.tag val val_uint63 : Uint63.t Val.tag val val_float : Float64.t Val.tag diff --git a/plugins/ring/ring.ml b/plugins/ring/ring.ml index fac2e0ab6f3bc..5422495ceb5d2 100644 --- a/plugins/ring/ring.ml +++ b/plugins/ring/ring.ml @@ -144,9 +144,9 @@ let ic_unsafe env sigma c = (*FIXME remove *) let decl_constant name univs c = let open Constr in - let vars = CVars.universes_of_constr c in + let _, vars = CVars.universes_of_constr c in let univs = UState.restrict_universe_context ~lbound:(Global.universes_lbound ()) univs vars in - let () = DeclareUctx.declare_universe_context ~poly:false univs in + let () = Global.push_context_set ~strict:true univs in let types = (Typeops.infer (Global.env ()) c).uj_type in let univs = UState.Monomorphic_entry Univ.ContextSet.empty, UnivNames.empty_binders in (* UnsafeMonomorphic: we always do poly:false *) diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index f36706cce281a..30b7b4176e111 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -47,7 +47,6 @@ end = struct open CVars open Declarations -open Univ open UVars open Constr @@ -79,7 +78,7 @@ let return_clause env sigma ind u params (nas, p) = let realdecls, _ = List.chop mip.mind_nrealdecls mip.mind_arity_ctxt in let self = let args = Context.Rel.instance mkRel 0 mip.mind_arity_ctxt in - let inst = Instance.of_array (Array.init (Instance.length u) Level.var) in + let inst = Instance.(abstract_instance (length u)) in mkApp (mkIndU (ind, inst), args) in let realdecls = LocalAssum (Context.anonR, self) :: realdecls in @@ -202,7 +201,7 @@ let print_universes = ref false let { Goptions.get = print_sort_quality } = Goptions.declare_bool_option_and_ref ~key:["Printing";"Sort";"Qualities"] - ~value:false + ~value:true () (** If true, prints local context of evars, whatever print_arguments *) @@ -760,6 +759,20 @@ let detype_level_name sigma l = | Some id -> GLocalUniv (CAst.make id) | None -> GUniv l +let detype_level sigma l = + UNamed (detype_level_name sigma l) + +let detype_qvar sigma q = + match UState.id_of_qvar (Evd.evar_universe_context sigma) q with + | Some id -> GLocalQVar (CAst.make (Name id)) + | None -> GQVar q + +let detype_quality sigma q = + let open Sorts.Quality in + match q with + | QConstant q -> GQConstant q + | QVar q -> detype_qvar sigma q + let detype_universe sigma u = List.map (on_fst (detype_level_name sigma)) (Univ.Universe.repr u) @@ -773,7 +786,7 @@ let detype_sort sigma = function else UAnonymous {rigid=true}) | QSort (q, u) -> if !print_universes then - let q = if print_sort_quality () then Some q else None in + let q = if print_sort_quality () then Some (detype_qvar sigma q) else None in UNamed (q, detype_universe sigma u) else UAnonymous {rigid=true} @@ -782,15 +795,16 @@ type binder_kind = BProd | BLambda | BLetIn (**********************************************************************) (* Main detyping function *) -let detype_level sigma l = - UNamed (detype_level_name sigma l) - let detype_instance sigma l = if not !print_universes then None else let l = EInstance.kind sigma l in if UVars.Instance.is_empty l then None - else Some (List.map (detype_level sigma) (Array.to_list (UVars.Instance.to_array l))) + else + let qs, us = UVars.Instance.to_array l in + let qs = List.map (detype_quality sigma) (Array.to_list qs) in + let us = List.map (detype_level sigma) (Array.to_list us) in + Some (qs, us) let delay (type a) (d : a delay) (f : a delay -> _ -> _ -> _ -> _ -> _ -> a glob_constr_r) flags env avoid sigma t : a glob_constr_g = match d with diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index 4f8ff1d958215..679f757953077 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -367,14 +367,20 @@ let compare_cumulative_instances pbty evd variances u u' = Success evd | Inr p -> UnifFailure (evd, UnifUnivInconsistency p) +let compare_constructor_instances evd u u' = + match Evarutil.compare_constructor_instances evd u u' with + | Inl evd -> + Success evd + | Inr p -> UnifFailure (evd, UnifUnivInconsistency p) + type application = FullyApplied | NumArgs of int let is_applied o n = match o with FullyApplied -> true | NumArgs m -> Int.equal m n let compare_heads pbty env evd ~nargs term term' = let check_strict evd u u' = - let cstrs = UVars.enforce_eq_instances u u' Univ.Constraints.empty in - try Success (Evd.add_constraints evd cstrs) + let cstrs = UVars.enforce_eq_instances u u' Sorts.QUConstraints.empty in + try Success (Evd.add_quconstraints evd cstrs) with UGraph.UniverseInconsistency p -> UnifFailure (evd, UnifUnivInconsistency p) in match EConstr.kind evd term, EConstr.kind evd term' with @@ -416,8 +422,7 @@ let compare_heads pbty env evd ~nargs term term' = let needed = Conversion.constructor_cumulativity_arguments (mind,ind,ctor) in if not (is_applied nargs needed) then check_strict evd u u' - else - Success (compare_constructor_instances evd u u') + else compare_constructor_instances evd u u' end | Construct _, Construct _ -> UnifFailure (evd, NotSameHead) | _, _ -> anomaly (Pp.str "") diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml index 30006d5884fa4..99099260c6ad3 100644 --- a/pretyping/evardefine.ml +++ b/pretyping/evardefine.ml @@ -196,7 +196,7 @@ let define_pure_evar_as_array env sigma evk = let s = destSort sigma concl in (* array@{u} ty : Type@{u} <= Type@{s} *) let sigma = Evd.set_leq_sort env sigma s' s in - let ar = Typeops.type_of_array env (UVars.Instance.of_array [|u|]) in + let ar = Typeops.type_of_array env (UVars.Instance.of_array ([||],[|u|])) in let sigma = Evd.define evk (mkApp (EConstr.of_constr ar, [| ty |])) sigma in sigma diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml index 9392a8c0feeae..faf9bf97c875e 100644 --- a/pretyping/glob_ops.ml +++ b/pretyping/glob_ops.ml @@ -44,6 +44,13 @@ let map_glob_decl_left_to_right f (na,k,obd,ty) = let comp2 = f ty in (na,k,comp1,comp2) +let glob_quality_eq (type a) (g1:a glob_quality_g) (g2:a glob_quality_g) = match g1, g2 with + | GQConstant q1, GQConstant q2 -> Sorts.Quality.Constants.equal q1 q2 + | GLocalQVar na1, GLocalQVar na2 -> CAst.eq Name.equal na1 na2 + | GQVar q1, GQVar q2 -> Sorts.QVar.equal q1 q2 + | GRawQVar q1, GRawQVar q2 -> Sorts.QVar.equal q1 q2 + | (GQConstant _ | GLocalQVar _ | GQVar _ | GRawQVar _), _ -> false + let glob_sort_name_eq g1 g2 = match g1, g2 with | GSProp, GSProp | GProp, GProp @@ -74,7 +81,7 @@ let glob_sort_gen_eq f u1 u2 = let glob_sort_eq u1 u2 = let eq (q1, l1) (q2, l2) = - Option.equal Sorts.QVar.equal q1 q2 && + Option.equal glob_quality_eq q1 q2 && List.equal (fun (x,m) (y,n) -> glob_sort_name_eq x y && Int.equal m n) l1 l2 in glob_sort_gen_eq eq u1 u2 @@ -82,6 +89,10 @@ let glob_sort_eq u1 u2 = let glob_level_eq u1 u2 = glob_sort_gen_eq glob_sort_name_eq u1 u2 +let instance_eq (q1,u1) (q2,u2) = + List.equal glob_quality_eq q1 q2 + && List.equal glob_level_eq u1 u2 + let binding_kind_eq bk1 bk2 = match bk1, bk2 with | Explicit, Explicit -> true | NonMaxImplicit, NonMaxImplicit -> true @@ -134,16 +145,16 @@ let fix_kind_eq k1 k2 = match k1, k2 with | GCoFix i1, GCoFix i2 -> Int.equal i1 i2 | (GFix _ | GCoFix _), _ -> false -let instance_eq f (x1,c1) (x2,c2) = +let evar_instance_eq f (x1,c1) (x2,c2) = Id.equal x1.CAst.v x2.CAst.v && f c1 c2 let mk_glob_constr_eq f c1 c2 = match DAst.get c1, DAst.get c2 with | GRef (gr1, u1), GRef (gr2, u2) -> GlobRef.CanOrd.equal gr1 gr2 && - Option.equal (List.equal glob_level_eq) u1 u2 + Option.equal instance_eq u1 u2 | GVar id1, GVar id2 -> Id.equal id1 id2 | GEvar (id1, arg1), GEvar (id2, arg2) -> - Id.equal id1.CAst.v id2.CAst.v && List.equal (instance_eq f) arg1 arg2 + Id.equal id1.CAst.v id2.CAst.v && List.equal (evar_instance_eq f) arg1 arg2 | GPatVar k1, GPatVar k2 -> matching_var_kind_eq k1 k2 | GApp (f1, arg1), GApp (f2, arg2) -> f f1 f2 && List.equal f arg1 arg2 @@ -176,13 +187,13 @@ let mk_glob_constr_eq f c1 c2 = match DAst.get c1, DAst.get c2 with f c1 c2 && Option.equal cast_kind_eq k1 k2 && f t1 t2 | GProj ((cst1, u1), args1, c1), GProj ((cst2, u2), args2, c2) -> GlobRef.(CanOrd.equal (ConstRef cst1) (ConstRef cst2)) && - Option.equal (List.equal glob_level_eq) u1 u2 && + Option.equal instance_eq u1 u2 && List.equal f args1 args2 && f c1 c2 | GInt i1, GInt i2 -> Uint63.equal i1 i2 | GFloat f1, GFloat f2 -> Float64.equal f1 f2 | GArray (u1, t1, def1, ty1), GArray (u2, t2, def2, ty2) -> Array.equal f t1 t2 && f def1 def2 && f ty1 ty2 && - Option.equal (List.equal glob_level_eq) u1 u2 + Option.equal instance_eq u1 u2 | (GRef _ | GVar _ | GEvar _ | GPatVar _ | GApp _ | GLambda _ | GProd _ | GLetIn _ | GCases _ | GLetTuple _ | GIf _ | GRec _ | GSort _ | GHole _ | GGenarg _ | GCast _ | GProj _ | GInt _ | GFloat _ | GArray _), _ -> false diff --git a/pretyping/glob_ops.mli b/pretyping/glob_ops.mli index 7132fb119c768..96d6ad244a18c 100644 --- a/pretyping/glob_ops.mli +++ b/pretyping/glob_ops.mli @@ -19,6 +19,8 @@ val glob_sort_gen_eq : ('a -> 'a -> bool) -> 'a glob_sort_gen -> 'a glob_sort_ge val glob_sort_eq : Glob_term.glob_sort -> Glob_term.glob_sort -> bool +val glob_quality_eq : 'a glob_quality_g -> 'a glob_quality_g -> bool + val glob_level_eq : Glob_term.glob_level -> Glob_term.glob_level -> bool val cases_pattern_eq : 'a cases_pattern_g -> 'a cases_pattern_g -> bool diff --git a/pretyping/glob_term.mli b/pretyping/glob_term.mli index 54c35d2d0bde6..2b2dbce09a145 100644 --- a/pretyping/glob_term.mli +++ b/pretyping/glob_term.mli @@ -22,6 +22,16 @@ type existential_name = Id.t (** Sorts *) +type _ glob_quality_g = + | GQConstant : Sorts.Quality.constant -> [`with_constant] glob_quality_g + | GLocalQVar of lname + | GQVar of Sorts.QVar.t + | GRawQVar of Sorts.QVar.t (* hack for funind *) + +type glob_quality = [`with_constant] glob_quality_g + +type glob_qvar = [`no_constant] glob_quality_g + type glob_sort_name = | GSProp (** representation of [SProp] literal *) | GProp (** representation of [Prop] level *) @@ -41,8 +51,10 @@ type 'a glob_sort_gen = (** levels, occurring in universe instances *) type glob_level = glob_sort_name glob_sort_gen +type glob_instance = glob_quality list * glob_level list + (** sort expressions *) -type glob_sort = (Sorts.QVar.t option * (glob_sort_name * int) list) glob_sort_gen +type glob_sort = (glob_qvar option * (glob_sort_name * int) list) glob_sort_gen type glob_constraint = glob_sort_name * Univ.constraint_type * glob_sort_name @@ -67,7 +79,7 @@ type binding_kind = Explicit | MaxImplicit | NonMaxImplicit (** Representation of an internalized (or in other words globalized) term. *) type 'a glob_constr_r = - | GRef of GlobRef.t * glob_level list option + | GRef of GlobRef.t * glob_instance option (** An identifier that represents a reference to an object defined either in the (global) environment or in the (local) context. *) | GVar of Id.t @@ -89,10 +101,10 @@ type 'a glob_constr_r = | GHole of Evar_kinds.t * Namegen.intro_pattern_naming_expr | GGenarg of Genarg.glob_generic_argument | GCast of 'a glob_constr_g * Constr.cast_kind option * 'a glob_constr_g - | GProj of (Constant.t * glob_level list option) * 'a glob_constr_g list * 'a glob_constr_g + | GProj of (Constant.t * glob_instance option) * 'a glob_constr_g list * 'a glob_constr_g | GInt of Uint63.t | GFloat of Float64.t - | GArray of glob_level list option * 'a glob_constr_g array * 'a glob_constr_g * 'a glob_constr_g + | GArray of glob_instance option * 'a glob_constr_g array * 'a glob_constr_g * 'a glob_constr_g and 'a glob_constr_g = ('a glob_constr_r, 'a) DAst.t and 'a glob_decl_g = Name.t * binding_kind * 'a glob_constr_g option * 'a glob_constr_g diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 292d919e514c9..d71238b8ba960 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -149,12 +149,50 @@ let level_name sigma = function let sigma, u = universe_level_name sigma l in Some (sigma, u) -let sort_info ?loc sigma l = match l with +let glob_level ?loc evd : glob_level -> _ = function + | UAnonymous {rigid} -> new_univ_level_variable ?loc (if rigid then univ_rigid else univ_flexible) evd + | UNamed s -> + match level_name evd s with + | None -> + user_err ?loc + (str "Universe instances cannot contain non-Set small levels, polymorphic" ++ + str " universe instances must be greater or equal to Set."); + | Some r -> r + +let glob_qvar ?loc evd : glob_qvar -> _ = function + | GQVar q -> evd, q + | GLocalQVar {v=Anonymous} -> + let evd, q = new_quality_variable ?loc evd in + evd, q + | GRawQVar q -> + let evd = Evd.merge_sort_variables ~sideff:true evd (Sorts.QVar.Set.singleton q) in + evd, q + | GLocalQVar {v=Name id; loc} -> + try evd, (Evd.quality_of_name evd id) + with Not_found -> + if not (is_strict_universe_declarations()) then + let evd, q = new_quality_variable ?loc evd in + evd, q + else user_err ?loc Pp.(str "Undeclared quality: " ++ Id.print id ++ str".") + +let glob_quality ?loc evd = let open Sorts.Quality in function + | GQConstant q -> evd, QConstant q + | GQVar _ | GLocalQVar _ | GRawQVar _ as q -> + let evd, q = glob_qvar ?loc evd q in + evd, QVar q + +let sort_info ?loc sigma q l = match l with | [] -> assert false -| [GSProp, 0] -> sigma, Sorts.sprop -| [GProp, 0] -> sigma, Sorts.prop +| [GSProp, 0] -> assert (Option.is_empty q); sigma, Sorts.sprop +| [GProp, 0] -> assert (Option.is_empty q); sigma, Sorts.prop | (u, n) :: us -> let open Pp in + let sigma, q = match q with + | None -> sigma, None + | Some q -> + let sigma, q = glob_qvar ?loc sigma q in + sigma, Some q + in let get_level sigma u n = match level_name sigma u with | None -> user_err ?loc @@ -176,7 +214,11 @@ let sort_info ?loc sigma l = match l with in let (sigma, u) = get_level sigma u n in let (sigma, u) = List.fold_left fold (sigma, u) us in - sigma, Sorts.sort_of_univ u + let s = match q with + | None -> Sorts.sort_of_univ u + | Some q -> Sorts.qsort q u + in + sigma, s type inference_hook = env -> evar_map -> Evar.t -> (evar_map * constr) option @@ -406,25 +448,22 @@ let pretype_id pretype loc env sigma id = (*************************************************************************) (* Main pretyping function *) -let glob_level ?loc evd : glob_level -> _ = function - | UAnonymous {rigid} -> new_univ_level_variable ?loc (if rigid then univ_rigid else univ_flexible) evd - | UNamed s -> - match level_name evd s with - | None -> - user_err ?loc - (str "Universe instances cannot contain non-Set small levels, polymorphic" ++ - str " universe instances must be greater or equal to Set."); - | Some r -> r - -let instance ?loc evd l = - let evd, l' = +let instance ?loc evd (ql,ul) = + let evd, ql' = + List.fold_left + (fun (evd, quals) l -> + let evd, l = glob_quality ?loc evd l in + (evd, l :: quals)) (evd, []) + ql + in + let evd, ul' = List.fold_left (fun (evd, univs) l -> let evd, l = glob_level ?loc evd l in (evd, l :: univs)) (evd, []) - l + ul in - evd, Some (UVars.Instance.of_array (Array.of_list (List.rev l'))) + evd, Some (UVars.Instance.of_array (Array.rev_of_list ql', Array.rev_of_list ul')) let pretype_global ?loc rigid env evd gr us = let evd, instance = @@ -441,12 +480,12 @@ let pretype_ref ?loc sigma env ref us = (try let ty = NamedDecl.get_type (lookup_named id !!env) in (match us with - | None | Some [] -> () - | Some us -> + | None | Some ([],[]) -> () + | Some (qs,us) -> let open UnivGen in Loc.raise ?loc (UniverseLengthMismatch { - actual = List.length us; - expect = 0; + actual = List.length qs, List.length us; + expect = 0, 0; })); sigma, make_judge (mkVar id) ty with Not_found -> @@ -464,9 +503,7 @@ let sort ?loc evd : glob_sort -> _ = function let evd, l = new_univ_level_variable ?loc (if rigid then univ_rigid else univ_flexible) evd in evd, ESorts.make (Sorts.sort_of_univ (Univ.Universe.make l)) | UNamed (q, l) -> - (* No user-facing syntax for qualities *) - let () = assert (Option.is_empty q) in - let evd, s = sort_info ?loc evd l in + let evd, s = sort_info ?loc evd q l in evd, ESorts.make s let judge_of_sort ?loc evd s = @@ -508,12 +545,12 @@ let mark_obligation_evar sigma k evc = type 'a pretype_fun = ?loc:Loc.t -> flags:pretype_flags -> type_constraint -> GlobEnv.t -> evar_map -> evar_map * 'a type pretyper = { - pretype_ref : pretyper -> GlobRef.t * glob_level list option -> unsafe_judgment pretype_fun; + pretype_ref : pretyper -> GlobRef.t * glob_instance option -> unsafe_judgment pretype_fun; pretype_var : pretyper -> Id.t -> unsafe_judgment pretype_fun; pretype_evar : pretyper -> existential_name CAst.t * (lident * glob_constr) list -> unsafe_judgment pretype_fun; pretype_patvar : pretyper -> Evar_kinds.matching_var_kind -> unsafe_judgment pretype_fun; pretype_app : pretyper -> glob_constr * glob_constr list -> unsafe_judgment pretype_fun; - pretype_proj : pretyper -> (Constant.t * glob_level list option) * glob_constr list * glob_constr -> unsafe_judgment pretype_fun; + pretype_proj : pretyper -> (Constant.t * glob_instance option) * glob_constr list * glob_constr -> unsafe_judgment pretype_fun; pretype_lambda : pretyper -> Name.t * binding_kind * glob_constr * glob_constr -> unsafe_judgment pretype_fun; pretype_prod : pretyper -> Name.t * binding_kind * glob_constr * glob_constr -> unsafe_judgment pretype_fun; pretype_letin : pretyper -> Name.t * glob_constr * glob_constr option * glob_constr -> unsafe_judgment pretype_fun; @@ -527,7 +564,7 @@ type pretyper = { pretype_cast : pretyper -> glob_constr * cast_kind option * glob_constr -> unsafe_judgment pretype_fun; pretype_int : pretyper -> Uint63.t -> unsafe_judgment pretype_fun; pretype_float : pretyper -> Float64.t -> unsafe_judgment pretype_fun; - pretype_array : pretyper -> glob_level list option * glob_constr array * glob_constr * glob_constr -> unsafe_judgment pretype_fun; + pretype_array : pretyper -> glob_instance option * glob_constr array * glob_constr * glob_constr -> unsafe_judgment pretype_fun; pretype_type : pretyper -> glob_constr -> unsafe_type_judgment pretype_fun; } @@ -1314,14 +1351,14 @@ let pretype_type self c ?loc ~flags valcon (env : GlobEnv.t) sigma = match DAst. fun ?loc ~flags tycon env sigma -> let sigma, u = match u with | None -> sigma, None - | Some [u] -> + | Some ([],[u]) -> let sigma, u = glob_level ?loc sigma u in sigma, Some u - | Some u -> + | Some (qs,us) -> let open UnivGen in Loc.raise ?loc (UniverseLengthMismatch { - actual = List.length u; - expect = 1; + actual = List.length qs, List.length us; + expect = 0, 1; }) in let sigma, tycon' = split_as_array !!env sigma tycon in @@ -1334,7 +1371,7 @@ let pretype_type self c ?loc ~flags valcon (env : GlobEnv.t) sigma = match DAst. let sigma, jdef = eval_pretyper self ~flags (mk_tycon jty.utj_val) env sigma def in let pretype_elem = eval_pretyper self ~flags (mk_tycon jty.utj_val) env in let sigma, jt = Array.fold_left_map pretype_elem sigma t in - let u = UVars.Instance.of_array [| u |] in + let u = UVars.Instance.of_array ([||],[| u |]) in let ta = EConstr.of_constr @@ Typeops.type_of_array !!env u in let j = { uj_val = EConstr.mkArray(EInstance.make u, Array.map (fun j -> j.uj_val) jt, jdef.uj_val, jty.utj_val); diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli index d34472af26f4d..817928d5bfe8a 100644 --- a/pretyping/pretyping.mli +++ b/pretyping/pretyping.mli @@ -158,12 +158,12 @@ type pretype_flags = { type 'a pretype_fun = ?loc:Loc.t -> flags:pretype_flags -> Evardefine.type_constraint -> GlobEnv.t -> evar_map -> evar_map * 'a type pretyper = { - pretype_ref : pretyper -> GlobRef.t * glob_level list option -> unsafe_judgment pretype_fun; + pretype_ref : pretyper -> GlobRef.t * glob_instance option -> unsafe_judgment pretype_fun; pretype_var : pretyper -> Id.t -> unsafe_judgment pretype_fun; pretype_evar : pretyper -> existential_name CAst.t * (lident * glob_constr) list -> unsafe_judgment pretype_fun; pretype_patvar : pretyper -> Evar_kinds.matching_var_kind -> unsafe_judgment pretype_fun; pretype_app : pretyper -> glob_constr * glob_constr list -> unsafe_judgment pretype_fun; - pretype_proj : pretyper -> (Constant.t * glob_level list option) * glob_constr list * glob_constr -> unsafe_judgment pretype_fun; + pretype_proj : pretyper -> (Constant.t * glob_instance option) * glob_constr list * glob_constr -> unsafe_judgment pretype_fun; pretype_lambda : pretyper -> Name.t * binding_kind * glob_constr * glob_constr -> unsafe_judgment pretype_fun; pretype_prod : pretyper -> Name.t * binding_kind * glob_constr * glob_constr -> unsafe_judgment pretype_fun; pretype_letin : pretyper -> Name.t * glob_constr * glob_constr option * glob_constr -> unsafe_judgment pretype_fun; @@ -177,7 +177,7 @@ type pretyper = { pretype_cast : pretyper -> glob_constr * Constr.cast_kind option * glob_constr -> unsafe_judgment pretype_fun; pretype_int : pretyper -> Uint63.t -> unsafe_judgment pretype_fun; pretype_float : pretyper -> Float64.t -> unsafe_judgment pretype_fun; - pretype_array : pretyper -> glob_level list option * glob_constr array * glob_constr * glob_constr -> unsafe_judgment pretype_fun; + pretype_array : pretyper -> glob_instance option * glob_constr array * glob_constr * glob_constr -> unsafe_judgment pretype_fun; pretype_type : pretyper -> glob_constr -> unsafe_type_judgment pretype_fun; } (** Type of pretyping algorithms in open-recursion style. A typical way to diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index f1af29c3432bc..1fdc9e6f31dac 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -1072,16 +1072,13 @@ let checked_sort_cmp_universes _env pb s0 s1 univs = check_sort_cmp_universes pb s0 s1 univs; univs let check_convert_instances ~flex:_ u u' univs = - let u = UVars.Instance.to_array u in - let u' = UVars.Instance.to_array u' in - let fold accu l1 l2 = Constraints.add (l1, Eq, l2) accu in - let cst = Array.fold_left2 fold Constraints.empty u u' in - if Evd.check_constraints univs cst then univs else raise NotConvertible + let csts = UVars.enforce_eq_instances u u' (Sorts.QConstraints.empty,Constraints.empty) in + if Evd.check_quconstraints univs csts then univs else raise NotConvertible (* general conversion and inference functions *) let check_inductive_instances cv_pb variance u1 u2 univs = let csts = get_cumulativity_constraints cv_pb variance u1 u2 in - if (Evd.check_constraints univs csts) then univs + if (Evd.check_quconstraints univs csts) then univs else raise NotConvertible let checked_universes = @@ -1147,17 +1144,7 @@ let univproblem_compare_sorts env pb s0 s1 uset = let univproblem_compare_instances ~flex i0 i1 uset = UnivProblem.enforce_eq_instances_univs flex i0 i1 uset -let univproblem_check_inductive_instances cv_pb variances u u' uset = - let open UnivProblem in - let open UVars.Variance in - let mk u = Sorts.sort_of_univ @@ Univ.Universe.make u in - let fold cstr v u u' = match v with - | Irrelevant -> Set.add (UWeak (u,u')) cstr - | Covariant when cv_pb == Conversion.CUMUL -> Set.add (ULe (mk u, mk u')) cstr - | Covariant | Invariant -> Set.add (UEq (mk u, mk u')) cstr - in - Array.fold_left3 fold uset - variances (UVars.Instance.to_array u) (UVars.Instance.to_array u') +let univproblem_check_inductive_instances = UnivProblem.compare_cumulative_instances let univproblem_univ_state = let open Conversion in @@ -1500,11 +1487,13 @@ let infer_convert_instances ~flex u u' (univs,cstrs as cuniv) = if UGraph.check_eq_instances univs u u' then cuniv else raise NotConvertible else - let cstrs' = UVars.enforce_eq_instances u u' Constraints.empty in + let qcstrs, cstrs' = UVars.enforce_eq_instances u u' Sorts.QUConstraints.empty in + let () = if not (Sorts.QConstraints.trivial qcstrs) then raise NotConvertible in (univs, Constraints.union cstrs cstrs') let infer_inductive_instances cv_pb variance u1 u2 (univs,csts) = - let csts' = get_cumulativity_constraints cv_pb variance u1 u2 in + let qcsts, csts' = get_cumulativity_constraints cv_pb variance u1 u2 in + let () = if not (Sorts.QConstraints.trivial qcsts) then raise NotConvertible in (UGraph.merge_constraints csts' univs, Univ.Constraints.union csts csts') let inferred_universes : (UGraph.t * Univ.Constraints.t) universe_compare = diff --git a/pretyping/structures.ml b/pretyping/structures.ml index 68d774afbbe13..f3a8c0e08505b 100644 --- a/pretyping/structures.ml +++ b/pretyping/structures.ml @@ -360,7 +360,7 @@ let find env sigma (proj,pat) = let bs' = List.map (EConstr.of_constr %> EConstr.Vars.subst_instance_constr u) bs in let params = List.map (fun c -> EConstr.Vars.subst_instance_constr u c) params in let us = List.map (fun c -> EConstr.Vars.subst_instance_constr u c) us in - let sigma = Evd.merge_context_set Evd.univ_flexible sigma ctx' in + let sigma = Evd.merge_full_context_set Evd.univ_flexible sigma ctx' in sigma, { body = t'; constant = c'; abstractions_ty = bs'; nparams; params; cvalue_arguments = us; cvalue_abstraction = n } diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml index 3f8dd043b5575..fecb9847f1220 100644 --- a/pretyping/typeclasses.ml +++ b/pretyping/typeclasses.ml @@ -90,7 +90,7 @@ let classes : typeclasses ref = Summary.ref GlobRef.Map.empty ~name:"classes" let instances : instances ref = Summary.ref GlobRef.Map.empty ~name:"instances" let typeclass_univ_instance (cl, u) = - assert (UVars.AbstractContext.size cl.cl_univs == UVars.Instance.length u); + assert (UVars.eq_sizes (UVars.AbstractContext.size cl.cl_univs) (UVars.Instance.length u)); let subst_ctx c = Context.Rel.map (subst_instance_constr u) c in { cl with cl_context = subst_ctx cl.cl_context; cl_props = subst_ctx cl.cl_props} diff --git a/pretyping/typing.ml b/pretyping/typing.ml index 5d34102c9c93d..68b24b07f4b30 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -360,7 +360,7 @@ let judge_of_float env v = let judge_of_array env sigma u tj defj tyj = let ulev = match UVars.Instance.to_array u with - | [|u|] -> u + | [||], [|u|] -> u | _ -> assert false in let sigma = Evd.set_leq_sort env sigma tyj.utj_type diff --git a/pretyping/vnorm.ml b/pretyping/vnorm.ml index f0d5c43e47afc..2d1cffab61d30 100644 --- a/pretyping/vnorm.ml +++ b/pretyping/vnorm.ml @@ -200,17 +200,15 @@ and nf_whd env sigma whd typ = and nf_univ_args ~nb_univs mk env sigma stk = let u = - if Int.equal nb_univs 0 then UVars.Instance.empty + if UVars.eq_sizes nb_univs (0,0) then UVars.Instance.empty else match stk with - | Zapp args :: _ -> - let inst = - Array.init nb_univs (fun i -> uni_lvl_val (arg args i)) - in - UVars.Instance.of_array inst - | _ -> assert false + | Zapp args :: _ -> + instance_of_args (fst nb_univs + snd nb_univs) args + | _ -> assert false in + assert (UVars.eq_sizes nb_univs (UVars.Instance.length u)); let (t,ty) = mk u in - nf_stk ~from:nb_univs env sigma t ty stk + nf_stk ~from:(fst nb_univs + snd nb_univs) env sigma t ty stk and nf_evar env sigma evk stk = let evi = try Evd.find_undefined sigma evk with Not_found -> assert false in diff --git a/printing/ppconstr.ml b/printing/ppconstr.ml index a1ca73169a7b4..0b74c1983133b 100644 --- a/printing/ppconstr.ml +++ b/printing/ppconstr.ml @@ -167,9 +167,15 @@ let tag_var = tag Tag.variable | [x] -> pr_univ_expr x | l -> str"max(" ++ prlist_with_sep (fun () -> str",") pr_univ_expr l ++ str")" + let pr_quality_expr (type a) (q: a quality_expr_g) = match q with + | CQConstant q -> tag_type (Sorts.Quality.Constants.pr q) + | CQAnon _ -> tag_type (str "_") + | CQVar qid -> tag_type (pr_qualid qid) + | CRawQVar q -> (* TODO names *) tag_type (Sorts.QVar.raw_pr q) + let pr_quality_univ (q, l) = match q with | None -> pr_univ l - | Some q -> (* TODO names *) Sorts.QVar.raw_pr q ++ spc() ++ str "|" ++ spc () ++ pr_univ l + | Some q -> pr_quality_expr q ++ spc() ++ str "|" ++ spc () ++ pr_univ l let pr_univ_annot pr x = str "@{" ++ pr x ++ str "}" @@ -205,8 +211,13 @@ let tag_var = tag Tag.variable let pr_qualid = pr_qualid let pr_patvar = pr_id + let pr_inside_universe_instance (ql,ul) = + (if List.is_empty ql then mt() + else prlist_with_sep spc pr_quality_expr ql ++ strbrk " | ") + ++ prlist_with_sep spc pr_univ_level_expr ul + let pr_universe_instance l = - pr_opt_no_spc (pr_univ_annot (prlist_with_sep spc pr_univ_level_expr)) l + pr_opt_no_spc (pr_univ_annot pr_inside_universe_instance) l let pr_reference qid = if qualid_is_ident qid then tag_var (pr_id @@ qualid_basename qid) diff --git a/printing/printer.ml b/printing/printer.ml index 971b9b01e62b1..32e48505b396a 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -177,41 +177,66 @@ let safe_gen f env sigma c = let safe_pr_lconstr_env = safe_gen pr_lconstr_env let safe_pr_constr_env = safe_gen pr_constr_env +let q_ident = Id.of_string "α" + let u_ident = Id.of_string "u" let universe_binders_with_opt_names orig names = let open Univ in - let orig = UVars.AbstractContext.names orig in - let orig = Array.to_list orig in - let udecl = match names with + let qorig, uorig = UVars.AbstractContext.names orig in + let qorig, uorig as orig = Array.to_list qorig, Array.to_list uorig in + let qdecl, udecl = match names with | None -> orig - | Some udecl -> + | Some (qdecl,udecl) -> try - List.map2 (fun orig {CAst.v = na} -> - match na with - | Anonymous -> orig - | Name id -> Name id) orig udecl + let qs = + List.map2 (fun orig {CAst.v = na} -> + match na with + | Anonymous -> orig + | Name id -> Name id) qorig qdecl + in + let us = + List.map2 (fun orig {CAst.v = na} -> + match na with + | Anonymous -> orig + | Name id -> Name id) uorig udecl + in + qs, us with Invalid_argument _ -> let open UnivGen in raise (UniverseLengthMismatch { - actual = List.length orig; - expect = List.length udecl; + actual = List.length qorig, List.length uorig; + expect = List.length qdecl, List.length udecl; }) in - let fold_named i (ubind,revubind as o) = function + let fold_qnamed i ((qbind,ubind),(revqbind,revubind) as o) = function + | Name id -> let ui = Sorts.QVar.make_var i in + (Id.Map.add id ui qbind, ubind), (Sorts.QVar.Map.add ui id revqbind, revubind) + | Anonymous -> o + in + let fold_unamed i ((qbind,ubind),(revqbind,revubind) as o) = function | Name id -> let ui = Level.var i in - Id.Map.add id ui ubind, Level.Map.add ui id revubind + (qbind, Id.Map.add id ui ubind), (revqbind, Level.Map.add ui id revubind) | Anonymous -> o in - let names = List.fold_left_i fold_named 0 (UnivNames.empty_binders,Level.Map.empty) udecl in - let fold_anons i (u_ident, (ubind, revubind) as o) = function + let names = List.fold_left_i fold_qnamed 0 UnivNames.(empty_binders,empty_rev_binders) qdecl in + let names = List.fold_left_i fold_unamed 0 names udecl in + let fold_qanons i (u_ident, ((qbind,ubind), (revqbind,revubind)) as o) = function + | Name _ -> o + | Anonymous -> + let ui = Sorts.QVar.make_var i in + let id = Namegen.next_ident_away_from u_ident (fun id -> Id.Map.mem id qbind) in + (id, ((Id.Map.add id ui qbind, ubind), (Sorts.QVar.Map.add ui id revqbind, revubind))) + in + let fold_uanons i (u_ident, ((qbind,ubind), (revqbind,revubind)) as o) = function | Name _ -> o | Anonymous -> let ui = Level.var i in let id = Namegen.next_ident_away_from u_ident (fun id -> Id.Map.mem id ubind) in - (id, (Id.Map.add id ui ubind, Level.Map.add ui id revubind)) + (id, ((qbind,Id.Map.add id ui ubind), (revqbind,Level.Map.add ui id revubind))) in - let (_, names) = List.fold_left_i fold_anons 0 (u_ident, names) udecl in + let (_, names) = List.fold_left_i fold_qanons 0 (q_ident, names) qdecl in + let (_, names) = List.fold_left_i fold_uanons 0 (u_ident, names) udecl in names let pr_universe_ctx_set sigma c = @@ -222,7 +247,11 @@ let pr_universe_ctx_set sigma c = let pr_universe_ctx sigma ?variance c = if !Detyping.print_universes && not (UVars.UContext.is_empty c) then - fnl()++pr_in_comment (v 0 (UVars.pr_universe_context (Termops.pr_evd_level sigma) ?variance c)) + fnl()++ + pr_in_comment + (v 0 + (UVars.pr_universe_context (Termops.pr_evd_qvar sigma) (Termops.pr_evd_level sigma) + ?variance c)) else mt() @@ -231,8 +260,9 @@ let pr_abstract_universe_ctx sigma ?variance ?priv c = let priv = Option.default Univ.ContextSet.empty priv in let has_priv = not (ContextSet.is_empty priv) in if !Detyping.print_universes && (not (UVars.AbstractContext.is_empty c) || has_priv) then + let prqvar u = Termops.pr_evd_qvar sigma u in let prlev u = Termops.pr_evd_level sigma u in - let pub = (if has_priv then str "Public universes:" ++ fnl() else mt()) ++ v 0 (UVars.pr_abstract_universe_context prlev ?variance c) in + let pub = (if has_priv then str "Public universes:" ++ fnl() else mt()) ++ v 0 (UVars.pr_abstract_universe_context prqvar prlev ?variance c) in let priv = if has_priv then fnl() ++ str "Private universes:" ++ fnl() ++ v 0 (Univ.pr_universe_context_set prlev priv) else mt() in fnl()++pr_in_comment (pub ++ priv) else @@ -250,6 +280,7 @@ let pr_global = pr_global_env Id.Set.empty let pr_universe_instance_constraints evd inst csts = let open Univ in + let prqvar = Termops.pr_evd_qvar evd in let prlev = Termops.pr_evd_level evd in let pcsts = if Constraints.is_empty csts then mt() else str " |= " ++ @@ -257,7 +288,7 @@ let pr_universe_instance_constraints evd inst csts = (fun (u,d,v) -> hov 0 (prlev u ++ pr_constraint_type d ++ prlev v)) (Constraints.elements csts) in - str"@{" ++ UVars.Instance.pr prlev inst ++ pcsts ++ str"}" + str"@{" ++ UVars.Instance.pr prqvar prlev inst ++ pcsts ++ str"}" let pr_universe_instance evd inst = pr_universe_instance_constraints evd inst Univ.Constraints.empty diff --git a/printing/printer.mli b/printing/printer.mli index 4520465fa148e..21a71f628ba3c 100644 --- a/printing/printer.mli +++ b/printing/printer.mli @@ -140,7 +140,7 @@ val pr_universes : evar_map -> Inefficient on large contexts due to name generation. *) val universe_binders_with_opt_names : UVars.AbstractContext.t -> - UnivNames.univ_name_list option -> UnivNames.universe_binders * Id.t Univ.Level.Map.t + UnivNames.full_name_list option -> UnivNames.universe_binders * UnivNames.rev_binders (** Printing global references using names as short as possible *) diff --git a/proofs/clenv.ml b/proofs/clenv.ml index ca35bb3197333..83af4ff2b5eac 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -76,9 +76,9 @@ let clenv_refresh env sigma ctx clenv = let evd = Evd.meta_merge (Evd.meta_list clenv.evd) (Evd.clear_metas sigma) in match ctx with | Some ctx -> - let (subst, ctx) = UnivGen.fresh_universe_context_set_instance ctx in + let (subst, ctx) = UnivGen.fresh_full_context_instance ctx in let emap c = Vars.subst_univs_level_constr subst c in - let evd = Evd.merge_context_set Evd.univ_flexible evd ctx in + let evd = Evd.merge_full_context_set Evd.univ_flexible evd ctx in (* Only metas are mentioning the old universes. *) mk_clausenv env (Evd.map_metas emap evd) clenv.metas (emap clenv.templval) diff --git a/proofs/clenv.mli b/proofs/clenv.mli index 04bc5bc3a0b75..ad756ad9be44a 100644 --- a/proofs/clenv.mli +++ b/proofs/clenv.mli @@ -27,7 +27,7 @@ val clenv_evd : clausenv -> Evd.evar_map (* Ad-hoc primitives *) val update_clenv_evd : clausenv -> evar_map -> clausenv val clenv_strip_proj_params : clausenv -> clausenv -val clenv_refresh : env -> evar_map -> Univ.ContextSet.t option -> clausenv -> clausenv +val clenv_refresh : env -> evar_map -> UnivGen.full_context_set option -> clausenv -> clausenv val clenv_arguments : clausenv -> metavariable list (** subject of clenv (instantiated) *) diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml index 73c21af51d3c4..29201399e8a22 100644 --- a/tactics/autorewrite.ml +++ b/tactics/autorewrite.ml @@ -367,6 +367,7 @@ let one_base where conds tac_main bas = Proofview.Goal.enter begin fun gl -> let sigma = Proofview.Goal.sigma gl in let subst, ctx' = UnivGen.fresh_universe_context_set_instance h.rew_ctx in + let subst = Sorts.QVar.Map.empty, subst in let c' = Vars.subst_univs_level_constr subst h.rew_lemma in let sigma = Evd.merge_context_set Evd.univ_flexible sigma ctx' in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma) (rewrite h.rew_l2r c' tc) diff --git a/tactics/cbn.ml b/tactics/cbn.ml index 4d86e0ad8d05f..a9319a69f464b 100644 --- a/tactics/cbn.ml +++ b/tactics/cbn.ml @@ -215,7 +215,7 @@ struct | Cst_const (c, u) -> if UVars.Instance.is_empty u then Constant.debug_print c else str"(" ++ Constant.debug_print c ++ str ", " ++ - UVars.Instance.pr Univ.Level.raw_pr u ++ str")" + UVars.Instance.pr Sorts.QVar.raw_pr Univ.Level.raw_pr u ++ str")" | Cst_proj p -> str".(" ++ Constant.debug_print (Projection.constant p) ++ str")" @@ -450,21 +450,28 @@ let magically_constant_of_fixbody env sigma reference bd = function let csts = EConstr.eq_constr_universes env sigma (EConstr.of_constr t) bd in begin match csts with | Some csts -> + let addqs l r (qs,us) = Sorts.QVar.Map.add l r qs, us in + let addus l r (qs,us) = qs, Univ.Level.Map.add l r us in let subst = Set.fold (fun cst acc -> - let l, r = match cst with - | ULub (u, v) | UWeak (u, v) -> u, v + match cst with + | QEq (a,b) | QLeq (a,b) -> + let a = match a with + | QVar q -> q + | _ -> assert false + in + addqs a b acc + | ULub (u, v) | UWeak (u, v) -> addus u v acc | UEq (u, v) | ULe (u, v) -> + (* XXX add something when qsort? *) let get u = match u with | Sorts.SProp | Sorts.Prop -> assert false | Sorts.Set -> Level.set | Sorts.Type u | Sorts.QSort (_, u) -> Option.get (Universe.level u) in - get u, get v - in - Univ.Level.Map.add l r acc) - csts Univ.Level.Map.empty + addus (get u) (get v) acc) + csts UVars.empty_full_subst in - let inst = UVars.subst_univs_level_instance subst u in + let inst = UVars.subst_full_level_instance subst u in mkConstU (cst, EInstance.make inst) | None -> bd end diff --git a/tactics/eqschemes.ml b/tactics/eqschemes.ml index 4a21f31295377..772bbaa17840e 100644 --- a/tactics/eqschemes.ml +++ b/tactics/eqschemes.ml @@ -72,7 +72,7 @@ let fresh env id avoid = let freshid = next_global_ident_away id avoid in freshid, Id.Set.add freshid avoid let with_context_set ctx (b, ctx') = - (b, Univ.ContextSet.union ctx ctx') + (b, UnivGen.full_context_union ctx ctx') let build_dependent_inductive ind (mib,mip) = let realargs,_ = List.chop mip.mind_nrealdecls mip.mind_arity_ctxt in @@ -408,7 +408,7 @@ let build_l2r_rew_scheme dep env handle ind kind = rel_vect 1 nrealargs; [|mkRel 1|]]) in let s, ctx' = UnivGen.fresh_sort_in_family kind in - let ctx = Univ.ContextSet.union ctx ctx' in + let ctx = UnivGen.full_context_union ctx ctx' in let s = mkSort s in let rci = Sorts.Relevant in (* TODO relevance *) let ci = make_case_info env ind rci RegularStyle in @@ -516,7 +516,7 @@ let build_l2r_forward_rew_scheme dep env ind kind = let realsign_ind_P n aP = name_context env ((LocalAssum (make_annot (Name varH) indr,aP))::realsign_P n) in let s, ctx' = UnivGen.fresh_sort_in_family kind in - let ctx = Univ.ContextSet.union ctx ctx' in + let ctx = UnivGen.full_context_union ctx ctx' in let s = mkSort s in let rci = Sorts.Relevant in let ci = make_case_info env ind rci RegularStyle in @@ -598,7 +598,7 @@ let build_r2l_forward_rew_scheme dep env ind kind = let realsign_ind = name_context env ((LocalAssum (make_annot (Name varH) indr,applied_ind))::realsign) in let s, ctx' = UnivGen.fresh_sort_in_family kind in - let ctx = Univ.ContextSet.union ctx ctx' in + let ctx = UnivGen.full_context_union ctx ctx' in let s = mkSort s in let rci = Sorts.Relevant in (* TODO relevance *) let ci = make_case_info env ind rci RegularStyle in @@ -809,8 +809,11 @@ let build_congr env (eq,refl,ctx) ind = let varf,avoid = fresh env (Id.of_string "f") avoid in let rci = Sorts.Relevant in (* TODO relevance *) let ci = make_case_info env ind rci RegularStyle in - let uni, ctx = Univ.extend_in_context_set (UnivGen.new_global_univ ()) ctx in - let ctx = (fst ctx, UnivSubst.enforce_leq uni (univ_of_eq env eq) (snd ctx)) in + let uni, ctx' = UnivGen.new_global_univ () in + let ctx = + let (qs,us),csts = ctx in + let us, csts = Univ.ContextSet.union (us,csts) ctx' in + ((qs, us), UnivSubst.enforce_leq uni (univ_of_eq env eq) csts) in let c = my_it_mkLambda_or_LetIn paramsctxt (mkNamedLambda (make_annot varB Sorts.Relevant) (mkType uni) @@ -845,4 +848,4 @@ let build_congr env (eq,refl,ctx) ind = let congr_scheme_kind = declare_individual_scheme_object "_congr" (fun env _ ind -> (* May fail if equality is not defined *) - build_congr env (get_coq_eq env Univ.ContextSet.empty) ind) + build_congr env (get_coq_eq env UnivGen.empty_full_context) ind) diff --git a/tactics/hints.ml b/tactics/hints.ml index f59bddb42396d..b39ac627ba85c 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -128,7 +128,7 @@ type hints_path = GlobRef.t hints_path_gen type hint_term = | IsGlobRef of GlobRef.t - | IsConstr of constr * Univ.ContextSet.t option (* None if monomorphic *) + | IsConstr of constr * UnivGen.full_context_set option (* None if monomorphic *) type 'a with_uid = { obj : 'a; @@ -138,14 +138,14 @@ type 'a with_uid = { type raw_hint = { rhint_term : constr; rhint_type : types; - rhint_uctx : Univ.ContextSet.t option; + rhint_uctx : UnivGen.full_context_set option; rhint_arty : int; (* Number of goals generated by the intended tactic *) } type hint = { hint_term : constr; hint_type : types; - hint_uctx : Univ.ContextSet.t option; (* None if monomorphic *) + hint_uctx : UnivGen.full_context_set option; (* None if monomorphic *) hint_clnv : Clenv.clausenv; hint_arty : int; (* Number of goals generated by the intended tactic *) } @@ -337,7 +337,7 @@ let lookup_tacs env sigma concl se = let merge_context_set_opt sigma ctx = match ctx with | None -> sigma -| Some ctx -> Evd.merge_context_set Evd.univ_flexible sigma ctx +| Some ctx -> Evd.merge_full_context_set Evd.univ_flexible sigma ctx let instantiate_hint env sigma p = let mk_clenv { rhint_term = c; rhint_type = cty; rhint_uctx = ctx; rhint_arty = ar } = @@ -1420,7 +1420,7 @@ let prepare_hint env init (sigma,c) = subst := (evar,mkVar id)::!subst; mkNamedLambda sigma (make_annot id Sorts.Relevant) t (iter (replace_term sigma evar (mkVar id) c)) in let c' = iter c in - let diff = Univ.ContextSet.diff (Evd.universe_context_set sigma) (Evd.universe_context_set init) in + let diff = UnivGen.diff_full_context (Evd.full_universe_context_set sigma) (Evd.full_universe_context_set init) in (c', diff) let warn_non_local_section_hint = @@ -1475,7 +1475,7 @@ let expand_constructor_hints env sigma lems = (fun i -> IsGlobRef (GlobRef.ConstructRef ((ind,i+1)))) | _ -> let (c, ctx) = prepare_hint env sigma (evd,lem) in - let ctx = if Univ.ContextSet.is_empty ctx then None else Some ctx in + let ctx = if UnivGen.is_empty_full_context ctx then None else Some ctx in [IsConstr (c, ctx)]) lems (* builds a hint database from a constr signature *) (* typically used with (lid, ltyp) = pf_hyps_types *) @@ -1772,9 +1772,9 @@ let fresh_hint env sigma h = | None -> sigma, c | Some ctx -> (* Refresh the instance of the hint *) - let (subst, ctx) = UnivGen.fresh_universe_context_set_instance ctx in + let (subst, ctx) = UnivGen.fresh_full_context_instance ctx in let c = Vars.subst_univs_level_constr subst c in - let sigma = Evd.merge_context_set Evd.univ_flexible sigma ctx in + let sigma = Evd.merge_full_context_set Evd.univ_flexible sigma ctx in sigma, c let hint_res_pf ?with_evars ?with_classes ?flags h = diff --git a/tactics/hints.mli b/tactics/hints.mli index cc6d3095fedbe..29083345e60d8 100644 --- a/tactics/hints.mli +++ b/tactics/hints.mli @@ -42,7 +42,7 @@ type 'a hint_ast = type hint -val hint_as_term : hint -> Univ.ContextSet.t option * constr +val hint_as_term : hint -> UnivGen.full_context_set option * constr type 'a hints_path_atom_gen = | PathHints of 'a list @@ -196,7 +196,7 @@ val add_hints : locality:hint_locality -> hint_db_name list -> hints_entry -> un val hint_globref : GlobRef.t -> hint_term -val hint_constr : constr * Univ.ContextSet.t option -> hint_term +val hint_constr : constr * UnivGen.full_context_set option -> hint_term [@ocaml.deprecated "Declare a hint constant instead"] (** A constr which is Hint'ed will be: diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 250910b1e27d2..2771ce06cc9cb 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -3386,7 +3386,8 @@ let constr_eq ~strict x y = match EConstr.eq_constr_universes env evd x y with | Some csts -> if strict then - if UnivProblem.Set.check (Evd.universes evd) csts then Proofview.tclUNIT () + if UState.check_universe_constraints (Evd.evar_universe_context evd) csts + then Proofview.tclUNIT () else let info = Exninfo.reify () in fail_universes ~info diff --git a/test-suite/misc/poly-capture-global-univs/src/evilImpl.ml b/test-suite/misc/poly-capture-global-univs/src/evilImpl.ml index b8060240b5e53..f89bcab62c6a4 100644 --- a/test-suite/misc/poly-capture-global-univs/src/evilImpl.ml +++ b/test-suite/misc/poly-capture-global-univs/src/evilImpl.ml @@ -2,6 +2,7 @@ open Names let evil name name_f = let open Univ in + let open UVars in let open Constr in let kind = Decls.(IsDefinition Definition) in let u = Level.var 0 in @@ -13,7 +14,7 @@ let evil name name_f = let tc = mkConst tc in let fe = Declare.definition_entry - ~univs:(UState.Polymorphic_entry (UContext.make [|Anonymous|] (Instance.of_array [|u|],Constraints.empty)), UnivNames.empty_binders) + ~univs:(UState.Polymorphic_entry (UContext.make ([||],[|Anonymous|]) (Instance.of_array ([||],[|u|]),Constraints.empty)), UnivNames.empty_binders) ~types:(Term.mkArrowR tc tu) (mkLambda (Context.nameR (Id.of_string "x"), tc, mkRel 1)) in diff --git a/test-suite/output/PrintGrammar.out b/test-suite/output/PrintGrammar.out index 02a8a3b9ac1a0..601ab81940bd9 100644 --- a/test-suite/output/PrintGrammar.out +++ b/test-suite/output/PrintGrammar.out @@ -144,7 +144,8 @@ Entry term is Entry univ_annot is [ LEFTA - [ "@{"; LIST0 universe_level; "}" + [ "@{"; LIST0 univ_level_or_quality; + OPT [ "|"; LIST0 univ_level_or_quality ]; "}" | ] ] Entry fix_decls is diff --git a/test-suite/output/SortQuality.out b/test-suite/output/SortQuality.out index d348761cec527..d229f2e81b18c 100644 --- a/test-suite/output/SortQuality.out +++ b/test-suite/output/SortQuality.out @@ -1 +1 @@ -Type@{(* α1 *) SortQuality.1} +Type@{α1 | SortQuality.1} diff --git a/test-suite/success/sort_poly.v b/test-suite/success/sort_poly.v new file mode 100644 index 0000000000000..e5391b2eb878d --- /dev/null +++ b/test-suite/success/sort_poly.v @@ -0,0 +1,36 @@ + +Fail Definition foo@{| Set < Set } := Type. + +Definition foo@{u| Set < u} := Type@{u}. + +Set Universe Polymorphism. + +Definition bar@{s | u | Set < u} := Type@{u}. +Set Printing Universes. +Print bar. + +Definition baz@{s | | } := Type@{s | Set}. +Print baz. + +Definition potato@{s | + | } := Type. + +Check eq_refl : Prop = baz@{Prop | }. + +Inductive bob@{s| |} : Prop := . + + +Definition qsort@{s | u |} := Type@{s | u}. + +Monomorphic Universe U. + +Definition tU := Type@{U}. +Definition qU := qsort@{Type | U}. + +Definition q1 := Eval lazy in qU. +Check eq_refl : q1 = tU. + +Definition q2 := Eval vm_compute in qU. +Check eq_refl : q2 = tU. + +Definition q3 := Eval native_compute in qU. +Check eq_refl : q3 = tU. diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml index ea891257e2600..c61bc080e9f44 100644 --- a/vernac/auto_ind_decl.ml +++ b/vernac/auto_ind_decl.ml @@ -285,7 +285,7 @@ let pred_context env ci params u nas = let realdecls, _ = List.chop mip.mind_nrealdecls mip.mind_arity_ctxt in let self = let args = Context.Rel.instance mkRel 0 mip.mind_arity_ctxt in - let inst = UVars.(Instance.of_array (Array.init (Instance.length u) Univ.Level.var)) in + let inst = UVars.Instance.(abstract_instance (length u)) in mkApp (mkIndU (ci.ci_ind, inst), args) in let realdecls = RelDecl.LocalAssum (Context.anonR, self) :: realdecls in @@ -1158,7 +1158,7 @@ let make_bl_scheme env handle mind = (* Setting universes *) let auctx = Declareops.universes_context mib.mind_universes in let u, uctx = UnivGen.fresh_instance_from auctx None in - let uctx = UState.merge ~sideff:false UState.univ_rigid (UState.from_env env) uctx in + let uctx = UState.merge_full ~sideff:false UState.univ_rigid (UState.from_env env) uctx in let ind = (mind,0) in let nparrec = mib.mind_nparams_rec in @@ -1289,7 +1289,7 @@ let make_lb_scheme env handle mind = (* Setting universes *) let auctx = Declareops.universes_context mib.mind_universes in let u, uctx = UnivGen.fresh_instance_from auctx None in - let uctx = UState.merge ~sideff:false UState.univ_rigid (UState.from_env env) uctx in + let uctx = UState.merge_full ~sideff:false UState.univ_rigid (UState.from_env env) uctx in let nparrec = mib.mind_nparams_rec in let lnonparrec,lnamesparrec = @@ -1476,7 +1476,7 @@ let make_eq_decidability env handle mind = (* Setting universes *) let auctx = Declareops.universes_context mib.mind_universes in let u, uctx = UnivGen.fresh_instance_from auctx None in - let uctx = UState.merge ~sideff:false UState.univ_rigid (UState.from_env env) uctx in + let uctx = UState.merge_full ~sideff:false UState.univ_rigid (UState.from_env env) uctx in let lnonparrec,lnamesparrec = Inductive.inductive_nonrec_rec_paramdecls (mib,u) in diff --git a/vernac/comAssumption.ml b/vernac/comAssumption.ml index f79f85b797055..3808dec0e894f 100644 --- a/vernac/comAssumption.ml +++ b/vernac/comAssumption.ml @@ -163,7 +163,7 @@ let do_assumptions ~program_mode ~poly ~scope ~kind ?deprecation nl l = let nf_evar c = EConstr.to_constr sigma c in let uvars, l = List.fold_left_map (fun uvars (coe,t,imps) -> let t = nf_evar t in - let uvars = Univ.Level.Set.union uvars (Vars.universes_of_constr t) in + let uvars = Univ.Level.Set.union uvars (snd (Vars.universes_of_constr t)) in uvars, (coe,t,imps)) Univ.Level.Set.empty l in diff --git a/vernac/comHints.ml b/vernac/comHints.ml index 731a18e96f12e..da86c022a6489 100644 --- a/vernac/comHints.ml +++ b/vernac/comHints.ml @@ -105,11 +105,10 @@ let interp_hints ~poly h = let c, uctx = Constrintern.interp_constr env sigma c in let uctx = UState.normalize_variables uctx in let c = Evarutil.nf_evar (Evd.from_ctx uctx) c in - let diff = UState.context_set uctx in let c = - if poly then (c, Some diff) + if poly then (c, Some (UState.full_context_set uctx)) else - let () = DeclareUctx.declare_universe_context ~poly:false diff in + let () = Global.push_context_set ~strict:true (UState.context_set uctx) in (c, None) in (PathAny, Hints.hint_constr c) [@ocaml.warning "-3"] diff --git a/vernac/comInductive.ml b/vernac/comInductive.ml index 9221b49769bbe..53d73c6b71094 100644 --- a/vernac/comInductive.ml +++ b/vernac/comInductive.ml @@ -208,7 +208,7 @@ let extract_level env evd min tys = in sup_list min sorts let is_flexible_sort evd s = match s with -| QSort _ -> assert false +| QSort _ -> false | Set | Prop | SProp -> false | Type u -> match Univ.Universe.level u with @@ -291,7 +291,7 @@ let inductive_levels env evd arities inds = | Prop | SProp -> None | Set -> Some Univ.Universe.type0 | Type u -> Some u - | QSort _ -> assert false + | QSort (_,u) -> Some u in let levels = List.map map destarities in let cstrs_levels, sizes = @@ -420,7 +420,7 @@ let template_polymorphism_candidate uctx params entry concl = match concl with | Some (Set | SProp | Prop) -> Univ.Level.Set.empty | Some (Type u) -> let ctor_levels = - let add_levels c levels = Univ.Level.Set.union levels (Vars.universes_of_constr c) in + let add_levels c levels = Univ.Level.Set.union levels (snd (Vars.universes_of_constr c)) in let param_levels = List.fold_left (fun levels d -> match d with | LocalAssum _ -> levels @@ -502,7 +502,7 @@ let check_param = function let restrict_inductive_universes sigma ctx_params arities constructors = let merge_universes_of_constr c = - Univ.Level.Set.union (EConstr.universes_of_constr sigma (EConstr.of_constr c)) in + Univ.Level.Set.union (snd (EConstr.universes_of_constr sigma (EConstr.of_constr c))) in let uvars = Univ.Level.Set.empty in let uvars = Context.Rel.(fold_outside (Declaration.fold_constr merge_universes_of_constr) ctx_params ~init:uvars) in let uvars = List.fold_right merge_universes_of_constr arities uvars in @@ -524,7 +524,7 @@ let variance_of_entry ~cumulative ~variances uctx = if not cumulative then begin check_trivial_variances variances; None end else let lvs = Array.length variances in - let lus = UVars.UContext.size uctx in + let _, lus = UVars.UContext.size uctx in assert (lvs <= lus); Some (Array.append variances (Array.make (lus - lvs) None)) @@ -823,7 +823,7 @@ let do_mutual_inductive ~template udecl indl ~cumulative ~poly ?typing_flags ~pr | Polymorphic_ind_entry uctx -> (UState.Polymorphic_entry uctx, UnivNames.empty_binders) in (* Declare the global universes *) - DeclareUctx.declare_universe_context ~poly:false uctx; + Global.push_context_set ~strict:true uctx; (* Declare the mutual inductive block with its associated schemes *) ignore (DeclareInd.declare_mutual_inductive_with_eliminations ?typing_flags ~indlocs mie binders implicits); (* Declare the possible notations of inductive types *) diff --git a/vernac/comPrimitive.ml b/vernac/comPrimitive.ml index 40bb2fda68702..3687b04663fb9 100644 --- a/vernac/comPrimitive.ml +++ b/vernac/comPrimitive.ml @@ -33,7 +33,7 @@ let do_primitive id udecl prim typopt = let env = Global.env () in let evd, udecl = Constrintern.interp_univ_decl_opt env udecl in let auctx = CPrimitives.op_or_type_univs prim in - let evd, u = Evd.with_context_set UState.univ_flexible evd (UnivGen.fresh_instance auctx) in + let evd, u = Evd.with_full_context_set UState.univ_flexible evd (UnivGen.fresh_instance auctx) in let expected_typ = EConstr.of_constr @@ Typeops.type_of_prim_or_type env u prim in let evd, (typ,impls) = Constrintern.(interp_type_evars_impls ~impls:empty_internalization_env) @@ -47,7 +47,7 @@ let do_primitive id udecl prim typopt = in Pretyping.check_evars_are_solved ~program_mode:false env evd; let evd = Evd.minimize_universes evd in - let uvars = EConstr.universes_of_constr evd typ in + let _qvars, uvars = EConstr.universes_of_constr evd typ in let evd = Evd.restrict_universe_context evd uvars in let typ = EConstr.to_constr evd typ in let univ_entry = Evd.check_univ_decl ~poly:(not (UVars.AbstractContext.is_empty auctx)) evd udecl in diff --git a/vernac/declare.ml b/vernac/declare.ml index 38207644d9fa2..5248fdf66d88a 100644 --- a/vernac/declare.ml +++ b/vernac/declare.ml @@ -387,7 +387,7 @@ let declare_constant_core ~name ~typing_flags cd = let ubinders = make_ubinders ctx de.proof_entry_universes in (* We register the global universes after exporting side-effects, since the latter depend on the former. *) - let () = DeclareUctx.declare_universe_context ~poly:false ctx in + let () = Global.push_context_set ~strict:true ctx in let cd = Entries.DefinitionEntry e in ConstantEntry cd, false, ubinders, None else @@ -397,12 +397,12 @@ let declare_constant_core ~name ~typing_flags cd = let de = { de with proof_entry_body = body } in let cd, ctx = cast_opaque_proof_entry EffectEntry de in let ubinders = make_ubinders ctx de.proof_entry_universes in - let () = DeclareUctx.declare_universe_context ~poly:false ctx in + let () = Global.push_context_set ~strict:true ctx in OpaqueEntry cd, false, ubinders, Some (body, feedback_id) | ParameterEntry e -> let univ_entry, ctx = extract_monomorphic (fst e.parameter_entry_universes) in let ubinders = make_ubinders ctx e.parameter_entry_universes in - let () = DeclareUctx.declare_universe_context ~poly:false ctx in + let () = Global.push_context_set ~strict:true ctx in let e = { Entries.parameter_entry_secctx = e.parameter_entry_secctx; Entries.parameter_entry_type = e.parameter_entry_type; @@ -417,7 +417,7 @@ let declare_constant_core ~name ~typing_flags cd = let univ_entry, ctx = extract_monomorphic univs in Some (typ, univ_entry), ubinders, ctx in - let () = DeclareUctx.declare_universe_context ~poly:false ctx in + let () = Global.push_context_set ~strict:true ctx in let e = { Entries.prim_entry_type = typ; Entries.prim_entry_content = e.prim_entry_content; @@ -499,28 +499,26 @@ let declare_variable_core ~name ~kind d = if Decls.variable_exists name then raise (DeclareUniv.AlreadyDeclared (None, name)); - let impl,opaque,univs = match d with (* Fails if not well-typed *) + let impl,opaque = match d with (* Fails if not well-typed *) | SectionLocalAssum {typ;impl;univs} -> - let poly, uctx = match fst univs with - | UState.Monomorphic_entry uctx -> false, uctx - | UState.Polymorphic_entry uctx -> true, UVars.UContext.to_context_set uctx + let () = match fst univs with + | UState.Monomorphic_entry uctx -> Global.push_context_set ~strict:true uctx + | UState.Polymorphic_entry uctx -> Global.push_section_context uctx in - let () = DeclareUctx.declare_universe_context ~poly uctx in let () = Global.push_named_assum (name,typ) in - impl, true, univs + impl, true | SectionLocalDef { clearbody; entry = de } -> (* The body should already have been forced upstream because it is a section-local definition, but it's not enforced by typing *) let ((body, body_uctx), eff) = Future.force de.proof_entry_body in let () = export_side_effects eff in - let poly, type_uctx = match fst de.proof_entry_universes with - | UState.Monomorphic_entry uctx -> false, uctx - | UState.Polymorphic_entry uctx -> true, UVars.UContext.to_context_set uctx + let () = match fst de.proof_entry_universes with + | UState.Monomorphic_entry uctx -> + Global.push_context_set ~strict:true (Univ.ContextSet.union uctx body_uctx) + | UState.Polymorphic_entry uctx -> + assert (Univ.ContextSet.is_empty body_uctx); + Global.push_section_context uctx in - let univs = Univ.ContextSet.union body_uctx type_uctx in - (* We must declare the universe constraints before type-checking the - term. *) - let () = DeclareUctx.declare_universe_context ~poly univs in (* NB: de.proof_entry_secctx is ignored *) let se = { Entries.secdef_body = body; @@ -529,7 +527,7 @@ let declare_variable_core ~name ~kind d = let () = Global.push_named_def (name, se) in let opaque = de.proof_entry_opaque in let () = if opaque then warn_opaque_let name in - Glob_term.Explicit, opaque || clearbody, de.proof_entry_universes + Glob_term.Explicit, opaque || clearbody in Nametab.push (Nametab.Until 1) (Libnames.make_path DirPath.empty name) (GlobRef.VarRef name); Decls.(add_variable_data name {opaque;kind}); @@ -705,7 +703,7 @@ let declare_mutually_recursive_core ~info ~cinfo ~opaque ~ntns ~uctx ~rec_declar (* XXX: Obligations don't do this, this seems like a bug? *) if restrict_ucontext then - let uctx = UState.restrict uctx vars in + let uctx = UState.restrict uctx (snd vars) in let univs = UState.check_univ_decl ~poly uctx udecl in uctx, univs else @@ -1705,7 +1703,7 @@ let prepare_proof ~unsafe_typ { proof } = let to_constr_body c = match EConstr.to_constr_opt evd c with | Some p -> - Vars.universes_of_constr p, p + snd (Vars.universes_of_constr p), p | None -> CErrors.user_err Pp.(str "Some unresolved existential variables remain") in @@ -1713,7 +1711,7 @@ let prepare_proof ~unsafe_typ { proof } = if unsafe_typ then let t = EConstr.Unsafe.to_constr t in - Vars.universes_of_constr t, t + snd (Vars.universes_of_constr t), t else to_constr_body t in (* ppedrot: FIXME, this is surely wrong. There is no reason to duplicate @@ -1825,8 +1823,8 @@ let close_proof_delayed ~feedback_id ps (fpl : closed_proof_output Future.comput this will prevent the body from adding universes and constraints. *) let uctx = UState.constrain_variables (fst (UState.context_set initial_euctx)) uctx in let used_univs = Univ.Level.Set.union - (Vars.universes_of_constr types) - (Vars.universes_of_constr pt) + (snd (Vars.universes_of_constr types)) + (snd (Vars.universes_of_constr pt)) in let uctx = UState.restrict uctx used_univs in let uctx = UState.check_mono_univ_decl uctx udecl in diff --git a/vernac/declareUctx.ml b/vernac/declareUctx.ml deleted file mode 100644 index 3db3450570989..0000000000000 --- a/vernac/declareUctx.ml +++ /dev/null @@ -1,33 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * Copyright INRIA, CNRS and contributors *) -(* (* Having Prop/Set/Var as section universes makes no sense *) - assert false - | Some na -> - try - let qid = Nametab.shortest_qualid_of_universe Names.Id.Map.empty na in - Names.Name (Libnames.qualid_basename qid) - with Not_found -> - (* Best-effort naming from the string representation of the level. - See univNames.ml for a similar hack. *) - Names.Name (Names.Id.of_string_soft (Univ.Level.to_string lvl)) - in - Array.map map (UVars.Instance.to_array inst) - -let declare_universe_context ~poly ctx = - if poly then - let uctx = UVars.UContext.of_context_set name_instance ctx in - Global.push_section_context uctx - else - Global.push_context_set ~strict:true ctx diff --git a/vernac/declareUctx.mli b/vernac/declareUctx.mli deleted file mode 100644 index 7ecfab04f2342..0000000000000 --- a/vernac/declareUctx.mli +++ /dev/null @@ -1,11 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * Copyright INRIA, CNRS and contributors *) -(* Univ.ContextSet.t -> unit diff --git a/vernac/declareUniv.ml b/vernac/declareUniv.ml index 2acd27932a53e..2dc393083b21d 100644 --- a/vernac/declareUniv.ml +++ b/vernac/declareUniv.ml @@ -29,8 +29,8 @@ type universe_source = type universe_name_decl = { udecl_src : universe_source; - udecl_named : (Id.t * Univ.UGlobal.t) list; - udecl_anon : Univ.UGlobal.t list; + udecl_named : (Id.t * UGlobal.t) list; + udecl_anon : UGlobal.t list; } let check_exists_universe sp = @@ -112,9 +112,11 @@ let declare_univ_binders gr (univs, pl) = match univs with | UState.Polymorphic_entry _ -> () | UState.Monomorphic_entry (levels, _) -> + let qs, pl = pl in + assert (Id.Map.is_empty qs); (* First the explicitly named universes *) let named, univs = Id.Map.fold (fun id univ (named,univs) -> - let univs = match Univ.Level.name univ with + let univs = match Level.name univ with | None -> assert false (* having Prop/Set/Var as binders is nonsense *) | Some univ -> (id,univ)::univs in @@ -135,12 +137,21 @@ let do_universe ~poly l = (Pp.str"Cannot declare polymorphic universes outside sections.") in let l = List.map (fun {CAst.v=id} -> (id, UnivGen.new_univ_global ())) l in - let ctx = List.fold_left (fun ctx (_,qid) -> Univ.Level.Set.add (Univ.Level.make qid) ctx) - Univ.Level.Set.empty l, Univ.Constraints.empty - in let src = if poly then BoundUniv else UnqualifiedUniv in let () = input_univ_names (src, l, []) in - DeclareUctx.declare_universe_context ~poly ctx + match poly with + | false -> + let ctx = List.fold_left (fun ctx (_,qid) -> Level.Set.add (Level.make qid) ctx) + Level.Set.empty l, Constraints.empty + in + Global.push_context_set ~strict:true ctx + | true -> + let names = CArray.map_of_list (fun (na,_) -> Name na) l in + let us = CArray.map_of_list (fun (_,l) -> Level.make l) l in + let ctx = + UVars.UContext.make ([||],names) (UVars.Instance.of_array ([||],us), Constraints.empty) + in + Global.push_section_context ctx let do_constraint ~poly l = let open Univ in @@ -150,5 +161,13 @@ let do_constraint ~poly l = Constraints.add cst acc) Constraints.empty l in - let uctx = ContextSet.add_constraints constraints ContextSet.empty in - DeclareUctx.declare_universe_context ~poly uctx + match poly with + | false -> + let uctx = ContextSet.add_constraints constraints ContextSet.empty in + Global.push_context_set ~strict:true uctx + | true -> + let uctx = UVars.UContext.make + ([||],[||]) + (UVars.Instance.empty,constraints) + in + Global.push_section_context uctx diff --git a/vernac/g_vernac.mlg b/vernac/g_vernac.mlg index dce3186501789..f117ddce2e291 100644 --- a/vernac/g_vernac.mlg +++ b/vernac/g_vernac.mlg @@ -207,6 +207,16 @@ let test_variance_ident = lk_kws ["=";"+";"*"] >> lk_ident end +let test_univ_decl = + let open Pcoq.Lookahead in + to_entry "test_univ_decl" (lk_ident_list >> lk_kw "|" >> lk_ident_list >> (lk_kw "+" <+> lk_empty) >> (lk_kw "|" <+> lk_kw "|}")) + +let test_cumul_univ_decl = + let open Pcoq.Lookahead in + let lk_list_variance_ident = lk_list (lk_kw "+" <+> lk_kw "*" <+> lk_kw "=" <+> lk_ident) in + to_entry "test_cumul_univ_decl" (lk_ident_list >> lk_kw "|" >> lk_list_variance_ident >> (lk_kw "+" <+> lk_empty) >> (lk_kw "|" <+> lk_kw "|}")) + + } (* Gallina declarations *) @@ -302,14 +312,27 @@ GRAMMAR EXTEND Gram [ [ l = universe_name; ord = [ "<" -> { Univ.Lt } | "=" -> { Univ.Eq } | "<=" -> { Univ.Le } ]; r = universe_name -> { (l, ord, r) } ] ] ; + univ_decl_constraints: + [ [ "|"; l' = LIST0 univ_constraint SEP ","; + ext = [ "+" -> { true } | -> { false } ]; "}" -> { (l',ext) } + | ext = [ "}" -> { true } | bar_cbrace -> { false } ] -> { ([], ext) } ] ] + ; univ_decl: - [ [ "@{" ; l = LIST0 identref; ext = [ "+" -> { true } | -> { false } ]; - cs = [ "|"; l' = LIST0 univ_constraint SEP ","; - ext = [ "+" -> { true } | -> { false } ]; "}" -> { (l',ext) } - | ext = [ "}" -> { true } | bar_cbrace -> { false } ] -> { ([], ext) } ] + [ [ "@{" ; test_univ_decl; l0 = LIST0 identref; "|"; l = LIST0 identref; ext = [ "+" -> { true } | -> { false } ]; + cs = univ_decl_constraints -> { let open UState in - { univdecl_instance = l; + { univdecl_qualities = l0; + univdecl_instance = l; + univdecl_extensible_instance = ext; + univdecl_constraints = fst cs; + univdecl_extensible_constraints = snd cs } } + | "@{" ; l = LIST0 identref; ext = [ "+" -> { true } | -> { false } ]; + cs = univ_decl_constraints + -> + { let open UState in + { univdecl_qualities = []; (* TODO *) + univdecl_instance = l; univdecl_extensible_instance = ext; univdecl_constraints = fst cs; univdecl_extensible_constraints = snd cs } } @@ -329,13 +352,23 @@ GRAMMAR EXTEND Gram ] ] ; cumul_univ_decl: - [ [ "@{" ; l = LIST0 variance_identref; ext = [ "+" -> { true } | -> { false } ]; - cs = [ "|"; l' = LIST0 univ_constraint SEP ","; - ext = [ "+" -> { true } | -> { false } ]; "}" -> { (l',ext) } - | ext = [ "}" -> { true } | bar_cbrace -> { false } ] -> { ([], ext) } ] + [ [ "@{" ; test_cumul_univ_decl; l0 = LIST0 identref; "|"; + l = LIST0 variance_identref; ext = [ "+" -> { true } | -> { false } ]; + cs = univ_decl_constraints + -> + { let open UState in + { univdecl_qualities = l0; + univdecl_instance = l; + univdecl_extensible_instance = ext; + univdecl_constraints = fst cs; + univdecl_extensible_constraints = snd cs } } + + | "@{" ; l = LIST0 variance_identref; ext = [ "+" -> { true } | -> { false } ]; + cs = univ_decl_constraints -> { let open UState in - { univdecl_instance = l; + { univdecl_qualities = []; + univdecl_instance = l; univdecl_extensible_instance = ext; univdecl_constraints = fst cs; univdecl_extensible_constraints = snd cs } } @@ -1152,7 +1185,7 @@ GRAMMAR EXTEND Gram ] ] ; univ_name_list: - [ [ "@{" ; l = LIST0 name; "}" -> { l } ] ] + [ [ "@{" ; l = LIST0 name; "}" -> { [],l } ] ] ; END diff --git a/vernac/himsg.ml b/vernac/himsg.ml index 369018a32cabd..ff2e54175316e 100644 --- a/vernac/himsg.ml +++ b/vernac/himsg.ml @@ -771,11 +771,22 @@ let explain_unsatisfied_constraints env sigma cst = Univ.Constraints.pr (Termops.pr_evd_level sigma) cst ++ spc () ++ str "(maybe a bugged tactic)." +let explain_unsatisfied_qconstraints env sigma cst = + strbrk "Unsatisfied quality constraints: " ++ + Sorts.QConstraints.pr (Termops.pr_evd_qvar sigma) cst ++ + spc() ++ str "(maybe a bugged tactic)." + let explain_undeclared_universe env sigma l = strbrk "Undeclared universe: " ++ Termops.pr_evd_level sigma l ++ spc () ++ str "(maybe a bugged tactic)." +let explain_undeclared_qualities env sigma l = + let n = Sorts.QVar.Set.cardinal l in + strbrk "Undeclared " ++ str (if n = 1 then "quality" else "qualities") ++ strbrk": " ++ + prlist_with_sep spc (Termops.pr_evd_qvar sigma) (Sorts.QVar.Set.elements l) ++ + spc () ++ str "(maybe a bugged tactic)." + let explain_disallowed_sprop () = Pp.(strbrk "SProp is disallowed because the " ++ str "\"Allow StrictProp\"" @@ -873,9 +884,13 @@ let explain_type_error env sigma err = | WrongCaseInfo (ind,ci) -> explain_wrong_case_info env ind ci | UnsatisfiedConstraints cst -> - explain_unsatisfied_constraints env sigma cst + explain_unsatisfied_constraints env sigma cst + | UnsatisfiedQConstraints cst -> + explain_unsatisfied_qconstraints env sigma cst | UndeclaredUniverse l -> - explain_undeclared_universe env sigma l + explain_undeclared_universe env sigma l + | UndeclaredQualities l -> + explain_undeclared_qualities env sigma l | DisallowedSProp -> explain_disallowed_sprop () | BadBinderRelevance (rlv, decl) -> explain_bad_binder_relevance env sigma rlv decl | BadCaseRelevance (rlv, case) -> explain_bad_case_relevance env sigma rlv case @@ -1067,7 +1082,7 @@ let explain_not_match_error = function status (not b) ++ str" declaration was found" | IncompatibleUniverses incon -> str"the universe constraints are inconsistent: " ++ - UGraph.explain_universe_inconsistency UnivNames.pr_with_global_universes incon + UGraph.explain_universe_inconsistency UnivNames.pr_level_with_global_universes incon | IncompatiblePolymorphism (env, t1, t2) -> str "conversion of polymorphic values generates additional constraints: " ++ quote (Printer.safe_pr_lconstr_env env (Evd.from_env env) t1) ++ spc () ++ @@ -1087,7 +1102,7 @@ let explain_not_match_error = function in str "incompatible polymorphic binders: got" ++ spc () ++ h (pr_auctx got) ++ spc() ++ str "but expected" ++ spc() ++ h (pr_auctx expect) ++ - (if not (Int.equal (AbstractContext.size got) (AbstractContext.size expect)) then mt() else + (if not (UVars.eq_sizes (AbstractContext.size got) (AbstractContext.size expect)) then mt() else fnl() ++ str "(incompatible constraints)") | IncompatibleVariance -> str "incompatible variance information" @@ -1528,7 +1543,7 @@ let _ = CErrors.register_handler (wrap_unhandled explain_exn_default) let rec vernac_interp_error_handler = function | UGraph.UniverseInconsistency i -> str "Universe inconsistency." ++ spc() ++ - UGraph.explain_universe_inconsistency UnivNames.pr_with_global_universes i ++ str "." + UGraph.explain_universe_inconsistency UnivNames.pr_level_with_global_universes i ++ str "." | TypeError(env,te) -> let te = map_ptype_error EConstr.of_constr te in explain_type_error env (Evd.from_env env) te diff --git a/vernac/ppvernac.ml b/vernac/ppvernac.ml index f3425d89793e7..0a04d0f14e79e 100644 --- a/vernac/ppvernac.ml +++ b/vernac/ppvernac.ml @@ -63,10 +63,12 @@ let pr_uconstraint (l, d, r) = pr_sort_name_expr l ++ spc () ++ Univ.pr_constraint_type d ++ spc () ++ pr_sort_name_expr r -let pr_univ_name_list = function - | None -> mt () - | Some l -> - str "@{" ++ prlist_with_sep spc pr_lname l ++ str"}" +let pr_full_univ_name_list = function + | None -> mt() + | Some (ql, ul) -> + str "@{" ++ prlist_with_sep spc pr_lname ql ++ + (if List.is_empty ql then mt() else strbrk " | ") ++ + prlist_with_sep spc pr_lname ul ++ str "}" let pr_variance_lident (lid,v) = let v = Option.cata UVars.Variance.pr (mt()) v in @@ -82,7 +84,7 @@ let pr_cumul_univdecl_instance l extensible = let pr_univdecl_constraints l extensible = if List.is_empty l && extensible then mt () - else str"|" ++ spc () ++ prlist_with_sep (fun () -> str",") pr_uconstraint l ++ + else str"|" ++ spc () ++ prlist_with_sep (fun () -> strbrk" | ") pr_uconstraint l ++ (if extensible then str"+" else mt()) let pr_universe_decl l = @@ -622,7 +624,7 @@ let pr_printable = function let pr_subgraph = prlist_with_sep spc pr_qualid in keyword cmd ++ pr_opt pr_subgraph g ++ pr_opt str fopt | PrintName (qid,udecl) -> - keyword "Print" ++ spc() ++ pr_smart_global qid ++ pr_univ_name_list udecl + keyword "Print" ++ spc() ++ pr_smart_global qid ++ pr_full_univ_name_list udecl | PrintModuleType qid -> keyword "Print Module Type" ++ spc() ++ pr_qualid qid | PrintModule qid -> @@ -637,7 +639,7 @@ let pr_printable = function keyword "Print Visibility" ++ pr_opt str s | PrintAbout (qid,l,gopt) -> pr_opt (fun g -> Goal_select.pr_goal_selector g ++ str ":"++ spc()) gopt - ++ keyword "About" ++ spc() ++ pr_smart_global qid ++ pr_univ_name_list l + ++ keyword "About" ++ spc() ++ pr_smart_global qid ++ pr_full_univ_name_list l | PrintImplicit qid -> keyword "Print Implicit" ++ spc() ++ pr_smart_global qid (* spiwack: command printing all the axioms and section variables used in a diff --git a/vernac/prettyp.ml b/vernac/prettyp.ml index ebba855f6904b..806e0878583ab 100644 --- a/vernac/prettyp.ml +++ b/vernac/prettyp.ml @@ -193,7 +193,7 @@ let print_if_is_coercion ref = let pr_template_variables = function | [] -> mt () - | vars -> str " on " ++ prlist_with_sep spc UnivNames.pr_with_global_universes vars + | vars -> str " on " ++ prlist_with_sep spc UnivNames.pr_level_with_global_universes vars let print_polymorphism ref = let poly = Global.is_polymorphic ref in diff --git a/vernac/prettyp.mli b/vernac/prettyp.mli index 23fe9589394e1..9b9720b17a78a 100644 --- a/vernac/prettyp.mli +++ b/vernac/prettyp.mli @@ -41,7 +41,7 @@ val print_safe_judgment : Safe_typing.judgment -> Pp.t val print_name : env -> Evd.evar_map -> qualid Constrexpr.or_by_notation - -> UnivNames.univ_name_list option + -> UnivNames.full_name_list option -> Pp.t val print_notation : env -> Evd.evar_map -> Constrexpr.notation_entry @@ -51,7 +51,7 @@ val print_notation : env -> Evd.evar_map val print_abbreviation : env -> KerName.t -> Pp.t val print_about : env -> Evd.evar_map -> qualid Constrexpr.or_by_notation -> - UnivNames.univ_name_list option -> Pp.t + UnivNames.full_name_list option -> Pp.t val print_impargs : GlobRef.t -> Pp.t (** Pretty-printing functions for classes and coercions *) diff --git a/vernac/printmod.mli b/vernac/printmod.mli index e3a5cea782c70..a34ed5f901b9a 100644 --- a/vernac/printmod.mli +++ b/vernac/printmod.mli @@ -12,7 +12,6 @@ open Names val pr_mutual_inductive_body : Environ.env -> MutInd.t -> Declarations.mutual_inductive_body -> - UnivNames.univ_name_list option -> Pp.t - + UnivNames.full_name_list option -> Pp.t val print_module : with_body:bool -> ModPath.t -> Pp.t val print_modtype : ModPath.t -> Pp.t diff --git a/vernac/record.ml b/vernac/record.ml index 2509012fac25d..81d6d64b08267 100644 --- a/vernac/record.ml +++ b/vernac/record.ml @@ -796,7 +796,7 @@ let interp_structure udecl kind ~template ~cumulative ~poly ~primitive_proj fini interp_structure_core ~cumulative finite ~univs ~variances ~primitive_proj impargs params template ~projections_kind ~indlocs data let declare_structure { Record_decl.mie; primitive_proj; impls; globnames; global_univ_decls; projunivs; ubinders; projections_kind; poly; records; indlocs } = - Option.iter (DeclareUctx.declare_universe_context ~poly:false) global_univ_decls; + Option.iter (Global.push_context_set ~strict:true) global_univ_decls; let kn = DeclareInd.declare_mutual_inductive_with_eliminations mie globnames impls ~primitive_expected:primitive_proj ~indlocs in diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml index c0efcfe28f484..4d60b300051c9 100644 --- a/vernac/vernacentries.ml +++ b/vernac/vernacentries.ml @@ -443,7 +443,7 @@ let print_universes ?loc ~sort ~subgraph dst = if Global.is_joined_environment () then mt () else str"There may remain asynchronous universe constraints" in - let prl = UnivNames.pr_with_global_universes in + let prl = UnivNames.pr_level_with_global_universes in begin match dst with | None -> UGraph.pr_universes prl univ ++ pr_remaining | Some s -> dump_universes_gen (fun u -> Pp.string_of_ppcmds (prl u)) univ s diff --git a/vernac/vernacexpr.mli b/vernac/vernacexpr.mli index c2f8657face14..7e59d0ba4b14e 100644 --- a/vernac/vernacexpr.mli +++ b/vernac/vernacexpr.mli @@ -40,7 +40,7 @@ type printable = | PrintMLLoadPath | PrintMLModules | PrintDebugGC - | PrintName of qualid or_by_notation * UnivNames.univ_name_list option + | PrintName of qualid or_by_notation * UnivNames.full_name_list option | PrintGraph | PrintClasses | PrintTypeclasses @@ -56,7 +56,7 @@ type printable = | PrintScopes | PrintScope of string | PrintVisibility of string option - | PrintAbout of qualid or_by_notation * UnivNames.univ_name_list option * Goal_select.t option + | PrintAbout of qualid or_by_notation * UnivNames.full_name_list option * Goal_select.t option | PrintImplicit of qualid or_by_notation | PrintAssumptions of bool * bool * qualid or_by_notation | PrintStrategy of qualid or_by_notation option