src/HOL/SPARK/Tools/spark_commands.ML
author wenzelm
Sat, 09 Nov 2013 18:00:36 +0100
changeset 54381 9c1f21365326
parent 51658 21c10672633b
child 55788 67699e08e969
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     1
(*  Title:      HOL/SPARK/Tools/spark_commands.ML
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     2
    Author:     Stefan Berghofer
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     3
    Copyright:  secunet Security Networks AG
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     4
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     5
Isar commands for handling SPARK/Ada verification conditions.
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     6
*)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     7
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     8
signature SPARK_COMMANDS =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     9
sig
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    10
  val setup: theory -> theory
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    11
end
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    12
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    13
structure SPARK_Commands: SPARK_COMMANDS =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    14
struct
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    15
48908
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
    16
fun spark_open header (prfx, files) thy =
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
    17
  let
49444
fad4724230ce made SML/NJ happy;
wenzelm
parents: 48908
diff changeset
    18
    val ([{src_path, text = vc_text, pos = vc_pos, ...}: Token.file,
48908
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
    19
      {text = fdl_text, pos = fdl_pos, ...},
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
    20
      {text = rls_text, pos = rls_pos, ...}], thy') = files thy;
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
    21
    val base = fst (Path.split_ext (File.full_path (Thy_Load.master_directory thy') src_path));
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
    22
  in
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
    23
    SPARK_VCs.set_vcs
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
    24
      (snd (Fdl_Parser.parse_declarations fdl_pos fdl_text))
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
    25
      (Fdl_Parser.parse_rules rls_pos rls_text)
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
    26
      (snd (Fdl_Parser.parse_vcs header vc_pos vc_text))
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
    27
      base prfx thy'
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
    28
  end;
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
    29
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
    30
(* FIXME *)
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
    31
fun spark_open_old (vc_name, prfx) thy =
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    32
  let
43702
24fb44c1086a more abstract Thy_Load.load_file/use_file for external theory resources;
wenzelm
parents: 43547
diff changeset
    33
    val ((vc_path, vc_id), vc_text) = Thy_Load.load_file thy (Path.explode vc_name);
41887
ececcbd08d35 simplified Thy_Info.check_file -- discontinued load path;
wenzelm
parents: 41592
diff changeset
    34
    val (base, header) =
ececcbd08d35 simplified Thy_Info.check_file -- discontinued load path;
wenzelm
parents: 41592
diff changeset
    35
      (case Path.split_ext vc_path of
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    36
        (base, "vcg") => (base, Fdl_Lexer.vcg_header >> K ())
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    37
      | (base, "siv") => (base, Fdl_Lexer.siv_header >> K ())
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    38
      | _ => error "File name must end with .vcg or .siv");
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    39
    val fdl_path = Path.ext "fdl" base;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    40
    val rls_path = Path.ext "rls" base;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    41
  in
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    42
    SPARK_VCs.set_vcs
41948
30732d2390c8 more basic use of Path.position/File.read;
wenzelm
parents: 41896
diff changeset
    43
      (snd (Fdl_Parser.parse_declarations (Path.position fdl_path) (File.read fdl_path)))
30732d2390c8 more basic use of Path.position/File.read;
wenzelm
parents: 41896
diff changeset
    44
      (Fdl_Parser.parse_rules (Path.position rls_path) (File.read rls_path))
43702
24fb44c1086a more abstract Thy_Load.load_file/use_file for external theory resources;
wenzelm
parents: 43547
diff changeset
    45
      (snd (Fdl_Parser.parse_vcs header (Path.position vc_path) vc_text))
42396
0869ce2006eb Package prefix is now taken into account when looking up user-defined
berghofe
parents: 42361
diff changeset
    46
      base prfx thy
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    47
  end;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    48
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    49
fun add_proof_fun_cmd pf thy =
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42356
diff changeset
    50
  let val ctxt = Proof_Context.init_global thy
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    51
  in SPARK_VCs.add_proof_fun
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    52
    (fn optT => Syntax.parse_term ctxt #>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    53
       the_default I (Option.map Type.constraint optT) #>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    54
       Syntax.check_term ctxt) pf thy
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    55
  end;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    56
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 43702
diff changeset
    57
fun add_spark_type_cmd (s, (raw_typ, cmap)) thy =
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 43702
diff changeset
    58
  SPARK_VCs.add_type (s, (Syntax.read_typ_global thy raw_typ, cmap)) thy;
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 42003
diff changeset
    59
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    60
fun get_vc thy vc_name =
48167
da1a1eae93fa Various improvements
berghofe
parents: 47067
diff changeset
    61
  (case SPARK_VCs.lookup_vc thy false vc_name of
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    62
    SOME (ctxt, (_, proved, ctxt', stmt)) =>
41896
582cccdda0ed spark_end now joins proofs of VCs before writing *.prv file.
berghofe
parents: 41887
diff changeset
    63
      if is_some proved then
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    64
        error ("The verification condition " ^
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    65
          quote vc_name ^ " has already been proved.")
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    66
      else (ctxt @ [ctxt'], stmt)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    67
  | NONE => error ("There is no verification condition " ^
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    68
      quote vc_name ^ "."));
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    69
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    70
fun prove_vc vc_name lthy =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    71
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42356
diff changeset
    72
    val thy = Proof_Context.theory_of lthy;
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    73
    val (ctxt, stmt) = get_vc thy vc_name
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    74
  in
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    75
    Specification.theorem Thm.theoremK NONE
41896
582cccdda0ed spark_end now joins proofs of VCs before writing *.prv file.
berghofe
parents: 41887
diff changeset
    76
      (fn thmss => (Local_Theory.background_theory
582cccdda0ed spark_end now joins proofs of VCs before writing *.prv file.
berghofe
parents: 41887
diff changeset
    77
         (SPARK_VCs.mark_proved vc_name (flat thmss))))
50787
065c684130ad disable interactive mode of Specification.theorem with its slow printing of results;
wenzelm
parents: 49444
diff changeset
    78
      (Binding.name vc_name, []) [] ctxt stmt false lthy
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    79
  end;
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    80
41896
582cccdda0ed spark_end now joins proofs of VCs before writing *.prv file.
berghofe
parents: 41887
diff changeset
    81
fun string_of_status NONE = "(unproved)"
582cccdda0ed spark_end now joins proofs of VCs before writing *.prv file.
berghofe
parents: 41887
diff changeset
    82
  | string_of_status (SOME _) = "(proved)";
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    83
51658
21c10672633b discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents: 50787
diff changeset
    84
fun show_status thy (p, f) =
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    85
  let
48167
da1a1eae93fa Various improvements
berghofe
parents: 47067
diff changeset
    86
    val (context, defs, vcs) = SPARK_VCs.get_vcs thy true;
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    87
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    88
    val vcs' = AList.coalesce (op =) (map_filter
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    89
      (fn (name, (trace, status, ctxt, stmt)) =>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    90
         if p status then
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    91
           SOME (trace, (name, status, ctxt, stmt))
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    92
         else NONE) vcs);
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    93
51658
21c10672633b discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents: 50787
diff changeset
    94
    val ctxt = thy |>
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42356
diff changeset
    95
      Proof_Context.init_global |>
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    96
      Context.proof_map (fold Element.init context)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    97
  in
41592
74c0629a11e9 Tuned show_status
berghofe
parents: 41586
diff changeset
    98
    [Pretty.str "Context:",
74c0629a11e9 Tuned show_status
berghofe
parents: 41586
diff changeset
    99
     Pretty.chunks (maps (Element.pretty_ctxt ctxt) context),
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   100
41592
74c0629a11e9 Tuned show_status
berghofe
parents: 41586
diff changeset
   101
     Pretty.str "Definitions:",
43547
f3a8476285c6 clarified Binding.pretty/print: no quotes, only markup -- Binding.str_of is rendered obsolete;
wenzelm
parents: 42396
diff changeset
   102
     Pretty.chunks (map (fn (b, th) => Pretty.block
f3a8476285c6 clarified Binding.pretty/print: no quotes, only markup -- Binding.str_of is rendered obsolete;
wenzelm
parents: 42396
diff changeset
   103
       [Binding.pretty b, Pretty.str ":",
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   104
        Pretty.brk 1,
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   105
        Display.pretty_thm ctxt th])
41592
74c0629a11e9 Tuned show_status
berghofe
parents: 41586
diff changeset
   106
          defs),
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   107
41592
74c0629a11e9 Tuned show_status
berghofe
parents: 41586
diff changeset
   108
     Pretty.str "Verification conditions:",
74c0629a11e9 Tuned show_status
berghofe
parents: 41586
diff changeset
   109
     Pretty.chunks2 (maps (fn (trace, vcs'') =>
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   110
       Pretty.str trace ::
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   111
       map (fn (name, status, context', stmt) =>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   112
         Pretty.big_list (name ^ " " ^ f status)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   113
           (Element.pretty_ctxt ctxt context' @
41592
74c0629a11e9 Tuned show_status
berghofe
parents: 41586
diff changeset
   114
            Element.pretty_stmt ctxt stmt)) vcs'') vcs')] |>
74c0629a11e9 Tuned show_status
berghofe
parents: 41586
diff changeset
   115
    Pretty.chunks2 |> Pretty.writeln
51658
21c10672633b discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents: 50787
diff changeset
   116
  end;
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   117
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   118
val _ =
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46725
diff changeset
   119
  Outer_Syntax.command @{command_spec "spark_open"}
41586
1f930561a560 more standard command descriptions;
wenzelm
parents: 41561
diff changeset
   120
    "open a new SPARK environment and load a SPARK-generated .vcg or .siv file"
48908
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
   121
    (Parse.name -- Parse.parname >> (Toplevel.theory o spark_open_old));
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
   122
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
   123
val _ =
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
   124
  Outer_Syntax.command @{command_spec "spark_open_vcg"}
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
   125
    "open a new SPARK environment and load a SPARK-generated .vcg file"
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
   126
    (Parse.parname -- Thy_Load.provide_parse_files "spark_open_vcg"
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
   127
      >> (Toplevel.theory o spark_open Fdl_Lexer.vcg_header));
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
   128
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
   129
val _ =
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
   130
  Outer_Syntax.command @{command_spec "spark_open_siv"}
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
   131
    "open a new SPARK environment and load a SPARK-generated .siv file"
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
   132
    (Parse.parname -- Thy_Load.provide_parse_files "spark_open_siv"
713f24d7a40f added specific 'spark_open_vcg' and 'spark_open_siv' with formal management of corresponding source files;
wenzelm
parents: 48167
diff changeset
   133
      >> (Toplevel.theory o spark_open Fdl_Lexer.siv_header));
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   134
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   135
val pfun_type = Scan.option
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   136
  (Args.parens (Parse.list1 Parse.name) --| Args.colon -- Parse.name);
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   137
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   138
val _ =
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46725
diff changeset
   139
  Outer_Syntax.command @{command_spec "spark_proof_functions"}
41586
1f930561a560 more standard command descriptions;
wenzelm
parents: 41561
diff changeset
   140
    "associate SPARK proof functions with terms"
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   141
    (Scan.repeat1 (Parse.name -- (pfun_type --| Args.$$$ "=" -- Parse.term)) >>
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46725
diff changeset
   142
      (Toplevel.theory o fold add_proof_fun_cmd));
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   143
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   144
val _ =
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46725
diff changeset
   145
  Outer_Syntax.command @{command_spec "spark_types"}
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 42003
diff changeset
   146
    "associate SPARK types with types"
46725
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 43702
diff changeset
   147
    (Scan.repeat1 (Parse.name -- Parse.!!! (Args.$$$ "=" |-- Parse.typ --
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 43702
diff changeset
   148
       Scan.optional (Args.parens (Parse.list1
d34ec0512dfb Added infrastructure for mapping SPARK field / constructor names
berghofe
parents: 43702
diff changeset
   149
         (Parse.name -- Parse.!!! (Args.$$$ "=" |-- Parse.xname)))) [])) >>
42356
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 42003
diff changeset
   150
       (Toplevel.theory o fold add_spark_type_cmd));
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 42003
diff changeset
   151
e8777e3ea6ef Added command for associating user-defined types with SPARK types.
berghofe
parents: 42003
diff changeset
   152
val _ =
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46725
diff changeset
   153
  Outer_Syntax.command @{command_spec "spark_vc"}
41586
1f930561a560 more standard command descriptions;
wenzelm
parents: 41561
diff changeset
   154
    "enter into proof mode for a specific verification condition"
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   155
    (Parse.name >> (fn name =>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   156
      (Toplevel.print o Toplevel.local_theory_to_proof NONE (prove_vc name))));
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   157
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   158
val _ =
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46725
diff changeset
   159
  Outer_Syntax.improper_command @{command_spec "spark_status"}
41586
1f930561a560 more standard command descriptions;
wenzelm
parents: 41561
diff changeset
   160
    "show the name and state of all loaded verification conditions"
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   161
    (Scan.optional
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   162
       (Args.parens
41896
582cccdda0ed spark_end now joins proofs of VCs before writing *.prv file.
berghofe
parents: 41887
diff changeset
   163
          (   Args.$$$ "proved" >> K (is_some, K "")
582cccdda0ed spark_end now joins proofs of VCs before writing *.prv file.
berghofe
parents: 41887
diff changeset
   164
           || Args.$$$ "unproved" >> K (is_none, K "")))
51658
21c10672633b discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents: 50787
diff changeset
   165
       (K true, string_of_status) >> (fn args =>
21c10672633b discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents: 50787
diff changeset
   166
        Toplevel.keep (fn state => show_status (Toplevel.theory_of state) args)))
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   167
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   168
val _ =
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46725
diff changeset
   169
  Outer_Syntax.command @{command_spec "spark_end"}
41586
1f930561a560 more standard command descriptions;
wenzelm
parents: 41561
diff changeset
   170
    "close the current SPARK environment"
48167
da1a1eae93fa Various improvements
berghofe
parents: 47067
diff changeset
   171
    (Scan.optional (@{keyword "("} |-- Parse.!!!
da1a1eae93fa Various improvements
berghofe
parents: 47067
diff changeset
   172
         (Parse.reserved "incomplete" --| @{keyword ")"}) >> K true) false >>
da1a1eae93fa Various improvements
berghofe
parents: 47067
diff changeset
   173
       (Toplevel.theory o SPARK_VCs.close));
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   174
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   175
val setup = Theory.at_end (fn thy =>
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   176
  let
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   177
    val _ = SPARK_VCs.is_closed thy
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   178
      orelse error ("Found the end of the theory, " ^ 
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   179
        "but the last SPARK environment is still open.")
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   180
  in NONE end);
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   181
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
   182
end;