| author | desharna | 
| Fri, 21 Mar 2025 15:20:13 +0100 | |
| changeset 82314 | c95eca07f6a0 | 
| parent 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 52722 | 1 | (* Title: HOL/SMT_Examples/Boogie.thy | 
| 2 | Author: Sascha Boehme, TU Muenchen | |
| 3 | *) | |
| 4 | ||
| 63167 | 5 | section \<open>Proving Boogie-generated verification conditions\<close> | 
| 52722 | 6 | |
| 7 | theory Boogie | |
| 8 | imports Main | |
| 72741 | 9 | keywords "boogie_file" :: thy_load | 
| 52722 | 10 | begin | 
| 11 | ||
| 63167 | 12 | text \<open> | 
| 52722 | 13 | HOL-Boogie and its applications are described in: | 
| 14 | \begin{itemize}
 | |
| 15 | ||
| 16 | \item Sascha B"ohme, K. Rustan M. Leino, and Burkhart Wolff. | |
| 17 | HOL-Boogie --- An Interactive Prover for the Boogie Program-Verifier. | |
| 18 | Theorem Proving in Higher Order Logics, 2008. | |
| 19 | ||
| 20 | \item Sascha B"ohme, Micha{\l} Moskal, Wolfram Schulte, and Burkhart Wolff.
 | |
| 21 | HOL-Boogie --- An Interactive Prover-Backend for the Verifying C Compiler. | |
| 22 | Journal of Automated Reasoning, 2009. | |
| 23 | ||
| 24 | \end{itemize}
 | |
| 63167 | 25 | \<close> | 
| 52722 | 26 | |
| 27 | ||
| 28 | ||
| 63167 | 29 | section \<open>Built-in types and functions of Boogie\<close> | 
| 52722 | 30 | |
| 63167 | 31 | subsection \<open>Integer arithmetics\<close> | 
| 52722 | 32 | |
| 63167 | 33 | text \<open> | 
| 34 | The operations \<open>div\<close> and \<open>mod\<close> are built-in in Boogie, but | |
| 52722 | 35 | without a particular semantics due to different interpretations in | 
| 36 | programming languages. We assume that each application comes with a | |
| 37 | proper axiomatization. | |
| 63167 | 38 | \<close> | 
| 52722 | 39 | |
| 40 | axiomatization | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
72741diff
changeset | 41 | boogie_div :: "int \<Rightarrow> int \<Rightarrow> int" (infixl \<open>div'_b\<close> 70) and | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
72741diff
changeset | 42 | boogie_mod :: "int \<Rightarrow> int \<Rightarrow> int" (infixl \<open>mod'_b\<close> 70) | 
| 52722 | 43 | |
| 44 | ||
| 45 | ||
| 63167 | 46 | section \<open>Setup\<close> | 
| 52722 | 47 | |
| 69605 | 48 | ML_file \<open>boogie.ML\<close> | 
| 52722 | 49 | |
| 50 | ||
| 51 | ||
| 63167 | 52 | section \<open>Verification condition proofs\<close> | 
| 52722 | 53 | |
| 58061 | 54 | declare [[smt_oracle = false]] | 
| 66741 | 55 | declare [[smt_read_only_certificates = true]] | 
| 52722 | 56 | |
| 57 | ||
| 69605 | 58 | external_file \<open>Boogie_Max.certs\<close> | 
| 58367 
8af1e68d7e1a
renamed SMT certificate files, following 'SMT2' -> 'SMT' renaming
 blanchet parents: 
58061diff
changeset | 59 | declare [[smt_certificates = "Boogie_Max.certs"]] | 
| 52722 | 60 | |
| 72741 | 61 | boogie_file \<open>Boogie_Max.b2i\<close> | 
| 52722 | 62 | |
| 63 | ||
| 69605 | 64 | external_file \<open>Boogie_Dijkstra.certs\<close> | 
| 58367 
8af1e68d7e1a
renamed SMT certificate files, following 'SMT2' -> 'SMT' renaming
 blanchet parents: 
58061diff
changeset | 65 | declare [[smt_certificates = "Boogie_Dijkstra.certs"]] | 
| 52722 | 66 | |
| 72741 | 67 | boogie_file \<open>Boogie_Dijkstra.b2i\<close> | 
| 52722 | 68 | |
| 69 | ||
| 58061 | 70 | declare [[z3_extensions = true]] | 
| 69605 | 71 | external_file \<open>VCC_Max.certs\<close> | 
| 58367 
8af1e68d7e1a
renamed SMT certificate files, following 'SMT2' -> 'SMT' renaming
 blanchet parents: 
58061diff
changeset | 72 | declare [[smt_certificates = "VCC_Max.certs"]] | 
| 52722 | 73 | |
| 72741 | 74 | boogie_file \<open>VCC_Max.b2i\<close> | 
| 52722 | 75 | |
| 76 | end |