@@ -814,6 +814,56 @@ module Internal = struct
814
814
815
815
end
816
816
817
+ (* The word [proof] is to be understood as [justification] *)
818
+ (* A possible alternatve would be [evidence]?? *)
819
+ type closed_proof_output = ((Constr .t * Evd .side_effects ) * Constr .t option ) list * UState .t
820
+
821
+ type proof_object =
822
+ | DefaultProof of
823
+ { proof : closed_proof_output
824
+ ; opaque : bool
825
+ ; using : Names.Id.Set .t option
826
+ ; keep_body_ucst_separate : UState .t option
827
+ }
828
+ | DeferredOpaqueProof of
829
+ { deferred_proof : closed_proof_output Future .computation
830
+ ; using : Names.Id.Set .t option
831
+ ; initial_proof_data : Proof .data
832
+ ; feedback_id : Stateid .t
833
+ ; initial_euctx : UState .t
834
+ }
835
+
836
+ let future_map2_pair_list_distribute p l f =
837
+ List. map_i (fun i c -> f (Future. chain p (fun (a , b ) -> (List. nth a i, b))) c) 0 l
838
+
839
+ let process_proof ~info :Info. ({ udecl; poly } ) ?(is_telescope =false ) = function
840
+ | DefaultProof { proof = (entries , uctx ); opaque; using; keep_body_ucst_separate } ->
841
+ (* Force transparency for Derive-like dependent statements *)
842
+ let opaques =
843
+ let n = List. length entries in
844
+ List. init n (fun i ->
845
+ if i < n-1 && is_telescope then (* waiting for addition of cinfo-based opacity in #19029 *) false
846
+ else opaque) in
847
+ let entries = List. map2 (fun ((body , eff ), typ ) opaque ->
848
+ let uctx, univs, body = make_univs_immediate ~poly ?keep_body_ucst_separate ~opaque ~uctx ~udecl ~eff body typ in
849
+ definition_entry_core ?using ~univs ?types:typ body) entries opaques in
850
+ entries, uctx
851
+ | DeferredOpaqueProof { deferred_proof = bodies ; using; initial_proof_data; feedback_id; initial_euctx } ->
852
+ let { Proof. poly; entry; sigma } = initial_proof_data in
853
+ let entries =
854
+ future_map2_pair_list_distribute bodies (Proofview. initial_goals entry)
855
+ (fun body_typ_uctx (_ , _ , initial_typ ) ->
856
+ (* Testing if evar-closed? *)
857
+ let initial_typ = Evarutil. nf_evars_universes sigma (EConstr.Unsafe. to_constr initial_typ) in
858
+ (* The flags keep_body_ucst_separate, opaque, etc. should be consistent with evar-closedness? *)
859
+ let univs = UState. univ_entry ~poly: false initial_euctx in
860
+ let body = Future. chain body_typ_uctx (fun (((body , eff ), typ ), uctx ) ->
861
+ let uctx = make_univs_deferred_private_mono ~initial_euctx ~uctx ~udecl body typ in
862
+ ((body, uctx), eff)) in
863
+ delayed_definition_entry ?using ~univs ~types: initial_typ ~feedback_id body)
864
+ in
865
+ entries, initial_euctx
866
+
817
867
let declare_definition_scheme ~internal ~univs ~role ~name ?loc c =
818
868
let kind = Decls. (IsDefinition Scheme ) in
819
869
let entry = pure_definition_entry ~univs c in
@@ -891,7 +941,8 @@ let gather_mutual_using_data =
891
941
(* XXX: this should be unified with the code for non-interactive
892
942
mutuals previously on this file. *)
893
943
894
- let declare_possibly_mutual_definitions ~info ~cinfo ~uctx entries =
944
+ let declare_possibly_mutual_definitions ~info ~cinfo proof_obj =
945
+ let entries, uctx = process_proof ~info proof_obj in
895
946
let { Info. hook; scope; clearbody; kind; typing_flags; user_warns; _ } = info in
896
947
let refs = List. map2 (fun CInfo. {name; impargs} ->
897
948
declare_entry ~name ~scope ~clearbody ~kind ?hook ~impargs ~typing_flags ~user_warns ~uctx ) cinfo entries in
@@ -1651,10 +1702,6 @@ module Proof_ending = struct
1651
1702
1652
1703
end
1653
1704
1654
- (* Alias *)
1655
- module Proof_ = Proof
1656
- module Proof = struct
1657
-
1658
1705
module Proof_info = struct
1659
1706
1660
1707
type t =
@@ -1678,6 +1725,22 @@ module Proof_info = struct
1678
1725
1679
1726
end
1680
1727
1728
+ module Proof_object = struct
1729
+
1730
+ type t =
1731
+ { proof_object : proof_object
1732
+ ; pinfo : Proof_info .t
1733
+ }
1734
+
1735
+ end
1736
+
1737
+ (* Alias *)
1738
+ module Proof_ = Proof
1739
+ module Proof = struct
1740
+
1741
+ type nonrec closed_proof_output = closed_proof_output
1742
+ type proof_object = Proof_object .t
1743
+
1681
1744
type t =
1682
1745
{ endline_tactic : Genarg .glob_generic_argument option
1683
1746
; using : Id.Set .t option
@@ -1877,12 +1940,6 @@ let get_open_goals ps =
1877
1940
(List. map (fun (l1 ,l2 ) -> List. length l1 + List. length l2) stack) +
1878
1941
List. length (Evd. shelf sigma)
1879
1942
1880
- type proof_object =
1881
- { entries : proof_entry list
1882
- ; uctx : UState .t
1883
- ; pinfo : Proof_info .t
1884
- }
1885
-
1886
1943
let warn_remaining_shelved_goals =
1887
1944
CWarnings. create ~name: " remaining-shelved-goals" ~category: CWarnings.CoreCategories. tactics
1888
1945
(fun () -> Pp. str" The proof has remaining shelved goals." )
@@ -1978,7 +2035,7 @@ let prepare_proof ?(warn_incomplete=true) { proof; pinfo } =
1978
2035
let rec_declaration = prepare_recursive_declaration pinfo.cinfo fixtypes fixrelevances fixbodies in
1979
2036
let typing_flags = pinfo.info.typing_flags in
1980
2037
fst (make_recursive_bodies env ~typing_flags ~possible_guard ~rec_declaration ) in
1981
- let proofs = List. map (fun (body , typ ) -> ((body, eff), typ)) proofs in
2038
+ let proofs = List. map (fun (body , typ ) -> ((body, eff), Some typ)) proofs in
1982
2039
let () = if warn_incomplete then check_incomplete_proof evd in
1983
2040
proofs, Evd. ustate evd
1984
2041
@@ -2029,76 +2086,29 @@ let control_only_guard { proof; pinfo } =
2029
2086
raise (NotGuarded (env, sigma, cofix_error, fix_errors, rec_declaration))
2030
2087
with Exit -> ()
2031
2088
2032
- let close_proof ?warn_incomplete ~opaque ~keep_body_ucst_separate ps =
2033
- NewProfile. profile " close_proof" (fun () ->
2089
+ let return_proof p = (prepare_proof p : closed_proof_output )
2034
2090
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
2039
- let opaques =
2040
- let n = List. length elist in
2041
- let is_derived = match CEphemeron. default pinfo.Proof_info. proof_ending Regular with End_derive -> true | _ -> false in
2042
- List. init n (fun i ->
2043
- if i < n-1 && is_derived then
2044
- (* Temporary code for setting opacity in Derive, waiting for addition of cinfo-based opacity in #19029 *)
2045
- false
2046
- else match opaque with
2047
- | Vernacexpr. Opaque -> true
2048
- | Vernacexpr. Transparent -> false ) in
2049
- let make_entry ((body , eff ), typ ) opaque =
2050
- let keep_body_ucst_separate = if keep_body_ucst_separate then Some initial_euctx else None in
2051
- let _, univs, body =
2052
- make_univs_immediate ~poly ?keep_body_ucst_separate ~opaque ~uctx ~udecl ~eff body (Some typ) in
2053
- definition_entry_core ?using ~univs ~types: typ body
2054
- in
2055
- let entries = List. map2 make_entry elist opaques in
2056
- { entries; uctx; pinfo })
2057
- ()
2058
-
2059
- type closed_proof_output = ((Constr .t * Evd .side_effects ) * Constr .t ) list * UState .t
2060
-
2061
- let close_proof_delayed ~feedback_id ps (fpl : closed_proof_output Future.computation ) =
2091
+ let close_proof ?warn_incomplete ~opaque ~keep_body_ucst_separate (proof : t ) : Proof_object.t =
2092
+ NewProfile. profile " close_proof" (fun () ->
2093
+ let opaque = match opaque with
2094
+ | Vernacexpr. Opaque -> true
2095
+ | Vernacexpr. Transparent -> false in
2096
+ let keep_body_ucst_separate = if keep_body_ucst_separate then Some proof.initial_euctx else None in
2097
+ { Proof_object. proof_object =
2098
+ DefaultProof { proof = prepare_proof ?warn_incomplete proof; opaque; using = proof.using; keep_body_ucst_separate }
2099
+ ; pinfo = proof.pinfo
2100
+ }) ()
2101
+
2102
+ let close_proof_delayed ~feedback_id proof (fpl : closed_proof_output Future.computation ) : Proof_object.t =
2062
2103
NewProfile. profile " close_proof_delayed" (fun () ->
2063
- let { using; proof; initial_euctx; pinfo } = ps in
2064
- let { Proof_info. info = { Info. udecl } } = pinfo in
2065
- let { Proof. poly; entry; sigma } = Proof. data proof in
2066
-
2067
- (* We don't allow poly = true in this path *)
2068
- if poly then
2069
- CErrors. anomaly (Pp. str " Cannot delay universe-polymorphic constants." );
2070
-
2071
- (* Because of dependent subgoals at the beginning of proofs, we could
2072
- have existential variables in the initial types of goals, we need to
2073
- normalise them for the kernel. *)
2074
- let nf = Evarutil. nf_evars_universes (Evd. set_universe_context sigma initial_euctx) in
2075
-
2076
- (* We only support opaque proofs, this will be enforced by using
2077
- different entries soon *)
2078
- let make_entry i (_ , _ , types ) =
2079
- (* Already checked the univ_decl for the type universes when starting the proof. *)
2080
- let univs = UState. univ_entry ~poly: false initial_euctx in
2081
- let typ = nf (EConstr.Unsafe. to_constr types) in
2082
-
2083
- (* NB: for Admitted proofs [fpl] is not valid (raises anomaly when forced) *)
2084
- Future. chain fpl (fun (pf , uctx ) ->
2085
- let (body, eff) = fst (List. nth pf i) in
2086
- (* Deferred proof, we already checked the universe declaration with
2087
- the initial universes, ensure that the final universes respect
2088
- the declaration as well. If the declaration is non-extensible,
2089
- this will prevent the body from adding universes and constraints. *)
2090
- let uctx = make_univs_deferred_private_mono ~initial_euctx ~uctx ~udecl body (Some typ) in
2091
- ((body, uctx), eff))
2092
- |> delayed_definition_entry ?using ~univs ~types: typ ~feedback_id
2093
- in
2094
- let entries = CList. map_i make_entry 0 (Proofview. initial_goals entry) in
2095
- { entries; uctx = initial_euctx; pinfo })
2096
- ()
2104
+ { Proof_object. proof_object =
2105
+ DeferredOpaqueProof { deferred_proof = fpl; using = proof.using; initial_proof_data = Proof. data proof.proof;
2106
+ feedback_id; initial_euctx = proof.initial_euctx }
2107
+ ; pinfo = proof.pinfo
2108
+ }) ()
2097
2109
2098
2110
let close_future_proof = close_proof_delayed
2099
2111
2100
- let return_proof p = prepare_proof p
2101
-
2102
2112
let update_sigma_univs ugraph p =
2103
2113
map ~f: (Proof. update_sigma_univs ugraph) p
2104
2114
@@ -2112,11 +2122,12 @@ let build_constant_by_tactic ~name ?warn_incomplete ~sigma ~sign ~poly (typ : EC
2112
2122
let pinfo = Proof_info. make ~cinfo ~info () in
2113
2123
let pf = start_proof_core ~name ~pinfo sigma [Some sign, typ] in
2114
2124
let pf, status = by tac pf in
2115
- let { entries; uctx } = close_proof ?warn_incomplete ~opaque: Vernacexpr. Transparent ~keep_body_ucst_separate: false pf in
2125
+ let proof = close_proof ?warn_incomplete ~keep_body_ucst_separate: false ~opaque: Vernacexpr. Transparent pf in
2126
+ let entries, uctx = process_proof ~info proof.proof_object in
2116
2127
let { Proof. sigma } = Proof. data pf.proof in
2117
- let sigma = Evd. set_universe_context sigma uctx in
2118
2128
match entries with
2119
2129
| [ { proof_entry_body = Default { body; opaque = Transparent } } as entry] ->
2130
+ let sigma = Evd. set_universe_context sigma uctx in
2120
2131
{ entry with proof_entry_body = body }, status, sigma
2121
2132
| _ ->
2122
2133
CErrors. anomaly Pp. (str " [build_constant_by_tactic] close_proof returned more than one proof term, or a non transparent one." )
@@ -2259,7 +2270,7 @@ let save_admitted ~pm ~proof =
2259
2270
(* Saving a lemma-like constant *)
2260
2271
(* ***********************************************************************)
2261
2272
2262
- let finish_derived {entries; pinfo; uctx} =
2273
+ let finish_derived pinfo uctx entries =
2263
2274
let n = List. length entries in
2264
2275
let { Proof_info. info = { Info. hook; scope; clearbody; kind; typing_flags; user_warns; poly; udecl; _ } } = pinfo in
2265
2276
let _, _, refs, _ =
@@ -2281,7 +2292,7 @@ let finish_derived {entries; pinfo; uctx} =
2281
2292
(0 , [] , [] , Univ.Level.Set. empty) pinfo.Proof_info. cinfo entries in
2282
2293
refs
2283
2294
2284
- let finish_proved_equations ~pm ~kind ~hook i proof_obj types sigma0 =
2295
+ let finish_proved_equations ~pm ~kind ~hook i entries types sigma0 =
2285
2296
2286
2297
let obls = ref 1 in
2287
2298
let sigma, recobls =
@@ -2298,14 +2309,14 @@ let finish_proved_equations ~pm ~kind ~hook i proof_obj types sigma0 =
2298
2309
let sigma, app = Evd. fresh_global (Global. env () ) sigma (GlobRef. ConstRef cst) in
2299
2310
let sigma = Evd. define ev (EConstr. applist (app, args)) sigma in
2300
2311
sigma, cst) sigma0
2301
- types proof_obj. entries
2312
+ types entries
2302
2313
in
2303
2314
let pm = hook ~pm recobls sigma in
2304
2315
pm, List. map (fun cst -> GlobRef. ConstRef cst) recobls
2305
2316
2306
- let check_single_entry { entries; uctx } label =
2317
+ let check_single_entry entries label =
2307
2318
match entries with
2308
- | [entry] -> entry, uctx
2319
+ | [entry] -> entry
2309
2320
| _ ->
2310
2321
CErrors. anomaly ~label Pp. (str " close_proof returned more than one proof term" )
2311
2322
@@ -2314,16 +2325,18 @@ let finish_proof ~pm proof_obj proof_info =
2314
2325
let { Proof_info. info; cinfo; possible_guard } = proof_info in
2315
2326
match CEphemeron. default proof_info.Proof_info. proof_ending Regular with
2316
2327
| Regular ->
2317
- let {entries; uctx} = proof_obj in
2318
- pm, declare_possibly_mutual_definitions ~uctx ~info ~cinfo entries
2328
+ pm, declare_possibly_mutual_definitions ~info ~cinfo proof_obj
2319
2329
| End_obligation oinfo ->
2320
- let entry, uctx = check_single_entry proof_obj " Obligation.save" in
2330
+ let entries, uctx = process_proof ~info proof_obj in
2331
+ let entry = check_single_entry entries " Obligation.save" in
2321
2332
Obls_. obligation_terminator ~pm ~entry ~uctx ~oinfo
2322
2333
| End_derive ->
2323
- pm, finish_derived proof_obj
2334
+ let entries, uctx = process_proof ~info ~is_telescope: true proof_obj in
2335
+ pm, finish_derived proof_info uctx entries
2324
2336
| End_equations { hook; i; types; sigma } ->
2325
2337
let kind = info.Info. kind in
2326
- finish_proved_equations ~pm ~kind ~hook i proof_obj types sigma
2338
+ let entries, uctx = process_proof ~info proof_obj in
2339
+ finish_proved_equations ~pm ~kind ~hook i entries types sigma
2327
2340
2328
2341
let err_save_forbidden_in_place_of_qed () =
2329
2342
CErrors. user_err (Pp. str " Cannot use Save with more than one constant or in this proof mode" )
@@ -2345,7 +2358,7 @@ let save ~pm ~proof ~opaque ~idopt =
2345
2358
(* Env and sigma are just used for error printing in save_remaining_recthms *)
2346
2359
let proof_obj = close_proof ~opaque ~keep_body_ucst_separate: false proof in
2347
2360
let proof_info = process_idopt_for_save ~idopt proof.pinfo in
2348
- finish_proof ~pm proof_obj proof_info
2361
+ finish_proof ~pm proof_obj.proof_object proof_info
2349
2362
2350
2363
let save_regular ~(proof : t ) ~opaque ~idopt =
2351
2364
let open Proof_ending in
@@ -2359,20 +2372,21 @@ let save_regular ~(proof : t) ~opaque ~idopt =
2359
2372
(* Special case to close a lemma without forcing a proof *)
2360
2373
(* **********************************************************************)
2361
2374
let save_lemma_admitted_delayed ~pm ~proof =
2362
- let { entries; uctx; pinfo } = proof in
2375
+ let { Proof_object. proof_object; pinfo } = proof in
2376
+ let entries, uctx = process_proof ~info: pinfo.info proof_object in
2363
2377
let typs = List. map (function { proof_entry_type } -> Option. get proof_entry_type) entries in
2364
2378
(* Note: an alternative would be to compute sec_vars of the partial
2365
2379
proof as a Future computation, as in compute_proof_using_for_admitted *)
2366
2380
let sec_vars = if get_keep_admitted_vars () then (List. hd entries).proof_entry_secctx else None in
2367
2381
(* If the proof is partial, do we want to take the (restriction on
2368
2382
visible uvars of) uctx so far or (as done below) the initial ones
2369
2383
that refers to only the types *)
2370
- finish_admitted ~pm ~uctx ~pinfo ~sec_vars typs
2384
+ finish_admitted ~pm ~uctx ~pinfo: proof.pinfo ~sec_vars typs
2371
2385
2372
2386
let save_lemma_proved_delayed ~pm ~proof ~idopt =
2373
2387
(* vio2vo used to call this with invalid [pinfo], now it should work fine. *)
2374
- let pinfo = process_idopt_for_save ~idopt proof.pinfo in
2375
- let pm, _ = finish_proof ~pm proof pinfo in
2388
+ let pinfo = process_idopt_for_save ~idopt proof.Proof_object. pinfo in
2389
+ let pm, _ = finish_proof ~pm proof.proof_object pinfo in
2376
2390
pm
2377
2391
2378
2392
end (* Proof module *)
0 commit comments