| author | wenzelm | 
| Sun, 27 Oct 2024 12:32:40 +0100 | |
| changeset 81275 | 5ed639c16ce7 | 
| parent 74097 | 6d7be1227d02 | 
| permissions | -rw-r--r-- | 
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 1 | (* Title: HOL/SMT_Examples/SMT_Word_Examples.thy | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 2 | Author: Sascha Boehme, TU Muenchen | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 3 | *) | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 4 | |
| 63167 | 5 | section \<open>Word examples for for SMT binding\<close> | 
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 6 | |
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 7 | theory SMT_Word_Examples | 
| 72515 
c7038c397ae3
moved most material from session HOL-Word to Word_Lib in the AFP
 haftmann parents: 
72514diff
changeset | 8 | imports "HOL-Library.Word" | 
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 9 | begin | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 10 | |
| 58061 | 11 | declare [[smt_oracle = true]] | 
| 12 | declare [[z3_extensions = true]] | |
| 58367 
8af1e68d7e1a
renamed SMT certificate files, following 'SMT2' -> 'SMT' renaming
 blanchet parents: 
58061diff
changeset | 13 | declare [[smt_certificates = "SMT_Word_Examples.certs"]] | 
| 58061 | 14 | declare [[smt_read_only_certificates = true]] | 
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 15 | |
| 63167 | 16 | text \<open> | 
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 17 | Currently, there is no proof reconstruction for words. | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 18 | All lemmas are proved using the oracle mechanism. | 
| 63167 | 19 | \<close> | 
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 20 | |
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 21 | |
| 63167 | 22 | section \<open>Bitvector numbers\<close> | 
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 23 | |
| 58061 | 24 | lemma "(27 :: 4 word) = -5" by smt | 
| 25 | lemma "(27 :: 4 word) = 11" by smt | |
| 26 | lemma "23 < (27::8 word)" by smt | |
| 27 | lemma "27 + 11 = (6::5 word)" by smt | |
| 28 | lemma "7 * 3 = (21::8 word)" by smt | |
| 29 | lemma "11 - 27 = (-16::8 word)" by smt | |
| 58410 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
58367diff
changeset | 30 | lemma "- (- 11) = (11::5 word)" by smt | 
| 58061 | 31 | lemma "-40 + 1 = (-39::7 word)" by smt | 
| 32 | lemma "a + 2 * b + c - b = (b + c) + (a :: 32 word)" by smt | |
| 33 | lemma "x = (5 :: 4 word) \<Longrightarrow> 4 * x = 4" by smt | |
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 34 | |
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 35 | |
| 63167 | 36 | section \<open>Bit-level logic\<close> | 
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 37 | |
| 74097 | 38 | context | 
| 39 | includes bit_operations_syntax | |
| 40 | begin | |
| 41 | ||
| 58061 | 42 | lemma "0b110 AND 0b101 = (0b100 :: 32 word)" by smt | 
| 43 | lemma "0b110 OR 0b011 = (0b111 :: 8 word)" by smt | |
| 44 | lemma "0xF0 XOR 0xFF = (0x0F :: 8 word)" by smt | |
| 45 | lemma "NOT (0xF0 :: 16 word) = 0xFF0F" by smt | |
| 46 | lemma "word_cat (27::4 word) (27::8 word) = (2843::12 word)" by smt | |
| 47 | lemma "word_cat (0b0011::4 word) (0b1111::6word) = (0b0011001111 :: 10 word)" by smt | |
| 48 | lemma "slice 1 (0b10110 :: 4 word) = (0b11 :: 2 word)" by smt | |
| 49 | lemma "ucast (0b1010 :: 4 word) = (0b1010 :: 10 word)" by smt | |
| 50 | lemma "scast (0b1010 :: 4 word) = (0b111010 :: 6 word)" by smt | |
| 72514 | 51 | lemma "push_bit 2 0b10011 = (0b1001100::8 word)" by smt | 
| 52 | lemma "drop_bit 2 0b11001 = (0b110::8 word)" by smt | |
| 53 | lemma "signed_drop_bit 2 0b10011 = (0b100::8 word)" by smt | |
| 58061 | 54 | lemma "word_rotr 2 0b0110 = (0b1001::4 word)" by smt | 
| 55 | lemma "word_rotl 1 0b1110 = (0b1101::4 word)" by smt | |
| 56 | lemma "(x AND 0xff00) OR (x AND 0x00ff) = (x::16 word)" by smt | |
| 57 | lemma "w < 256 \<Longrightarrow> (w :: 16 word) AND 0x00FF = w" by smt | |
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 58 | |
| 74097 | 59 | end | 
| 60 | ||
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 61 | |
| 63167 | 62 | section \<open>Combined integer-bitvector properties\<close> | 
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 63 | |
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 64 | lemma | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 65 | assumes "bv2int 0 = 0" | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 66 | and "bv2int 1 = 1" | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 67 | and "bv2int 2 = 2" | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 68 | and "bv2int 3 = 3" | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 69 | and "\<forall>x::2 word. bv2int x > 0" | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 70 | shows "\<forall>i::int. i < 0 \<longrightarrow> (\<forall>x::2 word. bv2int x > i)" | 
| 58061 | 71 | using assms by smt | 
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 72 | |
| 58061 | 73 | lemma "P (0 \<le> (a :: 4 word)) = P True" by smt | 
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 74 | |
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: diff
changeset | 75 | end |