aboutsummaryrefslogtreecommitdiffhomepage
path: root/stm/vernac_classifier.ml
diff options
context:
space:
mode:
Diffstat (limited to 'stm/vernac_classifier.ml')
-rw-r--r--stm/vernac_classifier.ml31
1 files changed, 17 insertions, 14 deletions
diff --git a/stm/vernac_classifier.ml b/stm/vernac_classifier.ml
index 2fa47ba43..cbbb54e45 100644
--- a/stm/vernac_classifier.ml
+++ b/stm/vernac_classifier.ml
@@ -48,18 +48,19 @@ let declare_vernac_classifier
=
classifiers := !classifiers @ [s,f]
+let idents_of_name : Names.Name.t -> Names.Id.t list =
+ function
+ | Names.Anonymous -> []
+ | Names.Name n -> [n]
+
let classify_vernac e =
- let rec static_classifier ~poly e = match e with
+ let static_classifier ~poly e = match e with
(* Univ poly compatibility: we run it now, so that we can just
* look at Flags in stm.ml. Would be nicer to have the stm
* look at the entire dag to detect this option. *)
| ( VernacSetOption (l,_) | VernacUnsetOption l)
when CList.equal String.equal l Vernacentries.universe_polymorphism_option_name ->
VtSideff [], VtNow
- (* Nested vernac exprs *)
- | VernacProgram e -> static_classifier ~poly e
- | VernacLocal (_,e) -> static_classifier ~poly e
- | VernacPolymorphic (b, e) -> static_classifier ~poly:b e
(* Qed *)
| VernacAbort _ -> VtQed VtDrop, VtLater
| VernacEndProof Admitted -> VtQed VtKeepAsAxiom, VtLater
@@ -87,18 +88,14 @@ let classify_vernac e =
| VernacSetOption (["Default";"Proof";"Using"],_) -> VtSideff [], VtNow
(* StartProof *)
| VernacDefinition ((Decl_kinds.DoDischarge,_),((_,i),_),ProveBody _) ->
- VtStartProof(default_proof_mode (),Doesn'tGuaranteeOpacity,[i]), VtLater
+ VtStartProof(default_proof_mode (),Doesn'tGuaranteeOpacity, idents_of_name i), VtLater
| VernacDefinition (_,((_,i),_),ProveBody _) ->
let guarantee = if poly then Doesn'tGuaranteeOpacity else GuaranteesOpacity in
- VtStartProof(default_proof_mode (),guarantee,[i]), VtLater
+ VtStartProof(default_proof_mode (),guarantee, idents_of_name i), VtLater
| VernacStartTheoremProof (_,l) ->
- let ids =
- CList.map_filter (function (Some ((_,i),pl), _) -> Some i | _ -> None) l in
+ let ids = List.map (fun (((_, i), _), _) -> i) l in
let guarantee = if poly then Doesn'tGuaranteeOpacity else GuaranteesOpacity in
VtStartProof (default_proof_mode (),guarantee,ids), VtLater
- | VernacGoal _ ->
- let guarantee = if poly then Doesn'tGuaranteeOpacity else GuaranteesOpacity in
- VtStartProof (default_proof_mode (),guarantee,[]), VtLater
| VernacFixpoint (discharge,l) ->
let guarantee =
if discharge = Decl_kinds.DoDischarge || poly then Doesn'tGuaranteeOpacity
@@ -125,7 +122,7 @@ let classify_vernac e =
| VernacAssumption (_,_,l) ->
let ids = List.flatten (List.map (fun (_,(l,_)) -> List.map (fun (id, _) -> snd id) l) l) in
VtSideff ids, VtLater
- | VernacDefinition (_,((_,id),_),DefineBody _) -> VtSideff [id], VtLater
+ | VernacDefinition (_,((_,id),_),DefineBody _) -> VtSideff (idents_of_name id), VtLater
| VernacInductive (_, _,_,l) ->
let ids = List.map (fun (((_,((_,id),_)),_,_,_,cl),_) -> id :: match cl with
| Constructors l -> List.map (fun (_,((_,id),_)) -> id) l
@@ -193,7 +190,13 @@ let classify_vernac e =
with Not_found -> anomaly(str"No classifier for"++spc()++str (fst s)++str".")
in
let rec static_control_classifier ~poly = function
- | VernacExpr e -> static_classifier ~poly e
+ | VernacExpr (f, e) ->
+ let poly = List.fold_left (fun poly f ->
+ match f with
+ | VernacPolymorphic b -> b
+ | (VernacProgram | VernacLocal _) -> poly
+ ) poly f in
+ static_classifier ~poly e
| VernacTimeout (_,e) -> static_control_classifier ~poly e
| VernacTime (_,(_,e)) | VernacRedirect (_, (_,e)) ->
static_control_classifier ~poly e