aboutsummaryrefslogtreecommitdiffhomepage
path: root/parsing/g_proofs.ml4
blob: dc545c691fc561898883181c2a7088db0bdee342 (plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
(************************************************************************)
(*  v      *   The Coq Proof Assistant  /  The Coq Development Team     *)
(* <O___,, *   INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012     *)
(*   \VV/  **************************************************************)
(*    //   *      This file is distributed under the terms of the       *)
(*         *       GNU Lesser General Public License Version 2.1        *)
(************************************************************************)

open Compat
open Constrexpr
open Vernacexpr
open Locality
open Misctypes
open Tok

open Pcoq
open Pcoq.Tactic
open Pcoq.Prim
open Pcoq.Constr
open Pcoq.Vernac_

let thm_token = G_vernac.thm_token

let only_identrefs =
  Gram.Entry.of_parser "test_only_identrefs"
    (fun strm ->
      let rec aux n =
      match get_tok (Util.stream_nth n strm) with
        | KEYWORD "." -> ()
        | KEYWORD ")" -> ()
        | IDENT _ -> aux (n+1)
        | _ -> raise Stream.Failure in
      aux 0)

let hint_proof_using e = function
  | Some _ as x -> x
  | None -> match Proof_using.get_default_proof_using () with
     | None -> None
     | Some s -> Some (Gram.entry_parse e (Gram.parsable (Stream.of_string s)))

(* Proof commands *)
GEXTEND Gram
  GLOBAL: command;

  opt_hintbases:
  [ [ -> []
    | ":"; l = LIST1 [id = IDENT -> id ] -> l ] ]
  ;
  command:
    [ [ IDENT "Goal"; c = lconstr -> VernacGoal c
      | IDENT "Proof" ->
          VernacProof (None,hint_proof_using section_subset_descr None)
      | IDENT "Proof" ; IDENT "Mode" ; mn = string -> VernacProofMode mn
      | IDENT "Proof"; "with"; ta = tactic; 
        l = OPT [ "using"; l = section_subset_descr -> l ] ->
          VernacProof (Some ta, hint_proof_using section_subset_descr l)
      | IDENT "Proof"; "using"; l = section_subset_descr;
        ta = OPT [ "with"; ta = tactic -> ta ] ->
          VernacProof (ta,Some l)
      | IDENT "Proof"; c = lconstr -> VernacExactProof c
      | IDENT "Abort" -> VernacAbort None
      | IDENT "Abort"; IDENT "All" -> VernacAbortAll
      | IDENT "Abort"; id = identref -> VernacAbort (Some id)
      | IDENT "Existential"; n = natural; c = constr_body ->
	  VernacSolveExistential (n,c)
      | IDENT "Admitted" -> VernacEndProof Admitted
      | IDENT "Qed" -> VernacEndProof (Proved (true,None))
      | IDENT "Save" -> VernacEndProof (Proved (true,None))
      | IDENT "Save"; tok = thm_token; id = identref ->
	  VernacEndProof (Proved (true,Some (id,Some tok)))
      | IDENT "Save"; id = identref ->
	  VernacEndProof (Proved (true,Some (id,None)))
      | IDENT "Defined" -> VernacEndProof (Proved (false,None))
      |	IDENT "Defined"; id=identref ->
	  VernacEndProof (Proved (false,Some (id,None)))
      | IDENT "Restart" -> VernacRestart
      | IDENT "Undo" -> VernacUndo 1
      | IDENT "Undo"; n = natural -> VernacUndo n
      | IDENT "Undo"; IDENT "To"; n = natural -> VernacUndoTo n
      | IDENT "Focus" -> VernacFocus None
      | IDENT "Focus"; n = natural -> VernacFocus (Some n)
      | IDENT "Unfocus" -> VernacUnfocus
      | IDENT "Unfocused" -> VernacUnfocused
      | IDENT "Show" -> VernacShow (ShowGoal OpenSubgoals)
      | IDENT "Show"; n = natural -> VernacShow (ShowGoal (NthGoal n))
      | IDENT "Show"; IDENT "Goal"; n = string ->
          VernacShow (ShowGoal (GoalId n))
      | IDENT "Show"; IDENT "Implicit"; IDENT "Arguments"; n = OPT natural ->
	  VernacShow (ShowGoalImplicitly n)
      | IDENT "Show"; IDENT "Node" -> VernacShow ShowNode
      | IDENT "Show"; IDENT "Script" -> VernacShow ShowScript
      | IDENT "Show"; IDENT "Existentials" -> VernacShow ShowExistentials
      | IDENT "Show"; IDENT "Tree" -> VernacShow ShowTree
      | IDENT "Show"; IDENT "Conjectures" -> VernacShow ShowProofNames
      | IDENT "Show"; IDENT "Proof" -> VernacShow ShowProof
      | IDENT "Show"; IDENT "Intro" -> VernacShow (ShowIntros false)
      | IDENT "Show"; IDENT "Intros" -> VernacShow (ShowIntros true)
      | IDENT "Show"; IDENT "Match"; id = identref -> VernacShow (ShowMatch id)
      | IDENT "Show"; IDENT "Thesis" -> VernacShow ShowThesis
      | IDENT "Guarded" -> VernacCheckGuard
      (* Hints for Auto and EAuto *)
      | IDENT "Create"; IDENT "HintDb" ;
	  id = IDENT ; b = [ "discriminated" -> true | -> false ] ->
	    VernacCreateHintDb (id, b)
      | IDENT "Remove"; IDENT "Hints"; ids = LIST1 global; dbnames = opt_hintbases ->
	  VernacRemoveHints (dbnames, ids)
      | IDENT "Hint"; local = obsolete_locality; h = hint;
	  dbnames = opt_hintbases ->
	  VernacHints (local,dbnames, h)
      (* Declare "Resolve" explicitly so as to be able to later extend with
         "Resolve ->" and "Resolve <-" *)
      | IDENT "Hint"; IDENT "Resolve"; lc = LIST1 reference_or_constr; 
	  pri = OPT [ "|"; i = natural -> i ];
	  dbnames = opt_hintbases ->
	  VernacHints (false,dbnames,
	    HintsResolve (List.map (fun x -> (pri, true, x)) lc))
      ] ];
  section_subset_descr:
    [ [ IDENT "All" -> SsAll
      | "Type" -> SsType
      | only_identrefs; l = LIST0 identref -> SsExpr (SsSet l)
      | e = section_subset_expr -> SsExpr e ] ]
  ;
  section_subset_expr:
    [ "35" 
      [ "-"; e = section_subset_expr -> SsCompl e ]
    | "50"
      [ e1 = section_subset_expr; "-"; e2 = section_subset_expr->SsSubstr(e1,e2)
      | e1 = section_subset_expr; "+"; e2 = section_subset_expr->SsUnion(e1,e2)]
    | "0"
      [ i = identref -> SsSet [i]
      | "("; only_identrefs; l = LIST0 identref; ")"-> SsSet l
      | "("; e = section_subset_expr; ")"-> e ] ]
  ;
  obsolete_locality:
    [ [ IDENT "Local" -> true | -> false ] ]
  ;
  reference_or_constr:
   [ [ r = global -> HintsReference r
     | c = constr -> HintsConstr c ] ]
  ;
  hint:
    [ [ IDENT "Resolve"; lc = LIST1 reference_or_constr; 
	pri = OPT [ "|"; i = natural -> i ] ->
          HintsResolve (List.map (fun x -> (pri, true, x)) lc)
      | IDENT "Immediate"; lc = LIST1 reference_or_constr -> HintsImmediate lc
      | IDENT "Transparent"; lc = LIST1 global -> HintsTransparency (lc, true)
      | IDENT "Opaque"; lc = LIST1 global -> HintsTransparency (lc, false)
      | IDENT "Unfold"; lqid = LIST1 global -> HintsUnfold lqid
      | IDENT "Constructors"; lc = LIST1 global -> HintsConstructors lc
      | IDENT "Extern"; n = natural; c = OPT constr_pattern ; "=>";
          tac = tactic ->
	  HintsExtern (n,c,tac) ] ]
    ;
  constr_body:
    [ [ ":="; c = lconstr -> c
      | ":"; t = lconstr; ":="; c = lconstr -> CCast(!@loc,c,CastConv t) ] ]
  ;
END