Skip to content

Commit

Permalink
Validate Proof command: runs typing.ml on the current proof
Browse files Browse the repository at this point in the history
  • Loading branch information
SkySkimmer committed Apr 6, 2023
1 parent 41b4309 commit f171149
Show file tree
Hide file tree
Showing 7 changed files with 50 additions and 1 deletion.
1 change: 1 addition & 0 deletions doc/tools/docgram/fullGrammar
Original file line number Diff line number Diff line change
Expand Up @@ -526,6 +526,7 @@ command: [
| "Show" "Intros"
| "Show" "Match" reference
| "Guarded"
| "Validate" "Proof"
| "Create" "HintDb" IDENT; [ "discriminated" | ]
| "Remove" "Hints" LIST1 global opt_hintbases
| "Hint" hint opt_hintbases
Expand Down
1 change: 1 addition & 0 deletions doc/tools/docgram/orderedGrammar
Original file line number Diff line number Diff line change
Expand Up @@ -977,6 +977,7 @@ command: [
| "Show" "Intros"
| "Show" "Match" qualid
| "Guarded"
| "Validate" "Proof"
| "Create" "HintDb" ident OPT "discriminated"
| "Remove" "Hints" LIST1 qualid OPT ( ":" LIST1 ident )
| "Comments" LIST0 [ one_term | string | natural ]
Expand Down
1 change: 1 addition & 0 deletions vernac/g_proofs.mlg
Original file line number Diff line number Diff line change
Expand Up @@ -95,6 +95,7 @@ GRAMMAR EXTEND Gram
| IDENT "Show"; IDENT "Intros" -> { VernacShow (ShowIntros true) }
| IDENT "Show"; IDENT "Match"; id = reference -> { VernacShow (ShowMatch id) }
| IDENT "Guarded" -> { VernacCheckGuard }
| IDENT "Validate"; IDENT "Proof" -> { VernacValidateProof }
(* Hints for Auto and EAuto *)
| IDENT "Create"; IDENT "HintDb" ;
id = IDENT ; b = [ IDENT "discriminated" -> { true } | -> { false } ] ->
Expand Down
2 changes: 2 additions & 0 deletions vernac/ppvernac.ml
Original file line number Diff line number Diff line change
Expand Up @@ -734,6 +734,8 @@ let pr_vernac_expr v =
| VernacCheckGuard ->
return (keyword "Guarded")

| VernacValidateProof -> return (keyword "Validate Proof")

(* Resetting *)
| VernacResetName id ->
return (keyword "Reset" ++ spc() ++ pr_lident id)
Expand Down
2 changes: 1 addition & 1 deletion vernac/vernac_classifier.ml
Original file line number Diff line number Diff line change
Expand Up @@ -71,7 +71,7 @@ let classify_vernac e =
| VernacProof _
| VernacFocus _ | VernacUnfocus
| VernacSubproof _
| VernacCheckGuard
| VernacCheckGuard | VernacValidateProof
| VernacUnfocused
| VernacBullet _ ->
VtProofStep { proof_block_detection = Some "bullet" }
Expand Down
43 changes: 43 additions & 0 deletions vernac/vernacentries.ml
Original file line number Diff line number Diff line change
Expand Up @@ -2345,6 +2345,45 @@ let vernac_check_guard ~pstate =
Inductiveops.control_only_guard env sigma pfterm;
str "The condition holds up to here."

let vernac_validate_proof ~pstate =
let pts = Declare.Proof.get pstate in
let { Proof.entry; Proof.sigma } = Proof.data pts in
let hyps, pfterm, pftyp = List.hd (Proofview.initial_goals entry) in
(* XXX can the initial hyps contain something broken? For now assume they're correct.
NB: in the "Lemma foo args : bla." case the args are part of the
term and intro'd after the proof is opened. Only the section
variables are in the hyps. *)
let env = Environ.reset_with_named_context hyps (Global.env ()) in
let sigma = Evarconv.solve_unif_constraints_with_heuristics env sigma in
let sigma' = Typing.check env sigma pfterm pftyp in
let evar_issues =
(* Use Evar.Map.merge as a kind of for_all2 *)
Evar.Map.merge (fun e orig now -> match orig, now with
| None, None -> assert false
| Some _, Some _ -> None (* assume same *)
| Some evi, None ->
let EvarInfo evi' = Evd.find sigma' e in
let body = match Evd.evar_body evi' with
| Evar_empty -> assert false
| Evar_defined body -> body
in
Some
Pp.(str "Evar " ++ Printer.pr_evar sigma (e, evi)
++ spc() ++ str "was inferred by unification to be" ++ spc()
++ pr_econstr_env (Evd.evar_env env evi') sigma' body)
| None, Some _ -> (* ignore new evar *)
assert (not (Evd.is_defined sigma e));
None
)
(Evd.undefined_map sigma)
(Evd.undefined_map sigma')
in
(* TODO check ustate *)

if Evar.Map.is_empty evar_issues then
str "No issues found."
else prlist_with_sep fnl snd (Evar.Map.bindings evar_issues)

(* We interpret vernacular commands to a DSL that specifies their
allowed actions on proof states *)
let translate_vernac ?loc ~atts v = let open Vernacextend in match v with
Expand Down Expand Up @@ -2649,6 +2688,10 @@ let translate_vernac ?loc ~atts v = let open Vernacextend in match v with
vtreadproof(fun ~pstate ->
unsupported_attributes atts;
Feedback.msg_notice @@ vernac_check_guard ~pstate)
| VernacValidateProof ->
vtreadproof(fun ~pstate ->
unsupported_attributes atts;
Feedback.msg_notice @@ vernac_validate_proof ~pstate)
| VernacProof (tac, using) ->
vtmodifyproof(fun ~pstate ->
unsupported_attributes atts;
Expand Down
1 change: 1 addition & 0 deletions vernac/vernacexpr.ml
Original file line number Diff line number Diff line change
Expand Up @@ -476,6 +476,7 @@ type nonrec vernac_expr =
| VernacEndSubproof
| VernacShow of showable
| VernacCheckGuard
| VernacValidateProof
| VernacProof of Genarg.raw_generic_argument option * section_subset_expr option
| VernacProofMode of string

Expand Down

0 comments on commit f171149

Please sign in to comment.