@@ -811,6 +811,50 @@ module Internal = struct
811
811
812
812
end
813
813
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
+
814
858
let declare_definition_scheme ~internal ~univs ~role ~name ?loc c =
815
859
let kind = Decls. (IsDefinition Scheme ) in
816
860
let entry = pure_definition_entry ~univs c in
@@ -888,7 +932,8 @@ let gather_mutual_using_data =
888
932
(* XXX: this should be unified with the code for non-interactive
889
933
mutuals previously on this file. *)
890
934
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
892
937
let { Info. hook; scope; clearbody; kind; typing_flags; user_warns; _ } = info in
893
938
let refs = List. map2 (fun CInfo. {name; impargs} ->
894
939
declare_entry ~name ~scope ~clearbody ~kind ?hook ~impargs ~typing_flags ~user_warns ~uctx ) cinfo entries in
@@ -1648,10 +1693,6 @@ module Proof_ending = struct
1648
1693
1649
1694
end
1650
1695
1651
- (* Alias *)
1652
- module Proof_ = Proof
1653
- module Proof = struct
1654
-
1655
1696
module Proof_info = struct
1656
1697
1657
1698
type t =
@@ -1675,6 +1716,22 @@ module Proof_info = struct
1675
1716
1676
1717
end
1677
1718
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
+
1678
1735
type t =
1679
1736
{ endline_tactic : Genarg .glob_generic_argument option
1680
1737
; using : Id.Set .t option
@@ -1874,12 +1931,6 @@ let get_open_goals ps =
1874
1931
(List. map (fun (l1 ,l2 ) -> List. length l1 + List. length l2) stack) +
1875
1932
List. length (Evd. shelf sigma)
1876
1933
1877
- type proof_object =
1878
- { entries : proof_entry list
1879
- ; uctx : UState .t
1880
- ; pinfo : Proof_info .t
1881
- }
1882
-
1883
1934
let warn_remaining_shelved_goals =
1884
1935
CWarnings. create ~name: " remaining-shelved-goals" ~category: CWarnings.CoreCategories. tactics
1885
1936
(fun () -> Pp. str" The proof has remaining shelved goals." )
@@ -1975,7 +2026,7 @@ let prepare_proof ?(warn_incomplete=true) { proof; pinfo } =
1975
2026
let rec_declaration = prepare_recursive_declaration pinfo.cinfo fixtypes fixrelevances fixbodies in
1976
2027
let typing_flags = pinfo.info.typing_flags in
1977
2028
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
1979
2030
let () = if warn_incomplete then check_incomplete_proof evd in
1980
2031
proofs, Evd. ustate evd
1981
2032
@@ -2026,70 +2077,29 @@ let control_only_guard { proof; pinfo } =
2026
2077
raise (NotGuarded (env, sigma, cofix_error, fix_errors, rec_declaration))
2027
2078
with Exit -> ()
2028
2079
2029
- let close_proof ?warn_incomplete ~opaque ~keep_body_ucst_separate ps =
2030
- NewProfile. profile " close_proof" (fun () ->
2080
+ let return_proof p = (prepare_proof p : closed_proof_output )
2031
2081
2032
- let { using; proof; initial_euctx; pinfo } = ps in
2033
- let { Proof_info. info = { Info. udecl } } = pinfo in
2034
- let { Proof. poly } = Proof. data proof in
2035
- let elist, uctx = prepare_proof ?warn_incomplete ps in
2082
+ let close_proof ?warn_incomplete ~opaque ~keep_body_ucst_separate (proof : t ) : Proof_object.t =
2083
+ NewProfile. profile " close_proof" (fun () ->
2036
2084
let opaque = match opaque with
2037
- | Vernacexpr. Opaque -> true
2038
- | Vernacexpr. Transparent -> false in
2039
-
2040
- let make_entry ((body , eff ), typ ) =
2041
- let keep_body_ucst_separate = if keep_body_ucst_separate then Some initial_euctx else None in
2042
- let _, univs, body =
2043
- make_univs_immediate ~poly ?keep_body_ucst_separate ~opaque ~uctx ~udecl ~eff body (Some typ) in
2044
- definition_entry_core ?using ~univs ~types: typ body
2045
- in
2046
- let entries = CList. map make_entry elist in
2047
- { entries; uctx; pinfo })
2048
- ()
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 ) =
2085
+ | Vernacexpr. Opaque -> true
2086
+ | Vernacexpr. Transparent -> false in
2087
+ let keep_body_ucst_separate = if keep_body_ucst_separate then Some proof.initial_euctx else None in
2088
+ { Proof_object. proof_object =
2089
+ DefaultProof { proof = prepare_proof ?warn_incomplete proof; opaque; using = proof.using; keep_body_ucst_separate }
2090
+ ; pinfo = proof.pinfo
2091
+ }) ()
2092
+
2093
+ let close_proof_delayed ~feedback_id proof (fpl : closed_proof_output Future.computation ) : Proof_object.t =
2053
2094
NewProfile. profile " close_proof_delayed" (fun () ->
2054
- let { using; proof; initial_euctx; pinfo } = ps in
2055
- let { Proof_info. info = { Info. udecl } } = pinfo in
2056
- let { Proof. poly; entry; sigma } = Proof. data proof in
2057
-
2058
- (* We don't allow poly = true in this path *)
2059
- if poly then
2060
- CErrors. anomaly (Pp. str " Cannot delay universe-polymorphic constants." );
2061
-
2062
- (* Because of dependent subgoals at the beginning of proofs, we could
2063
- have existential variables in the initial types of goals, we need to
2064
- normalise them for the kernel. *)
2065
- let nf = Evarutil. nf_evars_universes (Evd. set_universe_context sigma initial_euctx) in
2066
-
2067
- (* We only support opaque proofs, this will be enforced by using
2068
- different entries soon *)
2069
- let make_entry i (_ , _ , types ) =
2070
- (* Already checked the univ_decl for the type universes when starting the proof. *)
2071
- let univs = UState. univ_entry ~poly: false initial_euctx in
2072
- let typ = nf (EConstr.Unsafe. to_constr types) in
2073
-
2074
- (* NB: for Admitted proofs [fpl] is not valid (raises anomaly when forced) *)
2075
- Future. chain fpl (fun (pf , uctx ) ->
2076
- let (body, eff) = fst (List. nth pf i) in
2077
- (* Deferred proof, we already checked the universe declaration with
2078
- the initial universes, ensure that the final universes respect
2079
- the declaration as well. If the declaration is non-extensible,
2080
- this will prevent the body from adding universes and constraints. *)
2081
- let uctx = make_univs_deferred_private_mono ~initial_euctx ~uctx ~udecl body (Some typ) in
2082
- ((body, uctx), eff))
2083
- |> delayed_definition_entry ?using ~univs ~types: typ ~feedback_id
2084
- in
2085
- let entries = CList. map_i make_entry 0 (Proofview. initial_goals entry) in
2086
- { entries; uctx = initial_euctx; pinfo })
2087
- ()
2095
+ { Proof_object. 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
+ }) ()
2088
2100
2089
2101
let close_future_proof = close_proof_delayed
2090
2102
2091
- let return_proof p = prepare_proof p
2092
-
2093
2103
let update_sigma_univs ugraph p =
2094
2104
map ~f: (Proof. update_sigma_univs ugraph) p
2095
2105
@@ -2103,11 +2113,12 @@ let build_constant_by_tactic ~name ?warn_incomplete ~sigma ~sign ~poly (typ : EC
2103
2113
let pinfo = Proof_info. make ~cinfo ~info () in
2104
2114
let pf = start_proof_core ~name ~pinfo sigma [Some sign, typ] in
2105
2115
let pf, status = by tac pf in
2106
- 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
2107
2118
let { Proof. sigma } = Proof. data pf.proof in
2108
- let sigma = Evd. set_universe_context sigma uctx in
2109
2119
match entries with
2110
2120
| [ { proof_entry_body = Default { body; opaque = Transparent } } as entry] ->
2121
+ let sigma = Evd. set_universe_context sigma uctx in
2111
2122
{ entry with proof_entry_body = body }, status, sigma
2112
2123
| _ ->
2113
2124
CErrors. anomaly Pp. (str " [build_constant_by_tactic] close_proof returned more than one proof term, or a non transparent one." )
@@ -2250,7 +2261,7 @@ let save_admitted ~pm ~proof =
2250
2261
(* Saving a lemma-like constant *)
2251
2262
(* ***********************************************************************)
2252
2263
2253
- let finish_derived ~f ~name {entries; pinfo; uctx} =
2264
+ let finish_derived ~f ~name ~ uctx ~ pinfo entries =
2254
2265
(* [f] and [name] correspond to the proof of [f] and of [suchthat], respectively. *)
2255
2266
2256
2267
let { Proof_info. info = { Info. hook; scope; clearbody; kind; typing_flags; user_warns; poly; udecl; _ } } = pinfo in
@@ -2273,7 +2284,7 @@ let finish_derived ~f ~name {entries; pinfo; uctx} =
2273
2284
(0 , [] , [] , Univ.Level.Set. empty) pinfo.Proof_info. cinfo entries in
2274
2285
refs
2275
2286
2276
- let finish_proved_equations ~pm ~kind ~hook i proof_obj types sigma0 =
2287
+ let finish_proved_equations ~pm ~kind ~hook i entries types sigma0 =
2277
2288
2278
2289
let obls = ref 1 in
2279
2290
let sigma, recobls =
@@ -2290,14 +2301,14 @@ let finish_proved_equations ~pm ~kind ~hook i proof_obj types sigma0 =
2290
2301
let sigma, app = Evd. fresh_global (Global. env () ) sigma (GlobRef. ConstRef cst) in
2291
2302
let sigma = Evd. define ev (EConstr. applist (app, args)) sigma in
2292
2303
sigma, cst) sigma0
2293
- types proof_obj. entries
2304
+ types entries
2294
2305
in
2295
2306
let pm = hook ~pm recobls sigma in
2296
2307
pm, List. map (fun cst -> GlobRef. ConstRef cst) recobls
2297
2308
2298
- let check_single_entry { entries; uctx } label =
2309
+ let check_single_entry entries label =
2299
2310
match entries with
2300
- | [entry] -> entry, uctx
2311
+ | [entry] -> entry
2301
2312
| _ ->
2302
2313
CErrors. anomaly ~label Pp. (str " close_proof returned more than one proof term" )
2303
2314
@@ -2306,16 +2317,18 @@ let finish_proof ~pm proof_obj proof_info =
2306
2317
let { Proof_info. info; cinfo; possible_guard } = proof_info in
2307
2318
match CEphemeron. default proof_info.Proof_info. proof_ending Regular with
2308
2319
| Regular ->
2309
- let {entries; uctx} = proof_obj in
2310
- pm, declare_possibly_mutual_definitions ~uctx ~info ~cinfo entries
2320
+ pm, declare_possibly_mutual_definitions ~info ~cinfo proof_obj
2311
2321
| End_obligation oinfo ->
2312
- 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
2313
2324
Obls_. obligation_terminator ~pm ~entry ~uctx ~oinfo
2314
2325
| End_derive { f ; name } ->
2315
- 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
2316
2328
| End_equations { hook; i; types; sigma } ->
2317
2329
let kind = info.Info. kind in
2318
- 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
2319
2332
2320
2333
let err_save_forbidden_in_place_of_qed () =
2321
2334
CErrors. user_err (Pp. str " Cannot use Save with more than one constant or in this proof mode" )
@@ -2337,7 +2350,7 @@ let save ~pm ~proof ~opaque ~idopt =
2337
2350
(* Env and sigma are just used for error printing in save_remaining_recthms *)
2338
2351
let proof_obj = close_proof ~opaque ~keep_body_ucst_separate: false proof in
2339
2352
let proof_info = process_idopt_for_save ~idopt proof.pinfo in
2340
- finish_proof ~pm proof_obj proof_info
2353
+ finish_proof ~pm proof_obj.proof_object proof_info
2341
2354
2342
2355
let save_regular ~(proof : t ) ~opaque ~idopt =
2343
2356
let open Proof_ending in
@@ -2351,20 +2364,21 @@ let save_regular ~(proof : t) ~opaque ~idopt =
2351
2364
(* Special case to close a lemma without forcing a proof *)
2352
2365
(* **********************************************************************)
2353
2366
let save_lemma_admitted_delayed ~pm ~proof =
2354
- 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
2355
2369
let typs = List. map (function { proof_entry_type } -> Option. get proof_entry_type) entries in
2356
2370
(* Note: an alternative would be to compute sec_vars of the partial
2357
2371
proof as a Future computation, as in compute_proof_using_for_admitted *)
2358
2372
let sec_vars = if get_keep_admitted_vars () then (List. hd entries).proof_entry_secctx else None in
2359
2373
(* If the proof is partial, do we want to take the (restriction on
2360
2374
visible uvars of) uctx so far or (as done below) the initial ones
2361
2375
that refers to only the types *)
2362
- finish_admitted ~pm ~uctx ~pinfo ~sec_vars typs
2376
+ finish_admitted ~pm ~uctx ~pinfo: proof.pinfo ~sec_vars typs
2363
2377
2364
2378
let save_lemma_proved_delayed ~pm ~proof ~idopt =
2365
2379
(* vio2vo used to call this with invalid [pinfo], now it should work fine. *)
2366
- let pinfo = process_idopt_for_save ~idopt proof.pinfo in
2367
- 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
2368
2382
pm
2369
2383
2370
2384
end (* Proof module *)
0 commit comments