Skip to content

Commit d780de7

Browse files
committed
Factorization of close_proof and close_future_proof via some "process_proof".
1 parent 331f1d0 commit d780de7

File tree

1 file changed

+98
-82
lines changed

1 file changed

+98
-82
lines changed

vernac/declare.ml

Lines changed: 98 additions & 82 deletions
Original file line numberDiff line numberDiff line change
@@ -811,6 +811,50 @@ module Internal = struct
811811

812812
end
813813

814+
(* The word [proof] is to be understood as [justification] *)
815+
(* A possible alternatve would be [evidence]?? *)
816+
type closed_proof_output = ((Constr.t * Evd.side_effects) * Constr.t option) list * UState.t
817+
818+
type proof_object =
819+
| DefaultProof of
820+
{ proof : closed_proof_output
821+
; opaque : bool
822+
; using : Names.Id.Set.t option
823+
; keep_body_ucst_separate : UState.t option
824+
}
825+
| DeferredOpaqueProof of
826+
{ deferred_proof : closed_proof_output Future.computation
827+
; using : Names.Id.Set.t option
828+
; initial_proof_data : Proof.data
829+
; feedback_id : Stateid.t
830+
; initial_euctx : UState.t
831+
}
832+
833+
let future_map2_pair_list_distribute p l f =
834+
List.map_i (fun i c -> f (Future.chain p (fun (a, b) -> (List.nth a i, b))) c) 0 l
835+
836+
let process_proof ~info:Info.({ udecl; poly }) = function
837+
| DefaultProof { proof = (entries, uctx); opaque; using; keep_body_ucst_separate } ->
838+
let entries = List.map (fun ((body, eff), typ) ->
839+
let uctx, univs, body = make_univs_immediate ~poly ?keep_body_ucst_separate ~opaque ~uctx ~udecl ~eff body typ in
840+
definition_entry_core ?using ~univs ?types:typ body) entries in
841+
entries, uctx
842+
| DeferredOpaqueProof { deferred_proof = bodies; using; initial_proof_data; feedback_id; initial_euctx } ->
843+
let { Proof.poly; entry; sigma } = initial_proof_data in
844+
let entries =
845+
future_map2_pair_list_distribute bodies (Proofview.initial_goals entry)
846+
(fun body_typ_uctx (_, _, initial_typ) ->
847+
(* Testing if evar-closed? *)
848+
let initial_typ = Evarutil.nf_evars_universes sigma (EConstr.Unsafe.to_constr initial_typ) in
849+
(* The flags keep_body_ucst_separate, opaque, etc. should be consistent with evar-closedness? *)
850+
let univs = UState.univ_entry ~poly:false initial_euctx in
851+
let body = Future.chain body_typ_uctx (fun (((body, eff), typ), uctx) ->
852+
let uctx = make_univs_deferred_private_mono ~initial_euctx ~uctx ~udecl body typ in
853+
((body, uctx), eff)) in
854+
delayed_definition_entry ?using ~univs ~types:initial_typ ~feedback_id body)
855+
in
856+
entries, initial_euctx
857+
814858
let declare_definition_scheme ~internal ~univs ~role ~name ?loc c =
815859
let kind = Decls.(IsDefinition Scheme) in
816860
let entry = pure_definition_entry ~univs c in
@@ -888,7 +932,8 @@ let gather_mutual_using_data =
888932
(* XXX: this should be unified with the code for non-interactive
889933
mutuals previously on this file. *)
890934

891-
let declare_possibly_mutual_definitions ~info ~cinfo ~uctx entries =
935+
let declare_possibly_mutual_definitions ~info ~cinfo proof_obj =
936+
let entries, uctx = process_proof ~info proof_obj in
892937
let { Info.hook; scope; clearbody; kind; typing_flags; user_warns; _ } = info in
893938
let refs = List.map2 (fun CInfo.{name; impargs} ->
894939
declare_entry ~name ~scope ~clearbody ~kind ?hook ~impargs ~typing_flags ~user_warns ~uctx) cinfo entries in
@@ -1648,10 +1693,6 @@ module Proof_ending = struct
16481693

16491694
end
16501695

1651-
(* Alias *)
1652-
module Proof_ = Proof
1653-
module Proof = struct
1654-
16551696
module Proof_info = struct
16561697

16571698
type t =
@@ -1675,6 +1716,22 @@ module Proof_info = struct
16751716

16761717
end
16771718

