Skip to content

Commit ccb07f8

Browse files
committed
Factorization of close_proof and close_future_proof via some "process_proof".
1 parent 70f55d3 commit ccb07f8

File tree

2 files changed

+99
-83
lines changed

2 files changed

+99
-83
lines changed

vernac/declare.ml

Lines changed: 97 additions & 81 deletions
Original file line numberDiff line numberDiff line change
@@ -807,6 +807,50 @@ module Internal = struct
807807

808808
end
809809

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

887-
let declare_possibly_mutual_definitions ~info ~cinfo ~uctx entries =
931+
let declare_possibly_mutual_definitions ~info ~cinfo proof_obj =
932+
let entries, uctx = process_proof ~info proof_obj in
888933
let { Info.hook; scope; clearbody; kind; typing_flags; user_warns; _ } = info in
889934
let refs = List.map2 (fun CInfo.{name; impargs} ->
890935
declare_entry ~name ~scope ~clearbody ~kind ?hook ~impargs ~typing_flags ~user_warns ~uctx) cinfo entries in
@@ -1634,10 +1679,6 @@ module Proof_ending = struct
16341679

16351680
end
16361681

1637-
(* Alias *)
1638-
module Proof_ = Proof
1639-
module Proof = struct
1640-
16411682
module Proof_info = struct
16421683

16431684
type t =
@@ -1661,6 +1702,22 @@ module Proof_info = struct
16611702

16621703
end
16631704

