src/HOL/Boogie/Tools/boogie_commands.ML
author boehmes
Tue, 23 Feb 2010 15:20:19 +0100
changeset 35323 259931828ecc
parent 35125 acace7e30357
child 35356 5c937073e1d5
permissions -rw-r--r--
separated narrowing timeouts for intermediate and final steps
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
     1
(*  Title:      HOL/Boogie/Tools/boogie_commands.ML
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
     2
    Author:     Sascha Boehme, TU Muenchen
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
     3
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
     4
Isar commands to create a Boogie environment simulation.
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
     5
*)
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
     6
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
     7
signature BOOGIE_COMMANDS =
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
     8
sig
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
     9
  val setup: theory -> theory
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    10
end
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    11
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    12
structure Boogie_Commands: BOOGIE_COMMANDS =
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    13
struct
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    14
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    15
(* commands *)
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    16
35125
acace7e30357 optionally localize assertion labels (based on user-defined offsets) to reduce the effort of label adaptions after changes to the source program
boehmes
parents: 34181
diff changeset
    17
fun boogie_open ((quiet, base_name), offsets) thy =
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    18
  let
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    19
    val path = Path.explode (base_name ^ ".b2i")
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    20
    val _ = File.exists path orelse
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    21
      error ("Unable to find file " ^ quote (Path.implode path))
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    22
    val _ = Boogie_VCs.is_closed thy orelse
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    23
      error ("Found the beginning of a new Boogie environment, " ^
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    24
        "but another Boogie environment is still open.")
35125
acace7e30357 optionally localize assertion labels (based on user-defined offsets) to reduce the effort of label adaptions after changes to the source program
boehmes
parents: 34181
diff changeset
    25
  in Boogie_Loader.load_b2i (not quiet) offsets path thy end
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    26
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    27
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    28
datatype vc_opts =
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    29
  VC_Complete |
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    30
  VC_Take of int list * (bool * string list) option |
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    31
  VC_Only of string list |
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    32
  VC_Without of string list |
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    33
  VC_Examine of string list 
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    34
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    35
fun get_vc thy vc_name =
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    36
  (case Boogie_VCs.lookup thy vc_name of
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    37
    SOME vc => vc
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    38
  | NONE => 
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    39
      (case AList.lookup (op =) (Boogie_VCs.state_of thy) vc_name of
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    40
        SOME Boogie_VCs.Proved => error ("The verification condition " ^
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    41
          quote vc_name ^ " has already been proved.")
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    42
      | _ => error ("There is no verification condition " ^
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    43
          quote vc_name ^ ".")))
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    44
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    45
fun boogie_vc (vc_name, vc_opts) lthy =
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    46
  let
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    47
    val thy = ProofContext.theory_of lthy
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    48
    val vc = get_vc thy vc_name
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    49
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    50
    val vcs =
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    51
      (case vc_opts of
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    52
        VC_Complete => [vc]
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    53
      | VC_Take (ps, NONE) => [Boogie_VCs.paths_and ps [] vc]
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    54
      | VC_Take (ps, SOME (true, ls)) => [Boogie_VCs.paths_and ps ls vc]
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    55
      | VC_Take (ps, SOME (false, ls)) => [Boogie_VCs.paths_without ps ls vc]
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    56
      | VC_Only ls => [Boogie_VCs.only ls vc]
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    57
      | VC_Without ls => [Boogie_VCs.without ls vc]
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    58
      | VC_Examine ls => map_filter (Boogie_VCs.extract vc) ls)
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    59
    val ts = map Boogie_VCs.prop_of_vc vcs
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    60
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    61
    val discharge = fold (Boogie_VCs.discharge o pair vc_name)
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    62
    fun after_qed [thms] = Local_Theory.theory (discharge (vcs ~~ thms))
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    63
      | after_qed _ = I
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    64
  in
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    65
    lthy
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    66
    |> fold Variable.auto_fixes ts
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    67
    |> Proof.theorem_i NONE after_qed [map (rpair []) ts]
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    68
  end
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    69
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    70
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    71
fun write_list head =
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    72
  map Pretty.str o sort (dict_ord string_ord o pairself explode) #>
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    73
  Pretty.writeln o Pretty.big_list head
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    74
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    75
fun parens s = "(" ^ s ^ ")"
34079
3edfefaaf355 print assertions in a more natural order
boehmes
parents: 34068
diff changeset
    76
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    77
fun boogie_status thy =
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    78
  let
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    79
    fun string_of_state Boogie_VCs.Proved = "proved"
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    80
      | string_of_state Boogie_VCs.NotProved = "not proved"
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    81
      | string_of_state Boogie_VCs.PartiallyProved = "partially proved"
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    82
  in
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    83
    Boogie_VCs.state_of thy
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    84
    |> map (fn (name, proved) => name ^ " " ^ parens (string_of_state proved))
