@@ -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 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
@@ -1875,12 +1932,6 @@ let get_open_goals ps =
1875
1932
(List. map (fun (l1 ,l2 ) -> List. length l1 + List. length l2) stack) +
1876
1933
List. length (Evd. shelf sigma)
1877
1934
1878
- type proof_object =
1879
- { entries : proof_entry list
1880
- ; uctx : UState .t
1881
- ; pinfo : Proof_info .t
1882
- }
1883
-
1884
1935
let warn_remaining_shelved_goals =
1885
1936
CWarnings. create ~name: " remaining-shelved-goals" ~category: CWarnings.CoreCategories. tactics
1886
1937
(fun () -> Pp. str" The proof has remaining shelved goals." )
@@ -1976,7 +2027,7 @@ let prepare_proof ?(warn_incomplete=true) { proof; pinfo } =
1976
2027
let rec_declaration = prepare_recursive_declaration pinfo.cinfo fixtypes fixrelevances fixbodies in
1977
2028
let typing_flags = pinfo.info.typing_flags in
1978
2029
fst (make_recursive_bodies env ~typing_flags ~possible_guard ~rec_declaration ) in
1979
- let proofs = List. map (fun (body , typ ) -> ((body, eff), typ)) proofs in
2030
+ let proofs = List. map (fun (body , typ ) -> ((body, eff), Some typ)) proofs in
1980
2031
let () = if warn_incomplete then check_incomplete_proof evd in
1981
2032
proofs, Evd. evar_universe_context evd
1982
2033
@@ -2027,68 +2078,29 @@ let control_only_guard { proof; pinfo } =
2027
2078
raise (NotGuarded (env, sigma, cofix_error, fix_errors, rec_declaration))
2028
2079
with Exit -> ()
2029
2080
2081
+ let return_proof p = (prepare_proof p : closed_proof_output )
2082
+
2030
2083
let bool_of_opacity = function
2031
2084
| Vernacexpr. Opaque -> true
2032
2085
| Vernacexpr. Transparent -> false
2033
2086
2034
- let close_proof ?warn_incomplete ~opaque ~keep_body_ucst_separate ps =
2035
-
2036
- let { using; proof; initial_euctx; pinfo } = ps in
2037
- let { Proof_info. info = { Info. udecl } } = pinfo in
2038
- let { Proof. poly } = Proof. data proof in
2039
- let elist, uctx = prepare_proof ?warn_incomplete ps in
2087
+ let close_proof ?warn_incomplete ~opaque ~keep_body_ucst_separate (proof : t ) : Proof_object.t =
2040
2088
let opaque = bool_of_opacity opaque in
2089
+ let keep_body_ucst_separate = if keep_body_ucst_separate then Some proof.initial_euctx else None in
2090
+ { proof_object =
2091
+ DefaultProof { proof = prepare_proof ?warn_incomplete proof; opaque; using = proof.using; keep_body_ucst_separate }
2092
+ ; pinfo = proof.pinfo
2093
+ }
2041
2094
2042
- let make_entry ((body , eff ), typ ) =
2043
- let keep_body_ucst_separate = if keep_body_ucst_separate then Some initial_euctx else None in
2044
- let univs, body =
2045
- make_univs_immediate ~poly ?keep_body_ucst_separate ~opaque ~uctx ~udecl ~eff body (Some typ) in
2046
- definition_entry_core ?using ~univs ~types: typ body
2047
- in
2048
- let entries = CList. map make_entry elist in
2049
- { entries; uctx; pinfo }
2050
-
2051
- type closed_proof_output = ((Constr .t * Evd .side_effects ) * Constr .t ) list * UState .t
2052
-
2053
- let close_proof_delayed ~feedback_id ps (fpl : closed_proof_output Future.computation ) =
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 }
2095
+ let close_proof_delayed ~feedback_id proof (fpl : closed_proof_output Future.computation ) : Proof_object.t =
2096
+ { proof_object =
2097
+ DeferredOpaqueProof { deferred_proof = fpl; using = proof.using; initial_proof_data = Proof. data proof.proof;
2098
+ feedback_id; initial_euctx = proof.initial_euctx }
2099
+ ; pinfo = proof.pinfo
2100
+ }
2087
2101
2088
2102
let close_future_proof = close_proof_delayed
2089
2103
2090
- let return_proof p = prepare_proof p
2091
-
2092
2104
let update_sigma_univs ugraph p =
2093
2105
map ~f: (Proof. update_sigma_univs ugraph) p
2094
2106
@@ -2102,11 +2114,12 @@ let build_constant_by_tactic ~name ?warn_incomplete ~sigma ~sign ~poly (typ : EC
2102
2114
let pinfo = Proof_info. make ~cinfo ~info () in
2103
2115
let pf = start_proof_core ~name ~pinfo sigma [Some sign, typ] in
2104
2116
let pf, status = by tac pf in
2105
- let { entries; uctx } = close_proof ?warn_incomplete ~opaque: Vernacexpr. Transparent ~keep_body_ucst_separate: false pf in
2117
+ let proof = close_proof ?warn_incomplete ~keep_body_ucst_separate: false ~opaque: Vernacexpr. Transparent pf in
2118
+ let entries, uctx = process_proof ~info proof.proof_object in
2106
2119
let { Proof. sigma } = Proof. data pf.proof in
2107
- let sigma = Evd. set_universe_context sigma uctx in
2108
2120
match entries with
2109
2121
| [ { proof_entry_body = Default { body; opaque = Transparent } } as entry] ->
2122
+ let sigma = Evd. set_universe_context sigma uctx in
2110
2123
{ entry with proof_entry_body = body }, status, sigma
2111
2124
| _ ->
2112
2125
CErrors. anomaly Pp. (str " [build_constant_by_tactic] close_proof returned more than one proof term, or a non transparent one." )
@@ -2285,7 +2298,7 @@ let finish_derived ~f ~name ~entries =
2285
2298
let ct = declare_constant ~name ~typing_flags: None ~kind: Decls. (IsProof Proposition ) lemma_def in
2286
2299
[GlobRef. ConstRef f_kn; GlobRef. ConstRef ct]
2287
2300
2288
- let finish_proved_equations ~pm ~kind ~hook i proof_obj types sigma0 =
2301
+ let finish_proved_equations ~pm ~kind ~hook i entries types sigma0 =
2289
2302
2290
2303
let obls = ref 1 in
2291
2304
let sigma, recobls =
@@ -2302,14 +2315,14 @@ let finish_proved_equations ~pm ~kind ~hook i proof_obj types sigma0 =
2302
2315
let sigma, app = Evd. fresh_global (Global. env () ) sigma (GlobRef. ConstRef cst) in
2303
2316
let sigma = Evd. define ev (EConstr. applist (app, args)) sigma in
2304
2317
sigma, cst) sigma0
2305
- types proof_obj. entries
2318
+ types entries
2306
2319
in
2307
2320
let pm = hook ~pm recobls sigma in
2308
2321
pm, List. map (fun cst -> GlobRef. ConstRef cst) recobls
2309
2322
2310
- let check_single_entry { entries; uctx } label =
2323
+ let check_single_entry entries label =
2311
2324
match entries with
2312
- | [entry] -> entry, uctx
2325
+ | [entry] -> entry
2313
2326
| _ ->
2314
2327
CErrors. anomaly ~label Pp. (str " close_proof returned more than one proof term" )
2315
2328
@@ -2318,16 +2331,18 @@ let finish_proof ~pm proof_obj proof_info =
2318
2331
let { Proof_info. info; cinfo; possible_guard } = proof_info in
2319
2332
match CEphemeron. default proof_info.Proof_info. proof_ending Regular with
2320
2333
| Regular ->
2321
- let {entries; uctx} = proof_obj in
2322
- pm, declare_possibly_mutual_definitions ~uctx ~info ~cinfo entries
2334
+ pm, declare_possibly_mutual_definitions ~info ~cinfo proof_obj
2323
2335
| End_obligation oinfo ->
2324
- let entry, uctx = check_single_entry proof_obj " Obligation.save" in
2336
+ let entries, uctx = process_proof ~info proof_obj in
2337
+ let entry = check_single_entry entries " Obligation.save" in
2325
2338
Obls_. obligation_terminator ~pm ~entry ~uctx ~oinfo
2326
2339
| End_derive { f ; name } ->
2327
- pm, finish_derived ~f ~name ~entries: proof_obj.entries
2340
+ let entries, uctx = process_proof ~info proof_obj in
2341
+ pm, finish_derived ~f ~name ~entries
2328
2342
| End_equations { hook; i; types; sigma } ->
2329
2343
let kind = info.Info. kind in
2330
- finish_proved_equations ~pm ~kind ~hook i proof_obj types sigma
2344
+ let entries, uctx = process_proof ~info proof_obj in
2345
+ finish_proved_equations ~pm ~kind ~hook i entries types sigma
2331
2346
2332
2347
let err_save_forbidden_in_place_of_qed () =
2333
2348
CErrors. user_err (Pp. str " Cannot use Save with more than one constant or in this proof mode" )
@@ -2349,7 +2364,7 @@ let save ~pm ~proof ~opaque ~idopt =
2349
2364
(* Env and sigma are just used for error printing in save_remaining_recthms *)
2350
2365
let proof_obj = close_proof ~opaque ~keep_body_ucst_separate: false proof in
2351
2366
let proof_info = process_idopt_for_save ~idopt proof.pinfo in
2352
- finish_proof ~pm proof_obj proof_info
2367
+ finish_proof ~pm proof_obj.proof_object proof_info
2353
2368
2354
2369
let save_regular ~(proof : t ) ~opaque ~idopt =
2355
2370
let open Proof_ending in
@@ -2363,20 +2378,21 @@ let save_regular ~(proof : t) ~opaque ~idopt =
2363
2378
(* Special case to close a lemma without forcing a proof *)
2364
2379
(* **********************************************************************)
2365
2380
let save_lemma_admitted_delayed ~pm ~proof =
2366
- let { entries; uctx; pinfo } = proof in
2381
+ let { Proof_object. proof_object; pinfo } = proof in
2382
+ let entries, uctx = process_proof ~info: pinfo.info proof_object in
2367
2383
let typs = List. map (function { proof_entry_type } -> Option. get proof_entry_type) entries in
2368
2384
(* Note: an alternative would be to compute sec_vars of the partial
2369
2385
proof as a Future computation, as in compute_proof_using_for_admitted *)
2370
2386
let sec_vars = if get_keep_admitted_vars () then (List. hd entries).proof_entry_secctx else None in
2371
2387
(* If the proof is partial, do we want to take the (restriction on
2372
2388
visible uvars of) uctx so far or (as done below) the initial ones
2373
2389
that refers to only the types *)
2374
- finish_admitted ~pm ~uctx ~pinfo ~sec_vars typs
2390
+ finish_admitted ~pm ~uctx ~pinfo: proof.pinfo ~sec_vars typs
2375
2391
2376
2392
let save_lemma_proved_delayed ~pm ~proof ~idopt =
2377
2393
(* vio2vo used to call this with invalid [pinfo], now it should work fine. *)
2378
- let pinfo = process_idopt_for_save ~idopt proof.pinfo in
2379
- let pm, _ = finish_proof ~pm proof pinfo in
2394
+ let pinfo = process_idopt_for_save ~idopt proof.Proof_object. pinfo in
2395
+ let pm, _ = finish_proof ~pm proof.proof_object pinfo in
2380
2396
pm
2381
2397
2382
2398
end (* Proof module *)
0 commit comments