1705+
module Proof_object = struct
1706+
1707+
type t =
1708+
{ proof_object : proof_object
1709+
; pinfo : Proof_info.t
1710+
}
1711+
1712+
end
1713+
1714+
(* Alias *)
1715+
module Proof_ = Proof
1716+
module Proof = struct
1717+
1718+
type nonrec closed_proof_output = closed_proof_output
1719+
type proof_object = Proof_object.t
1720+
16641721
type t =
16651722
{ endline_tactic : Genarg.glob_generic_argument option
16661723
; using : Id.Set.t option
@@ -1861,12 +1918,6 @@ let get_open_goals ps =
18611918
(List.map (fun (l1,l2) -> List.length l1 + List.length l2) stack) +
18621919
List.length (Evd.shelf sigma)
18631920

1864-
type proof_object =
1865-
{ entries : proof_entry list
1866-
; uctx: UState.t
1867-
; pinfo : Proof_info.t
1868-
}
1869-
18701921
let warn_remaining_shelved_goals =
18711922
CWarnings.create ~name:"remaining-shelved-goals" ~category:CWarnings.CoreCategories.tactics
18721923
(fun () -> Pp.str"The proof has remaining shelved goals.")
@@ -1962,72 +2013,33 @@ let prepare_proof ?(warn_incomplete=true) { proof; pinfo } =
19622013
let rec_declaration = prepare_recursive_declaration pinfo.cinfo fixtypes fixrelevances fixbodies in
19632014
let typing_flags = pinfo.info.typing_flags in
19642015
fst (make_recursive_bodies env ~typing_flags ~possible_guard ~rec_declaration) in
1965-
let proofs = List.map (fun (body, typ) -> ((body, eff), typ)) proofs in
2016+
let proofs = List.map (fun (body, typ) -> ((body, eff), Some typ)) proofs in
19662017
let () = if warn_incomplete then check_incomplete_proof evd in
19672018
proofs, Evd.evar_universe_context evd
19682019

2020+
let return_proof p = (prepare_proof p : closed_proof_output)
2021+
19692022
let bool_of_opacity = function
19702023
| Vernacexpr.Opaque -> true
19712024
| Vernacexpr.Transparent -> false
19722025

1973-
let close_proof ?warn_incomplete ~opaque ~keep_body_ucst_separate ps =
1974-
1975-
let { using; proof; initial_euctx; pinfo } = ps in
1976-
let { Proof_info.info = { Info.udecl } } = pinfo in
1977-
let { Proof.poly } = Proof.data proof in
1978-
let elist, uctx = prepare_proof ?warn_incomplete ps in
2026+
let close_proof ?warn_incomplete ~opaque ~keep_body_ucst_separate (proof : t) : Proof_object.t =
19792027
let opaque = bool_of_opacity opaque in
2028+
let keep_body_ucst_separate = if keep_body_ucst_separate then Some proof.initial_euctx else None in
2029+
{ proof_object =
2030+
DefaultProof { proof = prepare_proof ?warn_incomplete proof; opaque; using = proof.using; keep_body_ucst_separate }
2031+
; pinfo = proof.pinfo
2032+
}
19802033

1981-
let make_entry ((body, eff), typ) =
1982-
let keep_body_ucst_separate = if keep_body_ucst_separate then Some initial_euctx else None in
1983-
let univs, body =
1984-
make_univs_immediate ~poly ?keep_body_ucst_separate ~opaque ~uctx ~udecl ~eff body (Some typ) in
1985-
definition_entry_core ?using ~univs ~types:typ body
1986-
in
1987-
let entries = CList.map make_entry elist in
1988-
{ entries; uctx; pinfo }
1989-
1990-
type closed_proof_output = ((Constr.t * Evd.side_effects) * Constr.t) list * UState.t
1991-
1992-
let close_proof_delayed ~feedback_id ps (fpl : closed_proof_output Future.computation) =
1993-
let { using; proof; initial_euctx; pinfo } = ps in
1994-
let { Proof_info.info = { Info.udecl } } = pinfo in
1995-
let { Proof.poly; entry; sigma } = Proof.data proof in
1996-
1997-
(* We don't allow poly = true in this path *)
1998-
if poly then
1999-
CErrors.anomaly (Pp.str "Cannot delay universe-polymorphic constants.");
2000-
2001-
(* Because of dependent subgoals at the beginning of proofs, we could
2002-
have existential variables in the initial types of goals, we need to
2003-
normalise them for the kernel. *)
2004-
let nf = Evarutil.nf_evars_universes (Evd.set_universe_context sigma initial_euctx) in
2005-
2006-
(* We only support opaque proofs, this will be enforced by using
2007-
different entries soon *)
2008-
let make_entry i (_, _, types) =
2009-
(* Already checked the univ_decl for the type universes when starting the proof. *)
2010-
let univs = UState.univ_entry ~poly:false initial_euctx in
2011-
let typ = nf (EConstr.Unsafe.to_constr types) in
2012-
2013-
(* NB: for Admitted proofs [fpl] is not valid (raises anomaly when forced) *)
2014-
Future.chain fpl (fun (pf, uctx) ->
2015-
let (body, eff) = fst (List.nth pf i) in
2016-
(* Deferred proof, we already checked the universe declaration with
2017-
the initial universes, ensure that the final universes respect
2018-
the declaration as well. If the declaration is non-extensible,
2019-
this will prevent the body from adding universes and constraints. *)
2020-
let uctx = make_univs_deferred_private_mono ~initial_euctx ~uctx ~udecl body (Some typ) in
2021-
((body, uctx), eff))
2022-
|> delayed_definition_entry ?using ~univs ~types:typ ~feedback_id
2023-
in
2024-
let entries = CList.map_i make_entry 0 (Proofview.initial_goals entry) in
2025-
{ entries; uctx = initial_euctx; pinfo }
2034+
let close_proof_delayed ~feedback_id proof (fpl : closed_proof_output Future.computation) : Proof_object.t =
2035+
{ proof_object =
2036+
DeferredOpaqueProof { deferred_proof = fpl; using = proof.using; initial_proof_data = Proof.data proof.proof;
2037+
feedback_id; initial_euctx = proof.initial_euctx }
2038+
; pinfo = proof.pinfo
2039+
}
20262040

20272041
let close_future_proof = close_proof_delayed
20282042

2029-
let return_proof p = prepare_proof p
2030-
20312043
let update_sigma_univs ugraph p =
20322044
map ~f:(Proof.update_sigma_univs ugraph) p
20332045

@@ -2041,11 +2053,12 @@ let build_constant_by_tactic ~name ?warn_incomplete ~sigma ~sign ~poly (typ : EC
20412053
let pinfo = Proof_info.make ~cinfo ~info () in
20422054
let pf = start_proof_core ~name ~pinfo sigma [Some sign, typ] in
20432055
let pf, status = by tac pf in
2044-
let { entries; uctx } = close_proof ?warn_incomplete ~opaque:Vernacexpr.Transparent ~keep_body_ucst_separate:false pf in
2056+
let proof = close_proof ?warn_incomplete ~keep_body_ucst_separate:false ~opaque:Vernacexpr.Transparent pf in
2057+
let entries, uctx = process_proof ~info proof.proof_object in
20452058
let { Proof.sigma } = Proof.data pf.proof in
2046-
let sigma = Evd.set_universe_context sigma uctx in
20472059
match entries with
20482060
| [ { proof_entry_body = Default { body; opaque = Transparent } } as entry] ->
2061+
let sigma = Evd.set_universe_context sigma uctx in
20492062
{ entry with proof_entry_body = body }, status, sigma
20502063
| _ ->
20512064
CErrors.anomaly Pp.(str "[build_constant_by_tactic] close_proof returned more than one proof term, or a non transparent one.")
@@ -2225,7 +2238,7 @@ let finish_derived ~f ~name ~entries =
22252238
let ct = declare_constant ~name ~typing_flags:None ~kind:Decls.(IsProof Proposition) lemma_def in
22262239
[GlobRef.ConstRef f_kn; GlobRef.ConstRef ct]
22272240

2228-
let finish_proved_equations ~pm ~kind ~hook i proof_obj types sigma0 =
2241+
let finish_proved_equations ~pm ~kind ~hook i entries types sigma0 =
22292242

22302243
let obls = ref 1 in
22312244
let sigma, recobls =
@@ -2242,14 +2255,14 @@ let finish_proved_equations ~pm ~kind ~hook i proof_obj types sigma0 =
22422255
let sigma, app = Evd.fresh_global (Global.env ()) sigma (GlobRef.ConstRef cst) in
22432256
let sigma = Evd.define ev (EConstr.applist (app, List.map EConstr.of_constr args)) sigma in
22442257
sigma, cst) sigma0
2245-
types proof_obj.entries
2258+
types entries
22462259
in
22472260
let pm = hook ~pm recobls sigma in
22482261
pm, List.map (fun cst -> GlobRef.ConstRef cst) recobls
22492262

2250-
let check_single_entry { entries; uctx } label =
2263+
let check_single_entry entries label =
22512264
match entries with
2252-
| [entry] -> entry, uctx
2265+
| [entry] -> entry
22532266
| _ ->
22542267
CErrors.anomaly ~label Pp.(str "close_proof returned more than one proof term")
22552268

@@ -2258,16 +2271,18 @@ let finish_proof ~pm proof_obj proof_info =
22582271
let { Proof_info.info; cinfo; possible_guard } = proof_info in
22592272
match CEphemeron.default proof_info.Proof_info.proof_ending Regular with
22602273
| Regular ->
2261-
let {entries; uctx} = proof_obj in
2262-
pm, declare_possibly_mutual_definitions ~uctx ~info ~cinfo entries
2274+
pm, declare_possibly_mutual_definitions ~info ~cinfo proof_obj
22632275
| End_obligation oinfo ->
2264-
let entry, uctx = check_single_entry proof_obj "Obligation.save" in
2276+
let entries, uctx = process_proof ~info proof_obj in
2277+
let entry = check_single_entry entries "Obligation.save" in
22652278
Obls_.obligation_terminator ~pm ~entry ~uctx ~oinfo
22662279
| End_derive { f ; name } ->
2267-
pm, finish_derived ~f ~name ~entries:proof_obj.entries
2280+
let entries, uctx = process_proof ~info proof_obj in
2281+
pm, finish_derived ~f ~name ~entries
22682282
| End_equations { hook; i; types; sigma } ->
22692283
let kind = info.Info.kind in
2270-
finish_proved_equations ~pm ~kind ~hook i proof_obj types sigma
2284+
let entries, uctx = process_proof ~info proof_obj in
2285+
finish_proved_equations ~pm ~kind ~hook i entries types sigma
22712286

22722287
let err_save_forbidden_in_place_of_qed () =
22732288
CErrors.user_err (Pp.str "Cannot use Save with more than one constant or in this proof mode")
@@ -2289,7 +2304,7 @@ let save ~pm ~proof ~opaque ~idopt =
22892304
(* Env and sigma are just used for error printing in save_remaining_recthms *)
22902305
let proof_obj = close_proof ~opaque ~keep_body_ucst_separate:false proof in
22912306
let proof_info = process_idopt_for_save ~idopt proof.pinfo in
2292-
finish_proof ~pm proof_obj proof_info
2307+
finish_proof ~pm proof_obj.proof_object proof_info
22932308

22942309
let save_regular ~(proof : t) ~opaque ~idopt =
22952310
let open Proof_ending in
@@ -2303,20 +2318,21 @@ let save_regular ~(proof : t) ~opaque ~idopt =
23032318
(* Special case to close a lemma without forcing a proof *)
23042319
(***********************************************************************)
23052320
let save_lemma_admitted_delayed ~pm ~proof =
2306-
let { entries; uctx; pinfo } = proof in
2321+
let { Proof_object.proof_object; pinfo } = proof in
2322+
let entries, uctx = process_proof ~info:pinfo.info proof_object in
23072323
let typs = List.map (function { proof_entry_type } -> Option.get proof_entry_type) entries in
23082324
(* Note: an alternative would be to compute sec_vars of the partial
23092325
proof as a Future computation, as in compute_proof_using_for_admitted *)
23102326
let sec_vars = if get_keep_admitted_vars () then (List.hd entries).proof_entry_secctx else None in
23112327
(* If the proof is partial, do we want to take the (restriction on
23122328
visible uvars of) uctx so far or (as done below) the initial ones
23132329
that refers to only the types *)
2314-
finish_admitted ~pm ~uctx ~pinfo ~sec_vars typs
2330+
finish_admitted ~pm ~uctx ~pinfo:proof.pinfo ~sec_vars typs
23152331

23162332
let save_lemma_proved_delayed ~pm ~proof ~idopt =
23172333
(* vio2vo used to call this with invalid [pinfo], now it should work fine. *)
2318-
let pinfo = process_idopt_for_save ~idopt proof.pinfo in
2319-
let pm, _ = finish_proof ~pm proof pinfo in
2334+
let pinfo = process_idopt_for_save ~idopt proof.Proof_object.pinfo in
2335+
let pm, _ = finish_proof ~pm proof.proof_object pinfo in
23202336
pm
23212337

23222338
end (* Proof module *)

vernac/vernacstate.ml

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -229,9 +229,9 @@ module Declare_ = struct
229229
cc (fun pt -> Declare.Proof.close_future_proof ~feedback_id pt pf))
230230
()
231231

232-
let close_proof ~opaque ~keep_body_ucst_separate =
232+
let close_proof ~opaque =
233233
NewProfile.profile "close_proof" (fun () ->
234-
cc (fun pt -> Declare.Proof.close_proof ~opaque ~keep_body_ucst_separate pt))
234+
cc (fun pt -> Declare.Proof.close_proof ~opaque pt))
235235
()
236236

237237
let discard_all () = s_lemmas := None

0 commit comments

Comments
 (0)