34080
a36d80e4e42e also sort verification conditions before printing
boehmes
parents: 34079
diff changeset
    85
    |> write_list "Boogie verification conditions:"
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    86
  end
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    87
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    88
fun boogie_status_vc full vc_name thy =
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    89
  let
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    90
    fun pretty tag s = s ^ " " ^ parens tag
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    91
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    92
    val (not_proved, proved) = Boogie_VCs.state_of_vc thy vc_name
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    93
  in
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    94
    if full
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    95
    then write_list ("Assertions of Boogie verification condition " ^
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    96
      quote vc_name ^ ":")
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    97
      (map (pretty "proved") proved @ map (pretty "not proved") not_proved)
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    98
    else write_list ("Unproved assertions of Boogie verification condition " ^
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    99
      quote vc_name ^ ":") not_proved
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   100
  end
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   101
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   102
fun boogie_status_vc_paths full vc_name thy =
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   103
  let
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   104
    fun labels ls = Pretty.blk (0, Pretty.breaks (map Pretty.str ls))
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   105
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   106
    fun pp (i, ns) =
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   107
      if full
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   108
      then
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   109
        [Pretty.big_list ("Path " ^ string_of_int (i+1) ^ ":")
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   110
          [labels (map (fn (n, true) => n | (n, _) => parens n) ns)]]
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   111
      else
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   112
        let val ns' = map_filter (fn (n, true) => SOME n | _ => NONE) ns
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   113
        in
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   114
          if null ns' then []
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   115
          else
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   116
            [Pretty.big_list ("Unproved assertions of path " ^
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   117
              string_of_int (i+1) ^ ":") [labels ns']]
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   118
        end
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   119
  in
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   120
    Pretty.writeln
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   121
      (Pretty.big_list
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   122
        ("Paths of Boogie verification condition " ^ quote vc_name ^ ":")
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   123
        (flat (map_index pp (Boogie_VCs.path_names_of (get_vc thy vc_name)))))
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   124
  end
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   125
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   126
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   127
local
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   128
  fun trying s = tracing ("Trying " ^ s ^ " ...")
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   129
  fun success_on s = tracing ("Succeeded on " ^ s ^ ".")
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   130
  fun failure_on s c = tracing ("Failed on " ^ s ^ c)
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   131
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   132
  fun string_of_asserts vc = space_implode ", " (fst (Boogie_VCs.names_of vc))
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   133
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   134
  fun string_of_path (i, n) =
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   135
    "path " ^ string_of_int i ^ " of " ^ string_of_int n
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   136
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   137
  fun itemize_paths ps =
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   138
    let val n = length ps
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   139
    in fst (fold_map (fn p => fn i => (((i, n), p), i+1)) ps 1) end
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   140
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   141
  fun par_map f = flat o Par_List.map f
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   142
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   143
  fun divide f vc =
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   144
    let val n = Boogie_VCs.size_of vc
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   145
    in
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   146
      if n <= 1 then fst (Boogie_VCs.names_of vc)
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   147
      else
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   148
        let val (vc1, vc2) = the (Boogie_VCs.split_path (n div 2) vc)
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   149
        in par_map f [vc1, vc2] end
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   150
    end
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   151
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   152
  fun prove thy meth vc =
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   153
    ProofContext.init thy
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   154
    |> Proof.theorem_i NONE (K I) [[(Boogie_VCs.prop_of_vc vc, [])]]
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   155
    |> Proof.apply meth
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   156
    |> Seq.hd
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   157
    |> Proof.global_done_proof
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   158
in
35323
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   159
fun boogie_narrow_vc (quick, timeout) vc_name meth thy =
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   160
  let
35323
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   161
    fun tp t = TimeLimit.timeLimit (Time.fromSeconds t) (prove thy meth)
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   162
35323
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   163
    fun try_vc t (tag, split_tag) split vc = (trying tag;
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   164
      (case try (tp t) vc of
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   165
        SOME _ => (success_on tag; [])
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   166
      | NONE => (failure_on tag split_tag; split vc)))
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   167
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   168
    fun some_asserts vc =
35323
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   169
      let
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   170
        val (t, sep) = if Boogie_VCs.size_of vc = 1 then (timeout, ".")
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   171
          else (quick, ", further splitting ...")
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   172
      in try_vc t (string_of_asserts vc, sep) (divide some_asserts) vc end
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   173
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   174
    fun single_path p =
35323
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   175
      try_vc quick (string_of_path p, ", splitting into assertions ...")
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   176
        (divide some_asserts)
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   177
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   178
    val complete_vc =
35323
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   179
      try_vc quick ("full goal", ", splitting into paths ...")
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   180
        (par_map (uncurry single_path) o itemize_paths o Boogie_VCs.paths_of)
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   181
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   182
    val unsolved = complete_vc (get_vc thy vc_name)
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   183
  in
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   184
    if null unsolved
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   185
    then writeln ("Completely solved Boogie verification condition " ^
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   186
      quote vc_name ^ ".")
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   187
    else write_list ("Unsolved assertions of Boogie verification condition " ^
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   188
      quote vc_name ^ ":") unsolved
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   189
  end
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   190
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   191
fun boogie_scan_vc timeout vc_name meth thy =
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   192
  let
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   193
    val tp = TimeLimit.timeLimit (Time.fromSeconds timeout) (prove thy meth)
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   194
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   195
    val vc = get_vc thy vc_name
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   196
    fun prove_assert name =
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   197
      (trying name; tp (the (Boogie_VCs.extract vc name)))
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   198
    val find_first_failure = find_first (is_none o try prove_assert)
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   199
  in
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   200
    (case find_first_failure (fst (Boogie_VCs.names_of vc)) of
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   201
      SOME name => writeln ("failed on " ^ quote name)
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   202
    | NONE => writeln "succeeded")
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   203
  end
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   204
end
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   205
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   206
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   207
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   208
fun boogie_end thy =
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   209
  let
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   210
    fun not_proved (_, Boogie_VCs.Proved) = NONE
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   211
      | not_proved (name, _) = SOME name
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   212
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   213
    val unproved = map_filter not_proved (Boogie_VCs.state_of thy)
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   214
  in
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   215
    if null unproved then Boogie_VCs.close thy
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   216
    else error (Pretty.string_of (Pretty.big_list 
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   217
      "The following verification conditions have not been proved:"
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   218
      (map Pretty.str unproved)))
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   219
  end
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   220
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   221
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   222
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   223
(* syntax and setup *)
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   224
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   225
fun scan_val n f = Args.$$$ n -- Args.colon |-- f
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   226
fun scan_arg f = Args.parens f
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   227
fun scan_opt n = Scan.optional (scan_arg (Args.$$$ n >> K true)) false
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   228
35125
acace7e30357 optionally localize assertion labels (based on user-defined offsets) to reduce the effort of label adaptions after changes to the source program
boehmes
parents: 34181
diff changeset
   229
val vc_offsets = Scan.optional (Args.bracks (OuterParse.list1
acace7e30357 optionally localize assertion labels (based on user-defined offsets) to reduce the effort of label adaptions after changes to the source program
boehmes
parents: 34181
diff changeset
   230
  (OuterParse.string --| Args.colon -- OuterParse.nat))) []
acace7e30357 optionally localize assertion labels (based on user-defined offsets) to reduce the effort of label adaptions after changes to the source program
boehmes
parents: 34181
diff changeset
   231
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   232
val _ =
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   233
  OuterSyntax.command "boogie_open"
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   234
    "Open a new Boogie environment and load a Boogie-generated .b2i file."
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   235
    OuterKeyword.thy_decl
35125
acace7e30357 optionally localize assertion labels (based on user-defined offsets) to reduce the effort of label adaptions after changes to the source program
boehmes
parents: 34181
diff changeset
   236
    (scan_opt "quiet" -- OuterParse.name -- vc_offsets >> 
acace7e30357 optionally localize assertion labels (based on user-defined offsets) to reduce the effort of label adaptions after changes to the source program
boehmes
parents: 34181
diff changeset
   237
      (Toplevel.theory o boogie_open))
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   238
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   239
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   240
val vc_name = OuterParse.name
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   241
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   242
val vc_labels = Scan.repeat1 OuterParse.name
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   243
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   244
val vc_paths =
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   245
  OuterParse.nat -- (Args.$$$ "-" |-- OuterParse.nat) >> (op upto) ||
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   246
  OuterParse.nat >> single
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   247
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   248
val vc_opts =
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   249
  scan_arg
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   250
   (scan_val "examine" vc_labels >> VC_Examine ||
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   251
    scan_val "take" ((OuterParse.list vc_paths >> flat) -- Scan.option (
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   252
      scan_val "without" vc_labels >> pair false ||
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   253
      scan_val "and_also" vc_labels >> pair true) >> VC_Take) ||
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   254
    scan_val "only" vc_labels >> VC_Only ||
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   255
    scan_val "without" vc_labels >> VC_Without) ||
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   256
  Scan.succeed VC_Complete  
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   257
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   258
val _ =
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   259
  OuterSyntax.command "boogie_vc"
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   260
    "Enter into proof mode for a specific verification condition."
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   261
    OuterKeyword.thy_goal
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   262
    (vc_name -- vc_opts >> (fn args =>
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   263
      (Toplevel.print o Toplevel.local_theory_to_proof NONE (boogie_vc args))))
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   264
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   265
35323
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   266
val quick_timeout = 5
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   267
val default_timeout = 20
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   268
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   269
fun timeout name = Scan.optional (scan_val name OuterParse.nat)
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   270
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   271
val status_test =
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   272
  scan_arg (
35323
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   273
    Args.$$$ "scan" |-- timeout "timeout" quick_timeout >> boogie_scan_vc ||
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   274
    Args.$$$ "narrow" |-- timeout "step_timeout" quick_timeout --
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   275
      timeout "final_timeout" default_timeout >> boogie_narrow_vc) --
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   276
  vc_name -- Method.parse >>
35323
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   277
  (fn ((f, vc_name), meth) => f vc_name meth)
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   278
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   279
val status_vc =
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   280
  (scan_arg
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   281
    (Args.$$$ "full" |--
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   282
      (Args.$$$ "paths" >> K (boogie_status_vc_paths true) ||
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   283
       Scan.succeed (boogie_status_vc true)) ||
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   284
     Args.$$$ "paths" >> K (boogie_status_vc_paths false)) ||
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   285
   Scan.succeed (boogie_status_vc false)) --
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   286
  vc_name >> (fn (f, vc_name) => f vc_name)
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   287
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   288
fun status_cmd f = Toplevel.no_timing o Toplevel.keep (fn state =>
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   289
  f (Toplevel.theory_of state))
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   290
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   291
val _ =
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   292
  OuterSyntax.improper_command "boogie_status"
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   293
    "Show the name and state of all loaded verification conditions."
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   294
    OuterKeyword.diag
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   295
    (status_test >> status_cmd ||
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   296
     status_vc >> status_cmd ||
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   297
     Scan.succeed (status_cmd boogie_status))
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   298
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   299
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   300
val _ =
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   301
  OuterSyntax.command "boogie_end"
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   302
    "Close the current Boogie environment."
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   303
    OuterKeyword.thy_decl
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   304
    (Scan.succeed (Toplevel.theory boogie_end))
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   305
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   306
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   307
val setup = Theory.at_end (fn thy =>
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   308
  let
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   309
    val _ = Boogie_VCs.is_closed thy
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   310
      orelse error ("Found the end of the theory, " ^ 
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   311
        "but the last Boogie environment is still open.")
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   312
  in NONE end)
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   313
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   314
end