src/HOL/Boogie/Boogie.thy
author Christian Sternagel
Thu, 30 Aug 2012 15:44:03 +0900
changeset 49093 fdc301f592c4
parent 48907 5c4275c3b5b8
permissions -rw-r--r--
forgot to add lemmas
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
     1
(*  Title:      HOL/Boogie/Boogie.thy
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
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
     5
header {* Integration of the Boogie program verifier *}
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
     6
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
     7
theory Boogie
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 35105
diff changeset
     8
imports Word
46950
d0181abdbdac declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents: 45294
diff changeset
     9
keywords
48907
5c4275c3b5b8 turned 'boogie_open' into thy_load command, without any declarations of 'uses';
wenzelm
parents: 48891
diff changeset
    10
  "boogie_open" :: thy_load and
5c4275c3b5b8 turned 'boogie_open' into thy_load command, without any declarations of 'uses';
wenzelm
parents: 48891
diff changeset
    11
  "boogie_end" :: thy_decl and
5c4275c3b5b8 turned 'boogie_open' into thy_load command, without any declarations of 'uses';
wenzelm
parents: 48891
diff changeset
    12
  "boogie_vc" :: thy_goal and
5c4275c3b5b8 turned 'boogie_open' into thy_load command, without any declarations of 'uses';
wenzelm
parents: 48891
diff changeset
    13
  "boogie_status" :: diag
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    14
begin
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    15
33444
cba964797872 added references to HOL-Boogie papers
boehmes
parents: 33419
diff changeset
    16
text {*
cba964797872 added references to HOL-Boogie papers
boehmes
parents: 33419
diff changeset
    17
HOL-Boogie and its applications are described in:
cba964797872 added references to HOL-Boogie papers
boehmes
parents: 33419
diff changeset
    18
\begin{itemize}
cba964797872 added references to HOL-Boogie papers
boehmes
parents: 33419
diff changeset
    19
cba964797872 added references to HOL-Boogie papers
boehmes
parents: 33419
diff changeset
    20
\item Sascha B"ohme, K. Rustan M. Leino, and Burkhart Wolff.
cba964797872 added references to HOL-Boogie papers
boehmes
parents: 33419
diff changeset
    21
HOL-Boogie --- An Interactive Prover for the Boogie Program-Verifier.
cba964797872 added references to HOL-Boogie papers
boehmes
parents: 33419
diff changeset
    22
Theorem Proving in Higher Order Logics, 2008.
cba964797872 added references to HOL-Boogie papers
boehmes
parents: 33419
diff changeset
    23
cba964797872 added references to HOL-Boogie papers
boehmes
parents: 33419
diff changeset
    24
\item Sascha B"ohme, Micha{\l} Moskal, Wolfram Schulte, and Burkhart Wolff.
cba964797872 added references to HOL-Boogie papers
boehmes
parents: 33419
diff changeset
    25
HOL-Boogie --- An Interactive Prover-Backend for the Verifying C Compiler.
cba964797872 added references to HOL-Boogie papers
boehmes
parents: 33419
diff changeset
    26
Journal of Automated Reasoning, 2009.
cba964797872 added references to HOL-Boogie papers
boehmes
parents: 33419
diff changeset
    27
cba964797872 added references to HOL-Boogie papers
boehmes
parents: 33419
diff changeset
    28
\end{itemize}
cba964797872 added references to HOL-Boogie papers
boehmes
parents: 33419
diff changeset
    29
*}
cba964797872 added references to HOL-Boogie papers
boehmes
parents: 33419
diff changeset
    30
cba964797872 added references to HOL-Boogie papers
boehmes
parents: 33419
diff changeset
    31
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33465
diff changeset
    32
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    33
section {* Built-in types and functions of Boogie *}
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    34
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    35
subsection {* Labels *}
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    36
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    37
text {*
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    38
See "Generating error traces from verification-condition counterexamples"
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    39
by Leino e.a. (2004) for a description of Boogie's labelling mechanism and
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    40
semantics.
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    41
*}
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    42
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    43
definition assert_at :: "bool \<Rightarrow> bool \<Rightarrow> bool" where "assert_at l P = P"
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    44
definition block_at :: "bool \<Rightarrow> bool \<Rightarrow> bool" where "block_at l P = P"
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    45
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    46
lemmas labels = assert_at_def block_at_def
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    47
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    48
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    49
subsection {* Integer arithmetics *}
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    50
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    51
text {*
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    52
The operations @{text div} and @{text mod} are built-in in Boogie, but
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    53
without a particular semantics due to different interpretations in
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    54
programming languages. We assume that each application comes with a
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    55
proper axiomatization.
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    56
*}
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    57
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    58
axiomatization
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    59
  boogie_div :: "int \<Rightarrow> int \<Rightarrow> int" (infixl "div'_b" 70) and
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    60
  boogie_mod :: "int \<Rightarrow> int \<Rightarrow> int" (infixl "mod'_b" 70)
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    61
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    62
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    63
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    64
section {* Boogie environment *}
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    65
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    66
text {*
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    67
Proving Boogie-generated verification conditions happens inside
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    68
a Boogie environment:
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    69
48907
5c4275c3b5b8 turned 'boogie_open' into thy_load command, without any declarations of 'uses';
wenzelm
parents: 48891
diff changeset
    70
  boogie_open "b2i file with extension"
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    71
  boogie_vc "verification condition 1" ...
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    72
  boogie_vc "verification condition 2" ...
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    73
  boogie_vc "verification condition 3" ...
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    74
  boogie_end
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    75
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    76
See the Boogie examples for more details.
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    77
 
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    78
At most one Boogie environment should occur per theory,
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    79
otherwise unexpected effects may arise.
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    80
*}
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    81
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    82
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents: 33465
diff changeset
    83
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    84
section {* Setup *}
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    85
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    86
ML {*
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    87
structure Boogie_Axioms = Named_Thms
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    88
(
45294
3c5d3d286055 tuned Named_Thms: proper binding;
wenzelm
parents: 36899
diff changeset
    89
  val name = @{binding boogie}
33465
boehmes
parents: 33444
diff changeset
    90
  val description =
boehmes
parents: 33444
diff changeset
    91
    "Boogie background axioms loaded along with Boogie verification conditions"
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    92
)
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    93
*}
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    94
setup Boogie_Axioms.setup
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    95
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 46950
diff changeset
    96
ML_file "Tools/boogie_vcs.ML"
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 46950
diff changeset
    97
ML_file "Tools/boogie_loader.ML"
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 46950
diff changeset
    98
ML_file "Tools/boogie_tactics.ML"
34181
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34068
diff changeset
    99
setup Boogie_Tactics.setup
003333ffa543 merged verification condition structure and term representation in one datatype,
boehmes
parents: 34068
diff changeset
   100
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 46950
diff changeset
   101
ML_file "Tools/boogie_commands.ML"
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   102
setup Boogie_Commands.setup
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   103
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
   104
end