diff options
| author | Gaëtan Gilbert | 2019-06-17 22:57:33 +0200 |
|---|---|---|
| committer | Gaëtan Gilbert | 2019-06-17 22:57:33 +0200 |
| commit | 459fc67f4a40ff82a7694f9afafb3ac303d74554 (patch) | |
| tree | 9c1607700b3689c36de0daf0427f5e95aeb5b55c /kernel/opaqueproof.ml | |
| parent | d886dff0857702fc4524779980ee6b7e9688c1d4 (diff) | |
| parent | 621201858125632496fd11f431529187d69cfdc6 (diff) | |
Merge PR #10362: Kernel-side delaying of polymorphic opaque constants
Reviewed-by: SkySkimmer
Reviewed-by: gares
Diffstat (limited to 'kernel/opaqueproof.ml')
| -rw-r--r-- | kernel/opaqueproof.ml | 92 |
1 files changed, 57 insertions, 35 deletions
diff --git a/kernel/opaqueproof.ml b/kernel/opaqueproof.ml index 335174f054..e256466112 100644 --- a/kernel/opaqueproof.ml +++ b/kernel/opaqueproof.ml @@ -20,18 +20,28 @@ type cooking_info = { modlist : work_list; abstract : Constr.named_context * Univ.Instance.t * Univ.AUContext.t } +type 'a delayed_universes = +| PrivateMonomorphic of 'a +| PrivatePolymorphic of int * Univ.ContextSet.t + +type opaque_proofterm = cooking_info list * (Constr.t * unit delayed_universes) option + type indirect_accessor = { - access_proof : DirPath.t -> int -> constr option; - access_discharge : cooking_info list -> int -> constr -> constr; + access_proof : DirPath.t -> int -> opaque_proofterm; + access_discharge : cooking_info list -> (Constr.t * unit delayed_universes) -> (Constr.t * unit delayed_universes); } -type proofterm = (constr * Univ.ContextSet.t) Future.computation -type universes = int +let drop_mono = function +| PrivateMonomorphic _ -> PrivateMonomorphic () +| PrivatePolymorphic _ as ctx -> ctx + +type proofterm = (constr * Univ.ContextSet.t delayed_universes) Future.computation + type opaque = | Indirect of substitution list * DirPath.t * int (* subst, lib, index *) - | Direct of universes * cooking_info list * proofterm + | Direct of cooking_info list * proofterm type opaquetab = { - opaque_val : (int * cooking_info list * proofterm) Int.Map.t; + opaque_val : (cooking_info list * proofterm) Int.Map.t; (** Actual proof terms *) opaque_len : int; (** Size of the above map *) @@ -46,25 +56,28 @@ let empty_opaquetab = { let not_here () = CErrors.user_err Pp.(str "Cannot access opaque delayed proof") -let create ~univs cu = Direct (univs, [],cu) +let create cu = Direct ([],cu) let turn_indirect dp o tab = match o with | Indirect (_,_,i) -> if not (Int.Map.mem i tab.opaque_val) then CErrors.anomaly (Pp.str "Indirect in a different table.") else CErrors.anomaly (Pp.str "Already an indirect opaque.") - | Direct (nunivs, d, cu) -> + | Direct (d, cu) -> (* Invariant: direct opaques only exist inside sections, we turn them indirect as soon as we are at toplevel. At this moment, we perform hashconsing of their contents, potentially as a future. *) let hcons (c, u) = let c = Constr.hcons c in - let u = Univ.hcons_universe_context_set u in + let u = match u with + | PrivateMonomorphic u -> PrivateMonomorphic (Univ.hcons_universe_context_set u) + | PrivatePolymorphic (n, u) -> PrivatePolymorphic (n, Univ.hcons_universe_context_set u) + in (c, u) in let cu = Future.chain cu hcons in let id = tab.opaque_len in - let opaque_val = Int.Map.add id (nunivs, d,cu) tab.opaque_val in + let opaque_val = Int.Map.add id (d,cu) tab.opaque_val in let opaque_dir = if DirPath.equal dp tab.opaque_dir then tab.opaque_dir else if DirPath.equal tab.opaque_dir DirPath.initial then dp @@ -79,8 +92,8 @@ let subst_opaque sub = function let discharge_direct_opaque ci = function | Indirect _ -> CErrors.anomaly (Pp.str "Not a direct opaque.") - | Direct (n, d, cu) -> - Direct (n, ci :: d, cu) + | Direct (d, cu) -> + Direct (ci :: d, cu) let join except cu = match except with | None -> ignore (Future.join cu) @@ -89,61 +102,70 @@ let join except cu = match except with else ignore (Future.join cu) let join_opaque ?except { opaque_val = prfs; opaque_dir = odp; _ } = function - | Direct (_,_,cu) -> join except cu + | Direct (_,cu) -> join except cu | Indirect (_,dp,i) -> if DirPath.equal dp odp then - let (_, _, fp) = Int.Map.find i prfs in + let (_, fp) = Int.Map.find i prfs in join except fp let force_proof access { opaque_val = prfs; opaque_dir = odp; _ } = function - | Direct (n, d, cu) -> - let (c, _) = Future.force cu in - access.access_discharge d n c + | Direct (d, cu) -> + let (c, u) = Future.force cu in + access.access_discharge d (c, drop_mono u) | Indirect (l,dp,i) -> - let c = + let c, u = if DirPath.equal dp odp then - let (n, d, cu) = Int.Map.find i prfs in - let (c, _) = Future.force cu in - access.access_discharge d n c - else match access.access_proof dp i with - | None -> not_here () - | Some v -> v + let (d, cu) = Int.Map.find i prfs in + let (c, u) = Future.force cu in + access.access_discharge d (c, drop_mono u) + else + let (d, cu) = access.access_proof dp i in + match cu with + | None -> not_here () + | Some (c, u) -> access.access_discharge d (c, u) in - force_constr (List.fold_right subst_substituted l (from_val c)) + let c = force_constr (List.fold_right subst_substituted l (from_val c)) in + (c, u) + +let get_mono (_, u) = match u with +| PrivateMonomorphic ctx -> ctx +| PrivatePolymorphic _ -> Univ.ContextSet.empty let force_constraints _access { opaque_val = prfs; opaque_dir = odp; _ } = function - | Direct (_,_,cu) -> - snd(Future.force cu) + | Direct (_,cu) -> + get_mono (Future.force cu) | Indirect (_,dp,i) -> if DirPath.equal dp odp then - let (_, _, cu) = Int.Map.find i prfs in - snd (Future.force cu) + let ( _, cu) = Int.Map.find i prfs in + get_mono (Future.force cu) else Univ.ContextSet.empty let get_direct_constraints = function | Indirect _ -> CErrors.anomaly (Pp.str "Not a direct opaque.") -| Direct (_, _, cu) -> Future.chain cu snd +| Direct (_, cu) -> + Future.chain cu get_mono module FMap = Future.UUIDMap let dump ?(except = Future.UUIDSet.empty) { opaque_val = otab; opaque_len = n; _ } = - let opaque_table = Array.make n ([], 0, None) in + let opaque_table = Array.make n ([], None) in let f2t_map = ref FMap.empty in - let iter n (univs, d, cu) = + let iter n (d, cu) = let uid = Future.uuid cu in let () = f2t_map := FMap.add (Future.uuid cu) n !f2t_map in let c = if Future.is_val cu then - let (c, _) = Future.force cu in - Some c + let (c, priv) = Future.force cu in + let priv = drop_mono priv in + Some (c, priv) else if Future.UUIDSet.mem uid except then None else CErrors.anomaly Pp.(str"Proof object "++int n++str" is not checked nor to be checked") in - opaque_table.(n) <- (d, univs, c) + opaque_table.(n) <- (d, c) in let () = Int.Map.iter iter otab in opaque_table, !f2t_map |
