src/HOL/Boogie/Tools/boogie_commands.ML
author wenzelm
Thu, 26 Aug 2010 13:09:12 +0200
changeset 38756 d07959fabde6
parent 37944 4b7afae88c57
child 40514 db5f14910dce
permissions -rw-r--r--
renamed ProofContext.theory(_result) to ProofContext.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
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 |
35356
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    33
  VC_Examine of string list |
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    34
  VC_Single of string
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    35
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    36
fun get_vc thy vc_name =
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    37
  (case Boogie_VCs.lookup thy vc_name of
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    38
    SOME vc => vc
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    39
  | NONE => 
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    40
      (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
    41
        SOME Boogie_VCs.Proved => error ("The verification condition " ^
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    42
          quote vc_name ^ " has already been proved.")
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    43
      | _ => error ("There is no verification condition " ^
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    44
          quote vc_name ^ ".")))
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    45
35356
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    46
local
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    47
  fun split_goal t =
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    48
    (case Boogie_Tactics.split t of
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    49
      [tp] => tp
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    50
    | _ => error "Multiple goals.")
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    51
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    52
  fun single_prep t =
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    53
    let
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    54
      val (us, u) = split_goal t
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    55
      val assms = [((@{binding vc_trace}, []), map (rpair []) us)]
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    56
    in
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    57
      pair [u] o snd o ProofContext.add_assms_i Assumption.assume_export assms
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    58
    end
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    59
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    60
  fun single_prove goal ctxt thm =
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    61
    Goal.prove ctxt [] [] goal (fn {context, ...} => HEADGOAL (
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    62
      Boogie_Tactics.split_tac
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    63
      THEN' Boogie_Tactics.drop_assert_at_tac
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    64
      THEN' SUBPROOF (fn _ => Tactic.rtac thm 1) context))
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    65
in
35864
d82c0dd51794 use a proof context instead of a local theory
boehmes
parents: 35356
diff changeset
    66
fun boogie_vc (vc_name, vc_opts) thy =
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    67
  let
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    68
    val vc = get_vc thy vc_name
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    69
35356
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    70
    fun extract vc l =
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    71
      (case Boogie_VCs.extract vc l of
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    72
        SOME vc' => vc'
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    73
      | NONE => error ("There is no assertion to be proved with label " ^
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    74
          quote l ^ "."))
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    75
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    76
    val vcs =
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    77
      (case vc_opts of
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    78
        VC_Complete => [vc]
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    79
      | 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
    80
      | 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
    81
      | 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
    82
      | VC_Only ls => [Boogie_VCs.only ls vc]
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    83
      | VC_Without ls => [Boogie_VCs.without ls vc]
35356
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    84
      | VC_Examine ls => map (extract vc) ls
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    85
      | VC_Single l => [extract vc l])
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
    86
    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
    87
35356
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    88
    val (prepare, finish) =
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    89
      (case vc_opts of
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    90
         VC_Single _ => (single_prep (hd ts), single_prove (hd ts))
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    91
      | _ => (pair ts, K I))
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
    92
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    93
    val discharge = fold (Boogie_VCs.discharge o pair vc_name)
38756
d07959fabde6 renamed ProofContext.theory(_result) to ProofContext.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
wenzelm
parents: 37944
diff changeset
    94
    fun after_qed [thms] = ProofContext.background_theory (discharge (vcs ~~ thms))
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    95
      | after_qed _ = I
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    96
  in
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36323
diff changeset
    97
    ProofContext.init_global thy
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
    98
    |> fold Variable.auto_fixes ts
35864
d82c0dd51794 use a proof context instead of a local theory
boehmes
parents: 35356
diff changeset
    99
    |> (fn ctxt1 => ctxt1
35356
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
   100
    |> prepare
35864
d82c0dd51794 use a proof context instead of a local theory
boehmes
parents: 35356
diff changeset
   101
    |-> (fn us => fn ctxt2 => ctxt2
36323
655e2d74de3a modernized naming conventions of main Isar proof elements;
wenzelm
parents: 35864
diff changeset
   102
    |> Proof.theorem NONE (fn thmss => fn ctxt =>
35864
d82c0dd51794 use a proof context instead of a local theory
boehmes
parents: 35356
diff changeset
   103
         let val export = map (finish ctxt1) o ProofContext.export ctxt ctxt2
35356
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
   104
         in after_qed (map export thmss) ctxt end) [map (rpair []) us]))
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   105
  end
35356
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
   106
end
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   107
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   108
fun write_list head =
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   109
  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
   110
  Pretty.writeln o Pretty.big_list head
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   111
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   112
fun parens s = "(" ^ s ^ ")"
34079
3edfefaaf355 print assertions in a more natural order
boehmes
parents: 34068
diff changeset
   113
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   114
fun boogie_status thy =
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   115
  let
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   116
    fun string_of_state Boogie_VCs.Proved = "proved"
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   117
      | string_of_state Boogie_VCs.NotProved = "not proved"
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   118
      | string_of_state Boogie_VCs.PartiallyProved = "partially proved"
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   119
  in
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   120
    Boogie_VCs.state_of thy
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   121
    |> map (fn (name, proved) => name ^ " " ^ parens (string_of_state proved))
34080
a36d80e4e42e also sort verification conditions before printing
boehmes
parents: 34079
diff changeset
   122
    |> write_list "Boogie verification conditions:"
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   123
  end
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   124
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   125
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
   126
  let
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   127
    fun pretty tag s = s ^ " " ^ parens tag
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   128
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   129
    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
   130
  in
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   131
    if full
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   132
    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
   133
      quote vc_name ^ ":")
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   134
      (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
   135
    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
   136
      quote vc_name ^ ":") not_proved
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   137
  end
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   138
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   139
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
   140
  let
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   141
    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
   142
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   143
    fun pp (i, ns) =
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   144
      if full
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   145
      then
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   146
        [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
   147
          [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
   148
      else
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   149
        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
   150
        in
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   151
          if null ns' then []
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   152
          else
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   153
            [Pretty.big_list ("Unproved assertions of path " ^
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   154
              string_of_int (i+1) ^ ":") [labels ns']]
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   155
        end
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   156
  in
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   157
    Pretty.writeln
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   158
      (Pretty.big_list
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   159
        ("Paths of Boogie verification condition " ^ quote vc_name ^ ":")
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   160
        (flat (map_index pp (Boogie_VCs.path_names_of (get_vc thy vc_name)))))
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   161
  end
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   162
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   163
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   164
local
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   165
  fun trying s = tracing ("Trying " ^ s ^ " ...")
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   166
  fun success_on s = tracing ("Succeeded on " ^ s ^ ".")
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   167
  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
   168
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   169
  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
   170
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   171
  fun string_of_path (i, n) =
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   172
    "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
   173
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   174
  fun itemize_paths ps =
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   175
    let val n = length ps
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   176
    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
   177
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   178
  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
   179
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   180
  fun divide f vc =
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   181
    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
   182
    in
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   183
      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
   184
      else
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   185
        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
   186
        in par_map f [vc1, vc2] end
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   187
    end
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   188
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   189
  fun prove thy meth vc =
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36323
diff changeset
   190
    ProofContext.init_global thy
36323
655e2d74de3a modernized naming conventions of main Isar proof elements;
wenzelm
parents: 35864
diff changeset
   191
    |> Proof.theorem 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
   192
    |> Proof.apply meth
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   193
    |> Seq.hd
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   194
    |> Proof.global_done_proof
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   195
in
35323
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   196
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
   197
  let
35323
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   198
    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
   199
35323
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   200
    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
   201
      (case try (tp t) vc of
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   202
        SOME _ => (success_on tag; [])
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   203
      | 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
   204
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   205
    fun some_asserts vc =
35323
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   206
      let
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   207
        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
   208
          else (quick, ", further splitting ...")
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   209
      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
   210
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   211
    fun single_path p =
35323
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   212
      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
   213
        (divide some_asserts)
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   214
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   215
    val complete_vc =
35323
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   216
      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
   217
        (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
   218
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   219
    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
   220
  in
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   221
    if null unsolved
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   222
    then writeln ("Completely solved Boogie verification condition " ^
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   223
      quote vc_name ^ ".")
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   224
    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
   225
      quote vc_name ^ ":") unsolved
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   226
  end
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   227
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   228
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
   229
  let
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   230
    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
   231
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   232
    val vc = get_vc thy vc_name
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   233
    fun prove_assert name =
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   234
      (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
   235
    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
   236
  in
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   237
    (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
   238
      SOME name => writeln ("failed on " ^ quote name)
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   239
    | NONE => writeln "succeeded")
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   240
  end
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   241
end
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   242
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   243
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   244
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   245
fun boogie_end thy =
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   246
  let
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   247
    fun not_proved (_, Boogie_VCs.Proved) = NONE
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   248
      | not_proved (name, _) = SOME name
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   249
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   250
    val unproved = map_filter not_proved (Boogie_VCs.state_of thy)
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   251
  in
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   252
    if null unproved then Boogie_VCs.close thy
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   253
    else error (Pretty.string_of (Pretty.big_list 
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   254
      "The following verification conditions have not been proved:"
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   255
      (map Pretty.str unproved)))
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   256
  end
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   257
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   258
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   259
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   260
(* syntax and setup *)
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   261
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   262
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
   263
fun scan_arg f = Args.parens f
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   264
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
   265
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36610
diff changeset
   266
val vc_offsets = Scan.optional (Args.bracks (Parse.list1
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36610
diff changeset
   267
  (Parse.string --| Args.colon -- Parse.nat))) []
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
   268
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   269
val _ =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36610
diff changeset
   270
  Outer_Syntax.command "boogie_open"
37944
4b7afae88c57 observe standard conventions for doc-strings;
wenzelm
parents: 36960
diff changeset
   271
    "open a new Boogie environment and load a Boogie-generated .b2i file"
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36610
diff changeset
   272
    Keyword.thy_decl
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36610
diff changeset
   273
    (scan_opt "quiet" -- Parse.name -- vc_offsets >> 
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
   274
      (Toplevel.theory o boogie_open))
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   275
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   276
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36610
diff changeset
   277
val vc_name = Parse.name
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   278
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36610
diff changeset
   279
val vc_label = Parse.name
35356
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
   280
val vc_labels = Scan.repeat1 vc_label
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   281
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   282
val vc_paths =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36610
diff changeset
   283
  Parse.nat -- (Args.$$$ "-" |-- Parse.nat) >> (op upto) ||
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36610
diff changeset
   284
  Parse.nat >> single
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   285
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   286
val vc_opts =
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   287
  scan_arg
35356
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
   288
   (scan_val "assertion" vc_label >> VC_Single ||
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
   289
    scan_val "examine" vc_labels >> VC_Examine ||
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36610
diff changeset
   290
    scan_val "take" ((Parse.list vc_paths >> flat) -- Scan.option (
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   291
      scan_val "without" vc_labels >> pair false ||
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   292
      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
   293
    scan_val "only" vc_labels >> VC_Only ||
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   294
    scan_val "without" vc_labels >> VC_Without) ||
35356
5c937073e1d5 added variant of boogie_vc to prove a single assertion: keep premises (i.e. the trace up this assertion) as facts in the context (and not as part of the goal) to increase performance when dealing with large goals
boehmes
parents: 35323
diff changeset
   295
  Scan.succeed VC_Complete
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   296
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   297
val _ =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36610
diff changeset
   298
  Outer_Syntax.command "boogie_vc"
37944
4b7afae88c57 observe standard conventions for doc-strings;
wenzelm
parents: 36960
diff changeset
   299
    "enter into proof mode for a specific verification condition"
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36610
diff changeset
   300
    Keyword.thy_goal
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   301
    (vc_name -- vc_opts >> (fn args =>
35864
d82c0dd51794 use a proof context instead of a local theory
boehmes
parents: 35356
diff changeset
   302
      (Toplevel.print o Toplevel.theory_to_proof (boogie_vc args))))
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   303
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   304
35323
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   305
val quick_timeout = 5
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   306
val default_timeout = 20
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   307
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36610
diff changeset
   308
fun timeout name = Scan.optional (scan_val name Parse.nat)
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   309
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   310
val status_test =
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   311
  scan_arg (
35323
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   312
    Args.$$$ "scan" |-- timeout "timeout" quick_timeout >> boogie_scan_vc ||
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   313
    Args.$$$ "narrow" |-- timeout "step_timeout" quick_timeout --
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   314
      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
   315
  vc_name -- Method.parse >>
35323
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 35125
diff changeset
   316
  (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
   317
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   318
val status_vc =
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   319
  (scan_arg
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   320
    (Args.$$$ "full" |--
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   321
      (Args.$$$ "paths" >> K (boogie_status_vc_paths true) ||
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   322
       Scan.succeed (boogie_status_vc true)) ||
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   323
     Args.$$$ "paths" >> K (boogie_status_vc_paths false)) ||
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   324
   Scan.succeed (boogie_status_vc false)) --
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   325
  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
   326
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   327
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
   328
  f (Toplevel.theory_of state))
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   329
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   330
val _ =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36610
diff changeset
   331
  Outer_Syntax.improper_command "boogie_status"
37944
4b7afae88c57 observe standard conventions for doc-strings;
wenzelm
parents: 36960
diff changeset
   332
    "show the name and state of all loaded verification conditions"
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36610
diff changeset
   333
    Keyword.diag
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   334
    (status_test >> status_cmd ||
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34080
diff changeset
   335
     status_vc >> status_cmd ||
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   336
     Scan.succeed (status_cmd boogie_status))
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   337
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   338
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   339
val _ =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36610
diff changeset
   340
  Outer_Syntax.command "boogie_end"
37944
4b7afae88c57 observe standard conventions for doc-strings;
wenzelm
parents: 36960
diff changeset
   341
    "close the current Boogie environment"
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36610
diff changeset
   342
    Keyword.thy_decl
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   343
    (Scan.succeed (Toplevel.theory boogie_end))
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   344
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33671
diff changeset
   345
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   346
val setup = Theory.at_end (fn thy =>
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   347
  let
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   348
    val _ = Boogie_VCs.is_closed thy
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   349
      orelse error ("Found the end of the theory, " ^ 
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   350
        "but the last Boogie environment is still open.")
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   351
  in NONE end)
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   352
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   353
end