1719+
module Proof_object = struct
1720+
1721+
type t =
1722+
{ proof_object : proof_object
1723+
; pinfo : Proof_info.t
1724+
}
1725+
1726+
end
1727+
1728+
(* Alias *)
1729+
module Proof_ = Proof
1730+
module Proof = struct
1731+
1732+
type nonrec closed_proof_output = closed_proof_output
1733+
type proof_object = Proof_object.t
1734+
16781735
type t =
16791736
{ endline_tactic : Genarg.glob_generic_argument option
16801737
; using : Id.Set.t option
@@ -1874,12 +1931,6 @@ let get_open_goals ps =
18741931
(List.map (fun (l1,l2) -> List.length l1 + List.length l2) stack) +
18751932
List.length (Evd.shelf sigma)
18761933

1877-
type proof_object =
1878-
{ entries : proof_entry list
1879-
; uctx: UState.t
1880-
; pinfo : Proof_info.t
1881-
}
1882-
18831934
let warn_remaining_shelved_goals =
18841935
CWarnings.create ~name:"remaining-shelved-goals" ~category:CWarnings.CoreCategories.tactics
18851936
(fun () -> Pp.str"The proof has remaining shelved goals.")
@@ -1975,7 +2026,7 @@ let prepare_proof ?(warn_incomplete=true) { proof; pinfo } =
19752026
let rec_declaration = prepare_recursive_declaration pinfo.cinfo fixtypes fixrelevances fixbodies in
19762027
let typing_flags = pinfo.info.typing_flags in
19772028
fst (make_recursive_bodies env ~typing_flags ~possible_guard ~rec_declaration) in
1978-
let proofs = List.map (fun (body, typ) -> ((body, eff), typ)) proofs in
2029+
let proofs = List.map (fun (body, typ) -> ((body, eff), Some typ)) proofs in
19792030
let () = if warn_incomplete then check_incomplete_proof evd in
19802031
proofs, Evd.ustate evd
19812032

@@ -2026,68 +2077,29 @@ let control_only_guard { proof; pinfo } =
20262077
raise (NotGuarded (env, sigma, cofix_error, fix_errors, rec_declaration))
20272078
with Exit -> ()
20282079

2080+
let return_proof p = (prepare_proof p : closed_proof_output)
2081+
20292082
let bool_of_opacity = function
20302083
| Vernacexpr.Opaque -> true
20312084
| Vernacexpr.Transparent -> false
20322085

2033-
let close_proof ?warn_incomplete ~opaque ~keep_body_ucst_separate ps =
2034-
2035-
let { using; proof; initial_euctx; pinfo } = ps in
2036-
let { Proof_info.info = { Info.udecl } } = pinfo in
2037-
let { Proof.poly } = Proof.data proof in
2038-
let elist, uctx = prepare_proof ?warn_incomplete ps in
2086+
let close_proof ?warn_incomplete ~opaque ~keep_body_ucst_separate (proof : t) : Proof_object.t =
20392087
let opaque = bool_of_opacity opaque in
2088+
let keep_body_ucst_separate = if keep_body_ucst_separate then Some proof.initial_euctx else None in
2089+
{ proof_object =
2090+
DefaultProof { proof = prepare_proof ?warn_incomplete proof; opaque; using = proof.using; keep_body_ucst_separate }
2091+
; pinfo = proof.pinfo
2092+
}
20402093

2041-
let make_entry ((body, eff), typ) =
2042-
let keep_body_ucst_separate = if keep_body_ucst_separate then Some initial_euctx else None in
2043-
let _, univs, body =
2044-
make_univs_immediate ~poly ?keep_body_ucst_separate ~opaque ~uctx ~udecl ~eff body (Some typ) in
2045-
definition_entry_core ?using ~univs ~types:typ body
2046-
in
2047-
let entries = CList.map make_entry elist in
2048-
{ entries; uctx; pinfo }
2049-
2050-
type closed_proof_output = ((Constr.t * Evd.side_effects) * Constr.t) list * UState.t
2051-
2052-
let close_proof_delayed ~feedback_id ps (fpl : closed_proof_output Future.computation) =
2053-
let { using; proof; initial_euctx; pinfo } = ps in
2054-
let { Proof_info.info = { Info.udecl } } = pinfo in
2055-
let { Proof.poly; entry; sigma } = Proof.data proof in
2056-
2057-
(* We don't allow poly = true in this path *)
2058-
if poly then
2059-
CErrors.anomaly (Pp.str "Cannot delay universe-polymorphic constants.");
2060-
2061-
(* Because of dependent subgoals at the beginning of proofs, we could
2062-
have existential variables in the initial types of goals, we need to
2063-
normalise them for the kernel. *)
2064-
let nf = Evarutil.nf_evars_universes (Evd.set_universe_context sigma initial_euctx) in
2065-
2066-
(* We only support opaque proofs, this will be enforced by using
2067-
different entries soon *)
2068-
let make_entry i (_, _, types) =
2069-
(* Already checked the univ_decl for the type universes when starting the proof. *)
2070-
let univs = UState.univ_entry ~poly:false initial_euctx in
2071-
let typ = nf (EConstr.Unsafe.to_constr types) in
2072-
2073-
(* NB: for Admitted proofs [fpl] is not valid (raises anomaly when forced) *)
2074-
Future.chain fpl (fun (pf, uctx) ->
2075-
let (body, eff) = fst (List.nth pf i) in
2076-
(* Deferred proof, we already checked the universe declaration with
2077-
the initial universes, ensure that the final universes respect
2078-
the declaration as well. If the declaration is non-extensible,
2079-
this will prevent the body from adding universes and constraints. *)
2080-
let uctx = make_univs_deferred_private_mono ~initial_euctx ~uctx ~udecl body (Some typ) in
2081-
((body, uctx), eff))
2082-
|> delayed_definition_entry ?using ~univs ~types:typ ~feedback_id
2083-
in
2084-
let entries = CList.map_i make_entry 0 (Proofview.initial_goals entry) in
2085-
{ entries; uctx = initial_euctx; pinfo }
2094+
let close_proof_delayed ~feedback_id proof (fpl : closed_proof_output Future.computation) : Proof_object.t =
2095+
{ proof_object =
2096+
DeferredOpaqueProof { deferred_proof = fpl; using = proof.using; initial_proof_data = Proof.data proof.proof;
2097+
feedback_id; initial_euctx = proof.initial_euctx }
2098+
; pinfo = proof.pinfo
2099+
}
20862100

20872101
let close_future_proof = close_proof_delayed
20882102

2089-
let return_proof p = prepare_proof p
2090-
20912103
let update_sigma_univs ugraph p =
20922104
map ~f:(Proof.update_sigma_univs ugraph) p
20932105

@@ -2101,11 +2113,12 @@ let build_constant_by_tactic ~name ?warn_incomplete ~sigma ~sign ~poly (typ : EC
21012113
let pinfo = Proof_info.make ~cinfo ~info () in
21022114
let pf = start_proof_core ~name ~pinfo sigma [Some sign, typ] in
21032115
let pf, status = by tac pf in
2104-
let { entries; uctx } = close_proof ?warn_incomplete ~opaque:Vernacexpr.Transparent ~keep_body_ucst_separate:false pf in
2116+
let proof = close_proof ?warn_incomplete ~keep_body_ucst_separate:false ~opaque:Vernacexpr.Transparent pf in
2117+
let entries, uctx = process_proof ~info proof.proof_object in
21052118
let { Proof.sigma } = Proof.data pf.proof in
2106-
let sigma = Evd.set_universe_context sigma uctx in
21072119
match entries with
21082120
| [ { proof_entry_body = Default { body; opaque = Transparent } } as entry] ->
2121+
let sigma = Evd.set_universe_context sigma uctx in
21092122
{ entry with proof_entry_body = body }, status, sigma
21102123
| _ ->
21112124
CErrors.anomaly Pp.(str "[build_constant_by_tactic] close_proof returned more than one proof term, or a non transparent one.")
@@ -2248,7 +2261,7 @@ let save_admitted ~pm ~proof =
22482261
(* Saving a lemma-like constant *)
22492262
(************************************************************************)
22502263

2251-
let finish_derived ~f ~name {entries; pinfo; uctx} =
2264+
let finish_derived ~f ~name ~uctx ~pinfo entries =
22522265
(* [f] and [name] correspond to the proof of [f] and of [suchthat], respectively. *)
22532266

22542267
let { Proof_info.info = { Info.hook; scope; clearbody; kind; typing_flags; user_warns; poly; udecl; _ } } = pinfo in
@@ -2271,7 +2284,7 @@ let finish_derived ~f ~name {entries; pinfo; uctx} =
22712284
(0, [], [], Univ.Level.Set.empty) pinfo.Proof_info.cinfo entries in
22722285
refs
22732286

2274-
let finish_proved_equations ~pm ~kind ~hook i proof_obj types sigma0 =
2287+
let finish_proved_equations ~pm ~kind ~hook i entries types sigma0 =
22752288

22762289
let obls = ref 1 in
22772290
let sigma, recobls =
@@ -2288,14 +2301,14 @@ let finish_proved_equations ~pm ~kind ~hook i proof_obj types sigma0 =
22882301
let sigma, app = Evd.fresh_global (Global.env ()) sigma (GlobRef.ConstRef cst) in
22892302
let sigma = Evd.define ev (EConstr.applist (app, args)) sigma in
22902303
sigma, cst) sigma0
2291-
types proof_obj.entries
2304+
types entries
22922305
in
22932306
let pm = hook ~pm recobls sigma in
22942307
pm, List.map (fun cst -> GlobRef.ConstRef cst) recobls
22952308

2296-
let check_single_entry { entries; uctx } label =
2309+
let check_single_entry entries label =
22972310
match entries with
2298-
| [entry] -> entry, uctx
2311+
| [entry] -> entry
22992312
| _ ->
23002313
CErrors.anomaly ~label Pp.(str "close_proof returned more than one proof term")
23012314

@@ -2304,16 +2317,18 @@ let finish_proof ~pm proof_obj proof_info =
23042317
let { Proof_info.info; cinfo; possible_guard } = proof_info in
23052318
match CEphemeron.default proof_info.Proof_info.proof_ending Regular with
23062319
| Regular ->
2307-
let {entries; uctx} = proof_obj in
2308-
pm, declare_possibly_mutual_definitions ~uctx ~info ~cinfo entries
2320+
pm, declare_possibly_mutual_definitions ~info ~cinfo proof_obj
23092321
| End_obligation oinfo ->
2310-
let entry, uctx = check_single_entry proof_obj "Obligation.save" in
2322+
let entries, uctx = process_proof ~info proof_obj in
2323+
let entry = check_single_entry entries "Obligation.save" in
23112324
Obls_.obligation_terminator ~pm ~entry ~uctx ~oinfo
23122325
| End_derive { f ; name } ->
2313-
pm, finish_derived ~f ~name proof_obj
2326+
let entries, uctx = process_proof ~info proof_obj in
2327+
pm, finish_derived ~f ~name ~uctx ~pinfo:proof_info entries
23142328
| End_equations { hook; i; types; sigma } ->
23152329
let kind = info.Info.kind in
2316-
finish_proved_equations ~pm ~kind ~hook i proof_obj types sigma
2330+
let entries, uctx = process_proof ~info proof_obj in
2331+
finish_proved_equations ~pm ~kind ~hook i entries types sigma
23172332

23182333
let err_save_forbidden_in_place_of_qed () =
23192334
CErrors.user_err (Pp.str "Cannot use Save with more than one constant or in this proof mode")
@@ -2335,7 +2350,7 @@ let save ~pm ~proof ~opaque ~idopt =
23352350
(* Env and sigma are just used for error printing in save_remaining_recthms *)
23362351
let proof_obj = close_proof ~opaque ~keep_body_ucst_separate:false proof in
23372352
let proof_info = process_idopt_for_save ~idopt proof.pinfo in
2338-
finish_proof ~pm proof_obj proof_info
2353+
finish_proof ~pm proof_obj.proof_object proof_info
23392354

23402355
let save_regular ~(proof : t) ~opaque ~idopt =
23412356
let open Proof_ending in
@@ -2349,20 +2364,21 @@ let save_regular ~(proof : t) ~opaque ~idopt =
23492364
(* Special case to close a lemma without forcing a proof *)
23502365
(***********************************************************************)
23512366
let save_lemma_admitted_delayed ~pm ~proof =
2352-
let { entries; uctx; pinfo } = proof in
2367+
let { Proof_object.proof_object; pinfo } = proof in
2368+
let entries, uctx = process_proof ~info:pinfo.info proof_object in
23532369
let typs = List.map (function { proof_entry_type } -> Option.get proof_entry_type) entries in
23542370
(* Note: an alternative would be to compute sec_vars of the partial
23552371
proof as a Future computation, as in compute_proof_using_for_admitted *)
23562372
let sec_vars = if get_keep_admitted_vars () then (List.hd entries).proof_entry_secctx else None in
23572373
(* If the proof is partial, do we want to take the (restriction on
23582374
visible uvars of) uctx so far or (as done below) the initial ones
23592375
that refers to only the types *)
2360-
finish_admitted ~pm ~uctx ~pinfo ~sec_vars typs
2376+
finish_admitted ~pm ~uctx ~pinfo:proof.pinfo ~sec_vars typs
23612377

23622378
let save_lemma_proved_delayed ~pm ~proof ~idopt =
23632379
(* vio2vo used to call this with invalid [pinfo], now it should work fine. *)
2364-
let pinfo = process_idopt_for_save ~idopt proof.pinfo in
2365-
let pm, _ = finish_proof ~pm proof pinfo in
2380+
let pinfo = process_idopt_for_save ~idopt proof.Proof_object.pinfo in
2381+
let pm, _ = finish_proof ~pm proof.proof_object pinfo in
23662382
pm
23672383

23682384
end (* Proof module *)

0 commit comments

Comments
 (0)